Module Glob_term
Untyped intermediate terms
type existential_name
= Names.Id.t
type glob_sort_name
=
|
GSProp
representation of
SProp
literal|
GProp
representation of
Prop
level|
GSet
representation of
Set
level|
GType of Libnames.qualid
representation of a
Type
leveltype 'a glob_sort_expr
=
|
UAnonymous of
{
rigid : bool;
}
not rigid = unifiable by minimization
|
UNamed of 'a
type glob_level
= glob_sort_name glob_sort_expr
levels, occurring in universe instances
type glob_sort
= (glob_sort_name * int) list glob_sort_expr
sort expressions
type glob_constraint
= glob_sort_name * Univ.constraint_type * glob_sort_name
type glob_recarg
= int option
and glob_fix_kind
=
|
GFix of glob_recarg array * int
|
GCoFix of int
type 'a cast_type
=
|
CastConv of 'a
|
CastVM of 'a
|
CastCoerce
Cast to a base type (eg, an underlying inductive type)
|
CastNative of 'a
type 'a cases_pattern_r
=
|
PatVar of Names.Name.t
|
PatCstr of Names.constructor * 'a cases_pattern_g list * Names.Name.t
PatCstr(p,C,l,x)
= "|'C' 'l' as 'x'"The kind of patterns that occurs in "match ... with ... end"
locs here refers to the ident's location, not whole pat
and 'a cases_pattern_g
= ('a cases_pattern_r, 'a) DAst.t
type cases_pattern
= [ `any ] cases_pattern_g
type binding_kind
=
|
Explicit
|
Implicit
type 'a glob_constr_r
=
|
GRef of Names.GlobRef.t * glob_level list option
An identifier that represents a reference to an object defined either in the (global) environment or in the (local) context.
|
GVar of Names.Id.t
An identifier that cannot be regarded as "GRef". Bound variables are typically represented this way.
|
GEvar of existential_name * (Names.Id.t * 'a glob_constr_g) list
|
GPatVar of Evar_kinds.matching_var_kind
Used for patterns only
|
GApp of 'a glob_constr_g * 'a glob_constr_g list
|
GLambda of Names.Name.t * binding_kind * 'a glob_constr_g * 'a glob_constr_g
|
GProd of Names.Name.t * binding_kind * 'a glob_constr_g * 'a glob_constr_g
|
GLetIn of Names.Name.t * 'a glob_constr_g * 'a glob_constr_g option * 'a glob_constr_g
|
GCases of Constr.case_style * 'a glob_constr_g option * 'a tomatch_tuples_g * 'a cases_clauses_g
GCases(style,r,tur,cc)
= "match 'tur' return 'r' with 'cc'" (inMatchStyle
)|
GLetTuple of Names.Name.t list * Names.Name.t * 'a glob_constr_g option * 'a glob_constr_g * 'a glob_constr_g
|
GIf of 'a glob_constr_g * Names.Name.t * 'a glob_constr_g option * 'a glob_constr_g * 'a glob_constr_g
|
GRec of glob_fix_kind * Names.Id.t array * 'a glob_decl_g list array * 'a glob_constr_g array * 'a glob_constr_g array
|
GSort of glob_sort
|
GHole of Evar_kinds.t * Namegen.intro_pattern_naming_expr * Genarg.glob_generic_argument option
|
GCast of 'a glob_constr_g * 'a glob_constr_g cast_type
|
GInt of Uint63.t
|
GFloat of Float64.t
Representation of an internalized (or in other words globalized) term.
and 'a glob_constr_g
= ('a glob_constr_r, 'a) DAst.t
and 'a glob_decl_g
= Names.Name.t * binding_kind * 'a glob_constr_g option * 'a glob_constr_g
and 'a predicate_pattern_g
= Names.Name.t * (Names.inductive * Names.Name.t list) CAst.t option
(na,id)
= "as 'na' in 'id'" where ifid
isSome(l,I,k,args)
.
and 'a tomatch_tuple_g
= 'a glob_constr_g * 'a predicate_pattern_g
and 'a tomatch_tuples_g
= 'a tomatch_tuple_g list
and 'a cases_clause_g
= (Names.Id.t list * 'a cases_pattern_g list * 'a glob_constr_g) CAst.t
(p,il,cl,t)
= "|'cl' => 't'". Precondition: the free variables oft
are members ofil
.
and 'a cases_clauses_g
= 'a cases_clause_g list
type glob_constr
= [ `any ] glob_constr_g
type tomatch_tuple
= [ `any ] tomatch_tuple_g
type tomatch_tuples
= [ `any ] tomatch_tuples_g
type cases_clause
= [ `any ] cases_clause_g
type cases_clauses
= [ `any ] cases_clauses_g
type glob_decl
= [ `any ] glob_decl_g
type predicate_pattern
= [ `any ] predicate_pattern_g
type any_glob_constr
=
|
AnyGlobConstr : 'r glob_constr_g -> any_glob_constr
type 'a disjunctive_cases_clause_g
= (Names.Id.t list * 'a cases_pattern_g list list * 'a glob_constr_g) CAst.t
type 'a disjunctive_cases_clauses_g
= 'a disjunctive_cases_clause_g list
type 'a cases_pattern_disjunction_g
= 'a cases_pattern_g list
type disjunctive_cases_clause
= [ `any ] disjunctive_cases_clause_g
type disjunctive_cases_clauses
= [ `any ] disjunctive_cases_clauses_g
type cases_pattern_disjunction
= [ `any ] cases_pattern_disjunction_g
type 'a extended_glob_local_binder_r
=
|
GLocalAssum of Names.Name.t * binding_kind * 'a glob_constr_g
|
GLocalDef of Names.Name.t * binding_kind * 'a glob_constr_g * 'a glob_constr_g option
|
GLocalPattern of 'a cases_pattern_disjunction_g * Names.Id.t list * Names.Id.t * binding_kind * 'a glob_constr_g
and 'a extended_glob_local_binder_g
= ('a extended_glob_local_binder_r, 'a) DAst.t
type extended_glob_local_binder
= [ `any ] extended_glob_local_binder_g