Ltac2_plugin.Tac2ffi
type valexpr =
| ValInt of int | (* Immediate integers *) |
| ValBlk of Tac2expr.tag * valexpr array | (* Structured blocks *) |
| ValStr of Stdlib.Bytes.t | (* Strings *) |
| ValCls of closure | (* Closures *) |
| ValOpn of Names.KerName.t * valexpr array | (* Open constructors *) |
| ValExt : 'a Tac2dyn.Val.tag * 'a -> valexpr | (* Arbitrary data *) |
| ValUint63 of Uint63.t | (* Primitive integers *) |
| ValFloat of Float64.t | (* Primitive floats *) |
val arity_one : (valexpr -> valexpr Proofview.tactic) arity
module Valexpr : sig ... end
These functions allow to convert back and forth between OCaml and Ltac2 data representation. The to_*
functions raise an anomaly whenever the data has not expected shape.
val of_unit : unit -> valexpr
val to_unit : valexpr -> unit
val unit : unit repr
val of_int : int -> valexpr
val to_int : valexpr -> int
val int : int repr
val of_bool : bool -> valexpr
val to_bool : valexpr -> bool
val bool : bool repr
val of_char : char -> valexpr
val to_char : valexpr -> char
val char : char repr
val of_bytes : Stdlib.Bytes.t -> valexpr
val to_bytes : valexpr -> Stdlib.Bytes.t
val bytes : Stdlib.Bytes.t repr
val of_string : string -> valexpr
WARNING these functions copy
val to_string : valexpr -> string
val string : string repr
val of_exn : Exninfo.iexn -> valexpr
val to_exn : valexpr -> Exninfo.iexn
val exn : Exninfo.iexn repr
val of_ident : Names.Id.t -> valexpr
val to_ident : valexpr -> Names.Id.t
val ident : Names.Id.t repr
val of_pattern : Pattern.constr_pattern -> valexpr
val to_pattern : valexpr -> Pattern.constr_pattern
val pattern : Pattern.constr_pattern repr
val of_constant : Names.Constant.t -> valexpr
val to_constant : valexpr -> Names.Constant.t
val constant : Names.Constant.t repr
val of_reference : Names.GlobRef.t -> valexpr
val to_reference : valexpr -> Names.GlobRef.t
val reference : Names.GlobRef.t repr
val of_ext : 'a Tac2dyn.Val.tag -> 'a -> valexpr
val to_ext : 'a Tac2dyn.Val.tag -> valexpr -> 'a
val repr_ext : 'a Tac2dyn.Val.tag -> 'a repr
val of_open : (Names.KerName.t * valexpr array) -> valexpr
val to_open : valexpr -> Names.KerName.t * valexpr array
val open_ : (Names.KerName.t * valexpr array) repr
val app_fun1 : ('a, 'b) fun1 -> 'a repr -> 'b repr -> 'a -> 'b Proofview.tactic
val val_constr : EConstr.t Tac2dyn.Val.tag
val val_ident : Names.Id.t Tac2dyn.Val.tag
val val_pattern : Pattern.constr_pattern Tac2dyn.Val.tag
val val_preterm : Ltac_pretype.closed_glob_constr Tac2dyn.Val.tag
val val_matching_context : Constr_matching.context Tac2dyn.Val.tag
val val_evar : Evar.t Tac2dyn.Val.tag
val val_pp : Pp.t Tac2dyn.Val.tag
val val_sort : EConstr.ESorts.t Tac2dyn.Val.tag
val val_cast : Constr.cast_kind Tac2dyn.Val.tag
val val_inductive : Names.inductive Tac2dyn.Val.tag
val val_constant : Names.Constant.t Tac2dyn.Val.tag
val val_constructor : Names.constructor Tac2dyn.Val.tag
val val_projection : Names.Projection.t Tac2dyn.Val.tag
val val_case : Constr.case_info Tac2dyn.Val.tag
val val_binder : (Names.Name.t Context.binder_annot * EConstr.types) Tac2dyn.Val.tag
val val_univ : Univ.Level.t Tac2dyn.Val.tag
val val_free : Names.Id.Set.t Tac2dyn.Val.tag
val val_ltac1 : Geninterp.Val.t Tac2dyn.Val.tag
val val_ind_data : (Names.Ind.t * Declarations.mutual_inductive_body) Tac2dyn.Val.tag
val val_exn : Exninfo.iexn Tac2dyn.Val.tag
Toplevel representation of OCaml exceptions. Invariant: no LtacError
should be put into a value with tag val_exn
.
Closures
val apply : closure -> valexpr list -> valexpr Proofview.tactic
Given a closure, apply it to some arguments. Handling of argument mismatches is done automatically, i.e. in case of over or under-application.
val abstract : int -> (valexpr list -> valexpr Proofview.tactic) -> closure
Turn a fixed-arity function into a closure. The inner function is guaranteed to be applied to a list whose size is the integer argument.
Exception
exception LtacError of Names.KerName.t * valexpr array
Ltac2-defined exceptions seen from OCaml side