try repeat tactic1 || tactic2;tactic3;[tactic31|...|tactic3n];tactic4.is understood as
(try (repeat (tactic1 || tactic2)));
((tactic3;[tactic31|...|tactic3n]);tactic4).
expr ::= expr ; expr | expr ; [ expr | ... | expr ] | tacexpr3 tacexpr3 ::= do (natural | ident) tacexpr3 | info tacexpr3 | progress tacexpr3 | repeat tacexpr3 | try tacexpr3 | tacexpr2 tacexpr2 ::= tacexpr1 || tacexpr3 | tacexpr1 tacexpr1 ::= fun name ... name => atom | let let_clause with ... with let_clause in atom | let rec rec_clause with ... with rec_clause in expr | match goal with context_rule | ... | context_rule end | match reverse goal with context_rule | ... | context_rule end | match expr with match_rule | ... | match_rule end | abstract atom | abstract atom using ident | first [ expr | ... | expr ] | solve [ expr | ... | expr ] | idtac | idtac string | fail | fail natural string | fresh | fresh string | context ident [ term ] | eval redexpr in term | type of term | constr : term | atomic_tactic | qualid tacarg ... tacarg | atom atom ::= qualid | () | ( expr )
Figure 9.1: Syntax of the tactic language
tacarg ::= qualid | () | ltac : atom | term let_clause ::= ident [name ... name] := expr rec_clause ::= ident name ... name := expr context_rule ::= context_hyps , ... , context_hyps |-cpattern => expr | |- cpattern => expr | _ => expr context_hyps ::= name : cpattern match_rule ::= cpattern => expr | context [ident] [ cpattern ] => expr | _ => expr
Figure 9.2: Syntax of the tactic language (continued)
top ::= Ltac ltac_def with ... with ltac_def ltac_def ::= ident [ident ... ident] := expr
Figure 9.3: Tactic toplevel definitions
expr1 ; expr2expr1 and expr2 are evaluated to v1 and v2. v1 and v2 must be tactic values. v1 is then applied and v2 is applied to every subgoal generated by the application of v1. Sequence is left associating.
expr0 ; [ expr1 | ... | exprn ]expri is evaluated to vi, for i=0,...,n. v0 is applied and vi is applied to the i-th generated subgoal by the application of v0, for =1,...,n. It fails if the application of v0 does not generate exactly n subgoals.
do num exprexpr is evaluated to v. v must be a tactic value. v is applied num times. Supposing num>1, after the first application of v, v is applied, at least once, to the generated subgoals and so on. It fails if the application of v fails before the num applications have been completed.
repeat exprexpr is evaluated to v. v must be a tactic value. v is applied until it fails. Supposing n>1, after the first application of v, v is applied, at least once, to the generated subgoals and so on. It stops when it fails for all the generated subgoals. It never fails.
try exprexpr is evaluated to v. v must be a tactic value. v is applied. If the application of v fails, it catches the error and leaves the goal unchanged. If the level of the exception is positive, then the exception is re-raised with its level decremented.
progress exprexpr is evaluated to v. v must be a tactic value. v is applied. If the application of v produced one subgoal equal to the initial goal (up to syntactical equality), then an error of level 0 is raised.
expr1 || expr2expr1 and expr2 are evaluated to v1 and v2. v1 and v2 must be tactic values. v1 is applied and if it fails then v2 is applied. Branching is left associating.
first [ expr1 | ... | exprn ]expri are evaluated to vi and vi must be tactic values, for i=1,...,n. Supposing n>1, it applies v1, if it works, it stops else it tries to apply v2 and so on. It fails when there is no applicable tactic.
solve [ expr1 | ... | exprn ]expri are evaluated to vi and vi must be tactic values, for i=1,...,n. Supposing n>1, it applies v1, if it solves, it stops else it tries to apply v2 and so on. It fails if there is no solving tactic.
idtac and idtac "message"The latter variant prints the string on the standard output.
fail n, fail "message" and fail n "message"The number n is the failure level. If no level is specified, it defaults to 0. The level is used by try and match goal. If 0, it makes match goal considering the next clause (backtracking). If non zero, the current match goal block or try command is aborted and the level is decremented.
let ident1 := expr1each expri is evaluated to vi, then, expr is evaluated by substituting vi to each occurrence of identi, for i=1,...,n. There is no dependencies between the expri and the identi.
with ident2 := expr2
...
with identn := exprn in
expr
qualid tacarg1 ... tacargnThe reference qualid must be bound to some defined tactic definition expecting at least n arguments. The expressions expri are evaluated to vi, for i=1,...,n.
fun ident1 ... identn => exprIndeed, local definitions of functions are a syntactic sugar for binding a fun tactic to an identifier.
match expr withThe expr is evaluated and should yield a term which is matched (non-linear first order unification) against cpattern1 then expr1 is evaluated into some value by substituting the pattern matching instantiations to the metavariables. If the matching with cpattern1 fails, cpattern2 is used and so on. The pattern _ matches any term and shunts all remaining patterns if any. If expr1 evaluates to a tactic, this tactic is not immediately applied to the current goal (in contrast with match goal). If all clauses fail (in particular, there is no pattern _) then a no-matching error is raised.
cpattern1 => expr1
| cpattern2 => expr2
...
| cpatternn => exprn
| _ => exprn+1
end
context ident [ cpattern ]It matches any term which one subterm matches cpattern. If there is a match, the optional ident is assign the ``matched context'', that is the initial term where the matched subterm is replaced by a hole. The definition of context in expressions below will show how to use such term contexts.
If each hypothesis pattern hyp1,i, with i=1,...,m1 is matched (non-linear first order unification) by an hypothesis of the goal and if cpattern1 is matched by the conclusion of the goal, then expr1 is evaluated to v1 by substituting the pattern matching to the metavariables and the real hypothesis names bound to the possible hypothesis names occurring in the hypothesis patterns. If v1 is a tactic value, then it is applied to the goal. If this application fails, then another combination of hypotheses is tried with the same proof context pattern. If there is no other combination of hypotheses then the second proof context pattern is tried and so on. If the next to last proof context pattern fails then exprn+1 is evaluated to vn+1 and vn+1 is applied.
match goal with | hyp1,1,...,hyp1,m1 |-cpattern1=> expr1 | hyp2,1,...,hyp2,m2 |-cpattern2=> expr2 ... | hypn,1,...,hypn,mn |-cpatternn=> exprn |_ => exprn+1 end
context ident [ expr ]ident must denote a context variable bound by a context pattern of a match expression. This expression evaluates replaces the hole of the value of ident by the value of expr.
fresh stringIt evaluates to an identifier unbound in the goal, which is obtained by padding string with a number if necessary. If no name is given, the prefix is H.
eval redexpr in termwhere redexpr is a reduction tactic among red, hnf, compute, simpl, cbv, lazy, unfold, fold, pattern.
Ltac ident ident1 ... identn := exprThis defines a new tactic that can be used in any tactic script or new tactic toplevel definition.
Ltac ident := fun ident1 ... identn => exprRecursive and mutual recursive function definitions are also possible with the syntax:
Ltac ident1 ident1,1 ... ident1,m1 := expr1
with ident2 ident2,1 ... ident2,m2 := expr2
...
with identn identn,1 ... identn,mn := exprn