Format the given tactic state. If target_lhs_only
is true and the target
is of the form lhs ~ rhs
, where ~
is a simplification relation,
then only the lhs
is displayed.
Remark: the parameter target_lhs_only
is a temporary hack used to implement
the conv
monad. It will be removed in the future.
Format expression with respect to the main goal in the tactic state. If the tactic state does not contain any goals, then format expression using an empty local context.
tactic
is the monad for building tactics.
You use this to:
- View and modify the local goals and hypotheses in the prover's state.
- Invoke type checking and elaboration of terms.
- View and modify the environment.
- Build new tactics out of existing ones such as
simp
andrewrite
.
- config_type : Type
- inhabited : inhabited (interactive.executor.config_type m)
- execute_with : interactive.executor.config_type m → m unit → tactic unit
Typeclass for custom interaction monads, which provides
the information required to convert an interactive-mode
construction to a tactic
which can actually be executed.
Given a [monad m]
, execute_with
explains how to turn a begin ... end
block, or a by ...
statement into a tactic α
which can actually be
executed. The inhabited
first argument facilitates the passing of an
optional configuration parameter config
, using the syntax:
lean
begin [custom_monad] with config,
...
end
Default executor
instance for tactic
s themselves
try_core t
acts like t
, but succeeds even if t
fails. It returns the
result of t
if t
succeeded and none
otherwise.
try t
acts like t
, but succeeds even if t
fails.
fail_if_success t
acts like t
, but succeeds if t
fails and fails if t
succeeds. Changes made by t
to the tactic_state
are preserved only if t
succeeds.
success_if_fail t
acts like t
, but succeeds if t
fails and fails if t
succeeds. Changes made by t
to the tactic_state
are preserved only if t
succeeds.
iterate t
repeats t
100.000 times or until t
fails, returning the
result of each iteration.
Decorate t's exceptions with msg.
Instances
- has_to_format_to_has_to_tactic_format
- expr.has_to_tactic_format
- list.has_to_tactic_format
- prod.has_to_tactic_format
- option.has_to_tactic_format
- reflected.has_to_tactic_format
- simp_lemmas.has_to_tactic_format
- cc_state.has_to_tactic_format
- tactic.has_to_tactic_format
- tactic.simp_arg_type_to_tactic_format
- hinst_lemma.has_to_tactic_format
- hinst_lemmas.has_to_tactic_format
- module_info.has_to_tactic_format
- array.has_to_tactic_format
- binder.has_to_tactic_format
- declaration.has_to_tactic_format
- native.rb_map.has_to_tactic_format
- buffer.has_to_tactic_format
- tactic.pformat.has_to_tactic_format
- tactic.tactic.has_to_tactic_format
- tactic.explode.has_to_tactic_format
- tactic.rcases_patt.has_to_tactic_format
- tactic.interactive.has_to_tactic_format_mono_function
- tactic.interactive.has_to_tactic_format_mono_ctx
- tactic.interactive.has_to_tactic_format_mono_law
- tactic.abel.has_to_tactic_format
- tactic.ring.has_to_tactic_format
- tactic.closure.has_to_tactic_format
- transfer.has_to_tactic_format_rel_data
- transfer.has_to_tactic_format_rule_data
- all : tactic.transparency
- semireducible : tactic.transparency
- instances : tactic.transparency
- reducible : tactic.transparency
- none : tactic.transparency
A parameter representing how aggressively definitions should be unfolded when trying to decide if two terms match, unify or are definitionally equal. By default, theorem declarations are never unfolded.
all
will unfold everything, including macros and theorems. Except projection macros.semireducible
will unfold everything except theorems and definitions tagged as irreducible.instances
will unfold all class instance definitions and definitions tagged with reducible.reducible
will only unfold definitions tagged with thereducible
attribute.none
will never unfold anything. [NOTE] You are not allowed to tag a definition with more than one ofreducible
,irreducible
,semireducible
attributes. [NOTE] there is a config flagm_unfold_lemmas
that will make it unfold theorems.
(eval_expr α e) evaluates 'e' IF 'e' has type 'α'.
Return the partial term/proof constructed so far. Note that the resultant expression may contain variables that are not declarate in the current main goal.
Display the partial term/proof constructed so far. This tactic is not equivalent to
do { r ← result, s ← read, return (format_expr s r) }
because this one will format the result with respect
to the current goal, and trace_result will do it with respect to the initial goal.
Return target type of the main goal. Fail if tactic_state does not have any goal left.
Clear the given local constant. The tactic fails if the given expression is not a local constant.
revert_lst : list expr → tactic nat
is the reverse of intron
. It takes a local constant c
and puts it back as bound by a pi
or elet
of the main target.
If there are other local constants that depend on c
, these are also reverted. Because of this, the nat
that is returned is the actual number of reverted local constants.
Example: with x : ℕ, h : P(x) ⊢ T(x)
, revert_lst [x]
returns 2
and produces the state ⊢ Π x, P(x) → T(x)
.
Return e
in weak head normal form with respect to the given transparency setting.
If unfold_ginductive
is tt
, then nested and/or mutually recursive inductive datatype constructors
and types are unfolded. Recall that nested and mutually recursive inductive datatype declarations
are compiled into primitive datatypes accepted by the Kernel.
(head) eta expand the given expression. f : α → β
head-eta-expands to λ a, f a
. If f
isn't a function then it just returns f
.
(head) beta reduction. (λ x, B) c
reduces to B[x/c]
.
(head) zeta reduction. Reduction of let bindings at the head of the expression. let x : a := b in c
reduces to c[x/b]
.
Zeta reduction. Reduction of let bindings. let x : a := b in c
reduces to c[x/b]
.
(head) eta reduction. (λ x, f x)
reduces to f
.
Succeeds if t
and s
can be unified using the given transparency setting.
Similar to unify
, but it treats metavariables as constants.
Infer the type of the given expression. Remark: transparency does not affect type inference
Resolve a name using the current local context, environment, aliases, etc.
Return the hypothesis in the main goal. Fail if tactic_state does not have any goal left.
Get a fresh name that is guaranteed to not be in use in the local context.
If n
is provided and n
is not in use, then n
is returned.
Otherwise a number i
is appended to give "n_i"
.
Helper tactic for creating simple applications where some arguments are inferred using type inference.
Example, given
lean
rel.{l_1 l_2} : Pi (α : Type.{l_1}) (β : α -> Type.{l_2}), (Pi x : α, β x) -> (Pi x : α, β x) -> , Prop
nat : Type
real : Type
vec.{l} : Pi (α : Type l) (n : nat), Type.{l1}
f g : Pi (n : nat), vec real n
then
lean
mk_app_core semireducible "rel" [f, g]
returns the application
lean
rel.{1 2} nat (fun n : nat, vec real n) f g
The unification constraints due to type inference are solved using the transparency md
.
Similar to mk_app
, but allows to specify which arguments are explicit/implicit.
Example, given (a b : nat)
then
lean
mk_mapp "ite" [some (a > b), none, none, some a, some b]
returns the application
lean
@ite.{1} (a > b) (nat.decidable_gt a b) nat a b
(mk_eq_refl h) is a more efficient version of (mk_app `eq.refl [h])
(mk_eq_symm h) is a more efficient version of (mk_app `eq.symm [h])
Close the current goal using e
. Fail is the type of e
is not definitionally equal to
the target type.
Elaborate the given quoted expression with respect to the current main goal.
Note that this means that any implicit arguments for the given pexpr
will be applied with fresh metavariables.
If allow_mvars
is tt, then metavariables are tolerated and become new goals if subgoals
is tt.
Return true if the given expression is a type class.
Try to create an instance of the given type class.
Change the target of the main goal.
The input expression must be definitionally equal to the current target.
If check
is ff
, then the tactic does not check whether e
is definitionally equal to the current target. If it is not,
then the error will only be detected by the kernel type checker.
Rotate goals to the left. That is, rotate_left 1
takes the main goal and puts it to the back of the subgoal list.
Gets a list of metavariables, one for each goal.
- non_dep_first : tactic.new_goals
- non_dep_only : tactic.new_goals
- all : tactic.new_goals
How to order the new goals made from an apply
tactic.
Supposing we were applying e : ∀ (a:α) (p : P(a)), Q
non_dep_first
would produce goals⊢ P(?m)
,⊢ α
. It puts the P goal at the front because none of the arguments afterp
ine
depend onp
. It doesn't matter what the resultQ
depends on.non_dep_only
would produce goal⊢ P(?m)
.all
would produce goals⊢ α
,⊢ P(?m)
.
- md : tactic.transparency
- approx : bool
- new_goals : tactic.new_goals
- instances : bool
- auto_param : bool
- opt_param : bool
- unify : bool
Configuration options for the apply
tactic.
md
sets how aggressively definitions are unfolded.new_goals
is the strategy for ordering new goals.instances
iftt
, thenapply
tries to synthesize unresolved[...]
arguments using type class resolution.auto_param
iftt
, thenapply
tries to synthesize unresolved(h : p . tac_id)
arguments using tactictac_id
.opt_param
iftt
, thenapply
tries to synthesize unresolved(a : t := v)
arguments by setting them tov
.unify
iftt
, thenapply
is free to assign existing metavariables in the goal when solving unification constraints. For example, in the goal|- ?x < succ 0
, the tacticapply succ_lt_succ
succeeds with the default configuration, butapply_with succ_lt_succ {unify := ff}
doesn't since it would require Lean to assign?x
tosucc ?y
where?y
is a fresh metavariable.
Apply the expression e
to the main goal, the unification is performed using the transparency mode in cfg
.
Supposing e : Π (a₁:α₁) ... (aₙ:αₙ), P(a₁,...,aₙ)
and the target is Q
, apply
will attempt to unify Q
with P(?a₁,...?aₙ)
.
All of the metavariables that are not assigned are added as new metavariables.
If cfg.approx
is tt
, then fallback to first-order unification, and approximate context during unification.
cfg.new_goals
specifies which unassigned metavariables become new goals, and their order.
If cfg.instances
is tt
, then use type class resolution to instantiate unassigned meta-variables.
The fields cfg.auto_param
and cfg.opt_param
are ignored by this tactic (See tactic.apply
).
It returns a list of all introduced meta variables and the parameter name associated with them, even the assigned ones.
Return the value assigned to the given universe meta-variable. Fail if argument is not an universe meta-variable or if it is not assigned.
Return the value assigned to the given meta-variable. Fail if argument is not a meta-variable or if it is not assigned.
Return true if the given meta-variable is assigned. Fail if argument is not a meta-variable.
Make a name that is guaranteed to be unique. Eg _fresh.1001.4667
. These will be different for each run of the tactic.
Induction on h
using recursor rec
, names for the new hypotheses
are retrieved from ns
. If ns
does not have sufficient names, then use the internal binder names
in the recursor.
It returns for each new goal the name of the constructor (if rec_name
is a builtin recursor),
a list of new hypotheses, and a list of substitutions for hypotheses
depending on h
. The substitutions map internal names to their replacement terms. If the
replacement is again a hypothesis the user name stays the same. The internal names are only valid
in the original goal, not in the type context of the new goal.
Remark: if rec_name
is not a builtin recursor, we use parameter names of rec_name
instead of
constructor names.
If rec
is none, then the type of h
is inferred, if it is of the form C ...
, tactic uses C.rec
Apply cases_on
recursor, names for the new hypotheses are retrieved from ns
.
h
must be a local constant. It returns for each new goal the name of the constructor, a list of new hypotheses, and a list of
substitutions for hypotheses depending on h
. The number of new goals may be smaller than the
number of constructors. Some goals may be discarded when the indices to not match.
See induction
for information on the list of substitutions.
The cases
tactic is implemented using this one, and it relaxes the restriction of h
.
Note: There is one "new hypothesis" for every constructor argument. These are usually local constants, but due to dependent pattern matching, they can also be arbitrary terms.
Similar to cases tactic, but does not revert/intro/clear hypotheses.
Generalizes the target with respect to e
.
instantiate assigned metavariables in the given expression
Add the given declaration to the environment
Changes the environment to the new_env
.
The new environment does not need to be a descendant of the old one.
Use with care.
Changes the environment to the new_env
. new_env
needs to be a descendant from the current environment.
doc_string env d k
returns the doc string for d
(if available)
Create an auxiliary definition with name c
where type
and value
may contain local constants and
meta-variables. This function collects all dependencies (universe parameters, universe metavariables,
local constants (aka hypotheses) and metavariables).
It updates the environment in the tactic_state, and returns an expression of the form
(c.{l_1 ... l_n} a_1 ... a_m)
where l_i's and a_j's are the collected dependencies.
Returns a list of all top-level (/-! ... -/
) docstrings in the active module and imported ones.
The returned object is a list of modules, indexed by (some filename)
for imported modules
and none
for the active one, where each module in the list is paired with a list
of (position_in_file, docstring)
pairs.
Returns a list of docstrings in the active module. An entry in the list can be either:
- a top-level (
/-! ... -/
) docstring, represented as(none, docstring)
- a declaration-specific (
/-- ... -/
) docstring, represented as(some decl_name, docstring)
copy_attribute attr_name c_name p d_name
copy attribute attr_name
from
src
to tgt
if it is defined for src
; make it persistent if p
is tt
;
if p
is none
, the copied attribute is made persistent iff it is persistent on src
Return list of currently open namespaces
Return tt iff t
"occurs" in e
. The occurrence checking is performed using
keyed matching with the given transparency setting.
We say t
occurs in e
by keyed matching iff there is a subterm s
s.t. t
and s
have the same head, and is_def_eq t s md
The main idea is to minimize the number of is_def_eq
checks
performed.
Abstracts all occurrences of the term t
in e
using keyed matching.
If unify
is ff
, then matching is used instead of unification.
That is, metavariables occurring in e
are not assigned.
Blocks the execution of the current thread for at least msecs
milliseconds.
This tactic is used mainly for debugging purposes.
Type check e
with respect to the current goal.
Fails if e
is not type correct.
A tag
is a list of names
. These are attached to goals to help tactics track them.
Equations
Tag goal g
with tag t
. It does nothing if goal tagging is disabled.
Remark: set_goal g []
removes the tag
Return tag associated with g
. Return []
if there is no tag.
By default, Lean only considers local instances in the header of declarations. This has two main benefits. 1- Results produced by the type class resolution procedure can be easily cached. 2- The set of local instances does not have to be recomputed.
This approach has the following disadvantages: 1- Frozen local instances cannot be reverted. 2- Local instances defined inside of a declaration are not considered during type class resolution.
This tactic resets the set of local instances. After executing this tactic,
the set of local instances will be recomputed and the cache will be frequently
reset. Note that, the cache is still used when executing a single tactic that
may generate many type class resolution problems (e.g., simp
).
Freeze the current set of local instances.
Auxiliary definition used to implement begin ... end blocks
Return true iff n is the name of declaration that is a proposition.
Return e
in weak head normal form with respect to the given transparency setting,
or e
head is a generalized constructor or inductive datatype.
Change the target of the main goal.
The input expression must be definitionally equal to the current target.
The tactic does not check whether e
is definitionally equal to the current target. The error will only be detected by the kernel type checker.
Pi or elet introduction.
Given the tactic state ⊢ Π x : α, Y
, intro `hello
will produce the state hello : α ⊢ Y[x/hello]
.
Returns the new local constant. Similarly for elet
expressions.
If the target is not a Pi or elet it will try to put it in WHNF.
A variant of intro
which makes sure that the introduced hypothesis's name is
unique in the context. If there is no hypothesis named n
in the context yet,
intro_fresh n
is the same as intro n
. If there is already a hypothesis named
n
, the new hypothesis is named n_1
(or n_2
if n_1
already exists, etc.).
If offset
is given, the new names are n_offset
, n_offset+1
etc.
If n
is _
, intro_fresh n
is the same as intro1
. The offset
is ignored
in this case.
Like intro
except the name is derived from the bound name in the Π.
Repeatedly apply intro1
and return the list of new local constants in order of introduction.
Introduces new hypotheses with forward dependencies.
Like intron'
but the introduced hypotheses' names are derived from base
,
i.e. base
, base_1
etc. The new names are unique in the context. If offset
is given, the new names will be base_offset
, base_offset+1
etc.
intron_with i ns base offset
introduces i
hypotheses using the names from
ns
. If ns
contains less than i
names, the remaining hypotheses' names are
derived from base
and offset
(as with intron_base
). If base
is _
, the
names are derived from the Π binder names.
Returns the introduced local constants and the remaining names from ns
(if
ns
contains more than i
names).
Returns n fully qualified if it refers to a constant, or else fails.
Example: with x : ℕ, h : P(x) ⊢ T(x)
, revert x
returns 2
and produces the state ⊢ Π x, P(x) → T(x)
.
Swap first two goals, do nothing if tactic state does not have at least two goals.
Add h : t := pr
to the current goal
Add h : t
to the current goal, given a proof pr : t
Rotate the goals to the left by n
. That is, put the main goal to the back n
times.
This tactic is applied to each goal. 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 tactic repeat
never fails.
first [t_1, ..., t_n]
applies the first tactic that doesn't fail.
The tactic fails if all t_i's fail.
Applies the given tactic to the main goal and fails if it is not solved.
solve [t_1, ... t_n]
applies the first tactic that solves the main goal.
Apply the given tactic to all goals. Return one result per goal.
Try to solve the main goal using type class resolution.
Return expr.const c [l_1, ..., l_n]
where l_i's are fresh universe meta-variables.
Apply the constant c
Create a fresh universe ?u
, a metavariable ?T : Type.{?u}
,
and return metavariable ?M : ?T
.
This action can be used to create a meta-variable when
we don't know its type at creation time
Return all hypotheses that depends on e
The dependency test is performed using kdepends_on
with the given transparency setting.
Revert all hypotheses that depend on e
Similar to cases_core
, but e
doesn't need to be a hypothesis.
Remark, it reverts dependencies using revert_kdeps
.
Two different transparency modes are used md
and dmd
.
The mode md
is used with cases_core
and dmd
with generalize
and revert_kdeps
.
It returns the constructor names associated with each new goal and the newly
introduced hypotheses. Note that while cases_core
may return "new
hypotheses" that are not local constants, this tactic only returns local
constants.
Return tt iff 'd' is a declaration in one of the current open namespaces
Execute tac for 'max' "heartbeats". The heartbeat is approx. the maximum number of memory allocations (in thousands) performed by 'tac'. This is a deterministic way of interrupting long running tactics.
add declaration d
as a protected declaration
check if n
is the name of a protected declaration
add_defn_equations
adds a definition specified by a list of equations.
The arguments:
lp
: list of universe parametersparams
: list of parameters (binders before the colon);fn
: a local constant giving the name and type of the declaration (withparams
in the local context);eqns
: a list of equations, each of which is a list of patterns (constructors applied to new local constants) and the branch expression;is_meta
: is the definition meta?
add_defn_equations
can be used as:
do my_add ← mk_local_def `my_add `(ℕ → ℕ),
a ← mk_local_def `a ℕ,
b ← mk_local_def `b ℕ,
add_defn_equations [a] my_add
[ ([``(nat.zero)], a),
([``(nat.succ %%b)], my_add b) ])
ff -- non-meta
to create the following definition:
def my_add (a : ℕ) : ℕ → ℕ
| nat.zero := a
| (nat.succ b) := my_add b
Get the revertible part of the local context. These are the hypotheses that
appear after the last frozen local instance in the local context. We call them
revertible because revert
can revert them, unlike those hypotheses which occur
before a frozen instance.
Rename local hypotheses according to the given name_map
. The name_map
contains as keys those hypotheses that should be renamed; the associated values
are the new names.
This tactic can only rename hypotheses which occur after the last frozen local
instance. If you need to rename earlier hypotheses, try
unfreeze_local_instances
.
If strict
is true, we fail if name_map
refers to hypotheses that do not
appear in the local context or that appear before a frozen local instance.
Conversely, if strict
is false, some entries of name_map
may be silently
ignored.
If use_unique_names
is true, the keys of name_map
should be the unique names
of hypotheses to be renamed. Otherwise, the keys should be display names.
Note that we allow shadowing, so renamed hypotheses may have the same name
as other hypotheses in the context. If use_unique_names
is false and there are
multiple hypotheses with the same display name in the context, they are all
renamed.
"Replace" hypothesis h : type
with h : new_type
where eq_pr
is a proof
that (type = new_type). The tactic actually creates a new hypothesis
with the same user facing name, and (tries to) clear h
.
The clear
step fails if h
has forward dependencies. In this case, the old h
will remain in the local context. The tactic returns the new hypothesis.