Ltac¶
This chapter documents the tactic language L
tac.
We start by giving the syntax, and next, we present the informal semantics. To learn more about the language and especially about its foundations, please refer to [Del00].
Example: Basic tactic macros
Here are some examples of simple tactic macros that the language lets you write.
See Section Examples of using Ltac for more advanced examples.
Syntax¶
The syntax of the tactic language is given below. See Chapter
The Gallina specification language for a description of the BNF metasyntax used
in these grammar rules. Various already defined entries will be used in this
chapter: entries num
, int
, ident
qualid
, term
, cpattern
and tactic
represent respectively natural and integer numbers,
identifiers, qualified names, Coq terms, patterns and the atomic
tactics described in Chapter Tactics.
The syntax of cpattern
is
the same as that of terms, but it is extended with pattern matching
metavariables. In cpattern
, a pattern matching metavariable is
represented with the syntax ?ident
. The
notation _
can also be used to denote metavariable whose instance is
irrelevant. In the notation ?ident
, the identifier 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 of pattern matching
clauses, the named metavariables are used without the question mark prefix. There
is also a special notation for second-order pattern matching problems: in an
applicative pattern of the form %@?ident ident1 … identn
,
the variable ident
matches any complex expression with (possible)
dependencies in the variables identi
and returns a functional term
of the form fun ident1 … identn => term
.
The main entry of the grammar is ltac_expr
. This language is used in proof
mode but it can also be used in toplevel definitions as shown below.
Note
The infix tacticals
… || …
,… + …
, and… ; …
are associative.Example
If you want that
tactic2; tactic3
be fully run on the first subgoal generated bytactic1
, before running on the other subgoals, then you should not writetactic1; (tactic2; tactic3)
but rathertactic1; [> tactic2; tactic3 .. ]
.In
tacarg
, there is an overlap betweenqualid
as a direct tactic argument andqualid
as a particular case ofterm
. The resolution is done by first looking for a reference of the tactic language and if it fails, for a reference to a term. To force the resolution as a reference of the tactic language, use the formltac:(qualid)
. To force the resolution as a reference to a term, use the syntax(qualid)
.As shown by the figure, tactical
… || …
binds more than the prefix tacticalstry
,repeat
,do
andabstract
which themselves bind more than the postfix tactical… ;[ … ]
which binds at the same level as… ; …
.Example
try repeat tactic1 || tactic2; tactic3; [ tactic+| ]; tactic4
is understood as:
((try (repeat (tactic1 || tactic2)); tactic3); [ tactic+| ]); tactic4
ltac_expr ::=ltac_expr
;ltac_expr
[>ltac_expr
| ... |ltac_expr
]ltac_expr
; [ltac_expr
| ... |ltac_expr
]ltac_expr3
ltac_expr3 ::= do (natural
|ident
)ltac_expr3
progressltac_expr3
repeatltac_expr3
tryltac_expr3
onceltac_expr3
exactly_onceltac_expr3
timeout (natural
|ident
)ltac_expr3
time [string
]ltac_expr3
onlyselector
:ltac_expr3
ltac_expr2
ltac_expr2 ::=ltac_expr1
||ltac_expr3
ltac_expr1
+ltac_expr3
tryifltac_expr1
thenltac_expr1
elseltac_expr1
ltac_expr1
ltac_expr1 ::= funname
...name
=>atom
let [rec]let_clause
with ... withlet_clause
inatom
match goal withcontext_rule
| ... |context_rule
end match reverse goal withcontext_rule
| ... |context_rule
end matchltac_expr
withmatch_rule
| ... |match_rule
end lazymatch goal withcontext_rule
| ... |context_rule
end lazymatch reverse goal withcontext_rule
| ... |context_rule
end lazymatchltac_expr
withmatch_rule
| ... |match_rule
end multimatch goal withcontext_rule
| ... |context_rule
end multimatch reverse goal withcontext_rule
| ... |context_rule
end multimatchltac_expr
withmatch_rule
| ... |match_rule
end abstractatom
abstractatom
usingident
first [ltac_expr
| ... |ltac_expr
] solve [ltac_expr
| ... |ltac_expr
] idtac [message_token
...message_token
] fail [natural
] [message_token
...message_token
] gfail [natural
] [message_token
...message_token
] fresh [component
…component
] contextident
[term
] evalredexpr
interm
type ofterm
constr :term
uconstr :term
type_termterm
numgoals guardtest
assert_failsltac_expr3
assert_succeedsltac_expr3
tactic
qualid
tacarg
...tacarg
atom
atom ::=qualid
()int
(ltac_expr
) component ::=string
|qualid
message_token ::=string
|ident
|int
tacarg ::=qualid
() ltac :atom
term
let_clause ::=ident
[name
...name
] :=ltac_expr
context_rule ::=context_hyp
, ...,context_hyp
|-cpattern
=>ltac_expr
cpattern
=>ltac_expr
|-cpattern
=>ltac_expr
_ =>ltac_expr
context_hyp ::=name
:cpattern
name
:=cpattern
[:cpattern
] match_rule ::=cpattern
=>ltac_expr
context [ident
] [cpattern
] =>ltac_expr
_ =>ltac_expr
test ::=int
=int
int
(< | <= | > | >=)int
selector ::= [ident
]int
(int
|int
-int
), ..., (int
|int
-int
) toplevel_selector ::=selector
all par !
top ::= [Local] Ltacltac_def
with ... withltac_def
ltac_def ::=ident
[ident
...ident
] :=ltac_expr
qualid
[ident
...ident
] ::=ltac_expr
Semantics¶
Tactic expressions can only be applied in the context of a proof. The evaluation yields either a term, an integer or a tactic. Intermediate results can be terms or integers but the final result must be a tactic which is then applied to the focused goals.
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 non-recursive local
definition or of an application).
The rest of this section explains the semantics of every construction of
L
tac.
Sequence¶
A sequence is an expression of the following form:
-
ltac_expr1 ; ltac_expr2
¶ The expression
ltac_expr1
is evaluated tov1
, which must be a tactic value. The tacticv1
is applied to the current goal, possibly producing more goals. Thenltac_expr2
is evaluated to producev2
, which must be a tactic value. The tacticv2
is applied to all the goals produced by the prior application. Sequence is associative.
Local application of tactics¶
Different tactics can be applied to the different goals using the following form:
-
[> ltac_expr*|]
¶ The expressions
ltac_expri
are evaluated tovi
, for i = 1, ..., n and all have to be tactics. Thevi
is applied to the i-th goal, for i = 1, ..., n. It fails if the number of focused goals is not exactly n.Note
If no tactic is given for the i-th goal, it behaves as if the tactic idtac were given. For instance,
[> | auto]
is a shortcut for[> idtac | auto ]
.-
Variant
[> ltac_expri*| | ltac_expr .. | ltac_exprj*|]
In this variant,
ltac_expr
is used for each goal coming after those covered by the list ofltac_expri
but before those covered by the list ofltac_exprj
.
-
Variant
[> ltac_expr*| | .. | ltac_expr*|]
In this variant, idtac is used for the goals not covered by the two lists of
ltac_expr
.
-
Variant
[> ltac_expr .. ]
In this variant, the tactic
ltac_expr
is applied independently to each of the goals, rather than globally. In particular, if there are no goals, the tactic is not run at all. A tactic which expects multiple goals, such asswap
, would act as if a single goal is focused.
-
Variant
ltac_expr0 ; [ltac_expri*|]
This variant of local tactic application is paired with a sequence. In this variant, there must be as many
ltac_expri
as goals generated by the application ofltac_expr0
to each of the individual goals independently. All the above variants work in this form too. Formally,ltac_expr ; [ ... ]
is equivalent to[> ltac_expr ; [> ... ] .. ]
.
-
Variant
Goal selectors¶
We can restrict the application of a tactic to a subset of the currently focused goals with:
-
toplevel_selector : ltac_expr
¶ We can also use selectors as a tactical, which allows to use them nested in a tactic expression, by using the keyword
only
:-
Variant
only selector : ltac_expr
¶ When selecting several goals, the tactic
ltac_expr
is applied globally to all selected goals.
-
Variant
[ident] : ltac_expr
In this variant,
ltac_expr
is applied locally to a goal previously named by the user (see Existential variables).
-
Variant
num-num+, : ltac_expr
In this variant,
ltac_expr
is applied globally to the subset of goals described by the given ranges. You can write a singlen
as a shortcut forn-n
when specifying multiple ranges.
-
Variant
all: ltac_expr
¶ In this variant,
ltac_expr
is applied to all focused goals.all:
can only be used at the toplevel of a tactic expression.
-
Variant
!: ltac_expr
In this variant, if exactly one goal is focused,
ltac_expr
is applied to it. Otherwise the tactic fails.!:
can only be used at the toplevel of a tactic expression.
-
Variant
par: ltac_expr
¶ In this variant,
ltac_expr
is applied to all focused goals in parallel. The number of workers can be controlled via the command line option-async-proofs-tac-j
taking as argument the desired number of workers. Limitations:par:
only works on goals containing no existential variables andltac_expr
must either solve the goal completely or do nothing (i.e. it cannot make some progress).par:
can only be used at the toplevel of a tactic expression.
-
Error
No such goal.
¶
-
Variant
For loop¶
There is a for loop that repeats a tactic num
times:
-
do num ltac_expr
¶ ltac_expr
is evaluated tov
which must be a tactic value. This tactic valuev
is appliednum
times. Supposingnum
> 1, after the first application ofv
,v
is applied, at least once, to the generated subgoals and so on. It fails if the application ofv
fails before the num applications have been completed.
Repeat loop¶
We have a repeat loop with:
-
repeat ltac_expr
¶ ltac_expr
is evaluated tov
. Ifv
denotes a tactic, this tactic is applied to each focused goal independently. If the application succeeds, the tactic is applied recursively to all the generated subgoals until it eventually fails. The recursion stops in a subgoal when the tactic has failed to make progress. The tacticrepeat ltac_expr
itself never fails.
Error catching¶
We can catch the tactic errors with:
-
try ltac_expr
¶ ltac_expr
is evaluated tov
which must be a tactic value. The tactic valuev
is applied to each focused goal independently. If the application ofv
fails in a goal, 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 ltac_expr
¶ ltac_expr
is evaluated to v which must be a tactic value. The tactic valuev
is applied to each focused subgoal independently. If the application ofv
to one of the focused subgoal produced subgoals equal to the initial goals (up to syntactical equality), then an error of level 0 is raised.-
Error
Failed to progress.
¶
-
Error
Backtracking branching¶
We can branch with the following structure:
-
ltac_expr1 + ltac_expr2
¶ ltac_expr1
andltac_expr2
are evaluated respectively tov1
andv2
which must be tactic values. The tactic valuev1
is applied to each focused goal independently and if it fails or a later tactic fails, then the proof backtracks to the current goal andv2
is applied.Tactics can be seen as having several successes. When a tactic fails it asks for more successes of the prior tactics.
ltac_expr1 + ltac_expr2
has all the successes ofv1
followed by all the successes ofv2
. Algebraically,(ltac_expr1 + ltac_expr2); ltac_expr3 = (ltac_expr1; ltac_expr3) + (ltac_expr2; ltac_expr3)
.Branching is left-associative.
First tactic to work¶
Backtracking branching may be too expensive. In this case we may restrict to a local, left biased, branching and consider the first tactic to work (i.e. which does not fail) among a panel of tactics:
-
first [ltac_expr*|]
¶ The
ltac_expri
are evaluated tovi
andvi
must be tactic values for i = 1, ..., n. Supposing n > 1,first [ltac_expr1 | ... | ltac_exprn]
appliesv1
in each focused goal independently and stops if it succeeds; otherwise it tries to applyv2
and so on. It fails when there is no applicable tactic. In other words,first [ltac_expr1 | ... | ltac_exprn]
behaves, in each goal, as the firstvi
to have at least one success.-
Error
No applicable tactic.
¶
-
Variant
first ltac_expr
This is an
L
tac alias that gives a primitive access to the first tactical as anL
tac definition without going through a parsing rule. It expects to be given a list of tactics through aTactic Notation
, allowing to write notations of the following form:Example
- Tactic Notation "foo" tactic_list(tacs) := first tacs.
-
Error
Left-biased branching¶
Yet another way of branching without backtracking is the following structure:
-
ltac_expr1 || ltac_expr2
¶ ltac_expr1
andltac_expr2
are evaluated respectively tov1
andv2
which must be tactic values. The tactic valuev1
is applied in each subgoal independently and if it fails to progress thenv2
is applied.ltac_expr1 || ltac_expr2
is equivalent tofirst [ progress ltac_expr1 | ltac_expr2 ]
(except that if it fails, it fails likev2
). Branching is left-associative.
Generalized biased branching¶
The tactic
-
tryif ltac_expr1 then ltac_expr2 else ltac_expr3
¶ is a generalization of the biased-branching tactics above. The expression
ltac_expr1
is evaluated tov1
, which is then applied to each subgoal independently. For each goal wherev1
succeeds at least once,ltac_expr2
is evaluated tov2
which is then applied collectively to the generated subgoals. Thev2
tactic can trigger backtracking points inv1
: wherev1
succeeds at least once,tryif ltac_expr1 then ltac_expr2 else ltac_expr3
is equivalent tov1; v2
. In each of the goals wherev1
does not succeed at least once,ltac_expr3
is evaluated inv3
which is is then applied to the goal.
Soft cut¶
Another way of restricting backtracking is to restrict a tactic to a single success a posteriori:
Checking the successes¶
Coq provides an experimental way to check that a tactic has exactly one success:
-
exactly_once ltac_expr
¶ ltac_expr
is evaluated tov
which must be a tactic value. The tactic valuev
is applied if it has at most one success. Ifv
fails,exactly_once ltac_expr
fails likev
. Ifv
has a exactly one success,exactly_once ltac_expr
succeeds likev
. Ifv
has two or more successes,exactly_once ltac_expr
fails.Warning
The experimental status of this tactic pertains to the fact if
v
performs side effects, they may occur in an unpredictable way. Indeed, normallyv
would only be executed up to the first success until backtracking is needed, however exactly_once needs to look ahead to see whether a second success exists, and may run further effects immediately.-
Error
This tactic has more than one success.
¶
-
Error
Checking the failure¶
Coq provides a derived tactic to check that a tactic fails:
Checking the success¶
Coq provides a derived tactic to check that a tactic has at least one success:
Solving¶
We may consider the first to solve (i.e. which generates no subgoal) among a panel of tactics:
-
solve [ltac_expr*|]
¶ The
ltac_expri
are evaluated tovi
andvi
must be tactic values, for i = 1, ..., n. Supposing n > 1,solve [ltac_expr1 | ... | ltac_exprn]
appliesv1
to each goal independently and stops if it succeeds; otherwise it tries to applyv2
and so on. It fails if there is no solving tactic.-
Error
Cannot solve the goal.
¶
-
Variant
solve ltac_expr
This is an
L
tac alias that gives a primitive access to thesolve:
tactical. See thefirst
tactical for more information.
-
Error
Identity¶
The constant idtac
is the identity tactic: it leaves any goal unchanged but
it appears in the proof script.
-
idtac message_token*
¶ This prints the given tokens. Strings and integers are printed literally. If a (term) variable is given, its contents are printed.
Failing¶
-
fail
¶ This is the always-failing tactic: it does not solve any goal. It is useful for defining other tacticals since it can be caught by
try
,repeat
,match goal
, or the branching tacticals.-
Variant
fail num
The number is the failure level. If no level is specified, it defaults to 0. The level is used by
try
,repeat
,match goal
and the branching tacticals. If 0, it makesmatch goal
consider the next clause (backtracking). If nonzero, the currentmatch goal
block,try
,repeat
, or branching command is aborted and the level is decremented. In the case of+
, a nonzero level skips the first backtrack point, even if the call tofail num
is not enclosed in a+
command, respecting the algebraic identity.
-
Variant
fail message_token*
The given tokens are used for printing the failure message.
-
Variant
fail num message_token*
This is a combination of the previous variants.
-
Variant
gfail
¶ This variant fails even when used after
;
and there are no goals left. Similarly,gfail
fails even when used afterall:
and there are no goals left. See the example for clarification.
-
Variant
gfail message_token*
-
Variant
gfail num message_token*
These variants fail with an error message or an error level even if there are no goals left. Be careful however if Coq terms have to be printed as part of the failure: term construction always forces the tactic into the goals, meaning that if there are no goals when it is evaluated, a tactic call like
let x := H in fail 0 x
will succeed.
-
Error
No such goal.
¶
Example
- Goal True.
- 1 subgoal ============================ True
- Proof.
- fail.
- Toplevel input, characters 0-5: > fail. > ^^^^^ Error: Tactic failure.
- Abort.
- Goal True.
- 1 subgoal ============================ True
- Proof.
- trivial; fail.
- No more subgoals.
- Qed.
- Goal True.
- 1 subgoal ============================ True
- Proof.
- trivial.
- No more subgoals.
- fail.
- Toplevel input, characters 0-5: > fail. > ^^^^^ Error: No such goal.
- Abort.
- Goal True.
- 1 subgoal ============================ True
- Proof.
- trivial.
- No more subgoals.
- all: fail.
- Qed.
- Goal True.
- 1 subgoal ============================ True
- Proof.
- gfail.
- Toplevel input, characters 0-6: > gfail. > ^^^^^^ Error: Tactic failure.
- Abort.
- Goal True.
- 1 subgoal ============================ True
- Proof.
- trivial; gfail.
- Toplevel input, characters 0-7: > trivial; gfail. > ^^^^^^^ Error: Tactic failure.
- Abort.
- Goal True.
- 1 subgoal ============================ True
- Proof.
- trivial.
- No more subgoals.
- gfail.
- Toplevel input, characters 0-6: > gfail. > ^^^^^^ Error: No such goal.
- Abort.
- Goal True.
- 1 subgoal ============================ True
- Proof.
- trivial.
- No more subgoals.
- all: gfail.
- Toplevel input, characters 0-11: > all: gfail. > ^^^^^^^^^^^ Error: Tactic failure.
- Abort.
-
Variant
Timeout¶
We can force a tactic to stop if it has not finished after a certain amount of time:
-
timeout num ltac_expr
¶ ltac_expr
is evaluated tov
which must be a tactic value. The tactic valuev
is applied normally, except that it is interrupted afternum
seconds if it is still running. In this case the outcome is a failure.Warning
For the moment, timeout is based on elapsed time in seconds, which is very machine-dependent: a script that works on a quick machine may fail on a slow one. The converse is even possible if you combine a timeout with some other tacticals. This tactical is hence proposed only for convenience during debugging or other development phases, we strongly advise you to not leave any timeout in final scripts. Note also that this tactical isn’t available on the native Windows port of Coq.
Timing a tactic¶
A tactic execution can be timed:
-
time string ltac_expr
¶ evaluates
ltac_expr
and displays the running time of the tactic expression, whether it fails or succeeds. In case of several successes, the time for each successive run is displayed. Time is in seconds and is machine-dependent. Thestring
argument is optional. When provided, it is used to identify this particular occurrence of time.
Timing a tactic that evaluates to a term¶
Tactic expressions that produce terms can be timed with the experimental tactic
-
time_constr ltac_expr
¶ which evaluates
ltac_expr ()
and displays the time the tactic expression evaluated, assuming successful evaluation. Time is in seconds and is machine-dependent.This tactic currently does not support nesting, and will report times based on the innermost execution. This is due to the fact that it is implemented using the following internal tactics:
-
finish_timing (string)? string
¶ Display an optionally named timer. The parenthesized string argument is also optional, and determines the label associated with the timer for printing.
By copying the definition of
time_constr
from the standard library, users can achieve support for a fixed pattern of nesting by passing differentstring
parameters torestart_timer
andfinish_timing
at each level of nesting.Example
- Ltac time_constr1 tac := let eval_early := match goal with _ => restart_timer "(depth 1)" end in let ret := tac () in let eval_early := match goal with _ => finish_timing ( "Tactic evaluation" ) "(depth 1)" end in ret.
- time_constr1 is defined
- Goal True.
- 1 subgoal ============================ True
- let v := time_constr ltac:(fun _ => let x := time_constr1 ltac:(fun _ => constr:(10 * 10)) in let y := time_constr1 ltac:(fun _ => eval compute in x) in y) in pose v.
- Tactic evaluation (depth 1) ran for 0. secs (0.u,0.s) Tactic evaluation (depth 1) ran for 0. secs (0.u,0.s) Tactic evaluation ran for 0.001 secs (0.u,0.s) 1 subgoal n := 100 : nat ============================ True
-
Local definitions¶
Local definitions can be done as follows:
-
let ident1 := ltac_expr1 with identi := ltac_expri* in ltac_expr
¶ each
ltac_expri
is evaluated tovi
, then,ltac_expr
is evaluated by substitutingvi
to each occurrence ofidenti
, for i = 1, ..., n. There are no dependencies between theltac_expri
and theidenti
.Local definitions can be made recursive by using
let rec
instead oflet
. In this latter case, the definitions are evaluated lazily so that the rec keyword can be used also in non-recursive cases so as to avoid the eager evaluation of local definitions.
Application¶
An application is an expression of the following form:
-
qualid tacarg+
The reference
qualid
must be bound to some defined tactic definition expecting at least as many arguments as the providedtacarg
. The expressionsltac_expri
are evaluated tovi
, for i = 1, ..., n.
Function construction¶
A parameterized tactic can be built anonymously (without resorting to local definitions) with:
Pattern matching on terms¶
We can carry out pattern matching on terms with:
-
match ltac_expr with cpatterni => ltac_expri+| end
The expression
ltac_expr
is evaluated and should yield a term which is matched againstcpattern1
. The matching is non-linear: if a metavariable occurs more than once, it should match the same expression every time. It is first-order except on the variables of the form@?id
that occur in head position of an application. For these variables, the matching is second-order and returns a functional term.Alternatively, when a metavariable of the form
?id
occurs under binders, sayx1, …, xn
and the expression matches, the metavariable is instantiated by a term which can then be used in any context which also binds the variablesx1, …, xn
with same types. This provides with a primitive form of matching under context which does not require manipulating a functional term.If the matching with
cpattern1
succeeds, thenltac_expr1
is evaluated into some value by substituting the pattern matching instantiations to the metavariables. Ifltac_expr1
evaluates to a tactic and the match expression is in position to be applied to a goal (e.g. it is not bound to a variable by alet in
), then this tactic is applied. If the tactic succeeds, the list of resulting subgoals is the result of the match expression. Ifltac_expr1
does not evaluate to a tactic or if the match expression is not in position to be applied to a goal, then the result of the evaluation ofltac_expr1
is the result of the match expression.If the matching with
cpattern1
fails, or if it succeeds but the evaluation ofltac_expr1
fails, or if the evaluation ofltac_expr1
succeeds but returns a tactic in execution position whose execution fails, thencpattern2
is used and so on. The pattern_
matches any term and shadows all remaining patterns if any. If all clauses fail (in particular, there is no pattern_
) then a no-matching-clause error is raised.Failures in subsequent tactics do not cause backtracking to select new branches or inside the right-hand side of the selected branch even if it has backtracking points.
-
Error
No matching clauses for match.
¶ No pattern can be used and, in particular, there is no
_
pattern.
-
Error
Argument of match does not evaluate to a term.
¶ This happens when
ltac_expr
does not denote a term.
-
Variant
multimatch ltac_expr with cpatterni => ltac_expri+| end
Using multimatch instead of match will allow subsequent tactics to backtrack into a right-hand side tactic which has backtracking points left and trigger the selection of a new matching branch when all the backtracking points of the right-hand side have been consumed.
The syntax
match …
is, in fact, a shorthand foronce multimatch …
.
-
Variant
lazymatch ltac_expr with cpatterni => ltac_expri+| end
Using lazymatch instead of match will perform the same pattern matching procedure but will commit to the first matching branch rather than trying a new matching if the right-hand side fails. If the right-hand side of the selected branch is a tactic with backtracking points, then subsequent failures cause this tactic to backtrack.
-
Variant
context ident [cpattern]
This special form of patterns matches any term with a subterm matching cpattern. If there is a match, the optional
ident
is assigned the "matched context", i.e. the initial term where the matched subterm is replaced by a hole. The example below will show how to use such term contexts.If the evaluation of the right-hand-side of a valid match fails, the next matching subterm is tried. If no further subterm matches, the next clause is tried. Matching subterms are considered top-bottom and from left to right (with respect to the raw printing obtained by setting the
Printing All
flag).
Example
- Ltac f x := match x with context f [S ?X] => idtac X; assert (p := eq_refl 1 : X=1); let x:= context f[O] in assert (x=O) end.
- f is defined
- Goal True.
- 1 subgoal ============================ True
- f (3+4).
- 2 1 2 subgoals p : 1 = 1 ============================ 1 + 4 = 0 subgoal 2 is: True
-
Error
Pattern matching on goals¶
We can perform pattern matching on goals using the following expression:
-
match goal with context_hyp+, |- cpattern => ltac_expr+| | _ => ltac_expr end
¶ If each hypothesis pattern
hyp
1,i, with i = 1, ..., m1 is matched (non-linear first-order unification) by a hypothesis of the goal and ifcpattern_1
is matched by the conclusion of the goal, thenltac_expr1
is evaluated tov1
by substituting the pattern matching to the metavariables and the real hypothesis names bound to the possible hypothesis names occurring in the hypothesis patterns. Ifv1
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 the lastltac_expr
is evaluated tov
andv
is applied. Note also that matching against subterms (using thecontext ident [ cpattern ]
) is available and is also subject to yielding several matchings.Failures in subsequent tactics do not cause backtracking to select new branches or combinations of hypotheses, or inside the right-hand side of the selected branch even if it has backtracking points.
-
Error
No matching clauses for match goal.
¶ No clause succeeds, i.e. all matching patterns, if any, fail at the application of the right-hand-side.
Note
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 right to left (i.e. hypi,mi` before hypi,1). For each hypothesis pattern, the goal hypotheses are matched in order (newest first), but it possible to reverse this order (oldest first) with the
match reverse goal with
variant.-
Variant
multimatch goal with context_hyp+, |- cpattern => ltac_expr+| | _ => ltac_expr end
Using
multimatch
instead ofmatch
will allow subsequent tactics to backtrack into a right-hand side tactic which has backtracking points left and trigger the selection of a new matching branch or combination of hypotheses when all the backtracking points of the right-hand side have been consumed.The syntax
match [reverse] goal …
is, in fact, a shorthand foronce multimatch [reverse] goal …
.
-
Variant
lazymatch goal with context_hyp+, |- cpattern => ltac_expr+| | _ => ltac_expr end
Using lazymatch instead of match will perform the same pattern matching procedure but will commit to the first matching branch with the first matching combination of hypotheses rather than trying a new matching if the right-hand side fails. If the right-hand side of the selected branch is a tactic with backtracking points, then subsequent failures cause this tactic to backtrack.
-
Error
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:
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. The following expression returns an identifier:
-
fresh component*
It evaluates to an identifier unbound in the goal. This fresh identifier is obtained by concatenating the value of the
component
s (each of them is, either aqualid
which has to refer to a (unqualified) name, or directly a name denoted by astring
).If the resulting name is already used, it is padded with a number so that it becomes fresh. If no component is given, the name is a fresh derivative of the name
H
.
Computing in a constr¶
Evaluation of a term can be performed with:
Manipulating untyped terms¶
-
uconstr : term
The terms built in
L
tac are well-typed by default. It may not be appropriate for building large terms using a recursiveL
tac function: the term has to be entirely type checked at each step, resulting in potentially very slow behavior. It is possible to build untyped terms usingL
tac with theuconstr : term
syntax.
-
type_term term
An untyped term, in
L
tac, can contain references to hypotheses or toL
tac variables containing typed or untyped terms. An untyped term can be type checked using the function type_term whose argument is parsed as an untyped term and returns a well-typed term which can be used in tactics.
Untyped terms built using uconstr :
can also be used as arguments to the
refine
tactic. In that case the untyped term is type
checked against the conclusion of the goal, and the holes which are not solved
by the typing procedure are turned into new subgoals.
Counting the goals¶
-
numgoals
The number of goals under focus can be recovered using the
numgoals
function. Combined with the guard command below, it can be used to branch over the number of goals produced by previous tactics.Example
- Ltac pr_numgoals := let n := numgoals in idtac "There are" n "goals".
- pr_numgoals is defined
- Goal True /\ True /\ True.
- 1 subgoal ============================ True /\ True /\ True
- split;[|split].
- 3 subgoals ============================ True subgoal 2 is: True subgoal 3 is: True
- all:pr_numgoals.
- There are 3 goals
Testing boolean expressions¶
-
guard test
¶ The
guard
tactic tests a boolean expression, and fails if the expression evaluates to false. If the expression evaluates to true, it succeeds without affecting the proof.The accepted tests are simple integer comparisons.
Example
- Goal True /\ True /\ True.
- 1 subgoal ============================ True /\ True /\ True
- split;[|split].
- 3 subgoals ============================ True subgoal 2 is: True subgoal 3 is: True
- all:let n:= numgoals in guard n<4.
- Fail all:let n:= numgoals in guard n=2.
- The command has indeed failed with message: Condition not satisfied: 3=2
-
Error
Condition not satisfied.
¶
Proving a subgoal as a separate lemma¶
-
abstract ltac_expr
¶ From the outside,
abstract ltac_expr
is the same assolve ltac_expr
. Internally it saves an auxiliary lemma calledident_subproofn
whereident
is the name of the current goal andn
is chosen so that this is a fresh name. Such an auxiliary lemma is inlined in the final proof term.This tactical is useful with tactics such as
omega
ordiscriminate
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.It may be useful to generate lemmas minimal w.r.t. the assumptions they depend on. This can be obtained thanks to the option below.
Warning
The abstract tactic, while very useful, still has some known limitations, see https://github.com/coq/coq/issues/9146 for more details. Thus we recommend using it caution in some "non-standard" contexts. In particular,
abstract
won't properly work when used inside quotationsltac:(...)
, or if used as part of typeclass resolution, it may produce wrong terms when in universe polymorphic mode.-
Variant
abstract ltac_expr using ident
Give explicitly the name of the auxiliary lemma.
Warning
Use this feature at your own risk; explicitly named and reused subterms don’t play well with asynchronous proofs.
-
Variant
transparent_abstract ltac_expr
¶ Save the subproof in a transparent lemma rather than an opaque one.
Warning
Use this feature at your own risk; building computationally relevant terms with tactics is fragile.
-
Variant
transparent_abstract ltac_expr using ident
Give explicitly the name of the auxiliary transparent lemma.
Warning
Use this feature at your own risk; building computationally relevant terms with tactics is fragile, and explicitly named and reused subterms don’t play well with asynchronous proofs.
-
Error
Proof is not complete.
¶
-
Variant
Tactic toplevel definitions¶
Defining L
tac functions¶
Basically, L
tac toplevel definitions are made as follows:
-
Command
Local? Ltac ident ident* := ltac_expr
¶ This defines a new
L
tac function that can be used in any tactic script or newL
tac toplevel definition.If preceded by the keyword
Local
, the tactic definition will not be exported outside the current module.
Examples of using L
tac¶
Proof that the natural numbers have at least two elements¶
Example: Proof that the natural numbers have at least two elements
The first example shows how to use pattern matching over the proof context to prove that natural numbers have at least two elements. This can be done as follows:
- Lemma card_nat : ~ exists x y : nat, forall z:nat, x = z \/ y = z.
- 1 subgoal ============================ ~ (exists x y : nat, forall z : nat, x = z \/ y = z)
- Proof.
- intros (x & y & Hz).
- 1 subgoal x, y : nat Hz : forall z : nat, x = z \/ y = z ============================ False
- destruct (Hz 0), (Hz 1), (Hz 2).
- 8 subgoals x, y : nat Hz : forall z : nat, x = z \/ y = z H : x = 0 H0 : x = 1 H1 : x = 2 ============================ False subgoal 2 is: False subgoal 3 is: False subgoal 4 is: False subgoal 5 is: False subgoal 6 is: False subgoal 7 is: False subgoal 8 is: False
At this point, the congruence
tactic would finish the job:
- all: congruence.
- No more subgoals.
But for the purpose of the example, let's craft our own custom tactic to solve this:
- Lemma card_nat : ~ exists x y : nat, forall z:nat, x = z \/ y = z.
- 1 subgoal ============================ ~ (exists x y : nat, forall z : nat, x = z \/ y = z)
- Proof.
- intros (x & y & Hz).
- 1 subgoal x, y : nat Hz : forall z : nat, x = z \/ y = z ============================ False
- destruct (Hz 0), (Hz 1), (Hz 2).
- 8 subgoals x, y : nat Hz : forall z : nat, x = z \/ y = z H : x = 0 H0 : x = 1 H1 : x = 2 ============================ False subgoal 2 is: False subgoal 3 is: False subgoal 4 is: False subgoal 5 is: False subgoal 6 is: False subgoal 7 is: False subgoal 8 is: False
- all: match goal with | _ : ?a = ?b, _ : ?a = ?c |- _ => assert (b = c) by now transitivity a end.
- 8 subgoals x, y : nat Hz : forall z : nat, x = z \/ y = z H : x = 0 H0 : x = 1 H1 : x = 2 H2 : 1 = 2 ============================ False subgoal 2 is: False subgoal 3 is: False subgoal 4 is: False subgoal 5 is: False subgoal 6 is: False subgoal 7 is: False subgoal 8 is: False
- all: discriminate.
- No more subgoals.
Notice that all the (very similar) cases coming from the three
eliminations (with three distinct natural numbers) are successfully
solved by a match goal
structure and, in particular, with only one
pattern (use of non-linear matching).
Proving that a list is a permutation of a second list¶
Example: Proving that a list is a permutation of a second list
Let's first define the permutation predicate:
- Section Sort.
- Variable A : Set.
- A is declared
- Inductive perm : list A -> list A -> Prop := | perm_refl : forall l, perm l l | perm_cons : forall a l0 l1, perm l0 l1 -> perm (a :: l0) (a :: l1) | perm_append : forall a l, perm (a :: l) (l ++ a :: nil) | perm_trans : forall l0 l1 l2, perm l0 l1 -> perm l1 l2 -> perm l0 l2.
- perm is defined perm_ind is defined perm_sind is defined
- End Sort.
- Require Import List.
Next we define an auxiliary tactic perm_aux
which takes an
argument used to control the recursion depth. This tactic works as
follows: If the lists are identical (i.e. convertible), it
completes the proof. Otherwise, if the lists have identical heads,
it looks at their tails. Finally, if the lists have different
heads, it rotates the first list by putting its head at the end.
Every time we perform a rotation, we decrement n
. When n
drops down to 1
, we stop performing rotations and we fail.
The idea is to give the length of the list as the initial value of
n
. This way of counting the number of rotations will avoid
going back to a head that had been considered before.
From Section Syntax we know that Ltac has a primitive
notion of integers, but they are only used as arguments for
primitive tactics and we cannot make computations with them. Thus,
instead, we use Coq's natural number type nat
.
- Ltac perm_aux n := match goal with | |- (perm _ ?l ?l) => apply perm_refl | |- (perm _ (?a :: ?l1) (?a :: ?l2)) => let newn := eval compute in (length l1) in (apply perm_cons; perm_aux newn) | |- (perm ?A (?a :: ?l1) ?l2) => match eval compute in n with | 1 => fail | _ => let l1' := constr:(l1 ++ a :: nil) in (apply (perm_trans A (a :: l1) l1' l2); [ apply perm_append | compute; perm_aux (pred n) ]) end end.
- perm_aux is defined
The main tactic is solve_perm
. It computes the lengths of the
two lists and uses them as arguments to call perm_aux
if the
lengths are equal. (If they aren't, the lists cannot be
permutations of each other.)
- Ltac solve_perm := match goal with | |- (perm _ ?l1 ?l2) => match eval compute in (length l1 = length l2) with | (?n = ?n) => perm_aux n end end.
- solve_perm is defined
And now, here is how we can use the tactic solve_perm
:
- Goal perm nat (1 :: 2 :: 3 :: nil) (3 :: 2 :: 1 :: nil).
- 1 subgoal ============================ perm nat (1 :: 2 :: 3 :: nil) (3 :: 2 :: 1 :: nil)
- solve_perm.
- No more subgoals.
- Goal perm nat (0 :: 1 :: 2 :: 3 :: 4 :: 5 :: 6 :: 7 :: 8 :: 9 :: nil) (0 :: 2 :: 4 :: 6 :: 8 :: 9 :: 7 :: 5 :: 3 :: 1 :: nil).
- 1 subgoal ============================ perm nat (0 :: 1 :: 2 :: 3 :: 4 :: 5 :: 6 :: 7 :: 8 :: 9 :: nil) (0 :: 2 :: 4 :: 6 :: 8 :: 9 :: 7 :: 5 :: 3 :: 1 :: nil)
- solve_perm.
- No more subgoals.
Deciding intuitionistic propositional logic¶
Pattern matching on goals allows powerful backtracking when returning tactic values. An interesting application is the problem of deciding intuitionistic propositional logic. Considering the contraction-free sequent calculi LJT* of Roy Dyckhoff [Dyc92], it is quite natural to code such a tactic using the tactic language as shown below.
- Ltac basic := match goal with | |- True => trivial | _ : False |- _ => contradiction | _ : ?A |- ?A => assumption end.
- basic is defined
- Ltac simplify := repeat (intros; match goal with | H : ~ _ |- _ => red in H | H : _ /\ _ |- _ => elim H; do 2 intro; clear H | H : _ \/ _ |- _ => elim H; intro; clear H | H : ?A /\ ?B -> ?C |- _ => cut (A -> B -> C); [ intro | intros; apply H; split; assumption ] | H: ?A \/ ?B -> ?C |- _ => cut (B -> C); [ cut (A -> C); [ intros; clear H | intro; apply H; left; assumption ] | intro; apply H; right; assumption ] | H0 : ?A -> ?B, H1 : ?A |- _ => cut B; [ intro; clear H0 | apply H0; assumption ] | |- _ /\ _ => split | |- ~ _ => red end).
- simplify is defined
- Ltac my_tauto := simplify; basic || match goal with | H : (?A -> ?B) -> ?C |- _ => cut (B -> C); [ intro; cut (A -> B); [ intro; cut C; [ intro; clear H | apply H; assumption ] | clear H ] | intro; apply H; intro; assumption ]; my_tauto | H : ~ ?A -> ?B |- _ => cut (False -> B); [ intro; cut (A -> False); [ intro; cut B; [ intro; clear H | apply H; assumption ] | clear H ] | intro; apply H; red; intro; assumption ]; my_tauto | |- _ \/ _ => (left; my_tauto) || (right; my_tauto) end.
- my_tauto is defined
The tactic basic
tries to reason using simple rules involving truth, falsity
and available assumptions. The tactic simplify
applies all the reversible
rules of Dyckhoff’s system. Finally, the tactic my_tauto
(the main
tactic to be called) simplifies with simplify
, tries to conclude with
basic
and tries several paths using the backtracking rules (one of the
four Dyckhoff’s rules for the left implication to get rid of the contraction
and the right or
).
Having defined my_tauto
, we can prove tautologies like these:
- Lemma my_tauto_ex1 : forall A B : Prop, A /\ B -> A \/ B.
- 1 subgoal ============================ forall A B : Prop, A /\ B -> A \/ B
- Proof.
- my_tauto.
- No more subgoals.
- Qed.
- Lemma my_tauto_ex2 : forall A B : Prop, (~ ~ B -> B) -> (A -> B) -> ~ ~ A -> B.
- 1 subgoal ============================ forall A B : Prop, (~ ~ B -> B) -> (A -> B) -> ~ ~ A -> B
- Proof.
- my_tauto.
- No more subgoals.
- Qed.
Deciding type isomorphisms¶
A trickier problem is to decide equalities between types modulo isomorphisms. Here, we choose to use the isomorphisms of the simply typed λ-calculus with Cartesian product and unit type (see, for example, [dC95]). The axioms of this λ-calculus are given below.
- Open Scope type_scope.
- Section Iso_axioms.
- Variables A B C : Set.
- A is declared B is declared C is declared
- Axiom Com : A * B = B * A.
- Com is declared
- Axiom Ass : A * (B * C) = A * B * C.
- Ass is declared
- Axiom Cur : (A * B -> C) = (A -> B -> C).
- Cur is declared
- Axiom Dis : (A -> B * C) = (A -> B) * (A -> C).
- Dis is declared
- Axiom P_unit : A * unit = A.
- P_unit is declared
- Axiom AR_unit : (A -> unit) = unit.
- AR_unit is declared
- Axiom AL_unit : (unit -> A) = A.
- AL_unit is declared
- Lemma Cons : B = C -> A * B = A * C.
- 1 subgoal A, B, C : Set ============================ B = C -> A * B = A * C
- Proof.
- intro Heq; rewrite Heq; reflexivity.
- No more subgoals.
- Qed.
- End Iso_axioms.
- Ltac simplify_type ty := match ty with | ?A * ?B * ?C => rewrite <- (Ass A B C); try simplify_type_eq | ?A * ?B -> ?C => rewrite (Cur A B C); try simplify_type_eq | ?A -> ?B * ?C => rewrite (Dis A B C); try simplify_type_eq | ?A * unit => rewrite (P_unit A); try simplify_type_eq | unit * ?B => rewrite (Com unit B); try simplify_type_eq | ?A -> unit => rewrite (AR_unit A); try simplify_type_eq | unit -> ?B => rewrite (AL_unit B); try simplify_type_eq | ?A * ?B => (simplify_type A; try simplify_type_eq) || (simplify_type B; try simplify_type_eq) | ?A -> ?B => (simplify_type A; try simplify_type_eq) || (simplify_type B; try simplify_type_eq) end with simplify_type_eq := match goal with | |- ?A = ?B => try simplify_type A; try simplify_type B end.
- simplify_type is defined simplify_type_eq is defined
- Ltac len trm := match trm with | _ * ?B => let succ := len B in constr:(S succ) | _ => constr:(1) end.
- len is defined
- Ltac assoc := repeat rewrite <- Ass.
- assoc is defined
- Ltac solve_type_eq n := match goal with | |- ?A = ?A => reflexivity | |- ?A * ?B = ?A * ?C => apply Cons; let newn := len B in solve_type_eq newn | |- ?A * ?B = ?C => match eval compute in n with | 1 => fail | _ => pattern (A * B) at 1; rewrite Com; assoc; solve_type_eq (pred n) end end.
- solve_type_eq is defined
- Ltac compare_structure := match goal with | |- ?A = ?B => let l1 := len A with l2 := len B in match eval compute in (l1 = l2) with | ?n = ?n => solve_type_eq n end end.
- compare_structure is defined
- Ltac solve_iso := simplify_type_eq; compare_structure.
- solve_iso is defined
The tactic to judge equalities modulo this axiomatization is shown above.
The algorithm is quite simple. First types are simplified using axioms that
can be oriented (this is done by simplify_type
and simplify_type_eq
).
The normal forms are sequences of Cartesian products without a Cartesian product
in the left component. These normal forms are then compared modulo permutation
of the components by the tactic compare_structure
. If they have the same
length, the tactic solve_type_eq
attempts to prove that the types are equal.
The main tactic that puts all these components together is solve_iso
.
Here are examples of what can be solved by solve_iso
.
- Lemma solve_iso_ex1 : forall A B : Set, A * unit * B = B * (unit * A).
- 1 subgoal ============================ forall A B : Set, A * unit * B = B * (unit * A)
- Proof.
- intros; solve_iso.
- No more subgoals.
- Qed.
- Lemma solve_iso_ex2 : forall A B C : Set, (A * unit -> B * (C * unit)) = (A * unit -> (C -> unit) * C) * (unit -> A -> B).
- 1 subgoal ============================ forall A B C : Set, (A * unit -> B * (C * unit)) = (A * unit -> (C -> unit) * C) * (unit -> A -> B)
- Proof.
- intros; solve_iso.
- No more subgoals.
- Qed.
Debugging L
tac tactics¶
Backtraces¶
-
Flag
Ltac Backtrace
¶ Setting this flag displays a backtrace on Ltac failures that can be useful to find out what went wrong. It is disabled by default for performance reasons.
Info trace¶
-
Command
Info num ltac_expr
¶ This command can be used to print the trace of the path eventually taken by an
L
tac script. That is, the list of executed tactics, discarding all the branches which have failed. To that end theInfo
command can be used with the following syntax.The number
num
is the unfolding level of tactics in the trace. At level 0, the trace contains a sequence of tactics in the actual script, at level 1, the trace will be the concatenation of the traces of these tactics, etc…Example
- Ltac t x := exists x; reflexivity.
- t is defined
- Goal exists n, n=0.
- 1 subgoal ============================ exists n : nat, n = 0
- Info 0 t 1||t 0.
- exists with 0;<ltac_plugin::reflexivity@0> No more subgoals.
- Undo.
- 1 subgoal ============================ exists n : nat, n = 0
- Info 1 t 1||t 0.
- <ltac_plugin::exists@1> with 0;simple refine ?X12;<unknown> No more subgoals.
The trace produced by
Info
tries its best to be a reparsableL
tac script, but this goal is not achievable in all generality. So some of the output traces will contain oddities.As an additional help for debugging, the trace produced by
Info
contains (in comments) the messages produced by theidtac
tactical at the right position in the script. In particular, the calls to idtac in branches which failed are not printed.
Interactive debugger¶
-
Flag
Ltac Debug
¶ This flag governs the step-by-step debugger that comes with the
L
tac interpreter.
When the debugger is activated, it stops at every step of the evaluation of
the current L
tac expression and prints information on what it is doing.
The debugger stops, prompting for a command which can be one of the
following:
simple newline: | go to the next step |
h: | get help |
x: | exit current evaluation |
s: | continue current evaluation without stopping |
r n: | advance n steps further |
r string: | advance up to the next call to “idtac string” |
-
Error
Debug mode not available in the IDE
¶
A non-interactive mode for the debugger is available via the flag:
-
Flag
Ltac Batch Debug
¶ This flag has the effect of presenting a newline at every prompt, when the debugger is on. The debug log thus created, which does not require user input to generate when this flag is set, can then be run through external tools such as diff.
Profiling L
tac tactics¶
It is possible to measure the time spent in invocations of primitive
tactics as well as tactics defined in L
tac and their inner
invocations. The primary use is the development of complex tactics,
which can sometimes be so slow as to impede interactive usage. The
reasons for the performance degradation can be intricate, like a slowly
performing L
tac match or a sub-tactic whose performance only
degrades in certain situations. The profiler generates a call tree and
indicates the time spent in a tactic depending on its calling context. Thus
it allows to locate the part of a tactic definition that contains the
performance issue.
-
Flag
Ltac Profiling
¶ This flag enables and disables the profiler.
-
Command
Show Ltac Profile
¶ Prints the profile
-
Command
Reset Ltac Profile
¶ Resets the profile, that is, deletes all accumulated information.
Warning
Backtracking across a
Reset Ltac Profile
will not restore the information.
- Require Import Coq.omega.Omega.
- [Loading ML file newring_plugin.cmxs ... done] [Loading ML file zify_plugin.cmxs ... done] [Loading ML file omega_plugin.cmxs ... done]
- Ltac mytauto := tauto.
- mytauto is defined
- Ltac tac := intros; repeat split; omega || mytauto.
- tac is defined
- Notation max x y := (x + (y - x)) (only parsing).
- Goal forall x y z A B C D E F G H I J K L M N O P Q R S T U V W X Y Z, max x (max y z) = max (max x y) z /\ max x (max y z) = max (max x y) z /\ (A /\ B /\ C /\ D /\ E /\ F /\ G /\ H /\ I /\ J /\ K /\ L /\ M /\ N /\ O /\ P /\ Q /\ R /\ S /\ T /\ U /\ V /\ W /\ X /\ Y /\ Z -> Z /\ Y /\ X /\ W /\ V /\ U /\ T /\ S /\ R /\ Q /\ P /\ O /\ N /\ M /\ L /\ K /\ J /\ I /\ H /\ G /\ F /\ E /\ D /\ C /\ B /\ A).
- 1 subgoal ============================ forall (x y z : nat) (A B C D E F G H I J K L M N O P Q R S T U V W X Y Z : Prop), x + (y + (z - y) - x) = x + (y - x) + (z - (x + (y - x))) /\ x + (y + (z - y) - x) = x + (y - x) + (z - (x + (y - x))) /\ (A /\ B /\ C /\ D /\ E /\ F /\ G /\ H /\ I /\ J /\ K /\ L /\ M /\ N /\ O /\ P /\ Q /\ R /\ S /\ T /\ U /\ V /\ W /\ X /\ Y /\ Z -> Z /\ Y /\ X /\ W /\ V /\ U /\ T /\ S /\ R /\ Q /\ P /\ O /\ N /\ M /\ L /\ K /\ J /\ I /\ H /\ G /\ F /\ E /\ D /\ C /\ B /\ A)
- Proof.
- Set Ltac Profiling.
- tac.
- No more subgoals.
- Show Ltac Profile.
- total time: 3.538s tactic local total calls max ────────────────────────────────────────┴──────┴──────┴───────┴─────────┘ ─tac ----------------------------------- 0.1% 100.0% 1 3.538s ─<Coq.Init.Tauto.with_uniform_flags> --- 0.0% 76.3% 26 0.205s ─<Coq.Init.Tauto.tauto_gen> ------------ 0.0% 76.3% 26 0.205s ─<Coq.Init.Tauto.tauto_intuitionistic> - 0.0% 76.2% 26 0.205s ─t_tauto_intuit ------------------------ 0.1% 76.2% 26 0.205s ─<Coq.Init.Tauto.simplif> -------------- 50.7% 73.5% 26 0.201s ─omega --------------------------------- 23.3% 23.3% 28 0.334s ─<Coq.Init.Tauto.is_conj> -------------- 14.1% 14.1% 28756 0.116s ─elim id ------------------------------- 5.9% 5.9% 650 0.027s ─<Coq.Init.Tauto.axioms> --------------- 2.2% 2.7% 0 0.012s tactic local total calls max ────────────────────────────────────────┴──────┴──────┴───────┴─────────┘ ─tac ----------------------------------- 0.1% 100.0% 1 3.538s ├─<Coq.Init.Tauto.with_uniform_flags> - 0.0% 76.3% 26 0.205s │└<Coq.Init.Tauto.tauto_gen> ---------- 0.0% 76.3% 26 0.205s │└<Coq.Init.Tauto.tauto_intuitionistic> 0.0% 76.2% 26 0.205s │└t_tauto_intuit ---------------------- 0.1% 76.2% 26 0.205s │ ├─<Coq.Init.Tauto.simplif> ---------- 50.7% 73.5% 26 0.201s │ │ ├─<Coq.Init.Tauto.is_conj> -------- 14.1% 14.1% 28756 0.116s │ │ └─elim id ------------------------- 5.9% 5.9% 650 0.027s │ └─<Coq.Init.Tauto.axioms> ----------- 2.2% 2.7% 0 0.012s └─omega ------------------------------- 23.3% 23.3% 28 0.334s
- Show Ltac Profile "omega".
- total time: 3.538s tactic local total calls max ────────────────────────────────────────┴──────┴──────┴───────┴─────────┘ ─omega --------------------------------- 23.3% 23.3% 28 0.334s tactic local total calls max
- Abort.
- Unset Ltac Profiling.
-
stop ltac profiling
¶ Similarly to
start ltac profiling
, this tactic behaves likeidtac
. Together, they allow you to exclude parts of a proof script from profiling.
-
reset ltac profile
¶ This tactic behaves like the corresponding vernacular command and allow displaying and resetting the profile from tactic scripts for benchmarking purposes.
-
show ltac profile
¶ This tactic behaves like the corresponding vernacular command and allow displaying and resetting the profile from tactic scripts for benchmarking purposes.
-
show ltac profile string
This tactic behaves like the corresponding vernacular command and allow displaying and resetting the profile from tactic scripts for benchmarking purposes.
You can also pass the -profile-ltac
command line option to coqc
, which
turns the Ltac Profiling
flag on at the beginning of each document,
and performs a Show Ltac Profile
at the end.
Warning
Note that the profiler currently does not handle backtracking into multi-success tactics, and issues a warning to this effect in many cases when such backtracking occurs.
Run-time optimization tactic¶
-
optimize_heap
¶ This tactic behaves like
idtac
, except that running it compacts the heap in the OCaml run-time system. It is analogous to the Vernacular commandOptimize Heap
.