Ltac2_plugin.Tac2intern
type context = (Names.Id.t * Tac2expr.type_scheme) list
val intern :
strict:bool ->
context ->
Tac2expr.raw_tacexpr ->
Tac2expr.glb_tacexpr * Tac2expr.type_scheme
val intern_typedef :
(Names.KerName.t * int) Names.Id.Map.t ->
Tac2expr.raw_quant_typedef ->
Tac2expr.glb_quant_typedef
val intern_open_type : Tac2expr.raw_typexpr -> Tac2expr.type_scheme
val is_value : Tac2expr.glb_tacexpr -> bool
Check that a term is a value. Only values are safe to marshall between processes.
val is_pure_constructor : Tac2expr.type_constant -> bool
val check_unit : ?loc:Loc.t -> Tac2expr.type_scheme -> unit
val check_subtype : Tac2expr.type_scheme -> Tac2expr.type_scheme -> bool
check_subtype t1 t2
returns true
iff all values of instances of type t1
also have type t2
.
val subst_type :
Mod_subst.substitution ->
'a Tac2expr.glb_typexpr ->
'a Tac2expr.glb_typexpr
val subst_expr :
Mod_subst.substitution ->
Tac2expr.glb_tacexpr ->
Tac2expr.glb_tacexpr
val subst_quant_typedef :
Mod_subst.substitution ->
Tac2expr.glb_quant_typedef ->
Tac2expr.glb_quant_typedef
val subst_type_scheme :
Mod_subst.substitution ->
Tac2expr.type_scheme ->
Tac2expr.type_scheme
val subst_rawexpr :
Mod_subst.substitution ->
Tac2expr.raw_tacexpr ->
Tac2expr.raw_tacexpr
val globalize : Names.Id.Set.t -> Tac2expr.raw_tacexpr -> Tac2expr.raw_tacexpr
Replaces all qualified identifiers by their corresponding kernel name. The set represents bound variables in the context.
Errors
val error_nargs_mismatch :
?loc:Loc.t ->
Tac2expr.ltac_constructor ->
int ->
int ->
'a
val error_nparams_mismatch : ?loc:Loc.t -> int -> int -> 'a
Misc
val drop_ltac2_env : Genintern.Store.t -> Genintern.Store.t