package coq

  1. Overview
  2. Docs
Legend:
Library
Module
Module type
Parameter
Class
Class type
type class_rawexpr =
  1. | FunClass
  2. | SortClass
  3. | RefClass of Libnames.qualid Constrexpr.or_by_notation
type goal_identifier = string
type scope_name = string
type goal_reference =
  1. | OpenSubgoals
  2. | NthGoal of int
  3. | GoalId of Names.Id.t
type printable =
  1. | PrintTables
  2. | PrintFullContext
  3. | PrintSectionContext of Libnames.qualid
  4. | PrintInspect of int
  5. | PrintGrammar of string
  6. | PrintCustomGrammar of string
  7. | PrintLoadPath of Names.DirPath.t option
  8. | PrintModules
  9. | PrintModule of Libnames.qualid
  10. | PrintModuleType of Libnames.qualid
  11. | PrintNamespace of Names.DirPath.t
  12. | PrintMLLoadPath
  13. | PrintMLModules
  14. | PrintDebugGC
  15. | PrintName of Libnames.qualid Constrexpr.or_by_notation * UnivNames.univ_name_list option
  16. | PrintGraph
  17. | PrintClasses
  18. | PrintTypeClasses
  19. | PrintInstances of Libnames.qualid Constrexpr.or_by_notation
  20. | PrintCoercions
  21. | PrintCoercionPaths of class_rawexpr * class_rawexpr
  22. | PrintCanonicalConversions
  23. | PrintUniverses of bool * Libnames.qualid list option * string option
  24. | PrintHint of Libnames.qualid Constrexpr.or_by_notation
  25. | PrintHintGoal
  26. | PrintHintDbName of string
  27. | PrintHintDb
  28. | PrintScopes
  29. | PrintScope of string
  30. | PrintVisibility of string option
  31. | PrintAbout of Libnames.qualid Constrexpr.or_by_notation * UnivNames.univ_name_list option * Goal_select.t option
  32. | PrintImplicit of Libnames.qualid Constrexpr.or_by_notation
  33. | PrintAssumptions of bool * bool * Libnames.qualid Constrexpr.or_by_notation
  34. | PrintStrategy of Libnames.qualid Constrexpr.or_by_notation option
  35. | PrintRegistered
type search_about_item =
  1. | SearchSubPattern of Constrexpr.constr_pattern_expr
  2. | SearchString of string * scope_name option
type searchable =
  1. | SearchPattern of Constrexpr.constr_pattern_expr
  2. | SearchRewrite of Constrexpr.constr_pattern_expr
  3. | SearchHead of Constrexpr.constr_pattern_expr
  4. | SearchAbout of (bool * search_about_item) list
type locatable =
  1. | LocateAny of Libnames.qualid Constrexpr.or_by_notation
  2. | LocateTerm of Libnames.qualid Constrexpr.or_by_notation
  3. | LocateLibrary of Libnames.qualid
  4. | LocateModule of Libnames.qualid
  5. | LocateOther of string * Libnames.qualid
  6. | LocateFile of string
type showable =
  1. | ShowGoal of goal_reference
  2. | ShowProof
  3. | ShowScript
  4. | ShowExistentials
  5. | ShowUniverses
  6. | ShowProofNames
  7. | ShowIntros of bool
  8. | ShowMatch of Libnames.qualid
type comment =
  1. | CommentConstr of Constrexpr.constr_expr
  2. | CommentString of string
  3. | CommentInt of int
type search_restriction =
  1. | SearchInside of Libnames.qualid list
  2. | SearchOutside of Libnames.qualid list
type rec_flag = bool
type verbose_flag = bool
type coercion_flag = bool
type instance_flag = bool option
type export_flag = bool
type inductive_flag = Declarations.recursivity_kind
type onlyparsing_flag = Flags.compat_version option
type locality_flag = bool
type option_setting =
  1. | OptionUnset
  2. | OptionSetTrue
  3. | OptionSetInt of int
  4. | OptionSetString of string
type option_ref_value =
  1. | StringRefValue of string
  2. | QualidRefValue of Libnames.qualid
type sort_expr = Sorts.family
type local_decl_expr =
  1. | AssumExpr of Names.lname * Constrexpr.constr_expr
  2. | DefExpr of Names.lname * Constrexpr.constr_expr * Constrexpr.constr_expr option
type inductive_kind =
  1. | Inductive_kw
  2. | CoInductive
  3. | Variant
  4. | Record
  5. | Structure
  6. | Class of bool
