ConstrexprSourcetype sort_name_expr = | CSProp| CProp| CSet| CType of Libnames.qualid| CRawType of Univ.Level.tUniverses like "foo.1" have no qualid form
*)Universes
Constraints don't have anonymous universes
type universe_decl_expr =
(Names.lident list, univ_constraint_expr list) UState.gen_universe_decltype cumul_univ_decl_expr =
((Names.lident * Univ.Variance.t option) list, univ_constraint_expr list)
UState.gen_universe_decltype notation_entry_level = | InConstrEntrySomeLevel| InCustomEntryLevel of string * entry_leveltype binder_kind = | Default of Glob_term.binding_kind| Generalized of Glob_term.binding_kind * bool(Inner binding always Implicit) Outer bindings, typeclass-specific flag for implicit generalization of superclasses
*)Some n = proj of the n-th visible argument
type cases_pattern_expr_r = | CPatAlias of cases_pattern_expr * Names.lname| CPatCstr of Libnames.qualid
* cases_pattern_expr list option
* cases_pattern_expr listCPatCstr (_, c, Some l1, l2) represents (@ c l1) l2
| CPatAtom of Libnames.qualid option| CPatOr of cases_pattern_expr list| CPatNotation of notation_with_optional_scope option
* notation
* cases_pattern_notation_substitution
* cases_pattern_expr listCPatNotation (_, n, l1 ,l2) represents (notation n applied with substitution l1) applied to arguments l2
*)| CPatPrim of prim_token| CPatRecord of (Libnames.qualid * cases_pattern_expr) list| CPatDelimiters of string * cases_pattern_expr| CPatCast of cases_pattern_expr * constr_exprconstr_expr is the abstract syntax tree produced by the parser
and cases_pattern_notation_substitution =
cases_pattern_expr list * cases_pattern_expr list listand constr_expr_r = | CRef of Libnames.qualid * instance_expr option| CFix of Names.lident * fix_expr list| CCoFix of Names.lident * cofix_expr list| CProdN of local_binder_expr list * constr_expr| CLambdaN of local_binder_expr list * constr_expr| CLetIn of Names.lname * constr_expr * constr_expr option * constr_expr| CAppExpl of proj_flag * Libnames.qualid * instance_expr option
* constr_expr list| CApp of proj_flag * constr_expr
* (constr_expr * explicitation CAst.t option) list| CRecord of (Libnames.qualid * constr_expr) list| CCases of Constr.case_style
* constr_expr option
* case_expr list
* branch_expr list| CLetTuple of Names.lname list
* Names.lname option * constr_expr option
* constr_expr
* constr_expr| CIf of constr_expr
* Names.lname option * constr_expr option
* constr_expr
* constr_expr| CHole of Evar_kinds.t option
* Namegen.intro_pattern_naming_expr
* Genarg.raw_generic_argument option| CPatVar of Pattern.patvar| CEvar of Glob_term.existential_name CAst.t * (Names.lident * constr_expr) list| CSort of sort_expr| CCast of constr_expr * constr_expr Glob_term.cast_type| CNotation of notation_with_optional_scope option
* notation
* constr_notation_substitution| CGeneralization of Glob_term.binding_kind
* abstraction_kind option
* constr_expr| CPrim of prim_token| CDelimiters of string * constr_expr| CArray of instance_expr option * constr_expr array * constr_expr * constr_exprand fix_expr =
Names.lident
* recursion_order_expr option
* local_binder_expr list
* constr_expr
* constr_exprand recursion_order_expr_r = | CStructRec of Names.lident| CWfRec of Names.lident * constr_expr| CMeasureRec of Names.lident option * constr_expr * constr_expr optionargument, measure, relation
*)and local_binder_expr = | CLocalAssum of Names.lname list * binder_kind * constr_expr| CLocalDef of Names.lname * constr_expr * constr_expr option| CLocalPattern of cases_pattern_exprand constr_notation_substitution =
constr_expr list
* constr_expr list list
* kinded_cases_pattern_expr list
* local_binder_expr list listConcrete syntax for modules and module types
type with_declaration_ast = | CWith_Module of Names.Id.t list CAst.t * Libnames.qualid| CWith_Definition of Names.Id.t list CAst.t
* universe_decl_expr option
* constr_exprtype module_ast_r = | CMident of Libnames.qualid| CMapply of module_ast * module_ast| CMwith of module_ast * with_declaration_ast