This file defines the most important datatype of Coq, namely kernel terms, as well as a handful of generic manipulation functions.
type case_style =
| LetStyle
| IfStyle
| LetPatternStyle
| MatchStyle
| RegularStyle
infer printing form from number of constructor
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 'constr pcase_invert =
| NoInvert
Normal reduction: match when the scrutinee is a constructor.
| CaseInvert of {
indices : 'constr array;
}
Reduce when the indices match those of the unique constructor. (SProp to non SProp only)
The type of constructions
types
is the same as constr
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 since type
is a reserved ML keyword)
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.
Constructs a de Bruijn index (DB indices begin at 1)
Constructs a machine integer
Constructs a machine float number
Constructs an patvar named "?n"
Constructs an existential variable
type cast_kind =
| VMcast
| NATIVEcast
| DEFAULTcast
| REVERTcast
This defines the strategy to use for verifiying a Cast
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).
Constructs the product (x:t1)t2
Constructs the abstraction [x:t1]t2
Constructs the product let x = t1 : t2 in t3
mkApp (f, [|t1; ...; tN|]
constructs the application (f t1 ... tn) .
Constructs a projection application
Inductive types
Constructs the ith (co)inductive type of the block named kn
Constructs the jth constructor of the ith (co)inductive type of the block named kn.
Make a constant, inductive, constructor or variable.
Constructs a destructor of inductive type.
mkCase ci params p c ac
stand for match c
as x
in I args
return p
with ac
presented as describe in ci
.
p
structure is args x |- "return clause"
ac
ith element is ith constructor case presented as construct_args |- case_term
Names of the indices + name of self
If recindxs = [|i1,...in|]
funnames = [|f1,.....fn|]
typarray = [|t1,...tn|]
bodies = [|b1,.....bn|]
then mkFix ((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
If funnames = [|f1,.....fn|]
typarray = [|t1,...tn|]
bodies = [b1,.....bn]
then mkCoFix (i, (funnames, typarray, bodies))
constructs the ith function of the block
CoFixpoint f1 = b1
with f2 = b2
...
with fn = bn.
Concrete type for making pattern-matching.
type 'constr pexistential = Evar.t * 'constr list
constr list
is an instance matching definitional named_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
, or cofix
.
| 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
or Let
).
| 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 as Prod (A,B,C)
.
| Lambda of Names.Name.t Context.binder_annot * 'types * 'constr
Concrete syntax "fun A:B => C"
is represented as Lambda (A,B,C)
.
| LetIn of Names.Name.t Context.binder_annot * 'constr * 'types * 'constr
Concrete syntax "let A:C := B in D"
is represented as LetIn (A,B,C,D)
.
| App of 'constr * 'constr array
Concrete syntax "(F P1 P2 ... Pn)"
is represented as App (F, [|P1; P2; ...; Pn|])
.
The mkApp
constructor also enforces the following invariant:
F
itself is not App
- 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
, or Axiom
, or Definition
, or Theorem
etc.)
| Ind of Names.inductive * 'univs
A name of an inductive type defined by Variant
, Inductive
or Record
Vernacular-commands.
| Construct of Names.constructor * 'univs
A constructor of an inductive type defined by Variant
, Inductive
or Record
Vernacular-commands.
| Case of case_info
* 'univs
* 'constr array
* 'types pcase_return
* 'constr pcase_invert
* 'constr
* 'constr pcase_branch 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
User view of constr
. For App
, it is ensured there is at least one argument and the function is not itself an applicative term
val isRelN : int -> constr -> bool
val isConstruct : constr -> bool
Term destructors
Destructor operations are partial functions and
Destructs a de Bruijn index
Destructs an existential variable
Destructs a sort. is_Prop
recognizes the sort Prop
, whether isprop
recognizes both Prop
and Set
.
Destructs the product $
(x:t_1)t_2 $
Destructs the abstraction $
x:t_1
t_2 $
Destructs the let $
x:=b:t_1
t_2 $
Decompose any term as an applicative term; the list of args can be empty
Same as decompose_app
, but returns an array.
Destructs an existential variable
Destructs a (co)inductive type
Destructs a match c as x in I args return P with ... |
Ci(...yij...) => ti | ... end
(or let (..y1i..) := c as x in I args
return P in t1
, or if c then t1 else t2
)
Destructs the $
i $
th function of the block Fixpoint 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 $
.
Equality
equal a b
is true if a
equals b
modulo alpha, casts, and application grouping
eq_constr_univs u a b
is true
if a
equals b
modulo alpha, casts, application grouping and the universe equalities in u
.
leq_constr_univs u a b
is true
if a
is convertible to b
modulo alpha, casts, application grouping and the universe inequalities in u
.
eq_constr_univs u a b
is true
if a
equals b
modulo alpha, casts, application grouping and the universe equalities in u
.
leq_constr_univs u a b
is true
if a
is convertible to b
modulo alpha, casts, application grouping and the universe inequalities in u
.
eq_constr_univs a b
true, c
if a
equals b
modulo alpha, casts, application grouping and ignoring universe instances.
Total ordering compatible with equal
Extension of Context with declarations on constr
Relocation and substitution
exliftn el c
lifts c
with lifting el
liftn n k c
lifts by n
indexes above or equal to k
in c
lift n c
lifts by n
the positive indexes in c
Functionals working on expressions canonically abstracted over a local context (possibly with let-ins)
map_branches f br
maps f
on the immediate subterms of an array of "match" branches br
in canonical eta-let-expanded form; it is not recursive and the order with which subterms are processed is not specified; it preserves sharing; the immediate subterms are the types and possibly terms occurring in the context of each branch as well as the body of each branch
map_return_predicate f p
maps f
on the immediate subterms of a return predicate of a "match" in canonical eta-let-expanded form; it is not recursive and the order with which subterms are processed is not specified; it preserves sharing; the immediate subterms are the types and possibly terms occurring in the context of each branch as well as the body of the predicate
map_branches_with_binders f br
maps f
on the immediate subterms of an array of "match" branches br
in canonical eta-let-expanded form; it carries an extra data n
(typically a lift index) which is processed by g
(which typically adds 1 to n
) at each binder traversal; it is not recursive and the order with which subterms are processed is not specified; it preserves sharing; the immediate subterms are the types and possibly terms occurring in the context of the branch as well as the body of the branch
map_return_predicate_with_binders f p
maps f
on the immediate subterms of a return predicate of a "match" in canonical eta-let-expanded form; it carries an extra data n
(typically a lift index) which is processed by g
(which typically adds 1 to n
) at each binder traversal; it is not recursive and the order with which subterms are processed is not specified; it preserves sharing; the immediate subterms are the types and possibly terms occurring in the context of each branch as well as the body of the predicate
fold f acc c
folds f
on the immediate subterms of c
starting from acc
and proceeding from left to right according to the usual representation of the constructions; it is not recursive
val fold_invert : ('a -> 'b -> 'a) -> 'a -> 'b pcase_invert -> 'a
map f c
maps f
on the immediate subterms of c
; it is not recursive and the order with which subterms are processed is not specified
Like map
, but also has an additional accumulator.
map_with_binders g f n c
maps f n
on the immediate subterms of c
; it carries an extra data n
(typically a lift index) which is processed by g
(which typically add 1 to n
) at each binder traversal; it is not recursive and the order with which subterms are processed is not specified
iter f c
iters f
on the immediate subterms of c
; it is not recursive and the order with which subterms are processed is not specified
iter_with_binders g f n c
iters f n
on the immediate subterms of c
; it carries an extra data n
(typically a lift index) which is processed by g
(which typically add 1 to n
) at each binder traversal; it is not recursive and the order with which subterms are processed is not specified
val iter_with_binders :
('a -> 'a) ->
('a -> constr -> unit) ->
'a ->
constr ->
unit
iter_with_binders g f n c
iters f n
on the immediate subterms of c
; it carries an extra data n
(typically a lift index) which is processed by g
(which typically add 1 to n
) at each binder traversal; it is not recursive and the order with which subterms are processed is not specified
val fold_constr_with_binders :
('a -> 'a) ->
('a -> 'b -> constr -> 'b) ->
'a ->
'b ->
constr ->
'b
type 'constr constr_compare_fn = int -> 'constr -> 'constr -> bool
compare_head f c1 c2
compare c1
and c2
using f
to compare the immediate subterms of c1
of c2
if needed; Cast's, binders name and Cases annotations are not taken into account
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) .
compare_head_gen u s f c1 c2
compare c1
and c2
using f
to compare the immediate subterms of c1
of c2
if needed, u
to compare universe instances, s
to compare sorts; Cast's, binders name and Cases annotations are not taken into account
compare_head_gen_with k1 k2 u s f c1 c2
compares c1
and c2
like compare_head_gen u s f c1 c2
, except that k1
(resp. k2
) is used,rather than kind
, to expose the immediate subterms of c1
(resp. c2
).
compare_head_gen_leq u s f fle c1 c2
compare c1
and c2
using f
to compare the immediate subterms of c1
of c2
for conversion, fle
for cumulativity, u
to compare universe instances (the first boolean tells if they belong to a Constant.t), s
to compare sorts for for subtyping; Cast's, binders name and Cases annotations are not taken into account
Hashconsing