type decl_notation = Names.lstring * Constrexpr.constr_expr * scope_name option
type simple_binder = Names.lident list * Constrexpr.constr_expr
type class_binder = Names.lident * Constrexpr.constr_expr list
type !'a with_coercion = coercion_flag * 'a
type !'a with_instance = instance_flag * 'a
type !'a with_notation = 'a * decl_notation list
type !'a with_priority = 'a * int option
type constructor_list_or_record_decl_expr =
  1. | Constructors of constructor_expr list
  2. | RecordDecl of Names.lident option * local_decl_expr with_instance with_priority with_notation list
type one_inductive_expr = Names.lident * Constrexpr.local_binder_expr list * Constrexpr.constr_expr option * constructor_expr list
and typeclass_context = typeclass_constraint list
type syntax_modifier =
  1. | SetItemLevel of string list * Notation_term.constr_as_binder_kind option * Extend.production_level option
  2. | SetLevel of int
  3. | SetCustomEntry of string * int option
  4. | SetAssoc of Gramlib.Gramext.g_assoc
  5. | SetEntryType of string * Extend.simple_constr_prod_entry_key
  6. | SetOnlyParsing
  7. | SetOnlyPrinting
  8. | SetCompatVersion of Flags.compat_version
  9. | SetFormat of string * Names.lstring
type proof_end =
  1. | Admitted
  2. | Proved of Proof_global.opacity_flag * Names.lident option
type scheme =
  1. | InductionScheme of bool * Libnames.qualid Constrexpr.or_by_notation * sort_expr
  2. | CaseScheme of bool * Libnames.qualid Constrexpr.or_by_notation * sort_expr
  3. | EqualityScheme of Libnames.qualid Constrexpr.or_by_notation
type section_subset_expr =
  1. | SsEmpty
  2. | SsType
  3. | SsSingl of Names.lident
  4. | SsCompl of section_subset_expr
  5. | SsUnion of section_subset_expr * section_subset_expr
  6. | SsSubstr of section_subset_expr * section_subset_expr
  7. | SsFwdClose of section_subset_expr
type register_kind =
  1. | RegisterInline
  2. | RegisterCoqlib of Libnames.qualid
type module_binder = bool option * Names.lident list * module_ast_inl
type vernac_cumulative =
  1. | VernacCumulative
  2. | VernacNonCumulative
