6. Tactics¶
6.1. Tactic Mode¶
Anywhere an expression is expected, Lean will accept a sequence of instructions bracketed by the keywords begin
and end
. The input between these keywords represents a tactic, typically a compound sequence of basic tactics, each possibly applied to suitable arguments, separated by commas. When processing such a tactic block, Lean’s elaborator executes the compound tactic with the expectation that it will produce an expression of the required type.
Individual tactics act on one or more goals, each of the form a : α ⊢ p
, where a : α
is a context and p
is the target type. Tactics are typically used to prove a theorem, in which case p
is a Prop
, but they can be used to construct an element of an arbitrary Type
as well.
At the outset, the elaborator presents the tactic block with a goal that consists of the local context in which the expression is being elaborated together with its expected type. Individual tactics can change goals and introduce new subgoals. A sequence of tactics is done when no subgoals remain, that is, when the compound tactic has succeeded in constructing an expression of the requisite type.
Tactics can fail. For example, a tactic may fail to make progress, or may not be appropriate to the goal. Other tactics can catch or handle those failures (see Section 6.6), but otherwise an error message is presented to the user.
Results produced by tactics are checked by the kernel for correctness. This provides another possible point of failure: a tactic block can, in principle, claim success but produce a term that fails to type check.
Tactics are themselves Lean expressions of a special tactic
type. This makes it possible to implement Lean tactics in Lean itself; see Chapter 8. Tactics in a begin ... end
block, however, are parsed in a special interactive mode that provides a more convenient manner of expression. In this section, we will focus exclusively on this interactive syntax.
You can use the keyword by
instead of begin ... end
to invoke a single tactic rather than a comma-separated sequence.
example (p q : Prop) : p ∧ q → q ∧ p :=
begin
intro h,
cases h,
split,
repeat { assumption }
end
example (p q : Prop) : p ∧ q → q ∧ p :=
assume ⟨h₁, h₂⟩,
and.intro (by assumption) (by assumption)
The documentation below coincides with documentation strings that are stored in the Lean source files and displayed by editors. The argument types are as follows:
id
: an identifierexpr
: an expression<binders>
: a sequence of identifiers and expressions(a : α)
wherea
is an identifier andα
is aType
or aProp
.
An annotation t?
means that the argument t
is optional, and an annotation t*
means any number of instances, possibly none. Many tactics parse arguments with additional tokens like with
, at
, only
, *
, or ⊢
, as indicated below. The token *
is typically used to denote all the hypotheses, and ⊢
is typically used to denote the goal, with ascii equivalent |-
.
6.2. Basic Tactics¶
intro id?
If the current goal is a Pi / forall
∀ x : t, u
(resp.let x := t in u
) thenintro
putsx : t
(resp.x := t
) in the local context. The new subgoal target isu
.If the goal is an arrow
t → u
, then it putsh : t
in the local context and the new goal target isu
.If the goal is neither a Pi/forall nor begins with a let binder, the tactic
intro
applies the tacticwhnf
until the tacticintro
can be applied or the goal is not head reducible. In the latter case, the tactic fails.
intros id*
Similar to
intro
tactic. The tacticintros
will keep introducing new hypotheses until the goal target is not a Pi/forall or let binder.The variant
intros h₁ ... hₙ
introducesn
new hypotheses using the given identifiers to name them.
introv id*
The tactic
introv
allows the user to automatically introduce the variables of a theorem and explicitly name the hypotheses involved. The given names are used to name non-dependent hypotheses.Examples:
example : ∀ a b : nat, a = b → b = a := begin introv h, exact h.symm endThe state after
introv h
isa b : ℕ, h : a = b ⊢ b = aexample : ∀ a b : nat, a = b → ∀ c, b = c → a = c := begin introv h₁ h₂, exact h₁.trans h₂ endThe state after
introv h₁ h₂
isa b : ℕ, h₁ : a = b, c : ℕ, h₂ : b = c ⊢ a = c
rename id id
The tacticrename h₁ h₂
renames hypothesish₁
toh₂
in the current local context.
apply expr
The
apply
tactic tries to match the current goal against the conclusion of the type of term. The argument term should be a term well-formed in the local context of the main goal. If it succeeds, then the tactic returns as many subgoals as the number of premises that have not been fixed by type inference or type class resolution. Non-dependent premises are added before dependent ones.The
apply
tactic uses higher-order pattern matching, type class resolution, and first-order unification with dependent types.
fapply expr
Similar to theapply
tactic, but does not reorder goals.
eapply expr
Similar to theapply
tactic, but only creates subgoals for non-dependent premises that have not been fixed by type inference or type class resolution.
apply_with expr (tactic.apply_cfg)
Similar to theapply
tactic, but allows the user to provide aapply_cfg
configuration object.
apply_instance
This tactic tries to close the main goal... ⊢ t
by generating a term of typet
using type class resolution.
refine expr
This tactic behaves like
exact
, but with a big difference: the user can put underscores_
in the expression as placeholders for holes that need to be filled, andrefine
will generate as many subgoals as there are holes.Note that some holes may be implicit. The type of each hole must either be synthesized by the system or declared by an explicit type ascription like
(_ : nat → Prop)
.
assumption
This tactic looks in the local context for a hypothesis whose type is equal to the goal target. If it finds one, it uses it to prove the goal, and otherwise it fails.
change expr (with expr)? (at (* | (⊢ | id)*))?
change u
replaces the targett
of the main goal tou
provided thatt
is well formed with respect to the local context of the main goal andt
andu
are definitionally equal.
change u at h
will change a local hypothesis tou
.
change t with u at h1 h2 ...
will replacet
withu
in all the supplied hypotheses (or*
), or in the goal if noat
clause is specified, provided thatt
andu
are definitionally equal.
exact expr
This tactic provides an exact proof term to solve the main goal. Ift
is the goal andp
is a term of typeu
thenexact p
succeeds if and only ift
andu
can be unified.
exacts ([expr, ...] | expr)
Likeexact
, but takes a list of terms and checks that all goals are discharged after the tactic.
revert id*
revert h₁ ... hₙ
applies to any goal with hypothesesh₁ ... hₙ
. It moves the hypotheses and their dependencies to the target of the goal. This tactic is the inverse of intro.
generalize id? : expr = id
generalize : e = x
replaces all occurrences ofe
in the target with a new hypothesisx
of the same type.
generalize h : e = x
in addition registers the hypothesish : e = x
.
admit
Closes the main goal usingsorry
.
contradiction
The contradiction tactic attempts to find in the current local context an hypothesis that is equivalent to an empty inductive type (e.g.false
), a hypothesis of the formc_1 ... = c_2 ...
wherec_1
andc_2
are distinct constructors, or two contradictory hypotheses.
trivial
Tries to solve the current goal using a canonical proof oftrue
, or thereflexivity
tactic, or the contradiction tactic.
exfalso
Replaces the target of the main goal byfalse
.
clear id*
clear h₁ ... hₙ
tries to clear each hypothesishᵢ
from the local context.
specialize expr
The tacticspecialize h a₁ ... aₙ
works on local hypothesish
. The premises of this hypothesis, either universal quantifications or non-dependent implications, are instantiated by concrete terms coming either from argumentsa₁
…aₙ
. The tactic adds a new hypothesis with the same nameh := h a₁ ... aₙ
and tries to clear the previous one.
by_cases expr (with id)?
by_cases p with h
splits the main goal into two cases, assumingh : p
in the first branch, andh : ¬ p
in the second branch.This tactic requires that
p
is decidable. To ensure that all propositions are decidable via classical reasoning, uselocal attribute classical.prop_decidable [instance]
.
by_contradiction id?
If the target of the main goal is a proposition
p
,by_contradiction h
reduces to goal to provingfalse
using the additional hypothesish : ¬ p
. Ifh
is omitted, a name is generated automatically.This tactic requires that
p
is decidable. To ensure that all propositions are decidable via classical reasoning, uselocal attribute classical.prop_decidable [instance]
.
by_contra id?
An abbreviation forby_contradiction
.
6.3. Equality and Other Relations¶
reflexivity
This tactic applies to a goal whose goal has the formt ~ u
where~
is a reflexive relation, that is, a relation which has a reflexivity lemma tagged with the attribute[refl]
. The tactic checks whethert
andu
are definitionally equal and then solves the goal.
refl
Shorter name for the tacticreflexivity
.
symmetry
This tactic applies to a goal whose target has the formt ~ u
where~
is a symmetric relation, that is, a relation which has a symmetry lemma tagged with the attribute[symm]
. It replaces the goal withu ~ t
.
transitivity ?expr
This tactic applies to a goal whose target has the form
t ~ u
where~
is a transitive relation, that is, a relation which has a transitivity lemma tagged with the attribute[trans]
.
transitivity s
replaces the goal with the two subgoalst ~ s
ands ~ u
. Ifs
is omitted, then a metavariable is used instead.
6.4. Structured Tactic Proofs¶
Tactic blocks can have nested begin ... end
blocks and, equivalently, blocks { ... }
enclosed with curly braces. Opening such a block focuses on the current goal, so that no other goals are visible within the nested block. Closing a block while any subgoals remain results in an error.
assume (: expr | <binders>)
Assuming the target of the goal is a Pi or a let,
assume h : t
unifies the type of the binder witht
and introduces it with nameh
, just likeintro h
. Ifh
is absent, the tactic uses the namethis
. IfT
is omitted, it will be inferred.
assume (h₁ : t₁) ... (hₙ : tₙ)
introduces multiple hypotheses. Any of the types may be omitted, but the names must be present.
have id? (: expr)? (:= expr)?
have h : t := p
adds the hypothesish : t
to the current goal ifp
a term of typet
. Ift
is omitted, it will be inferred.
have h : t
adds the hypothesish : t
to the current goal and opens a new subgoal with targett
. The new subgoal becomes the main goal. Ift
is omitted, it will be replaced by a fresh metavariable.If
h
is omitted, the namethis
is used.
let id? (: expr)? (:= expr)?
let h : T := p
adds the hypothesish : t := p
to the current goal ifp
a term of typet
. If t is omitted, it will be inferred.
let h : t
adds the hypothesish : t := ?M
to the current goal and opens a new subgoal?M : t
. The new subgoal becomes the main goal. Ift
is omitted, it will be replaced by a fresh metavariable.If
h
is omitted, the namethis
is used.
suffices id? (: expr)?
suffices h : t
is the same ashave h : t, tactic.swap
. In other words, it adds the hypothesish : t
to the current goal and opens a new subgoal with targett
.
show expr
show t
finds the first goal whose target unifies witht
. It makes that the main goal, performs the unification, and replaces the target with the unified version oft
.
from expr
A synonym forexact
that allows writinghave/suffices/show ..., from ...
in tactic mode.
variables (p q : Prop)
example : p ∧ (p → q) → q ∧ p :=
begin
assume h : p ∧ (p → q),
have h₁ : p, from and.left h,
have : p → q := and.right h,
suffices : q, from and.intro this h₁,
show q, from ‹p → q› h₁
end
example (p q : Prop) : p → p → p :=
begin
assume h (h' : p),
from h
end
example : ∃ x, x = 5 :=
begin
let u := 3 + 2,
existsi u, reflexivity
end
6.5. Inductive Types¶
The following tactics are designed specifically to work with elements on an inductive type.
induction expr (using id)? (with id*)? (generalizing id*)?
Assuming
x
is a variable in the local context with an inductive type,induction x
applies induction onx
to the main goal, producing one goal for each constructor of the inductive type, in which the target is replaced by a general instance of that constructor and an inductive hypothesis is added for each recursive argument to the constructor. If the type of an element in the local context depends onx
, that element is reverted and reintroduced afterward, so that the inductive hypothesis incorporates that hypothesis as well.For example, given
n : nat
and a goal with a hypothesish : P n
and targetQ n
,induction n
produces one goal with hypothesish : P 0
and targetQ 0
, and one goal with hypothesesh : P (nat.succ a)
andih₁ : P a → Q a
and targetQ (nat.succ a)
. Here the namesa
andih₁
ire chosen automatically.
induction e
, wheree
is an expression instead of a variable, generalizese
in the goal, and then performs induction on the resulting variable.
induction e with y₁ ... yₙ
, wheree
is a variable or an expression, specifies that the sequence of namesy₁ ... yₙ
should be used for the arguments to the constructors and inductive hypotheses, including implicit arguments. If the list does not include enough names for all of the arguments, additional names are generated automatically. If too many names are given, the extra ones are ignored. Underscores can be used in the list, in which case the corresponding names are generated automatically.
induction e using r
allows the user to specify the principle of induction that should be used. Herer
should be a theorem whose result type must be of the formC t
, whereC
is a bound variable andt
is a (possibly empty) sequence of bound variables
induction e generalizing z₁ ... zₙ
, wherez₁ ... zₙ
are variables in the local context, generalizes overz₁ ... zₙ
before applying the induction but then introduces them in each goal. In other words, the net effect is that each inductive hypothesis is generalized.
cases (id :)? expr (with id*)?
Assuming
x
is a variable in the local context with an inductive type,cases x
splits the main goal, producing one goal for each constructor of the inductive type, in which the target is replaced by a general instance of that constructor. If the type of an element in the local context depends onx
, that element is reverted and reintroduced afterward, so that the case split affects that hypothesis as well.For example, given
n : nat
and a goal with a hypothesish : P n
and targetQ n
,cases n
produces one goal with hypothesish : P 0
and targetQ 0
, and one goal with hypothesish : P (nat.succ a)
and targetQ (nat.succ a)
. Here the namea
is chosen automatically.
cases e
, wheree
is an expression instead of a variable, generalizese
in the goal, and then cases on the resulting variable.
cases e with y₁ ... yₙ
, wheree
is a variable or an expression, specifies that the sequence of namesy₁ ... yₙ
should be used for the arguments to the constructors, including implicit arguments. If the list does not include enough names for all of the arguments, additional names are generated automatically. If too many names are given, the extra ones are ignored. Underscores can be used in the list, in which case the corresponding names are generated automatically.
cases h : e
, wheree
is a variable or an expression, performs cases one
as above, but also adds a hypothesish : e = ...
to each hypothesis, where...
is the constructor instance for that particular case.
case id id* { tactic }
Focuses on the
induction
/cases
subgoal corresponding to the given introduction rule, optionally renaming introduced locals.example (n : ℕ) : n = n := begin induction n, case nat.zero { reflexivity }, case nat.succ a ih { reflexivity } end
destruct expr
Assuming
x
is a variable in the local context with an inductive type,destruct x
splits the main goal, producing one goal for each constructor of the inductive type, in whichx
is assumed to be a general instance of that constructor. In contrast tocases
, the local context is unchanged, i.e. no elements are reverted or introduced.For example, given
n : nat
and a goal with a hypothesish : P n
and targetQ n
,destruct n
produces one goal with targetn = 0 → Q n
, and one goal with target∀ (a : ℕ), (λ (w : ℕ), n = w → Q n) (nat.succ a)
. Here the namea
is chosen automatically.
existsi
existsi e
will instantiate an existential quantifier in the target withe
and leave the instantiated body as the new target. More generally, it applies to any inductive type with one constructor and at least two arguments, applying the constructor withe
as the first argument and leaving the remaining arguments as goals.
existsi [e₁, ..., eₙ]
iteratively does the same for each expression in the list.
constructor
This tactic applies to a goal such that its conclusion is an inductive type (sayI
). It tries to apply each constructor ofI
until it succeeds.
econstructor
Similar toconstructor
, but only non-dependent premises are added as new goals.
left
Applies the first constructor when the type of the target is an inductive data type with two constructors.
right
Applies the second constructor when the type of the target is an inductive data type with two constructors.
split
Applies the constructor when the type of the target is an inductive data type with one constructor.
injection expr (with id*)?
The
injection
tactic is based on the fact that constructors of inductive data types are injections. That means that ifc
is a constructor of an inductive datatype, and if(c t₁)
and(c t₂)
are two terms that are equal thent₁
andt₂
are equal too.If
q
is a proof of a statement of conclusiont₁ = t₂
, then injection applies injectivity to derive the equality of all arguments oft₁
andt₂
placed in the same positions. For example, from(a::b) = (c::d)
we derivea=c
andb=d
. To use this tactict₁
andt₂
should be constructor applications of the same constructor.Given
h : a::b = c::d
, the tacticinjection h
adds two new hypothesis with typesa = c
andb = d
to the main goal. The tacticinjection h with h₁ h₂
uses the namesh₁
anh₂
to name the new hypotheses.
injections (with id*)?
injections with h₁ ... hₙ
iteratively appliesinjection
to hypotheses using the namesh₁ ... hₙ
.
6.6. Tactic Combinators¶
Tactic combinators build compound tactics from simpler ones.
repeat { tactic }
repeat { t }
repeatedly appliest
untilt
fails. The compound tactic always succeeds.
try { tactic }
try { t }
tries to apply tactict
, but succeeds whether or nott
succeeds.
skip
A do-nothing tactic that always succeeds.
solve1 { tactic }
solve1 { t }
applies the tactict
to the main goal and fails if it is not solved.
abstract id? { tactic }
abstract id { t }
tries to use tactict
to solve the main goal. If it succeeds, it abstracts the goal as an independent definition or theorem with nameid
. Ifid
is omitted, a name is generated automatically.
all_goals { tactic }
all_goals { t }
applies the tactict
to every goal, and succeeds if each application succeeds.
any_goals { tactic }
any_goals { t }
applies the tactict
to every goal, and succeeds if at least one application succeeds.
done
Fail if there are unsolved goals.
fail_if_success { tactic }
Fails if the given tactic succeeds.
success_if_fail { tactic }
Succeeds if the given tactic succeeds.
guard_target expr
guard_target t
fails if the target of the main goal is nott
.
guard_hyp id := expr
guard_hyp h := t
fails if the hypothesish
does not have typet
.
6.7. The Rewriter¶
rewrite ([ (←? expr), ... ] | ←? expr) (at (* | (⊢ | id)*))? tactic.rewrite_cfg?
rewrite e
applies identitye
as a rewrite rule to the target of the main goal. Ife
is preceded by left arrow (←
or<-
), the rewrite is applied in the reverse direction. Ife
is a defined constant, then the equational lemmas associated withe
are used. This provides a convenient way to unfolde
.
rewrite [e₁, ..., eₙ]
applies the given rules sequentially.
rewrite e at l
rewritese
at location(s)l
, wherel
is either*
or a list of hypotheses in the local context. In the latter case, a turnstile⊢
or|-
can also be used, to signify the target of the goal.
rw
An abbreviation forrewrite
.
rwa
rewrite
followed byassumption
.
erewrite
A variant ofrewrite
that uses the unifier more aggressively, unfolding semireducible definitions.
erw
An abbreviation forerewrite
.
subst expr
Given hypothesish : x = t
orh : t = x
, wherex
is a local constant,subst h
substitutesx
byt
everywhere in the main goal and then clearsh
.
6.8. The Simplifier and Congruence Closure¶
simp only? (* | [(* | (- id | expr)), ...]?) (with id*)? (at (* | (⊢ | id)*))? tactic.simp_config_ext?
The
simp
tactic uses lemmas and hypotheses to simplify the main goal target or non-dependent hypotheses. It has many variants.
simp
simplifies the main goal target using lemmas tagged with the attribute[simp]
.
simp [h₁ h₂ ... hₙ]
simplifies the main goal target using the lemmas tagged with the attribute[simp]
and the givenhᵢ
’s, where thehᵢ
’s are expressions. These expressions may contain underscores, in which case they are replaced by metavariables thatsimp
tries to instantiate. If ahᵢ
is a defined constantf
, then the equational lemmas associated withf
are used. This provides a convenient way to unfoldf
.
simp [*]
simplifies the main goal target using the lemmas tagged with the attribute[simp]
and all hypotheses.
simp *
is a shorthand forsimp [*]
.
simp only [h₁ h₂ ... hₙ]
is likesimp [h₁ h₂ ... hₙ]
but does not use[simp]
lemmas
simp [-id₁, ... -idₙ]
simplifies the main goal target using the lemmas tagged with the attribute[simp]
, but removes the ones namedidᵢ
.
simp at h₁ h₂ ... hₙ
simplifies the non-dependent hypothesesh₁ : T₁
…hₙ : Tₙ
. The tactic fails if the target or another hypothesis depends on one of them. The token⊢
or|-
can be added to the list to include the target.
simp at *
simplifies all the hypotheses and the target.
simp * at *
simplifies target and all (non-dependent propositional) hypotheses using the other hypotheses.
simp with attr₁ ... attrₙ
simplifies the main goal target using the lemmas tagged with any of the attributes[attr₁]
, …,[attrₙ]
or[simp]
.
dsimp only? (* | [(* | (- id | expr)), ...]?) (with id*)? (at (* | (⊢ | id)*))? tactic.dsimp_config?
dsimp
is similar tosimp
, except that it only uses definitional equalities.
simp_intros id* only? (* | [(* | (- id | expr)), ...]?) (with id*)? tactic.simp_intros_config?
simp_intros h₁ h₂ ... hₙ
is similar tointros h₁ h₂ ... hₙ
except that each hypothesis is simplified as it is introduced, and each introduced hypothesis is used to simplify later ones and the final target.As with
simp
, a list of simplification lemmas can be provided. The modifiersonly
andwith
behave as withsimp
.
unfold id* (at (* | (⊢ | id)*))? tactic.unfold_config?
Given defined constants
e₁ ... eₙ
,unfold e₁ ... eₙ
iteratively unfolds all occurrences in the target of the main goal, using equational lemmas associated with the definitions.As with
simp
, theat
modifier can be used to specify locations for the unfolding.
unfold1 id* (at (* | (⊢ | id)*))? tactic.unfold_config?
Similar tounfold
, but does not iterate the unfolding.
dunfold id* (at (* | (⊢ | id)*))? tactic.dunfold_config?
Similar tounfold
, but only uses definitional equalities.
delta id* (at (* | (⊢ | id)*))?
Similar todunfold
, but performs a raw delta reduction, rather than using an equation associated with the defined constants.
unfold_projs
This tactic unfolds all structure projections.
trace_simp_set
Just construct the simp set and trace it. Used for debugging.
ac_reflexivity
Proves a goal with targets = t
whens
andt
are equal up to the associativity and commutativity of their binary operations.
ac_refl
An abbreviation forac_reflexivity
.
cc
Tries to prove the main goal using congruence closure.
6.9. Other Tactics¶
trace_state
This tactic displays the current state in the tracing buffer.
trace a
trace a
displaysa
in the tracing buffer.
type_check expr
Type check the given expression, and trace its type.
apply_opt_param
If the target of the main goal is an opt_param, assigns the default value.
apply_auto_param
If the target of the main goal is an auto_param, executes the associated tactic.
dedup
Renames hypotheses with the same name.