Module Constr
This file defines the most important datatype of Coq, namely kernel terms, as well as a handful of generic manipulation functions.
type pconstant= Names.Constant.t Univ.puniversesSimply type aliases
type pinductive= Names.inductive Univ.puniversestype pconstructor= Names.constructor Univ.puniversestype metavariable= intExistential variables
type case_style=|LetStyle|IfStyle|LetPatternStyle|MatchStyle|RegularStyleinfer printing form from number of constructor
Case annotation
type case_printing={ind_tags : bool list;tell whether letin or lambda in the arity of the inductive type
cstr_tags : bool list array;tell whether letin or lambda in the signature of each constructor
style : case_style;}type case_info={ci_ind : Names.inductive;ci_npar : int;ci_cstr_ndecls : int array;ci_cstr_nargs : int array;ci_relevance : Sorts.relevance;ci_pp_info : case_printing;}type ('constr, 'univs) case_invert=
The type of constructions
type ttype constr= ttypesis the same asconstrbut is intended to be used for documentation to indicate that such or such function specifically works with types (i.e. terms of type a sort). (Rem:plurial form sincetypeis a reserved ML keyword)
type types= constr
Functions for dealing with constr terms.
The following functions are intended to simplify and to uniform the manipulation of terms. Some of these functions may be overlapped with previous ones.
Term constructors.
val mkRel : int -> constrConstructs a de Bruijn index (DB indices begin at 1)
val mkVar : Names.Id.t -> constrConstructs a Variable
val mkArray : (Univ.Instance.t * constr array * constr * types) -> constrConstructs an array
val mkMeta : metavariable -> constrConstructs an patvar named "?n"
val mkEvar : existential -> constrval mkSort : Sorts.t -> typesConstruct a sort
val mkSProp : typesval mkProp : typesval mkSet : typesval mkType : Univ.Universe.t -> types
type cast_kind=|VMcast|NATIVEcast|DEFAULTcast|REVERTcastThis defines the strategy to use for verifiying a Cast
val mkCast : (constr * cast_kind * constr) -> constrConstructs the term
t1::t2, i.e. the term t1 casted with the type t2 (that means t2 is declared as the type of t1).
val mkProd : (Names.Name.t Context.binder_annot * types * types) -> typesConstructs the product
(x:t1)t2
val mkLambda : (Names.Name.t Context.binder_annot * types * constr) -> constrConstructs the abstraction [x:t1]t2
val mkLetIn : (Names.Name.t Context.binder_annot * constr * types * constr) -> constrConstructs the product
let x = t1 : t2 in t3
val mkApp : (constr * constr array) -> constrmkApp (f, [|t1; ...; tN|]constructs the application (f t1 ... tn)$(f~t_1\dots f_n)$.
val map_puniverses : ('a -> 'b) -> 'a Univ.puniverses -> 'b Univ.puniversesval mkConst : Names.Constant.t -> constrConstructs a Constant.t
val mkConstU : pconstant -> constrval mkProj : (Names.Projection.t * constr) -> constrConstructs a projection application
val mkInd : Names.inductive -> constrConstructs the ith (co)inductive type of the block named kn
val mkIndU : pinductive -> constrval mkConstruct : Names.constructor -> constrConstructs the jth constructor of the ith (co)inductive type of the block named kn.
val mkConstructU : pconstructor -> constrval mkConstructUi : (pinductive * int) -> constrval mkRef : Names.GlobRef.t Univ.puniverses -> constrMake a constant, inductive, constructor or variable.
val mkCase : (case_info * constr * (constr, Univ.Instance.t) case_invert * constr * constr array) -> constrConstructs a destructor of inductive type.
mkCase ci p c acstand for matchcasxinI argsreturnpwithacpresented as describe inci.pstructure isfun args x -> "return clause"acith element is ith constructor case presented as lambda construct_args (without params). case_term
type ('constr, 'types) prec_declaration= Names.Name.t Context.binder_annot array * 'types array * 'constr arrayIf
recindxs = [|i1,...in|]funnames = [|f1,.....fn|]typarray = [|t1,...tn|]bodies = [|b1,.....bn|]thenmkFix ((recindxs,i), funnames, typarray, bodies)constructs the$i$th function of the block (counting from 0)Fixpoint f1 [ctx1] = b1 with f2 [ctx2] = b2 ... with fn [ctxn] = bn.where the length of the
$j$th context is$ij$.
type ('constr, 'types) pfixpoint= (int array * int) * ('constr, 'types) prec_declarationtype ('constr, 'types) pcofixpoint= int * ('constr, 'types) prec_declarationtype rec_declaration= (constr, types) prec_declarationtype fixpoint= (constr, types) pfixpoint
type cofixpoint= (constr, types) pcofixpointIf
funnames = [|f1,.....fn|]typarray = [|t1,...tn|]bodies = [b1,.....bn]thenmkCoFix (i, (funnames, typarray, bodies))constructs the ith function of the blockCoFixpoint f1 = b1 with f2 = b2 ... with fn = bn.
val mkCoFix : cofixpoint -> constr
Concrete type for making pattern-matching.
type 'constr pexistential= Evar.t * 'constr listconstr listis an instance matching definitionalnamed_contextin the same order (i.e. last argument first)
type ('constr, 'types, 'sort, 'univs) kind_of_term=|Rel of intGallina-variable introduced by
forall,fun,let-in,fix, orcofix.|Var of Names.Id.tGallina-variable that was introduced by Vernacular-command that extends the local context of the currently open section (i.e.
VariableorLet).|Meta of metavariable|Evar of 'constr pexistential|Sort of 'sort|Cast of 'constr * cast_kind * 'types|Prod of Names.Name.t Context.binder_annot * 'types * 'typesConcrete syntax
"forall A:B,C"is represented asProd (A,B,C).|Lambda of Names.Name.t Context.binder_annot * 'types * 'constrConcrete syntax
"fun A:B => C"is represented asLambda (A,B,C).|LetIn of Names.Name.t Context.binder_annot * 'constr * 'types * 'constrConcrete syntax
"let A:C := B in D"is represented asLetIn (A,B,C,D).|App of 'constr * 'constr arrayConcrete syntax
"(F P1 P2 ... Pn)"is represented asApp (F, [|P1; P2; ...; Pn|]).The
mkAppconstructor also enforces the following invariant:Fitself is notApp- and
[|P1;..;Pn|]is not empty.
|Const of Names.Constant.t * 'univsGallina-variable that was introduced by Vernacular-command that extends the global environment (i.e.
Parameter, orAxiom, orDefinition, orTheoremetc.)|Ind of Names.inductive * 'univsA name of an inductive type defined by
Variant,InductiveorRecordVernacular-commands.|Construct of Names.constructor * 'univsA constructor of an inductive type defined by
Variant,InductiveorRecordVernacular-commands.|Case of case_info * 'constr * ('constr, 'univs) case_invert * 'constr * 'constr array|Fix of ('constr, 'types) pfixpoint|CoFix of ('constr, 'types) pcofixpoint|Proj of Names.Projection.t * 'constr|Int of Uint63.t|Float of Float64.t|Array of 'univs * 'constr array * 'constr * 'types
val kind : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_termval of_kind : (constr, types, Sorts.t, Univ.Instance.t) kind_of_term -> constrval kind_nocast_gen : ('v -> ('v, 'v, 'sort, 'univs) kind_of_term) -> 'v -> ('v, 'v, 'sort, 'univs) kind_of_termval kind_nocast : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_termval isRel : constr -> boolSimple case analysis
val isRelN : int -> constr -> boolval isVar : constr -> boolval isVarId : Names.Id.t -> constr -> boolval isRef : constr -> boolval isRefX : Names.GlobRef.t -> constr -> boolval isInd : constr -> boolval isEvar : constr -> boolval isMeta : constr -> boolval isEvar_or_Meta : constr -> boolval isSort : constr -> boolval isCast : constr -> boolval isApp : constr -> boolval isLambda : constr -> boolval isLetIn : constr -> boolval isProd : constr -> boolval isConst : constr -> boolval isConstruct : constr -> boolval isFix : constr -> boolval isCoFix : constr -> boolval isCase : constr -> boolval isProj : constr -> boolval is_Prop : constr -> boolval is_Set : constr -> boolval isprop : constr -> boolval is_Type : constr -> boolval iskind : constr -> boolval is_small : Sorts.t -> bool
Term destructors
val destRel : constr -> intDestructs a de Bruijn index
val destMeta : constr -> metavariableDestructs an existential variable
val destVar : constr -> Names.Id.tDestructs a variable
val destSort : constr -> Sorts.tDestructs a sort.
is_Proprecognizes the sortProp, whetherisproprecognizes bothPropandSet.
val destProd : types -> Names.Name.t Context.binder_annot * types * typesDestructs the product
$(x:t_1)t_2$
val destLambda : constr -> Names.Name.t Context.binder_annot * types * constrDestructs the abstraction
$x:t_1t_2$
val destLetIn : constr -> Names.Name.t Context.binder_annot * constr * types * constrDestructs the let
$x:=b:t_1t_2$
val decompose_app : constr -> constr * constr listDecompose any term as an applicative term; the list of args can be empty
val destConst : constr -> Names.Constant.t Univ.puniversesDestructs a constant
val destEvar : constr -> existentialDestructs an existential variable
val destInd : constr -> Names.inductive Univ.puniversesDestructs a (co)inductive type
val destConstruct : constr -> Names.constructor Univ.puniversesDestructs a constructor
val destCase : constr -> case_info * constr * (constr, Univ.Instance.t) case_invert * constr * constr arrayDestructs a
match c as x in I args return P with ... | Ci(...yij...) => ti | ... end(orlet (..y1i..) := c as x in I args return P in t1, orif c then t1 else t2)- returns
(info,c,fun args x => P,[|...|fun yij => ti| ...|])whereinfois pretty-printing information
val destProj : constr -> Names.Projection.t * constrDestructs a projection
val destFix : constr -> fixpointDestructs the
$i$th function of the blockFixpoint f{_ 1} ctx{_ 1} = b{_ 1} with f{_ 2} ctx{_ 2} = b{_ 2} ... with f{_ n} ctx{_ n} = b{_ n}, where the length of the$j$th context is$ij$.
val destCoFix : constr -> cofixpointval destRef : constr -> Names.GlobRef.t Univ.puniverses
Equality
val equal : constr -> constr -> boolequal a bis true ifaequalsbmodulo alpha, casts, and application grouping
val eq_constr_univs : constr UGraph.check_functioneq_constr_univs u a bistrueifaequalsbmodulo alpha, casts, application grouping and the universe equalities inu.
val leq_constr_univs : constr UGraph.check_functionleq_constr_univs u a bistrueifais convertible tobmodulo alpha, casts, application grouping and the universe inequalities inu.
val eq_constr_univs_infer : UGraph.t -> constr -> constr -> bool Univ.constrainedeq_constr_univs u a bistrueifaequalsbmodulo alpha, casts, application grouping and the universe equalities inu.
val leq_constr_univs_infer : UGraph.t -> constr -> constr -> bool Univ.constrainedleq_constr_univs u a bistrueifais convertible tobmodulo alpha, casts, application grouping and the universe inequalities inu.
Extension of Context with declarations on constr
type rel_declaration= (constr, types) Context.Rel.Declaration.pttype named_declaration= (constr, types) Context.Named.Declaration.pttype compacted_declaration= (constr, types) Context.Compacted.Declaration.pttype rel_context= rel_declaration listtype named_context= named_declaration listtype compacted_context= compacted_declaration list
Relocation and substitution
val exliftn : Esubst.lift -> constr -> constrexliftn el cliftscwith liftingel
Functionals working on expressions canonically abstracted over a local context (possibly with let-ins)
val map_under_context_with_binders : ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> int -> constr -> constr
val map_branches_with_binders : ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> case_info -> constr array -> constr array
val map_return_predicate_with_binders : ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> case_info -> constr -> constr
val map_under_context_with_full_binders : ((constr, constr) Context.Rel.Declaration.pt -> 'a -> 'a) -> ('a -> constr -> constr) -> 'a -> int -> constr -> constr
Functionals working on the immediate subterm of a construction
val fold : ('a -> constr -> 'a) -> 'a -> constr -> 'aval fold_with_full_binders : (rel_declaration -> 'a -> 'a) -> ('a -> 'b -> constr -> 'b) -> 'a -> 'b -> constr -> 'bval fold_invert : ('a -> 'b -> 'a) -> 'a -> ('b, 'c) case_invert -> 'a
val map : (constr -> constr) -> constr -> constrval map_invert : ('a -> 'a) -> ('a, 'b) case_invert -> ('a, 'b) case_invert
val fold_map : ('a -> constr -> 'a * constr) -> 'a -> constr -> 'a * constrval fold_map_invert : ('a -> 'b -> 'a * 'b) -> 'a -> ('b, 'c) case_invert -> 'a * ('b, 'c) case_invert
val iter : (constr -> unit) -> constr -> unitval iter_invert : ('a -> unit) -> ('a, 'b) case_invert -> unit
val compare_head : constr constr_compare_fn -> constr constr_compare_fn
type 'univs instance_compare_fn= (Names.GlobRef.t * int) option -> 'univs -> 'univs -> boolConvert a global reference applied to 2 instances. The int says how many arguments are given (as we can only use cumulativity for fully applied inductives/constructors) .
val compare_head_gen : Univ.Instance.t instance_compare_fn -> (Sorts.t -> Sorts.t -> bool) -> constr constr_compare_fn -> constr constr_compare_fnval compare_head_gen_leq_with : ('v -> ('v, 'v, 'sort, 'univs) kind_of_term) -> ('v -> ('v, 'v, 'sort, 'univs) kind_of_term) -> 'univs instance_compare_fn -> ('sort -> 'sort -> bool) -> 'v constr_compare_fn -> 'v constr_compare_fn -> 'v constr_compare_fnval compare_head_gen_with : ('v -> ('v, 'v, 'sort, 'univs) kind_of_term) -> ('v -> ('v, 'v, 'sort, 'univs) kind_of_term) -> 'univs instance_compare_fn -> ('sort -> 'sort -> bool) -> 'v constr_compare_fn -> 'v constr_compare_fncompare_head_gen_with k1 k2 u s f c1 c2comparesc1andc2likecompare_head_gen u s f c1 c2, except thatk1(resp.k2) is used,rather thankind, to expose the immediate subterms ofc1(resp.c2).
val compare_head_gen_leq : Univ.Instance.t instance_compare_fn -> (Sorts.t -> Sorts.t -> bool) -> constr constr_compare_fn -> constr constr_compare_fn -> constr constr_compare_fnval eq_invert : ('a -> 'a -> bool) -> ('b -> 'b -> bool) -> ('a, 'b) case_invert -> ('a, 'b) case_invert -> bool