Various tactics related to local definitions (local constants of the form x : α := t
) #
We call t
the value of x
.
Just like split
, fsplit
applies the constructor when the type of the target is
an inductive data type with one constructor.
However it does not reorder goals or invoke auto_param
tactics.
Calls injection
on each hypothesis, and then, for each hypothesis on which injection
succeeds, clears the old hypothesis.
run_parser p
is like run_cmd
but for the parser monad. It executes parser p
at the
top level, giving access to operations like emit_code_here
.
Hole command used to fill in a structure's field when specifying an instance.
In the following:
invoking the hole command "Instance Stub" ("Generate a skeleton for the structure under construction.") produces:
Hole command used to generate a match
expression.
In the following:
invoking hole command "Match Stub" ("Generate a list of equations for a match
expression")
produces:
meta def foo (e : expr) : tactic unit :=
match e with
| (expr.var a) := _
| (expr.sort a) := _
| (expr.const a a_1) := _
| (expr.mvar a a_1 a_2) := _
| (expr.local_const a a_1 a_2 a_3) := _
| (expr.app a a_1) := _
| (expr.lam a a_1 a_2 a_3) := _
| (expr.pi a a_1 a_2 a_3) := _
| (expr.elet a a_1 a_2 a_3) := _
| (expr.macro a a_1) := _
end
Invoking hole command "Equations Stub" ("Generate a list of equations for a recursive definition") in the following:
produces:
meta def foo : expr → tactic unit
| (expr.var a) := _
| (expr.sort a) := _
| (expr.const a a_1) := _
| (expr.mvar a a_1 a_2) := _
| (expr.local_const a a_1 a_2 a_3) := _
| (expr.app a a_1) := _
| (expr.lam a a_1 a_2 a_3) := _
| (expr.pi a a_1 a_2 a_3) := _
| (expr.elet a a_1 a_2 a_3) := _
| (expr.macro a a_1) := _
A similar result can be obtained by invoking "Equations Stub" on the following:
meta def foo : expr → tactic unit := -- don't forget to erase `:=`!!
| (expr.var a) := _
| (expr.sort a) := _
| (expr.const a a_1) := _
| (expr.mvar a a_1 a_2) := _
| (expr.local_const a a_1 a_2 a_3) := _
| (expr.app a a_1) := _
| (expr.lam a a_1 a_2 a_3) := _
| (expr.pi a a_1 a_2 a_3) := _
| (expr.elet a a_1 a_2 a_3) := _
| (expr.macro a a_1) := _
This command lists the constructors that can be used to satisfy the expected type.
Invoking "List Constructors" ("Show the list of constructors of the expected type") in the following hole:
def foo : ℤ ⊕ ℕ :=
{! !}
produces:
and will display:
A user attribute that applies to lemmas of the shape ∀ x, f (g x) = h x
.
It derives an auxiliary lemma of the form f ∘ g = h
for reasoning about higher-order functions.
Copies a definition into the tactic.interactive
namespace to make it usable
in proof scripts. It allows one to write
@[interactive]
meta def my_tactic := ...
instead of
meta def my_tactic := ...
run_cmd add_interactive [``my_tactic]
```
setup_tactic_parser
is a user command that opens the namespaces used in writing
interactive tactics, and declares the local postfix notation ?
for optional
and *
for many
.
It does not use the namespace
command, so it will typically be used after
namespace tactic.interactive
.
import_private foo from bar
finds a private declaration foo
in the same file as bar
and creates a local notation to refer to it.
import_private foo
looks for foo
in all imported files.
When possible, make foo
non-private rather than using this feature.