Previous Up Next

Chapter 9  The tactic language

This chapter gives a compact documentation of Ltac, the tactic language available in Coq. We start by giving the syntax, and next, we present the informal semantics. If you want to know more regarding this language and especially about its fundations, you can refer to [36]. Chapter 10 is devoted to giving examples of use of this language on small but also with non-trivial problems.

9.1  Syntax

The syntax of the tactic language is given Figures 9.1 and 9.2. See page ?? for a description of the BNF metasyntax used in these grammar rules. Various already defined entries will be used in this chapter: entries natural, integer, ident, qualid, term, cpattern and atomic_tactic represent respectively the natural and integer numbers, the authorized identificators and qualified names, Coq's terms and patterns and all the atomic tactics described in chapter 8. The syntax of cpattern is the same as that of terms, but there can be specific variables like ?id where id is a ident or _, which are metavariables for pattern matching. ?id allows us to keep instantiations and to make constraints whereas _ shows that we are not interested in what will be matched. On the right hand side, they are used without the question mark.

The main entry of the grammar is expr. This language is used in proof mode but it can also be used in toplevel definitions as shown in Figure 9.3.


Remarks:
  1. The infix tacticals ``... || ...'' and ``... ; ...'' are associative.

  2. As shown by the figure, tactical || binds more than the prefix tacticals try, repeat, do, info and abstract which themselves bind more than the postfix tactical ``... ;[ ... ]'' which binds more than ``... ; ...''.

    For instance
    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


9.2  Semantics

Tactic expressions can only be applied in the context of a goal. The evaluation yields either a term, an integer or a tactic. Intermediary results can be terms or integers but the final result must be a tactic which is then applied to the current goal.

There is a special case for match goal expressions of which the clauses evaluate to tactics. Such expressions can only be used as end result of a tactic expression (never as argument of a local definition or of an application).

The rest of this section explains the semantics of every construction of Ltac.

Sequence

A sequence is an expression of the following form:
expr1 ; expr2
expr1 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.

General sequence

We can generalize the previous sequence operator as
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.

For loop

There is a for loop that repeats a tactic num times:
do num expr
expr 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 loop

We have a repeat loop with:
repeat expr
expr 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.

Error catching

We can catch the tactic errors with:
try expr
expr 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.

Detecting progress

We can check if a tactic made progress with:
progress expr
expr 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.


Error message: Failed to progress

Branching

We can easily branch with the following structure:
expr1 || expr2
expr1 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 tactic to work

We may consider the first tactic to work (i.e. which does not fail) among a panel of tactics:
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.


Error message: No applicable tactic

Solving

We may consider the first to solve (i.e. which generates no subgoal) among a panel of tactics:
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.


Error message: Cannot solve the goal

Identity

The constant idtac is the identity tactic: it leaves any goal unchanged but it appears in the proof script.
idtac and idtac "message"
The latter variant prints the string on the standard output.

Failing

The tactic fail is the always-failing tactic: it does not solve any goal. It is useful for defining other tacticals since it can be catched by try or match goal. There are three variants:
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.


Error message: Tactic Failure "message" (level n).

Local definitions

Local definitions can be done as follows:
let ident1 := expr1
with ident2 := expr2
...
with identn := exprn in
expr
each 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.

Local definitions can be recursive by using let rec instead of let. Only functions can be defined by recursion, so at least one argument is required.

Application

An application is an expression of the following form:
qualid tacarg1 ... tacargn
The 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.

Function construction

A parameterized tactic can be built anonymously (without resorting to local definitions) with:
fun ident1 ... identn => expr
Indeed, local definitions of functions are a syntactic sugar for binding a fun tactic to an identifier.

Pattern matching on terms

We can carry out pattern matching on terms with:
match expr with
   cpattern1 => expr1
 | cpattern2 => expr2
 ...
 | cpatternn => exprn
 | _ => exprn+1
end
The 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.


Error messages:
  1. No matching clauses for match

    No pattern can be used and, in particular, there is no _ pattern.

  2. Argument of match does not evaluate to a term

    This happens when expr does not denote a term.
There is a special form of patterns to match a subterm against the pattern:
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.

This operator never makes backtracking. If there are several subterms matching the pattern, only the first match is considered. Note that the order of matching is left unspecified.

Pattern matching on goals

We can make pattern matching on goals using the following expression:
match goal with
   | hyp1,1,...,hyp1,m1   |-cpattern1=> expr1
| hyp2,1,...,hyp2,m2   |-cpattern2=> expr2
  ...
| hypn,1,...,hypn,mn   |-cpatternn=> exprn
|_    => exprn+1
end
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.


Error message: No matching clauses for match goal

No goal pattern can be used and, in particular, there is no _ goal pattern.




It is important to know that each hypothesis of the goal can be matched by at most one hypothesis pattern. The order of matching is the following: hypothesis patterns are examined from the right to the left (i.e. hypi,mi before hypi,1). For each hypothesis pattern, the goal hypothesis are matched in order (fresher hypothesis first), but it possible to reverse this order (older first) with the match reverse goal with variant.

Filling a term context

The following expression is not a tactic in the sense that it does not produce subgoals but generates a term to be used in tactic expressions:
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.


Error message: not a context variable

Generating fresh hypothesis names

Tactics sometimes have to generate new names for hypothesis. Letting the system decide a name with the intro tactic is not so good since it is very awkward to retrieve the name the system gave.

As before, the following expression returns a term:
fresh string
It 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.

type of term

This tactic computes the type of term.

Computing in a constr

Evaluation of a term can be performed with:
eval redexpr in term
where redexpr is a reduction tactic among red, hnf, compute, simpl, cbv, lazy, unfold, fold, pattern.

Accessing tactic decomposition

Tactical ``info expr'' is not really a tactical. For elementary tactics, this is equivalent to expr. For complex tactic like auto, it displays the operations performed by the tactic.

Proving a subgoal as a separate lemma

From the outside ``abstract expr'' is the same as solve expr. Internally it saves an auxiliary lemma called ident_subproofn where ident is the name of the current goal and n is chosen so that this is a fresh name.

This tactical is useful with tactics such as omega or discriminate that generate huge proof terms. With that tool the user can avoid the explosion at time of the Save command without having to cut manually the proof in smaller lemmas.


Variants:
  1. abstract expr using ident.
    Give explicitly the name of the auxiliary lemma.

Error message: Proof is not complete

9.3  Tactic toplevel definitions

Basically, tactics toplevel definitions are made as follows:
Ltac ident ident1 ... identn := expr
This defines a new tactic that can be used in any tactic script or new tactic toplevel definition.


Remark: The preceding definition can equivalently be written:
Ltac ident := fun ident1 ... identn => expr
Recursive 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

Previous Up Next