Module Ltac2_plugin.Tac2externals
type (_, _) spec
Type
('v,'f) spec
represents a high-level Ltac2 tactic specification. It indicates how to turn a value of type'f
into an Ltac2 tactic, which may involve converting between OCaml and Ltac2 value representations, and also lifting a pure function to the tactic monad. The type parameter'v
gives the type of value produced by interpreting the specification.
val tac : 'r Tac2ffi.repr -> (Tac2ffi.valexpr Proofview.tactic, 'r Proofview.tactic) spec
tac r
is the specification of a tactic (in the tactic monad sense) whose return type is specified (and converted into an Ltac2 value) viar
.
val tac' : ('r -> Tac2ffi.valexpr) -> (Tac2ffi.valexpr Proofview.tactic, 'r Proofview.tactic) spec
tac'
is similar totac
, but only needs a conversion function.
val ret : 'r Tac2ffi.repr -> (Tac2ffi.valexpr Proofview.tactic, 'r) spec
ret r
is the specification of a pure tactic (i.e., a tactic defined as a pure OCaml value, not needing the tactic monad) whose return type is given byr
(seetac
).
val ret' : ('r -> Tac2ffi.valexpr) -> (Tac2ffi.valexpr Proofview.tactic, 'r) spec
ret'
is similar toret
, but only needs a conversion function.
val eret : 'r Tac2ffi.repr -> (Tac2ffi.valexpr Proofview.tactic, Environ.env -> Evd.evar_map -> 'r) spec
eret
is similar toret
, but for tactics that can be implemented with a pure OCaml value, provided extra argumentsenv
andsigma
, computed viatclENV
andtclEVARMAP
.
val eret' : ('r -> Tac2ffi.valexpr) -> (Tac2ffi.valexpr Proofview.tactic, Environ.env -> Evd.evar_map -> 'r) spec
eret'
is similar toeret
, but only needs a conversion function.
val gret : 'r Tac2ffi.repr -> (Tac2ffi.valexpr Proofview.tactic, Proofview.Goal.t -> 'r) spec
gret
is similar toret
, but for tactics that can be implemented with a pure OCaml value, provided the current goalg
as extra argument. A fatal error is produced when there is not exactly one goal in focus at the point of using an Ltac2 value defined with this specification. Indeed, the value ofg
is computed usingGoal.enter_one
.
val gret' : ('r -> Tac2ffi.valexpr) -> (Tac2ffi.valexpr Proofview.tactic, Proofview.Goal.t -> 'r) spec
gret'
is similar togret
, but only needs a conversion function.
val (@->) : 'a Tac2ffi.repr -> ('t, 'f) spec -> (Tac2ffi.valexpr -> 't, 'a -> 'f) spec
r @-> s
extends the specifications
with a closure argument whose type is specified by (and converted from an Ltac2 value via)r
.
val (@-->) : (Tac2ffi.valexpr -> 'a) -> ('t, 'f) spec -> (Tac2ffi.valexpr -> 't, 'a -> 'f) spec
(@-->)
is similar to(@->)
, but only needs a conversion function.
val define : Tac2expr.ml_tactic_name -> (_, 'f) spec -> 'f -> unit
define tac s f
defines an external tactictac
by interpretingf
with the specifications
. TheInvalid_argument
exception is raised when the givens
does not produce a "pure" tactic, that is, something that can be turned into an Ltac2 value (i.e., a closure, or a pure value).