Module Vernacexpr
type class_rawexpr=|FunClass|SortClass|RefClass of Libnames.qualid Constrexpr.or_by_notationVernac expressions, produced by the parser
type goal_identifier= stringtype scope_name= stringtype goal_reference=|OpenSubgoals|NthGoal of int|GoalId of Names.Id.ttype printable=|PrintTables|PrintFullContext|PrintSectionContext of Libnames.qualid|PrintInspect of int|PrintGrammar of string|PrintCustomGrammar of string|PrintLoadPath of Names.DirPath.t option|PrintModules|PrintModule of Libnames.qualid|PrintModuleType of Libnames.qualid|PrintNamespace of Names.DirPath.t|PrintMLLoadPath|PrintMLModules|PrintDebugGC|PrintName of Libnames.qualid Constrexpr.or_by_notation * UnivNames.univ_name_list option|PrintGraph|PrintClasses|PrintTypeClasses|PrintInstances of Libnames.qualid Constrexpr.or_by_notation|PrintCoercions|PrintCoercionPaths of class_rawexpr * class_rawexpr|PrintCanonicalConversions|PrintUniverses of bool * Libnames.qualid list option * string option|PrintHint of Libnames.qualid Constrexpr.or_by_notation|PrintHintGoal|PrintHintDbName of string|PrintHintDb|PrintScopes|PrintScope of string|PrintVisibility of string option|PrintAbout of Libnames.qualid Constrexpr.or_by_notation * UnivNames.univ_name_list option * Goal_select.t option|PrintImplicit of Libnames.qualid Constrexpr.or_by_notation|PrintAssumptions of bool * bool * Libnames.qualid Constrexpr.or_by_notation|PrintStrategy of Libnames.qualid Constrexpr.or_by_notation option|PrintRegisteredtype search_about_item=|SearchSubPattern of Constrexpr.constr_pattern_expr|SearchString of string * scope_name optiontype searchable=|SearchPattern of Constrexpr.constr_pattern_expr|SearchRewrite of Constrexpr.constr_pattern_expr|SearchHead of Constrexpr.constr_pattern_expr|SearchAbout of (bool * search_about_item) listtype locatable=|LocateAny of Libnames.qualid Constrexpr.or_by_notation|LocateTerm of Libnames.qualid Constrexpr.or_by_notation|LocateLibrary of Libnames.qualid|LocateModule of Libnames.qualid|LocateOther of string * Libnames.qualid|LocateFile of stringtype showable=|ShowGoal of goal_reference|ShowProof|ShowScript|ShowExistentials|ShowUniverses|ShowProofNames|ShowIntros of bool|ShowMatch of Libnames.qualidtype comment=|CommentConstr of Constrexpr.constr_expr|CommentString of string|CommentInt of inttype search_restriction=|SearchInside of Libnames.qualid list|SearchOutside of Libnames.qualid listtype rec_flag= booltype verbose_flag= booltype coercion_flag= booltype instance_flag= bool optiontype export_flag= booltype inductive_flag= Declarations.recursivity_kindtype onlyparsing_flag= Flags.compat_version optiontype locality_flag= booltype option_setting=|OptionUnset|OptionSetTrue|OptionSetInt of int|OptionSetString of stringtype option_ref_value=|StringRefValue of string|QualidRefValue of Libnames.qualid
type sort_expr= Sorts.familytype definition_expr=|ProveBody of Constrexpr.local_binder_expr list * Constrexpr.constr_expr|DefineBody of Constrexpr.local_binder_expr list * Genredexpr.raw_red_expr option * Constrexpr.constr_expr * Constrexpr.constr_expr optiontype fixpoint_expr= Constrexpr.ident_decl * Constrexpr.recursion_order_expr option * Constrexpr.local_binder_expr list * Constrexpr.constr_expr * Constrexpr.constr_expr optiontype cofixpoint_expr= Constrexpr.ident_decl * Constrexpr.local_binder_expr list * Constrexpr.constr_expr * Constrexpr.constr_expr optiontype local_decl_expr=|AssumExpr of Names.lname * Constrexpr.constr_expr|DefExpr of Names.lname * Constrexpr.constr_expr * Constrexpr.constr_expr optiontype inductive_kind=|Inductive_kw|CoInductive|Variant|Record|Structure|Class of booltype decl_notation= Names.lstring * Constrexpr.constr_expr * scope_name optiontype simple_binder= Names.lident list * Constrexpr.constr_exprtype class_binder= Names.lident * Constrexpr.constr_expr listtype 'a with_coercion= coercion_flag * 'atype 'a with_instance= instance_flag * 'atype 'a with_notation= 'a * decl_notation listtype 'a with_priority= 'a * int optiontype constructor_expr= (Names.lident * Constrexpr.constr_expr) with_coerciontype constructor_list_or_record_decl_expr=|Constructors of constructor_expr list|RecordDecl of Names.lident option * local_decl_expr with_instance with_priority with_notation listtype inductive_expr= Constrexpr.ident_decl with_coercion * Constrexpr.local_binder_expr list * Constrexpr.constr_expr option * inductive_kind * constructor_list_or_record_decl_exprtype one_inductive_expr= Names.lident * Constrexpr.local_binder_expr list * Constrexpr.constr_expr option * constructor_expr listtype typeclass_constraint= Constrexpr.name_decl * Decl_kinds.binding_kind * Constrexpr.constr_exprand typeclass_context= typeclass_constraint listtype proof_expr= Constrexpr.ident_decl * (Constrexpr.local_binder_expr list * Constrexpr.constr_expr)type syntax_modifier=|SetItemLevel of string list * Notation_term.constr_as_binder_kind option * Extend.production_level option|SetLevel of int|SetCustomEntry of string * int option|SetAssoc of Gramlib.Gramext.g_assoc|SetEntryType of string * Extend.simple_constr_prod_entry_key|SetOnlyParsing|SetOnlyPrinting|SetCompatVersion of Flags.compat_version|SetFormat of string * Names.lstringtype proof_end=|Admitted|Proved of Proof_global.opacity_flag * Names.lident optiontype scheme=|InductionScheme of bool * Libnames.qualid Constrexpr.or_by_notation * sort_expr|CaseScheme of bool * Libnames.qualid Constrexpr.or_by_notation * sort_expr|EqualityScheme of Libnames.qualid Constrexpr.or_by_notationtype section_subset_expr=|SsEmpty|SsType|SsSingl of Names.lident|SsCompl of section_subset_expr|SsUnion of section_subset_expr * section_subset_expr|SsSubstr of section_subset_expr * section_subset_expr|SsFwdClose of section_subset_expr
type register_kind=|RegisterInline|RegisterCoqlib of Libnames.qualid
Types concerning the module layer
type module_ast_inl= Constrexpr.module_ast * Declaremods.inlinetype module_binder= bool option * Names.lident list * module_ast_inltype vernac_cumulative=|VernacCumulative|VernacNonCumulativeSome bif locally enabled/disabled according tob,Noneif we should use the global flag.
The type of vernacular expressions
type vernac_argument_status={name : Names.Name.t;recarg_like : bool;notation_scope : string CAst.t option;implicit_status : Impargs.implicit_kind;}type extend_name= string * inttype nonrec vernac_expr=|VernacLoad of verbose_flag * string|VernacSyntaxExtension of bool * Names.lstring * syntax_modifier list|VernacOpenCloseScope of bool * scope_name|VernacDeclareScope of scope_name|VernacDelimiters of scope_name * string option|VernacBindScope of scope_name * class_rawexpr list|VernacInfix of Names.lstring * syntax_modifier list * Constrexpr.constr_expr * scope_name option|VernacNotation of Constrexpr.constr_expr * Names.lstring * syntax_modifier list * scope_name option|VernacNotationAddFormat of string * string * string|VernacDeclareCustomEntry of string|VernacDefinition of Decl_kinds.discharge * Decl_kinds.definition_object_kind * Constrexpr.name_decl * definition_expr|VernacStartTheoremProof of Decl_kinds.theorem_kind * proof_expr list|VernacEndProof of proof_end|VernacExactProof of Constrexpr.constr_expr|VernacAssumption of Decl_kinds.discharge * Decl_kinds.assumption_object_kind * Declaremods.inline * (Constrexpr.ident_decl list * Constrexpr.constr_expr) with_coercion list|VernacInductive of vernac_cumulative option * Decl_kinds.private_flag * inductive_flag * (inductive_expr * decl_notation list) list|VernacFixpoint of Decl_kinds.discharge * (fixpoint_expr * decl_notation list) list|VernacCoFixpoint of Decl_kinds.discharge * (cofixpoint_expr * decl_notation list) list|VernacScheme of (Names.lident option * scheme) list|VernacCombinedScheme of Names.lident * Names.lident list|VernacUniverse of Names.lident list|VernacConstraint of Glob_term.glob_constraint list|VernacBeginSection of Names.lident|VernacEndSegment of Names.lident|VernacRequire of Libnames.qualid option * export_flag option * Libnames.qualid list|VernacImport of export_flag * Libnames.qualid list|VernacCanonical of Libnames.qualid Constrexpr.or_by_notation|VernacCoercion of Libnames.qualid Constrexpr.or_by_notation * class_rawexpr * class_rawexpr|VernacIdentityCoercion of Names.lident * class_rawexpr * class_rawexpr|VernacNameSectionHypSet of Names.lident * section_subset_expr|VernacInstance of Constrexpr.local_binder_expr list * typeclass_constraint * (bool * Constrexpr.constr_expr) option * Hints.hint_info_expr|VernacDeclareInstance of Constrexpr.local_binder_expr list * Constrexpr.ident_decl * Decl_kinds.binding_kind * Constrexpr.constr_expr * Hints.hint_info_expr|VernacContext of Constrexpr.local_binder_expr list|VernacExistingInstance of (Libnames.qualid * Hints.hint_info_expr) list|VernacExistingClass of Libnames.qualid|VernacDeclareModule of bool option * Names.lident * module_binder list * module_ast_inl|VernacDefineModule of bool option * Names.lident * module_binder list * module_ast_inl Declaremods.module_signature * module_ast_inl list|VernacDeclareModuleType of Names.lident * module_binder list * module_ast_inl list * module_ast_inl list|VernacInclude of module_ast_inl list|VernacSolveExistential of int * Constrexpr.constr_expr|VernacAddLoadPath of rec_flag * string * Names.DirPath.t option|VernacRemoveLoadPath of string|VernacAddMLPath of rec_flag * string|VernacDeclareMLModule of string list|VernacChdir of string option|VernacWriteState of string|VernacRestoreState of string|VernacResetName of Names.lident|VernacResetInitial|VernacBack of int|VernacBackTo of int|VernacCreateHintDb of string * bool|VernacRemoveHints of string list * Libnames.qualid list|VernacHints of string list * Hints.hints_expr|VernacSyntacticDefinition of Names.lident * Names.Id.t list * Constrexpr.constr_expr * onlyparsing_flag|VernacArguments of Libnames.qualid Constrexpr.or_by_notation * vernac_argument_status list * (Names.Name.t * Impargs.implicit_kind) list list * int option * [ `ReductionDontExposeCase | `ReductionNeverUnfold | `Rename | `ExtraScopes | `Assert | `ClearImplicits | `ClearScopes | `DefaultImplicits ] list|VernacReserve of simple_binder list|VernacGeneralizable of Names.lident list option|VernacSetOpacity of Conv_oracle.level * Libnames.qualid Constrexpr.or_by_notation list|VernacSetStrategy of (Conv_oracle.level * Libnames.qualid Constrexpr.or_by_notation list) list|VernacSetOption of export_flag * Goptions.option_name * option_setting|VernacAddOption of Goptions.option_name * option_ref_value list|VernacRemoveOption of Goptions.option_name * option_ref_value list|VernacMemOption of Goptions.option_name * option_ref_value list|VernacPrintOption of Goptions.option_name|VernacCheckMayEval of Genredexpr.raw_red_expr option * Goal_select.t option * Constrexpr.constr_expr|VernacGlobalCheck of Constrexpr.constr_expr|VernacDeclareReduction of string * Genredexpr.raw_red_expr|VernacPrint of printable|VernacSearch of searchable * Goal_select.t option * search_restriction|VernacLocate of locatable|VernacRegister of Libnames.qualid * register_kind|VernacPrimitive of Names.lident * CPrimitives.op_or_type * Constrexpr.constr_expr option|VernacComments of comment list|VernacAbort of Names.lident option|VernacAbortAll|VernacRestart|VernacUndo of int|VernacUndoTo of int|VernacFocus of int option|VernacUnfocus|VernacUnfocused|VernacBullet of Proof_bullet.t|VernacSubproof of Goal_select.t option|VernacEndSubproof|VernacShow of showable|VernacCheckGuard|VernacProof of Genarg.raw_generic_argument option * section_subset_expr option|VernacProofMode of string|VernacExtend of extend_name * Genarg.raw_generic_argument listtype vernac_control=|VernacExpr of Attributes.vernac_flags * vernac_expr|VernacTime of bool * vernac_control CAst.t|VernacRedirect of string * vernac_control CAst.t|VernacTimeout of int * vernac_control CAst.t|VernacFail of vernac_control CAst.t