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 UVars.puniverses
Simply type aliases
type pinductive
= Names.inductive UVars.puniverses
type pconstructor
= Names.constructor UVars.puniverses
type metavariable
= int
Existential variables
type case_style
=
|
LetStyle
|
IfStyle
|
LetPatternStyle
|
MatchStyle
|
RegularStyle
infer 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_pp_info : case_printing;
}
The type of constructions
type t
type constr
= t
types
is the same asconstr
but 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 sincetype
is 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 -> constr
Constructs a de Bruijn index (DB indices begin at 1)
val mkVar : Names.Id.t -> constr
Constructs a Variable
val mkArray : (UVars.Instance.t * constr array * constr * types) -> constr
Constructs an array
val mkMeta : metavariable -> constr
Constructs an patvar named "?n"
val mkEvar : existential -> constr
val mkSort : Sorts.t -> types
Construct a sort
val mkSProp : types
val mkProp : types
val mkSet : types
val mkType : Univ.Universe.t -> types
type cast_kind
=
|
VMcast
|
NATIVEcast
|
DEFAULTcast
This defines the strategy to use for verifiying a Cast
val mkCast : (constr * cast_kind * constr) -> constr
Constructs 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) -> types
Constructs the product
(x:t1)t2
val mkLambda : (Names.Name.t Context.binder_annot * types * constr) -> constr
Constructs the abstraction [x:t1]t2
val mkLetIn : (Names.Name.t Context.binder_annot * constr * types * constr) -> constr
Constructs the product
let x = t1 : t2 in t3
val mkApp : (constr * constr array) -> constr
mkApp (f, [|t1; ...; tN|]
constructs the application (f t1 ... tn)$(f~t_1\dots f_n)$
.
val map_puniverses : ('a -> 'b) -> 'a UVars.puniverses -> 'b UVars.puniverses
val mkConstU : pconstant -> constr
Constructs a Constant.t
val mkProj : (Names.Projection.t * Sorts.relevance * constr) -> constr
Constructs a projection application
val mkIndU : pinductive -> constr
Constructs the ith (co)inductive type of the block named kn
val mkConstructU : pconstructor -> constr
Constructs the jth constructor of the ith (co)inductive type of the block named kn.
val mkConstructUi : (pinductive * int) -> constr
val mkRef : Names.GlobRef.t UVars.puniverses -> constr
Make a constant, inductive, constructor or variable.
module UnsafeMonomorphic : sig ... end
type 'constr pcase_branch
= Names.Name.t Context.binder_annot array * 'constr
Names bound by matching the constructor for this branch.
type 'types pcase_return
= (Names.Name.t Context.binder_annot array * 'types) * Sorts.relevance
Names of the indices + name of self
type 'constr pcase_invert
=
type ('constr, 'types, 'univs) pcase
= case_info * 'univs * 'constr array * 'types pcase_return * 'constr pcase_invert * 'constr * 'constr pcase_branch array
type case_invert
= constr pcase_invert
type case_return
= types pcase_return
type case_branch
= constr pcase_branch
type case
= (constr, types, UVars.Instance.t) pcase
type ('constr, 'types) prec_declaration
= Names.Name.t Context.binder_annot array * 'types array * 'constr array
If
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_declaration
The array of
int
's tells for each component of the array of mutual fixpoints the number of lambdas to skip before finding the recursive argument (e.g., value is 2 in "fix f (x:A) (y:=t) (z:B) (v:=u) (w:I)struct w
"), telling to skip x and z and that w is the recursive argument); The second componentint
tells which component of the block is returned
type ('constr, 'types) pcofixpoint
= int * ('constr, 'types) prec_declaration
The component
int
tells which component of the block of cofixpoint is returned
type cofixpoint
= (constr, types) pcofixpoint
If
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 SList.t
constr list
is an instance matching definitionalnamed_context
in the same order (i.e. last argument first)
type ('constr, 'types, 'sort, 'univs) kind_of_term
=
|
Rel of int
Gallina-variable introduced by
forall
,fun
,let-in
,fix
, orcofix
.|
Var of Names.Id.t
Gallina-variable that was introduced by Vernacular-command that extends the local context of the currently open section (i.e.
Variable
orLet
).|
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 * 'types
Concrete syntax
"forall A:B,C"
is represented asProd (A,B,C)
.|
Lambda of Names.Name.t Context.binder_annot * 'types * 'constr
Concrete syntax
"fun A:B => C"
is represented asLambda (A,B,C)
.|
LetIn of Names.Name.t Context.binder_annot * 'constr * 'types * 'constr
Concrete syntax
"let A:C := B in D"
is represented asLetIn (A,B,C,D)
.|
App of 'constr * 'constr array
Concrete syntax
"(F P1 P2 ... Pn)"
is represented asApp (F, [|P1; P2; ...; Pn|])
. ThemkApp
constructor also enforces the following invariant:F
itself is notApp
- and
[|P1;..;Pn|]
is not empty.
|
Const of Names.Constant.t * 'univs
Gallina-variable that was introduced by Vernacular-command that extends the global environment (i.e.
Parameter
, orAxiom
, orDefinition
, orTheorem
etc.)|
Ind of Names.inductive * 'univs
A name of an inductive type defined by
Variant
,Inductive
orRecord
Vernacular-commands.|
Construct of Names.constructor * 'univs
A constructor of an inductive type defined by
Variant
,Inductive
orRecord
Vernacular-commands.|
Case of case_info * 'univs * 'constr array * 'types pcase_return * 'constr pcase_invert * 'constr * 'constr pcase_branch array
Case (ci,u,params,p,iv,c,brs)
is amatch c return p with brs
expression.c
lives in inductiveci.ci_ind
at universe instanceu
and parametersparams
. If this match has case inversion (ie match on a 1 constructor SProp inductive with proof relevant return type) the indices are iniv
.The names in
p
are the names of the bound indices and inductive value (ie thein
andas
clauses).The names in the
brs
are the names of the variables bound in the respective branch.|
Fix of ('constr, 'types) pfixpoint
|
CoFix of ('constr, 'types) pcofixpoint
|
Proj of Names.Projection.t * Sorts.relevance * 'constr
The relevance is the relevance of the whole term
|
Int of Uint63.t
|
Float of Float64.t
|
Array of 'univs * 'constr array * 'constr * 'types
Array (u,vals,def,t)
is an array ofvals
in typet
with default valuedef
.u
is a universe containingt
.
val kind : constr -> (constr, types, Sorts.t, UVars.Instance.t) kind_of_term
val of_kind : (constr, types, Sorts.t, UVars.Instance.t) kind_of_term -> constr
val kind_nocast_gen : ('v -> ('v, 'v, 'sort, 'univs) kind_of_term) -> 'v -> ('v, 'v, 'sort, 'univs) kind_of_term
val kind_nocast : constr -> (constr, types, Sorts.t, UVars.Instance.t) kind_of_term
val isRel : constr -> bool
Simple case analysis
val isRelN : int -> constr -> bool
val isVar : constr -> bool
val isVarId : Names.Id.t -> constr -> bool
val isRef : constr -> bool
val isRefX : Names.GlobRef.t -> constr -> bool
val isInd : constr -> bool
val isEvar : constr -> bool
val isMeta : constr -> bool
val isEvar_or_Meta : constr -> bool
val isSort : constr -> bool
val isCast : constr -> bool
val isApp : constr -> bool
val isLambda : constr -> bool
val isLetIn : constr -> bool
val isProd : constr -> bool
val isConst : constr -> bool
val isConstruct : constr -> bool
val isFix : constr -> bool
val isCoFix : constr -> bool
val isCase : constr -> bool
val isProj : constr -> bool
val is_Prop : constr -> bool
val is_Set : constr -> bool
val isprop : constr -> bool
val is_Type : constr -> bool
val iskind : constr -> bool
val is_small : Sorts.t -> bool
Term destructors
val destRel : constr -> int
Destructs a de Bruijn index
val destMeta : constr -> metavariable
Destructs an existential variable
val destVar : constr -> Names.Id.t
Destructs a variable
val destSort : constr -> Sorts.t
Destructs a sort.
is_Prop
recognizes the sortProp
, whetherisprop
recognizes bothProp
andSet
.
val destProd : types -> Names.Name.t Context.binder_annot * types * types
Destructs the product
$
(x:t_1)t_2$
val destLambda : constr -> Names.Name.t Context.binder_annot * types * constr
Destructs the abstraction
$
x:t_1
t_2$
val destLetIn : constr -> Names.Name.t Context.binder_annot * constr * types * constr
Destructs the let
$
x:=b:t_1
t_2$
val decompose_app_list : constr -> constr * constr list
Decompose any term as an applicative term; the list of args can be empty
val destConst : constr -> Names.Constant.t UVars.puniverses
Destructs a constant
val destEvar : constr -> existential
Destructs an existential variable
val destInd : constr -> Names.inductive UVars.puniverses
Destructs a (co)inductive type
val destConstruct : constr -> Names.constructor UVars.puniverses
Destructs a constructor
val destCase : constr -> case
Destructs 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| ...|])
whereinfo
is pretty-printing information
val destProj : constr -> Names.Projection.t * Sorts.relevance * constr
Destructs a projection
val destFix : constr -> fixpoint
Destructs 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 -> cofixpoint
val destRef : constr -> Names.GlobRef.t UVars.puniverses
Equality
val equal : constr -> constr -> bool
equal a b
is true ifa
equalsb
modulo alpha, casts, and application grouping
val eq_constr_univs : constr UGraph.check_function
eq_constr_univs u a b
istrue
ifa
equalsb
modulo alpha, casts, application grouping and the universe equalities inu
.
val leq_constr_univs : constr UGraph.check_function
leq_constr_univs u a b
istrue
ifa
is convertible tob
modulo alpha, casts, application grouping and the universe inequalities inu
.
Extension of Context with declarations on constr
type rel_declaration
= (constr, types) Context.Rel.Declaration.pt
type named_declaration
= (constr, types) Context.Named.Declaration.pt
type compacted_declaration
= (constr, types) Context.Compacted.Declaration.pt
type rel_context
= rel_declaration list
type named_context
= named_declaration list
type compacted_context
= compacted_declaration list
Relocation and substitution
val exliftn : Esubst.lift -> constr -> constr
exliftn el c
liftsc
with liftingel
Functionals working on expressions canonically abstracted over a local context (possibly with let-ins)
val map_branches : (constr -> constr) -> case_branch array -> case_branch array
val map_return_predicate : (constr -> constr) -> case_return -> case_return
val map_branches_with_binders : ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> case_branch array -> case_branch array
val map_return_predicate_with_binders : ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> case_return -> case_return
Functionals working on the immediate subterm of a construction
val fold : ('a -> constr -> 'a) -> 'a -> constr -> 'a
val fold_invert : ('a -> 'b -> 'a) -> 'a -> 'b pcase_invert -> 'a
val map : (constr -> constr) -> constr -> constr
val map_invert : ('a -> 'a) -> 'a pcase_invert -> 'a pcase_invert
val fold_map : ('a -> constr -> 'a * constr) -> 'a -> constr -> 'a * constr
val fold_map_invert : ('a -> 'b -> 'a * 'b) -> 'a -> 'b pcase_invert -> 'a * 'b pcase_invert
val iter : (constr -> unit) -> constr -> unit
val iter_invert : ('a -> unit) -> 'a pcase_invert -> unit
val compare_head : (existential -> existential -> bool) -> constr constr_compare_fn -> constr constr_compare_fn
type 'univs instance_compare_fn
= (Names.GlobRef.t * int) option -> 'univs -> 'univs -> bool
Convert 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 : UVars.Instance.t instance_compare_fn -> (Sorts.t -> Sorts.t -> bool) -> (existential -> existential -> bool) -> constr constr_compare_fn -> constr constr_compare_fn
val 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 pexistential -> 'v pexistential -> bool) -> 'v constr_compare_fn -> 'v constr_compare_fn -> 'v constr_compare_fn
val 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 pexistential -> 'v pexistential -> bool) -> 'v constr_compare_fn -> 'v constr_compare_fn
compare_head_gen_with k1 k2 u s f c1 c2
comparesc1
andc2
likecompare_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 : UVars.Instance.t instance_compare_fn -> (Sorts.t -> Sorts.t -> bool) -> (existential -> existential -> bool) -> constr constr_compare_fn -> constr constr_compare_fn -> constr constr_compare_fn
val eq_invert : ('a -> 'a -> bool) -> 'a pcase_invert -> 'a pcase_invert -> bool
Hashconsing
val hash : constr -> int
val case_info_hash : case_info -> int
val hcons : constr -> constr
val debug_print : constr -> Pp.t
val debug_print_fix : ('a -> Pp.t) -> ('a, 'a) pfixpoint -> Pp.t
val mkConst : Names.Constant.t -> constr
val mkInd : Names.inductive -> constr
val mkConstruct : Names.constructor -> constr