type vernac_argument_status = {
  1. name : Names.Name.t;
  2. recarg_like : bool;
  3. notation_scope : string CAst.t option;
  4. implicit_status : Impargs.implicit_kind;
}
type extend_name = string * int
type nonrec vernac_expr =
  1. | VernacLoad of verbose_flag * string
  2. | VernacSyntaxExtension of bool * Names.lstring * syntax_modifier list
  3. | VernacOpenCloseScope of bool * scope_name
  4. | VernacDeclareScope of scope_name
  5. | VernacDelimiters of scope_name * string option
  6. | VernacBindScope of scope_name * class_rawexpr list
  7. | VernacInfix of Names.lstring * syntax_modifier list * Constrexpr.constr_expr * scope_name option
  8. | VernacNotation of Constrexpr.constr_expr * Names.lstring * syntax_modifier list * scope_name option
  9. | VernacNotationAddFormat of string * string * string
  10. | VernacDeclareCustomEntry of string
  11. | VernacDefinition of Decl_kinds.discharge * Decl_kinds.definition_object_kind * Constrexpr.name_decl * definition_expr
  12. | VernacStartTheoremProof of Decl_kinds.theorem_kind * proof_expr list
  13. | VernacEndProof of proof_end
  14. | VernacExactProof of Constrexpr.constr_expr
  15. | VernacAssumption of Decl_kinds.discharge * Decl_kinds.assumption_object_kind * Declaremods.inline * (Constrexpr.ident_decl list * Constrexpr.constr_expr) with_coercion list
  16. | VernacInductive of vernac_cumulative option * Decl_kinds.private_flag * inductive_flag * (inductive_expr * decl_notation list) list
  17. | VernacFixpoint of Decl_kinds.discharge * (fixpoint_expr * decl_notation list) list
  18. | VernacCoFixpoint of Decl_kinds.discharge * (cofixpoint_expr * decl_notation list) list
  19. | VernacScheme of (Names.lident option * scheme) list
  20. | VernacCombinedScheme of Names.lident * Names.lident list
  21. | VernacUniverse of Names.lident list
  22. | VernacConstraint of Glob_term.glob_constraint list
  23. | VernacBeginSection of Names.lident
  24. | VernacEndSegment of Names.lident
  25. | VernacRequire of Libnames.qualid option * export_flag option * Libnames.qualid list
  26. | VernacImport of export_flag * Libnames.qualid list
  27. | VernacCanonical of Libnames.qualid Constrexpr.or_by_notation
  28. | VernacCoercion of Libnames.qualid Constrexpr.or_by_notation * class_rawexpr * class_rawexpr
  29. | VernacIdentityCoercion of Names.lident * class_rawexpr * class_rawexpr
  30. | VernacNameSectionHypSet of Names.lident * section_subset_expr
  31. | VernacInstance of Constrexpr.local_binder_expr list * typeclass_constraint * (bool * Constrexpr.constr_expr) option * Hints.hint_info_expr
  32. | VernacDeclareInstance of Constrexpr.local_binder_expr list * Constrexpr.ident_decl * Decl_kinds.binding_kind * Constrexpr.constr_expr * Hints.hint_info_expr
  33. | VernacContext of Constrexpr.local_binder_expr list
  34. | VernacExistingInstance of (Libnames.qualid * Hints.hint_info_expr) list
  35. | VernacExistingClass of Libnames.qualid
  36. | VernacDeclareModule of bool option * Names.lident * module_binder list * module_ast_inl
  37. | VernacDefineModule of bool option * Names.lident * module_binder list * module_ast_inl Declaremods.module_signature * module_ast_inl list
  38. | VernacDeclareModuleType of Names.lident * module_binder list * module_ast_inl list * module_ast_inl list
  39. | VernacInclude of module_ast_inl list
  40. | VernacSolveExistential of int * Constrexpr.constr_expr
  41. | VernacAddLoadPath of rec_flag * string * Names.DirPath.t option
  42. | VernacRemoveLoadPath of string
  43. | VernacAddMLPath of rec_flag * string
  44. | VernacDeclareMLModule of string list
  45. | VernacChdir of string option
  46. | VernacWriteState of string
  47. | VernacRestoreState of string
  48. | VernacResetName of Names.lident
  49. | VernacResetInitial
  50. | VernacBack of int
  51. | VernacBackTo of int
  52. | VernacCreateHintDb of string * bool
  53. | VernacRemoveHints of string list * Libnames.qualid list
  54. | VernacHints of string list * Hints.hints_expr
  55. | VernacSyntacticDefinition of Names.lident * Names.Id.t list * Constrexpr.constr_expr * onlyparsing_flag
  56. | VernacArguments of Libnames.qualid Constrexpr.or_by_notation * vernac_argument_status list * (Names.Name.t * Impargs.implicit_kind) list list * int option * [ `Assert | `ClearImplicits | `ClearScopes | `DefaultImplicits | `ExtraScopes | `ReductionDontExposeCase | `ReductionNeverUnfold | `Rename ] list
  57. | VernacReserve of simple_binder list
  58. | VernacGeneralizable of Names.lident list option
  59. | VernacSetOpacity of Conv_oracle.level * Libnames.qualid Constrexpr.or_by_notation list
  60. | VernacSetStrategy of (Conv_oracle.level * Libnames.qualid Constrexpr.or_by_notation list) list
  61. | VernacSetOption of export_flag * Goptions.option_name * option_setting
  62. | VernacAddOption of Goptions.option_name * option_ref_value list
  63. | VernacRemoveOption of Goptions.option_name * option_ref_value list
  64. | VernacMemOption of Goptions.option_name * option_ref_value list
  65. | VernacPrintOption of Goptions.option_name
  66. | VernacCheckMayEval of Genredexpr.raw_red_expr option * Goal_select.t option * Constrexpr.constr_expr
  67. | VernacGlobalCheck of Constrexpr.constr_expr
  68. | VernacDeclareReduction of string * Genredexpr.raw_red_expr
  69. | VernacPrint of printable
  70. | VernacSearch of searchable * Goal_select.t option * search_restriction
  71. | VernacLocate of locatable
  72. | VernacRegister of Libnames.qualid * register_kind
  73. | VernacPrimitive of Names.lident * CPrimitives.op_or_type * Constrexpr.constr_expr option
  74. | VernacComments of comment list
  75. | VernacAbort of Names.lident option
  76. | VernacAbortAll
  77. | VernacRestart
  78. | VernacUndo of int
  79. | VernacUndoTo of int
  80. | VernacFocus of int option
  81. | VernacUnfocus
  82. | VernacUnfocused
  83. | VernacBullet of Proof_bullet.t
  84. | VernacSubproof of Goal_select.t option
  85. | VernacEndSubproof
  86. | VernacShow of showable
  87. | VernacCheckGuard
  88. | VernacProof of Genarg.raw_generic_argument option * section_subset_expr option
  89. | VernacProofMode of string
  90. | VernacExtend of extend_name * Genarg.raw_generic_argument list
type vernac_control =
  1. | VernacExpr of Attributes.vernac_flags * vernac_expr
  2. | VernacTime of bool * vernac_control CAst.t
  3. | VernacRedirect of string * vernac_control CAst.t
  4. | VernacTimeout of int * vernac_control CAst.t
  5. | VernacFail of vernac_control CAst.t
OCaml

Innovation. Community. Security.