aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--.gitlab-ci.yml5
-rw-r--r--API/API.ml427
-rw-r--r--API/API.mli6236
-rw-r--r--API/API.mllib1
-rw-r--r--API/grammar_API.ml63
-rw-r--r--API/grammar_API.mli249
-rw-r--r--Makefile.build6
-rw-r--r--Makefile.ide8
-rw-r--r--configure.ml25
-rw-r--r--dev/doc/naming-conventions.tex2
-rw-r--r--dev/top_printers.ml3
-rw-r--r--doc/stdlib/index-list.html.template1
-rw-r--r--grammar/argextend.mlp34
-rw-r--r--grammar/q_util.mlp22
-rw-r--r--grammar/tacextend.mlp22
-rw-r--r--grammar/vernacextend.mlp18
-rw-r--r--interp/constrintern.ml4
-rw-r--r--kernel/entries.ml (renamed from kernel/entries.mli)0
-rw-r--r--kernel/kernel.mllib5
-rw-r--r--lib/flags.ml6
-rw-r--r--lib/flags.mli2
-rw-r--r--lib/profile.ml29
-rw-r--r--plugins/extraction/extract_env.ml82
-rw-r--r--plugins/extraction/extract_env.mli4
-rw-r--r--plugins/extraction/g_extraction.ml46
-rw-r--r--plugins/extraction/modutil.ml9
-rw-r--r--plugins/extraction/modutil.mli1
-rw-r--r--plugins/firstorder/g_ground.ml41
-rw-r--r--plugins/funind/g_indfun.ml41
-rw-r--r--plugins/ltac/extraargs.ml41
-rw-r--r--plugins/ltac/extraargs.mli1
-rw-r--r--plugins/ltac/extratactics.ml41
-rw-r--r--plugins/ltac/g_auto.ml41
-rw-r--r--plugins/ltac/g_ltac.ml42
-rw-r--r--plugins/ltac/g_obligations.ml41
-rw-r--r--plugins/ltac/g_rewrite.ml41
-rw-r--r--plugins/ltac/g_tactic.ml41
-rw-r--r--plugins/ltac/pltac.ml1
-rw-r--r--plugins/ltac/pltac.mli1
-rw-r--r--plugins/ltac/tacentries.ml1
-rw-r--r--plugins/ltac/tacentries.mli1
-rw-r--r--plugins/ltac/tacintern.ml1
-rw-r--r--plugins/ltac/tacintern.mli1
-rw-r--r--plugins/ltac/tacinterp.ml1
-rw-r--r--plugins/ltac/tacsubst.ml1
-rw-r--r--plugins/setoid_ring/g_newring.ml41
-rw-r--r--plugins/ssr/ssrcommon.ml1
-rw-r--r--plugins/ssr/ssrparser.ml41
-rw-r--r--plugins/ssr/ssrparser.mli2
-rw-r--r--plugins/ssr/ssrvernac.ml41
-rw-r--r--plugins/ssrmatching/ssrmatching.ml42
-rw-r--r--plugins/ssrmatching/ssrmatching.mli1
-rw-r--r--proofs/proof_type.ml (renamed from proofs/proof_type.mli)0
-rw-r--r--proofs/proofs.mllib1
-rw-r--r--tactics/class_tactics.ml2
-rw-r--r--test-suite/bugs/closed/4720.v46
-rw-r--r--test-suite/bugs/closed/5177.v21
-rw-r--r--theories/Compat/Coq86.v2
-rw-r--r--theories/Compat/Coq87.v9
-rw-r--r--theories/FSets/FMapPositive.v8
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivEucl.v8
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivFloor.v8
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivTrunc.v8
-rw-r--r--theories/ZArith/Zdiv.v7
-rw-r--r--toplevel/coqinit.ml3
-rw-r--r--toplevel/coqtop.ml2
66 files changed, 4131 insertions, 3292 deletions
diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml
index 8b43d975a..b47494d3a 100644
--- a/.gitlab-ci.yml
+++ b/.gitlab-ci.yml
@@ -40,6 +40,11 @@ before_script:
- if [ ! "(" -d .opamcache ")" ]; then mv ~/.opam .opamcache; else mv ~/.opam ~/.opam-old; fi
- ln -s $(readlink -f .opamcache) ~/.opam
+ # the default repo in this docker image is a local directory
+ # at the time of 4aaeb8abf it lagged behind the official
+ # repository such that camlp5 7.01 was not available
+ - opam repository set-url default https://opam.ocaml.org
+ - opam update
- opam switch ${COMPILER}
- eval $(opam config env)
- opam config list
diff --git a/API/API.ml b/API/API.ml
index 32c664d7b..c952e123d 100644
--- a/API/API.ml
+++ b/API/API.ml
@@ -6,199 +6,280 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-module Ppvernac = Ppvernac
-module Command = Command
-module States = States
-module Kindops = Kindops
+(* Warning, this file respects the dependency order established in Coq.
+
+ To see such order issue the comand:
+
+```
+bash -c 'for i in kernel intf library engine pretyping interp proofs parsing printing tactics vernac stm toplevel; do echo -e "\n## $i files" && cat ${i}/${i}.mllib; done && echo -e "\n## highparsing files" && cat parsing/highparsing.mllib' > API/link
+```
+ *)
+
+(******************************************************************************)
+(* config *)
+(******************************************************************************)
module Coq_config = Coq_config
+
+(******************************************************************************)
+(* Kernel *)
+(******************************************************************************)
+(* "mli" files *)
+module Declarations = Declarations
+module Entries = Entries
+
+module Names = Names
+(* module Uint31 *)
+module Univ = Univ
+module UGraph = UGraph
module Esubst = Esubst
+module Sorts = Sorts
module Evar = Evar
-module Constrexpr = Constrexpr
-module Libobject = Libobject
-module Evd = Evd
-module Libnames = Libnames
-module Nameops = Nameops
-module Topfmt = Topfmt
-module Locus = Locus
-module Locusops = Locusops
-module Lemmas = Lemmas
-module Clenv = Clenv
-module Elimschemes = Elimschemes
-module Classes = Classes
-module Class_tactics = Class_tactics
-module Eauto = Eauto
-module Keys = Keys
-module Vernac_classifier = Vernac_classifier
-module Autorewrite = Autorewrite
-module Redops = Redops
-module Elim = Elim
-module Geninterp = Geninterp
-module Obligations = Obligations
-module Retroknowledge = Retroknowledge
-module Evar_refiner = Evar_refiner
-module Hipattern = Hipattern
-module Auto = Auto
-module Hints = Hints
-module Contradiction = Contradiction
-module Tacticals = Tacticals
-module Tactics = Tactics
-module Inv = Inv
-module Leminv = Leminv
-module Equality = Equality
-module Redexpr = Redexpr
-module Pfedit = Pfedit
-module Stm = Stm
-module Stateid = Stateid
-module Declaremods = Declaremods
-module Miscops = Miscops
-module Miscprint = Miscprint
-module Genprint = Genprint
-module Ppconstr = Ppconstr
-module Pputils = Pputils
-module Extend = Extend
-module Logic = Logic
-module Himsg = Himsg
-module Tacred = Tacred
-module Names = Names
-module Indrec = Indrec
-module Glob_ops = Glob_ops
-module Constrexpr_ops = Constrexpr_ops
-module Eqdecide = Eqdecide
-module Genredexpr = Genredexpr
-module Detyping = Detyping
-module Tactypes = Tactypes
-module ExplainErr = ExplainErr
-module Printer = Printer
-module Constrextern = Constrextern
-module Locality = Locality
-module Impargs = Impargs
-module Termops = Termops
-module Refiner = Refiner
-module Ppextend = Ppextend
-module Nametab = Nametab
-module Vernacentries = Vernacentries
-module Mltop = Mltop
-module Goal = Goal
-module Proof_bullet = Proof_bullet
-module Proof_global = Proof_global
-module Proof = Proof
-module Smartlocate = Smartlocate
-module Dumpglob = Dumpglob
-module Constrintern = Constrintern
-module Topconstr = Topconstr
-module Notation_ops = Notation_ops
-module Patternops = Patternops
-module Mod_typing = Mod_typing
-module Modops = Modops
-module Opaqueproof = Opaqueproof
-module Ind_tables = Ind_tables
-module Typeops = Typeops
-module Inductive = Inductive
+module Constr = Constr
+module Context = Context
module Vars = Vars
-module Reduction = Reduction
+module Term = Term
module Mod_subst = Mod_subst
-module Sorts = Sorts
-module Univ = Univ
-module Constr = Constr
+module Cbytecodes = Cbytecodes
+(* module Copcodes *)
+module Cemitcodes = Cemitcodes
+(* module Nativevalues *)
+(* module Primitives *)
+module Opaqueproof = Opaqueproof
+module Declareops = Declareops
+module Retroknowledge = Retroknowledge
+module Conv_oracle = Conv_oracle
+(* module Pre_env *)
+(* module Cbytegen *)
+(* module Nativelambda *)
+(* module Nativecode *)
+(* module Nativelib *)
+module Environ = Environ
module CClosure = CClosure
+module Reduction = Reduction
+(* module Nativeconv *)
module Type_errors = Type_errors
+module Modops = Modops
+module Inductive = Inductive
+module Typeops = Typeops
+(* module Indtypes *)
+(* module Cooking *)
+(* module Term_typing *)
+(* module Subtyping *)
+module Mod_typing = Mod_typing
+(* module Nativelibrary *)
module Safe_typing = Safe_typing
-module UGraph = UGraph
-module Namegen = Namegen
-module Ftactic = Ftactic
-module UState = UState
-module Proofview_monad = Proofview_monad
-module Classops = Classops
+(* module Vm *)
+(* module Csymtable *)
+(* module Vconv *)
+
+(******************************************************************************)
+(* Intf *)
+(******************************************************************************)
+module Constrexpr = Constrexpr
+module Locus = Locus
+module Glob_term = Glob_term
+module Extend = Extend
+module Misctypes = Misctypes
+module Decl_kinds = Decl_kinds
+module Vernacexpr = Vernacexpr
+module Notation_term = Notation_term
+module Evar_kinds = Evar_kinds
+module Genredexpr = Genredexpr
+
+(******************************************************************************)
+(* Library *)
+(******************************************************************************)
+module Univops = Univops
+module Nameops = Nameops
+module Libnames = Libnames
+module Globnames = Globnames
+module Libobject = Libobject
+module Summary = Summary
+module Nametab = Nametab
module Global = Global
-module Goptions = Goptions
module Lib = Lib
+module Declaremods = Declaremods
+(* module Loadpath *)
module Library = Library
-module Summary = Summary
+module States = States
+module Kindops = Kindops
+(* module Dischargedhypsmap *)
+module Goptions = Goptions
+(* module Decls *)
+(* module Heads *)
+module Keys = Keys
+module Coqlib = Coqlib
+
+(******************************************************************************)
+(* Engine *)
+(******************************************************************************)
+(* module Logic_monad *)
module Universes = Universes
-module Declare = Declare
-module Refine = Refine
-module Find_subterm = Find_subterm
-module Evar_kinds = Evar_kinds
-module Decl_kinds = Decl_kinds
-module Misctypes = Misctypes
+module UState = UState
+module Evd = Evd
+module EConstr = EConstr
+module Tactypes = Tactypes
module Pattern = Pattern
-module Vernacexpr = Vernacexpr
-module Search = Search
-module Notation_term = Notation_term
+module Namegen = Namegen
+module Termops = Termops
+module Proofview_monad = Proofview_monad
+module Evarutil = Evarutil
+module Proofview = Proofview
+module Ftactic = Ftactic
+module Geninterp = Geninterp
+
+(******************************************************************************)
+(* Pretyping *)
+(******************************************************************************)
+module Locusops = Locusops
+module Pretype_errors = Pretype_errors
module Reductionops = Reductionops
module Inductiveops = Inductiveops
-module Recordops = Recordops
+(* module Vnorm *)
+(* module Arguments_renaming *)
+module Impargs = Impargs
+(* module Nativenorm *)
module Retyping = Retyping
-module Typing = Typing
+(* module Cbv *)
+module Find_subterm = Find_subterm
+(* module Evardefine *)
module Evarsolve = Evarsolve
+module Recordops = Recordops
+module Evarconv = Evarconv
+module Typing = Typing
+module Miscops = Miscops
+module Glob_ops = Glob_ops
+module Redops = Redops
+module Patternops = Patternops
module Constr_matching = Constr_matching
+module Tacred = Tacred
+module Typeclasses = Typeclasses
+module Classops = Classops
+(* module Program *)
+(* module Coercion *)
+module Detyping = Detyping
+module Indrec = Indrec
+(* module Cases *)
module Pretyping = Pretyping
-module Evarconv = Evarconv
module Unification = Unification
-module Typeclasses = Typeclasses
-module Pretype_errors = Pretype_errors
-module Notation = Notation
-module Declarations = Declarations
-module Univops = Univops
-module Declareops = Declareops
-module Globnames = Globnames
-module Environ = Environ
-module Term = Term
-module Coqlib = Coqlib
-module Glob_term = Glob_term
-module Context = Context
+(******************************************************************************)
+(* interp *)
+(******************************************************************************)
module Stdarg = Stdarg
+module Genintern = Genintern
+module Constrexpr_ops = Constrexpr_ops
+module Notation_ops = Notation_ops
+module Ppextend = Ppextend
+module Notation = Notation
+module Dumpglob = Dumpglob
+(* module Syntax_def *)
+module Smartlocate = Smartlocate
+module Topconstr = Topconstr
+(* module Reserve *)
+(* module Implicit_quantifiers *)
+module Constrintern = Constrintern
+(* module Modintern *)
+module Constrextern = Constrextern
+(* module Discharge *)
+module Declare = Declare
+
+(******************************************************************************)
+(* Proofs *)
+(******************************************************************************)
+module Miscprint = Miscprint
+module Goal = Goal
+module Evar_refiner = Evar_refiner
+(* module Proof_using *)
+module Proof_type = Proof_type
+module Logic = Logic
+module Refine = Refine
+module Proof = Proof
+module Proof_bullet = Proof_bullet
+module Proof_global = Proof_global
+module Redexpr = Redexpr
+module Refiner = Refiner
module Tacmach = Tacmach
-module Proofview = Proofview
-module Evarutil = Evarutil
-module EConstr = EConstr
+module Pfedit = Pfedit
+module Clenv = Clenv
+(* module Clenvtac *)
+(* "mli" file *)
+
+(******************************************************************************)
+(* Printing *)
+(******************************************************************************)
+module Genprint = Genprint
+module Pputils = Pputils
+module Ppconstr = Ppconstr
+module Printer = Printer
+(* module Printmod *)
+(* module Prettyp *)
+module Ppvernac = Ppvernac
+
+(******************************************************************************)
+(* Parsing *)
+(******************************************************************************)
+module Tok = Tok
+module CLexer = CLexer
+module Pcoq = Pcoq
+module Egramml = Egramml
+(* Egramcoq *)
+
+(******************************************************************************)
+(* Tactics *)
+(******************************************************************************)
+(* module Dnet *)
+(* module Dn *)
+(* module Btermdn *)
+module Tacticals = Tacticals
+module Hipattern = Hipattern
+module Ind_tables = Ind_tables
+(* module Eqschemes *)
+module Elimschemes = Elimschemes
+module Tactics = Tactics
+module Elim = Elim
+module Equality = Equality
+module Contradiction = Contradiction
+module Inv = Inv
+module Leminv = Leminv
+module Hints = Hints
+module Auto = Auto
+module Eauto = Eauto
+module Class_tactics = Class_tactics
+(* module Term_dnet *)
+module Eqdecide = Eqdecide
+module Autorewrite = Autorewrite
+
+(******************************************************************************)
+(* Vernac *)
+(******************************************************************************)
+(* module Vernacprop *)
+module Lemmas = Lemmas
+module Himsg = Himsg
+module ExplainErr = ExplainErr
+(* module Class *)
+module Locality = Locality
+module Metasyntax = Metasyntax
+(* module Auto_ind_decl *)
+module Search = Search
+(* module Indschemes *)
+module Obligations = Obligations
+module Command = Command
+module Classes = Classes
+(* module Record *)
+(* module Assumptions *)
+module Vernacinterp = Vernacinterp
+module Mltop = Mltop
+module Topfmt = Topfmt
+module Vernacentries = Vernacentries
+
+(******************************************************************************)
+(* Stm *)
+(******************************************************************************)
+module Vernac_classifier = Vernac_classifier
+module Stm = Stm
-module Prelude =
- struct
- type global_reference = Globnames.global_reference
- type metavariable = int
- type meta_value_map = (metavariable * Constr.constr) list
- type named_context_val = Environ.named_context_val
- type conv_pb = Reduction.conv_pb =
- | CONV
- | CUMUL
- type constr = Constr.constr
- type types = Constr.types
- type evar = Constr.existential_key
- type 'constr pexistential = 'constr Constr.pexistential
- type env = Environ.env
- type evar_map = Evd.evar_map
- type rigid = Evd.rigid =
- | UnivRigid
- | UnivFlexible of bool
- type reference = Libnames.reference =
- | Qualid of Libnames.qualid Loc.located
- | Ident of Names.Id.t Loc.located
- end
-
-(* NOTE: It does not make sense to replace the following "module expression"
- simply with "module Proof_type = Proof_type" because
- there is only "kernel/entries.mli";
- there is no "kernel/entries.ml" file *)
-module Entries =
- struct
- type mutual_inductive_entry = Entries.mutual_inductive_entry
- type inline = int option
- type 'a proof_output = Constr.constr Univ.in_universe_context_set * 'a
- type 'a const_entry_body = 'a proof_output Future.computation
- type 'a definition_entry = 'a Entries.definition_entry =
- { const_entry_body : 'a const_entry_body;
- const_entry_secctx : Context.Named.t option;
- const_entry_feedback : Stateid.t option;
- const_entry_type : Term.types option;
- const_entry_polymorphic : bool;
- const_entry_universes : Univ.universe_context;
- const_entry_opaque : bool;
- const_entry_inline_code : bool }
- type parameter_entry = Entries.parameter_entry
- type projection_entry = Entries.projection_entry
- type 'a constant_entry = 'a Entries.constant_entry =
- | DefinitionEntry of 'a definition_entry
- | ParameterEntry of parameter_entry
- | ProjectionEntry of projection_entry
- end
+(******************************************************************************)
+(* Highparsing *)
+(******************************************************************************)
+module G_vernac = G_vernac
+module G_proofs = G_proofs
diff --git a/API/API.mli b/API/API.mli
index b1a746e02..bb24d5768 100644
--- a/API/API.mli
+++ b/API/API.mli
@@ -6,181 +6,78 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-module Prelude :
-sig
- (* None of the items in this modules are meant to be used by plugin-writers.
- This module is here only for "technical reasons"
- (it will disappear when we take advantage of mutually-recursive modules) *)
-
- (* API.Term.constr *)
- type constr = Constr.t
-
- (* API.Term.types *)
- type types = Constr.t
-
- (* API.Evar.t *)
- type evar = Evar.t
-
- (* 'constr API.Term.pexistential *)
- type 'constr pexistential = evar * 'constr array
-
- (* API.Environ.env *)
- type env = Environ.env
+(* Warning, this file should respect the dependency order established
+ in Coq. To see such order issue the comand:
- (* API.Evar.Map.t *)
- type evar_map = Evd.evar_map
+ ```
+ bash -c 'for i in kernel intf library engine pretyping interp proofs parsing printing tactics vernac stm toplevel; do echo -e "\n## $i files" && cat ${i}/${i}.mllib; done && echo -e "\n## highparsing files" && cat parsing/highparsing.mllib' > API/link
+ ```
- (* API.Globnames.global_reference *)
- type global_reference = Globnames.global_reference
+ Note however that files in intf/ are located manually now as their
+ conceptual linking order in the Coq codebase is incorrect (but it
+ works due to these files being implementation-free.
- type rigid = Evd.rigid =
- | UnivRigid
- | UnivFlexible of bool
+ See below in the file for their concrete position.
+*)
- type conv_pb = Reduction.conv_pb =
- | CONV
- | CUMUL
-
- type named_context_val = Environ.named_context_val
-
- type metavariable = int
-
- (* Termops.meta_value_map *)
- type meta_value_map = (metavariable * constr) list
-
- (* API.Libnames.reference *)
- type reference = Libnames.reference =
- | Qualid of Libnames.qualid Loc.located
- | Ident of Names.Id.t Loc.located
+(************************************************************************)
+(* Modules from config/ *)
+(************************************************************************)
+module Coq_config :
+sig
+ val exec_extension : string
end
-module Univ :
+(************************************************************************)
+(* Modules from kernel/ *)
+(************************************************************************)
+module Names :
sig
- module Level :
- sig
- type t = Univ.Level.t
- val set : t
- val pr : t -> Pp.std_ppcmds
- end
- module Instance :
- sig
- type t = Univ.Instance.t
- val empty : t
- val of_array : Level.t array -> t
- val to_array : t -> Level.t array
- val pr : (Level.t -> Pp.std_ppcmds) -> t -> Pp.std_ppcmds
- end
- type 'a puniverses = 'a * Instance.t
- val out_punivs : 'a puniverses -> 'a
-
- module Constraint : module type of struct include Univ.Constraint end
-
- type 'a constrained = 'a * Constraint.t
-
- module UContext :
- sig
- type t = Univ.UContext.t
- val empty : t
- end
+ open Util
- module AUContext :
+ module Id :
sig
- type t = Univ.AUContext.t
- end
-
- type universe_context = UContext.t
- [@@ocaml.deprecated "alias of API.Univ.UContext.t"]
-
- type abstract_universe_context = Univ.AUContext.t
- type cumulativity_info = Univ.CumulativityInfo.t
- type abstract_cumulativity_info = Univ.ACumulativityInfo.t
+ type t
+ val equal : t -> t -> bool
+ val compare : t -> t -> int
+ val hash : t -> int
+ val is_valid : string -> bool
+ val of_bytes : bytes -> t
+ val of_string : string -> t
+ val of_string_soft : string -> t
+ val to_string : t -> string
+ val print : t -> Pp.std_ppcmds
- module LSet : module type of struct include Univ.LSet end
- module ContextSet :
- sig
- type t = Univ.ContextSet.t
- val empty : t
- val of_context : UContext.t -> t
- val to_context : t -> UContext.t
+ module Set : Set.S with type elt = t
+ module Map : Map.ExtS with type key = t and module Set := Set
+ module Pred : Predicate.S with type elt = t
+ module List : List.MonoS with type elt = t
+ val hcons : t -> t
end
- type 'a in_universe_context_set = 'a * ContextSet.t
- type 'a in_universe_context = 'a * UContext.t
- type constraint_type = Univ.constraint_type
-
- module Universe :
+ module Name :
sig
- type t = Univ.Universe.t
- val pr : t -> Pp.std_ppcmds
- end
-
- type universe_context_set = ContextSet.t
- [@@ocaml.deprecated "alias of API.Names.ContextSet.t"]
-
- type universe_set = LSet.t
- [@@ocaml.deprecated "alias of API.Names.LSet.t"]
-
- type 'a constraint_function = 'a -> 'a -> Constraint.t -> Constraint.t
- type universe_subst = Univ.universe_subst
- type universe_level_subst = Univ.universe_level_subst
-
- val enforce_leq : Universe.t constraint_function
- val pr_uni : Universe.t -> Pp.std_ppcmds
- val pr_universe_context : (Level.t -> Pp.std_ppcmds) -> UContext.t -> Pp.std_ppcmds
- val pr_universe_context_set : (Level.t -> Pp.std_ppcmds) -> ContextSet.t -> Pp.std_ppcmds
- val pr_universe_subst : universe_subst -> Pp.std_ppcmds
- val pr_universe_level_subst : universe_level_subst -> Pp.std_ppcmds
- val pr_constraints : (Level.t -> Pp.std_ppcmds) -> Constraint.t -> Pp.std_ppcmds
-end
-
-module UState :
-sig
- type t = UState.t
- val context : t -> Univ.UContext.t
- val context_set : t -> Univ.ContextSet.t
- val of_context_set : Univ.ContextSet.t -> t
-end
-
-module Sorts :
-sig
- type contents = Sorts.contents = Pos | Null
- type t = Sorts.t =
- | Prop of contents
- | Type of Univ.Universe.t
- val is_prop : t -> bool
- val hash : t -> int
-
- type family = Sorts.family = InProp | InSet | InType
- val family : t -> family
-end
-
-module Names :
-sig
- module Id : module type of struct include Names.Id end
-
- module MBId : sig
- type t = Names.MBId.t
+ type t = Anonymous (** anonymous identifier *)
+ | Name of Id.t (** non-anonymous identifier *)
+ val mk_name : Id.t -> t
+ val is_anonymous : t -> bool
+ val is_name : t -> bool
+ val compare : t -> t -> int
val equal : t -> t -> bool
- val to_id : t -> Names.Id.t
- val repr : t -> int * Names.Id.t * Names.DirPath.t
- val debug_to_string : t -> string
+ val hash : t -> int
+ val hcons : t -> t
+ val print : t -> Pp.std_ppcmds
end
- type evaluable_global_reference = Names.evaluable_global_reference =
- | EvalVarRef of Id.t
- | EvalConstRef of Names.Constant.t
-
- module Name : module type of struct include Names.Name end
-
type name = Name.t =
- | Anonymous
+ | Anonymous
| Name of Id.t
[@@ocaml.deprecated "alias of API.Name.t"]
module DirPath :
sig
- type t = Names.DirPath.t
+ type t
val empty : t
val make : Id.t list -> t
val repr : t -> Id.t list
@@ -188,23 +85,31 @@ sig
val to_string : t -> string
end
+ module MBId : sig
+ type t
+ val equal : t -> t -> bool
+ val to_id : t -> Id.t
+ val repr : t -> int * Id.t * DirPath.t
+ val debug_to_string : t -> string
+ end
+
module Label :
sig
- type t = Names.Label.t
+ type t
val make : string -> t
val equal : t -> t -> bool
val compare : t -> t -> int
- val of_id : Names.Id.t -> t
- val to_id : t -> Names.Id.t
+ val of_id : Id.t -> t
+ val to_id : t -> Id.t
val to_string : t -> string
end
module ModPath :
sig
- type t = Names.ModPath.t =
- | MPfile of Names.DirPath.t
- | MPbound of MBId.t
- | MPdot of t * Label.t
+ type t =
+ | MPfile of DirPath.t
+ | MPbound of MBId.t
+ | MPdot of t * Label.t
val compare : t -> t -> int
val equal : t -> t -> bool
val hash : t -> int
@@ -215,7 +120,7 @@ sig
module KerName :
sig
- type t = Names.KerName.t
+ type t
val make : ModPath.t -> DirPath.t -> Label.t -> t
val make2 : ModPath.t -> Label.t -> t
val modpath : t -> ModPath.t
@@ -232,40 +137,44 @@ sig
module Constant :
sig
- type t = Names.Constant.t
+ type t
val equal : t -> t -> bool
- val make1 : Names.KerName.t -> t
- val make2 : Names.ModPath.t -> Label.t -> t
- val make3 : Names.ModPath.t -> Names.DirPath.t -> Label.t -> t
- val repr3 : t -> Names.ModPath.t * Names.DirPath.t * Label.t
- val canonical : t -> Names.KerName.t
- val user : t -> Names.KerName.t
+ val make1 : KerName.t -> t
+ val make2 : ModPath.t -> Label.t -> t
+ val make3 : ModPath.t -> DirPath.t -> Label.t -> t
+ val repr3 : t -> ModPath.t * DirPath.t * Label.t
+ val canonical : t -> KerName.t
+ val user : t -> KerName.t
val label : t -> Label.t
end
module MutInd :
sig
- type t = Names.MutInd.t
- val make1 : Names.KerName.t -> t
- val make2 : Names.ModPath.t -> Label.t -> t
+ type t
+ val make1 : KerName.t -> t
+ val make2 : ModPath.t -> Label.t -> t
val equal : t -> t -> bool
- val repr3 : t -> Names.ModPath.t * Names.DirPath.t * Label.t
- val canonical : t -> Names.KerName.t
- val modpath : t -> Names.ModPath.t
+ val repr3 : t -> ModPath.t * DirPath.t * Label.t
+ val canonical : t -> KerName.t
+ val modpath : t -> ModPath.t
val label : t -> Label.t
- val user : t -> Names.KerName.t
+ val user : t -> KerName.t
val print : t -> Pp.std_ppcmds
end
module Projection :
sig
- type t = Names.Projection.t
+ type t
val make : Constant.t -> bool -> t
val map : (Constant.t -> Constant.t) -> t -> t
val constant : t -> Constant.t
val equal : t -> t -> bool
end
+ type evaluable_global_reference =
+ | EvalVarRef of Id.t
+ | EvalConstRef of Constant.t
+
type inductive = MutInd.t * int
val eq_ind : inductive -> inductive -> bool
@@ -273,25 +182,31 @@ sig
val eq_constructor : constructor -> constructor -> bool
val constructor_hash : constructor -> int
- module MPset : module type of struct include Names.MPset end
- module MPmap : module type of struct include Names.MPmap end
- module KNset : module type of struct include Names.KNset end
- module KNmap : module type of struct include Names.KNmap end
- module Cset : module type of struct include Names.Cset end
- module Cset_env : module type of struct include Names.Cset_env end
- module Cmap : module type of struct include Names.Cmap end
- module Cmap_env : module type of struct include Names.Cmap_env end
- module Cpred : module type of struct include Names.Cpred end
- module Mindset : module type of struct include Names.Mindset end
- module Mindmap : module type of struct include Names.Mindmap end
- module Mindmap_env : module type of struct include Names.Mindmap_env end
- module Indmap : module type of struct include Names.Indmap end
- with type key = inductive
- module Indmap_env : module type of struct include Names.Indmap_env end
- module Constrmap : module type of struct include Names.Constrmap end
- module Constrmap_env : module type of struct include Names.Constrmap_env end
+ module MPset : Set.S with type elt = ModPath.t
+ module MPmap : Map.ExtS with type key = ModPath.t and module Set := MPset
+
+ module KNset : CSig.SetS with type elt = KerName.t
+ module KNpred : Predicate.S with type elt = KerName.t
+ module KNmap : Map.ExtS with type key = KerName.t and module Set := KNset
+
+ module Cpred : Predicate.S with type elt = Constant.t
+ module Cset : CSig.SetS with type elt = Constant.t
+ module Cset_env : CSig.SetS with type elt = Constant.t
+
+ module Cmap : Map.ExtS with type key = Constant.t and module Set := Cset
+ module Cmap_env : Map.ExtS with type key = Constant.t and module Set := Cset_env
+
+ module Mindset : CSig.SetS with type elt = MutInd.t
+ module Mindmap : Map.ExtS with type key = MutInd.t and module Set := Mindset
+ module Mindmap_env : CSig.MapS with type key = MutInd.t
+
+ module Indmap : CSig.MapS with type key = inductive
+ module Constrmap : CSig.MapS with type key = constructor
+ module Indmap_env : CSig.MapS with type key = inductive
+ module Constrmap_env : CSig.MapS with type key = constructor
type transparent_state = Id.Pred.t * Cpred.t
+
val empty_transparent_state : transparent_state
val full_transparent_state : transparent_state
val var_full_transparent_state : transparent_state
@@ -310,9 +225,8 @@ sig
[@@ocaml.deprecated "alias of API.Names.ModPath.t"]
type variable = Id.t
- [@@ocaml.deprecated "alias of API.Names.Id.t"]
- type 'a tableKey = 'a Names.tableKey =
+ type 'a tableKey =
| ConstKey of 'a
| VarKey of Id.t
| RelKey of Int.t
@@ -397,23 +311,297 @@ sig
val debug_string_of_con : Constant.t -> string
- module Idset : module type of struct include Id.Set end
+ type identifier = Id.t
+ module Idset : Set.S with type elt = identifier and type t = Id.Set.t
+
end
-module Context :
+module Univ :
sig
+ module Level :
+ sig
+ type t
+ val set : t
+ val pr : t -> Pp.std_ppcmds
+ end
+
+ type universe_level = Level.t
+
+ module LSet :
+ sig
+ include CSig.SetS with type elt = universe_level
+ val pr : (Level.t -> Pp.std_ppcmds) -> t -> Pp.std_ppcmds
+ end
+
+ module Universe :
+ sig
+ type t
+ val pr : t -> Pp.std_ppcmds
+ end
+
+ type universe = Universe.t
+
+ module Instance :
+ sig
+ type t
+ val empty : t
+ val of_array : Level.t array -> t
+ val to_array : t -> Level.t array
+ val pr : (Level.t -> Pp.std_ppcmds) -> t -> Pp.std_ppcmds
+ end
+
+ type 'a puniverses = 'a * Instance.t
+
+ val out_punivs : 'a puniverses -> 'a
+
+ type constraint_type = Lt | Le | Eq
+
+ type univ_constraint = universe_level * constraint_type * universe_level
+
+ module Constraint : sig
+ include Set.S with type elt = univ_constraint
+ end
+
+ type 'a constrained = 'a * Constraint.t
+
+ module UContext :
+ sig
+ type t
+ val empty : t
+ end
+
+ type universe_context = UContext.t
+
+ module AUContext :
+ sig
+ type t
+ val empty : t
+ end
+
+ type abstract_universe_context = AUContext.t
+
+ module CumulativityInfo :
+ sig
+ type t
+ end
+
+ type cumulativity_info = CumulativityInfo.t
+
+ module ACumulativityInfo :
+ sig
+ type t
+ end
+ type abstract_cumulativity_info = ACumulativityInfo.t
+
+ module ContextSet :
+ sig
+ type t
+ val empty : t
+ val of_context : UContext.t -> t
+ val to_context : t -> UContext.t
+ end
+
+ type 'a in_universe_context_set = 'a * ContextSet.t
+ type 'a in_universe_context = 'a * UContext.t
+
+ type universe_context_set = ContextSet.t
+
+ type universe_set = LSet.t
+
+ type 'a constraint_function = 'a -> 'a -> Constraint.t -> Constraint.t
+
+ module LMap :
+ sig
+ include CMap.ExtS with type key = universe_level and module Set := LSet
+
+ val union : 'a t -> 'a t -> 'a t
+ val diff : 'a t -> 'a t -> 'a t
+ val subst_union : 'a option t -> 'a option t -> 'a option t
+ val pr : ('a -> Pp.std_ppcmds) -> 'a t -> Pp.std_ppcmds
+ end
+
+ type 'a universe_map = 'a LMap.t
+ type universe_subst = universe universe_map
+ type universe_level_subst = universe_level universe_map
+
+ val enforce_leq : Universe.t constraint_function
+ val pr_uni : Universe.t -> Pp.std_ppcmds
+ val pr_universe_context : (Level.t -> Pp.std_ppcmds) -> UContext.t -> Pp.std_ppcmds
+ val pr_universe_context_set : (Level.t -> Pp.std_ppcmds) -> ContextSet.t -> Pp.std_ppcmds
+ val pr_universe_subst : universe_subst -> Pp.std_ppcmds
+ val pr_universe_level_subst : universe_level_subst -> Pp.std_ppcmds
+ val pr_constraints : (Level.t -> Pp.std_ppcmds) -> Constraint.t -> Pp.std_ppcmds
+end
+
+module UGraph :
+sig
+ type t
+ val pr_universes : (Univ.Level.t -> Pp.std_ppcmds) -> t -> Pp.std_ppcmds
+end
+
+module Esubst :
+sig
+ type 'a subs
+ val subs_id : int -> 'a subs
+end
+
+module Sorts :
+sig
+ type contents = Pos | Null
+ type t =
+ | Prop of contents
+ | Type of Univ.Universe.t
+ val is_prop : t -> bool
+ val hash : t -> int
+
+ type family = InProp | InSet | InType
+ val family : t -> family
+end
+
+module Evar :
+sig
+ (** Unique identifier of some {i evar} *)
+ type t
+
+ (** Recover the underlying integer. *)
+ val repr : t -> int
+
+ val equal : t -> t -> bool
+
+ (** a set of unique identifiers of some {i evars} *)
+ module Set : Set.S with type elt = t
+ module Map : CMap.ExtS with type key = t and module Set := Set
+
+end
+
+module Constr :
+sig
+ open Names
+
+ type t
+
+ type constr = t
+ type types = t
+
+ type cast_kind =
+ | VMcast
+ | NATIVEcast
+ | DEFAULTcast
+ | REVERTcast
+
+ type metavariable = int
+
+ type existential_key = Evar.t
+ type 'constr pexistential = existential_key * 'constr array
+
+ type 'a puniverses = 'a Univ.puniverses
+ type pconstant = Constant.t puniverses
+ type pinductive = inductive puniverses
+ type pconstructor = constructor puniverses
+
+ type ('constr, 'types) prec_declaration =
+ Name.t array * 'types array * 'constr array
+
+ type ('constr, 'types) pfixpoint =
+ (int array * int) * ('constr, 'types) prec_declaration
+
+ type ('constr, 'types) pcofixpoint =
+ int * ('constr, 'types) prec_declaration
+
+ 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 case_info =
+ { ci_ind : inductive; (* inductive type to which belongs the value that is being matched *)
+ ci_npar : int; (* number of parameters of the above inductive type *)
+ ci_cstr_ndecls : int array; (* For each constructor, the corresponding integer determines
+ the number of values that can be bound in a match-construct.
+ NOTE: parameters of the inductive type are therefore excluded from the count *)
+ ci_cstr_nargs : int array; (* for each constructor, the corresponding integers determines
+ the number of values that can be applied to the constructor,
+ in addition to the parameters of the related inductive type
+ NOTE: "lets" are therefore excluded from the count
+ NOTE: parameters of the inductive type are also excluded from the count *)
+ ci_pp_info : case_printing (* not interpreted by the kernel *)
+ }
+
+ type ('constr, 'types, 'sort, 'univs) kind_of_term =
+ | Rel of int
+ | Var of Id.t
+ | Meta of metavariable
+ | Evar of 'constr pexistential
+ | Sort of 'sort
+ | Cast of 'constr * cast_kind * 'types
+ | Prod of Name.t * 'types * 'types
+ | Lambda of Name.t * 'types * 'constr
+ | LetIn of Name.t * 'constr * 'types * 'constr
+ | App of 'constr * 'constr array
+ | Const of (Constant.t * 'univs)
+ | Ind of (inductive * 'univs)
+ | Construct of (constructor * 'univs)
+ | Case of case_info * 'constr * 'constr * 'constr array
+ | Fix of ('constr, 'types) pfixpoint
+ | CoFix of ('constr, 'types) pcofixpoint
+ | Proj of Projection.t * 'constr
+
+ val equal : t -> t -> bool
+ val eq_constr_nounivs : t -> t -> bool
+ val compare : t -> t -> int
+
+ val hash : t -> int
+
+ val mkRel : int -> t
+ val mkVar : Id.t -> t
+ val mkMeta : metavariable -> t
+ type existential = existential_key * constr array
+ val mkEvar : existential -> t
+ val mkSort : Sorts.t -> t
+ val mkProp : t
+ val mkSet : t
+ val mkType : Univ.Universe.t -> t
+ val mkCast : t * cast_kind * t -> t
+ val mkProd : Name.t * types * types -> types
+ val mkLambda : Name.t * types * t -> t
+ val mkLetIn : Name.t * t * types * t -> t
+ val mkApp : t * t array -> t
+ val map_puniverses : ('a -> 'b) -> 'a puniverses -> 'b puniverses
+
+ val mkConst : Constant.t -> t
+ val mkConstU : pconstant -> t
+
+ val mkProj : (Projection.t * t) -> t
+
+ val mkInd : inductive -> t
+ val mkIndU : pinductive -> t
+
+ val mkConstruct : constructor -> t
+ val mkConstructU : pconstructor -> t
+ val mkConstructUi : pinductive * int -> t
+
+ val mkCase : case_info * t * t * t array -> t
+
+end
+
+module Context :
+sig
module Rel :
sig
module Declaration :
sig
(* local declaration *)
(* local declaration *)
- type ('constr, 'types) pt = ('constr, 'types) Context.Rel.Declaration.pt =
+ type ('constr, 'types) pt =
| LocalAssum of Names.Name.t * 'types (** name, type *)
| LocalDef of Names.Name.t * 'constr * 'types (** name, value, type *)
- type t = (Prelude.constr, Prelude.types) pt
+ type t = (Constr.constr, Constr.types) pt
(** Return the name bound by a given declaration. *)
val get_name : ('c, 't) pt -> Names.Name.t
@@ -450,7 +638,7 @@ sig
val map_constr : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt
(** Perform a given action on all terms in a given declaration. *)
- val iter_constr : ('c -> unit) -> ('c, 'c) pt -> unit
+ val iter_constr : ('c -> unit) -> ('c, 'c) pt -> unit
(** Reduce all terms in a given declaration to a single value. *)
val fold_constr : ('c -> 'a -> 'a) -> ('c, 'c) pt -> 'a -> 'a
@@ -503,11 +691,11 @@ sig
module Declaration :
sig
(** local declaration *)
- type ('constr, 'types) pt = ('constr, 'types) Context.Named.Declaration.pt =
+ type ('constr, 'types) pt =
| LocalAssum of Names.Id.t * 'types (** identifier, type *)
| LocalDef of Names.Id.t * 'constr * 'types (** identifier, value, type *)
- type t = (Prelude.constr, Prelude.types) pt
+ type t = (Constr.constr, Constr.types) pt
(** Return the identifier bound by a given declaration. *)
val get_id : ('c, 't) pt -> Names.Id.t
@@ -604,12 +792,32 @@ sig
end
end
+module Vars :
+sig
+ type substl = Constr.t list
+
+ val substl : substl -> Constr.t -> Constr.t
+
+ val subst1 : Constr.t -> Constr.t -> Constr.t
+
+ val lift : int -> Constr.t -> Constr.t
+
+ val closed0 : Constr.t -> bool
+
+ val closedn : int -> Constr.t -> bool
+
+ val replace_vars : (Names.Id.t * Constr.t) list -> Constr.t -> Constr.t
+
+ val noccurn : int -> Constr.t -> bool
+ val subst_var : Names.Id.t -> Constr.t -> Constr.t
+ val subst_vars : Names.Id.t list -> Constr.t -> Constr.t
+ val substnl : substl -> int -> Constr.t -> Constr.t
+end
+
module Term :
sig
- type sorts_family = Sorts.family = InProp | InSet | InType
- [@@deprecated "alias of API.Sorts.family"]
- type metavariable = Prelude.metavariable
+ type sorts_family = Sorts.family = InProp | InSet | InType
type contents = Sorts.contents = Pos | Null
@@ -618,45 +826,55 @@ sig
| Type of Univ.Universe.t
[@@ocaml.deprecated "alias of API.Sorts.t"]
- type constr = Prelude.constr
- type types = Prelude.types
+ type constr = Constr.t
+ type types = Constr.t
+
+ type metavariable = int
+
type ('constr, 'types) prec_declaration = Names.Name.t array * 'types array * 'constr array
- type 'constr pexistential = 'constr Prelude.pexistential
- type cast_kind = Term.cast_kind =
+
+ type 'constr pexistential = 'constr Constr.pexistential
+ type cast_kind = Constr.cast_kind =
| VMcast
| NATIVEcast
| DEFAULTcast
| REVERTcast
+
type 'a puniverses = 'a Univ.puniverses
type pconstant = Names.Constant.t puniverses
type pinductive = Names.inductive puniverses
type pconstructor = Names.constructor puniverses
- type case_style = Term.case_style =
- | LetStyle
- | IfStyle
- | LetPatternStyle
- | MatchStyle
- | RegularStyle
- type case_printing = Term.case_printing =
- { ind_tags : bool list;
- cstr_tags : bool list array;
- style : case_style
- }
- type case_info = Term.case_info =
- { ci_ind : Names.inductive;
- ci_npar : int;
- ci_cstr_ndecls : int array;
- ci_cstr_nargs : int array;
- ci_pp_info : case_printing
- }
+ type case_style = Constr.case_style =
+ | LetStyle
+ | IfStyle
+ | LetPatternStyle
+ | MatchStyle
+ | RegularStyle
+
+ type case_printing = Constr.case_printing =
+ { ind_tags : bool list;
+ cstr_tags : bool list array;
+ style : case_style
+ }
+
+ type case_info = Constr.case_info =
+ { ci_ind : Names.inductive;
+ ci_npar : int;
+ ci_cstr_ndecls: int array;
+ ci_cstr_nargs : int array;
+ ci_pp_info : case_printing
+ }
+
type ('constr, 'types) pfixpoint =
(int array * int) * ('constr, 'types) prec_declaration
+
type ('constr, 'types) pcofixpoint =
int * ('constr, 'types) prec_declaration
- type ('constr, 'types, 'sort, 'univs) kind_of_term = ('constr, 'types, 'sort, 'univs) Term.kind_of_term =
+
+ type ('constr, 'types, 'sort, 'univs) kind_of_term = ('constr, 'types, 'sort, 'univs) Constr.kind_of_term =
| Rel of int
| Var of Names.Id.t
- | Meta of metavariable
+ | Meta of Constr.metavariable
| Evar of 'constr pexistential
| Sort of 'sort
| Cast of 'constr * cast_kind * 'types
@@ -671,7 +889,7 @@ sig
| Fix of ('constr, 'types) pfixpoint
| CoFix of ('constr, 'types) pcofixpoint
| Proj of Names.Projection.t * 'constr
- type existential = Prelude.evar * constr array
+ type existential = Constr.existential_key * constr array
type rec_declaration = Names.Name.t array * constr array * constr array
type fixpoint = (int array * int) * rec_declaration
type cofixpoint = int * rec_declaration
@@ -685,7 +903,7 @@ sig
val mkRel : int -> constr
val mkVar : Names.Id.t -> constr
- val mkMeta : Prelude.metavariable -> constr
+ val mkMeta : Constr.metavariable -> constr
val mkEvar : existential -> constr
val mkSort : Sorts.t -> types
@@ -765,18 +983,19 @@ sig
*)
val eq_constr_nounivs : constr -> constr -> bool
- type ('constr, 'types) kind_of_type = ('constr, 'types) Term.kind_of_type =
- | SortType of Sorts.t
- | CastType of 'types * 'types
- | ProdType of Names.Name.t * 'types * 'types
- | LetInType of Names.Name.t * 'constr * 'types * 'types
- | AtomicType of 'constr * 'constr array
+ type ('constr, 'types) kind_of_type =
+ | SortType of Sorts.t
+ | CastType of 'types * 'types
+ | ProdType of Names.Name.t * 'types * 'types
+ | LetInType of Names.Name.t * 'constr * 'types * 'types
+ | AtomicType of 'constr * 'constr array
+
val kind_of_type : types -> (constr, types) kind_of_type
val is_prop_sort : Sorts.t -> bool
[@@ocaml.deprecated "alias of API.Sorts.is_prop"]
- type existential_key = Prelude.evar
+ type existential_key = Constr.existential_key
val family_of_sort : Sorts.t -> Sorts.family
@@ -797,181 +1016,13 @@ sig
val compare_constr : (constr -> constr -> bool) -> constr -> constr -> bool
end
-module EConstr :
-sig
- type t = EConstr.t
- type constr = t
- type types = t
- type unsafe_judgment = EConstr.unsafe_judgment
- type named_declaration = (constr, types) Context.Named.Declaration.pt
- type named_context = (constr, types) Context.Named.pt
- type rel_context = (constr, types) Context.Rel.pt
- type rel_declaration = (constr, types) Context.Rel.Declaration.pt
- type existential = constr Term.pexistential
- module ESorts :
- sig
- type t = EConstr.ESorts.t
- (** Type of sorts up-to universe unification. Essentially a wrapper around
- Sorts.t so that normalization is ensured statically. *)
-
- val make : Sorts.t -> t
- (** Turn a sort into an up-to sort. *)
-
- val kind : Prelude.evar_map -> t -> Sorts.t
- (** Returns the view into the current sort. Note that the kind of a variable
- may change if the unification state of the evar map changes. *)
-
- end
-
- module EInstance :
- sig
- type t = EConstr.EInstance.t
- (** Type of universe instances up-to universe unification. Similar to
- {ESorts.t} for {Univ.Instance.t}. *)
-
- val make : Univ.Instance.t -> t
- val kind : Prelude.evar_map -> t -> Univ.Instance.t
- val empty : t
- val is_empty : t -> bool
- end
-
- val of_constr : Term.constr -> constr
-
- val kind : Prelude.evar_map -> constr -> (constr, constr, ESorts.t, EInstance.t) Term.kind_of_term
-
- val mkArrow : constr -> constr -> constr
- val mkInd : Names.inductive -> t
- val mkProp : constr
- val mkProd : Names.Name.t * constr * constr -> constr
- val mkRel : int -> constr
- val mkSort : Sorts.t -> constr
- val mkVar : Names.Id.t -> constr
- val mkLambda : Names.Name.t * constr * constr -> constr
- val mkLambda_or_LetIn : rel_declaration -> constr -> constr
- val mkApp : constr * constr array -> constr
- val mkEvar : constr Term.pexistential -> constr
-
- val mkMeta : Prelude.metavariable -> constr
-
- val mkConstructU : Names.constructor * EInstance.t -> constr
- val mkLetIn : Names.Name.t * constr * constr * constr -> constr
- val mkProd_or_LetIn : rel_declaration -> constr -> constr
- val mkCast : constr * Term.cast_kind * constr -> constr
- val mkNamedLambda : Names.Id.t -> types -> constr -> constr
- val mkNamedProd : Names.Id.t -> types -> types -> types
-
- val isCast : Evd.evar_map -> t -> bool
- val isEvar : Prelude.evar_map -> constr -> bool
- val isInd : Prelude.evar_map -> constr -> bool
- val isRel : Prelude.evar_map -> constr -> bool
- val isSort : Prelude.evar_map -> constr -> bool
- val isVar : Prelude.evar_map -> constr -> bool
- val isConst : Prelude.evar_map -> constr -> bool
- val isConstruct : Prelude.evar_map -> constr -> bool
-
- val destInd : Prelude.evar_map -> constr -> Names.inductive * EInstance.t
- val destVar : Prelude.evar_map -> constr -> Names.Id.t
- val destEvar : Prelude.evar_map -> constr -> constr Term.pexistential
- val destRel : Prelude.evar_map -> constr -> int
- val destProd : Prelude.evar_map -> constr -> Names.Name.t * types * types
- val destLambda : Prelude.evar_map -> constr -> Names.Name.t * types * constr
- val destApp : Prelude.evar_map -> constr -> constr * constr array
- val destConst : Prelude.evar_map -> constr -> Names.Constant.t * EInstance.t
- val destConstruct : Prelude.evar_map -> constr -> Names.constructor * EInstance.t
- val destFix : Evd.evar_map -> t -> (t, t) Term.pfixpoint
- val destCast : Evd.evar_map -> t -> t * Term.cast_kind * t
-
- val mkConstruct : Names.constructor -> constr
-
- val compose_lam : (Names.Name.t * constr) list -> constr -> constr
-
- val decompose_lam : Prelude.evar_map -> constr -> (Names.Name.t * constr) list * constr
- val decompose_lam_n_assum : Prelude.evar_map -> int -> constr -> rel_context * constr
- val decompose_app : Prelude.evar_map -> constr -> constr * constr list
- val decompose_prod : Prelude.evar_map -> constr -> (Names.Name.t * constr) list * constr
- val decompose_prod_assum : Prelude.evar_map -> constr -> rel_context * constr
-
- val applist : constr * constr list -> constr
-
- val to_constr : Prelude.evar_map -> constr -> Constr.t
-
- val push_rel : rel_declaration -> Prelude.env -> Prelude.env
-
- module Unsafe :
- sig
- val to_constr : constr -> Term.constr
-
- val to_rel_decl : (constr, types) Context.Rel.Declaration.pt -> (Prelude.constr, Prelude.types) Context.Rel.Declaration.pt
-
- (** Physical identity. Does not care for defined evars. *)
-
- val to_named_decl : (constr, types) Context.Named.Declaration.pt -> (Prelude.constr, Prelude.types) Context.Named.Declaration.pt
-
- val to_instance : EInstance.t -> Univ.Instance.t
- end
-
- module Vars :
- sig
- val substnl : t list -> int -> t -> t
- val noccurn : Prelude.evar_map -> int -> constr -> bool
- val closed0 : Prelude.evar_map -> constr -> bool
- val subst1 : constr -> constr -> constr
- val substl : constr list -> constr -> constr
- val lift : int -> constr -> constr
- val liftn : int -> int -> t -> t
- val subst_var : Names.Id.t -> t -> t
- val subst_vars : Names.Id.t list -> t -> t
- end
-
- val fresh_global :
- ?loc:Loc.t -> ?rigid:Prelude.rigid -> ?names:Univ.Instance.t -> Environ.env ->
- Evd.evar_map -> Prelude.global_reference -> Evd.evar_map * t
-
-val of_named_decl : (Term.constr, Term.types) Context.Named.Declaration.pt -> (constr, types) Context.Named.Declaration.pt
- val of_rel_decl : (Term.constr, Term.types) Context.Rel.Declaration.pt -> (constr, types) Context.Rel.Declaration.pt
- val kind_of_type : Prelude.evar_map -> constr -> (constr, constr) Term.kind_of_type
- val to_lambda : Prelude.evar_map -> int -> constr -> constr
- val it_mkLambda_or_LetIn : constr -> rel_context -> constr
- val push_rel_context : rel_context -> Prelude.env -> Prelude.env
- val eq_constr : Prelude.evar_map -> constr -> constr -> bool
- val iter_with_binders : Prelude.evar_map -> ('a -> 'a) -> ('a -> constr -> unit) -> 'a -> constr -> unit
- val fold : Prelude.evar_map -> ('a -> constr -> 'a) -> 'a -> constr -> 'a
- val existential_type : Prelude.evar_map -> existential -> types
- val iter : Prelude.evar_map -> (constr -> unit) -> constr -> unit
- val eq_constr_universes : Prelude.evar_map -> constr -> constr -> Universes.universe_constraints option
- val eq_constr_nounivs : Prelude.evar_map -> constr -> constr -> bool
- val compare_constr : Evd.evar_map -> (constr -> constr -> bool) -> constr -> constr -> bool
- val isApp : Prelude.evar_map -> constr -> bool
- val it_mkProd_or_LetIn : constr -> rel_context -> constr
- val push_named : named_declaration -> Prelude.env -> Prelude.env
- val destCase : Prelude.evar_map -> constr -> Term.case_info * constr * constr * constr array
- val decompose_lam_assum : Prelude.evar_map -> constr -> rel_context * constr
- val mkConst : Names.Constant.t -> constr
- val mkCase : Term.case_info * constr * constr * constr array -> constr
- val named_context : Prelude.env -> named_context
- val val_of_named_context : named_context -> Prelude.named_context_val
- val mkFix : (t, t) Term.pfixpoint -> t
- val decompose_prod_n_assum : Evd.evar_map -> int -> t -> rel_context * t
- val isMeta : Evd.evar_map -> t -> bool
-
- val destMeta : Evd.evar_map -> t -> Term.metavariable
-
- val map_with_binders : Evd.evar_map -> ('a -> 'a) -> ('a -> t -> t) -> 'a -> t -> t
- val mkNamedLetIn : Names.Id.t -> constr -> types -> constr -> constr
- val map : Evd.evar_map -> (t -> t) -> t -> t
- val mkConstU : Names.Constant.t * EInstance.t -> t
- val isProd : Evd.evar_map -> t -> bool
- val mkConstructUi : (Names.inductive * EInstance.t) * int -> t
- val isLambda : Evd.evar_map -> t -> bool
-end
-
module Mod_subst :
sig
- type substitution = Mod_subst.substitution
- type 'a substituted = 'a Mod_subst.substituted
- type delta_resolver = Mod_subst.delta_resolver
+ type delta_resolver
+ type substitution
+ type 'a substituted
- val force_constr : Term.constr substituted -> Term.constr
+ val force_constr : Constr.t substituted -> Constr.t
val empty_delta_resolver : delta_resolver
val constant_of_delta_kn : delta_resolver -> Names.KerName.t -> Names.Constant.t
@@ -979,21 +1030,94 @@ sig
val subst_kn : substitution -> Names.KerName.t -> Names.KerName.t
val subst_evaluable_reference :
substitution -> Names.evaluable_global_reference -> Names.evaluable_global_reference
- val subst_mps : substitution -> Term.constr -> Term.constr
+ val subst_mps : substitution -> Constr.t -> Constr.t
val subst_constant : substitution -> Names.Constant.t -> Names.Constant.t
val subst_ind : substitution -> Names.inductive -> Names.inductive
val debug_pr_subst : substitution -> Pp.std_ppcmds
val debug_pr_delta : delta_resolver -> Pp.std_ppcmds
end
+module Opaqueproof :
+sig
+ type opaquetab
+ type opaque
+ val empty_opaquetab : opaquetab
+ val force_proof : opaquetab -> opaque -> Constr.t
+end
+
+module Cbytecodes :
+sig
+ type tag = int
+ type reloc_table = (tag * int) array
+end
+
+module Cemitcodes :
+sig
+ type to_patch_substituted
+end
+
+module Decl_kinds :
+sig
+ type polymorphic = bool
+ type cumulative_inductive_flag = bool
+ type recursivity_kind =
+ | Finite
+ | CoFinite
+ | BiFinite
+
+ type locality =
+ | Discharge
+ | Local
+ | Global
+
+ type definition_object_kind =
+ | Definition
+ | Coercion
+ | SubClass
+ | CanonicalStructure
+ | Example
+ | Fixpoint
+ | CoFixpoint
+ | Scheme
+ | StructureComponent
+ | IdentityCoercion
+ | Instance
+ | Method
+ type theorem_kind =
+ | Theorem
+ | Lemma
+ | Fact
+ | Remark
+ | Property
+ | Proposition
+ | Corollary
+ type goal_object_kind =
+ | DefinitionBody of definition_object_kind
+ | Proof of theorem_kind
+ type goal_kind = locality * polymorphic * goal_object_kind
+ type assumption_object_kind =
+ | Definitional
+ | Logical
+ | Conjectural
+ type logical_kind =
+ | IsAssumption of assumption_object_kind
+ | IsDefinition of definition_object_kind
+ | IsProof of theorem_kind
+ type binding_kind =
+ | Explicit
+ | Implicit
+ type private_flag = bool
+ type definition_kind = locality * polymorphic * definition_object_kind
+end
+
module Retroknowledge :
sig
- type action = Retroknowledge.action
- type nat_field = Retroknowledge.nat_field =
+ type action
+ type nat_field =
| NatType
| NatPlus
| NatTimes
- type n_field = Retroknowledge.n_field =
+ type n_field =
| NPositive
| NType
| NTwice
@@ -1002,7 +1126,7 @@ sig
| NPhiInv
| NPlus
| NTimes
- type int31_field = Retroknowledge.int31_field =
+ type int31_field =
| Int31Bits
| Int31Type
| Int31Constructor
@@ -1028,44 +1152,60 @@ sig
| Int31Lor
| Int31Land
| Int31Lxor
- type field = Retroknowledge.field =
+ type field =
| KInt31 of string * int31_field
end
+module Conv_oracle :
+sig
+ type level
+end
+
module Declarations :
sig
- type recarg = Declarations.recarg =
+
+ open Names
+
+ type recarg =
| Norec
| Mrec of Names.inductive
| Imbr of Names.inductive
type wf_paths = recarg Rtree.t
- type inline = Declarations.inline
- type constant_def = Declarations.constant_def =
+ type inline = int option
+ type constant_def =
| Undef of inline
- | Def of Term.constr Mod_subst.substituted
+ | Def of Constr.t Mod_subst.substituted
| OpaqueDef of Opaqueproof.opaque
- type template_arity = Declarations.template_arity = {
+ type template_arity = {
template_param_levels : Univ.Level.t option list;
template_level : Univ.Universe.t;
}
- type ('a, 'b) declaration_arity = ('a, 'b) Declarations.declaration_arity =
+ type ('a, 'b) declaration_arity =
| RegularArity of 'a
| TemplateArity of 'b
- type constant_type = (Prelude.types, Context.Rel.t * template_arity) declaration_arity
- type constant_universes = Declarations.constant_universes
- type projection_body = Declarations.projection_body = {
+ type constant_type = (Constr.types, Context.Rel.t * template_arity) declaration_arity
+
+ type constant_universes =
+ | Monomorphic_const of Univ.universe_context
+ | Polymorphic_const of Univ.abstract_universe_context
+
+ type projection_body = {
proj_ind : Names.MutInd.t;
proj_npars : int;
proj_arg : int;
- proj_type : Term.types;
- proj_eta : Term.constr * Term.types;
- proj_body : Term.constr;
+ proj_type : Constr.types;
+ proj_eta : Constr.t * Constr.types;
+ proj_body : Constr.t;
}
- type typing_flags = Declarations.typing_flags
- type constant_body = Declarations.constant_body = {
+ type typing_flags = {
+ check_guarded : bool;
+ check_universes : bool;
+ }
+
+ type constant_body = {
const_hyps : Context.Named.t;
const_body : constant_def;
const_type : constant_type;
@@ -1075,16 +1215,24 @@ sig
const_inline_code : bool;
const_typing_flags : typing_flags;
}
- type one_inductive_body = Declarations.one_inductive_body = {
+
+ type regular_inductive_arity = {
+ mind_user_arity : Constr.types;
+ mind_sort : Sorts.t;
+ }
+
+ type inductive_arity = (regular_inductive_arity, template_arity) declaration_arity
+
+ type one_inductive_body = {
mind_typename : Names.Id.t;
mind_arity_ctxt : Context.Rel.t;
- mind_arity : Declarations.inductive_arity;
+ mind_arity : inductive_arity;
mind_consnames : Names.Id.t array;
- mind_user_lc : Term.types array;
+ mind_user_lc : Constr.types array;
mind_nrealargs : int;
mind_nrealdecls : int;
mind_kelim : Sorts.family list;
- mind_nf_lc : Term.types array;
+ mind_nf_lc : Constr.types array;
mind_consnrealargs : int array;
mind_consnrealdecls : int array;
mind_recargs : wf_paths;
@@ -1092,42 +1240,47 @@ sig
mind_nb_args : int;
mind_reloc_tbl : Cbytecodes.reloc_table;
}
- type ('ty,'a) functorize = ('ty,'a) Declarations.functorize =
+
+ type ('ty,'a) functorize =
| NoFunctor of 'a
| MoreFunctor of Names.MBId.t * 'ty * ('ty,'a) functorize
- type with_declaration = Declarations.with_declaration =
+
+ type with_declaration =
| WithMod of Names.Id.t list * Names.ModPath.t
- | WithDef of Names.Id.t list * Term.constr Univ.in_universe_context
- type module_alg_expr = Declarations.module_alg_expr =
+ | WithDef of Names.Id.t list * Constr.t Univ.in_universe_context
+
+ type module_alg_expr =
| MEident of Names.ModPath.t
| MEapply of module_alg_expr * Names.ModPath.t
| MEwith of module_alg_expr * with_declaration
- type abstract_inductive_universes = Declarations.abstract_inductive_universes =
- | Monomorphic_ind of Univ.UContext.t
- | Polymorphic_ind of Univ.abstract_universe_context
- | Cumulative_ind of Univ.abstract_cumulativity_info
+ type abstract_inductive_universes =
+ | Monomorphic_ind of Univ.universe_context
+ | Polymorphic_ind of Univ.abstract_universe_context
+ | Cumulative_ind of Univ.abstract_cumulativity_info
- type mutual_inductive_body = Declarations.mutual_inductive_body = {
+ type record_body = (Id.t * Constant.t array * projection_body array) option
+
+ type mutual_inductive_body = {
mind_packets : one_inductive_body array;
- mind_record : Declarations.record_body option;
+ mind_record : record_body option;
mind_finite : Decl_kinds.recursivity_kind;
mind_ntypes : int;
mind_hyps : Context.Named.t;
mind_nparams : int;
mind_nparams_rec : int;
mind_params_ctxt : Context.Rel.t;
- mind_universes : Declarations.abstract_inductive_universes;
+ mind_universes : abstract_inductive_universes;
mind_private : bool option;
- mind_typing_flags : Declarations.typing_flags;
+ mind_typing_flags : typing_flags;
}
and module_expression = (module_type_body,module_alg_expr) functorize
- and module_implementation = Declarations.module_implementation =
+ and module_implementation =
| Abstract
| Algebraic of module_expression
| Struct of module_signature
| FullStruct
- and module_body = Declarations.module_body =
+ and module_body =
{ mod_mp : Names.ModPath.t;
mod_expr : module_implementation;
mod_type : module_signature;
@@ -1139,28 +1292,97 @@ sig
and module_signature = (module_type_body,structure_body) functorize
and module_type_body = module_body
and structure_body = (Names.Label.t * structure_field_body) list
- and structure_field_body = Declarations.structure_field_body =
+ and structure_field_body =
| SFBconst of constant_body
| SFBmind of mutual_inductive_body
| SFBmodule of module_body
| SFBmodtype of module_type_body
end
-module Univops : sig
- val universes_of_constr : Term.constr -> Univ.LSet.t
- val restrict_universe_context : Univ.ContextSet.t -> Univ.LSet.t -> Univ.ContextSet.t
+module Declareops :
+sig
+ val constant_has_body : Declarations.constant_body -> bool
+ val is_opaque : Declarations.constant_body -> bool
+ val eq_recarg : Declarations.recarg -> Declarations.recarg -> bool
+end
+
+module Entries :
+sig
+
+ open Names
+ open Constr
+
+ type local_entry =
+ | LocalDefEntry of constr
+ | LocalAssumEntry of constr
+
+ type inductive_universes =
+ | Monomorphic_ind_entry of Univ.universe_context
+ | Polymorphic_ind_entry of Univ.universe_context
+ | Cumulative_ind_entry of Univ.cumulativity_info
+
+ type one_inductive_entry = {
+ mind_entry_typename : Id.t;
+ mind_entry_arity : constr;
+ mind_entry_template : bool; (* Use template polymorphism *)
+ mind_entry_consnames : Id.t list;
+ mind_entry_lc : constr list }
+
+ type mutual_inductive_entry = {
+ mind_entry_record : (Names.Id.t option) option;
+ (** Some (Some id): primitive record with id the binder name of the record
+ in projections.
+ Some None: non-primitive record *)
+ mind_entry_finite : Decl_kinds.recursivity_kind;
+ mind_entry_params : (Id.t * local_entry) list;
+ mind_entry_inds : one_inductive_entry list;
+ mind_entry_universes : inductive_universes;
+ (* universe constraints and the constraints for subtyping of
+ inductive types in the block. *)
+ mind_entry_private : bool option;
+ }
+
+ type inline = int option
+ type 'a proof_output = Constr.t Univ.in_universe_context_set * 'a
+ type 'a const_entry_body = 'a proof_output Future.computation
+ type 'a definition_entry =
+ { const_entry_body : 'a const_entry_body;
+ (* List of section variables *)
+ const_entry_secctx : Context.Named.t option;
+ (* State id on which the completion of type checking is reported *)
+ const_entry_feedback : Stateid.t option;
+ const_entry_type : Constr.types option;
+ const_entry_polymorphic : bool;
+ const_entry_universes : Univ.UContext.t;
+ const_entry_opaque : bool;
+ const_entry_inline_code : bool }
+ type parameter_entry = Context.Named.t option * bool * Constr.types Univ.in_universe_context * inline
+
+ type projection_entry = {
+ proj_entry_ind : MutInd.t;
+ proj_entry_arg : int }
+
+ type 'a constant_entry =
+ | DefinitionEntry of 'a definition_entry
+ | ParameterEntry of parameter_entry
+ | ProjectionEntry of projection_entry
+ type module_struct_entry = Declarations.module_alg_expr
+ type module_params_entry =
+ (Names.MBId.t * module_struct_entry) list
+ type module_type_entry = module_params_entry * module_struct_entry
end
module Environ :
sig
- type env = Prelude.env
- type named_context_val = Prelude.named_context_val
- type ('constr, 'types) punsafe_judgment = ('constr, 'types) Environ.punsafe_judgment =
+ type env
+ type named_context_val
+
+ type ('constr, 'types) punsafe_judgment =
{
uj_val : 'constr;
uj_type : 'types
}
- type 'types punsafe_type_judgment = 'types Environ.punsafe_type_judgment = {
+ type 'types punsafe_type_judgment = {
utj_val : 'types;
utj_type : Sorts.t }
@@ -1172,7 +1394,7 @@ sig
val push_rec_types : Term.rec_declaration -> env -> env
val lookup_rel : int -> env -> Context.Rel.Declaration.t
val lookup_named : Names.Id.t -> env -> Context.Named.Declaration.t
- val lookup_named_val : Names.Id.t -> Environ.named_context_val -> Context.Named.Declaration.t
+ val lookup_named_val : Names.Id.t -> named_context_val -> Context.Named.Declaration.t
val lookup_constant : Names.Constant.t -> env -> Declarations.constant_body
val opaque_tables : env -> Opaqueproof.opaquetab
val is_projection : Names.Constant.t -> env -> bool
@@ -1184,63 +1406,170 @@ sig
val push_named_context_val : Context.Named.Declaration.t -> named_context_val -> named_context_val
val reset_with_named_context : named_context_val -> env -> env
val rel_context : env -> Context.Rel.t
- val constant_value_in : env -> Names.Constant.t Univ.puniverses -> Term.constr
- val named_type : Names.Id.t -> env -> Term.types
- val constant_opt_value_in : env -> Names.Constant.t Univ.puniverses -> Term.constr option
+ val constant_value_in : env -> Names.Constant.t Univ.puniverses -> Constr.t
+ val named_type : Names.Id.t -> env -> Constr.types
+ val constant_opt_value_in : env -> Names.Constant.t Univ.puniverses -> Constr.t option
val fold_named_context_reverse :
('a -> Context.Named.Declaration.t -> 'a) -> init:'a -> env -> 'a
- val evaluable_named : Names.Id.t -> Environ.env -> bool
+ val evaluable_named : Names.Id.t -> env -> bool
val push_context_set : ?strict:bool -> Univ.ContextSet.t -> env -> env
end
-module UGraph :
+module CClosure :
sig
- type t = UGraph.t
- val pr_universes : (Univ.Level.t -> Pp.std_ppcmds) -> t -> Pp.std_ppcmds
+
+ type table_key = Names.Constant.t Univ.puniverses Names.tableKey
+
+ type fconstr
+
+ type fterm =
+ | FRel of int
+ | FAtom of Constr.t (** Metas and Sorts *)
+ | FCast of fconstr * Constr.cast_kind * fconstr
+ | FFlex of table_key
+ | FInd of Names.inductive Univ.puniverses
+ | FConstruct of Names.constructor Univ.puniverses
+ | FApp of fconstr * fconstr array
+ | FProj of Names.Projection.t * fconstr
+ | FFix of Term.fixpoint * fconstr Esubst.subs
+ | FCoFix of Term.cofixpoint * fconstr Esubst.subs
+ | FCaseT of Term.case_info * Constr.t * fconstr * Constr.t array * fconstr Esubst.subs (* predicate and branches are closures *)
+ | FLambda of int * (Names.Name.t * Constr.t) list * Constr.t * fconstr Esubst.subs
+ | FProd of Names.Name.t * fconstr * fconstr
+ | FLetIn of Names.Name.t * fconstr * fconstr * Constr.t * fconstr Esubst.subs
+ | FEvar of Term.existential * fconstr Esubst.subs
+ | FLIFT of int * fconstr
+ | FCLOS of Constr.t * fconstr Esubst.subs
+ | FLOCKED
+
+ module RedFlags : sig
+ type reds
+ type red_kind
+ val mkflags : red_kind list -> reds
+ val fBETA : red_kind
+ val fCOFIX : red_kind
+ val fCONST : Names.Constant.t -> red_kind
+ val fFIX : red_kind
+ val fMATCH : red_kind
+ val fZETA : red_kind
+ val red_add_transparent : reds -> Names.transparent_state -> reds
+ end
+
+ type 'a infos_cache
+ type 'a infos = {
+ i_flags : RedFlags.reds;
+ i_cache : 'a infos_cache }
+
+ type clos_infos = fconstr infos
+
+ val mk_clos : fconstr Esubst.subs -> Constr.t -> fconstr
+ val mk_atom : Constr.t -> fconstr
+ val mk_clos_deep :
+ (fconstr Esubst.subs -> Constr.t -> fconstr) ->
+ fconstr Esubst.subs -> Constr.t -> fconstr
+ val mk_red : fterm -> fconstr
+ val all : RedFlags.reds
+ val beta : RedFlags.reds
+ val betaiota : RedFlags.reds
+ val betaiotazeta : RedFlags.reds
+
+ val create_clos_infos : ?evars:(Term.existential -> Constr.t option) -> RedFlags.reds -> Environ.env -> clos_infos
+
+ val whd_val : clos_infos -> fconstr -> Constr.t
+
+ val inject : Constr.t -> fconstr
+
+ val kl : clos_infos -> fconstr -> Constr.t
+ val term_of_fconstr : fconstr -> Constr.t
end
module Reduction :
sig
exception NotConvertible
- type conv_pb = Prelude.conv_pb =
+ type conv_pb =
| CONV
| CUMUL
- val whd_all : Environ.env -> Term.constr -> Term.constr
+ val whd_all : Environ.env -> Constr.t -> Constr.t
- val whd_betaiotazeta : Environ.env -> Term.constr -> Term.constr
+ val whd_betaiotazeta : Environ.env -> Constr.t -> Constr.t
val is_arity : Environ.env -> Term.types -> bool
val dest_prod : Environ.env -> Term.types -> Context.Rel.t * Term.types
- type 'a extended_conversion_function =
+ type 'a extended_conversion_function =
?l2r:bool -> ?reds:Names.transparent_state -> Environ.env ->
- ?evars:((Term.existential->Term.constr option) * UGraph.t) ->
+ ?evars:((Term.existential->Constr.t option) * UGraph.t) ->
'a -> 'a -> unit
- val conv : Term.constr extended_conversion_function
+ val conv : Constr.t extended_conversion_function
end
-module Vars :
+module Type_errors :
sig
- type substl = Term.constr list
-
- val substl : substl -> Term.constr -> Term.constr
-
- val subst1 : Term.constr -> Term.constr -> Term.constr
-
- val lift : int -> Term.constr -> Term.constr
- val closed0 : Term.constr -> bool
+ open Names
+ open Term
+ open Environ
+
+ type 'constr pguard_error =
+ (** Fixpoints *)
+ | NotEnoughAbstractionInFixBody
+ | RecursionNotOnInductiveType of 'constr
+ | RecursionOnIllegalTerm of int * (env * 'constr) * int list * int list
+ | NotEnoughArgumentsForFixCall of int
+ (** CoFixpoints *)
+ | CodomainNotInductiveType of 'constr
+ | NestedRecursiveOccurrences
+ | UnguardedRecursiveCall of 'constr
+ | RecCallInTypeOfAbstraction of 'constr
+ | RecCallInNonRecArgOfConstructor of 'constr
+ | RecCallInTypeOfDef of 'constr
+ | RecCallInCaseFun of 'constr
+ | RecCallInCaseArg of 'constr
+ | RecCallInCasePred of 'constr
+ | NotGuardedForm of 'constr
+ | ReturnPredicateNotCoInductive of 'constr
+
+ type arity_error =
+ | NonInformativeToInformative
+ | StrongEliminationOnNonSmallType
+ | WrongArity
+
+ type ('constr, 'types) ptype_error =
+ | UnboundRel of int
+ | UnboundVar of variable
+ | NotAType of ('constr, 'types) punsafe_judgment
+ | BadAssumption of ('constr, 'types) punsafe_judgment
+ | ReferenceVariables of identifier * 'constr
+ | ElimArity of pinductive * sorts_family list * 'constr * ('constr, 'types) punsafe_judgment
+ * (sorts_family * sorts_family * arity_error) option
+ | CaseNotInductive of ('constr, 'types) punsafe_judgment
+ | WrongCaseInfo of pinductive * case_info
+ | NumberBranches of ('constr, 'types) punsafe_judgment * int
+ | IllFormedBranch of 'constr * pconstructor * 'constr * 'constr
+ | Generalization of (Name.t * 'types) * ('constr, 'types) punsafe_judgment
+ | ActualType of ('constr, 'types) punsafe_judgment * 'types
+ | CantApplyBadType of
+ (int * 'constr * 'constr) * ('constr, 'types) punsafe_judgment * ('constr, 'types) punsafe_judgment array
+ | CantApplyNonFunctional of ('constr, 'types) punsafe_judgment * ('constr, 'types) punsafe_judgment array
+ | IllFormedRecBody of 'constr pguard_error * Name.t array * int * env * ('constr, 'types) punsafe_judgment array
+ | IllTypedRecBody of
+ int * Name.t array * ('constr, 'types) punsafe_judgment array * 'types array
+ | UnsatisfiedConstraints of Univ.Constraint.t
+
+ type type_error = (constr, types) ptype_error
- val closedn : int -> Term.constr -> bool
-
- val replace_vars : (Names.Id.t * Term.constr) list -> Term.constr -> Term.constr
+ exception TypeError of Environ.env * type_error
+end
- val noccurn : int -> Term.constr -> bool
- val subst_var : Names.Id.t -> Term.constr -> Term.constr
- val subst_vars : Names.Id.t list -> Term.constr -> Term.constr
- val substnl : substl -> int -> Term.constr -> Term.constr
+module Modops :
+sig
+ val destr_nofunctor : ('ty,'a) Declarations.functorize -> 'a
+ val add_structure :
+ Names.ModPath.t -> Declarations.structure_body -> Mod_subst.delta_resolver ->
+ Environ.env -> Environ.env
+ val add_module_type : Names.ModPath.t -> Declarations.module_type_body -> Environ.env -> Environ.env
end
module Inductive :
@@ -1249,7 +1578,7 @@ sig
val type_of_inductive : Environ.env -> mind_specif Univ.puniverses -> Term.types
exception SingletonInductiveBecomesProp of Names.Id.t
val lookup_mind_specif : Environ.env -> Names.inductive -> mind_specif
- val find_inductive : Environ.env -> Term.types -> Term.pinductive * Term.constr list
+ val find_inductive : Environ.env -> Term.types -> Term.pinductive * Constr.t list
end
module Typeops :
@@ -1259,144 +1588,606 @@ sig
val type_of_constant_in : Environ.env -> Term.pconstant -> Term.types
end
-module Opaqueproof :
+module Mod_typing :
sig
- type opaquetab = Opaqueproof.opaquetab
- type opaque = Opaqueproof.opaque
- val empty_opaquetab : opaquetab
- val force_proof : opaquetab -> opaque -> Term.constr
+ type 'alg translation =
+ Declarations.module_signature * 'alg * Mod_subst.delta_resolver * Univ.ContextSet.t
+ val translate_modtype :
+ Environ.env -> Names.ModPath.t -> Entries.inline ->
+ Entries.module_type_entry -> Declarations.module_type_body
+ val translate_mse :
+ Environ.env -> Names.ModPath.t option -> Entries.inline -> Declarations.module_alg_expr ->
+ Declarations.module_alg_expr translation
end
-module Modops :
+module Safe_typing :
sig
- val destr_nofunctor : ('ty,'a) Declarations.functorize -> 'a
- val add_structure :
- Names.ModPath.t -> Declarations.structure_body -> Mod_subst.delta_resolver ->
- Environ.env -> Environ.env
- val add_module_type : Names.ModPath.t -> Declarations.module_type_body -> Environ.env -> Environ.env
+ type private_constants
+ val mk_pure_proof : Constr.t -> private_constants Entries.proof_output
end
-module Entries :
+(************************************************************************)
+(* End of modules from kernel/ *)
+(************************************************************************)
+
+(************************************************************************)
+(* Modules from intf/ *)
+(************************************************************************)
+
+module Misctypes :
sig
- type mutual_inductive_entry = Entries.mutual_inductive_entry
- type inline = int option
- type 'a proof_output = Term.constr Univ.in_universe_context_set * 'a
- type 'a const_entry_body = 'a proof_output Future.computation
- type 'a definition_entry = 'a Entries.definition_entry =
- { const_entry_body : 'a const_entry_body;
- (* List of section variables *)
- const_entry_secctx : Context.Named.t option;
- (* State id on which the completion of type checking is reported *)
- const_entry_feedback : Stateid.t option;
- const_entry_type : Term.types option;
- const_entry_polymorphic : bool;
- const_entry_universes : Univ.UContext.t;
- const_entry_opaque : bool;
- const_entry_inline_code : bool }
- type parameter_entry = Context.Named.t option * bool * Term.types Univ.in_universe_context * inline
- type projection_entry = Entries.projection_entry
- type 'a constant_entry = 'a Entries.constant_entry =
- | DefinitionEntry of 'a definition_entry
- | ParameterEntry of parameter_entry
- | ProjectionEntry of projection_entry
+ type evars_flag = bool
+ type clear_flag = bool option
+ type advanced_flag = bool
+ type rec_flag = bool
+
+ type 'a or_by_notation =
+ | AN of 'a
+ | ByNotation of (string * string option) Loc.located
+
+ type 'a or_var =
+ | ArgArg of 'a
+ | ArgVar of Names.Id.t Loc.located
+
+ type 'a and_short_name = 'a * Names.Id.t Loc.located option
+
+ type 'a glob_sort_gen =
+ | GProp (** representation of [Prop] literal *)
+ | GSet (** representation of [Set] literal *)
+ | GType of 'a (** representation of [Type] literal *)
+
+ type level_info = Names.Name.t Loc.located option
+ type glob_level = level_info glob_sort_gen
+
+ type sort_info = Names.Name.t Loc.located list
+ type glob_sort = sort_info glob_sort_gen
+
+ type case_style = Term.case_style =
+ | LetStyle
+ | IfStyle
+ | LetPatternStyle
+ | MatchStyle
+ | RegularStyle (** infer printing form from number of constructor *)
+
+ type 'a cast_type =
+ | CastConv of 'a
+ | CastVM of 'a
+ | CastCoerce
+ | CastNative of 'a
+
+ type 'constr intro_pattern_expr =
+ | IntroForthcoming of bool
+ | IntroNaming of intro_pattern_naming_expr
+ | IntroAction of 'constr intro_pattern_action_expr
+ and intro_pattern_naming_expr =
+ | IntroIdentifier of Names.Id.t
+ | IntroFresh of Names.Id.t
+ | IntroAnonymous
+ and 'constr intro_pattern_action_expr =
+ | IntroWildcard
+ | IntroOrAndPattern of 'constr or_and_intro_pattern_expr
+ | IntroInjection of ('constr intro_pattern_expr) Loc.located list
+ | IntroApplyOn of 'constr Loc.located * 'constr intro_pattern_expr Loc.located
+ | IntroRewrite of bool
+ and 'constr or_and_intro_pattern_expr =
+ | IntroOrPattern of ('constr intro_pattern_expr) Loc.located list list
+ | IntroAndPattern of ('constr intro_pattern_expr) Loc.located list
+
+ type quantified_hypothesis =
+ | AnonHyp of int
+ | NamedHyp of Names.Id.t
+
+ type 'a explicit_bindings = (quantified_hypothesis * 'a) Loc.located list
+
+ type 'a bindings =
+ | ImplicitBindings of 'a list
+ | ExplicitBindings of 'a explicit_bindings
+ | NoBindings
+
+ type 'a with_bindings = 'a * 'a bindings
+
+ type 'a core_destruction_arg =
+ | ElimOnConstr of 'a
+ | ElimOnIdent of Names.Id.t Loc.located
+ | ElimOnAnonHyp of int
+
+ type inversion_kind =
+ | SimpleInversion
+ | FullInversion
+ | FullInversionClear
+
+ type multi =
+ | Precisely of int
+ | UpTo of int
+ | RepeatStar
+ | RepeatPlus
+ type 'id move_location =
+ | MoveAfter of 'id
+ | MoveBefore of 'id
+ | MoveFirst
+ | MoveLast
+
+ type 'a destruction_arg = clear_flag * 'a core_destruction_arg
+
end
-module Mod_typing :
+module Locus :
sig
- type 'alg translation =
- Declarations.module_signature * 'alg * Mod_subst.delta_resolver * Univ.ContextSet.t
- val translate_mse :
- Environ.env -> Names.ModPath.t option -> Entries.inline -> Declarations.module_alg_expr ->
- Declarations.module_alg_expr translation
+ type 'a occurrences_gen =
+ | AllOccurrences
+ | AllOccurrencesBut of 'a list (** non-empty *)
+ | NoOccurrences
+ | OnlyOccurrences of 'a list (** non-empty *)
+ type occurrences = int occurrences_gen
+ type occurrences_expr = (int Misctypes.or_var) occurrences_gen
+ type 'a with_occurrences = occurrences_expr * 'a
+ type hyp_location_flag =
+ InHyp | InHypTypeOnly | InHypValueOnly
+ type 'a hyp_location_expr = 'a with_occurrences * hyp_location_flag
+ type 'id clause_expr =
+ { onhyps : 'id hyp_location_expr list option;
+ concl_occs : occurrences_expr }
+ type clause = Names.Id.t clause_expr
+ type hyp_location = Names.Id.t * hyp_location_flag
+ type goal_location = hyp_location option
end
-module Esubst :
+(************************************************************************)
+(* End Modules from intf/ *)
+(************************************************************************)
+
+(************************************************************************)
+(* Modules from library/ *)
+(************************************************************************)
+
+module Univops :
sig
- type 'a subs = 'a Esubst.subs
- val subs_id : int -> 'a subs
+ val universes_of_constr : Term.constr -> Univ.universe_set
+ val restrict_universe_context : Univ.universe_context_set -> Univ.universe_set -> Univ.universe_context_set
end
-module CClosure :
+module Nameops :
sig
- type fconstr = CClosure.fconstr
- type clos_infos = CClosure.clos_infos
- type table_key = Names.Constant.t Univ.puniverses Names.tableKey
- type fterm = CClosure.fterm =
- | FRel of int
- | FAtom of Term.constr (** Metas and Sorts *)
- | FCast of fconstr * Term.cast_kind * fconstr
- | FFlex of table_key
- | FInd of Names.inductive Univ.puniverses
- | FConstruct of Names.constructor Univ.puniverses
- | FApp of fconstr * fconstr array
- | FProj of Names.Projection.t * fconstr
- | FFix of Term.fixpoint * fconstr Esubst.subs
- | FCoFix of Term.cofixpoint * fconstr Esubst.subs
- | FCaseT of Term.case_info * Term.constr * fconstr * Term.constr array * fconstr Esubst.subs (* predicate and branches are closures *)
- | FLambda of int * (Names.Name.t * Term.constr) list * Term.constr * fconstr Esubst.subs
- | FProd of Names.Name.t * fconstr * fconstr
- | FLetIn of Names.Name.t * fconstr * fconstr * Term.constr * fconstr Esubst.subs
- | FEvar of Term.existential * fconstr Esubst.subs
- | FLIFT of int * fconstr
- | FCLOS of Term.constr * fconstr Esubst.subs
- | FLOCKED
- module RedFlags : sig
- type reds = CClosure.RedFlags.reds
- type red_kind = CClosure.RedFlags.red_kind
- val mkflags : red_kind list -> reds
- val fBETA : red_kind
- val fCOFIX : red_kind
- val fCONST : Names.Constant.t -> CClosure.RedFlags.red_kind
- val fFIX : red_kind
- val fMATCH : red_kind
- val fZETA : red_kind
- val red_add_transparent : reds -> Names.transparent_state -> reds
+ val atompart_of_id : Names.Id.t -> string
+
+ val pr_id : Names.Id.t -> Pp.std_ppcmds
+ [@@ocaml.deprecated "alias of API.Names.Id.print"]
+
+ val pr_name : Names.Name.t -> Pp.std_ppcmds
+ [@@ocaml.deprecated "alias of API.Names.Name.print"]
+
+ val name_fold : (Names.Id.t -> 'a -> 'a) -> Names.Name.t -> 'a -> 'a
+ val name_app : (Names.Id.t -> Names.Id.t) -> Names.Name.t -> Names.Name.t
+ val add_suffix : Names.Id.t -> string -> Names.Id.t
+ val increment_subscript : Names.Id.t -> Names.Id.t
+ val make_ident : string -> int option -> Names.Id.t
+ val out_name : Names.Name.t -> Names.Id.t
+ val pr_lab : Names.Label.t -> Pp.std_ppcmds
+ module Name :
+ sig
+ include module type of struct include Names.Name end
+ val get_id : t -> Names.Id.t
+ val fold_right : (Names.Id.t -> 'a -> 'a) -> t -> 'a -> 'a
end
- val mk_clos : fconstr Esubst.subs -> Term.constr -> fconstr
- val mk_atom : Term.constr -> fconstr
- val mk_clos_deep :
- (fconstr Esubst.subs -> Term.constr -> fconstr) ->
- fconstr Esubst.subs -> Term.constr -> fconstr
- val mk_red : fterm -> fconstr
- val all : RedFlags.reds
- val beta : RedFlags.reds
- val betaiota : RedFlags.reds
- val betaiotazeta : RedFlags.reds
+end
+
+module Libnames :
+sig
- val create_clos_infos : ?evars:(Term.existential -> Term.constr option) -> RedFlags.reds -> Environ.env -> clos_infos
+ open Util
+ open Names
- val whd_val : clos_infos -> fconstr -> Term.constr
+ type full_path
+ val pr_path : full_path -> Pp.std_ppcmds
+ val make_path : Names.DirPath.t -> Names.Id.t -> full_path
+ val eq_full_path : full_path -> full_path -> bool
+ val dirpath : full_path -> Names.DirPath.t
+ val path_of_string : string -> full_path
- val inject : Term.constr -> fconstr
+ type qualid
+ val make_qualid : Names.DirPath.t -> Names.Id.t -> qualid
+ val qualid_eq : qualid -> qualid -> bool
+ val repr_qualid : qualid -> Names.DirPath.t * Names.Id.t
+ val pr_qualid : qualid -> Pp.std_ppcmds
+ val string_of_qualid : qualid -> string
+ val qualid_of_string : string -> qualid
+ val qualid_of_path : full_path -> qualid
+ val qualid_of_dirpath : Names.DirPath.t -> qualid
+ val qualid_of_ident : Names.Id.t -> qualid
- val kl : clos_infos -> fconstr -> Term.constr
- val term_of_fconstr : fconstr -> Term.constr
+ type reference =
+ | Qualid of qualid Loc.located
+ | Ident of Names.Id.t Loc.located
+ val loc_of_reference : reference -> Loc.t option
+ val qualid_of_reference : reference -> qualid Loc.located
+ val pr_reference : reference -> Pp.std_ppcmds
+
+ val is_dirpath_prefix_of : Names.DirPath.t -> Names.DirPath.t -> bool
+ val split_dirpath : Names.DirPath.t -> Names.DirPath.t * Names.Id.t
+ val dirpath_of_string : string -> Names.DirPath.t
+ val pr_dirpath : Names.DirPath.t -> Pp.std_ppcmds
+
+ val string_of_path : full_path -> string
+ val basename : full_path -> Names.Id.t
+
+ type object_name = full_path * Names.KerName.t
+ type object_prefix = Names.DirPath.t * (Names.ModPath.t * Names.DirPath.t)
+
+ module Dirset : Set.S with type elt = DirPath.t
+ module Dirmap : Map.ExtS with type key = DirPath.t and module Set := Dirset
+ module Spmap : CSig.MapS with type key = full_path
end
-module Type_errors :
+module Globnames :
sig
- type type_error = Type_errors.type_error
- exception TypeError of Environ.env * type_error
+
+ open Util
+
+ type global_reference =
+ | VarRef of Names.Id.t
+ | ConstRef of Names.Constant.t
+ | IndRef of Names.inductive
+ | ConstructRef of Names.constructor
+
+ type extended_global_reference =
+ | TrueGlobal of global_reference
+ | SynDef of Names.KerName.t
+
+ (* Long term: change implementation so that only 1 kind of order is needed.
+ * Today: _env ones are fine grained, which one to pick depends. Eg.
+ * - conversion rule are implemented by the non_env ones
+ * - pretty printing (of user provided names/aliases) are implemented by
+ * the _env ones
+ *)
+ module Refset : CSig.SetS with type elt = global_reference
+ module Refmap : Map.ExtS
+ with type key = global_reference and module Set := Refset
+
+ module Refset_env : CSig.SetS with type elt = global_reference
+ module Refmap_env : Map.ExtS
+ with type key = global_reference and module Set := Refset_env
+
+ module RefOrdered :
+ sig
+ type t = global_reference
+ val compare : t -> t -> int
+ end
+
+ val pop_global_reference : global_reference -> global_reference
+ val eq_gr : global_reference -> global_reference -> bool
+ val destIndRef : global_reference -> Names.inductive
+
+ val encode_mind : Names.DirPath.t -> Names.Id.t -> Names.MutInd.t
+ val encode_con : Names.DirPath.t -> Names.Id.t -> Names.Constant.t
+
+ val global_of_constr : Constr.t -> global_reference
+
+ val subst_global : Mod_subst.substitution -> global_reference -> global_reference * Constr.t
+ val destConstructRef : global_reference -> Names.constructor
+
+ val reference_of_constr : Constr.t -> global_reference
+ [@@ocaml.deprecated "alias of API.Globnames.global_of_constr"]
+
+ val is_global : global_reference -> Constr.t -> bool
end
-module Evar :
+module Libobject :
sig
- (** Unique identifier of some {i evar} *)
- type t = Prelude.evar
+ type obj
+ type 'a substitutivity =
+ | Dispose
+ | Substitute of 'a
+ | Keep of 'a
+ | Anticipate of 'a
- (** Recover the underlying integer. *)
- val repr : t -> int
+ type 'a object_declaration = {
+ object_name : string;
+ cache_function : Libnames.object_name * 'a -> unit;
+ load_function : int -> Libnames.object_name * 'a -> unit;
+ open_function : int -> Libnames.object_name * 'a -> unit;
+ classify_function : 'a -> 'a substitutivity;
+ subst_function : Mod_subst.substitution * 'a -> 'a;
+ discharge_function : Libnames.object_name * 'a -> 'a option;
+ rebuild_function : 'a -> 'a
+ }
+ val declare_object : 'a object_declaration -> ('a -> obj)
+ val default_object : string -> 'a object_declaration
+ val object_tag : obj -> string
+end
- val equal : t -> t -> bool
+module Summary :
+sig
- (** a set of unique identifiers of some {i evars} *)
- module Set : module type of struct include Evar.Set end
+ type frozen
+ type marshallable
+
+ type 'a summary_declaration =
+ { freeze_function : marshallable -> 'a;
+ unfreeze_function : 'a -> unit;
+ init_function : unit -> unit; }
+
+ val ref : ?freeze:(marshallable -> 'a -> 'a) -> name:string -> 'a -> 'a ref
+ val declare_summary : string -> 'a summary_declaration -> unit
+ module Local :
+ sig
+ type 'a local_ref
+ val ref : ?freeze:('a -> 'a) -> name:string -> 'a -> 'a local_ref
+ val (:=) : 'a local_ref -> 'a -> unit
+ val (!) : 'a local_ref -> 'a
+ end
+end
+
+module Nametab :
+sig
+ exception GlobalizationError of Libnames.qualid
+
+ type ltac_constant = Names.KerName.t
+
+ val global : Libnames.reference -> Globnames.global_reference
+ val global_of_path : Libnames.full_path -> Globnames.global_reference
+ val shortest_qualid_of_global : Names.Id.Set.t -> Globnames.global_reference -> Libnames.qualid
+ val path_of_global : Globnames.global_reference -> Libnames.full_path
+ val locate_extended : Libnames.qualid -> Globnames.extended_global_reference
+ val full_name_module : Libnames.qualid -> Names.DirPath.t
+ val locate_tactic : Libnames.qualid -> Names.KerName.t
+ val pr_global_env : Names.Id.Set.t -> Globnames.global_reference -> Pp.std_ppcmds
+ val shortest_qualid_of_tactic : Names.KerName.t -> Libnames.qualid
+ val basename_of_global : Globnames.global_reference -> Names.Id.t
+
+ type visibility =
+ | Until of int
+ | Exactly of int
+
+ val push_tactic : visibility -> Libnames.full_path -> Names.KerName.t -> unit
+ val error_global_not_found : ?loc:Loc.t -> Libnames.qualid -> 'a
+ val shortest_qualid_of_module : Names.ModPath.t -> Libnames.qualid
+ val dirpath_of_module : Names.ModPath.t -> Names.DirPath.t
+ val locate_module : Libnames.qualid -> Names.ModPath.t
+ val dirpath_of_global : Globnames.global_reference -> Names.DirPath.t
+ val locate : Libnames.qualid -> Globnames.global_reference
+ val locate_constant : Libnames.qualid -> Names.Constant.t
+end
+
+module Global :
+sig
+ val env : unit -> Environ.env
+ val lookup_mind : Names.MutInd.t -> Declarations.mutual_inductive_body
+ val lookup_constant : Names.Constant.t -> Declarations.constant_body
+ val lookup_module : Names.ModPath.t -> Declarations.module_body
+ val lookup_modtype : Names.ModPath.t -> Declarations.module_type_body
+ val lookup_inductive : Names.inductive -> Declarations.mutual_inductive_body * Declarations.one_inductive_body
+ val constant_of_delta_kn : Names.KerName.t -> Names.Constant.t
+ val register :
+ Retroknowledge.field -> Constr.t -> Constr.t -> unit
+ val env_of_context : Environ.named_context_val -> Environ.env
+ val is_polymorphic : Globnames.global_reference -> bool
+
+ val constr_of_global_in_context : Environ.env ->
+ Globnames.global_reference -> Constr.types * Univ.AUContext.t
+
+ val type_of_global_in_context : Environ.env ->
+ Globnames.global_reference -> Constr.types * Univ.AUContext.t
+
+ val current_dirpath : unit -> Names.DirPath.t
+ val body_of_constant_body : Declarations.constant_body -> (Constr.t * Univ.AUContext.t) option
+ val body_of_constant : Names.Constant.t -> (Constr.t * Univ.AUContext.t) option
+ val add_constraints : Univ.Constraint.t -> unit
+end
+
+module Lib : sig
+ type is_type = bool
+ type export = bool option
+ type node =
+ | Leaf of Libobject.obj (* FIX: horrible hack (wrt. Enrico) *)
+ | CompilingLibrary of Libnames.object_prefix
+ | OpenedModule of is_type * export * Libnames.object_prefix * Summary.frozen
+ | ClosedModule of library_segment
+ | OpenedSection of Libnames.object_prefix * Summary.frozen
+ | ClosedSection of library_segment
+
+ and library_segment = (Libnames.object_name * node) list
+
+ val current_mp : unit -> Names.ModPath.t
+ val is_modtype : unit -> bool
+ val is_module : unit -> bool
+ val sections_are_opened : unit -> bool
+ val add_anonymous_leaf : ?cache_first:bool -> Libobject.obj -> unit
+ val contents : unit -> library_segment
+ val cwd : unit -> Names.DirPath.t
+ val add_leaf : Names.Id.t -> Libobject.obj -> Libnames.object_name
+ val make_kn : Names.Id.t -> Names.KerName.t
+ val make_path : Names.Id.t -> Libnames.full_path
+ val discharge_con : Names.Constant.t -> Names.Constant.t
+ val discharge_inductive : Names.inductive -> Names.inductive
+end
+
+module Declaremods :
+sig
+
+ val append_end_library_hook : (unit -> unit) -> unit
+
+end
+
+module Library :
+sig
+ val library_is_loaded : Names.DirPath.t -> bool
+ val loaded_libraries : unit -> Names.DirPath.t list
+end
+
+module States :
+sig
+ val with_state_protection_on_exception : ('a -> 'b) -> 'a -> 'b
+ val with_state_protection : ('a -> 'b) -> 'a -> 'b
+end
+
+module Kindops :
+sig
+ val logical_kind_of_goal_kind : Decl_kinds.goal_object_kind -> Decl_kinds.logical_kind
+end
+
+module Goptions :
+sig
+ type option_name = string list
+ type 'a option_sig =
+ {
+ optdepr : bool;
+ optname : string;
+ optkey : option_name;
+ optread : unit -> 'a;
+ optwrite : 'a -> unit
+ }
+
+ type 'a write_function = 'a -> unit
+
+ val declare_bool_option : ?preprocess:(bool -> bool) ->
+ bool option_sig -> bool write_function
+ val declare_int_option : ?preprocess:(int option -> int option) ->
+ int option option_sig -> int option write_function
+ val declare_string_option: ?preprocess:(string -> string) ->
+ string option_sig -> string write_function
+ val set_bool_option_value : option_name -> bool -> unit
+end
+
+module Keys :
+sig
+ type key
+ val constr_key : ('a -> ('a, 't, 'u, 'i) Constr.kind_of_term) -> 'a -> key option
+ val declare_equiv_keys : key -> key -> unit
+ val pr_keys : (Globnames.global_reference -> Pp.std_ppcmds) -> Pp.std_ppcmds
+end
+
+module Coqlib :
+sig
+
+ type coq_eq_data = { eq : Globnames.global_reference;
+ ind : Globnames.global_reference;
+ refl : Globnames.global_reference;
+ sym : Globnames.global_reference;
+ trans: Globnames.global_reference;
+ congr: Globnames.global_reference;
+ }
+
+ type coq_sigma_data = {
+ proj1 : Globnames.global_reference;
+ proj2 : Globnames.global_reference;
+ elim : Globnames.global_reference;
+ intro : Globnames.global_reference;
+ typ : Globnames.global_reference }
+ val gen_reference : string -> string list -> string -> Globnames.global_reference
+ val find_reference : string -> string list -> string -> Globnames.global_reference
+ val check_required_library : string list -> unit
+ val logic_module_name : string list
+ val glob_true : Globnames.global_reference
+ val glob_false : Globnames.global_reference
+ val glob_O : Globnames.global_reference
+ val glob_S : Globnames.global_reference
+ val nat_path : Libnames.full_path
+ val datatypes_module_name : string list
+ val glob_eq : Globnames.global_reference
+ val build_coq_eq_sym : Globnames.global_reference Util.delayed
+ val build_coq_False : Globnames.global_reference Util.delayed
+ val build_coq_not : Globnames.global_reference Util.delayed
+ val build_coq_eq : Globnames.global_reference Util.delayed
+ val build_coq_eq_data : coq_eq_data Util.delayed
+ val path_of_O : Names.constructor
+ val path_of_S : Names.constructor
+ val build_prod : coq_sigma_data Util.delayed
+ val build_coq_True : Globnames.global_reference Util.delayed
+ val coq_iff_ref : Globnames.global_reference lazy_t
+ val build_coq_iff_left_proj : Globnames.global_reference Util.delayed
+ val build_coq_iff_right_proj : Globnames.global_reference Util.delayed
+ val init_modules : string list list
+ val build_coq_eq_refl : Globnames.global_reference Util.delayed
+ val arith_modules : string list list
+ val zarith_base_modules : string list list
+ val gen_reference_in_modules : string -> string list list-> string -> Globnames.global_reference
+ val jmeq_module_name : string list
+ val coq_eq_ref : Globnames.global_reference lazy_t
+ val coq_not_ref : Globnames.global_reference lazy_t
+ val coq_or_ref : Globnames.global_reference lazy_t
+ val build_coq_and : Globnames.global_reference Util.delayed
+ val build_coq_I : Globnames.global_reference Util.delayed
+ val coq_reference : string -> string list -> string -> Globnames.global_reference
+end
+
+(************************************************************************)
+(* End of modules from library/ *)
+(************************************************************************)
+
+(************************************************************************)
+(* Modules from engine/ *)
+(************************************************************************)
+
+module Universes :
+sig
+ type universe_binders
+ type universe_opt_subst
+ val fresh_inductive_instance : Environ.env -> Names.inductive -> Term.pinductive Univ.in_universe_context_set
+ val new_Type : Names.DirPath.t -> Term.types
+ val type_of_global : Globnames.global_reference -> Term.types Univ.in_universe_context_set
+ val constr_of_global : Globnames.global_reference -> Constr.t
+ val new_univ_level : Names.DirPath.t -> Univ.Level.t
+ val new_sort_in_family : Sorts.family -> Sorts.t
+ val pr_with_global_universes : Univ.Level.t -> Pp.std_ppcmds
+ val pr_universe_opt_subst : universe_opt_subst -> Pp.std_ppcmds
+ type universe_constraint
+
+ module Constraints :
+ sig
+ type t
+ val pr : t -> Pp.std_ppcmds
+ end
+
+ type universe_constraints = Constraints.t
+end
+
+module UState :
+sig
+ type t
+ val context : t -> Univ.UContext.t
+ val context_set : t -> Univ.ContextSet.t
+ val of_context_set : Univ.ContextSet.t -> t
+
+ type rigid =
+ | UnivRigid
+ | UnivFlexible of bool
+
+end
+
+(* XXX: Moved from intf *)
+module Evar_kinds :
+sig
+ type obligation_definition_status =
+ | Define of bool
+ | Expand
+
+ type matching_var_kind =
+ | FirstOrderPatVar of Names.Id.t
+ | SecondOrderPatVar of Names.Id.t
+
+ type t =
+ | ImplicitArg of Globnames.global_reference * (int * Names.Id.t option)
+ * bool (** Force inference *)
+ | BinderType of Names.Name.t
+ | NamedHole of Names.Id.t (* coming from some ?[id] syntax *)
+ | QuestionMark of obligation_definition_status * Names.Name.t
+ | CasesType of bool (* true = a subterm of the type *)
+ | InternalHole
+ | TomatchTypeParameter of Names.inductive * int
+ | GoalEvar
+ | ImpossibleCase
+ | MatchingVar of matching_var_kind
+ | VarInstance of Names.Id.t
+ | SubEvar of Constr.existential_key
end
module Evd :
sig
+
+ type evar = Constr.existential_key
+
val string_of_existential : Evar.t -> string
- type evar_constraint = Prelude.conv_pb * Environ.env * Term.constr * Term.constr
+ type evar_constraint = Reduction.conv_pb * Environ.env * Constr.t * Constr.t
(* --------------------------------- *)
@@ -1404,37 +2195,37 @@ sig
module Store :
sig
- type t = Evd.Store.t
+ type t
val empty : t
end
module Filter :
sig
- type t = Evd.Filter.t
+ type t
val repr : t -> bool list option
end
(** This value defines the refinement of a given {i evar} *)
- type evar_body = Evd.evar_body =
+ type evar_body =
| Evar_empty (** given {i evar} was not yet refined *)
- | Evar_defined of Term.constr (** given {i var} was refined to the indicated term *)
+ | Evar_defined of Constr.t (** given {i var} was refined to the indicated term *)
(** all the information we have concerning some {i evar} *)
- type evar_info = Evd.evar_info =
+ type evar_info =
{
- evar_concl : Term.constr;
+ evar_concl : Constr.t;
evar_hyps : Environ.named_context_val;
evar_body : evar_body;
evar_filter : Filter.t;
evar_source : Evar_kinds.t Loc.located;
- evar_candidates : Term.constr list option; (* if not None, list of allowed instances *)
+ evar_candidates : Constr.t list option; (* if not None, list of allowed instances *)
evar_extra : Store.t
}
- val evar_concl : evar_info -> Term.constr
+ val evar_concl : evar_info -> Constr.t
val evar_body : evar_info -> evar_body
val evar_context : evar_info -> Context.Named.t
- val instantiate_evar_array : evar_info -> Term.constr -> Term.constr array -> Term.constr
+ val instantiate_evar_array : evar_info -> Constr.t -> Constr.t array -> Constr.t
val evar_filtered_env : evar_info -> Environ.env
val evar_hyps : evar_info -> Environ.named_context_val
@@ -1442,39 +2233,50 @@ sig
(* evar map *)
- type evar_map = Prelude.evar_map
- type open_constr = evar_map * Term.constr
+ type evar_map
+ type open_constr = evar_map * Constr.t
+
+ open Util
- type rigid = Prelude.rigid =
+ module Metaset : Set.S with type elt = Constr.metavariable
+
+ type rigid = UState.rigid =
| UnivRigid
- | UnivFlexible of bool
+ | UnivFlexible of bool
-
- type 'a freelisted = 'a Evd.freelisted = {
+ type 'a freelisted = {
rebus : 'a;
- freemetas : Evd.Metaset.t
+ freemetas : Metaset.t
}
- type instance_status = Evd.instance_status
- type clbinding = Evd.clbinding =
- | Cltyp of Names.Name.t * Term.constr freelisted
- | Clval of Names.Name.t * (Term.constr freelisted * instance_status) * Term.constr freelisted
+
+ type instance_constraint = IsSuperType | IsSubType | Conv
+
+ type instance_typing_status =
+ CoerceToType | TypeNotProcessed | TypeProcessed
+
+ type instance_status = instance_constraint * instance_typing_status
+
+ type clbinding =
+ | Cltyp of Names.Name.t * Constr.t freelisted
+ | Clval of Names.Name.t * (Constr.t freelisted * instance_status) * Constr.t freelisted
+
val empty : evar_map
val from_env : Environ.env -> evar_map
val find : evar_map -> Evar.t -> evar_info
- val find_undefined : evar_map -> Prelude.evar -> evar_info
+ val find_undefined : evar_map -> evar -> evar_info
val is_defined : evar_map -> Evar.t -> bool
val mem : evar_map -> Evar.t -> bool
val add : evar_map -> Evar.t -> evar_info -> evar_map
val evar_universe_context : evar_map -> UState.t
val set_universe_context : evar_map -> UState.t -> evar_map
val universes : evar_map -> UGraph.t
- val define : Evar.t -> Term.constr -> evar_map -> evar_map
+ val define : Evar.t -> Constr.t -> evar_map -> evar_map
val fold : (Evar.t -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'a
val evar_key : Names.Id.t -> evar_map -> Evar.t
val create_evar_defs : evar_map -> evar_map
- val meta_declare : Prelude.metavariable -> Term.types -> ?name:Names.Name.t -> evar_map -> evar_map
+ val meta_declare : Constr.metavariable -> Term.types -> ?name:Names.Name.t -> evar_map -> evar_map
val clear_metas : evar_map -> evar_map
@@ -1483,24 +2285,24 @@ sig
val remove : evar_map -> Evar.t -> evar_map
val fresh_global : ?loc:Loc.t -> ?rigid:rigid -> ?names:Univ.Instance.t -> Environ.env ->
- evar_map -> Prelude.global_reference -> evar_map * Term.constr
+ evar_map -> Globnames.global_reference -> evar_map * Constr.t
val evar_filtered_context : evar_info -> Context.Named.t
val fresh_inductive_instance : ?loc:Loc.t -> Environ.env -> evar_map -> Names.inductive -> evar_map * Term.pinductive
val fold_undefined : (Evar.t -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'a
val universe_context_set : evar_map -> Univ.ContextSet.t
- val evar_ident : Prelude.evar -> evar_map -> Names.Id.t option
+ val evar_ident : evar -> evar_map -> Names.Id.t option
val extract_all_conv_pbs : evar_map -> evar_map * evar_constraint list
val universe_context : ?names:(Names.Id.t Loc.located) list -> evar_map ->
(Names.Id.t * Univ.Level.t) list * Univ.UContext.t
val nf_constraints : evar_map -> evar_map
val from_ctx : UState.t -> evar_map
- val meta_list : evar_map -> (Prelude.metavariable * clbinding) list
+ val meta_list : evar_map -> (Constr.metavariable * clbinding) list
- val meta_defined : evar_map -> Prelude.metavariable -> bool
+ val meta_defined : evar_map -> Constr.metavariable -> bool
- val meta_name : evar_map -> Prelude.metavariable -> Names.Name.t
+ val meta_name : evar_map -> Constr.metavariable -> Names.Name.t
module MonadR :
sig
@@ -1510,7 +2312,7 @@ sig
end
end
- type 'a sigma = 'a Evd.sigma = {
+ type 'a sigma = {
it : 'a ;
sigma : evar_map
}
@@ -1527,14 +2329,11 @@ sig
val union_evar_universe_context : UState.t -> UState.t -> UState.t
val merge_universe_context : evar_map -> UState.t -> evar_map
- type unsolvability_explanation = Evd.unsolvability_explanation =
- | SeveralInstancesFound of int
-
- module Metaset : module type of struct include Evd.Metaset end
- with type elt = Prelude.metavariable
+ type unsolvability_explanation =
+ | SeveralInstancesFound of int
(** Return {i ids} of all {i evars} that occur in a given term. *)
- val evars_of_term : Term.constr -> Evar.Set.t
+ val evars_of_term : Constr.t -> Evar.Set.t
val evar_universe_context_of : Univ.ContextSet.t -> UState.t
[@@ocaml.deprecated "alias of API.UState.of_context_set"]
@@ -1545,14 +2344,379 @@ sig
type evar_universe_context = UState.t
[@@ocaml.deprecated "alias of API.UState.t"]
- val existential_opt_value : evar_map -> Term.existential -> Term.constr option
- val existential_value : evar_map -> Term.existential -> Term.constr
+ val existential_opt_value : evar_map -> Term.existential -> Constr.t option
+ val existential_value : evar_map -> Term.existential -> Constr.t
exception NotInstantiatedEvar
val fresh_sort_in_family : ?loc:Loc.t -> ?rigid:rigid -> Environ.env -> evar_map -> Sorts.family -> evar_map * Sorts.t
end
+(* XXX: moved from intf *)
+module Constrexpr :
+sig
+
+ type binder_kind =
+ | Default of Decl_kinds.binding_kind
+ | Generalized of Decl_kinds.binding_kind * Decl_kinds.binding_kind * bool
+
+ type explicitation =
+ | ExplByPos of int * Names.Id.t option
+ | ExplByName of Names.Id.t
+ type sign = bool
+ type raw_natural_number = string
+ type prim_token =
+ | Numeral of raw_natural_number * sign
+ | String of string
+
+ type notation = string
+ type instance_expr = Misctypes.glob_level list
+ type proj_flag = int option
+ type abstraction_kind =
+ | AbsLambda
+ | AbsPi
+
+ type cases_pattern_expr_r =
+ | CPatAlias of cases_pattern_expr * Names.Id.t
+ | CPatCstr of Libnames.reference
+ * cases_pattern_expr list option * cases_pattern_expr list
+ (** [CPatCstr (_, c, Some l1, l2)] represents (@c l1) l2 *)
+ | CPatAtom of Libnames.reference option
+ | CPatOr of cases_pattern_expr list
+ | CPatNotation of notation * cases_pattern_notation_substitution
+ * cases_pattern_expr list
+ | CPatPrim of prim_token
+ | CPatRecord of (Libnames.reference * cases_pattern_expr) list
+ | CPatDelimiters of string * cases_pattern_expr
+ | CPatCast of cases_pattern_expr * constr_expr
+ and cases_pattern_expr = cases_pattern_expr_r CAst.t
+
+ and cases_pattern_notation_substitution =
+ cases_pattern_expr list * cases_pattern_expr list list
+
+ and constr_expr_r =
+ | CRef of Libnames.reference * instance_expr option
+ | CFix of Names.Id.t Loc.located * fix_expr list
+ | CCoFix of Names.Id.t Loc.located * cofix_expr list
+ | CProdN of binder_expr list * constr_expr
+ | CLambdaN of binder_expr list * constr_expr
+ | CLetIn of Names.Name.t Loc.located * constr_expr * constr_expr option * constr_expr
+ | CAppExpl of (proj_flag * Libnames.reference * instance_expr option) * constr_expr list
+ | CApp of (proj_flag * constr_expr) *
+ (constr_expr * explicitation Loc.located option) list
+ | CRecord of (Libnames.reference * constr_expr) list
+ | CCases of Term.case_style
+ * constr_expr option
+ * case_expr list
+ * branch_expr list
+ | CLetTuple of Names.Name.t Loc.located list * (Names.Name.t Loc.located option * constr_expr option) *
+ constr_expr * constr_expr
+ | CIf of constr_expr * (Names.Name.t Loc.located option * constr_expr option)
+ * constr_expr * constr_expr
+ | CHole of Evar_kinds.t option * Misctypes.intro_pattern_naming_expr * Genarg.raw_generic_argument option
+ | CPatVar of Names.Id.t
+ | CEvar of Names.Id.t * (Names.Id.t * constr_expr) list
+ | CSort of Misctypes.glob_sort
+ | CCast of constr_expr * constr_expr Misctypes.cast_type
+ | CNotation of notation * constr_notation_substitution
+ | CGeneralization of Decl_kinds.binding_kind * abstraction_kind option * constr_expr
+ | CPrim of prim_token
+ | CDelimiters of string * constr_expr
+ and constr_expr = constr_expr_r CAst.t
+
+ and case_expr = constr_expr * Names.Name.t Loc.located option * cases_pattern_expr option
+
+ and branch_expr =
+ (cases_pattern_expr list Loc.located list * constr_expr) Loc.located
+
+ and binder_expr =
+ Names.Name.t Loc.located list * binder_kind * constr_expr
+
+ and fix_expr =
+ Names.Id.t Loc.located * (Names.Id.t Loc.located option * recursion_order_expr) *
+ local_binder_expr list * constr_expr * constr_expr
+
+ and cofix_expr =
+ Names.Id.t Loc.located * local_binder_expr list * constr_expr * constr_expr
+
+ and recursion_order_expr =
+ | CStructRec
+ | CWfRec of constr_expr
+ | CMeasureRec of constr_expr * constr_expr option
+
+ and local_binder_expr =
+ | CLocalAssum of Names.Name.t Loc.located list * binder_kind * constr_expr
+ | CLocalDef of Names.Name.t Loc.located * constr_expr * constr_expr option
+ | CLocalPattern of (cases_pattern_expr * constr_expr option) Loc.located
+
+ and constr_notation_substitution =
+ constr_expr list *
+ constr_expr list list *
+ local_binder_expr list list
+
+ type typeclass_constraint = (Names.Name.t Loc.located * Names.Id.t Loc.located list option) * Decl_kinds.binding_kind * constr_expr
+ type constr_pattern_expr = constr_expr
+end
+
+module Genredexpr :
+sig
+
+ (** The parsing produces initially a list of [red_atom] *)
+ type 'a red_atom =
+ | FBeta
+ | FMatch
+ | FFix
+ | FCofix
+ | FZeta
+ | FConst of 'a list
+ | FDeltaBut of 'a list
+
+ (** This list of atoms is immediately converted to a [glob_red_flag] *)
+ type 'a glob_red_flag = {
+ rBeta : bool;
+ rMatch : bool;
+ rFix : bool;
+ rCofix : bool;
+ rZeta : bool;
+ rDelta : bool; (** true = delta all but rConst; false = delta only on rConst*)
+ rConst : 'a list
+ }
+
+ (** Generic kinds of reductions *)
+ type ('a,'b,'c) red_expr_gen =
+ | Red of bool
+ | Hnf
+ | Simpl of 'b glob_red_flag*('b,'c) Util.union Locus.with_occurrences option
+ | Cbv of 'b glob_red_flag
+ | Cbn of 'b glob_red_flag
+ | Lazy of 'b glob_red_flag
+ | Unfold of 'b Locus.with_occurrences list
+ | Fold of 'a list
+ | Pattern of 'a Locus.with_occurrences list
+ | ExtraRedExpr of string
+ | CbvVm of ('b,'c) Util.union Locus.with_occurrences option
+ | CbvNative of ('b,'c) Util.union Locus.with_occurrences option
+
+ type ('a,'b,'c) may_eval =
+ | ConstrTerm of 'a
+ | ConstrEval of ('a,'b,'c) red_expr_gen * 'a
+ | ConstrContext of Names.Id.t Loc.located * 'a
+ | ConstrTypeOf of 'a
+
+ type r_trm = Constrexpr.constr_expr
+ type r_pat = Constrexpr.constr_pattern_expr
+ type r_cst = Libnames.reference Misctypes.or_by_notation
+ type raw_red_expr = (r_trm, r_cst, r_pat) red_expr_gen
+end
+
+(* XXX: end of moved from intf *)
+
+module EConstr :
+sig
+ type t
+ type constr = t
+ type types = t
+ type unsafe_judgment = (constr, types) Environ.punsafe_judgment
+ type named_declaration = (constr, types) Context.Named.Declaration.pt
+ type named_context = (constr, types) Context.Named.pt
+ type rel_context = (constr, types) Context.Rel.pt
+ type rel_declaration = (constr, types) Context.Rel.Declaration.pt
+ type existential = constr Constr.pexistential
+ module ESorts :
+ sig
+ type t
+ (** Type of sorts up-to universe unification. Essentially a wrapper around
+ Sorts.t so that normalization is ensured statically. *)
+
+ val make : Sorts.t -> t
+ (** Turn a sort into an up-to sort. *)
+
+ val kind : Evd.evar_map -> t -> Sorts.t
+ (** Returns the view into the current sort. Note that the kind of a variable
+ may change if the unification state of the evar map changes. *)
+
+ end
+
+ module EInstance :
+ sig
+ type t
+ (** Type of universe instances up-to universe unification. Similar to
+ {ESorts.t} for {Univ.Instance.t}. *)
+
+ val make : Univ.Instance.t -> t
+ val kind : Evd.evar_map -> t -> Univ.Instance.t
+ val empty : t
+ val is_empty : t -> bool
+ end
+
+ val of_constr : Constr.t -> constr
+
+ val kind : Evd.evar_map -> constr -> (constr, constr, ESorts.t, EInstance.t) Constr.kind_of_term
+
+ val mkArrow : constr -> constr -> constr
+ val mkInd : Names.inductive -> t
+ val mkProp : constr
+ val mkProd : Names.Name.t * constr * constr -> constr
+ val mkRel : int -> constr
+ val mkSort : Sorts.t -> constr
+ val mkVar : Names.Id.t -> constr
+ val mkLambda : Names.Name.t * constr * constr -> constr
+ val mkLambda_or_LetIn : rel_declaration -> constr -> constr
+ val mkApp : constr * constr array -> constr
+ val mkEvar : constr Constr.pexistential -> constr
+
+ val mkMeta : Constr.metavariable -> constr
+
+ val mkConstructU : Names.constructor * EInstance.t -> constr
+ val mkLetIn : Names.Name.t * constr * constr * constr -> constr
+ val mkProd_or_LetIn : rel_declaration -> constr -> constr
+ val mkCast : constr * Constr.cast_kind * constr -> constr
+ val mkNamedLambda : Names.Id.t -> types -> constr -> constr
+ val mkNamedProd : Names.Id.t -> types -> types -> types
+
+ val isCast : Evd.evar_map -> t -> bool
+ val isEvar : Evd.evar_map -> constr -> bool
+ val isInd : Evd.evar_map -> constr -> bool
+ val isRel : Evd.evar_map -> constr -> bool
+ val isSort : Evd.evar_map -> constr -> bool
+ val isVar : Evd.evar_map -> constr -> bool
+ val isConst : Evd.evar_map -> constr -> bool
+ val isConstruct : Evd.evar_map -> constr -> bool
+
+ val destInd : Evd.evar_map -> constr -> Names.inductive * EInstance.t
+ val destVar : Evd.evar_map -> constr -> Names.Id.t
+ val destEvar : Evd.evar_map -> constr -> constr Constr.pexistential
+ val destRel : Evd.evar_map -> constr -> int
+ val destProd : Evd.evar_map -> constr -> Names.Name.t * types * types
+ val destLambda : Evd.evar_map -> constr -> Names.Name.t * types * constr
+ val destApp : Evd.evar_map -> constr -> constr * constr array
+ val destConst : Evd.evar_map -> constr -> Names.Constant.t * EInstance.t
+ val destConstruct : Evd.evar_map -> constr -> Names.constructor * EInstance.t
+ val destFix : Evd.evar_map -> t -> (t, t) Constr.pfixpoint
+ val destCast : Evd.evar_map -> t -> t * Constr.cast_kind * t
+
+ val mkConstruct : Names.constructor -> constr
+
+ val compose_lam : (Names.Name.t * constr) list -> constr -> constr
+
+ val decompose_lam : Evd.evar_map -> constr -> (Names.Name.t * constr) list * constr
+ val decompose_lam_n_assum : Evd.evar_map -> int -> constr -> rel_context * constr
+ val decompose_app : Evd.evar_map -> constr -> constr * constr list
+ val decompose_prod : Evd.evar_map -> constr -> (Names.Name.t * constr) list * constr
+ val decompose_prod_assum : Evd.evar_map -> constr -> rel_context * constr
+
+ val applist : constr * constr list -> constr
+
+ val to_constr : Evd.evar_map -> constr -> Constr.t
+
+ val push_rel : rel_declaration -> Environ.env -> Environ.env
+
+ module Unsafe :
+ sig
+ val to_constr : constr -> Constr.t
+
+ val to_rel_decl : (constr, types) Context.Rel.Declaration.pt -> (Constr.constr, Constr.types) Context.Rel.Declaration.pt
+
+ (** Physical identity. Does not care for defined evars. *)
+
+ val to_named_decl : (constr, types) Context.Named.Declaration.pt -> (Constr.constr, Constr.types) Context.Named.Declaration.pt
+
+ val to_instance : EInstance.t -> Univ.Instance.t
+ end
+
+ module Vars :
+ sig
+ val substnl : t list -> int -> t -> t
+ val noccurn : Evd.evar_map -> int -> constr -> bool
+ val closed0 : Evd.evar_map -> constr -> bool
+ val subst1 : constr -> constr -> constr
+ val substl : constr list -> constr -> constr
+ val lift : int -> constr -> constr
+ val liftn : int -> int -> t -> t
+ val subst_var : Names.Id.t -> t -> t
+ val subst_vars : Names.Id.t list -> t -> t
+ end
+
+ val fresh_global :
+ ?loc:Loc.t -> ?rigid:UState.rigid -> ?names:Univ.Instance.t -> Environ.env ->
+ Evd.evar_map -> Globnames.global_reference -> Evd.evar_map * t
+
+ val of_named_decl : (Constr.t, Constr.types) Context.Named.Declaration.pt -> (constr, types) Context.Named.Declaration.pt
+ val of_rel_decl : (Constr.t, Constr.types) Context.Rel.Declaration.pt -> (constr, types) Context.Rel.Declaration.pt
+ val kind_of_type : Evd.evar_map -> constr -> (constr, constr) Term.kind_of_type
+ val to_lambda : Evd.evar_map -> int -> constr -> constr
+ val it_mkLambda_or_LetIn : constr -> rel_context -> constr
+ val push_rel_context : rel_context -> Environ.env -> Environ.env
+ val eq_constr : Evd.evar_map -> constr -> constr -> bool
+ val iter_with_binders : Evd.evar_map -> ('a -> 'a) -> ('a -> constr -> unit) -> 'a -> constr -> unit
+ val fold : Evd.evar_map -> ('a -> constr -> 'a) -> 'a -> constr -> 'a
+ val existential_type : Evd.evar_map -> existential -> types
+ val iter : Evd.evar_map -> (constr -> unit) -> constr -> unit
+ val eq_constr_universes : Evd.evar_map -> constr -> constr -> Universes.universe_constraints option
+ val eq_constr_nounivs : Evd.evar_map -> constr -> constr -> bool
+ val compare_constr : Evd.evar_map -> (constr -> constr -> bool) -> constr -> constr -> bool
+ val isApp : Evd.evar_map -> constr -> bool
+ val it_mkProd_or_LetIn : constr -> rel_context -> constr
+ val push_named : named_declaration -> Environ.env -> Environ.env
+ val destCase : Evd.evar_map -> constr -> Constr.case_info * constr * constr * constr array
+ val decompose_lam_assum : Evd.evar_map -> constr -> rel_context * constr
+ val mkConst : Names.Constant.t -> constr
+ val mkCase : Constr.case_info * constr * constr * constr array -> constr
+ val named_context : Environ.env -> named_context
+ val val_of_named_context : named_context -> Environ.named_context_val
+ val mkFix : (t, t) Constr.pfixpoint -> t
+ val decompose_prod_n_assum : Evd.evar_map -> int -> t -> rel_context * t
+ val isMeta : Evd.evar_map -> t -> bool
+
+ val destMeta : Evd.evar_map -> t -> Constr.metavariable
+
+ val map_with_binders : Evd.evar_map -> ('a -> 'a) -> ('a -> t -> t) -> 'a -> t -> t
+ val mkNamedLetIn : Names.Id.t -> constr -> types -> constr -> constr
+ val map : Evd.evar_map -> (t -> t) -> t -> t
+ val mkConstU : Names.Constant.t * EInstance.t -> t
+ val isProd : Evd.evar_map -> t -> bool
+ val mkConstructUi : (Names.inductive * EInstance.t) * int -> t
+ val isLambda : Evd.evar_map -> t -> bool
+end
+
+(* XXX: Located manually from intf *)
+module Pattern :
+sig
+
+ type case_info_pattern =
+ { cip_style : Misctypes.case_style;
+ cip_ind : Names.inductive option;
+ cip_ind_tags : bool list option; (** indicates LetIn/Lambda in arity *)
+ cip_extensible : bool (** does this match end with _ => _ ? *) }
+
+ type constr_pattern =
+ | PRef of Globnames.global_reference
+ | PVar of Names.Id.t
+ | PEvar of Evar.t * constr_pattern array
+ | PRel of int
+ | PApp of constr_pattern * constr_pattern array
+ | PSoApp of Names.Id.t * constr_pattern list
+ | PProj of Names.Projection.t * constr_pattern
+ | PLambda of Names.Name.t * constr_pattern * constr_pattern
+ | PProd of Names.Name.t * constr_pattern * constr_pattern
+ | PLetIn of Names.Name.t * constr_pattern * constr_pattern option * constr_pattern
+ | PSort of Misctypes.glob_sort
+ | PMeta of Names.Id.t option
+ | PIf of constr_pattern * constr_pattern * constr_pattern
+ | PCase of case_info_pattern * constr_pattern * constr_pattern *
+ (int * bool list * constr_pattern) list (** index of constructor, nb of args *)
+ | PFix of Term.fixpoint
+ | PCoFix of Term.cofixpoint
+
+ type constr_under_binders = Names.Id.t list * EConstr.constr
+
+ (** Types of substitutions with or w/o bound variables *)
+
+ type patvar_map = EConstr.constr Names.Id.Map.t
+ type extended_patvar_map = constr_under_binders Names.Id.Map.t
+
+end
+
module Namegen :
sig
(** *)
@@ -1584,10 +2748,82 @@ sig
Evd.evar_map -> Names.Id.t list -> Names.Name.t list -> EConstr.types -> EConstr.types
end
-module Safe_typing :
+module Termops :
sig
- type private_constants = Safe_typing.private_constants
- val mk_pure_proof : Term.constr -> Safe_typing.private_constants Entries.proof_output
+ val it_mkLambda_or_LetIn : Constr.t -> Context.Rel.t -> Constr.t
+ val local_occur_var : Evd.evar_map -> Names.Id.t -> EConstr.constr -> bool
+ val occur_var : Environ.env -> Evd.evar_map -> Names.Id.t -> EConstr.constr -> bool
+ val pr_evar_info : Evd.evar_info -> Pp.std_ppcmds
+
+ val print_constr : EConstr.constr -> Pp.std_ppcmds
+
+ (** [dependent m t] tests whether [m] is a subterm of [t] *)
+ val dependent : Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool
+
+ (** [pop c] returns a copy of [c] with decremented De Bruijn indexes *)
+ val pop : EConstr.constr -> EConstr.constr
+
+ (** Does a given term contain an existential variable? *)
+ val occur_existential : Evd.evar_map -> EConstr.constr -> bool
+
+ (** [map_constr_with_binders_left_to_right g f acc c] maps [f updated_acc] on all the immediate subterms of [c].
+ {ul {- if a given immediate subterm of [c] is not below a binder, then [updated_acc] is the same as [acc].}
+ {- if a given immediate subterm of [c] is below a binder [b], then [updated_acc] is computed as [g b acc].}} *)
+ val map_constr_with_binders_left_to_right :
+ Evd.evar_map -> (EConstr.rel_declaration -> 'a -> 'a) -> ('a -> EConstr.constr -> EConstr.constr) -> 'a -> EConstr.constr -> EConstr.constr
+
+ (** Remove the outer-most {!Term.kind_of_term.Cast} from a given term. *)
+ val strip_outer_cast : Evd.evar_map -> EConstr.constr -> EConstr.constr
+
+ (** [nb_lam] ⟦[fun (x1:t1)...(xn:tn) => c]⟧ where [c] is not an abstraction gives [n].
+ Casts are ignored. *)
+ val nb_lam : Evd.evar_map -> EConstr.constr -> int
+
+ (** [push_rel_assum env_assumtion env] adds a given {i env assumption} to the {i env context} of a given {i environment}. *)
+ val push_rel_assum : Names.Name.t * EConstr.types -> Environ.env -> Environ.env
+
+ (** [push_rels_assum env_assumptions env] adds given {i env assumptions} to the {i env context} of a given {i environment}. *)
+ val push_rels_assum : (Names.Name.t * Term.types) list -> Environ.env -> Environ.env
+
+ type meta_value_map = (Constr.metavariable * Constr.t) list
+
+ val last_arg : Evd.evar_map -> EConstr.constr -> EConstr.constr
+ val assums_of_rel_context : ('c, 't) Context.Rel.pt -> (Names.Name.t * 't) list
+ val prod_applist : Evd.evar_map -> EConstr.constr -> EConstr.constr list -> EConstr.constr
+ val nb_prod : Evd.evar_map -> EConstr.constr -> int
+ val is_section_variable : Names.Id.t -> bool
+ val ids_of_rel_context : ('c, 't) Context.Rel.pt -> Names.Id.t list
+ val subst_term : Evd.evar_map -> EConstr.constr -> EConstr.constr -> EConstr.constr
+ val global_vars_set_of_decl : Environ.env -> Evd.evar_map -> EConstr.named_declaration -> Names.Id.Set.t
+ val vars_of_env: Environ.env -> Names.Id.Set.t
+ val ids_of_named_context : ('c, 't) Context.Named.pt -> Names.Id.t list
+ val ids_of_context : Environ.env -> Names.Id.t list
+ val global_of_constr : Evd.evar_map -> EConstr.constr -> Globnames.global_reference * EConstr.EInstance.t
+ val print_named_context : Environ.env -> Pp.std_ppcmds
+ val print_constr_env : Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.std_ppcmds
+ val clear_named_body : Names.Id.t -> Environ.env -> Environ.env
+ val is_Prop : Evd.evar_map -> EConstr.constr -> bool
+ val is_global : Evd.evar_map -> Globnames.global_reference -> EConstr.constr -> bool
+
+ val eq_constr : Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool
+
+ val occur_var_in_decl :
+ Environ.env -> Evd.evar_map ->
+ Names.Id.t -> EConstr.named_declaration -> bool
+
+ val subst_meta : meta_value_map -> Constr.t -> Constr.t
+
+ val free_rels : Evd.evar_map -> EConstr.constr -> Int.Set.t
+
+ val occur_term : Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool
+ [@@ocaml.deprecated "alias of API.Termops.dependent"]
+
+ val replace_term : Evd.evar_map -> EConstr.constr -> EConstr.constr -> EConstr.constr -> EConstr.constr
+ val map_named_decl : ('a -> 'b) -> ('a, 'a) Context.Named.Declaration.pt -> ('b, 'b) Context.Named.Declaration.pt
+ val map_rel_decl : ('a -> 'b) -> ('a, 'a) Context.Rel.Declaration.pt -> ('b, 'b) Context.Rel.Declaration.pt
+ val pr_metaset : Evd.Metaset.t -> Pp.std_ppcmds
+ val pr_evar_map : ?with_univs:bool -> int option -> Evd.evar_map -> Pp.std_ppcmds
+ val pr_evar_universe_context : UState.t -> Pp.std_ppcmds
end
module Proofview_monad :
@@ -1595,56 +2831,74 @@ sig
type lazy_msg = unit -> Pp.std_ppcmds
module Info :
sig
- type tree = Proofview_monad.Info.tree
+ type tree
end
end
-(* All items in the Goal modules are deprecated. *)
-module Goal :
+module Evarutil :
sig
- type goal = Evar.t
-
- val pr_goal : goal -> Pp.std_ppcmds
-
- module V82 :
- sig
- val new_goal_with : Evd.evar_map -> goal -> Context.Named.t -> goal Evd.sigma
+ val e_new_global : Evd.evar_map ref -> Globnames.global_reference -> EConstr.constr
- val nf_hyps : Evd.evar_map -> goal -> Environ.named_context_val
+ val nf_evars_and_universes : Evd.evar_map -> Evd.evar_map * (Constr.t -> Constr.t)
+ val nf_evar : Evd.evar_map -> EConstr.constr -> EConstr.constr
+ val nf_evar_info : Evd.evar_map -> Evd.evar_info -> Evd.evar_info
- val env : Evd.evar_map -> goal -> Environ.env
+ val mk_new_meta : unit -> EConstr.constr
- val concl : Evd.evar_map -> goal -> EConstr.constr
+ (** [new_meta] is a generator of unique meta variables *)
+ val new_meta : unit -> Constr.metavariable
- val mk_goal : Evd.evar_map ->
- Environ.named_context_val ->
- EConstr.constr ->
- Evd.Store.t ->
- goal * EConstr.constr * Evd.evar_map
+ val new_Type : ?rigid:Evd.rigid -> Environ.env -> Evd.evar_map -> Evd.evar_map * EConstr.constr
+ val new_global : Evd.evar_map -> Globnames.global_reference -> Evd.evar_map * EConstr.constr
- val extra : Evd.evar_map -> goal -> Evd.Store.t
+ val new_evar :
+ Environ.env -> Evd.evar_map -> ?src:Evar_kinds.t Loc.located -> ?filter:Evd.Filter.t ->
+ ?candidates:EConstr.constr list -> ?store:Evd.Store.t ->
+ ?naming:Misctypes.intro_pattern_naming_expr ->
+ ?principal:bool -> EConstr.types -> Evd.evar_map * EConstr.constr
- val partial_solution_to : Evd.evar_map -> goal -> goal -> EConstr.constr -> Evd.evar_map
+ val new_evar_instance :
+ Environ.named_context_val -> Evd.evar_map -> EConstr.types ->
+ ?src:Evar_kinds.t Loc.located -> ?filter:Evd.Filter.t -> ?candidates:EConstr.constr list ->
+ ?store:Evd.Store.t -> ?naming:Misctypes.intro_pattern_naming_expr ->
+ ?principal:bool ->
+ EConstr.constr list -> Evd.evar_map * EConstr.constr
- val partial_solution : Evd.evar_map -> goal -> EConstr.constr -> Evd.evar_map
+ val clear_hyps_in_evi : Environ.env -> Evd.evar_map ref -> Environ.named_context_val ->
+ EConstr.types -> Names.Id.Set.t -> Environ.named_context_val * EConstr.types
- val hyps : Evd.evar_map -> goal -> Environ.named_context_val
+ type clear_dependency_error =
+ | OccurHypInSimpleClause of Names.Id.t option
+ | EvarTypingBreak of Constr.existential
- val abstract_type : Evd.evar_map -> goal -> EConstr.types
- end
+ exception ClearDependencyError of Names.Id.t * clear_dependency_error
+ val undefined_evars_of_term : Evd.evar_map -> EConstr.constr -> Evar.Set.t
+ val e_new_evar :
+ Environ.env -> Evd.evar_map ref -> ?src:Evar_kinds.t Loc.located -> ?filter:Evd.Filter.t ->
+ ?candidates:EConstr.constr list -> ?store:Evd.Store.t ->
+ ?naming:Misctypes.intro_pattern_naming_expr ->
+ ?principal:bool -> EConstr.types -> EConstr.constr
+ val new_type_evar :
+ Environ.env -> Evd.evar_map -> ?src:Evar_kinds.t Loc.located -> ?filter:Evd.Filter.t ->
+ ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> Evd.rigid ->
+ Evd.evar_map * (EConstr.constr * Sorts.t)
+ val nf_evars_universes : Evd.evar_map -> Constr.t -> Constr.t
+ val safe_evar_value : Evd.evar_map -> Term.existential -> Constr.t option
+ val evd_comb1 : (Evd.evar_map -> 'b -> Evd.evar_map * 'a) -> Evd.evar_map ref -> 'b -> 'a
end
module Proofview :
sig
- type proofview = Proofview.proofview
- type entry = Proofview.entry
- type +'a tactic = 'a Proofview.tactic
- type telescope = Proofview.telescope =
+ type proofview
+ type entry
+ type +'a tactic
+ type telescope =
| TNil of Evd.evar_map
| TCons of Environ.env * Evd.evar_map * EConstr.types * (Evd.evar_map -> EConstr.constr -> telescope)
+
module NonLogical :
sig
- type +'a t = 'a Proofview.NonLogical.t
+ type +'a t
val make : (unit -> 'a) -> 'a t
val return : 'a -> 'a t
val ( >> ) : unit t -> 'a t -> 'a t
@@ -1655,7 +2909,7 @@ sig
val print_notice : Pp.std_ppcmds -> unit t
val print_info : Pp.std_ppcmds -> unit t
val run : 'a t -> 'a
- type 'a ref = 'a Proofview.NonLogical.ref
+ type 'a ref
val ref : 'a -> 'a ref t
val ( := ) : 'a ref -> 'a -> unit t
val ( ! ) : 'a ref -> 'a t
@@ -1663,7 +2917,7 @@ sig
val catch : 'a t -> (Exninfo.iexn -> 'a t) -> 'a t
val read_line : string t
end
- val proofview : proofview -> Goal.goal list * Evd.evar_map
+ val proofview : proofview -> Evd.evar list * Evd.evar_map
val cycle : int -> unit tactic
val swap : int -> int -> unit tactic
val revgoals : unit tactic
@@ -1690,25 +2944,25 @@ sig
val shelve_unifiable : unit tactic
val apply : Environ.env -> 'a tactic -> proofview -> 'a
* proofview
- * (bool*Goal.goal list*Goal.goal list)
+ * (bool * Evd.evar list * Evd.evar list)
* Proofview_monad.Info.tree
val numgoals : int tactic
- val with_shelf : 'a tactic -> (Goal.goal list * 'a) tactic
+ val with_shelf : 'a tactic -> (Evd.evar list * 'a) tactic
module Unsafe :
sig
val tclEVARS : Evd.evar_map -> unit tactic
- val tclGETGOALS : Goal.goal list tactic
+ val tclGETGOALS : Evd.evar list tactic
- val tclSETGOALS : Goal.goal list -> unit tactic
+ val tclSETGOALS : Evd.evar list -> unit tactic
- val tclNEWGOALS : Goal.goal list -> unit tactic
+ val tclNEWGOALS : Evd.evar list -> unit tactic
end
module Goal :
sig
- type 'a t = 'a Proofview.Goal.t
+ type 'a t
val enter : ([ `LZ ] t -> unit tactic) -> unit tactic
val hyps : 'a t -> EConstr.named_context
val nf_enter : ([ `NF ] t -> unit tactic) -> unit tactic
@@ -1749,7 +3003,7 @@ end
module Ftactic :
sig
- type +'a focus = 'a Ftactic.focus
+ type +'a focus
type +'a t = 'a focus Proofview.tactic
val return : 'a -> 'a t
val run : 'a t -> ('a -> unit Proofview.tactic) -> unit Proofview.tactic
@@ -1771,61 +3025,13 @@ sig
end
end
-module Evarutil :
-sig
- val e_new_global : Evd.evar_map ref -> Globnames.global_reference -> EConstr.constr
-
- val nf_evars_and_universes : Evd.evar_map -> Evd.evar_map * (Term.constr -> Term.constr)
- val nf_evar : Evd.evar_map -> EConstr.constr -> EConstr.constr
- val nf_evar_info : Evd.evar_map -> Evd.evar_info -> Evd.evar_info
-
- val mk_new_meta : unit -> EConstr.constr
-
- (** [new_meta] is a generator of unique meta variables *)
- val new_meta : unit -> Prelude.metavariable
-
- val new_Type : ?rigid:Evd.rigid -> Environ.env -> Evd.evar_map -> Evd.evar_map * EConstr.constr
- val new_global : Evd.evar_map -> Prelude.global_reference -> Evd.evar_map * EConstr.constr
-
- val new_evar :
- Environ.env -> Evd.evar_map -> ?src:Evar_kinds.t Loc.located -> ?filter:Evd.Filter.t ->
- ?candidates:EConstr.constr list -> ?store:Evd.Store.t ->
- ?naming:Misctypes.intro_pattern_naming_expr ->
- ?principal:bool -> EConstr.types -> Evd.evar_map * EConstr.constr
-
- val new_evar_instance :
- Environ.named_context_val -> Evd.evar_map -> EConstr.types ->
- ?src:Evar_kinds.t Loc.located -> ?filter:Evd.Filter.t -> ?candidates:EConstr.constr list ->
- ?store:Evd.Store.t -> ?naming:Misctypes.intro_pattern_naming_expr ->
- ?principal:bool ->
- EConstr.constr list -> Evd.evar_map * EConstr.constr
-
- val clear_hyps_in_evi : Environ.env -> Evd.evar_map ref -> Environ.named_context_val ->
- EConstr.types -> Names.Id.Set.t -> Environ.named_context_val * EConstr.types
-
- exception ClearDependencyError of Names.Id.t * Evarutil.clear_dependency_error
- val undefined_evars_of_term : Evd.evar_map -> EConstr.constr -> Evar.Set.t
- val e_new_evar :
- Environ.env -> Evd.evar_map ref -> ?src:Evar_kinds.t Loc.located -> ?filter:Evd.Filter.t ->
- ?candidates:EConstr.constr list -> ?store:Evd.Store.t ->
- ?naming:Misctypes.intro_pattern_naming_expr ->
- ?principal:bool -> EConstr.types -> EConstr.constr
- val new_type_evar :
- Environ.env -> Evd.evar_map -> ?src:Evar_kinds.t Loc.located -> ?filter:Evd.Filter.t ->
- ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> Evd.rigid ->
- Evd.evar_map * (EConstr.constr * Sorts.t)
- val nf_evars_universes : Evd.evar_map -> Term.constr -> Term.constr
- val safe_evar_value : Evd.evar_map -> Term.existential -> Term.constr option
- val evd_comb1 : (Evd.evar_map -> 'b -> Evd.evar_map * 'a) -> Evd.evar_map ref -> 'b -> 'a
-end
-
module Geninterp :
sig
module Val :
sig
- type 'a typ = 'a Geninterp.Val.typ
- type t = Geninterp.Val.t = Dyn : 'a typ * 'a -> t
- type 'a tag = 'a Geninterp.Val.tag =
+ type 'a typ
+ type t = Dyn : 'a typ * 'a -> t
+ type 'a tag =
| Base : 'a typ -> 'a tag
| List : 'a tag -> 'a list tag
| Opt : 'a tag -> 'a option tag
@@ -1841,8 +3047,8 @@ sig
end
module TacStore :
sig
- type t = Geninterp.TacStore.t
- type 'a field = 'a Geninterp.TacStore.field
+ type t
+ type 'a field
val empty : t
val field : unit -> 'a field
val get : t -> 'a field -> 'a option
@@ -1850,9 +3056,10 @@ sig
val remove : t -> 'a field -> t
val merge : t -> t -> t
end
- type interp_sign = Geninterp.interp_sign =
- {lfun : Val.t Names.Id.Map.t;
- extra : TacStore.t }
+ type interp_sign = {
+ lfun : Val.t Names.Id.Map.t;
+ extra : TacStore.t
+ }
type ('glb, 'top) interp_fun = interp_sign -> 'glb -> 'top Ftactic.t
val register_interp0 :
('raw, 'glb, 'top) Genarg.genarg_type -> ('glb, Val.t) interp_fun -> unit
@@ -1861,443 +3068,582 @@ sig
val interp : ('raw, 'glb, 'top) Genarg.genarg_type -> ('glb, Val.t) interp_fun
end
-module Globnames :
+(* XXX: Located manually from intf *)
+module Glob_term :
sig
- type global_reference = Globnames.global_reference =
- | VarRef of Names.Id.t
- | ConstRef of Names.Constant.t
- | IndRef of Names.inductive
- | ConstructRef of Names.constructor
+ type cases_pattern_r =
+ | PatVar of Names.Name.t
+ | PatCstr of Names.constructor * cases_pattern list * Names.Name.t
+ and cases_pattern = cases_pattern_r CAst.t
+ type existential_name = Names.Id.t
+ type glob_constr_r =
+ | GRef of Globnames.global_reference * Misctypes.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 * glob_constr) list
+ | GPatVar of Evar_kinds.matching_var_kind
+ | GApp of glob_constr * glob_constr list
+ | GLambda of Names.Name.t * Decl_kinds.binding_kind * glob_constr * glob_constr
+ | GProd of Names.Name.t * Decl_kinds.binding_kind * glob_constr * glob_constr
+ | GLetIn of Names.Name.t * glob_constr * glob_constr option * glob_constr
+ | GCases of Term.case_style * glob_constr option * tomatch_tuples * cases_clauses
+ | GLetTuple of Names.Name.t list * (Names.Name.t * glob_constr option) * glob_constr * glob_constr
+ | GIf of glob_constr * (Names.Name.t * glob_constr option) * glob_constr * glob_constr
+ | GRec of fix_kind * Names.Id.t array * glob_decl list array *
+ glob_constr array * glob_constr array
+ | GSort of Misctypes.glob_sort
+ | GHole of Evar_kinds.t * Misctypes.intro_pattern_naming_expr * Genarg.glob_generic_argument option
+ | GCast of glob_constr * glob_constr Misctypes.cast_type
- type extended_global_reference = Globnames.extended_global_reference =
- | TrueGlobal of global_reference
- | SynDef of Names.KerName.t
+ and glob_constr = glob_constr_r CAst.t
- (* Long term: change implementation so that only 1 kind of order is needed.
- * Today: _env ones are fine grained, which one to pick depends. Eg.
- * - conversion rule are implemented by the non_env ones
- * - pretty printing (of user provided names/aliases) are implemented by
- * the _env ones
- *)
- module Refset : module type of struct include Globnames.Refset end
- module Refmap : module type of struct include Globnames.Refmap end
- module Refset_env : module type of struct include Globnames.Refset_env end
- module Refmap_env : module type of struct include Globnames.Refmap_env end
- module RefOrdered :
- sig
- type t = global_reference
- val compare : t -> t -> int
- end
+ and glob_decl = Names.Name.t * Decl_kinds.binding_kind * glob_constr option * glob_constr
- val pop_global_reference : global_reference -> global_reference
- val eq_gr : global_reference -> global_reference -> bool
- val destIndRef : global_reference -> Names.inductive
+ and fix_recursion_order =
+ | GStructRec
+ | GWfRec of glob_constr
+ | GMeasureRec of glob_constr * glob_constr option
- val encode_mind : Names.DirPath.t -> Names.Id.t -> Names.MutInd.t
- val encode_con : Names.DirPath.t -> Names.Id.t -> Names.Constant.t
+ and fix_kind =
+ | GFix of ((int option * fix_recursion_order) array * int)
+ | GCoFix of int
- val global_of_constr : Term.constr -> global_reference
+ and predicate_pattern =
+ Names.Name.t * (Names.inductive * Names.Name.t list) Loc.located option
- val subst_global : Mod_subst.substitution -> global_reference -> global_reference * Term.constr
- val destConstructRef : Globnames.global_reference -> Names.constructor
+ and tomatch_tuple = (glob_constr * predicate_pattern)
- val reference_of_constr : Term.constr -> global_reference
- [@@ocaml.deprecated "alias of API.Globnames.global_of_constr"]
+ and tomatch_tuples = tomatch_tuple list
+
+ and cases_clause = (Names.Id.t list * cases_pattern list * glob_constr) Loc.located
+ and cases_clauses = cases_clause list
+
+ (** A globalised term together with a closure representing the value
+ of its free variables. Intended for use when these variables are taken
+ from the Ltac environment. *)
+
+ type closure = {
+ idents : Names.Id.t Names.Id.Map.t;
+ typed : Pattern.constr_under_binders Names.Id.Map.t ;
+ untyped: closed_glob_constr Names.Id.Map.t }
+ and closed_glob_constr = {
+ closure: closure;
+ term: glob_constr }
+
+ (** Ltac variable maps *)
+ type var_map = Pattern.constr_under_binders Names.Id.Map.t
+ type uconstr_var_map = closed_glob_constr Names.Id.Map.t
+ type unbound_ltac_var_map = Geninterp.Val.t Names.Id.Map.t
+
+ type ltac_var_map = {
+ ltac_constrs : var_map;
+ (** Ltac variables bound to constrs *)
+ ltac_uconstrs : uconstr_var_map;
+ (** Ltac variables bound to untyped constrs *)
+ ltac_idents: Names.Id.t Names.Id.Map.t;
+ (** Ltac variables bound to identifiers *)
+ ltac_genargs : unbound_ltac_var_map;
+ (** Ltac variables bound to other kinds of arguments *)
+ }
- val is_global : global_reference -> Term.constr -> bool
end
-module Evar_kinds :
+module Notation_term :
sig
- type obligation_definition_status = Evar_kinds.obligation_definition_status =
- | Define of bool
- | Expand
+ type scope_name = string
+ type notation_var_instance_type =
+ | NtnTypeConstr | NtnTypeOnlyBinder | NtnTypeConstrList | NtnTypeBinderList
+ type tmp_scope_name = scope_name
- type matching_var_kind = Evar_kinds.matching_var_kind =
- | FirstOrderPatVar of Names.Id.t
- | SecondOrderPatVar of Names.Id.t
+ type subscopes = tmp_scope_name option * scope_name list
+ type notation_constr =
+ | NRef of Globnames.global_reference
+ | NVar of Names.Id.t
+ | NApp of notation_constr * notation_constr list
+ | NHole of Evar_kinds.t * Misctypes.intro_pattern_naming_expr * Genarg.glob_generic_argument option
+ | NList of Names.Id.t * Names.Id.t * notation_constr * notation_constr * bool
+ | NLambda of Names.Name.t * notation_constr * notation_constr
+ | NProd of Names.Name.t * notation_constr * notation_constr
+ | NBinderList of Names.Id.t * Names.Id.t * notation_constr * notation_constr
+ | NLetIn of Names.Name.t * notation_constr * notation_constr option * notation_constr
+ | NCases of Term.case_style * notation_constr option *
+ (notation_constr * (Names.Name.t * (Names.inductive * Names.Name.t list) option)) list *
+ (Glob_term.cases_pattern list * notation_constr) list
+ | NLetTuple of Names.Name.t list * (Names.Name.t * notation_constr option) *
+ notation_constr * notation_constr
+ | NIf of notation_constr * (Names.Name.t * notation_constr option) *
+ notation_constr * notation_constr
+ | NRec of Glob_term.fix_kind * Names.Id.t array *
+ (Names.Name.t * notation_constr option * notation_constr) list array *
+ notation_constr array * notation_constr array
+ | NSort of Misctypes.glob_sort
+ | NCast of notation_constr * notation_constr Misctypes.cast_type
+ type interpretation = (Names.Id.t * (subscopes * notation_var_instance_type)) list *
+ notation_constr
+end
- type t = Evar_kinds.t =
- | ImplicitArg of Globnames.global_reference * (int * Names.Id.t option)
- * bool (** Force inference *)
- | BinderType of Names.Name.t
- | NamedHole of Names.Id.t (* coming from some ?[id] syntax *)
- | QuestionMark of obligation_definition_status * Names.Name.t
- | CasesType of bool (* true = a subterm of the type *)
- | InternalHole
- | TomatchTypeParameter of Names.inductive * int
- | GoalEvar
- | ImpossibleCase
- | MatchingVar of matching_var_kind
- | VarInstance of Names.Id.t
- | SubEvar of Prelude.evar
+module Tactypes :
+sig
+ type glob_constr_and_expr = Glob_term.glob_constr * Constrexpr.constr_expr option
+ type glob_constr_pattern_and_expr = Names.Id.Set.t * glob_constr_and_expr * Pattern.constr_pattern
+ type 'a delayed_open = Environ.env -> Evd.evar_map -> Evd.evar_map * 'a
+ type delayed_open_constr = EConstr.constr delayed_open
+ type delayed_open_constr_with_bindings = EConstr.constr Misctypes.with_bindings delayed_open
+ type intro_pattern = delayed_open_constr Misctypes.intro_pattern_expr Loc.located
+ type intro_patterns = delayed_open_constr Misctypes.intro_pattern_expr Loc.located list
+ type intro_pattern_naming = Misctypes.intro_pattern_naming_expr Loc.located
+ type or_and_intro_pattern = delayed_open_constr Misctypes.or_and_intro_pattern_expr Loc.located
end
-module Decl_kinds :
+(* XXX: end of moved from intf *)
+
+(************************************************************************)
+(* End of modules from engine/ *)
+(************************************************************************)
+
+(************************************************************************)
+(* Modules from pretyping/ *)
+(************************************************************************)
+
+module Locusops :
sig
- type polymorphic = bool
- type cumulative_inductive_flag = bool
- type recursivity_kind = Decl_kinds.recursivity_kind =
- | Finite
- | CoFinite
- | BiFinite
- type locality = Decl_kinds.locality =
- | Discharge
- | Local
- | Global
- type definition_object_kind = Decl_kinds.definition_object_kind =
- | Definition
- | Coercion
- | SubClass
- | CanonicalStructure
- | Example
- | Fixpoint
- | CoFixpoint
- | Scheme
- | StructureComponent
- | IdentityCoercion
- | Instance
- | Method
- type theorem_kind = Decl_kinds.theorem_kind =
- | Theorem
- | Lemma
- | Fact
- | Remark
- | Property
- | Proposition
- | Corollary
- type goal_object_kind = Decl_kinds.goal_object_kind =
- | DefinitionBody of definition_object_kind
- | Proof of theorem_kind
- type goal_kind = locality * polymorphic * goal_object_kind
- type assumption_object_kind = Decl_kinds.assumption_object_kind =
- | Definitional
- | Logical
- | Conjectural
- type logical_kind = Decl_kinds.logical_kind =
- | IsAssumption of assumption_object_kind
- | IsDefinition of definition_object_kind
- | IsProof of theorem_kind
- type binding_kind = Decl_kinds.binding_kind =
- | Explicit
- | Implicit
- type private_flag = bool
- type definition_kind = locality * polymorphic * definition_object_kind
+ val clause_with_generic_occurrences : 'a Locus.clause_expr -> bool
+ val nowhere : 'a Locus.clause_expr
+ val allHypsAndConcl : 'a Locus.clause_expr
+ val is_nowhere : 'a Locus.clause_expr -> bool
+ val occurrences_map :
+ ('a list -> 'b list) -> 'a Locus.occurrences_gen -> 'b Locus.occurrences_gen
+ val convert_occs : Locus.occurrences -> bool * int list
+ val onConcl : 'a Locus.clause_expr
+ val onHyp : 'a -> 'a Locus.clause_expr
end
-module Misctypes :
+module Pretype_errors :
sig
- type evars_flag = bool
- type clear_flag = bool option
- type advanced_flag = bool
- type rec_flag = bool
+ type unification_error
+ type subterm_unification_error
+
+ type type_error = (EConstr.t, EConstr.types) Type_errors.ptype_error
+
+ type pretype_error =
+ | CantFindCaseType of EConstr.constr
+ | ActualTypeNotCoercible of EConstr.unsafe_judgment * EConstr.types * unification_error
+ | UnifOccurCheck of Evar.t * EConstr.constr
+ | UnsolvableImplicit of Evar.t * Evd.unsolvability_explanation option
+ | CannotUnify of EConstr.constr * EConstr.constr * unification_error option
+ | CannotUnifyLocal of EConstr.constr * EConstr.constr * EConstr.constr
+ | CannotUnifyBindingType of EConstr.constr * EConstr.constr
+ | CannotGeneralize of EConstr.constr
+ | NoOccurrenceFound of EConstr.constr * Names.Id.t option
+ | CannotFindWellTypedAbstraction of EConstr.constr * EConstr.constr list * (Environ.env * type_error) option
+ | WrongAbstractionType of Names.Name.t * EConstr.constr * EConstr.types * EConstr.types
+ | AbstractionOverMeta of Names.Name.t * Names.Name.t
+ | NonLinearUnification of Names.Name.t * EConstr.constr
+ | VarNotFound of Names.Id.t
+ | UnexpectedType of EConstr.constr * EConstr.constr
+ | NotProduct of EConstr.constr
+ | TypingError of type_error
+ | CannotUnifyOccurrences of subterm_unification_error
+ | UnsatisfiableConstraints of
+ (Evar.t * Evar_kinds.t) option * Evar.Set.t option
- type 'a or_by_notation = 'a Misctypes.or_by_notation =
- | AN of 'a
- | ByNotation of (string * string option) Loc.located
- type 'a or_var = 'a Misctypes.or_var =
- | ArgArg of 'a
- | ArgVar of Names.Id.t Loc.located
- type 'a and_short_name = 'a * Names.Id.t Loc.located option
- type glob_level = Misctypes.glob_level
- type 'a glob_sort_gen = 'a Misctypes.glob_sort_gen =
- | GProp
- | GSet
- | GType of 'a
- type sort_info = Names.Name.t Loc.located list
- type glob_sort = sort_info glob_sort_gen
- type 'a cast_type = 'a Misctypes.cast_type =
- | CastConv of 'a
- | CastVM of 'a
- | CastCoerce
- | CastNative of 'a
- type 'constr intro_pattern_expr = 'constr Misctypes.intro_pattern_expr =
- | IntroForthcoming of bool
- | IntroNaming of intro_pattern_naming_expr
- | IntroAction of 'constr intro_pattern_action_expr
- and intro_pattern_naming_expr = Misctypes.intro_pattern_naming_expr =
- | IntroIdentifier of Names.Id.t
- | IntroFresh of Names.Id.t
- | IntroAnonymous
- and 'constr intro_pattern_action_expr = 'constr Misctypes.intro_pattern_action_expr =
- | IntroWildcard
- | IntroOrAndPattern of 'constr or_and_intro_pattern_expr
- | IntroInjection of ('constr intro_pattern_expr) Loc.located list
- | IntroApplyOn of 'constr Loc.located * 'constr intro_pattern_expr Loc.located
- | IntroRewrite of bool
- and 'constr or_and_intro_pattern_expr = 'constr Misctypes.or_and_intro_pattern_expr =
- | IntroOrPattern of ('constr intro_pattern_expr) Loc.located list list
- | IntroAndPattern of ('constr intro_pattern_expr) Loc.located list
- type quantified_hypothesis = Misctypes.quantified_hypothesis =
- | AnonHyp of int
- | NamedHyp of Names.Id.t
- type 'a explicit_bindings = (quantified_hypothesis * 'a) Loc.located list
- type 'a bindings = 'a Misctypes.bindings =
- | ImplicitBindings of 'a list
- | ExplicitBindings of 'a explicit_bindings
- | NoBindings
- type 'a with_bindings = 'a * 'a bindings
- type 'a core_destruction_arg = 'a Misctypes.core_destruction_arg =
- | ElimOnConstr of 'a
- | ElimOnIdent of Names.Id.t Loc.located
- | ElimOnAnonHyp of int
- type inversion_kind = Misctypes.inversion_kind =
- | SimpleInversion
- | FullInversion
- | FullInversionClear
- type multi = Misctypes.multi =
- | Precisely of int
- | UpTo of int
- | RepeatStar
- | RepeatPlus
- type 'id move_location = 'id Misctypes.move_location =
- | MoveAfter of 'id
- | MoveBefore of 'id
- | MoveFirst
- | MoveLast
- type 'a destruction_arg = clear_flag * 'a core_destruction_arg
+ exception PretypeError of Environ.env * Evd.evar_map * pretype_error
+ val error_var_not_found : ?loc:Loc.t -> Names.Id.t -> 'b
+ val precatchable_exception : exn -> bool
end
-module Pattern :
+module Reductionops :
sig
- type case_info_pattern = Pattern.case_info_pattern
- type constr_pattern = Pattern.constr_pattern =
- | PRef of Globnames.global_reference
- | PVar of Names.Id.t
- | PEvar of Evar.t * constr_pattern array
- | PRel of int
- | PApp of constr_pattern * constr_pattern array
- | PSoApp of Names.Id.t * constr_pattern list
- | PProj of Names.Projection.t * constr_pattern
- | PLambda of Names.Name.t * constr_pattern * constr_pattern
- | PProd of Names.Name.t * constr_pattern * constr_pattern
- | PLetIn of Names.Name.t * constr_pattern * constr_pattern option * constr_pattern
- | PSort of Misctypes.glob_sort
- | PMeta of Names.Id.t option
- | PIf of constr_pattern * constr_pattern * constr_pattern
- | PCase of case_info_pattern * constr_pattern * constr_pattern *
- (int * bool list * constr_pattern) list (** index of constructor, nb of args *)
- | PFix of Term.fixpoint
- | PCoFix of Term.cofixpoint
- type constr_under_binders = Names.Id.t list * EConstr.constr
- type extended_patvar_map = constr_under_binders Names.Id.Map.t
- type patvar_map = EConstr.constr Names.Id.Map.t
+ type local_reduction_function = Evd.evar_map -> EConstr.constr -> EConstr.constr
+
+ type reduction_function = Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr
+
+ type local_stack_reduction_function =
+ Evd.evar_map -> EConstr.constr -> EConstr.constr * EConstr.constr list
+
+ type e_reduction_function = Environ.env -> Evd.evar_map -> EConstr.constr -> Evd.evar_map * EConstr.constr
+ type state
+
+ val clos_whd_flags : CClosure.RedFlags.reds -> reduction_function
+ val nf_beta : local_reduction_function
+ val nf_betaiota : local_reduction_function
+ val splay_prod : Environ.env -> Evd.evar_map -> EConstr.constr ->
+ (Names.Name.t * EConstr.constr) list * EConstr.constr
+ val splay_prod_n : Environ.env -> Evd.evar_map -> int -> EConstr.constr -> EConstr.rel_context * EConstr.constr
+ val whd_all : reduction_function
+ val whd_beta : local_reduction_function
+
+ val whd_betaiotazeta : local_reduction_function
+
+ val whd_betaiota_stack : local_stack_reduction_function
+
+ val clos_norm_flags : CClosure.RedFlags.reds -> reduction_function
+ val is_conv : ?reds:Names.transparent_state -> Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool
+ val beta_applist : Evd.evar_map -> EConstr.constr * EConstr.constr list -> EConstr.constr
+ val sort_of_arity : Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.ESorts.t
+ val is_conv_leq : ?reds:Names.transparent_state -> Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool
+ val whd_betaiota : local_reduction_function
+ val is_arity : Environ.env -> Evd.evar_map -> EConstr.constr -> bool
+ val nf_evar : Evd.evar_map -> EConstr.constr -> EConstr.constr
+ val nf_meta : Evd.evar_map -> EConstr.constr -> EConstr.constr
+ val hnf_prod_appvect : Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr array -> EConstr.constr
+ val pr_state : state -> Pp.std_ppcmds
+ module Stack :
+ sig
+ type 'a t
+ val pr : ('a -> Pp.std_ppcmds) -> 'a t -> Pp.std_ppcmds
+ end
+ module Cst_stack :
+ sig
+ type t
+ val pr : t -> Pp.std_ppcmds
+ end
end
-module Constrexpr :
+module Inductiveops :
sig
- type binder_kind = Constrexpr.binder_kind =
- | Default of Decl_kinds.binding_kind
- | Generalized of Decl_kinds.binding_kind * Decl_kinds.binding_kind * bool
- type explicitation = Constrexpr.explicitation =
- | ExplByPos of int * Names.Id.t option
- | ExplByName of Names.Id.t
- type sign = bool
- type raw_natural_number = string
- type prim_token = Constrexpr.prim_token =
- | Numeral of raw_natural_number * sign
- | String of string
- type notation = string
- type instance_expr = Misctypes.glob_level list
- type proj_flag = int option
- type abstraction_kind = Constrexpr.abstraction_kind =
- | AbsLambda
- | AbsPi
- type cases_pattern_expr_r = Constrexpr.cases_pattern_expr_r =
- | CPatAlias of cases_pattern_expr * Names.Id.t
- | CPatCstr of Prelude.reference
- * cases_pattern_expr list option * cases_pattern_expr list
- (** [CPatCstr (_, c, Some l1, l2)] represents (@c l1) l2 *)
- | CPatAtom of Prelude.reference option
- | CPatOr of cases_pattern_expr list
- | CPatNotation of notation * cases_pattern_notation_substitution
- * cases_pattern_expr list
- | CPatPrim of prim_token
- | CPatRecord of (Prelude.reference * cases_pattern_expr) list
- | CPatDelimiters of string * cases_pattern_expr
- | CPatCast of cases_pattern_expr * constr_expr
- and cases_pattern_expr = cases_pattern_expr_r CAst.t
+ type inductive_family
+ type inductive_type =
+ | IndType of inductive_family * EConstr.constr list
+ type constructor_summary =
+ {
+ cs_cstr : Term.pconstructor;
+ cs_params : Constr.t list;
+ cs_nargs : int;
+ cs_args : Context.Rel.t;
+ cs_concl_realargs : Constr.t array;
+ }
- and cases_pattern_notation_substitution =
- cases_pattern_expr list * cases_pattern_expr list list
+ val arities_of_constructors : Environ.env -> Term.pinductive -> Term.types array
+ val constructors_nrealargs_env : Environ.env -> Names.inductive -> int array
+ val constructor_nallargs_env : Environ.env -> Names.constructor -> int
- and constr_expr_r = Constrexpr.constr_expr_r =
- | CRef of Prelude.reference * instance_expr option
- | CFix of Names.Id.t Loc.located * fix_expr list
- | CCoFix of Names.Id.t Loc.located * cofix_expr list
- | CProdN of binder_expr list * constr_expr
- | CLambdaN of binder_expr list * constr_expr
- | CLetIn of Names.Name.t Loc.located * constr_expr * constr_expr option * constr_expr
- | CAppExpl of (proj_flag * Prelude.reference * instance_expr option) * constr_expr list
- | CApp of (proj_flag * constr_expr) *
- (constr_expr * explicitation Loc.located option) list
- | CRecord of (Prelude.reference * constr_expr) list
- | CCases of Term.case_style
- * constr_expr option
- * case_expr list
- * branch_expr list
- | CLetTuple of Names.Name.t Loc.located list * (Names.Name.t Loc.located option * constr_expr option) *
- constr_expr * constr_expr
- | CIf of constr_expr * (Names.Name.t Loc.located option * constr_expr option)
- * constr_expr * constr_expr
- | CHole of Evar_kinds.t option * Misctypes.intro_pattern_naming_expr * Genarg.raw_generic_argument option
- | CPatVar of Names.Id.t
- | CEvar of Glob_term.existential_name * (Names.Id.t * constr_expr) list
- | CSort of Misctypes.glob_sort
- | CCast of constr_expr * constr_expr Misctypes.cast_type
- | CNotation of notation * constr_notation_substitution
- | CGeneralization of Decl_kinds.binding_kind * abstraction_kind option * constr_expr
- | CPrim of prim_token
- | CDelimiters of string * constr_expr
- and constr_expr = constr_expr_r CAst.t
+ val inductive_nparams : Names.inductive -> int
- and case_expr = constr_expr * Names.Name.t Loc.located option * cases_pattern_expr option
+ val inductive_nparamdecls : Names.inductive -> int
- and branch_expr =
- (cases_pattern_expr list Loc.located list * constr_expr) Loc.located
+ val type_of_constructors : Environ.env -> Term.pinductive -> Term.types array
+ val find_mrectype : Environ.env -> Evd.evar_map -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * EConstr.constr list
+ val mis_is_recursive :
+ Names.inductive * Declarations.mutual_inductive_body * Declarations.one_inductive_body -> bool
+ val nconstructors : Names.inductive -> int
+ val find_rectype : Environ.env -> Evd.evar_map -> EConstr.types -> inductive_type
+ val get_constructors : Environ.env -> inductive_family -> constructor_summary array
+ val dest_ind_family : inductive_family -> Names.inductive Term.puniverses * Constr.t list
+ val find_inductive : Environ.env -> Evd.evar_map -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * Constr.t list
+ val type_of_inductive : Environ.env -> Term.pinductive -> Term.types
+end
- and binder_expr =
- Names.Name.t Loc.located list * binder_kind * constr_expr
+module Impargs :
+sig
+ type implicit_status
+ type implicit_side_condition
+ type implicits_list = implicit_side_condition * implicit_status list
+ type manual_explicitation = Constrexpr.explicitation * (bool * bool * bool)
+ type manual_implicits = manual_explicitation list
+ val is_status_implicit : implicit_status -> bool
+ val name_of_implicit : implicit_status -> Names.Id.t
+ val implicits_of_global : Globnames.global_reference -> implicits_list list
+ val declare_manual_implicits : bool -> Globnames.global_reference -> ?enriching:bool ->
+ manual_implicits list -> unit
+ val is_implicit_args : unit -> bool
+ val is_strict_implicit_args : unit -> bool
+ val is_contextual_implicit_args : unit -> bool
+ val make_implicit_args : bool -> unit
+ val make_strict_implicit_args : bool -> unit
+ val make_contextual_implicit_args : bool -> unit
+end
- and fix_expr =
- Names.Id.t Loc.located * (Names.Id.t Loc.located option * recursion_order_expr) *
- local_binder_expr list * constr_expr * constr_expr
+module Retyping : (* reconstruct the type of a term knowing that it was already typechecked *)
+sig
+ val get_type_of : ?polyprop:bool -> ?lax:bool -> Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.types
+ val get_sort_family_of : ?polyprop:bool -> Environ.env -> Evd.evar_map -> EConstr.types -> Sorts.family
+ val expand_projection : Environ.env -> Evd.evar_map -> Names.Projection.t -> EConstr.constr -> EConstr.constr list -> EConstr.constr
+ val get_sort_of :
+ ?polyprop:bool -> Environ.env -> Evd.evar_map -> EConstr.types -> Sorts.t
+end
- and cofix_expr =
- Names.Id.t Loc.located * local_binder_expr list * constr_expr * constr_expr
+module Find_subterm :
+sig
+ val error_invalid_occurrence : int list -> 'a
+end
- and recursion_order_expr = Constrexpr.recursion_order_expr =
- | CStructRec
- | CWfRec of constr_expr
- | CMeasureRec of constr_expr * constr_expr option
+module Evarsolve :
+sig
+ val refresh_universes :
+ ?status:Evd.rigid -> ?onlyalg:bool -> ?refreshset:bool -> bool option ->
+ Environ.env -> Evd.evar_map -> EConstr.types -> Evd.evar_map * EConstr.types
+end
- and local_binder_expr = Constrexpr.local_binder_expr =
- | CLocalAssum of Names.Name.t Loc.located list * binder_kind * constr_expr
- | CLocalDef of Names.Name.t Loc.located * constr_expr * constr_expr option
- | CLocalPattern of (cases_pattern_expr * constr_expr option) Loc.located
+module Recordops :
+sig
- and constr_notation_substitution =
- constr_expr list *
- constr_expr list list *
- local_binder_expr list list
+ type cs_pattern =
+ | Const_cs of Globnames.global_reference
+ | Prod_cs
+ | Sort_cs of Sorts.family
+ | Default_cs
- type typeclass_constraint = (Names.Name.t Loc.located * Names.Id.t Loc.located list option) * Decl_kinds.binding_kind * constr_expr
- type constr_pattern_expr = constr_expr
+ type obj_typ = {
+ o_DEF : Constr.t;
+ o_CTX : Univ.AUContext.t;
+ o_INJ : int option; (** position of trivial argument *)
+ o_TABS : Constr.t list; (** ordered *)
+ o_TPARAMS : Constr.t list; (** ordered *)
+ o_NPARAMS : int;
+ o_TCOMPS : Constr.t list }
+
+ val lookup_projections : Names.inductive -> Names.Constant.t option list
+ val lookup_canonical_conversion : (Globnames.global_reference * cs_pattern) -> Constr.t * obj_typ
+ val find_projection_nparams : Globnames.global_reference -> int
end
-module Goptions :
+module Evarconv :
sig
- type option_name = string list
- type 'a option_sig = 'a Goptions.option_sig =
- {
- optdepr : bool;
- optname : string;
- optkey : option_name;
- optread : unit -> 'a;
- optwrite : 'a -> unit
- }
- type 'a write_function = 'a Goptions.write_function
- val declare_bool_option : ?preprocess:(bool -> bool) ->
- bool option_sig -> bool write_function
- val declare_int_option : ?preprocess:(int option -> int option) ->
- int option option_sig -> int option write_function
- val declare_string_option: ?preprocess:(string -> string) ->
- string option_sig -> string write_function
- val set_bool_option_value : option_name -> bool -> unit
+ val e_conv : Environ.env -> ?ts:Names.transparent_state -> Evd.evar_map ref -> EConstr.constr -> EConstr.constr -> bool
+ val the_conv_x : Environ.env -> ?ts:Names.transparent_state -> EConstr.constr -> EConstr.constr -> Evd.evar_map -> Evd.evar_map
+ val the_conv_x_leq : Environ.env -> ?ts:Names.transparent_state -> EConstr.constr -> EConstr.constr -> Evd.evar_map -> Evd.evar_map
+ val solve_unif_constraints_with_heuristics : Environ.env -> ?ts:Names.transparent_state -> Evd.evar_map -> Evd.evar_map
end
-module Locus :
+module Typing :
sig
- type 'a occurrences_gen = 'a Locus.occurrences_gen =
- | AllOccurrences
- | AllOccurrencesBut of 'a list (** non-empty *)
- | NoOccurrences
- | OnlyOccurrences of 'a list (** non-empty *)
- type occurrences = int occurrences_gen
- type occurrences_expr = (int Misctypes.or_var) occurrences_gen
- type 'a with_occurrences = occurrences_expr * 'a
- type hyp_location_flag = Locus.hyp_location_flag =
- InHyp | InHypTypeOnly | InHypValueOnly
- type 'a hyp_location_expr = 'a with_occurrences * hyp_location_flag
- type 'id clause_expr = 'id Locus.clause_expr =
- { onhyps : 'id hyp_location_expr list option;
- concl_occs : occurrences_expr }
- type clause = Names.Id.t clause_expr
- type hyp_location = Names.Id.t * hyp_location_flag
- type goal_location = hyp_location option
+ val e_sort_of : Environ.env -> Evd.evar_map ref -> EConstr.types -> Sorts.t
+
+ val type_of : ?refresh:bool -> Environ.env -> Evd.evar_map -> EConstr.constr -> Evd.evar_map * EConstr.types
+ val e_solve_evars : Environ.env -> Evd.evar_map ref -> EConstr.constr -> EConstr.constr
+
+ val unsafe_type_of : Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.types
+
+ val e_check : Environ.env -> Evd.evar_map ref -> EConstr.constr -> EConstr.types -> unit
+
+ val e_type_of : ?refresh:bool -> Environ.env -> Evd.evar_map ref -> EConstr.constr -> EConstr.types
end
-module Genredexpr :
+module Miscops :
sig
+ val map_red_expr_gen : ('a -> 'd) -> ('b -> 'e) -> ('c -> 'f) ->
+ ('a,'b,'c) Genredexpr.red_expr_gen -> ('d,'e,'f) Genredexpr.red_expr_gen
+ val map_cast_type : ('a -> 'b) -> 'a Misctypes.cast_type -> 'b Misctypes.cast_type
+end
- (** The parsing produces initially a list of [red_atom] *)
+module Glob_ops :
+sig
+ val map_glob_constr_left_to_right : (Glob_term.glob_constr -> Glob_term.glob_constr) -> Glob_term.glob_constr -> Glob_term.glob_constr
+ val loc_of_glob_constr : Glob_term.glob_constr -> Loc.t option
+ val glob_constr_eq : Glob_term.glob_constr -> Glob_term.glob_constr -> bool
+ val bound_glob_vars : Glob_term.glob_constr -> Names.Id.Set.t
- type 'a red_atom = 'a Genredexpr.red_atom =
- | FBeta
- | FMatch
- | FFix
- | FCofix
- | FZeta
- | FConst of 'a list
- | FDeltaBut of 'a list
+ (** Conversion from glob_constr to cases pattern, if possible
- (** This list of atoms is immediately converted to a [glob_red_flag] *)
+ Take the current alias as parameter,
+ @raise Not_found if translation is impossible *)
+ val cases_pattern_of_glob_constr : Names.Name.t -> Glob_term.glob_constr -> Glob_term.cases_pattern
+ val map_glob_constr :
+ (Glob_term.glob_constr -> Glob_term.glob_constr) -> Glob_term.glob_constr -> Glob_term.glob_constr
- type 'a glob_red_flag = 'a Genredexpr.glob_red_flag = {
- rBeta : bool;
- rMatch : bool;
- rFix : bool;
- rCofix : bool;
- rZeta : bool;
- rDelta : bool; (** true = delta all but rConst; false = delta only on rConst*)
- rConst : 'a list
- }
+ val empty_lvar : Glob_term.ltac_var_map
- (** Generic kinds of reductions *)
+end
- type ('a,'b,'c) red_expr_gen = ('a,'b,'c) Genredexpr.red_expr_gen =
- | Red of bool
- | Hnf
- | Simpl of 'b glob_red_flag*('b,'c) Util.union Locus.with_occurrences option
- | Cbv of 'b glob_red_flag
- | Cbn of 'b glob_red_flag
- | Lazy of 'b glob_red_flag
- | Unfold of 'b Locus.with_occurrences list
- | Fold of 'a list
- | Pattern of 'a Locus.with_occurrences list
- | ExtraRedExpr of string
- | CbvVm of ('b,'c) Util.union Locus.with_occurrences option
- | CbvNative of ('b,'c) Util.union Locus.with_occurrences option
+module Redops :
+sig
+ val all_flags : 'a Genredexpr.glob_red_flag
+ val make_red_flag : 'a Genredexpr.red_atom list -> 'a Genredexpr.glob_red_flag
+end
- type ('a,'b,'c) may_eval = ('a,'b,'c) Genredexpr.may_eval =
- | ConstrTerm of 'a
- | ConstrEval of ('a,'b,'c) red_expr_gen * 'a
- | ConstrContext of Names.Id.t Loc.located * 'a
- | ConstrTypeOf of 'a
+module Patternops :
+sig
+ val pattern_of_glob_constr : Glob_term.glob_constr -> Names.Id.t list * Pattern.constr_pattern
+ val subst_pattern : Mod_subst.substitution -> Pattern.constr_pattern -> Pattern.constr_pattern
+ val pattern_of_constr : Environ.env -> Evd.evar_map -> Constr.t -> Pattern.constr_pattern
+ val instantiate_pattern : Environ.env ->
+ Evd.evar_map -> Pattern.extended_patvar_map ->
+ Pattern.constr_pattern -> Pattern.constr_pattern
+end
- type r_trm = Constrexpr.constr_expr
- type r_pat = Constrexpr.constr_pattern_expr
- type r_cst = Prelude.reference Misctypes.or_by_notation
- type raw_red_expr = (r_trm, r_cst, r_pat) red_expr_gen
+module Constr_matching :
+sig
+ val special_meta : Constr.metavariable
+
+ type binding_bound_vars = Names.Id.Set.t
+ type bound_ident_map = Names.Id.t Names.Id.Map.t
+ val is_matching : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> EConstr.constr -> bool
+ val extended_matches :
+ Environ.env -> Evd.evar_map -> binding_bound_vars * Pattern.constr_pattern ->
+ EConstr.constr -> bound_ident_map * Pattern.extended_patvar_map
+ exception PatternMatchingFailure
+ type matching_result =
+ { m_sub : bound_ident_map * Pattern.patvar_map;
+ m_ctx : EConstr.constr }
+ val match_subterm_gen : Environ.env -> Evd.evar_map ->
+ bool ->
+ binding_bound_vars * Pattern.constr_pattern -> EConstr.constr ->
+ matching_result IStream.t
+ val matches : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> EConstr.constr -> Pattern.patvar_map
+end
+
+module Tacred :
+sig
+ val try_red_product : Reductionops.reduction_function
+ val simpl : Reductionops.reduction_function
+ val unfoldn :
+ (Locus.occurrences * Names.evaluable_global_reference) list -> Reductionops.reduction_function
+ val hnf_constr : Reductionops.reduction_function
+ val red_product : Reductionops.reduction_function
+ val is_evaluable : Environ.env -> Names.evaluable_global_reference -> bool
+ val evaluable_of_global_reference :
+ Environ.env -> Globnames.global_reference -> Names.evaluable_global_reference
+ val error_not_evaluable : Globnames.global_reference -> 'a
+ val reduce_to_quantified_ref :
+ Environ.env -> Evd.evar_map -> Globnames.global_reference -> EConstr.types -> EConstr.types
+ val pattern_occs : (Locus.occurrences * EConstr.constr) list -> Reductionops.e_reduction_function
+ val cbv_norm_flags : CClosure.RedFlags.reds -> Reductionops.reduction_function
end
+(* XXX: Located manually from intf *)
+module Tok :
+sig
+
+ type t =
+ | KEYWORD of string
+ | PATTERNIDENT of string
+ | IDENT of string
+ | FIELD of string
+ | INT of string
+ | STRING of string
+ | LEFTQMARK
+ | BULLET of string
+ | EOI
+
+end
+
+module CLexer :
+sig
+ val add_keyword : string -> unit
+ val remove_keyword : string -> unit
+ val is_keyword : string -> bool
+ val keywords : unit -> CString.Set.t
+
+ type keyword_state
+ val set_keyword_state : keyword_state -> unit
+ val get_keyword_state : unit -> keyword_state
+
+ val check_ident : string -> unit
+ val terminal : string -> Tok.t
+
+ include Grammar.GLexerType with type te = Tok.t
+end
+
+module Extend :
+sig
+
+ type gram_assoc = NonA | RightA | LeftA
+
+ type gram_position =
+ | First
+ | Last
+ | Before of string
+ | After of string
+ | Level of string
+
+ type production_level =
+ | NextLevel
+ | NumLevel of int
+
+ type 'a entry = 'a Grammar.GMake(CLexer).Entry.e
+
+ type 'a user_symbol =
+ | Ulist1 of 'a user_symbol
+ | Ulist1sep of 'a user_symbol * string
+ | Ulist0 of 'a user_symbol
+ | Ulist0sep of 'a user_symbol * string
+ | Uopt of 'a user_symbol
+ | Uentry of 'a
+ | Uentryl of 'a * int
+
+ type ('self, 'a) symbol =
+ | Atoken : Tok.t -> ('self, string) symbol
+ | Alist1 : ('self, 'a) symbol -> ('self, 'a list) symbol
+ | Alist1sep : ('self, 'a) symbol * ('self, _) symbol -> ('self, 'a list) symbol
+ | Alist0 : ('self, 'a) symbol -> ('self, 'a list) symbol
+ | Alist0sep : ('self, 'a) symbol * ('self, _) symbol -> ('self, 'a list) symbol
+ | Aopt : ('self, 'a) symbol -> ('self, 'a option) symbol
+ | Aself : ('self, 'self) symbol
+ | Anext : ('self, 'self) symbol
+ | Aentry : 'a entry -> ('self, 'a) symbol
+ | Aentryl : 'a entry * int -> ('self, 'a) symbol
+ | Arules : 'a rules list -> ('self, 'a) symbol
+
+ and ('self, _, 'r) rule =
+ | Stop : ('self, 'r, 'r) rule
+ | Next : ('self, 'a, 'r) rule * ('self, 'b) symbol -> ('self, 'b -> 'a, 'r) rule
+
+ and ('a, 'r) norec_rule = { norec_rule : 's. ('s, 'a, 'r) rule }
+
+ and 'a rules =
+ | Rules : ('act, Loc.t -> 'a) norec_rule * 'act -> 'a rules
+
+ type ('lev,'pos) constr_entry_key_gen =
+ | ETName | ETReference | ETBigint
+ | ETBinder of bool
+ | ETConstr of ('lev * 'pos)
+ | ETPattern
+ | ETOther of string * string
+ | ETConstrList of ('lev * 'pos) * Tok.t list
+ | ETBinderList of bool * Tok.t list
+
+ type side = Left | Right
+
+ type production_position =
+ | BorderProd of side * gram_assoc option
+ | InternalProd
+
+ type constr_prod_entry_key =
+ (production_level,production_position) constr_entry_key_gen
+
+ type simple_constr_prod_entry_key =
+ (production_level,unit) constr_entry_key_gen
+
+ type 'a production_rule =
+ | Rule : ('a, 'act, Loc.t -> 'a) rule * 'act -> 'a production_rule
+
+ type 'a single_extend_statment =
+ string option *
+ (** Level *)
+ gram_assoc option *
+ (** Associativity *)
+ 'a production_rule list
+ (** Symbol list with the interpretation function *)
+
+ type 'a extend_statment =
+ gram_position option *
+ 'a single_extend_statment list
+end
+
+(* XXX: Located manually from intf *)
module Vernacexpr :
sig
+ open Misctypes
+ open Constrexpr
+ open Libnames
+
type instance_flag = bool option
type coercion_flag = bool
type inductive_flag = Decl_kinds.recursivity_kind
type lname = Names.Name.t Loc.located
type lident = Names.Id.t Loc.located
- type opacity_flag = Vernacexpr.opacity_flag =
+ type opacity_flag =
| Opaque of lident list option
| Transparent
type locality_flag = bool
- type inductive_kind = Vernacexpr.inductive_kind =
+ type inductive_kind =
| Inductive_kw | CoInductive | Variant | Record | Structure | Class of bool
- type 'a hint_info_gen = 'a Vernacexpr.hint_info_gen =
- { hint_priority : int option;
- hint_pattern : 'a option }
- type vernac_type = Vernacexpr.vernac_type =
+
+ type vernac_type =
| VtStartProof of vernac_start
| VtSideff of vernac_sideff_type
| VtQed of vernac_qed_type
@@ -2306,94 +3652,136 @@ sig
| VtQuery of vernac_part_of_script * Feedback.route_id
| VtStm of vernac_control * vernac_part_of_script
| VtUnknown
- and vernac_qed_type = Vernacexpr.vernac_qed_type =
- | VtKeep
- | VtKeepAsAxiom
- | VtDrop
+ and vernac_qed_type =
+ | VtKeep
+ | VtKeepAsAxiom
+ | VtDrop
and vernac_start = string * opacity_guarantee * Names.Id.t list
and vernac_sideff_type = Names.Id.t list
and vernac_part_of_script = bool
- and vernac_control = Vernacexpr.vernac_control =
- | VtWait
- | VtJoinDocument
- | VtBack of Stateid.t
- and opacity_guarantee = Vernacexpr.opacity_guarantee =
- | GuaranteesOpacity
- | Doesn'tGuaranteeOpacity
- and proof_step = Vernacexpr.proof_step = {
- parallel : [ `Yes of solving_tac * anon_abstracting_tac | `No ];
- proof_block_detection : proof_block_name option
- }
+ and vernac_control =
+ | VtWait
+ | VtJoinDocument
+ | VtBack of Stateid.t
+ and opacity_guarantee =
+ | GuaranteesOpacity
+ | Doesn'tGuaranteeOpacity
+ and proof_step = {
+ parallel : [ `Yes of solving_tac * anon_abstracting_tac | `No ];
+ proof_block_detection : proof_block_name option
+ }
and solving_tac = bool
and anon_abstracting_tac = bool
and proof_block_name = string
- type vernac_when = Vernacexpr.vernac_when =
- | VtNow
- | VtLater
+
+ type vernac_when =
+ | VtNow
+ | VtLater
+
type verbose_flag = bool
type obsolete_locality = bool
- type lstring = Vernacexpr.lstring
+ type lstring
type 'a with_coercion = coercion_flag * 'a
type scope_name = string
type decl_notation = lstring * Constrexpr.constr_expr * scope_name option
type constructor_expr = (lident * Constrexpr.constr_expr) with_coercion
type 'a with_notation = 'a * decl_notation list
- type local_decl_expr = Vernacexpr.local_decl_expr =
+
+ type local_decl_expr =
| AssumExpr of lname * Constrexpr.constr_expr
| DefExpr of lname * Constrexpr.constr_expr * Constrexpr.constr_expr option
+
type 'a with_priority = 'a * int option
type 'a with_instance = instance_flag * 'a
- type constructor_list_or_record_decl_expr = Vernacexpr.constructor_list_or_record_decl_expr =
+ type constructor_list_or_record_decl_expr =
| Constructors of constructor_expr list
| RecordDecl of lident option * local_decl_expr with_instance with_priority with_notation list
+
type plident = lident * lident list option
+
type inductive_expr = plident with_coercion * Constrexpr.local_binder_expr list * Constrexpr.constr_expr option * inductive_kind * constructor_list_or_record_decl_expr
- type syntax_modifier = Vernacexpr.syntax_modifier
- type class_rawexpr = Vernacexpr.class_rawexpr
- type definition_expr = Vernacexpr.definition_expr
- type hint_info_expr = Constrexpr.constr_pattern_expr hint_info_gen
- type proof_expr = Vernacexpr.proof_expr
- type proof_end = Vernacexpr.proof_end =
+ type syntax_modifier =
+ | SetItemLevel of string list * Extend.production_level
+ | SetLevel of int
+ | SetAssoc of Extend.gram_assoc
+ | SetEntryType of string * Extend.simple_constr_prod_entry_key
+ | SetOnlyParsing
+ | SetOnlyPrinting
+ | SetCompatVersion of Flags.compat_version
+ | SetFormat of string * string Loc.located
+
+ type class_rawexpr = FunClass | SortClass | RefClass of reference or_by_notation
+
+ type definition_expr =
+ | ProveBody of local_binder_expr list * constr_expr
+ | DefineBody of local_binder_expr list * Genredexpr.raw_red_expr option * constr_expr
+ * constr_expr option
+ type proof_expr =
+ plident option * (local_binder_expr list * constr_expr)
+
+ type proof_end =
| Admitted
| Proved of opacity_flag * lident option
- type inline = Vernacexpr.inline
+
type fixpoint_expr = plident * (Names.Id.t Loc.located option * Constrexpr.recursion_order_expr) * Constrexpr.local_binder_expr list * Constrexpr.constr_expr * Constrexpr.constr_expr option
- type cofixpoint_expr = Vernacexpr.cofixpoint_expr
- type scheme = Vernacexpr.scheme
- type section_subset_expr = Vernacexpr.section_subset_expr
- type module_binder = Vernacexpr.module_binder
- type vernac_argument_status = Vernacexpr.vernac_argument_status
- type vernac_implicit_status = Vernacexpr.vernac_implicit_status
- type module_ast_inl = Vernacexpr.module_ast_inl
- type 'a module_signature = 'a Vernacexpr.module_signature
+
+ type cofixpoint_expr
+
+ type scheme
+
+ type section_subset_expr
+
+ type module_binder
+
+ type vernac_argument_status
+ type vernac_implicit_status
+ type module_ast_inl
type extend_name = string * int
- type simple_binder = Vernacexpr.simple_binder
- type option_value = Vernacexpr.option_value
- type showable = Vernacexpr.showable
- type bullet = Vernacexpr.bullet
- type stm_vernac = Vernacexpr.stm_vernac
- type comment = Vernacexpr.comment
- type register_kind = Vernacexpr.register_kind
- type locatable = Vernacexpr.locatable
- type search_restriction = Vernacexpr.search_restriction
- type searchable = Vernacexpr.searchable
- type printable = Vernacexpr.printable
- type option_ref_value = Vernacexpr.option_ref_value
- type onlyparsing_flag = Vernacexpr.onlyparsing_flag
- type reference_or_constr = Vernacexpr.reference_or_constr
- type hint_mode = Vernacexpr.hint_mode
- type hints_expr = Vernacexpr.hints_expr =
+ type simple_binder
+ type option_value
+ type showable
+ type bullet
+ type stm_vernac
+ type comment
+ type register_kind
+ type locatable
+ type search_restriction
+ type searchable
+ type printable
+ type option_ref_value
+ type onlyparsing_flag
+ type reference_or_constr
+
+ type hint_mode
+
+ type 'a hint_info_gen =
+ { hint_priority : int option;
+ hint_pattern : 'a option }
+
+ type hint_info_expr = Constrexpr.constr_pattern_expr hint_info_gen
+
+ type hints_expr =
| HintsResolve of (hint_info_expr * bool * reference_or_constr) list
| HintsImmediate of reference_or_constr list
- | HintsUnfold of Prelude.reference list
- | HintsTransparency of Prelude.reference list * bool
- | HintsMode of Prelude.reference * hint_mode list
- | HintsConstructors of Prelude.reference list
+ | HintsUnfold of Libnames.reference list
+ | HintsTransparency of Libnames.reference list * bool
+ | HintsMode of Libnames.reference * hint_mode list
+ | HintsConstructors of Libnames.reference list
| HintsExtern of int * Constrexpr.constr_expr option * Genarg.raw_generic_argument
- type vernac_expr = Vernacexpr.vernac_expr =
+
+ type 'a module_signature =
+ | Enforce of 'a (** ... : T *)
+ | Check of 'a list (** ... <: T1 <: T2, possibly empty *)
+
+ type inline =
+ | NoInline
+ | DefaultInline
+ | InlineAt of int
+
+ type vernac_expr =
| VernacLoad of verbose_flag * string
| VernacTime of vernac_expr Loc.located
| VernacRedirect of string * vernac_expr Loc.located
@@ -2429,10 +3817,10 @@ sig
| VernacBeginSection of lident
| VernacEndSegment of lident
| VernacRequire of
- Prelude.reference option * bool option * Prelude.reference list
- | VernacImport of bool * Prelude.reference list
- | VernacCanonical of Prelude.reference Misctypes.or_by_notation
- | VernacCoercion of obsolete_locality * Prelude.reference Misctypes.or_by_notation *
+ Libnames.reference option * bool option * Libnames.reference list
+ | VernacImport of bool * Libnames.reference list
+ | VernacCanonical of Libnames.reference Misctypes.or_by_notation
+ | VernacCoercion of obsolete_locality * Libnames.reference Misctypes.or_by_notation *
class_rawexpr * class_rawexpr
| VernacIdentityCoercion of obsolete_locality * lident *
class_rawexpr * class_rawexpr
@@ -2442,11 +3830,11 @@ sig
Constrexpr.local_binder_expr list *
Constrexpr.typeclass_constraint *
(bool * Constrexpr.constr_expr) option *
- hint_info_expr
+ hint_info_expr
| VernacContext of Constrexpr.local_binder_expr list
| VernacDeclareInstances of
- (Prelude.reference * hint_info_expr) list
- | VernacDeclareClass of Prelude.reference
+ (Libnames.reference * hint_info_expr) list
+ | VernacDeclareClass of Libnames.reference
| VernacDeclareModule of bool option * lident *
module_binder list * module_ast_inl
| VernacDefineModule of bool option * lident * module_binder list *
@@ -2467,26 +3855,26 @@ sig
| VernacBack of int
| VernacBackTo of int
| VernacCreateHintDb of string * bool
- | VernacRemoveHints of string list * Prelude.reference list
+ | VernacRemoveHints of string list * Libnames.reference list
| VernacHints of obsolete_locality * string list * hints_expr
| VernacSyntacticDefinition of Names.Id.t Loc.located * (Names.Id.t list * Constrexpr.constr_expr) *
obsolete_locality * onlyparsing_flag
- | VernacDeclareImplicits of Prelude.reference Misctypes.or_by_notation *
+ | VernacDeclareImplicits of Libnames.reference Misctypes.or_by_notation *
(Constrexpr.explicitation * bool * bool) list list
- | VernacArguments of Prelude.reference Misctypes.or_by_notation *
+ | VernacArguments of Libnames.reference Misctypes.or_by_notation *
vernac_argument_status list *
(Names.Name.t * vernac_implicit_status) list list *
int option *
[ `ReductionDontExposeCase | `ReductionNeverUnfold | `Rename |
`ExtraScopes | `Assert | `ClearImplicits | `ClearScopes |
`DefaultImplicits ] list
- | VernacArgumentsScope of Prelude.reference Misctypes.or_by_notation *
+ | VernacArgumentsScope of Libnames.reference Misctypes.or_by_notation *
scope_name option list
| VernacReserve of simple_binder list
| VernacGeneralizable of (lident list) option
- | VernacSetOpacity of (Conv_oracle.level * Prelude.reference Misctypes.or_by_notation list)
+ | VernacSetOpacity of (Conv_oracle.level * Libnames.reference Misctypes.or_by_notation list)
| VernacSetStrategy of
- (Conv_oracle.level * Prelude.reference Misctypes.or_by_notation list) list
+ (Conv_oracle.level * Libnames.reference Misctypes.or_by_notation list) list
| VernacUnsetOption of Goptions.option_name
| VernacSetOption of Goptions.option_name * option_value
| VernacSetAppendOption of Goptions.option_name * string
@@ -2525,7 +3913,7 @@ sig
| VernacProgram of vernac_expr
| VernacPolymorphic of bool * vernac_expr
| VernacLocal of bool * vernac_expr
- and goal_selector = Vernacexpr.goal_selector =
+ and goal_selector =
| SelectNth of int
| SelectList of (int * int) list
| SelectId of Names.Id.t
@@ -2535,443 +3923,82 @@ sig
plident * Constrexpr.local_binder_expr list * Constrexpr.constr_expr option * constructor_expr list
end
-module Glob_term :
-sig
- type cases_pattern_r = Glob_term.cases_pattern_r =
- | PatVar of Names.Name.t
- | PatCstr of Names.constructor * cases_pattern list * Names.Name.t
- and cases_pattern = cases_pattern_r CAst.t
- type existential_name = Names.Id.t
- type glob_constr_r = Glob_term.glob_constr_r =
- | GRef of Globnames.global_reference * Misctypes.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 * glob_constr) list
- | GPatVar of Evar_kinds.matching_var_kind
- | GApp of glob_constr * glob_constr list
- | GLambda of Names.Name.t * Decl_kinds.binding_kind * glob_constr * glob_constr
- | GProd of Names.Name.t * Decl_kinds.binding_kind * glob_constr * glob_constr
- | GLetIn of Names.Name.t * glob_constr * glob_constr option * glob_constr
- | GCases of Term.case_style * glob_constr option * tomatch_tuples * cases_clauses
- | GLetTuple of Names.Name.t list * (Names.Name.t * glob_constr option) * glob_constr * glob_constr
- | GIf of glob_constr * (Names.Name.t * glob_constr option) * glob_constr * glob_constr
- | GRec of fix_kind * Names.Id.t array * glob_decl list array *
- glob_constr array * glob_constr array
- | GSort of Misctypes.glob_sort
- | GHole of Evar_kinds.t * Misctypes.intro_pattern_naming_expr * Genarg.glob_generic_argument option
- | GCast of glob_constr * glob_constr Misctypes.cast_type
-
- and glob_constr = glob_constr_r CAst.t
-
- and glob_decl = Names.Name.t * Decl_kinds.binding_kind * glob_constr option * glob_constr
-
- and fix_recursion_order = Glob_term.fix_recursion_order =
- | GStructRec
- | GWfRec of glob_constr
- | GMeasureRec of glob_constr * glob_constr option
-
- and fix_kind = Glob_term.fix_kind =
- | GFix of ((int option * fix_recursion_order) array * int)
- | GCoFix of int
+(* XXX: end manual intf move *)
- and predicate_pattern =
- Names.Name.t * (Names.inductive * Names.Name.t list) Loc.located option
-
- and tomatch_tuple = (glob_constr * predicate_pattern)
-
- and tomatch_tuples = tomatch_tuple list
-
- and cases_clause = (Names.Id.t list * cases_pattern list * glob_constr) Loc.located
- and cases_clauses = cases_clause list
-
- type closure = Glob_term.closure =
- { idents:Names.Id.t Names.Id.Map.t;
- typed: Pattern.constr_under_binders Names.Id.Map.t ;
- untyped:closed_glob_constr Names.Id.Map.t }
- and closed_glob_constr = Glob_term.closed_glob_constr = {
- closure: closure;
- term: glob_constr }
-
- type var_map = Pattern.constr_under_binders Names.Id.Map.t
- type uconstr_var_map = Glob_term.closed_glob_constr Names.Id.Map.t
- type unbound_ltac_var_map = Geninterp.Val.t Names.Id.Map.t
- type ltac_var_map = Glob_term.ltac_var_map = {
- ltac_constrs : var_map;
- (** Ltac variables bound to constrs *)
- ltac_uconstrs : uconstr_var_map;
- (** Ltac variables bound to untyped constrs *)
- ltac_idents: Names.Id.t Names.Id.Map.t;
- (** Ltac variables bound to identifiers *)
- ltac_genargs : unbound_ltac_var_map;
- (** Ltac variables bound to other kinds of arguments *)
- }
-end
-
-module Libnames :
-sig
- type full_path = Libnames.full_path
- val pr_path : Libnames.full_path -> Pp.std_ppcmds
- val make_path : Names.DirPath.t -> Names.Id.t -> full_path
- val eq_full_path : full_path -> full_path -> bool
- val dirpath : full_path -> Names.DirPath.t
- val path_of_string : string -> full_path
-
- type qualid = Libnames.qualid
- val make_qualid : Names.DirPath.t -> Names.Id.t -> qualid
- val qualid_eq : qualid -> qualid -> bool
- val repr_qualid : qualid -> Names.DirPath.t * Names.Id.t
- val pr_qualid : qualid -> Pp.std_ppcmds
- val string_of_qualid : qualid -> string
- val qualid_of_string : string -> qualid
- val qualid_of_path : full_path -> qualid
- val qualid_of_dirpath : Names.DirPath.t -> qualid
- val qualid_of_ident : Names.Id.t -> qualid
-
- type reference = Prelude.reference =
- | Qualid of Libnames.qualid Loc.located
- | Ident of Names.Id.t Loc.located
- val loc_of_reference : reference -> Loc.t option
- val qualid_of_reference : reference -> qualid Loc.located
- val pr_reference : reference -> Pp.std_ppcmds
-
- val is_dirpath_prefix_of : Names.DirPath.t -> Names.DirPath.t -> bool
- val split_dirpath : Names.DirPath.t -> Names.DirPath.t * Names.Id.t
- val dirpath_of_string : string -> Names.DirPath.t
- val pr_dirpath : Names.DirPath.t -> Pp.std_ppcmds
-
- val string_of_path : full_path -> string
- val basename : full_path -> Names.Id.t
-
- type object_name = Libnames.full_path * Names.KerName.t
- type object_prefix = Names.DirPath.t * (Names.ModPath.t * Names.DirPath.t)
-
- module Dirset : module type of struct include Libnames.Dirset end
- module Dirmap : module type of struct include Libnames.Dirmap end
- module Spmap : module type of struct include Libnames.Spmap end
-end
-
-module Libobject :
-sig
- type obj = Libobject.obj
- type 'a substitutivity = 'a Libobject.substitutivity =
- | Dispose
- | Substitute of 'a
- | Keep of 'a
- | Anticipate of 'a
- type 'a object_declaration = 'a Libobject.object_declaration =
- {
- object_name : string;
- cache_function : Libnames.object_name * 'a -> unit;
- load_function : int -> Libnames.object_name * 'a -> unit;
- open_function : int -> Libnames.object_name * 'a -> unit;
- classify_function : 'a -> 'a substitutivity;
- subst_function : Mod_subst.substitution * 'a -> 'a;
- discharge_function : Libnames.object_name * 'a -> 'a option;
- rebuild_function : 'a -> 'a
- }
- val declare_object : 'a object_declaration -> ('a -> obj)
- val default_object : string -> 'a object_declaration
- val object_tag : obj -> string
-end
-
-module Universes :
-sig
- type universe_binders = Universes.universe_binders
- type universe_opt_subst = Universes.universe_opt_subst
- val fresh_inductive_instance : Environ.env -> Names.inductive -> Term.pinductive Univ.in_universe_context_set
- val new_Type : Names.DirPath.t -> Term.types
- val type_of_global : Globnames.global_reference -> Term.types Univ.in_universe_context_set
- val constr_of_global : Prelude.global_reference -> Term.constr
- val new_univ_level : Names.DirPath.t -> Univ.Level.t
- val new_sort_in_family : Sorts.family -> Sorts.t
- val pr_with_global_universes : Univ.Level.t -> Pp.std_ppcmds
- val pr_universe_opt_subst : universe_opt_subst -> Pp.std_ppcmds
- type universe_constraint = Universes.universe_constraint
- module Constraints :
- sig
- type t = Universes.Constraints.t
- val pr : t -> Pp.std_ppcmds
- end
-end
-
-module Global :
-sig
- val env : unit -> Environ.env
- val lookup_mind : Names.MutInd.t -> Declarations.mutual_inductive_body
- val lookup_constant : Names.Constant.t -> Declarations.constant_body
- val lookup_module : Names.ModPath.t -> Declarations.module_body
- val lookup_modtype : Names.ModPath.t -> Declarations.module_type_body
- val lookup_inductive : Names.inductive -> Declarations.mutual_inductive_body * Declarations.one_inductive_body
- val constant_of_delta_kn : Names.KerName.t -> Names.Constant.t
- val register :
- Retroknowledge.field -> Term.constr -> Term.constr -> unit
- val env_of_context : Environ.named_context_val -> Environ.env
- val is_polymorphic : Globnames.global_reference -> bool
-
- val constr_of_global_in_context : Environ.env -> Globnames.global_reference -> Constr.t * Univ.AUContext.t
- val type_of_global_in_context : Environ.env -> Globnames.global_reference -> Constr.t * Univ.AUContext.t
-
- val current_dirpath : unit -> Names.DirPath.t
- val body_of_constant_body : Declarations.constant_body -> (Term.constr * Univ.AUContext.t) option
- val body_of_constant : Names.Constant.t -> (Term.constr * Univ.AUContext.t) option
- val add_constraints : Univ.Constraint.t -> unit
-end
-
-module Lib : sig
- type is_type = bool
- type export = bool option
- type node = Lib.node =
- | Leaf of Libobject.obj (* FIX: horrible hack (wrt. Enrico) *)
- | CompilingLibrary of Libnames.object_prefix
- | OpenedModule of is_type * export * Libnames.object_prefix * Summary.frozen
- | ClosedModule of library_segment
- | OpenedSection of Libnames.object_prefix * Summary.frozen
- | ClosedSection of library_segment
-
- and library_segment = (Libnames.object_name * node) list
-
- val current_mp : unit -> Names.ModPath.t
- val is_modtype : unit -> bool
- val is_module : unit -> bool
- val sections_are_opened : unit -> bool
- val add_anonymous_leaf : ?cache_first:bool -> Libobject.obj -> unit
- val contents : unit -> library_segment
- val cwd : unit -> Names.DirPath.t
- val add_leaf : Names.Id.t -> Libobject.obj -> Libnames.object_name
- val make_kn : Names.Id.t -> Names.KerName.t
- val make_path : Names.Id.t -> Libnames.full_path
- val discharge_con : Names.Constant.t -> Names.Constant.t
- val discharge_inductive : Names.inductive -> Names.inductive
-end
-
-module Library :
-sig
- val library_is_loaded : Names.DirPath.t -> bool
- val loaded_libraries : unit -> Names.DirPath.t list
-end
-
-module Summary :
-sig
- type marshallable = Summary.marshallable
- type 'a summary_declaration = 'a Summary.summary_declaration =
- { freeze_function : marshallable -> 'a;
- unfreeze_function : 'a -> unit;
- init_function : unit -> unit; }
- val ref : ?freeze:(marshallable -> 'a -> 'a) -> name:string -> 'a -> 'a ref
- val declare_summary : string -> 'a summary_declaration -> unit
- module Local :
- sig
- type 'a local_ref = 'a Summary.Local.local_ref
- val ref : ?freeze:('a -> 'a) -> name:string -> 'a -> 'a local_ref
- val (:=) : 'a local_ref -> 'a -> unit
- val (!) : 'a local_ref -> 'a
- end
-end
-
-module Declare :
-sig
- type internal_flag = Declare.internal_flag =
- | UserAutomaticRequest
- | InternalTacticRequest
- | UserIndividualRequest
- type constant_declaration = Safe_typing.private_constants Entries.constant_entry * Decl_kinds.logical_kind
- type section_variable_entry = Declare.section_variable_entry =
- | SectionLocalDef of Safe_typing.private_constants Entries.definition_entry
- | SectionLocalAssum of Term.types Univ.in_universe_context_set * Decl_kinds.polymorphic * bool
- type variable_declaration = Names.DirPath.t * section_variable_entry * Decl_kinds.logical_kind
- val declare_constant :
- ?internal:internal_flag -> ?local:bool -> Names.Id.t -> ?export_seff:bool -> constant_declaration -> Names.Constant.t
- val declare_universe_context : Decl_kinds.polymorphic -> Univ.ContextSet.t -> unit
- val declare_definition :
- ?internal:internal_flag -> ?opaque:bool -> ?kind:Decl_kinds.definition_object_kind ->
- ?local:bool -> ?poly:Decl_kinds.polymorphic -> Names.Id.t -> ?types:Term.constr ->
- Term.constr Univ.in_universe_context_set -> Names.Constant.t
- val definition_entry : ?fix_exn:Future.fix_exn ->
- ?opaque:bool -> ?inline:bool -> ?types:Term.types ->
- ?poly:Decl_kinds.polymorphic -> ?univs:Univ.UContext.t ->
- ?eff:Safe_typing.private_constants -> Term.constr -> Safe_typing.private_constants Entries.definition_entry
- val definition_message : Names.Id.t -> unit
- val declare_variable : Names.Id.t -> variable_declaration -> Libnames.object_name
-end
-
-module Reductionops :
-sig
- type local_reduction_function = Evd.evar_map -> EConstr.constr -> EConstr.constr
-
- type reduction_function = Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr
-
- type local_stack_reduction_function =
- Evd.evar_map -> EConstr.constr -> EConstr.constr * EConstr.constr list
-
- type e_reduction_function = Environ.env -> Evd.evar_map -> EConstr.constr -> Evd.evar_map * EConstr.constr
- type state = Reductionops.state
-
- val clos_whd_flags : CClosure.RedFlags.reds -> reduction_function
- val nf_beta : local_reduction_function
- val nf_betaiota : local_reduction_function
- val splay_prod : Environ.env -> Evd.evar_map -> EConstr.constr ->
- (Names.Name.t * EConstr.constr) list * EConstr.constr
- val splay_prod_n : Environ.env -> Evd.evar_map -> int -> EConstr.constr -> EConstr.rel_context * EConstr.constr
- val whd_all : reduction_function
- val whd_beta : local_reduction_function
-
- val whd_betaiotazeta : local_reduction_function
-
- val whd_betaiota_stack : local_stack_reduction_function
-
- val clos_norm_flags : CClosure.RedFlags.reds -> reduction_function
- val is_conv : ?reds:Names.transparent_state -> Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool
- val beta_applist : Evd.evar_map -> EConstr.constr * EConstr.constr list -> EConstr.constr
- val sort_of_arity : Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.ESorts.t
- val is_conv_leq : ?reds:Names.transparent_state -> Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool
- val whd_betaiota : local_reduction_function
- val is_arity : Environ.env -> Evd.evar_map -> EConstr.constr -> bool
- val nf_evar : Evd.evar_map -> EConstr.constr -> EConstr.constr
- val nf_meta : Evd.evar_map -> EConstr.constr -> EConstr.constr
- val hnf_prod_appvect : Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr array -> EConstr.constr
- val pr_state : state -> Pp.std_ppcmds
- module Stack :
- sig
- type 'a t = 'a Reductionops.Stack.t
- val pr : ('a -> Pp.std_ppcmds) -> 'a t -> Pp.std_ppcmds
- end
- module Cst_stack :
- sig
- type t = Reductionops.Cst_stack.t
- val pr : t -> Pp.std_ppcmds
- end
-end
-
-module Inductiveops :
+module Typeclasses :
sig
- type inductive_family = Inductiveops.inductive_family
- type inductive_type = Inductiveops.inductive_type =
- | IndType of inductive_family * EConstr.constr list
- type constructor_summary = Inductiveops.constructor_summary =
- {
- cs_cstr : Term.pconstructor;
- cs_params : Term.constr list;
- cs_nargs : int;
- cs_args : Context.Rel.t;
- cs_concl_realargs : Term.constr array;
- }
-
- val arities_of_constructors : Environ.env -> Term.pinductive -> Term.types array
- val constructors_nrealargs_env : Environ.env -> Names.inductive -> int array
- val constructor_nallargs_env : Environ.env -> Names.constructor -> int
-
- val inductive_nparams : Names.inductive -> int
-
- val inductive_nparamdecls : Names.inductive -> int
-
- val type_of_constructors : Environ.env -> Term.pinductive -> Term.types array
- val find_mrectype : Environ.env -> Evd.evar_map -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * EConstr.constr list
- val mis_is_recursive :
- Names.inductive * Declarations.mutual_inductive_body * Declarations.one_inductive_body -> bool
- val nconstructors : Names.inductive -> int
- val find_rectype : Environ.env -> Evd.evar_map -> EConstr.types -> inductive_type
- val get_constructors : Environ.env -> inductive_family -> constructor_summary array
- val dest_ind_family : inductive_family -> Names.inductive Term.puniverses * Term.constr list
- val find_inductive : Environ.env -> Evd.evar_map -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * Term.constr list
- val type_of_inductive : Environ.env -> Term.pinductive -> Term.types
-end
+ type typeclass = {
+ cl_univs : Univ.AUContext.t;
+ cl_impl : Globnames.global_reference;
+ cl_context : (Globnames.global_reference * bool) option list * Context.Rel.t;
+ cl_props : Context.Rel.t;
+ cl_projs : (Names.Name.t * (direction * Vernacexpr.hint_info_expr) option
+ * Names.Constant.t option) list;
+ cl_strict : bool;
+ cl_unique : bool;
+ }
+ and direction
-module Recordops :
-sig
- type cs_pattern = Recordops.cs_pattern =
- | Const_cs of Globnames.global_reference
- | Prod_cs
- | Sort_cs of Sorts.family
- | Default_cs
- type obj_typ = Recordops.obj_typ = {
- o_DEF : Term.constr;
- o_CTX : Univ.AUContext.t;
- o_INJ : int option; (** position of trivial argument *)
- o_TABS : Term.constr list; (** ordered *)
- o_TPARAMS : Term.constr list; (** ordered *)
- o_NPARAMS : int;
- o_TCOMPS : Term.constr list }
- val lookup_projections : Names.inductive -> Names.Constant.t option list
- val lookup_canonical_conversion : (Globnames.global_reference * cs_pattern) -> Term.constr * obj_typ
- val find_projection_nparams : Globnames.global_reference -> int
-end
+ type instance
+ type evar_filter = Evar.t -> Evar_kinds.t -> bool
-module Retyping : (* reconstruct the type of a term knowing that it was already typechecked *)
-sig
- val get_type_of : ?polyprop:bool -> ?lax:bool -> Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.types
- val get_sort_family_of : ?polyprop:bool -> Environ.env -> Evd.evar_map -> EConstr.types -> Sorts.family
- val expand_projection : Environ.env -> Evd.evar_map -> Names.Projection.t -> EConstr.constr -> EConstr.constr list -> EConstr.constr
- val get_sort_of :
- ?polyprop:bool -> Environ.env -> Evd.evar_map -> EConstr.types -> Sorts.t
+ val resolve_typeclasses : ?fast_path:bool -> ?filter:evar_filter -> ?unique:bool ->
+ ?split:bool -> ?fail:bool -> Environ.env -> Evd.evar_map -> Evd.evar_map
+ val set_resolvable : Evd.Store.t -> bool -> Evd.Store.t
+ val resolve_one_typeclass : ?unique:bool -> Environ.env -> Evd.evar_map -> EConstr.types -> Evd.evar_map * EConstr.constr
+ val class_info : Globnames.global_reference -> typeclass
+ val mark_resolvables : ?filter:evar_filter -> Evd.evar_map -> Evd.evar_map
+ val add_instance : instance -> unit
+ val new_instance : typeclass -> Vernacexpr.hint_info_expr -> bool -> Decl_kinds.polymorphic ->
+ Globnames.global_reference -> instance
end
-module Typing :
+module Classops :
sig
- val e_sort_of : Environ.env -> Evd.evar_map ref -> EConstr.types -> Sorts.t
-
- val type_of : ?refresh:bool -> Environ.env -> Evd.evar_map -> EConstr.constr -> Evd.evar_map * EConstr.types
- val e_solve_evars : Environ.env -> Evd.evar_map ref -> EConstr.constr -> EConstr.constr
-
- val unsafe_type_of : Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.types
-
- val e_check : Environ.env -> Evd.evar_map ref -> EConstr.constr -> EConstr.types -> unit
+ type coe_index
+ type inheritance_path = coe_index list
+ type cl_index
- val e_type_of : ?refresh:bool -> Environ.env -> Evd.evar_map ref -> EConstr.constr -> EConstr.types
+ val hide_coercion : Globnames.global_reference -> int option
+ val lookup_path_to_sort_from : Environ.env -> Evd.evar_map -> EConstr.types ->
+ EConstr.types * inheritance_path
+ val get_coercion_value : coe_index -> Constr.t
+ val coercions : unit -> coe_index list
+ val pr_cl_index : cl_index -> Pp.std_ppcmds
end
-module Evarsolve :
+module Detyping :
sig
- val refresh_universes :
- ?status:Evd.rigid -> ?onlyalg:bool -> ?refreshset:bool -> bool option ->
- Environ.env -> Evd.evar_map -> EConstr.types -> Evd.evar_map * EConstr.types
+ val print_universes : bool ref
+ val print_evar_arguments : bool ref
+ val detype : ?lax:bool -> bool -> Names.Id.t list -> Environ.env -> Evd.evar_map -> EConstr.constr -> Glob_term.glob_constr
+ val subst_glob_constr : Mod_subst.substitution -> Glob_term.glob_constr -> Glob_term.glob_constr
+ val set_detype_anonymous : (?loc:Loc.t -> int -> Glob_term.glob_constr) -> unit
end
-module Constr_matching :
-sig
- val special_meta : Prelude.metavariable
-
- type binding_bound_vars = Names.Id.Set.t
- type bound_ident_map = Names.Id.t Names.Id.Map.t
- val is_matching : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> EConstr.constr -> bool
- val extended_matches :
- Environ.env -> Evd.evar_map -> binding_bound_vars * Pattern.constr_pattern ->
- EConstr.constr -> bound_ident_map * Pattern.extended_patvar_map
- exception PatternMatchingFailure
- type matching_result =
- { m_sub : bound_ident_map * Pattern.patvar_map;
- m_ctx : EConstr.constr }
- val match_subterm_gen : Environ.env -> Evd.evar_map ->
- bool ->
- binding_bound_vars * Pattern.constr_pattern -> EConstr.constr ->
- matching_result IStream.t
- val matches : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> EConstr.constr -> Pattern.patvar_map
-end
-
-module Tactypes :
+module Indrec :
sig
- type glob_constr_and_expr = Glob_term.glob_constr * Constrexpr.constr_expr option
- type glob_constr_pattern_and_expr = Names.Id.Set.t * glob_constr_and_expr * Pattern.constr_pattern
- type 'a delayed_open = Environ.env -> Evd.evar_map -> Evd.evar_map * 'a
- type delayed_open_constr = EConstr.constr delayed_open
- type delayed_open_constr_with_bindings = EConstr.constr Misctypes.with_bindings delayed_open
- type intro_pattern = delayed_open_constr Misctypes.intro_pattern_expr Loc.located
- type intro_patterns = delayed_open_constr Misctypes.intro_pattern_expr Loc.located list
- type intro_pattern_naming = Misctypes.intro_pattern_naming_expr Loc.located
- type or_and_intro_pattern = delayed_open_constr Misctypes.or_and_intro_pattern_expr Loc.located
+ type dep_flag = bool
+ val lookup_eliminator : Names.inductive -> Sorts.family -> Globnames.global_reference
+ val build_case_analysis_scheme : Environ.env -> Evd.evar_map -> Term.pinductive ->
+ dep_flag -> Sorts.family -> Evd.evar_map * Constr.t
+ val make_elimination_ident : Names.Id.t -> Sorts.family -> Names.Id.t
+ val build_mutual_induction_scheme :
+ Environ.env -> Evd.evar_map -> (Term.pinductive * dep_flag * Sorts.family) list -> Evd.evar_map * Constr.t list
+ val build_case_analysis_scheme_default : Environ.env -> Evd.evar_map -> Term.pinductive ->
+ Sorts.family -> Evd.evar_map * Constr.t
end
module Pretyping :
sig
- type typing_constraint = Pretyping.typing_constraint =
- | OfType of EConstr.types
- | IsType
- | WithoutTypeConstraint
+ type typing_constraint =
+ | OfType of EConstr.types
+ | IsType
+ | WithoutTypeConstraint
type inference_hook = Environ.env -> Evd.evar_map -> Evar.t -> Evd.evar_map * EConstr.constr
- type inference_flags = Pretyping.inference_flags = {
+
+ type inference_flags = {
use_typeclasses : bool;
solve_unification_constraints : bool;
use_hook : inference_hook option;
@@ -2992,7 +4019,7 @@ sig
?expected_type:typing_constraint ->
Geninterp.interp_sign -> Glob_term.closed_glob_constr -> EConstr.constr Tactypes.delayed_open
val understand : ?flags:inference_flags -> ?expected_type:typing_constraint ->
- Environ.env -> Evd.evar_map -> Glob_term.glob_constr -> Term.constr Evd.in_evar_universe_context
+ Environ.env -> Evd.evar_map -> Glob_term.glob_constr -> Constr.t Evd.in_evar_universe_context
val check_evars : Environ.env -> Evd.evar_map -> Evd.evar_map -> EConstr.constr -> unit
val interp_elimination_sort : Misctypes.glob_sort -> Sorts.family
val register_constr_interp0 :
@@ -3004,39 +4031,30 @@ sig
Glob_term.ltac_var_map -> typing_constraint -> Glob_term.glob_constr -> Evd.evar_map * EConstr.constr
end
-module Evarconv :
-sig
- val e_conv : Environ.env -> ?ts:Names.transparent_state -> Evd.evar_map ref -> EConstr.constr -> EConstr.constr -> bool
- val the_conv_x : Environ.env -> ?ts:Names.transparent_state -> EConstr.constr -> EConstr.constr -> Evd.evar_map -> Evd.evar_map
- val the_conv_x_leq : Environ.env -> ?ts:Names.transparent_state -> EConstr.constr -> EConstr.constr -> Evd.evar_map -> Evd.evar_map
- val solve_unif_constraints_with_heuristics : Environ.env -> ?ts:Names.transparent_state -> Evd.evar_map -> Evd.evar_map
-end
-
module Unification :
sig
- type core_unify_flags = Unification.core_unify_flags =
- {
- modulo_conv_on_closed_terms : Names.transparent_state option;
- use_metas_eagerly_in_conv_on_closed_terms : bool;
- use_evars_eagerly_in_conv_on_closed_terms : bool;
- modulo_delta : Names.transparent_state;
- modulo_delta_types : Names.transparent_state;
- check_applied_meta_types : bool;
- use_pattern_unification : bool;
- use_meta_bound_pattern_unification : bool;
- frozen_evars : Evar.Set.t;
- restrict_conv_on_strict_subterms : bool;
- modulo_betaiota : bool;
- modulo_eta : bool;
- }
- type unify_flags = Unification.unify_flags =
- {
- core_unify_flags : core_unify_flags;
- merge_unify_flags : core_unify_flags;
- subterm_unify_flags : core_unify_flags;
- allow_K_in_toplevel_higher_order_unification : bool;
- resolve_evars : bool
- }
+ type core_unify_flags = {
+ modulo_conv_on_closed_terms : Names.transparent_state option;
+ use_metas_eagerly_in_conv_on_closed_terms : bool;
+ use_evars_eagerly_in_conv_on_closed_terms : bool;
+ modulo_delta : Names.transparent_state;
+ modulo_delta_types : Names.transparent_state;
+ check_applied_meta_types : bool;
+ use_pattern_unification : bool;
+ use_meta_bound_pattern_unification : bool;
+ frozen_evars : Evar.Set.t;
+ restrict_conv_on_strict_subterms : bool;
+ modulo_betaiota : bool;
+ modulo_eta : bool;
+ }
+ type unify_flags =
+ {
+ core_unify_flags : core_unify_flags;
+ merge_unify_flags : core_unify_flags;
+ subterm_unify_flags : core_unify_flags;
+ allow_K_in_toplevel_higher_order_unification : bool;
+ resolve_evars : bool
+ }
val default_no_delta_unify_flags : unit -> unify_flags
val w_unify : Environ.env -> Evd.evar_map -> Reduction.conv_pb -> ?flags:unify_flags -> EConstr.constr -> EConstr.constr -> Evd.evar_map
val elim_flags : unit -> unify_flags
@@ -3044,77 +4062,47 @@ sig
Environ.env -> Evd.evar_map -> ?flags:unify_flags -> EConstr.constr * EConstr.constr -> Evd.evar_map * EConstr.constr
end
-module Typeclasses :
+(************************************************************************)
+(* End of modules from pretyping/ *)
+(************************************************************************)
+
+(************************************************************************)
+(* Modules from interp/ *)
+(************************************************************************)
+
+module Genintern :
sig
- type typeclass = Typeclasses.typeclass = {
- cl_univs : Univ.AUContext.t;
- cl_impl : Globnames.global_reference;
- cl_context : (Globnames.global_reference * bool) option list * Context.Rel.t;
- cl_props : Context.Rel.t;
- cl_projs : (Names.Name.t * (direction * Vernacexpr.hint_info_expr) option
- * Names.Constant.t option) list;
- cl_strict : bool;
- cl_unique : bool;
+ open Genarg
+
+ module Store : Store.S
+
+ type glob_sign = {
+ ltacvars : Names.Id.Set.t;
+ genv : Environ.env;
+ extra : Store.t;
}
- and direction = Typeclasses.direction
- type instance = Typeclasses.instance
- type evar_filter = Evar.t -> Evar_kinds.t -> bool
- val resolve_typeclasses : ?fast_path:bool -> ?filter:evar_filter -> ?unique:bool ->
- ?split:bool -> ?fail:bool -> Environ.env -> Evd.evar_map -> Evd.evar_map
- val set_resolvable : Evd.Store.t -> bool -> Evd.Store.t
- val resolve_one_typeclass : ?unique:bool -> Environ.env -> Evd.evar_map -> EConstr.types -> Evd.evar_map * EConstr.constr
- val class_info : Globnames.global_reference -> typeclass
- val mark_resolvables : ?filter:evar_filter -> Evd.evar_map -> Evd.evar_map
- val add_instance : instance -> unit
- val new_instance : typeclass -> Vernacexpr.hint_info_expr -> bool -> Decl_kinds.polymorphic ->
- Globnames.global_reference -> instance
-end
-module Pretype_errors :
-sig
- type unification_error = Pretype_errors.unification_error
- type subterm_unification_error = Pretype_errors.subterm_unification_error
- type pretype_error = Pretype_errors.pretype_error =
- | CantFindCaseType of EConstr.constr
- | ActualTypeNotCoercible of EConstr.unsafe_judgment * EConstr.types * unification_error
- | UnifOccurCheck of Evar.t * EConstr.constr
- | UnsolvableImplicit of Evar.t * Evd.unsolvability_explanation option
- | CannotUnify of EConstr.constr * EConstr.constr * unification_error option
- | CannotUnifyLocal of EConstr.constr * EConstr.constr * EConstr.constr
- | CannotUnifyBindingType of EConstr.constr * EConstr.constr
- | CannotGeneralize of EConstr.constr
- | NoOccurrenceFound of EConstr.constr * Names.Id.t option
- | CannotFindWellTypedAbstraction of EConstr.constr * EConstr.constr list * (Environ.env * Pretype_errors.type_error) option
- | WrongAbstractionType of Names.Name.t * EConstr.constr * EConstr.types * EConstr.types
- | AbstractionOverMeta of Names.Name.t * Names.Name.t
- | NonLinearUnification of Names.Name.t * EConstr.constr
- | VarNotFound of Names.Id.t
- | UnexpectedType of EConstr.constr * EConstr.constr
- | NotProduct of EConstr.constr
- | TypingError of Pretype_errors.type_error
- | CannotUnifyOccurrences of subterm_unification_error
- | UnsatisfiableConstraints of
- (Evar.t * Evar_kinds.t) option * Evar.Set.t option
+ val empty_glob_sign : Environ.env -> glob_sign
- exception PretypeError of Environ.env * Evd.evar_map * pretype_error
- val error_var_not_found : ?loc:Loc.t -> Names.Id.t -> 'b
- val precatchable_exception : exn -> bool
-end
+ type ('raw, 'glb) intern_fun = glob_sign -> 'raw -> glob_sign * 'glb
-module Smartlocate :
-sig
- val locate_global_with_alias : ?head:bool -> Libnames.qualid Loc.located -> Globnames.global_reference
- val global_with_alias : ?head:bool -> Prelude.reference -> Globnames.global_reference
- val global_of_extended_global : Globnames.extended_global_reference -> Globnames.global_reference
- val loc_of_smart_reference : Prelude.reference Misctypes.or_by_notation -> Loc.t option
- val smart_global : ?head:bool -> Prelude.reference Misctypes.or_by_notation -> Globnames.global_reference
-end
-module Dumpglob :
-sig
- val add_glob : ?loc:Loc.t -> Globnames.global_reference -> unit
- val pause : unit -> unit
- val continue : unit -> unit
+ val generic_intern : (raw_generic_argument, glob_generic_argument) intern_fun
+
+ type 'glb subst_fun = Mod_subst.substitution -> 'glb -> 'glb
+ val generic_substitute : Genarg.glob_generic_argument subst_fun
+
+ type 'glb ntn_subst_fun = Tactypes.glob_constr_and_expr Names.Id.Map.t -> 'glb -> 'glb
+
+ val register_intern0 : ('raw, 'glb, 'top) genarg_type ->
+ ('raw, 'glb) intern_fun -> unit
+
+ val register_subst0 : ('raw, 'glb, 'top) genarg_type ->
+ 'glb subst_fun -> unit
+
+ val register_ntn_subst0 : ('raw, 'glb, 'top) genarg_type ->
+ 'glb ntn_subst_fun -> unit
+
end
module Stdarg :
@@ -3126,18 +4114,18 @@ sig
val wit_bool : bool Genarg.uniform_genarg_type
val wit_string : string Genarg.uniform_genarg_type
val wit_pre_ident : string Genarg.uniform_genarg_type
- val wit_global : (Prelude.reference, Globnames.global_reference Loc.located Misctypes.or_var, Globnames.global_reference) Genarg.genarg_type
+ val wit_global : (Libnames.reference, Globnames.global_reference Loc.located Misctypes.or_var, Globnames.global_reference) Genarg.genarg_type
val wit_ident : Names.Id.t Genarg.uniform_genarg_type
val wit_integer : int Genarg.uniform_genarg_type
val wit_constr : (Constrexpr.constr_expr, Tactypes.glob_constr_and_expr, EConstr.constr) Genarg.genarg_type
val wit_open_constr : (Constrexpr.constr_expr, Tactypes.glob_constr_and_expr, EConstr.constr) Genarg.genarg_type
val wit_intro_pattern : (Constrexpr.constr_expr Misctypes.intro_pattern_expr Loc.located, Tactypes.glob_constr_and_expr Misctypes.intro_pattern_expr Loc.located, Tactypes.intro_pattern) Genarg.genarg_type
val wit_int_or_var : (int Misctypes.or_var, int Misctypes.or_var, int) Genarg.genarg_type
- val wit_ref : (Prelude.reference, Globnames.global_reference Loc.located Misctypes.or_var, Globnames.global_reference) Genarg.genarg_type
+ val wit_ref : (Libnames.reference, Globnames.global_reference Loc.located Misctypes.or_var, Globnames.global_reference) Genarg.genarg_type
val wit_clause_dft_concl : (Names.Id.t Loc.located Locus.clause_expr,Names.Id.t Loc.located Locus.clause_expr,Names.Id.t Locus.clause_expr) Genarg.genarg_type
val wit_uconstr : (Constrexpr.constr_expr , Tactypes.glob_constr_and_expr, Glob_term.closed_glob_constr) Genarg.genarg_type
val wit_red_expr :
- ((Constrexpr.constr_expr,Prelude.reference Misctypes.or_by_notation,Constrexpr.constr_expr) Genredexpr.red_expr_gen,
+ ((Constrexpr.constr_expr,Libnames.reference Misctypes.or_by_notation,Constrexpr.constr_expr) Genredexpr.red_expr_gen,
(Tactypes.glob_constr_and_expr,Names.evaluable_global_reference Misctypes.and_short_name Misctypes.or_var,Tactypes.glob_constr_pattern_and_expr) Genredexpr.red_expr_gen,
(EConstr.constr,Names.evaluable_global_reference,Pattern.constr_pattern) Genredexpr.red_expr_gen) Genarg.genarg_type
val wit_quant_hyp : Misctypes.quantified_hypothesis Genarg.uniform_genarg_type
@@ -3153,94 +4141,105 @@ sig
val wit_quantified_hypothesis : Misctypes.quantified_hypothesis Genarg.uniform_genarg_type
val wit_clause : (Names.Id.t Loc.located Locus.clause_expr,Names.Id.t Loc.located Locus.clause_expr,Names.Id.t Locus.clause_expr) Genarg.genarg_type
val wit_preident : string Genarg.uniform_genarg_type
- val wit_reference : (Prelude.reference, Globnames.global_reference Loc.located Misctypes.or_var, Globnames.global_reference) Genarg.genarg_type
+ val wit_reference : (Libnames.reference, Globnames.global_reference Loc.located Misctypes.or_var, Globnames.global_reference) Genarg.genarg_type
val wit_open_constr_with_bindings :
(Constrexpr.constr_expr Misctypes.with_bindings,
Tactypes.glob_constr_and_expr Misctypes.with_bindings,
EConstr.constr Misctypes.with_bindings Tactypes.delayed_open) Genarg.genarg_type
end
-module Coqlib :
+module Constrexpr_ops :
sig
- type coq_eq_data = Coqlib.coq_eq_data = { eq : Globnames.global_reference;
- ind : Globnames.global_reference;
- refl : Globnames.global_reference;
- sym : Globnames.global_reference;
- trans: Globnames.global_reference;
- congr: Globnames.global_reference;
- }
- type coq_sigma_data = Coqlib.coq_sigma_data = {
- proj1 : Globnames.global_reference;
- proj2 : Globnames.global_reference;
- elim : Globnames.global_reference;
- intro : Globnames.global_reference;
- typ : Globnames.global_reference }
- val gen_reference : string -> string list -> string -> Globnames.global_reference
- val find_reference : string -> string list -> string -> Globnames.global_reference
- val check_required_library : string list -> unit
- val logic_module_name : string list
- val glob_true : Globnames.global_reference
- val glob_false : Globnames.global_reference
- val glob_O : Globnames.global_reference
- val glob_S : Globnames.global_reference
- val nat_path : Libnames.full_path
- val datatypes_module_name : string list
- val glob_eq : Globnames.global_reference
- val build_coq_eq_sym : Globnames.global_reference Util.delayed
- val build_coq_False : Globnames.global_reference Util.delayed
- val build_coq_not : Globnames.global_reference Util.delayed
- val build_coq_eq : Globnames.global_reference Util.delayed
- val build_coq_eq_data : coq_eq_data Util.delayed
- val path_of_O : Names.constructor
- val path_of_S : Names.constructor
- val build_prod : coq_sigma_data Util.delayed
- val build_coq_True : Globnames.global_reference Util.delayed
- val coq_iff_ref : Globnames.global_reference lazy_t
- val build_coq_iff_left_proj : Globnames.global_reference Util.delayed
- val build_coq_iff_right_proj : Globnames.global_reference Util.delayed
- val init_modules : string list list
- val build_coq_eq_refl : Globnames.global_reference Util.delayed
- val arith_modules : string list list
- val zarith_base_modules : string list list
- val gen_reference_in_modules : string -> string list list-> string -> Globnames.global_reference
- val jmeq_module_name : string list
- val coq_eq_ref : Globnames.global_reference lazy_t
- val coq_not_ref : Globnames.global_reference lazy_t
- val coq_or_ref : Globnames.global_reference lazy_t
- val build_coq_and : Globnames.global_reference Util.delayed
- val build_coq_I : Globnames.global_reference Util.delayed
- val coq_reference : string -> string list -> string -> Globnames.global_reference
+ val mkIdentC : Names.Id.t -> Constrexpr.constr_expr
+ val mkAppC : Constrexpr.constr_expr * Constrexpr.constr_expr list -> Constrexpr.constr_expr
+ val names_of_local_assums : Constrexpr.local_binder_expr list -> Names.Name.t Loc.located list
+ val coerce_reference_to_id : Libnames.reference -> Names.Id.t
+ val coerce_to_id : Constrexpr.constr_expr -> Names.Id.t Loc.located
+ val constr_loc : Constrexpr.constr_expr -> Loc.t option
+ val mkRefC : Libnames.reference -> Constrexpr.constr_expr
+ val mkLambdaC : Names.Name.t Loc.located list * Constrexpr.binder_kind * Constrexpr.constr_expr * Constrexpr.constr_expr -> Constrexpr.constr_expr
+ val default_binder_kind : Constrexpr.binder_kind
+ val mkLetInC : Names.Name.t Loc.located * Constrexpr.constr_expr * Constrexpr.constr_expr option * Constrexpr.constr_expr -> Constrexpr.constr_expr
+ val mkCProdN : ?loc:Loc.t -> Constrexpr.local_binder_expr list -> Constrexpr.constr_expr -> Constrexpr.constr_expr
end
-module Impargs :
+module Notation_ops :
sig
- type implicit_status = Impargs.implicit_status
- type implicit_side_condition = Impargs.implicit_side_condition
- type implicits_list = implicit_side_condition * implicit_status list
- type manual_explicitation = Constrexpr.explicitation * (bool * bool * bool)
- type manual_implicits = manual_explicitation list
- val is_status_implicit : implicit_status -> bool
- val name_of_implicit : implicit_status -> Names.Id.t
- val implicits_of_global : Globnames.global_reference -> implicits_list list
- val declare_manual_implicits : bool -> Globnames.global_reference -> ?enriching:bool ->
- manual_implicits list -> unit
- val is_implicit_args : unit -> bool
- val is_strict_implicit_args : unit -> bool
- val is_contextual_implicit_args : unit -> bool
- val make_implicit_args : bool -> unit
- val make_strict_implicit_args : bool -> unit
- val make_contextual_implicit_args : bool -> unit
+ val glob_constr_of_notation_constr : ?loc:Loc.t -> Notation_term.notation_constr -> Glob_term.glob_constr
+ val glob_constr_of_notation_constr_with_binders : ?loc:Loc.t ->
+ ('a -> Names.Name.t -> 'a * Names.Name.t) ->
+ ('a -> Notation_term.notation_constr -> Glob_term.glob_constr) ->
+ 'a -> Notation_term.notation_constr -> Glob_term.glob_constr
+end
+
+module Ppextend :
+sig
+
+ type precedence = int
+ type parenRelation =
+ | L | E | Any | Prec of precedence
+ type tolerability = precedence * parenRelation
+
+end
+
+module Notation :
+sig
+ type cases_pattern_status = bool
+ type required_module = Libnames.full_path * string list
+ type 'a prim_token_interpreter = ?loc:Loc.t -> 'a -> Glob_term.glob_constr
+ type 'a prim_token_uninterpreter = Glob_term.glob_constr list * (Glob_term.glob_constr -> 'a option) * cases_pattern_status
+ type delimiters = string
+ type local_scopes = Notation_term.tmp_scope_name option * Notation_term.scope_name list
+ type notation_location = (Names.DirPath.t * Names.DirPath.t) * string
+ val declare_string_interpreter : Notation_term.scope_name -> required_module ->
+ string prim_token_interpreter -> string prim_token_uninterpreter -> unit
+ val declare_numeral_interpreter : Notation_term.scope_name -> required_module ->
+ Bigint.bigint prim_token_interpreter -> Bigint.bigint prim_token_uninterpreter -> unit
+ val interp_notation_as_global_reference : ?loc:Loc.t -> (Globnames.global_reference -> bool) ->
+ Constrexpr.notation -> delimiters option -> Globnames.global_reference
+ val locate_notation : (Glob_term.glob_constr -> Pp.std_ppcmds) -> Constrexpr.notation ->
+ Notation_term.scope_name option -> Pp.std_ppcmds
+ val find_delimiters_scope : ?loc:Loc.t -> delimiters -> Notation_term.scope_name
+ val pr_scope : (Glob_term.glob_constr -> Pp.std_ppcmds) -> Notation_term.scope_name -> Pp.std_ppcmds
+ val pr_scopes : (Glob_term.glob_constr -> Pp.std_ppcmds) -> Pp.std_ppcmds
+ val interp_notation : ?loc:Loc.t -> Constrexpr.notation -> local_scopes ->
+ Notation_term.interpretation * (notation_location * Notation_term.scope_name option)
+ val uninterp_prim_token : Glob_term.glob_constr -> Notation_term.scope_name * Constrexpr.prim_token
+end
+
+module Dumpglob :
+sig
+ val add_glob : ?loc:Loc.t -> Globnames.global_reference -> unit
+ val pause : unit -> unit
+ val continue : unit -> unit
+end
+
+module Smartlocate :
+sig
+ val locate_global_with_alias : ?head:bool -> Libnames.qualid Loc.located -> Globnames.global_reference
+ val global_with_alias : ?head:bool -> Libnames.reference -> Globnames.global_reference
+ val global_of_extended_global : Globnames.extended_global_reference -> Globnames.global_reference
+ val loc_of_smart_reference : Libnames.reference Misctypes.or_by_notation -> Loc.t option
+ val smart_global : ?head:bool -> Libnames.reference Misctypes.or_by_notation -> Globnames.global_reference
+end
+
+module Topconstr :
+sig
+ val replace_vars_constr_expr :
+ Names.Id.t Names.Id.Map.t -> Constrexpr.constr_expr -> Constrexpr.constr_expr
end
module Constrintern :
sig
- type ltac_sign = Constrintern.ltac_sign = {
- ltac_vars : Names.Id.Set.t;
- ltac_bound : Names.Id.Set.t;
- ltac_extra : Genintern.Store.t;
- }
- type var_internalization_data = Constrintern.var_internalization_data
- type var_internalization_type = Constrintern.var_internalization_type =
+ type ltac_sign = {
+ ltac_vars : Names.Id.Set.t;
+ ltac_bound : Names.Id.Set.t;
+ ltac_extra : Genintern.Store.t;
+ }
+
+ type var_internalization_data
+
+ type var_internalization_type =
| Inductive of Names.Id.t list * bool
| Recursive
| Method
@@ -3262,9 +4261,9 @@ sig
Constrexpr.constr_pattern_expr -> Names.Id.t list * Pattern.constr_pattern
val intern_constr : Environ.env -> Constrexpr.constr_expr -> Glob_term.glob_constr
val for_grammar : ('a -> 'b) -> 'a -> 'b
- val interp_reference : ltac_sign -> Prelude.reference -> Glob_term.glob_constr
+ val interp_reference : ltac_sign -> Libnames.reference -> Glob_term.glob_constr
val interp_constr : Environ.env -> Evd.evar_map -> ?impls:internalization_env ->
- Constrexpr.constr_expr -> Term.constr Evd.in_evar_universe_context
+ Constrexpr.constr_expr -> Constr.t Evd.in_evar_universe_context
val interp_open_constr : Environ.env -> Evd.evar_map -> Constrexpr.constr_expr -> Evd.evar_map * EConstr.constr
val locate_reference : Libnames.qualid -> Globnames.global_reference
val interp_type : Environ.env -> Evd.evar_map -> ?impls:internalization_env ->
@@ -3279,163 +4278,155 @@ sig
val global_reference : Names.Id.t -> Globnames.global_reference
end
-module Notation_term :
-sig
- type scope_name = string
- type notation_var_instance_type = Notation_term.notation_var_instance_type =
- | NtnTypeConstr | NtnTypeOnlyBinder | NtnTypeConstrList | NtnTypeBinderList
- type tmp_scope_name = Notation_term.tmp_scope_name
- type subscopes = tmp_scope_name option * scope_name list
- type notation_constr = Notation_term.notation_constr =
- | NRef of Globnames.global_reference
- | NVar of Names.Id.t
- | NApp of notation_constr * notation_constr list
- | NHole of Evar_kinds.t * Misctypes.intro_pattern_naming_expr * Genarg.glob_generic_argument option
- | NList of Names.Id.t * Names.Id.t * notation_constr * notation_constr * bool
- | NLambda of Names.Name.t * notation_constr * notation_constr
- | NProd of Names.Name.t * notation_constr * notation_constr
- | NBinderList of Names.Id.t * Names.Id.t * notation_constr * notation_constr
- | NLetIn of Names.Name.t * notation_constr * notation_constr option * notation_constr
- | NCases of Term.case_style * notation_constr option *
- (notation_constr * (Names.Name.t * (Names.inductive * Names.Name.t list) option)) list *
- (Glob_term.cases_pattern list * notation_constr) list
- | NLetTuple of Names.Name.t list * (Names.Name.t * notation_constr option) *
- notation_constr * notation_constr
- | NIf of notation_constr * (Names.Name.t * notation_constr option) *
- notation_constr * notation_constr
- | NRec of Glob_term.fix_kind * Names.Id.t array *
- (Names.Name.t * notation_constr option * notation_constr) list array *
- notation_constr array * notation_constr array
- | NSort of Misctypes.glob_sort
- | NCast of notation_constr * notation_constr Misctypes.cast_type
- type interpretation = (Names.Id.t * (subscopes * notation_var_instance_type)) list *
- notation_constr
-end
-
-module Notation :
-sig
- type cases_pattern_status = bool
- type required_module = Libnames.full_path * string list
- type 'a prim_token_interpreter = ?loc:Loc.t -> 'a -> Glob_term.glob_constr
- type 'a prim_token_uninterpreter = Glob_term.glob_constr list * (Glob_term.glob_constr -> 'a option) * cases_pattern_status
- type delimiters = string
- type local_scopes = Notation_term.tmp_scope_name option * Notation_term.scope_name list
- type notation_location = (Names.DirPath.t * Names.DirPath.t) * string
- val declare_string_interpreter : Notation_term.scope_name -> required_module ->
- string prim_token_interpreter -> string prim_token_uninterpreter -> unit
- val declare_numeral_interpreter : Notation_term.scope_name -> required_module ->
- Bigint.bigint prim_token_interpreter -> Bigint.bigint prim_token_uninterpreter -> unit
- val interp_notation_as_global_reference : ?loc:Loc.t -> (Globnames.global_reference -> bool) ->
- Constrexpr.notation -> delimiters option -> Globnames.global_reference
- val locate_notation : (Glob_term.glob_constr -> Pp.std_ppcmds) -> Constrexpr.notation ->
- Notation_term.scope_name option -> Pp.std_ppcmds
- val find_delimiters_scope : ?loc:Loc.t -> delimiters -> Notation_term.scope_name
- val pr_scope : (Glob_term.glob_constr -> Pp.std_ppcmds) -> Notation_term.scope_name -> Pp.std_ppcmds
- val pr_scopes : (Glob_term.glob_constr -> Pp.std_ppcmds) -> Pp.std_ppcmds
- val interp_notation : ?loc:Loc.t -> Constrexpr.notation -> local_scopes ->
- Notation_term.interpretation * (notation_location * Notation_term.scope_name option)
- val uninterp_prim_token : Glob_term.glob_constr -> Notation_term.scope_name * Constrexpr.prim_token
-end
-
-module Mltop :
+module Constrextern :
sig
- val declare_cache_obj : (unit -> unit) -> string -> unit
- val add_known_plugin : (unit -> unit) -> string -> unit
- val add_known_module : string -> unit
- val module_is_known : string -> bool
+ val extern_glob_constr : Names.Id.Set.t -> Glob_term.glob_constr -> Constrexpr.constr_expr
+ val extern_glob_type : Names.Id.Set.t -> Glob_term.glob_constr -> Constrexpr.constr_expr
+ val extern_constr : ?lax:bool -> bool -> Environ.env -> Evd.evar_map -> Constr.t -> Constrexpr.constr_expr
+ val without_symbols : ('a -> 'b) -> 'a -> 'b
+ val print_universes : bool ref
+ val extern_type : bool -> Environ.env -> Evd.evar_map -> Term.types -> Constrexpr.constr_expr
+ val with_universes : ('a -> 'b) -> 'a -> 'b
+ val set_extern_reference :
+ (?loc:Loc.t -> Names.Id.Set.t -> Globnames.global_reference -> Libnames.reference) -> unit
end
-module Redexpr :
+module Declare :
sig
- type red_expr =
- (EConstr.constr, Names.evaluable_global_reference, Pattern.constr_pattern) Genredexpr.red_expr_gen
- val reduction_of_red_expr :
- Environ.env -> red_expr -> Reductionops.e_reduction_function * Term.cast_kind
- val declare_reduction : string -> Reductionops.reduction_function -> unit
-end
+ type internal_flag =
+ | UserAutomaticRequest
+ | InternalTacticRequest
+ | UserIndividualRequest
-module Tacmach :
-sig
- type tactic = Goal.goal Evd.sigma -> Goal.goal list Evd.sigma
+ type constant_declaration = Safe_typing.private_constants Entries.constant_entry * Decl_kinds.logical_kind
- type 'a sigma = 'a Evd.sigma
- [@@ocaml.deprecated "alias of API.Evd.sigma"]
+ type section_variable_entry =
+ | SectionLocalDef of Safe_typing.private_constants Entries.definition_entry
+ | SectionLocalAssum of Term.types Univ.in_universe_context_set * Decl_kinds.polymorphic * bool
- val re_sig : 'a -> Evd.evar_map -> 'a Evd.sigma
+ type variable_declaration = Names.DirPath.t * section_variable_entry * Decl_kinds.logical_kind
- val pf_reduction_of_red_expr : Goal.goal Evd.sigma -> Redexpr.red_expr -> EConstr.constr -> Evd.evar_map * EConstr.constr
+ val declare_constant :
+ ?internal:internal_flag -> ?local:bool -> Names.Id.t -> ?export_seff:bool -> constant_declaration -> Names.Constant.t
- val pf_unsafe_type_of : Goal.goal Evd.sigma -> EConstr.constr -> EConstr.types
+ val declare_universe_context : Decl_kinds.polymorphic -> Univ.ContextSet.t -> unit
- val pf_get_new_id : Names.Id.t -> Goal.goal Evd.sigma -> Names.Id.t
+ val declare_definition :
+ ?internal:internal_flag -> ?opaque:bool -> ?kind:Decl_kinds.definition_object_kind ->
+ ?local:bool -> ?poly:Decl_kinds.polymorphic -> Names.Id.t -> ?types:Constr.t ->
+ Constr.t Univ.in_universe_context_set -> Names.Constant.t
+ val definition_entry : ?fix_exn:Future.fix_exn ->
+ ?opaque:bool -> ?inline:bool -> ?types:Term.types ->
+ ?poly:Decl_kinds.polymorphic -> ?univs:Univ.UContext.t ->
+ ?eff:Safe_typing.private_constants -> Constr.t -> Safe_typing.private_constants Entries.definition_entry
+ val definition_message : Names.Id.t -> unit
+ val declare_variable : Names.Id.t -> variable_declaration -> Libnames.object_name
+end
- val pf_env : Goal.goal Evd.sigma -> Environ.env
+(************************************************************************)
+(* End of modules from interp/ *)
+(************************************************************************)
- val pf_concl : Goal.goal Evd.sigma -> EConstr.types
+(************************************************************************)
+(* Modules from proofs/ *)
+(************************************************************************)
- val pf_apply : (Environ.env -> Evd.evar_map -> 'a) -> Goal.goal Evd.sigma -> 'a
+module Miscprint :
+sig
+ val pr_or_and_intro_pattern :
+ ('a -> Pp.std_ppcmds) -> 'a Misctypes.or_and_intro_pattern_expr -> Pp.std_ppcmds
+ val pr_intro_pattern_naming : Misctypes.intro_pattern_naming_expr -> Pp.std_ppcmds
+ val pr_intro_pattern :
+ ('a -> Pp.std_ppcmds) -> 'a Misctypes.intro_pattern_expr Loc.located -> Pp.std_ppcmds
+ val pr_bindings :
+ ('a -> Pp.std_ppcmds) ->
+ ('a -> Pp.std_ppcmds) -> 'a Misctypes.bindings -> Pp.std_ppcmds
+ val pr_bindings_no_with :
+ ('a -> Pp.std_ppcmds) ->
+ ('a -> Pp.std_ppcmds) -> 'a Misctypes.bindings -> Pp.std_ppcmds
+ val pr_with_bindings :
+ ('a -> Pp.std_ppcmds) ->
+ ('a -> Pp.std_ppcmds) -> 'a * 'a Misctypes.bindings -> Pp.std_ppcmds
+end
- val pf_get_hyp : Goal.goal Evd.sigma -> Names.Id.t -> EConstr.named_declaration
- val pf_get_hyp_typ : Goal.goal Evd.sigma -> Names.Id.t -> EConstr.types
- val project : Goal.goal Evd.sigma -> Evd.evar_map
- val refine : EConstr.constr -> tactic
- val refine_no_check : EConstr.constr -> tactic
- val pf_type_of : Goal.goal Evd.sigma -> EConstr.constr -> Evd.evar_map * EConstr.types
+(* All items in the Goal modules are deprecated. *)
+module Goal :
+sig
+ type goal = Evar.t
- val pf_hyps : Goal.goal Evd.sigma -> EConstr.named_context
+ val pr_goal : goal -> Pp.std_ppcmds
- val pf_ids_of_hyps : Goal.goal Evd.sigma -> Names.Id.t list
+ module V82 :
+ sig
+ val new_goal_with : Evd.evar_map -> goal -> Context.Named.t -> goal Evd.sigma
- val pf_reduce_to_atomic_ind : Goal.goal Evd.sigma -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * EConstr.types
+ val nf_hyps : Evd.evar_map -> goal -> Environ.named_context_val
- val pf_reduce_to_quantified_ind : Goal.goal Evd.sigma -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * EConstr.types
+ val env : Evd.evar_map -> goal -> Environ.env
- val pf_eapply : (Environ.env -> Evd.evar_map -> 'a -> Evd.evar_map * 'b) ->
- Evar.t Evd.sigma -> 'a -> Evar.t Evd.sigma * 'b
+ val concl : Evd.evar_map -> goal -> EConstr.constr
- val pf_unfoldn : (Locus.occurrences * Names.evaluable_global_reference) list
- -> Goal.goal Evd.sigma -> EConstr.constr -> EConstr.constr
+ val mk_goal : Evd.evar_map ->
+ Environ.named_context_val ->
+ EConstr.constr ->
+ Evd.Store.t ->
+ goal * EConstr.constr * Evd.evar_map
- val pf_reduce : (Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr) -> Goal.goal Evd.sigma -> EConstr.constr -> EConstr.constr
+ val extra : Evd.evar_map -> goal -> Evd.Store.t
- val pf_conv_x : Goal.goal Evd.sigma -> EConstr.constr -> EConstr.constr -> bool
+ val partial_solution_to : Evd.evar_map -> goal -> goal -> EConstr.constr -> Evd.evar_map
- val pf_is_matching : Goal.goal Evd.sigma -> Pattern.constr_pattern -> EConstr.constr -> bool
+ val partial_solution : Evd.evar_map -> goal -> EConstr.constr -> Evd.evar_map
- val pf_hyps_types : Goal.goal Evd.sigma -> (Names.Id.t * EConstr.types) list
+ val hyps : Evd.evar_map -> goal -> Environ.named_context_val
- val pr_gls : Goal.goal Evd.sigma -> Pp.std_ppcmds
+ val abstract_type : Evd.evar_map -> goal -> EConstr.types
+ end
+end
- val pf_nf_betaiota : Goal.goal Evd.sigma -> EConstr.constr -> EConstr.constr
+module Evar_refiner :
+sig
+ val w_refine : Evar.t * Evd.evar_info ->
+ Pretyping.glob_constr_ltac_closure -> Evd.evar_map -> Evd.evar_map
+end
- val pf_last_hyp : Goal.goal Evd.sigma -> EConstr.named_declaration
- val pf_nth_hyp_id : Goal.goal Evd.sigma -> int -> Names.Id.t
+module Proof_type :
+sig
+ type prim_rule =
+ | Cut of bool * bool * Names.Id.t * Term.types
+ | Refine of Constr.t
- val sig_it : 'a Evd.sigma -> 'a
+ type tactic = Goal.goal Evd.sigma -> Goal.goal list Evd.sigma
+end
- module New :
- sig
- val pf_apply : (Environ.env -> Evd.evar_map -> 'a) -> 'b Proofview.Goal.t -> 'a
- val project : 'a Proofview.Goal.t -> Evd.evar_map
- val pf_unsafe_type_of : 'a Proofview.Goal.t -> EConstr.constr -> EConstr.types
- val of_old : (Goal.goal Evd.sigma -> 'a) -> [ `NF ] Proofview.Goal.t -> 'a
+module Logic :
+sig
+ type refiner_error =
+ | BadType of Constr.t * Constr.t * Constr.t
+ | UnresolvedBindings of Names.Name.t list
+ | CannotApply of Constr.t * Constr.t
+ | NotWellTyped of Constr.t
+ | NonLinearProof of Constr.t
+ | MetaInType of EConstr.constr
+ | IntroNeedsProduct
+ | DoesNotOccurIn of Constr.t * Names.Id.t
+ | NoSuchHyp of Names.Id.t
+ exception RefinerError of refiner_error
+ val catchable_exception : exn -> bool
+end
- val pf_env : 'a Proofview.Goal.t -> Environ.env
- val pf_ids_of_hyps : 'a Proofview.Goal.t -> Names.Id.t list
- val pf_concl : 'a Proofview.Goal.t -> EConstr.types
- val pf_get_new_id : Names.Id.t -> 'a Proofview.Goal.t -> Names.Id.t
- val pf_get_hyp_typ : Names.Id.t -> 'a Proofview.Goal.t -> EConstr.types
- val pf_get_type_of : 'a Proofview.Goal.t -> EConstr.constr -> EConstr.types
- val pf_global : Names.Id.t -> 'a Proofview.Goal.t -> Globnames.global_reference
- val pf_hyps_types : 'a Proofview.Goal.t -> (Names.Id.t * EConstr.types) list
- end
+module Refine :
+sig
+ val refine : typecheck:bool -> (Evd.evar_map -> Evd.evar_map * EConstr.t) -> unit Proofview.tactic
+ val solve_constraints : unit Proofview.tactic
end
module Proof :
sig
- type proof = Proof.proof
- type 'a focus_kind = 'a Proof.focus_kind
+ type proof
+ type 'a focus_kind
+
val run_tactic : Environ.env ->
unit Proofview.tactic -> proof -> proof * (bool * Proofview_monad.Info.tree)
val unshelve : proof -> proof
@@ -3456,28 +4447,31 @@ end
module Proof_global :
sig
- type proof_mode = Proof_global.proof_mode = {
+ type proof_mode = {
name : string;
set : unit -> unit ;
reset : unit -> unit
}
type proof_universes = UState.t * Universes.universe_binders option
- type proof_object = Proof_global.proof_object = {
- id : Names.Id.t;
- entries : Safe_typing.private_constants Entries.definition_entry list;
- persistence : Decl_kinds.goal_kind;
- universes: proof_universes;
- }
- type proof_ending = Proof_global.proof_ending =
+ type proof_object = {
+ id : Names.Id.t;
+ entries : Safe_typing.private_constants Entries.definition_entry list;
+ persistence : Decl_kinds.goal_kind;
+ universes: proof_universes;
+ }
+
+ type proof_ending =
| Admitted of Names.Id.t * Decl_kinds.goal_kind * Entries.parameter_entry *
proof_universes
| Proved of Vernacexpr.opacity_flag *
Vernacexpr.lident option *
proof_object
- type proof_terminator = Proof_global.proof_terminator
- type lemma_possible_guards = Proof_global.lemma_possible_guards
- type universe_binders = Proof_global.universe_binders
+
+ type proof_terminator
+ type lemma_possible_guards
+ type universe_binders
type closed_proof = proof_object * proof_terminator
+
val make_terminator : (proof_ending -> unit) -> proof_terminator
val start_dependent_proof :
Names.Id.t -> ?pl:universe_binders -> Decl_kinds.goal_kind ->
@@ -3498,382 +4492,361 @@ sig
val get_current_proof_name : unit -> Names.Id.t
end
-module Nametab :
-sig
- exception GlobalizationError of Libnames.qualid
-
- type ltac_constant = Names.KerName.t
-
- val global : Libnames.reference -> Globnames.global_reference
- val global_of_path : Libnames.full_path -> Globnames.global_reference
- val shortest_qualid_of_global : Names.Id.Set.t -> Globnames.global_reference -> Libnames.qualid
- val path_of_global : Globnames.global_reference -> Libnames.full_path
- val locate_extended : Libnames.qualid -> Globnames.extended_global_reference
- val full_name_module : Libnames.qualid -> Names.DirPath.t
- val locate_tactic : Libnames.qualid -> Names.KerName.t
- val pr_global_env : Names.Id.Set.t -> Globnames.global_reference -> Pp.std_ppcmds
- val shortest_qualid_of_tactic : Names.KerName.t -> Libnames.qualid
- val basename_of_global : Globnames.global_reference -> Names.Id.t
-
- type visibility = Nametab.visibility =
- | Until of int
- | Exactly of int
-
- val push_tactic : visibility -> Libnames.full_path -> Names.KerName.t -> unit
- val error_global_not_found : ?loc:Loc.t -> Libnames.qualid -> 'a
- val shortest_qualid_of_module : Names.ModPath.t -> Libnames.qualid
- val dirpath_of_module : Names.ModPath.t -> Names.DirPath.t
- val locate_module : Libnames.qualid -> Names.ModPath.t
- val dirpath_of_global : Globnames.global_reference -> Names.DirPath.t
- val locate : Libnames.qualid -> Globnames.global_reference
- val locate_constant : Libnames.qualid -> Names.Constant.t
-end
-
-module Ppextend :
+module Redexpr :
sig
- type precedence = int
- type parenRelation = Ppextend.parenRelation =
- | L | E | Any | Prec of precedence
- type tolerability = precedence * parenRelation
+ type red_expr =
+ (EConstr.constr, Names.evaluable_global_reference, Pattern.constr_pattern) Genredexpr.red_expr_gen
+ val reduction_of_red_expr :
+ Environ.env -> red_expr -> Reductionops.e_reduction_function * Constr.cast_kind
+ val declare_reduction : string -> Reductionops.reduction_function -> unit
end
module Refiner :
sig
val project : 'a Evd.sigma -> Evd.evar_map
-
+
val unpackage : 'a Evd.sigma -> Evd.evar_map ref * 'a
val repackage : Evd.evar_map ref -> 'a -> 'a Evd.sigma
- val tclSHOWHYPS : Tacmach.tactic -> Tacmach.tactic
+ val tclSHOWHYPS : Proof_type.tactic -> Proof_type.tactic
exception FailError of int * Pp.std_ppcmds Lazy.t
- val tclEVARS : Evd.evar_map -> Tacmach.tactic
- val tclMAP : ('a -> Tacmach.tactic) -> 'a list -> Tacmach.tactic
- val tclREPEAT : Tacmach.tactic -> Tacmach.tactic
- val tclORELSE : Tacmach.tactic -> Tacmach.tactic -> Tacmach.tactic
- val tclFAIL : int -> Pp.std_ppcmds -> Tacmach.tactic
- val tclIDTAC : Tacmach.tactic
- val tclTHEN : Tacmach.tactic -> Tacmach.tactic -> Tacmach.tactic
- val tclTHENLIST : Tacmach.tactic list -> Tacmach.tactic
- val tclTRY : Tacmach.tactic -> Tacmach.tactic
- val tclAT_LEAST_ONCE : Tacmach.tactic -> Tacmach.tactic
+ val tclEVARS : Evd.evar_map -> Proof_type.tactic
+ val tclMAP : ('a -> Proof_type.tactic) -> 'a list -> Proof_type.tactic
+ val tclREPEAT : Proof_type.tactic -> Proof_type.tactic
+ val tclORELSE : Proof_type.tactic -> Proof_type.tactic -> Proof_type.tactic
+ val tclFAIL : int -> Pp.std_ppcmds -> Proof_type.tactic
+ val tclIDTAC : Proof_type.tactic
+ val tclTHEN : Proof_type.tactic -> Proof_type.tactic -> Proof_type.tactic
+ val tclTHENLIST : Proof_type.tactic list -> Proof_type.tactic
+ val tclTRY : Proof_type.tactic -> Proof_type.tactic
+ val tclAT_LEAST_ONCE : Proof_type.tactic -> Proof_type.tactic
end
-module Termops :
+module Tacmach :
sig
- val it_mkLambda_or_LetIn : Term.constr -> Context.Rel.t -> Term.constr
- val local_occur_var : Evd.evar_map -> Names.Id.t -> EConstr.constr -> bool
- val occur_var : Environ.env -> Evd.evar_map -> Names.Id.t -> EConstr.constr -> bool
- val pr_evar_info : Evd.evar_info -> Pp.std_ppcmds
- val print_constr : EConstr.constr -> Pp.std_ppcmds
+ type tactic = Proof_type.tactic
- (** [dependent m t] tests whether [m] is a subterm of [t] *)
- val dependent : Prelude.evar_map -> EConstr.constr -> EConstr.constr -> bool
+ type 'a sigma = 'a Evd.sigma
+ [@@ocaml.deprecated "alias of API.Evd.sigma"]
- (** [pop c] returns a copy of [c] with decremented De Bruijn indexes *)
- val pop : EConstr.constr -> EConstr.constr
+ val re_sig : 'a -> Evd.evar_map -> 'a Evd.sigma
- (** Does a given term contain an existential variable? *)
- val occur_existential : Prelude.evar_map -> EConstr.constr -> bool
+ val pf_reduction_of_red_expr : Goal.goal Evd.sigma -> Redexpr.red_expr -> EConstr.constr -> Evd.evar_map * EConstr.constr
- (** [map_constr_with_binders_left_to_right g f acc c] maps [f updated_acc] on all the immediate subterms of [c].
- {ul {- if a given immediate subterm of [c] is not below a binder, then [updated_acc] is the same as [acc].}
- {- if a given immediate subterm of [c] is below a binder [b], then [updated_acc] is computed as [g b acc].}} *)
- val map_constr_with_binders_left_to_right :
- Prelude.evar_map -> (EConstr.rel_declaration -> 'a -> 'a) -> ('a -> EConstr.constr -> EConstr.constr) -> 'a -> EConstr.constr -> EConstr.constr
+ val pf_unsafe_type_of : Goal.goal Evd.sigma -> EConstr.constr -> EConstr.types
- (** Remove the outer-most {!Term.kind_of_term.Cast} from a given term. *)
- val strip_outer_cast : Prelude.evar_map -> EConstr.constr -> EConstr.constr
+ val pf_get_new_id : Names.Id.t -> Goal.goal Evd.sigma -> Names.Id.t
- (** [nb_lam] ⟦[fun (x1:t1)...(xn:tn) => c]⟧ where [c] is not an abstraction gives [n].
- Casts are ignored. *)
- val nb_lam : Prelude.evar_map -> EConstr.constr -> int
+ val pf_env : Goal.goal Evd.sigma -> Environ.env
- (** [push_rel_assum env_assumtion env] adds a given {i env assumption} to the {i env context} of a given {i environment}. *)
- val push_rel_assum : Names.Name.t * EConstr.types -> Environ.env -> Environ.env
+ val pf_concl : Goal.goal Evd.sigma -> EConstr.types
- (** [push_rels_assum env_assumptions env] adds given {i env assumptions} to the {i env context} of a given {i environment}. *)
- val push_rels_assum : (Names.Name.t * Term.types) list -> Environ.env -> Environ.env
+ val pf_apply : (Environ.env -> Evd.evar_map -> 'a) -> Goal.goal Evd.sigma -> 'a
- type meta_value_map = Prelude.meta_value_map
+ val pf_get_hyp : Goal.goal Evd.sigma -> Names.Id.t -> EConstr.named_declaration
+ val pf_get_hyp_typ : Goal.goal Evd.sigma -> Names.Id.t -> EConstr.types
+ val project : Goal.goal Evd.sigma -> Evd.evar_map
+ val refine : EConstr.constr -> Proof_type.tactic
+ val refine_no_check : EConstr.constr -> Proof_type.tactic
+ val pf_type_of : Goal.goal Evd.sigma -> EConstr.constr -> Evd.evar_map * EConstr.types
- val last_arg : Evd.evar_map -> EConstr.constr -> EConstr.constr
- val assums_of_rel_context : ('c, 't) Context.Rel.pt -> (Names.Name.t * 't) list
- val prod_applist : Evd.evar_map -> EConstr.constr -> EConstr.constr list -> EConstr.constr
- val nb_prod : Evd.evar_map -> EConstr.constr -> int
- val is_section_variable : Names.Id.t -> bool
- val ids_of_rel_context : ('c, 't) Context.Rel.pt -> Names.Id.t list
- val subst_term : Evd.evar_map -> EConstr.constr -> EConstr.constr -> EConstr.constr
- val global_vars_set_of_decl : Environ.env -> Evd.evar_map -> EConstr.named_declaration -> Names.Id.Set.t
- val vars_of_env: Environ.env -> Names.Id.Set.t
- val ids_of_named_context : ('c, 't) Context.Named.pt -> Names.Id.t list
- val ids_of_context : Environ.env -> Names.Id.t list
- val global_of_constr : Evd.evar_map -> EConstr.constr -> Globnames.global_reference * EConstr.EInstance.t
- val print_named_context : Environ.env -> Pp.std_ppcmds
- val print_constr_env : Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.std_ppcmds
- val clear_named_body : Names.Id.t -> Environ.env -> Environ.env
- val is_Prop : Evd.evar_map -> EConstr.constr -> bool
- val is_global : Evd.evar_map -> Globnames.global_reference -> EConstr.constr -> bool
+ val pf_hyps : Goal.goal Evd.sigma -> EConstr.named_context
- val eq_constr : Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool
+ val pf_ids_of_hyps : Goal.goal Evd.sigma -> Names.Id.t list
- val occur_var_in_decl :
- Environ.env -> Evd.evar_map ->
- Names.Id.t -> EConstr.named_declaration -> bool
+ val pf_reduce_to_atomic_ind : Goal.goal Evd.sigma -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * EConstr.types
- val subst_meta : Prelude.meta_value_map -> Term.constr -> Term.constr
+ val pf_reduce_to_quantified_ind : Goal.goal Evd.sigma -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * EConstr.types
- val free_rels : Evd.evar_map -> EConstr.constr -> Int.Set.t
+ val pf_eapply : (Environ.env -> Evd.evar_map -> 'a -> Evd.evar_map * 'b) ->
+ Evar.t Evd.sigma -> 'a -> Evar.t Evd.sigma * 'b
- val occur_term : Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool
- [@@ocaml.deprecated "alias of API.Termops.dependent"]
+ val pf_unfoldn : (Locus.occurrences * Names.evaluable_global_reference) list
+ -> Goal.goal Evd.sigma -> EConstr.constr -> EConstr.constr
- val replace_term : Evd.evar_map -> EConstr.constr -> EConstr.constr -> EConstr.constr -> EConstr.constr
- val map_named_decl : ('a -> 'b) -> ('a, 'a) Context.Named.Declaration.pt -> ('b, 'b) Context.Named.Declaration.pt
- val map_rel_decl : ('a -> 'b) -> ('a, 'a) Context.Rel.Declaration.pt -> ('b, 'b) Context.Rel.Declaration.pt
- val pr_metaset : Evd.Metaset.t -> Pp.std_ppcmds
- val pr_evar_map : ?with_univs:bool -> int option -> Evd.evar_map -> Pp.std_ppcmds
- val pr_evar_universe_context : UState.t -> Pp.std_ppcmds
-end
+ val pf_reduce : (Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr) -> Goal.goal Evd.sigma -> EConstr.constr -> EConstr.constr
-module Locality :
-sig
- val make_section_locality : bool option -> bool
- module LocalityFixme : sig
- val consume : unit -> bool option
+ val pf_conv_x : Goal.goal Evd.sigma -> EConstr.constr -> EConstr.constr -> bool
+
+ val pf_is_matching : Goal.goal Evd.sigma -> Pattern.constr_pattern -> EConstr.constr -> bool
+
+ val pf_hyps_types : Goal.goal Evd.sigma -> (Names.Id.t * EConstr.types) list
+
+ val pr_gls : Goal.goal Evd.sigma -> Pp.std_ppcmds
+
+ val pf_nf_betaiota : Goal.goal Evd.sigma -> EConstr.constr -> EConstr.constr
+
+ val pf_last_hyp : Goal.goal Evd.sigma -> EConstr.named_declaration
+
+ val pf_nth_hyp_id : Goal.goal Evd.sigma -> int -> Names.Id.t
+
+ val sig_it : 'a Evd.sigma -> 'a
+
+ module New :
+ sig
+ val pf_apply : (Environ.env -> Evd.evar_map -> 'a) -> 'b Proofview.Goal.t -> 'a
+ val project : 'a Proofview.Goal.t -> Evd.evar_map
+ val pf_unsafe_type_of : 'a Proofview.Goal.t -> EConstr.constr -> EConstr.types
+ val of_old : (Goal.goal Evd.sigma -> 'a) -> [ `NF ] Proofview.Goal.t -> 'a
+
+ val pf_env : 'a Proofview.Goal.t -> Environ.env
+ val pf_ids_of_hyps : 'a Proofview.Goal.t -> Names.Id.t list
+ val pf_concl : 'a Proofview.Goal.t -> EConstr.types
+ val pf_get_new_id : Names.Id.t -> 'a Proofview.Goal.t -> Names.Id.t
+ val pf_get_hyp_typ : Names.Id.t -> 'a Proofview.Goal.t -> EConstr.types
+ val pf_get_type_of : 'a Proofview.Goal.t -> EConstr.constr -> EConstr.types
+ val pf_global : Names.Id.t -> 'a Proofview.Goal.t -> Globnames.global_reference
+ val pf_hyps_types : 'a Proofview.Goal.t -> (Names.Id.t * EConstr.types) list
end
- val make_module_locality : bool option -> bool
end
-module Search :
+module Pfedit :
sig
- type glob_search_about_item = Search.glob_search_about_item =
- | GlobSearchSubPattern of Pattern.constr_pattern
- | GlobSearchString of string
- type filter_function = Globnames.global_reference -> Environ.env -> Term.constr -> bool
- type display_function = Globnames.global_reference -> Environ.env -> Term.constr -> unit
- val search_about_filter : glob_search_about_item -> filter_function
- val module_filter : Names.DirPath.t list * bool -> filter_function
- val generic_search : int option -> display_function -> unit
-end
+ val solve_by_implicit_tactic : unit -> Pretyping.inference_hook option
+ val refine_by_tactic : Environ.env -> Evd.evar_map -> EConstr.types -> unit Proofview.tactic ->
+ Constr.t * Evd.evar_map
+ val declare_implicit_tactic : unit Proofview.tactic -> unit
+ val clear_implicit_tactic : unit -> unit
+ val by : unit Proofview.tactic -> bool
+ val solve : ?with_end_tac:unit Proofview.tactic ->
+ Vernacexpr.goal_selector -> int option -> unit Proofview.tactic ->
+ Proof.proof -> Proof.proof * bool
+ val cook_proof :
+ unit -> (Names.Id.t * (Safe_typing.private_constants Entries.definition_entry * Proof_global.proof_universes * Decl_kinds.goal_kind))
-module Notation_ops :
-sig
- val glob_constr_of_notation_constr : ?loc:Loc.t -> Notation_term.notation_constr -> Glob_term.glob_constr
- val glob_constr_of_notation_constr_with_binders : ?loc:Loc.t ->
- ('a -> Names.Name.t -> 'a * Names.Name.t) ->
- ('a -> Notation_term.notation_constr -> Glob_term.glob_constr) ->
- 'a -> Notation_term.notation_constr -> Glob_term.glob_constr
-end
+ val get_current_context : unit -> Evd.evar_map * Environ.env
+
+ (* Deprecated *)
+ val delete_current_proof : unit -> unit
+ [@@ocaml.deprecated "use Proof_global.discard_current"]
+
+ val get_current_proof_name : unit -> Names.Id.t
+ [@@ocaml.deprecated "use Proof_global.get_current_proof_name"]
-module Constrextern :
-sig
- val extern_glob_constr : Names.Id.Set.t -> Glob_term.glob_constr -> Constrexpr.constr_expr
- val extern_glob_type : Names.Id.Set.t -> Glob_term.glob_constr -> Constrexpr.constr_expr
- val extern_constr : ?lax:bool -> bool -> Environ.env -> Evd.evar_map -> Term.constr -> Constrexpr.constr_expr
- val without_symbols : ('a -> 'b) -> 'a -> 'b
- val print_universes : bool ref
- val extern_type : bool -> Environ.env -> Evd.evar_map -> Term.types -> Constrexpr.constr_expr
- val with_universes : ('a -> 'b) -> 'a -> 'b
- val set_extern_reference :
- (?loc:Loc.t -> Names.Id.Set.t -> Globnames.global_reference -> Libnames.reference) -> unit
end
-module Patternops :
+module Clenv :
sig
- val pattern_of_glob_constr : Glob_term.glob_constr -> Names.Id.t list * Pattern.constr_pattern
- val subst_pattern : Mod_subst.substitution -> Pattern.constr_pattern -> Pattern.constr_pattern
- val pattern_of_constr : Environ.env -> Evd.evar_map -> Term.constr -> Pattern.constr_pattern
- val instantiate_pattern : Environ.env ->
- Evd.evar_map -> Pattern.extended_patvar_map ->
- Pattern.constr_pattern -> Pattern.constr_pattern
+
+ type hole = {
+ hole_evar : EConstr.constr;
+ hole_type : EConstr.types;
+ hole_deps : bool;
+ hole_name : Names.Name.t;
+ }
+
+ type clause = {
+ cl_holes : hole list;
+ cl_concl : EConstr.types;
+ }
+
+ val make_evar_clause : Environ.env -> Evd.evar_map -> ?len:int -> EConstr.types ->
+ (Evd.evar_map * clause)
+ val solve_evar_clause : Environ.env -> Evd.evar_map -> bool -> clause -> EConstr.constr Misctypes.bindings ->
+ Evd.evar_map
+ type clausenv
+ val pr_clenv : clausenv -> Pp.std_ppcmds
end
-module Printer :
+(************************************************************************)
+(* End of modules from proofs/ *)
+(************************************************************************)
+
+(************************************************************************)
+(* Modules from parsing/ *)
+(************************************************************************)
+
+module Pcoq :
sig
- val pr_named_context : Environ.env -> Evd.evar_map -> Context.Named.t -> Pp.std_ppcmds
- val pr_rel_context : Environ.env -> Evd.evar_map -> Context.Rel.t -> Pp.std_ppcmds
- val pr_goal : Goal.goal Evd.sigma -> Pp.std_ppcmds
- val pr_constr_env : Prelude.env -> Prelude.evar_map -> Term.constr -> Pp.std_ppcmds
- val pr_lconstr_env : Prelude.env -> Prelude.evar_map -> Term.constr -> Pp.std_ppcmds
+ open Loc
+ open Names
+ open Extend
+ open Vernacexpr
+ open Genarg
+ open Constrexpr
+ open Libnames
+ open Misctypes
+ open Genredexpr
- val pr_constr : Term.constr -> Pp.std_ppcmds
+ module Gram : sig
+ include Grammar.S with type te = Tok.t
- val pr_lconstr : Term.constr -> Pp.std_ppcmds
+ type 'a entry = 'a Entry.e
+ type internal_entry = Tok.t Gramext.g_entry
+ type symbol = Tok.t Gramext.g_symbol
+ type action = Gramext.g_action
+ type production_rule = symbol list * action
+ type single_extend_statment =
+ string option * Gramext.g_assoc option * production_rule list
+ type extend_statment =
+ Gramext.position option * single_extend_statment list
- val pr_econstr : EConstr.constr -> Pp.std_ppcmds
- val pr_glob_constr : Glob_term.glob_constr -> Pp.std_ppcmds
- val pr_constr_pattern : Pattern.constr_pattern -> Pp.std_ppcmds
- val pr_glob_constr_env : Environ.env -> Glob_term.glob_constr -> Pp.std_ppcmds
- val pr_lglob_constr_env : Environ.env -> Glob_term.glob_constr -> Pp.std_ppcmds
- val pr_econstr_env : Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.std_ppcmds
- val pr_constr_pattern_env : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> Pp.std_ppcmds
- val pr_lconstr_pattern_env : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> Pp.std_ppcmds
- val pr_closed_glob : Glob_term.closed_glob_constr -> Pp.std_ppcmds
- val pr_lglob_constr : Glob_term.glob_constr -> Pp.std_ppcmds
- val pr_leconstr_env : Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.std_ppcmds
- val pr_leconstr : EConstr.constr -> Pp.std_ppcmds
- val pr_global : Globnames.global_reference -> Pp.std_ppcmds
- val pr_lconstr_under_binders : Pattern.constr_under_binders -> Pp.std_ppcmds
- val pr_lconstr_under_binders_env : Environ.env -> Evd.evar_map -> Pattern.constr_under_binders -> Pp.std_ppcmds
+ type coq_parsable
- val pr_constr_under_binders_env : Environ.env -> Evd.evar_map -> Pattern.constr_under_binders -> Pp.std_ppcmds
- val pr_closed_glob_env : Environ.env -> Evd.evar_map -> Glob_term.closed_glob_constr -> Pp.std_ppcmds
- val pr_rel_context_of : Environ.env -> Evd.evar_map -> Pp.std_ppcmds
- val pr_named_context_of : Environ.env -> Evd.evar_map -> Pp.std_ppcmds
- val pr_ltype : Term.types -> Pp.std_ppcmds
- val pr_ljudge : EConstr.unsafe_judgment -> Pp.std_ppcmds * Pp.std_ppcmds
- val pr_idpred : Names.Id.Pred.t -> Pp.std_ppcmds
- val pr_cpred : Names.Cpred.t -> Pp.std_ppcmds
- val pr_transparent_state : Names.transparent_state -> Pp.std_ppcmds
-end
+ val parsable : ?file:string -> char Stream.t -> coq_parsable
+ val action : 'a -> action
+ val entry_create : string -> 'a entry
+ val entry_parse : 'a entry -> coq_parsable -> 'a
+ val entry_print : Format.formatter -> 'a entry -> unit
+ val with_parsable : coq_parsable -> ('a -> 'b) -> 'a -> 'b
-module Classes :
-sig
- val set_typeclass_transparency : Names.evaluable_global_reference -> bool -> bool -> unit
- val new_instance :
- ?abstract:bool ->
- ?global:bool ->
- ?refine:bool ->
- Decl_kinds.polymorphic ->
- Constrexpr.local_binder_expr list ->
- Constrexpr.typeclass_constraint ->
- (bool * Constrexpr.constr_expr) option ->
- ?generalize:bool ->
- ?tac:unit Proofview.tactic ->
- ?hook:(Globnames.global_reference -> unit) ->
- Vernacexpr.hint_info_expr ->
- Names.Id.t
-end
+ (* Apparently not used *)
+ val srules' : production_rule list -> symbol
+ val parse_tokens_after_filter : 'a entry -> Tok.t Stream.t -> 'a
-module Classops :
-sig
- type coe_index = Classops.coe_index
- type inheritance_path = coe_index list
- type cl_index = Classops.cl_index
+ end with type 'a Entry.e = 'a Grammar.GMake(CLexer).Entry.e
- val hide_coercion : Globnames.global_reference -> int option
- val lookup_path_to_sort_from : Environ.env -> Evd.evar_map -> EConstr.types ->
- EConstr.types * inheritance_path
- val get_coercion_value : coe_index -> Constr.t
- val coercions : unit -> coe_index list
- val pr_cl_index : cl_index -> Pp.std_ppcmds
-end
+ val parse_string : 'a Gram.entry -> string -> 'a
+ val eoi_entry : 'a Gram.entry -> 'a Gram.entry
+ val map_entry : ('a -> 'b) -> 'a Gram.entry -> 'b Gram.entry
-module ExplainErr :
-sig
- val process_vernac_interp_error : ?allow_uncaught:bool -> Util.iexn -> Util.iexn
- val register_additional_error_info : (Util.iexn -> Pp.std_ppcmds option Loc.located option) -> unit
-end
+ type gram_universe
-module Tacred :
-sig
- val try_red_product : Reductionops.reduction_function
- val simpl : Reductionops.reduction_function
- val unfoldn :
- (Locus.occurrences * Names.evaluable_global_reference) list -> Reductionops.reduction_function
- val hnf_constr : Reductionops.reduction_function
- val red_product : Reductionops.reduction_function
- val is_evaluable : Environ.env -> Names.evaluable_global_reference -> bool
- val evaluable_of_global_reference :
- Environ.env -> Globnames.global_reference -> Names.evaluable_global_reference
- val error_not_evaluable : Globnames.global_reference -> 'a
- val reduce_to_quantified_ref :
- Environ.env -> Evd.evar_map -> Globnames.global_reference -> EConstr.types -> EConstr.types
- val pattern_occs : (Locus.occurrences * EConstr.constr) list -> Reductionops.e_reduction_function
- val cbv_norm_flags : CClosure.RedFlags.reds -> Reductionops.reduction_function
-end
+ val uprim : gram_universe
+ val uconstr : gram_universe
+ val utactic : gram_universe
+ val uvernac : gram_universe
-module Detyping :
-sig
- val print_universes : bool ref
- val print_evar_arguments : bool ref
- val detype : ?lax:bool -> bool -> Names.Id.t list -> Environ.env -> Evd.evar_map -> EConstr.constr -> Glob_term.glob_constr
- val subst_glob_constr : Mod_subst.substitution -> Glob_term.glob_constr -> Glob_term.glob_constr
- val set_detype_anonymous : (?loc:Loc.t -> int -> Glob_term.glob_constr) -> unit
-end
+ val register_grammar : ('raw, 'glb, 'top) genarg_type -> 'raw Gram.entry -> unit
-module Constrexpr_ops :
-sig
- val mkIdentC : Names.Id.t -> Constrexpr.constr_expr
- val mkAppC : Constrexpr.constr_expr * Constrexpr.constr_expr list -> Constrexpr.constr_expr
- val names_of_local_assums : Constrexpr.local_binder_expr list -> Names.Name.t Loc.located list
- val coerce_reference_to_id : Prelude.reference -> Names.Id.t
- val coerce_to_id : Constrexpr.constr_expr -> Names.Id.t Loc.located
- val constr_loc : Constrexpr.constr_expr -> Loc.t option
- val mkRefC : Prelude.reference -> Constrexpr.constr_expr
- val mkLambdaC : Names.Name.t Loc.located list * Constrexpr.binder_kind * Constrexpr.constr_expr * Constrexpr.constr_expr -> Constrexpr.constr_expr
- val default_binder_kind : Constrexpr.binder_kind
- val mkLetInC : Names.Name.t Loc.located * Constrexpr.constr_expr * Constrexpr.constr_expr option * Constrexpr.constr_expr -> Constrexpr.constr_expr
- val mkCProdN : ?loc:Loc.t -> Constrexpr.local_binder_expr list -> Constrexpr.constr_expr -> Constrexpr.constr_expr
-end
+ val genarg_grammar : ('raw, 'glb, 'top) genarg_type -> 'raw Gram.entry
-module Glob_ops :
-sig
- val map_glob_constr_left_to_right : (Glob_term.glob_constr -> Glob_term.glob_constr) -> Glob_term.glob_constr -> Glob_term.glob_constr
- val loc_of_glob_constr : Glob_term.glob_constr -> Loc.t option
- val glob_constr_eq : Glob_term.glob_constr -> Glob_term.glob_constr -> bool
- val bound_glob_vars : Glob_term.glob_constr -> Names.Id.Set.t
+ val create_generic_entry : gram_universe -> string ->
+ ('a, rlevel) abstract_argument_type -> 'a Gram.entry
- (** Conversion from glob_constr to cases pattern, if possible
+ module Prim :
+ sig
+ open Names
+ open Libnames
+ val preident : string Gram.entry
+ val ident : Id.t Gram.entry
+ val name : Name.t located Gram.entry
+ val identref : Id.t located Gram.entry
+ val pidentref : (Id.t located * (Id.t located list) option) Gram.entry
+ val pattern_ident : Id.t Gram.entry
+ val pattern_identref : Id.t located Gram.entry
+ val base_ident : Id.t Gram.entry
+ val natural : int Gram.entry
+ val bigint : Constrexpr.raw_natural_number Gram.entry
+ val integer : int Gram.entry
+ val string : string Gram.entry
+ val lstring : string located Gram.entry
+ val qualid : qualid located Gram.entry
+ val fullyqualid : Id.t list located Gram.entry
+ val reference : reference Gram.entry
+ val by_notation : (string * string option) Loc.located Gram.entry
+ val smart_global : reference or_by_notation Gram.entry
+ val dirpath : DirPath.t Gram.entry
+ val ne_string : string Gram.entry
+ val ne_lstring : string located Gram.entry
+ val var : Id.t located Gram.entry
+ end
- Take the current alias as parameter,
- @raise Not_found if translation is impossible *)
- val cases_pattern_of_glob_constr : Names.Name.t -> Glob_term.glob_constr -> Glob_term.cases_pattern
- val map_glob_constr :
- (Glob_term.glob_constr -> Glob_term.glob_constr) -> Glob_term.glob_constr -> Glob_term.glob_constr
- val empty_lvar : Glob_term.ltac_var_map
-end
+ module Constr :
+ sig
+ val constr : constr_expr Gram.entry
+ val constr_eoi : constr_expr Gram.entry
+ val lconstr : constr_expr Gram.entry
+ val binder_constr : constr_expr Gram.entry
+ val operconstr : constr_expr Gram.entry
+ val ident : Id.t Gram.entry
+ val global : reference Gram.entry
+ val universe_level : glob_level Gram.entry
+ val sort : glob_sort Gram.entry
+ val pattern : cases_pattern_expr Gram.entry
+ val constr_pattern : constr_expr Gram.entry
+ val lconstr_pattern : constr_expr Gram.entry
+ val closed_binder : local_binder_expr list Gram.entry
+ val binder : local_binder_expr list Gram.entry (* closed_binder or variable *)
+ val binders : local_binder_expr list Gram.entry (* list of binder *)
+ val open_binders : local_binder_expr list Gram.entry
+ val binders_fixannot : (local_binder_expr list * (Id.t located option * recursion_order_expr)) Gram.entry
+ val typeclass_constraint : (Name.t located * bool * constr_expr) Gram.entry
+ val record_declaration : constr_expr Gram.entry
+ val appl_arg : (constr_expr * explicitation located option) Gram.entry
+ end
-module Indrec :
-sig
- type dep_flag = bool
- val lookup_eliminator : Names.inductive -> Sorts.family -> Globnames.global_reference
- val build_case_analysis_scheme : Environ.env -> Evd.evar_map -> Term.pinductive ->
- dep_flag -> Sorts.family -> Evd.evar_map * Term.constr
- val make_elimination_ident : Names.Id.t -> Sorts.family -> Names.Id.t
- val build_mutual_induction_scheme :
- Environ.env -> Evd.evar_map -> (Term.pinductive * dep_flag * Sorts.family) list -> Evd.evar_map * Term.constr list
- val build_case_analysis_scheme_default : Environ.env -> Evd.evar_map -> Term.pinductive ->
- Sorts.family -> Evd.evar_map * Term.constr
-end
+ module Vernac_ :
+ sig
+ val gallina : vernac_expr Gram.entry
+ val gallina_ext : vernac_expr Gram.entry
+ val command : vernac_expr Gram.entry
+ val syntax : vernac_expr Gram.entry
+ val vernac : vernac_expr Gram.entry
+ val rec_definition : (fixpoint_expr * decl_notation list) Gram.entry
+ val vernac_eoi : vernac_expr Gram.entry
+ val noedit_mode : vernac_expr Gram.entry
+ val command_entry : vernac_expr Gram.entry
+ val red_expr : raw_red_expr Gram.entry
+ val hint_info : Vernacexpr.hint_info_expr Gram.entry
+ end
+
+ val epsilon_value : ('a -> 'self) -> ('self, 'a) Extend.symbol -> 'self option
+
+ val get_command_entry : unit -> vernac_expr Gram.entry
+ val set_command_entry : vernac_expr Gram.entry -> unit
+
+ type gram_reinit = gram_assoc * gram_position
+ val grammar_extend : 'a Gram.entry -> gram_reinit option ->
+ 'a Extend.extend_statment -> unit
+
+ module GramState : Store.S
+
+ type 'a grammar_command
+
+ type extend_rule =
+ | ExtendRule : 'a Gram.entry * gram_reinit option * 'a extend_statment -> extend_rule
+
+ type 'a grammar_extension = 'a -> GramState.t -> extend_rule list * GramState.t
+
+ val create_grammar_command : string -> 'a grammar_extension -> 'a grammar_command
+
+ val extend_grammar_command : 'a grammar_command -> 'a -> unit
+ val recover_grammar_command : 'a grammar_command -> 'a list
+ val with_grammar_rule_protection : ('a -> 'b) -> 'a -> 'b
+
+ val to_coqloc : Ploc.t -> Loc.t
+ val (!@) : Ploc.t -> Loc.t
-module Logic :
-sig
- type refiner_error = Logic.refiner_error =
- | BadType of Term.constr * Term.constr * Term.constr
- | UnresolvedBindings of Names.Name.t list
- | CannotApply of Term.constr * Term.constr
- | NotWellTyped of Term.constr
- | NonLinearProof of Term.constr
- | MetaInType of EConstr.constr
- | IntroNeedsProduct
- | DoesNotOccurIn of Term.constr * Names.Id.t
- | NoSuchHyp of Names.Id.t
- exception RefinerError of refiner_error
- val catchable_exception : exn -> bool
end
-module Himsg :
+module Egramml :
sig
- val explain_refiner_error : Logic.refiner_error -> Pp.std_ppcmds
- val explain_pretype_error : Environ.env -> Evd.evar_map -> Pretype_errors.pretype_error -> Pp.std_ppcmds
+ open Vernacexpr
+
+ type 's grammar_prod_item =
+ | GramTerminal of string
+ | GramNonTerminal : ('a Genarg.raw_abstract_argument_type option *
+ ('s, 'a) Extend.symbol) Loc.located -> 's grammar_prod_item
+
+ val extend_vernac_command_grammar :
+ extend_name -> vernac_expr Pcoq.Gram.entry option ->
+ vernac_expr grammar_prod_item list -> unit
+
+ val make_rule :
+ (Loc.t -> Genarg.raw_generic_argument list -> 'a) ->
+ 'a grammar_prod_item list -> 'a Extend.production_rule
+
end
-module Extend :
+(************************************************************************)
+(* End of modules from parsing/ *)
+(************************************************************************)
+
+(************************************************************************)
+(* Modules from printing/ *)
+(************************************************************************)
+
+module Genprint :
sig
- type ('self, 'a) symbol = ('self, 'a) Extend.symbol
- type 'a user_symbol = 'a Extend.user_symbol =
- | Ulist1 of 'a user_symbol
- | Ulist1sep of 'a user_symbol * string
- | Ulist0 of 'a user_symbol
- | Ulist0sep of 'a user_symbol * string
- | Uopt of 'a user_symbol
- | Uentry of 'a
- | Uentryl of 'a * int
+ type 'a printer = 'a -> Pp.std_ppcmds
+ val generic_top_print : Genarg.tlevel Genarg.generic_argument printer
+ val register_print0 : ('raw, 'glb, 'top) Genarg.genarg_type ->
+ 'raw printer -> 'glb printer -> 'top printer -> unit
end
module Pputils :
@@ -3908,80 +4881,170 @@ sig
val pr_glob_sort : Misctypes.glob_sort -> Pp.std_ppcmds
end
-module Genprint :
+module Printer :
sig
- type 'a printer = 'a -> Pp.std_ppcmds
- val generic_top_print : Genarg.tlevel Genarg.generic_argument printer
- val register_print0 : ('raw, 'glb, 'top) Genarg.genarg_type ->
- 'raw printer -> 'glb printer -> 'top printer -> unit
-end
+ val pr_named_context : Environ.env -> Evd.evar_map -> Context.Named.t -> Pp.std_ppcmds
+ val pr_rel_context : Environ.env -> Evd.evar_map -> Context.Rel.t -> Pp.std_ppcmds
+ val pr_goal : Goal.goal Evd.sigma -> Pp.std_ppcmds
-module Miscprint :
-sig
- val pr_or_and_intro_pattern :
- ('a -> Pp.std_ppcmds) -> 'a Misctypes.or_and_intro_pattern_expr -> Pp.std_ppcmds
- val pr_intro_pattern_naming : Misctypes.intro_pattern_naming_expr -> Pp.std_ppcmds
- val pr_intro_pattern :
- ('a -> Pp.std_ppcmds) -> 'a Misctypes.intro_pattern_expr Loc.located -> Pp.std_ppcmds
- val pr_bindings :
- ('a -> Pp.std_ppcmds) ->
- ('a -> Pp.std_ppcmds) -> 'a Misctypes.bindings -> Pp.std_ppcmds
- val pr_bindings_no_with :
- ('a -> Pp.std_ppcmds) ->
- ('a -> Pp.std_ppcmds) -> 'a Misctypes.bindings -> Pp.std_ppcmds
- val pr_with_bindings :
- ('a -> Pp.std_ppcmds) ->
- ('a -> Pp.std_ppcmds) -> 'a * 'a Misctypes.bindings -> Pp.std_ppcmds
-end
+ val pr_constr_env : Environ.env -> Evd.evar_map -> Constr.t -> Pp.std_ppcmds
+ val pr_lconstr_env : Environ.env -> Evd.evar_map -> Constr.t -> Pp.std_ppcmds
-module Miscops :
-sig
- val map_red_expr_gen : ('a -> 'd) -> ('b -> 'e) -> ('c -> 'f) ->
- ('a,'b,'c) Genredexpr.red_expr_gen -> ('d,'e,'f) Genredexpr.red_expr_gen
- val map_cast_type : ('a -> 'b) -> 'a Misctypes.cast_type -> 'b Misctypes.cast_type
-end
+ val pr_constr : Constr.t -> Pp.std_ppcmds
-module Stateid :
-sig
- type t = Stateid.t
- module Self : module type of struct include Stateid.Self end
+ val pr_lconstr : Constr.t -> Pp.std_ppcmds
+
+ val pr_econstr : EConstr.constr -> Pp.std_ppcmds
+ val pr_glob_constr : Glob_term.glob_constr -> Pp.std_ppcmds
+ val pr_constr_pattern : Pattern.constr_pattern -> Pp.std_ppcmds
+ val pr_glob_constr_env : Environ.env -> Glob_term.glob_constr -> Pp.std_ppcmds
+ val pr_lglob_constr_env : Environ.env -> Glob_term.glob_constr -> Pp.std_ppcmds
+ val pr_econstr_env : Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.std_ppcmds
+ val pr_constr_pattern_env : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> Pp.std_ppcmds
+ val pr_lconstr_pattern_env : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> Pp.std_ppcmds
+ val pr_closed_glob : Glob_term.closed_glob_constr -> Pp.std_ppcmds
+ val pr_lglob_constr : Glob_term.glob_constr -> Pp.std_ppcmds
+ val pr_leconstr_env : Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.std_ppcmds
+ val pr_leconstr : EConstr.constr -> Pp.std_ppcmds
+ val pr_global : Globnames.global_reference -> Pp.std_ppcmds
+ val pr_lconstr_under_binders : Pattern.constr_under_binders -> Pp.std_ppcmds
+ val pr_lconstr_under_binders_env : Environ.env -> Evd.evar_map -> Pattern.constr_under_binders -> Pp.std_ppcmds
+
+ val pr_constr_under_binders_env : Environ.env -> Evd.evar_map -> Pattern.constr_under_binders -> Pp.std_ppcmds
+ val pr_closed_glob_env : Environ.env -> Evd.evar_map -> Glob_term.closed_glob_constr -> Pp.std_ppcmds
+ val pr_rel_context_of : Environ.env -> Evd.evar_map -> Pp.std_ppcmds
+ val pr_named_context_of : Environ.env -> Evd.evar_map -> Pp.std_ppcmds
+ val pr_ltype : Term.types -> Pp.std_ppcmds
+ val pr_ljudge : EConstr.unsafe_judgment -> Pp.std_ppcmds * Pp.std_ppcmds
+ val pr_idpred : Names.Id.Pred.t -> Pp.std_ppcmds
+ val pr_cpred : Names.Cpred.t -> Pp.std_ppcmds
+ val pr_transparent_state : Names.transparent_state -> Pp.std_ppcmds
end
-module Stm :
+(************************************************************************)
+(* End of modules from printing/ *)
+(************************************************************************)
+
+(************************************************************************)
+(* Modules from tactics/ *)
+(************************************************************************)
+
+module Tacticals :
sig
- type state = Stm.state
- val state_of_id :
- Stateid.t -> [ `Valid of state option | `Expired | `Error of exn ]
+ open Proof_type
+
+ val tclORELSE : tactic -> tactic -> tactic
+ val tclDO : int -> tactic -> tactic
+ val tclIDTAC : tactic
+ val tclFAIL : int -> Pp.std_ppcmds -> tactic
+ val tclTHEN : tactic -> tactic -> tactic
+ val tclTHENLIST : tactic list -> tactic
+ val pf_constr_of_global :
+ Globnames.global_reference -> (EConstr.constr -> Proof_type.tactic) -> Proof_type.tactic
+ val tclMAP : ('a -> tactic) -> 'a list -> tactic
+ val tclTRY : tactic -> tactic
+ val tclCOMPLETE : tactic -> tactic
+ val tclTHENS : tactic -> tactic list -> tactic
+ val tclFIRST : tactic list -> tactic
+ val tclTHENFIRST : tactic -> tactic -> tactic
+ val tclTHENLAST : tactic -> tactic -> tactic
+ val tclTHENSFIRSTn : tactic -> tactic array -> tactic -> tactic
+ val tclTHENSLASTn : tactic -> tactic -> tactic array -> tactic
+ val tclSOLVE : tactic list -> tactic
+
+ val onClause : (Names.Id.t option -> tactic) -> Locus.clause -> tactic
+ val onAllHypsAndConcl : (Names.Id.t option -> tactic) -> tactic
+ val onLastHypId : (Names.Id.t -> tactic) -> tactic
+ val onNthHypId : int -> (Names.Id.t -> tactic) -> tactic
+ val onNLastHypsId : int -> (Names.Id.t list -> tactic) -> tactic
+
+ val tclTHENSEQ : tactic list -> tactic
+ [@@ocaml.deprecated "alias of API.Tacticals.tclTHENLIST"]
+
+ val nLastDecls : int -> Goal.goal Evd.sigma -> EConstr.named_context
+
+ val tclTHEN_i : tactic -> (int -> tactic) -> tactic
+
+ val tclPROGRESS : tactic -> tactic
+
+ val elimination_sort_of_goal : Goal.goal Evd.sigma -> Sorts.family
+
+ module New :
+ sig
+ open Proofview
+ val tclORELSE0 : unit tactic -> unit tactic -> unit tactic
+ val tclFAIL : int -> Pp.std_ppcmds -> 'a tactic
+ val pf_constr_of_global : Globnames.global_reference -> EConstr.constr tactic
+ val tclTHEN : unit tactic -> unit tactic -> unit tactic
+ val tclTHENS : unit tactic -> unit tactic list -> unit tactic
+ val tclFIRST : unit tactic list -> unit tactic
+ val tclZEROMSG : ?loc:Loc.t -> Pp.std_ppcmds -> 'a tactic
+ val tclORELSE : unit tactic -> unit tactic -> unit tactic
+ val tclREPEAT : unit tactic -> unit tactic
+ val tclTRY : unit tactic -> unit tactic
+ val tclTHENFIRST : unit tactic -> unit tactic -> unit tactic
+ val tclPROGRESS : unit Proofview.tactic -> unit Proofview.tactic
+ val tclTHENS3PARTS : unit tactic -> unit tactic array -> unit tactic -> unit tactic array -> unit tactic
+ val tclDO : int -> unit tactic -> unit tactic
+ val tclTIMEOUT : int -> unit tactic -> unit tactic
+ val tclTIME : string option -> 'a tactic -> 'a tactic
+ val tclOR : unit tactic -> unit tactic -> unit tactic
+ val tclONCE : unit tactic -> unit tactic
+ val tclEXACTLY_ONCE : unit tactic -> unit tactic
+ val tclIFCATCH :
+ unit tactic ->
+ (unit -> unit tactic) ->
+ (unit -> unit tactic) -> unit tactic
+ val tclSOLVE : unit tactic list -> unit tactic
+ val tclCOMPLETE : 'a tactic -> 'a tactic
+ val tclSELECT : Vernacexpr.goal_selector -> 'a tactic -> 'a tactic
+ val tclWITHHOLES : bool -> 'a tactic -> Evd.evar_map -> 'a tactic
+ val tclDELAYEDWITHHOLES : bool -> 'a Tactypes.delayed_open -> ('a -> unit tactic) -> unit tactic
+ val tclTHENLIST : unit tactic list -> unit tactic
+ val tclTHENLAST : unit tactic -> unit tactic -> unit tactic
+ val tclMAP : ('a -> unit tactic) -> 'a list -> unit tactic
+ val tclIDTAC : unit tactic
+ val tclIFTHENELSE : unit tactic -> unit tactic -> unit tactic -> unit tactic
+ val tclIFTHENSVELSE : unit tactic -> unit tactic array -> unit tactic -> unit tactic
+ end
end
-module Declaremods :
+module Hipattern :
sig
- val append_end_library_hook : (unit -> unit) -> unit
+ exception NoEquationFound
+ type 'a matching_function = Evd.evar_map -> EConstr.constr -> 'a option
+ type testing_function = Evd.evar_map -> EConstr.constr -> bool
+ val is_disjunction : ?strict:bool -> ?onlybinary:bool -> testing_function
+ val match_with_disjunction : ?strict:bool -> ?onlybinary:bool -> (EConstr.constr * EConstr.constr list) matching_function
+ val match_with_equality_type : (EConstr.constr * EConstr.constr list) matching_function
+ val is_empty_type : testing_function
+ val is_unit_type : testing_function
+ val is_unit_or_eq_type : testing_function
+ val is_conjunction : ?strict:bool -> ?onlybinary:bool -> testing_function
+ val match_with_conjunction : ?strict:bool -> ?onlybinary:bool -> (EConstr.constr * EConstr.constr list) matching_function
+ val match_with_imp_term : (EConstr.constr * EConstr.constr) matching_function
+ val match_with_forall_term : (Names.Name.t * EConstr.constr * EConstr.constr) matching_function
+ val match_with_nodep_ind : (EConstr.constr * EConstr.constr list * int) matching_function
+ val match_with_sigma_type : (EConstr.constr * EConstr.constr list) matching_function
end
-module Pfedit :
+module Ind_tables :
sig
- val solve_by_implicit_tactic : unit -> Pretyping.inference_hook option
- val refine_by_tactic : Environ.env -> Evd.evar_map -> EConstr.types -> unit Proofview.tactic ->
- Term.constr * Evd.evar_map
- val declare_implicit_tactic : unit Proofview.tactic -> unit
- val clear_implicit_tactic : unit -> unit
- val by : unit Proofview.tactic -> bool
- val solve : ?with_end_tac:unit Proofview.tactic ->
- Vernacexpr.goal_selector -> int option -> unit Proofview.tactic ->
- Proof.proof -> Proof.proof * bool
- val cook_proof :
- unit -> (Names.Id.t * (Safe_typing.private_constants Entries.definition_entry * Proof_global.proof_universes * Decl_kinds.goal_kind))
-
- val get_current_context : unit -> Evd.evar_map * Environ.env
+ type individual
+ type 'a scheme_kind
- (* Deprecated *)
- val delete_current_proof : unit -> unit
- [@@ocaml.deprecated "use Proof_global.discard_current"]
-
- val get_current_proof_name : unit -> Names.Id.t
- [@@ocaml.deprecated "use Proof_global.get_current_proof_name"]
+ val check_scheme : 'a scheme_kind -> Names.inductive -> bool
+ val find_scheme : ?mode:Declare.internal_flag -> 'a scheme_kind -> Names.inductive -> Names.Constant.t * Safe_typing.private_constants
+ val pr_scheme_kind : 'a scheme_kind -> Pp.std_ppcmds
+end
+module Elimschemes :
+sig
+ val case_scheme_kind_from_prop : Ind_tables.individual Ind_tables.scheme_kind
+ val case_dep_scheme_kind_from_type_in_prop : Ind_tables.individual Ind_tables.scheme_kind
+ val case_scheme_kind_from_type : Ind_tables.individual Ind_tables.scheme_kind
+ val case_dep_scheme_kind_from_type : Ind_tables.individual Ind_tables.scheme_kind
+ val case_dep_scheme_kind_from_prop : Ind_tables.individual Ind_tables.scheme_kind
end
module Tactics :
@@ -3990,7 +5053,8 @@ sig
type change_arg = Pattern.patvar_map -> Evd.evar_map -> Evd.evar_map * EConstr.constr
type tactic_reduction = Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr
- type elim_scheme = Tactics.elim_scheme =
+
+ type elim_scheme =
{
elimc: EConstr.constr Misctypes.with_bindings option;
elimt: EConstr.types;
@@ -4020,7 +5084,7 @@ sig
val simplest_elim : EConstr.constr -> unit tactic
val introf : unit tactic
val cut : EConstr.types -> unit tactic
- val convert_concl : ?check:bool -> EConstr.types -> Term.cast_kind -> unit tactic
+ val convert_concl : ?check:bool -> EConstr.types -> Constr.cast_kind -> unit tactic
val intro_using : Names.Id.t -> unit tactic
val intro : unit tactic
val fresh_id_in_env : Names.Id.t list -> Names.Id.t -> Environ.env -> Names.Id.t
@@ -4030,7 +5094,7 @@ sig
val apply_with_delayed_bindings_gen :
Misctypes.advanced_flag -> Misctypes.evars_flag -> (Misctypes.clear_flag * Tactypes.delayed_open_constr_with_bindings Loc.located) list -> unit Proofview.tactic
val apply_delayed_in :
- Misctypes.advanced_flag -> Misctypes.evars_flag -> Names.Id.t ->
+ Misctypes.advanced_flag -> Misctypes.evars_flag -> Names.Id.t ->
(Misctypes.clear_flag * Tactypes.delayed_open_constr_with_bindings Loc.located) list ->
Tactypes.intro_pattern option -> unit Proofview.tactic
val elim :
@@ -4095,8 +5159,8 @@ sig
val generalize : EConstr.constr list -> unit Proofview.tactic
val simplest_case : EConstr.constr -> unit Proofview.tactic
val introduction : ?check:bool -> Names.Id.t -> unit Proofview.tactic
- val convert_concl_no_check : EConstr.types -> Term.cast_kind -> unit Proofview.tactic
- val reduct_in_concl : tactic_reduction * Term.cast_kind -> unit Proofview.tactic
+ val convert_concl_no_check : EConstr.types -> Constr.cast_kind -> unit Proofview.tactic
+ val reduct_in_concl : tactic_reduction * Constr.cast_kind -> unit Proofview.tactic
val reduct_in_hyp : ?check:bool -> tactic_reduction -> Locus.hyp_location -> unit Proofview.tactic
val convert_hyp_no_check : EConstr.named_declaration -> unit Proofview.tactic
val reflexivity_red : bool -> unit Proofview.tactic
@@ -4113,7 +5177,7 @@ sig
(Locus.occurrences * Names.evaluable_global_reference) list -> unit Proofview.tactic
val intros_using : Names.Id.t list -> unit Proofview.tactic
val simpl_in_concl : unit Proofview.tactic
- val reduct_option : ?check:bool -> tactic_reduction * Term.cast_kind -> Locus.goal_location -> unit Proofview.tactic
+ val reduct_option : ?check:bool -> tactic_reduction * Constr.cast_kind -> Locus.goal_location -> unit Proofview.tactic
val simplest_split : unit Proofview.tactic
val unfold_in_hyp :
(Locus.occurrences * Names.evaluable_global_reference) list -> Locus.hyp_location -> unit Proofview.tactic
@@ -4149,83 +5213,12 @@ sig
end
end
-module Tacticals :
+module Elim :
sig
- open Tacmach
- val tclORELSE : tactic -> tactic -> tactic
- val tclDO : int -> tactic -> tactic
- val tclIDTAC : tactic
- val tclFAIL : int -> Pp.std_ppcmds -> tactic
- val tclTHEN : tactic -> tactic -> tactic
- val tclTHENLIST : tactic list -> tactic
- val pf_constr_of_global :
- Globnames.global_reference -> (EConstr.constr -> Tacmach.tactic) -> Tacmach.tactic
- val tclMAP : ('a -> tactic) -> 'a list -> tactic
- val tclTRY : tactic -> tactic
- val tclCOMPLETE : tactic -> tactic
- val tclTHENS : tactic -> tactic list -> tactic
- val tclFIRST : tactic list -> tactic
- val tclTHENFIRST : tactic -> tactic -> tactic
- val tclTHENLAST : tactic -> tactic -> tactic
- val tclTHENSFIRSTn : tactic -> tactic array -> tactic -> tactic
- val tclTHENSLASTn : tactic -> tactic -> tactic array -> tactic
- val tclSOLVE : tactic list -> tactic
-
- val onClause : (Names.Id.t option -> tactic) -> Locus.clause -> tactic
- val onAllHypsAndConcl : (Names.Id.t option -> tactic) -> tactic
- val onLastHypId : (Names.Id.t -> tactic) -> tactic
- val onNthHypId : int -> (Names.Id.t -> tactic) -> tactic
- val onNLastHypsId : int -> (Names.Id.t list -> tactic) -> tactic
-
- val tclTHENSEQ : tactic list -> tactic
- [@@ocaml.deprecated "alias of API.Tacticals.tclTHENLIST"]
-
- val nLastDecls : int -> Goal.goal Evd.sigma -> EConstr.named_context
-
- val tclTHEN_i : tactic -> (int -> tactic) -> tactic
-
- val tclPROGRESS : tactic -> tactic
-
- val elimination_sort_of_goal : Goal.goal Evd.sigma -> Sorts.family
-
- module New :
- sig
- open Proofview
- val tclORELSE0 : unit tactic -> unit tactic -> unit tactic
- val tclFAIL : int -> Pp.std_ppcmds -> 'a tactic
- val pf_constr_of_global : Globnames.global_reference -> EConstr.constr tactic
- val tclTHEN : unit tactic -> unit tactic -> unit tactic
- val tclTHENS : unit tactic -> unit tactic list -> unit tactic
- val tclFIRST : unit tactic list -> unit tactic
- val tclZEROMSG : ?loc:Loc.t -> Pp.std_ppcmds -> 'a tactic
- val tclORELSE : unit tactic -> unit tactic -> unit tactic
- val tclREPEAT : unit tactic -> unit tactic
- val tclTRY : unit tactic -> unit tactic
- val tclTHENFIRST : unit tactic -> unit tactic -> unit tactic
- val tclPROGRESS : unit Proofview.tactic -> unit Proofview.tactic
- val tclTHENS3PARTS : unit tactic -> unit tactic array -> unit tactic -> unit tactic array -> unit tactic
- val tclDO : int -> unit tactic -> unit tactic
- val tclTIMEOUT : int -> unit tactic -> unit tactic
- val tclTIME : string option -> 'a tactic -> 'a tactic
- val tclOR : unit tactic -> unit tactic -> unit tactic
- val tclONCE : unit tactic -> unit tactic
- val tclEXACTLY_ONCE : unit tactic -> unit tactic
- val tclIFCATCH :
- unit tactic ->
- (unit -> unit tactic) ->
- (unit -> unit tactic) -> unit tactic
- val tclSOLVE : unit tactic list -> unit tactic
- val tclCOMPLETE : 'a tactic -> 'a tactic
- val tclSELECT : Vernacexpr.goal_selector -> 'a tactic -> 'a tactic
- val tclWITHHOLES : bool -> 'a tactic -> Evd.evar_map -> 'a tactic
- val tclDELAYEDWITHHOLES : bool -> 'a Tactypes.delayed_open -> ('a -> unit tactic) -> unit tactic
- val tclTHENLIST : unit tactic list -> unit tactic
- val tclTHENLAST : unit tactic -> unit tactic -> unit tactic
- val tclMAP : ('a -> unit tactic) -> 'a list -> unit tactic
- val tclIDTAC : unit tactic
- val tclIFTHENELSE : unit tactic -> unit tactic -> unit tactic -> unit tactic
- val tclIFTHENSVELSE : unit tactic -> unit tactic array -> unit tactic -> unit tactic
- end
+ val h_decompose : Names.inductive list -> EConstr.constr -> unit Proofview.tactic
+ val h_double_induction : Misctypes.quantified_hypothesis -> Misctypes.quantified_hypothesis-> unit Proofview.tactic
+ val h_decompose_or : EConstr.constr -> unit Proofview.tactic
+ val h_decompose_and : EConstr.constr -> unit Proofview.tactic
end
module Equality :
@@ -4269,14 +5262,15 @@ sig
orientation -> Locus.occurrences -> freeze_evars_flag -> dep_proof_flag ->
?tac:(unit Proofview.tactic * conditions) -> EConstr.constr Misctypes.with_bindings -> Misctypes.evars_flag -> unit Proofview.tactic
val subst : Names.Id.t list -> unit Proofview.tactic
- type subst_tactic_flags = Equality.subst_tactic_flags = {
+
+ type subst_tactic_flags = {
only_leibniz : bool;
rewrite_dependent_proof : bool
}
val subst_all : ?flags:subst_tactic_flags -> unit -> unit Proofview.tactic
val general_rewrite_in :
- orientation -> Locus.occurrences -> freeze_evars_flag -> dep_proof_flag ->
+ orientation -> Locus.occurrences -> freeze_evars_flag -> dep_proof_flag ->
?tac:(unit Proofview.tactic * conditions) -> Names.Id.t -> EConstr.constr -> Misctypes.evars_flag -> unit Proofview.tactic
val general_setoid_rewrite_clause :
@@ -4302,43 +5296,63 @@ sig
val absurd : EConstr.constr -> unit Proofview.tactic
end
-module Clenv :
+module Inv :
sig
- type hole = Clenv.hole = {
- hole_evar : EConstr.constr;
- hole_type : EConstr.types;
- hole_deps : bool;
- hole_name : Names.Name.t;
- }
- type clause = Clenv.clause = {
- cl_holes : hole list;
- cl_concl : EConstr.types;
- }
- val make_evar_clause : Environ.env -> Evd.evar_map -> ?len:int -> EConstr.types ->
- (Evd.evar_map * clause)
- val solve_evar_clause : Environ.env -> Evd.evar_map -> bool -> clause -> EConstr.constr Misctypes.bindings ->
- Evd.evar_map
- type clausenv = Clenv.clausenv
- val pr_clenv : Clenv.clausenv -> Pp.std_ppcmds
+ val dinv :
+ Misctypes.inversion_kind -> EConstr.constr option ->
+ Tactypes.or_and_intro_pattern option -> Misctypes.quantified_hypothesis -> unit Proofview.tactic
+ val inv_clause :
+ Misctypes.inversion_kind -> Tactypes.or_and_intro_pattern option -> Names.Id.t list ->
+ Misctypes.quantified_hypothesis -> unit Proofview.tactic
+ val inv_clear_tac : Names.Id.t -> unit Proofview.tactic
+ val inv_tac : Names.Id.t -> unit Proofview.tactic
+ val dinv_tac : Names.Id.t -> unit Proofview.tactic
+ val dinv_clear_tac : Names.Id.t -> unit Proofview.tactic
+ val inv : Misctypes.inversion_kind -> Tactypes.or_and_intro_pattern option ->
+ Misctypes.quantified_hypothesis -> unit Proofview.tactic
+end
+
+module Leminv :
+sig
+ val lemInv_clause :
+ Misctypes.quantified_hypothesis -> EConstr.constr -> Names.Id.t list -> unit Proofview.tactic
+ val add_inversion_lemma_exn :
+ Names.Id.t -> Constrexpr.constr_expr -> Misctypes.glob_sort -> bool -> (Names.Id.t -> unit Proofview.tactic) ->
+ unit
end
module Hints :
sig
- type hint = Hints.hint
- type debug = Hints.debug =
- | Debug | Info | Off
- type 'a hints_path_atom_gen = 'a Hints.hints_path_atom_gen =
+
+ type raw_hint = EConstr.t * EConstr.types * Univ.universe_context_set
+
+ type 'a hint_ast =
+ | Res_pf of 'a (* Hint Apply *)
+ | ERes_pf of 'a (* Hint EApply *)
+ | Give_exact of 'a
+ | Res_pf_THEN_trivial_fail of 'a (* Hint Immediate *)
+ | Unfold_nth of Names.evaluable_global_reference (* Hint Unfold *)
+ | Extern of Genarg.glob_generic_argument (* Hint Extern *)
+
+ type hint
+
+ type debug =
+ | Debug | Info | Off
+
+ type 'a hints_path_atom_gen =
| PathHints of 'a list
| PathAny
- type hint_term = Hints.hint_term =
+
+ type hint_term =
| IsGlobRef of Globnames.global_reference
| IsConstr of EConstr.constr * Univ.ContextSet.t
+
type hint_db_name = string
type hint_info = (Names.Id.t list * Pattern.constr_pattern) Vernacexpr.hint_info_gen
type hnf = bool
type hints_path_atom = Globnames.global_reference hints_path_atom_gen
- type 'a hints_path_gen = 'a Hints.hints_path_gen =
+ type 'a hints_path_gen =
| PathAtom of 'a hints_path_atom_gen
| PathStar of 'a hints_path_gen
| PathSeq of 'a hints_path_gen * 'a hints_path_gen
@@ -4348,7 +5362,7 @@ sig
type hints_path = Globnames.global_reference hints_path_gen
- type hints_entry = Hints.hints_entry =
+ type hints_entry =
| HintsResolveEntry of (hint_info * Decl_kinds.polymorphic * hnf * hints_path_atom * hint_term) list
| HintsImmediateEntry of (hints_path_atom * Decl_kinds.polymorphic * hint_term) list
| HintsCutEntry of hints_path
@@ -4357,15 +5371,7 @@ sig
| HintsModeEntry of Globnames.global_reference * Vernacexpr.hint_mode list
| HintsExternEntry of hint_info * Genarg.glob_generic_argument
- type 'a hint_ast = 'a Hints.hint_ast =
- | Res_pf of 'a
- | ERes_pf of 'a
- | Give_exact of 'a
- | Res_pf_THEN_trivial_fail of 'a
- | Unfold_nth of Names.evaluable_global_reference
- | Extern of Genarg.glob_generic_argument
- type raw_hint = EConstr.constr * EConstr.types * Univ.ContextSet.t
- type 'a with_metadata = 'a Hints.with_metadata = private {
+ type 'a with_metadata = private {
pri : int;
poly : Decl_kinds.polymorphic;
pat : Pattern.constr_pattern option;
@@ -4378,7 +5384,7 @@ sig
module Hint_db :
sig
- type t = Hints.Hint_db.t
+ type t
val empty : ?name:hint_db_name -> Names.transparent_state -> bool -> t
val transparent_state : t -> Names.transparent_state
val iter : (Globnames.global_reference option ->
@@ -4386,15 +5392,15 @@ sig
end
type hint_db = Hint_db.t
- val add_hints : Vernacexpr.locality_flag -> hint_db_name list -> hints_entry -> unit
+ val add_hints : bool -> hint_db_name list -> hints_entry -> unit
val searchtable_map : hint_db_name -> hint_db
val pp_hints_path_atom : ('a -> Pp.std_ppcmds) -> 'a hints_path_atom_gen -> Pp.std_ppcmds
val pp_hints_path_gen : ('a -> Pp.std_ppcmds) -> 'a hints_path_gen -> Pp.std_ppcmds
val glob_hints_path_atom :
- Prelude.reference hints_path_atom_gen -> Globnames.global_reference hints_path_atom_gen
+ Libnames.reference hints_path_atom_gen -> Globnames.global_reference hints_path_atom_gen
val pp_hints_path : hints_path -> Pp.std_ppcmds
val glob_hints_path :
- Prelude.reference hints_path_gen -> Globnames.global_reference hints_path_gen
+ Libnames.reference hints_path_gen -> Globnames.global_reference hints_path_gen
val typeclasses_db : hint_db_name
val add_hints_init : (unit -> unit) -> unit
val create_hint_db : bool -> hint_db_name -> Names.transparent_state -> bool -> unit
@@ -4421,140 +5427,6 @@ sig
val default_full_auto : unit Proofview.tactic
end
-module Hipattern :
-sig
- exception NoEquationFound
- type 'a matching_function = Evd.evar_map -> EConstr.constr -> 'a option
- type testing_function = Evd.evar_map -> EConstr.constr -> bool
- val is_disjunction : ?strict:bool -> ?onlybinary:bool -> testing_function
- val match_with_disjunction : ?strict:bool -> ?onlybinary:bool -> (EConstr.constr * EConstr.constr list) matching_function
- val match_with_equality_type : (EConstr.constr * EConstr.constr list) matching_function
- val is_empty_type : testing_function
- val is_unit_type : testing_function
- val is_unit_or_eq_type : testing_function
- val is_conjunction : ?strict:bool -> ?onlybinary:bool -> testing_function
- val match_with_conjunction : ?strict:bool -> ?onlybinary:bool -> (EConstr.constr * EConstr.constr list) matching_function
- val match_with_imp_term : (EConstr.constr * EConstr.constr) matching_function
- val match_with_forall_term : (Names.Name.t * EConstr.constr * EConstr.constr) matching_function
- val match_with_nodep_ind : (EConstr.constr * EConstr.constr list * int) matching_function
- val match_with_sigma_type : (EConstr.constr * EConstr.constr list) matching_function
-end
-
-module Inv :
-sig
- val dinv :
- Misctypes.inversion_kind -> EConstr.constr option ->
- Tactypes.or_and_intro_pattern option -> Misctypes.quantified_hypothesis -> unit Proofview.tactic
- val inv_clause :
- Misctypes.inversion_kind -> Tactypes.or_and_intro_pattern option -> Names.Id.t list ->
- Misctypes.quantified_hypothesis -> unit Proofview.tactic
- val inv_clear_tac : Names.Id.t -> unit Proofview.tactic
- val inv_tac : Names.Id.t -> unit Proofview.tactic
- val dinv_tac : Names.Id.t -> unit Proofview.tactic
- val dinv_clear_tac : Names.Id.t -> unit Proofview.tactic
- val inv : Misctypes.inversion_kind -> Tactypes.or_and_intro_pattern option ->
- Misctypes.quantified_hypothesis -> unit Proofview.tactic
-end
-
-module Leminv :
-sig
- val lemInv_clause :
- Misctypes.quantified_hypothesis -> EConstr.constr -> Names.Id.t list -> unit Proofview.tactic
- val add_inversion_lemma_exn :
- Names.Id.t -> Constrexpr.constr_expr -> Misctypes.glob_sort -> bool -> (Names.Id.t -> unit Proofview.tactic) ->
- unit
-end
-
-module Vernacentries :
-sig
- val dump_global : Prelude.reference Misctypes.or_by_notation -> unit
- val interp_redexp_hook : (Environ.env -> Evd.evar_map -> Genredexpr.raw_red_expr ->
- Evd.evar_map * Redexpr.red_expr) Hook.t
- val command_focus : unit Proof.focus_kind
-end
-
-module Evar_refiner :
-sig
- val w_refine : Evar.t * Evd.evar_info ->
- Pretyping.glob_constr_ltac_closure -> Evd.evar_map -> Evd.evar_map
-end
-
-module Obligations :
-sig
- val default_tactic : unit Proofview.tactic ref
- val obligation : int * Names.Id.t option * Constrexpr.constr_expr option ->
- Genarg.glob_generic_argument option -> unit
- val next_obligation : Names.Id.t option -> Genarg.glob_generic_argument option -> unit
- val try_solve_obligation : int -> Names.Id.t option -> unit Proofview.tactic option -> unit
- val try_solve_obligations : Names.Id.t option -> unit Proofview.tactic option -> unit
- val solve_all_obligations : unit Proofview.tactic option -> unit
- val admit_obligations : Names.Id.t option -> unit
- val show_obligations : ?msg:bool -> Names.Id.t option -> unit
- val show_term : Names.Id.t option -> Pp.std_ppcmds
-end
-
-module Elim :
-sig
- val h_decompose : Names.inductive list -> EConstr.constr -> unit Proofview.tactic
- val h_double_induction : Misctypes.quantified_hypothesis -> Misctypes.quantified_hypothesis-> unit Proofview.tactic
- val h_decompose_or : EConstr.constr -> unit Proofview.tactic
- val h_decompose_and : EConstr.constr -> unit Proofview.tactic
-end
-
-module Redops :
-sig
- val all_flags : 'a Genredexpr.glob_red_flag
- val make_red_flag : 'a Genredexpr.red_atom list -> 'a Genredexpr.glob_red_flag
-end
-
-module Autorewrite :
-sig
- type rew_rule = { rew_lemma: Term.constr;
- rew_type: Term.types;
- rew_pat: Term.constr;
- rew_ctx: Univ.ContextSet.t;
- rew_l2r: bool;
- rew_tac: Genarg.glob_generic_argument option }
- type raw_rew_rule = (Term.constr Univ.in_universe_context_set * bool *
- Genarg.raw_generic_argument option)
- Loc.located
- val auto_multi_rewrite : ?conds:Equality.conditions -> string list -> Locus.clause -> unit Proofview.tactic
- val auto_multi_rewrite_with : ?conds:Equality.conditions -> unit Proofview.tactic -> string list -> Locus.clause -> unit Proofview.tactic
- val add_rew_rules : string -> raw_rew_rule list -> unit
- val find_rewrites : string -> rew_rule list
- val find_matches : string -> Term.constr -> rew_rule list
- val print_rewrite_hintdb : string -> Pp.std_ppcmds
-end
-
-module Refine :
-sig
- val refine : typecheck:bool -> (Evd.evar_map -> Evd.evar_map * EConstr.t) -> unit Proofview.tactic
- val solve_constraints : unit Proofview.tactic
-end
-
-module Find_subterm :
-sig
- val error_invalid_occurrence : int list -> 'a
-end
-
-module Vernac_classifier :
-sig
- val declare_vernac_classifier :
- Vernacexpr.extend_name -> (Genarg.raw_generic_argument list -> unit -> Vernacexpr.vernac_classification) -> unit
- val classify_as_proofstep : Vernacexpr.vernac_classification
- val classify_as_query : Vernacexpr.vernac_classification
- val classify_as_sideeff : Vernacexpr.vernac_classification
- val classify_vernac : Vernacexpr.vernac_expr -> Vernacexpr.vernac_classification
-end
-
-module Keys :
-sig
- type key = Keys.key
- val constr_key : ('a -> ('a, 't, 'u, 'i) Term.kind_of_term) -> 'a -> key option
- val declare_equiv_keys : key -> key -> unit
- val pr_keys : (Globnames.global_reference -> Pp.std_ppcmds) -> Pp.std_ppcmds
-end
-
module Eauto :
sig
val e_assumption : unit Proofview.tactic
@@ -4566,14 +5438,16 @@ sig
val autounfold_tac : Hints.hint_db_name list option -> Locus.clause -> unit Proofview.tactic
val autounfold_one : Hints.hint_db_name list -> Locus.hyp_location option -> unit Proofview.tactic
val eauto_with_bases :
- ?debug:Hints.debug -> bool * int -> Tactypes.delayed_open_constr list -> Hints.hint_db list -> Tacmach.tactic
+ ?debug:Hints.debug -> bool * int -> Tactypes.delayed_open_constr list -> Hints.hint_db list -> Proof_type.tactic
end
module Class_tactics :
sig
- type search_strategy = Class_tactics.search_strategy =
+
+ type search_strategy =
| Dfs
| Bfs
+
val set_typeclasses_debug : bool -> unit
val set_typeclasses_strategy : search_strategy -> unit
val set_typeclasses_depth : int option -> unit
@@ -4587,34 +5461,56 @@ sig
val catchable : exn -> bool
end
-module Ind_tables :
+module Eqdecide :
sig
- type individual = Ind_tables.individual
- type 'a scheme_kind = 'a Ind_tables.scheme_kind
+ val compare : EConstr.constr -> EConstr.constr -> unit Proofview.tactic
+ val decideEqualityGoal : unit Proofview.tactic
+end
- val check_scheme : 'a scheme_kind -> Names.inductive -> bool
- val find_scheme : ?mode:Declare.internal_flag -> 'a scheme_kind -> Names.inductive -> Names.Constant.t * Safe_typing.private_constants
- val pr_scheme_kind : 'a scheme_kind -> Pp.std_ppcmds
+module Autorewrite :
+sig
+ type rew_rule = { rew_lemma: Constr.t;
+ rew_type: Term.types;
+ rew_pat: Constr.t;
+ rew_ctx: Univ.ContextSet.t;
+ rew_l2r: bool;
+ rew_tac: Genarg.glob_generic_argument option }
+ type raw_rew_rule = (Constr.t Univ.in_universe_context_set * bool *
+ Genarg.raw_generic_argument option)
+ Loc.located
+ val auto_multi_rewrite : ?conds:Equality.conditions -> string list -> Locus.clause -> unit Proofview.tactic
+ val auto_multi_rewrite_with : ?conds:Equality.conditions -> unit Proofview.tactic -> string list -> Locus.clause -> unit Proofview.tactic
+ val add_rew_rules : string -> raw_rew_rule list -> unit
+ val find_rewrites : string -> rew_rule list
+ val find_matches : string -> Constr.t -> rew_rule list
+ val print_rewrite_hintdb : string -> Pp.std_ppcmds
end
-module Elimschemes :
+(************************************************************************)
+(* End of modules from tactics/ *)
+(************************************************************************)
+
+(************************************************************************)
+(* Modules from vernac/ *)
+(************************************************************************)
+
+module Ppvernac :
sig
- val case_scheme_kind_from_prop : Ind_tables.individual Ind_tables.scheme_kind
- val case_dep_scheme_kind_from_type_in_prop : Ind_tables.individual Ind_tables.scheme_kind
- val case_scheme_kind_from_type : Ind_tables.individual Ind_tables.scheme_kind
- val case_dep_scheme_kind_from_type : Ind_tables.individual Ind_tables.scheme_kind
- val case_dep_scheme_kind_from_prop : Ind_tables.individual Ind_tables.scheme_kind
+ val pr_vernac : Vernacexpr.vernac_expr -> Pp.std_ppcmds
+ val pr_rec_definition : (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) -> Pp.std_ppcmds
end
module Lemmas :
sig
- type 'a declaration_hook = 'a Lemmas.declaration_hook
+
+ type 'a declaration_hook
+
val mk_hook :
(Decl_kinds.locality -> Globnames.global_reference -> 'a) -> 'a declaration_hook
val start_proof : Names.Id.t -> ?pl:Proof_global.universe_binders -> Decl_kinds.goal_kind -> Evd.evar_map ->
?terminator:(Proof_global.lemma_possible_guards -> unit declaration_hook -> Proof_global.proof_terminator) ->
?sign:Environ.named_context_val -> EConstr.types ->
- ?init_tac:unit Proofview.tactic -> ?compute_guard:Proof_global.lemma_possible_guards ->
+ ?init_tac:unit Proofview.tactic -> ?compute_guard:Proof_global.lemma_possible_guards ->
unit declaration_hook -> unit
val call_hook :
Future.fix_exn -> 'a declaration_hook -> Decl_kinds.locality -> Globnames.global_reference -> 'a
@@ -4622,150 +5518,94 @@ sig
val get_current_context : unit -> Evd.evar_map * Environ.env
end
-module Eqdecide :
+module Himsg :
sig
- val compare : EConstr.constr -> EConstr.constr -> unit Proofview.tactic
- val decideEqualityGoal : unit Proofview.tactic
+ val explain_refiner_error : Logic.refiner_error -> Pp.std_ppcmds
+ val explain_pretype_error : Environ.env -> Evd.evar_map -> Pretype_errors.pretype_error -> Pp.std_ppcmds
end
-module Locusops :
+module ExplainErr :
sig
- val clause_with_generic_occurrences : 'a Locus.clause_expr -> bool
- val nowhere : 'a Locus.clause_expr
- val allHypsAndConcl : 'a Locus.clause_expr
- val is_nowhere : 'a Locus.clause_expr -> bool
- val occurrences_map :
- ('a list -> 'b list) -> 'a Locus.occurrences_gen -> 'b Locus.occurrences_gen
- val convert_occs : Locus.occurrences -> bool * int list
- val onConcl : 'a Locus.clause_expr
- val onHyp : 'a -> 'a Locus.clause_expr
+ val process_vernac_interp_error : ?allow_uncaught:bool -> Util.iexn -> Util.iexn
+ val register_additional_error_info : (Util.iexn -> Pp.std_ppcmds option Loc.located option) -> unit
end
-module Topfmt :
+module Locality :
sig
- val std_ft : Format.formatter ref
- val with_output_to : out_channel -> Format.formatter
- val get_margin : unit -> int option
+ val make_section_locality : bool option -> bool
+ module LocalityFixme : sig
+ val consume : unit -> bool option
+ end
+ val make_module_locality : bool option -> bool
end
-module Nameops :
+module Metasyntax :
sig
- val atompart_of_id : Names.Id.t -> string
- val pr_id : Names.Id.t -> Pp.std_ppcmds
- [@@ocaml.deprecated "alias of API.Names.Id.print"]
+ val add_token_obj : string -> unit
- val pr_name : Names.Name.t -> Pp.std_ppcmds
- [@@ocaml.deprecated "alias of API.Names.Name.print"]
+ type any_entry = AnyEntry : 'a Pcoq.Gram.entry -> any_entry
+ val register_grammar : string -> any_entry list -> unit
- val name_fold : (Names.Id.t -> 'a -> 'a) -> Names.Name.t -> 'a -> 'a
- val name_app : (Names.Id.t -> Names.Id.t) -> Names.Name.t -> Names.Name.t
- val add_suffix : Names.Id.t -> string -> Names.Id.t
- val increment_subscript : Names.Id.t -> Names.Id.t
- val make_ident : string -> int option -> Names.Id.t
- val out_name : Names.Name.t -> Names.Id.t
- val pr_lab : Names.Label.t -> Pp.std_ppcmds
- module Name :
- sig
- include module type of struct include Names.Name end
- val get_id : t -> Names.Id.t
- val fold_right : (Names.Id.t -> 'a -> 'a) -> t -> 'a -> 'a
- end
end
-module Declareops :
+module Search :
sig
- val constant_has_body : Declarations.constant_body -> bool
- val is_opaque : Declarations.constant_body -> bool
- val eq_recarg : Declarations.recarg -> Declarations.recarg -> bool
+ type glob_search_about_item =
+ | GlobSearchSubPattern of Pattern.constr_pattern
+ | GlobSearchString of string
+ type filter_function = Globnames.global_reference -> Environ.env -> Constr.t -> bool
+ type display_function = Globnames.global_reference -> Environ.env -> Constr.t -> unit
+ val search_about_filter : glob_search_about_item -> filter_function
+ val module_filter : Names.DirPath.t list * bool -> filter_function
+ val generic_search : int option -> display_function -> unit
end
-module Constr :
+module Obligations :
sig
- type t = Term.constr
- [@@ocaml.deprecated "alias of API.Term.constr"]
-
- type constr = Term.constr
- [@@ocaml.deprecated "alias of API.Term.constr"]
-
- type types = Term.constr
- [@@ocaml.deprecated "alias of API.Term.types"]
-
- type cast_kind = Term.cast_kind =
- | VMcast
- | NATIVEcast
- | DEFAULTcast
- | REVERTcast
- type ('constr, 'types, 'sort, 'univs) kind_of_term = ('constr, 'types, 'sort, 'univs) Term.kind_of_term =
- | Rel of int
- | Var of Names.Id.t
- | Meta of Term.metavariable
- | Evar of 'constr Term.pexistential
- | Sort of 'sort
- | Cast of 'constr * cast_kind * 'types
- | Prod of Names.Name.t * 'types * 'types
- | Lambda of Names.Name.t * 'types * 'constr
- | LetIn of Names.Name.t * 'constr * 'types * 'constr
- | App of 'constr * 'constr array
- | Const of (Names.Constant.t * 'univs)
- | Ind of (Names.inductive * 'univs)
- | Construct of (Names.constructor * 'univs)
- | Case of Term.case_info * 'constr * 'constr * 'constr array
- | Fix of ('constr, 'types) Term.pfixpoint
- | CoFix of ('constr, 'types) Term.pcofixpoint
- | Proj of Names.Projection.t * 'constr
- [@@ocaml.deprecated "alias of API.Term.cast_kind"]
-
- val equal : Term.constr -> Term.constr -> bool
- [@@ocaml.deprecated "alias of API.Term.eq_constr"]
-
- val mkIndU : Term.pinductive -> Term.constr
- [@@ocaml.deprecated "alias of API.Term.mkIndU"]
-
- val mkConstU : Term.pconstant -> Term.constr
- [@@ocaml.deprecated "alias of API.Term.mkConstU"]
-
- val mkConst : Names.Constant.t -> Term.constr
- [@@ocaml.deprecated "alias of API.Term.mkConst"]
-
- val mkVar : Names.Id.t -> Term.constr
- [@@ocaml.deprecated "alias of API.Term.mkVar"]
-
- val compare : Term.constr -> Term.constr -> int
- [@@ocaml.deprecated "alias of API.Term.constr_ord"]
-
- val mkApp : Term.constr * Term.constr array -> Term.constr
- [@@ocaml.deprecated "alias of API.Term.mkApp"]
+ val default_tactic : unit Proofview.tactic ref
+ val obligation : int * Names.Id.t option * Constrexpr.constr_expr option ->
+ Genarg.glob_generic_argument option -> unit
+ val next_obligation : Names.Id.t option -> Genarg.glob_generic_argument option -> unit
+ val try_solve_obligation : int -> Names.Id.t option -> unit Proofview.tactic option -> unit
+ val try_solve_obligations : Names.Id.t option -> unit Proofview.tactic option -> unit
+ val solve_all_obligations : unit Proofview.tactic option -> unit
+ val admit_obligations : Names.Id.t option -> unit
+ val show_obligations : ?msg:bool -> Names.Id.t option -> unit
+ val show_term : Names.Id.t option -> Pp.std_ppcmds
end
-[@@ocaml.deprecated "alias of API.Term"]
-module Coq_config :
+module Command :
sig
- val exec_extension : string
-end
+ open Names
+ open Constrexpr
+ open Vernacexpr
+
+ type structured_fixpoint_expr = {
+ fix_name : Id.t;
+ fix_univs : lident list option;
+ fix_annot : Id.t Loc.located option;
+ fix_binders : local_binder_expr list;
+ fix_body : constr_expr option;
+ fix_type : constr_expr
+ }
-module Kindops :
-sig
- val logical_kind_of_goal_kind : Decl_kinds.goal_object_kind -> Decl_kinds.logical_kind
-end
+ type structured_one_inductive_expr = {
+ ind_name : Id.t;
+ ind_univs : lident list option;
+ ind_arity : constr_expr;
+ ind_lc : (Id.t * constr_expr) list
+ }
-module States :
-sig
- val with_state_protection_on_exception : ('a -> 'b) -> 'a -> 'b
- val with_state_protection : ('a -> 'b) -> 'a -> 'b
-end
+ type structured_inductive_expr =
+ local_binder_expr list * structured_one_inductive_expr list
-module Command :
-sig
- type structured_fixpoint_expr = Command.structured_fixpoint_expr
- type recursive_preentry = Names.Id.t list * Term.constr option list * Term.types list
- type structured_inductive_expr = Command.structured_inductive_expr
- type one_inductive_impls = Command.one_inductive_impls
+ type recursive_preentry = Names.Id.t list * Constr.t option list * Constr.types list
+
+ type one_inductive_impls
val do_mutual_inductive :
- (Vernacexpr.one_inductive_expr * Vernacexpr.decl_notation list) list ->
- Decl_kinds.cumulative_inductive_flag ->
- Decl_kinds.polymorphic ->
+ (Vernacexpr.one_inductive_expr * Vernacexpr.decl_notation list) list -> Decl_kinds.cumulative_inductive_flag -> Decl_kinds.polymorphic ->
Decl_kinds.private_flag -> Decl_kinds.recursivity_kind -> unit
val do_definition : Names.Id.t -> Decl_kinds.definition_kind -> Vernacexpr.lident list option ->
@@ -4781,7 +5621,7 @@ sig
val interp_fixpoint :
structured_fixpoint_expr list -> Vernacexpr.decl_notation list ->
- recursive_preentry * Vernacexpr.lident list option * UState.t *
+ recursive_preentry * Vernacexpr.lident list option * UState.t *
(EConstr.rel_context * Impargs.manual_implicits * int option) list
val extract_mutual_inductive_declaration_components :
@@ -4800,14 +5640,108 @@ sig
Names.MutInd.t
end
-module Ppvernac :
+module Classes :
sig
- val pr_vernac : Vernacexpr.vernac_expr -> Pp.std_ppcmds
- val pr_rec_definition : (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) -> Pp.std_ppcmds
+ val set_typeclass_transparency : Names.evaluable_global_reference -> bool -> bool -> unit
+ val new_instance :
+ ?abstract:bool ->
+ ?global:bool ->
+ ?refine:bool ->
+ Decl_kinds.polymorphic ->
+ Constrexpr.local_binder_expr list ->
+ Constrexpr.typeclass_constraint ->
+ (bool * Constrexpr.constr_expr) option ->
+ ?generalize:bool ->
+ ?tac:unit Proofview.tactic ->
+ ?hook:(Globnames.global_reference -> unit) ->
+ Vernacexpr.hint_info_expr ->
+ Names.Id.t
end
-module Topconstr :
+module Vernacinterp :
sig
- val replace_vars_constr_expr :
- Names.Id.t Names.Id.Map.t -> Constrexpr.constr_expr -> Constrexpr.constr_expr
+ type deprecation = bool
+
+ type vernac_command = Genarg.raw_generic_argument list -> unit -> unit
+
+ val vinterp_add : deprecation -> Vernacexpr.extend_name ->
+ vernac_command -> unit
+
+end
+
+module Mltop :
+sig
+ val declare_cache_obj : (unit -> unit) -> string -> unit
+ val add_known_plugin : (unit -> unit) -> string -> unit
+ val add_known_module : string -> unit
+ val module_is_known : string -> bool
+end
+
+module Topfmt :
+sig
+ val std_ft : Format.formatter ref
+ val with_output_to : out_channel -> Format.formatter
+ val get_margin : unit -> int option
+end
+
+module Vernacentries :
+sig
+ val dump_global : Libnames.reference Misctypes.or_by_notation -> unit
+ val interp_redexp_hook : (Environ.env -> Evd.evar_map -> Genredexpr.raw_red_expr ->
+ Evd.evar_map * Redexpr.red_expr) Hook.t
+ val command_focus : unit Proof.focus_kind
+end
+
+(************************************************************************)
+(* End of modules from vernac/ *)
+(************************************************************************)
+
+(************************************************************************)
+(* Modules from stm/ *)
+(************************************************************************)
+
+module Vernac_classifier :
+sig
+ val declare_vernac_classifier :
+ Vernacexpr.extend_name -> (Genarg.raw_generic_argument list -> unit -> Vernacexpr.vernac_classification) -> unit
+ val classify_as_proofstep : Vernacexpr.vernac_classification
+ val classify_as_query : Vernacexpr.vernac_classification
+ val classify_as_sideeff : Vernacexpr.vernac_classification
+ val classify_vernac : Vernacexpr.vernac_expr -> Vernacexpr.vernac_classification
+end
+
+module Stm :
+sig
+ type state
+ val state_of_id :
+ Stateid.t -> [ `Valid of state option | `Expired | `Error of exn ]
end
+
+(************************************************************************)
+(* End of modules from stm/ *)
+(************************************************************************)
+
+(************************************************************************)
+(* Modules from highparsing/ *)
+(************************************************************************)
+
+module G_vernac :
+sig
+
+ val def_body : Vernacexpr.definition_expr Pcoq.Gram.entry
+ val section_subset_expr : Vernacexpr.section_subset_expr Pcoq.Gram.entry
+ val query_command : (Vernacexpr.goal_selector option -> Vernacexpr.vernac_expr) Pcoq.Gram.entry
+
+end
+
+module G_proofs :
+sig
+
+ val hint : Vernacexpr.hints_expr Pcoq.Gram.entry
+ val hint_proof_using : 'a Pcoq.Gram.entry -> 'a option -> 'a option
+
+end
+
+(************************************************************************)
+(* End of modules from highparsing/ *)
+(************************************************************************)
diff --git a/API/API.mllib b/API/API.mllib
index f4bdf83db..25275c704 100644
--- a/API/API.mllib
+++ b/API/API.mllib
@@ -1,2 +1 @@
API
-Grammar_API
diff --git a/API/grammar_API.ml b/API/grammar_API.ml
deleted file mode 100644
index 485c16665..000000000
--- a/API/grammar_API.ml
+++ /dev/null
@@ -1,63 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-module G_proofs = G_proofs
-module Metasyntax = Metasyntax
-module Egramcoq = Egramcoq
-module G_vernac = G_vernac
-module Pcoq = Pcoq
-module Tok = Tok
-module CLexer = CLexer
-module Egramml = Egramml
-module Mltop = Mltop
-module Vernacinterp = Vernacinterp
-module Genintern = Genintern
-
-module Extend =
- struct
- type 'a entry = 'a Extend.entry
- type ('self, 'a) symbol = ('self, 'a) Extend.symbol =
- | Atoken : Tok.t -> ('self, string) symbol
- | Alist1 : ('self, 'a) symbol -> ('self, 'a list) symbol
- | Alist1sep : ('self, 'a) symbol * ('self, _) symbol -> ('self, 'a list) symbol
- | Alist0 : ('self, 'a) symbol -> ('self, 'a list) symbol
- | Alist0sep : ('self, 'a) symbol * ('self, _) symbol -> ('self, 'a list) symbol
- | Aopt : ('self, 'a) symbol -> ('self, 'a option) symbol
- | Aself : ('self, 'self) symbol
- | Anext : ('self, 'self) symbol
- | Aentry : 'a entry -> ('self, 'a) symbol
- | Aentryl : 'a entry * int -> ('self, 'a) symbol
- | Arules : 'a rules list -> ('self, 'a) symbol
- and ('self, 'a, 'r) rule = ('self, 'a, 'r) Extend.rule =
- | Stop : ('self, 'r, 'r) rule
- | Next : ('self, 'a, 'r) rule * ('self, 'b) symbol -> ('self, 'b -> 'a, 'r) rule
- and ('a, 'r) norec_rule = ('a, 'r) Extend.norec_rule =
- { norec_rule : 's. ('s, 'a, 'r) rule }
- and 'a rules = 'a Extend.rules =
- | Rules : ('act, Loc.t -> 'a) norec_rule * 'act -> 'a rules
- type gram_assoc = Extend.gram_assoc = NonA | RightA | LeftA
- type 'a production_rule = 'a Extend.production_rule =
- | Rule : ('a, 'act, Loc.t -> 'a) rule * 'act -> 'a production_rule
- type 'a single_extend_statment = string option * gram_assoc option * 'a production_rule list
- type gram_position = Extend.gram_position =
- | First
- | Last
- | Before of string
- | After of string
- | Level of string
- type 'a extend_statment = Extend.gram_position option * 'a single_extend_statment list
-
- type 'a user_symbol = 'a Extend.user_symbol =
- | Ulist1 of 'a user_symbol
- | Ulist1sep of 'a user_symbol * string
- | Ulist0 of 'a user_symbol
- | Ulist0sep of 'a user_symbol * string
- | Uopt of 'a user_symbol
- | Uentry of 'a
- | Uentryl of 'a * int
- end
diff --git a/API/grammar_API.mli b/API/grammar_API.mli
deleted file mode 100644
index c2115a506..000000000
--- a/API/grammar_API.mli
+++ /dev/null
@@ -1,249 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-module Extend :
-sig
- type 'a entry = 'a Pcoq.Gram.Entry.e
- type ('self, 'a) symbol = ('self, 'a) Extend.symbol =
- | Atoken : Tok.t -> ('self, string) symbol
- | Alist1 : ('self, 'a) symbol -> ('self, 'a list) symbol
- | Alist1sep : ('self, 'a) symbol * ('self, _) symbol -> ('self, 'a list) symbol
- | Alist0 : ('self, 'a) symbol -> ('self, 'a list) symbol
- | Alist0sep : ('self, 'a) symbol * ('self, _) symbol -> ('self, 'a list) symbol
- | Aopt : ('self, 'a) symbol -> ('self, 'a option) symbol
- | Aself : ('self, 'self) symbol
- | Anext : ('self, 'self) symbol
- | Aentry : 'a entry -> ('self, 'a) symbol
- | Aentryl : 'a entry * int -> ('self, 'a) symbol
- | Arules : 'a rules list -> ('self, 'a) symbol
- and ('self, 'a, 'r) rule = ('self, 'a, 'r) Extend.rule =
- | Stop : ('self, 'r, 'r) rule
- | Next : ('self, 'a, 'r) rule * ('self, 'b) symbol -> ('self, 'b -> 'a, 'r) rule
- and ('a, 'r) norec_rule = ('a, 'r) Extend.norec_rule =
- { norec_rule : 's. ('s, 'a, 'r) rule }
- and 'a rules = 'a Extend.rules =
- | Rules : ('act, Loc.t -> 'a) norec_rule * 'act -> 'a rules
- type gram_assoc = Extend.gram_assoc = NonA | RightA | LeftA
- type 'a production_rule = 'a Extend.production_rule =
- | Rule : ('a, 'act, Loc.t -> 'a) rule * 'act -> 'a production_rule
- type 'a single_extend_statment = string option * gram_assoc option * 'a production_rule list
- type gram_position = Extend.gram_position =
- | First
- | Last
- | Before of string
- | After of string
- | Level of string
- type 'a extend_statment = gram_position option * 'a single_extend_statment list
- type 'a user_symbol = 'a Extend.user_symbol =
- | Ulist1 of 'a user_symbol
- | Ulist1sep of 'a user_symbol * string
- | Ulist0 of 'a user_symbol
- | Ulist0sep of 'a user_symbol * string
- | Uopt of 'a user_symbol
- | Uentry of 'a
- | Uentryl of 'a * int
-end
-
-module Genintern :
-sig
- open API
- module Store : module type of struct include Genintern.Store end
- type glob_sign = Genintern.glob_sign =
- { ltacvars : Names.Id.Set.t;
- genv : Environ.env;
- extra : Store.t }
- type ('raw, 'glb) intern_fun = glob_sign -> 'raw -> glob_sign * 'glb
- type 'glb subst_fun = Mod_subst.substitution -> 'glb -> 'glb
- type 'glb ntn_subst_fun = Tactypes.glob_constr_and_expr Names.Id.Map.t -> 'glb -> 'glb
- val empty_glob_sign : Environ.env -> glob_sign
- val register_intern0 : ('raw, 'glb, 'top) Genarg.genarg_type ->
- ('raw, 'glb) intern_fun -> unit
- val register_subst0 : ('raw, 'glb, 'top) Genarg.genarg_type ->
- 'glb subst_fun -> unit
- val register_ntn_subst0 : ('raw, 'glb, 'top) Genarg.genarg_type ->
- 'glb ntn_subst_fun -> unit
- val generic_substitute : Genarg.glob_generic_argument subst_fun
- val generic_intern : (Genarg.raw_generic_argument, Genarg.glob_generic_argument) intern_fun
-end
-
-module Tok :
-sig
- type t = Tok.t =
- | KEYWORD of string
- | PATTERNIDENT of string
- | IDENT of string
- | FIELD of string
- | INT of string
- | STRING of string
- | LEFTQMARK
- | BULLET of string
- | EOI
-end
-
-module Pcoq :
-sig
- type gram_universe = Pcoq.gram_universe
- module Gram :
- sig
- type te = Tok.t
- module Entry :
- sig
- type 'a e = 'a Extend.entry
- val of_parser : string -> (te Stream.t -> 'a) -> 'a e
- val obj : 'a e -> te Gramext.g_entry
- val create : string -> 'a e
- end
- type 'a entry = 'a Entry.e
- val extend : 'a Pcoq.Gram.Entry.e -> Gramext.position option ->
- (string option * Gramext.g_assoc option *
- (Tok.t Gramext.g_symbol list * Gramext.g_action) list) list -> unit
- val entry_create : string -> 'a Entry.e
- end
- module Prim : sig
- open Names
- open Loc
- val preident : string Gram.Entry.e
- val ident : Names.Id.t Gram.Entry.e
- val name : Name.t located Gram.Entry.e
- val identref : Names.Id.t located Gram.Entry.e
- val pidentref : (Names.Id.t located * (Names.Id.t located list) option) Gram.Entry.e
- val pattern_ident : Names.Id.t Gram.Entry.e
- val pattern_identref : Names.Id.t located Gram.Entry.e
- val base_ident : Names.Id.t Gram.Entry.e
- val natural : int Gram.Entry.e
- val bigint : Constrexpr.raw_natural_number Gram.Entry.e
- val integer : int Gram.Entry.e
- val string : string Gram.Entry.e
- val qualid : API.Libnames.qualid located Gram.Entry.e
- val fullyqualid : Names.Id.t list located Gram.Entry.e
- val reference : API.Libnames.reference Gram.Entry.e
- val by_notation : (string * string option) Loc.located Gram.entry
- val smart_global : API.Libnames.reference API.Misctypes.or_by_notation Gram.Entry.e
- val dirpath : DirPath.t Gram.Entry.e
- val ne_string : string Gram.Entry.e
- val ne_lstring : string located Gram.Entry.e
- val var : Names.Id.t located Gram.Entry.e
- end
-
- val eoi_entry : 'a Gram.Entry.e -> 'a Gram.Entry.e
- val create_generic_entry : gram_universe -> string ->
- ('a, Genarg.rlevel) Genarg.abstract_argument_type -> 'a Gram.Entry.e
- val utactic : gram_universe
- type gram_reinit = Extend.gram_assoc * Extend.gram_position
- val grammar_extend : 'a Gram.Entry.e -> gram_reinit option ->
- 'a Extend.extend_statment -> unit
- val genarg_grammar : ('raw, 'glb, 'top) Genarg.genarg_type -> 'raw Gram.Entry.e
- val register_grammar : ('raw, 'glb, 'top) Genarg.genarg_type -> 'raw Gram.Entry.e -> unit
- module Constr :
- sig
- val sort : API.Misctypes.glob_sort Gram.Entry.e
- val lconstr : API.Constrexpr.constr_expr Gram.Entry.e
- val lconstr_pattern : API.Constrexpr.constr_expr Gram.Entry.e
- val ident : API.Names.Id.t Gram.Entry.e
- val constr : API.Constrexpr.constr_expr Gram.Entry.e
- val closed_binder : API.Constrexpr.local_binder_expr list Gram.Entry.e
- val constr_pattern : API.Constrexpr.constr_expr Gram.Entry.e
- val global : API.Libnames.reference Gram.Entry.e
- val binder_constr : API.Constrexpr.constr_expr Gram.Entry.e
- val operconstr : API.Constrexpr.constr_expr Gram.Entry.e
- val pattern : API.Constrexpr.cases_pattern_expr Gram.Entry.e
- val binders : API.Constrexpr.local_binder_expr list Gram.Entry.e
- end
- module Vernac_ :
- sig
- val gallina : API.Vernacexpr.vernac_expr Gram.Entry.e
- val gallina_ext : API.Vernacexpr.vernac_expr Gram.Entry.e
- val red_expr : API.Genredexpr.raw_red_expr Gram.Entry.e
- val noedit_mode : API.Vernacexpr.vernac_expr Gram.Entry.e
- val command : API.Vernacexpr.vernac_expr Gram.Entry.e
- val rec_definition : (API.Vernacexpr.fixpoint_expr * API.Vernacexpr.decl_notation list) Gram.Entry.e
- val vernac : API.Vernacexpr.vernac_expr Gram.Entry.e
- end
-
- type extend_rule =
- | ExtendRule : 'a Gram.Entry.e * gram_reinit option * 'a Extend.extend_statment -> extend_rule
-
- module GramState : module type of struct include Pcoq.GramState end
- type 'a grammar_command = 'a Pcoq.grammar_command
- type 'a grammar_extension = 'a -> GramState.t -> extend_rule list * GramState.t
- val create_grammar_command : string -> 'a grammar_extension -> 'a grammar_command
- val extend_grammar_command : 'a grammar_command -> 'a -> unit
- val epsilon_value : ('a -> 'self) -> ('self, 'a) Extend.symbol -> 'self option
- val parse_string : 'a Gram.Entry.e -> string -> 'a
- val (!@) : Ploc.t -> Loc.t
- val set_command_entry : API.Vernacexpr.vernac_expr Gram.Entry.e -> unit
- val to_coqloc : Ploc.t -> Loc.t
-end
-
-module CLexer :
-sig
- type keyword_state = CLexer.keyword_state
- val terminal : string -> Tok.t
- val add_keyword : string -> unit
- val is_keyword : string -> bool
- val check_ident : string -> unit
- val get_keyword_state : unit -> keyword_state
- val set_keyword_state : keyword_state -> unit
-end
-
-module Egramml :
-sig
- type 's grammar_prod_item = 's Egramml.grammar_prod_item =
- | GramTerminal of string
- | GramNonTerminal : ('a Genarg.raw_abstract_argument_type option *
- ('s, 'a) Extend.symbol) Loc.located -> 's grammar_prod_item
-
-
- val extend_vernac_command_grammar :
- API.Vernacexpr.extend_name -> Vernacexpr.vernac_expr Pcoq.Gram.Entry.e option ->
- Vernacexpr.vernac_expr grammar_prod_item list -> unit
-
- val make_rule :
- (Loc.t -> Genarg.raw_generic_argument list -> 'a) ->
- 'a grammar_prod_item list -> 'a Extend.production_rule
-end
-
-module Mltop :
-sig
- val add_known_module : string -> unit
- val module_is_known : string -> bool
- val declare_cache_obj : (unit -> unit) -> string -> unit
-end
-module Vernacinterp :
-sig
- type deprecation = bool
- type vernac_command = Genarg.raw_generic_argument list -> unit -> unit
- val vinterp_add : deprecation -> API.Vernacexpr.extend_name ->
- vernac_command -> unit
-end
-
-module G_vernac :
-sig
- val def_body : API.Vernacexpr.definition_expr Pcoq.Gram.Entry.e
- val section_subset_expr : API.Vernacexpr.section_subset_expr Pcoq.Gram.Entry.e
- val query_command : (Vernacexpr.goal_selector option -> Vernacexpr.vernac_expr)
- Pcoq.Gram.Entry.e
-end
-
-module G_proofs :
-sig
- val hint : Vernacexpr.hints_expr Pcoq.Gram.Entry.e
- val hint_proof_using : 'a Pcoq.Gram.Entry.e -> 'a option -> 'a option
-end
-
-module Egramcoq :
-sig
-end
-
-module Metasyntax :
-sig
- type any_entry = Metasyntax.any_entry =
- | AnyEntry : 'a Pcoq.Gram.Entry.e -> any_entry
- val register_grammar : string -> any_entry list -> unit
- val add_token_obj : string -> unit
-end
diff --git a/Makefile.build b/Makefile.build
index 7703df08f..e3316df91 100644
--- a/Makefile.build
+++ b/Makefile.build
@@ -564,6 +564,12 @@ kernel/kernel.cma: kernel/kernel.mllib
$(SHOW)'OCAMLC -a -o $@'
$(HIDE)$(OCAMLC) $(MLINCLUDES) $(BYTEFLAGS) $(VMBYTEFLAGS) -a -o $@ $(filter-out %.mllib, $^)
+# Specific rule for API/API.cmi
+# Make sure that API/API.mli cannot leak types from the Coq codebase.
+API/API.cmi : API/API.mli
+ $(SHOW)'OCAMLC $<'
+ $(HIDE)$(OCAMLC) -I lib -I $(MYCAMLP4LIB) -c $<
+
%.cma: %.mllib
$(SHOW)'OCAMLC -a -o $@'
$(HIDE)$(OCAMLC) $(MLINCLUDES) $(BYTEFLAGS) -a -o $@ $(filter-out %.mllib, $^)
diff --git a/Makefile.ide b/Makefile.ide
index 0cfbdeb4e..b534b385b 100644
--- a/Makefile.ide
+++ b/Makefile.ide
@@ -61,12 +61,16 @@ GTKLIBS=$(shell pkg-config --variable=libdir gtk+-2.0)
# CoqIde special targets
###########################################################################
-.PHONY: coqide coqide-opt coqide-byte coqide-files
+.PHONY: coqide coqide-opt coqide-byte coqide-files coqide-binaries
.PHONY: ide-toploop ide-byteloop ide-optloop
-# target to build CoqIde
+# target to build CoqIde (native version) and the stuff needed to lauch it
coqide: coqide-files coqide-opt theories/Init/Prelude.vo
+# target to build CoqIde (in native and byte versions), and no more
+# NB: this target is used in the opam package coq-coqide
+coqide-binaries: coqide-opt coqide-byte
+
ifeq ($(HASCOQIDE),opt)
coqide-opt: $(COQIDE) ide-toploop
else
diff --git a/configure.ml b/configure.ml
index e13fa80fd..4eac8eacc 100644
--- a/configure.ml
+++ b/configure.ml
@@ -300,27 +300,17 @@ let args_options = Arg.align [
"<dir> Where to install doc files";
"-emacslib", arg_string_option Prefs.emacslib,
"<dir> Where to install emacs files";
- "-emacs", Arg.String (fun s ->
- prerr_endline "Warning: -emacs option is deprecated. Use -emacslib instead.";
- Prefs.emacslib := Some s),
- "<dir> Deprecated: same as -emacslib";
"-coqdocdir", arg_string_option Prefs.coqdocdir,
"<dir> Where to install Coqdoc style files";
"-ocamlfind", arg_string_option Prefs.ocamlfindcmd,
"<dir> Specifies the ocamlfind command to use";
"-lablgtkdir", arg_string_option Prefs.lablgtkdir,
"<dir> Specifies the path to the Lablgtk library";
- "-usecamlp5", Arg.Unit (fun () ->
- prerr_endline "Warning: -usecamlp5 option is deprecated. Camlp5 is already a required dependency."),
- " Deprecated: Camlp5 is a required dependency (Camlp4 is not supported anymore)";
"-camlp5dir",
Arg.String (fun s -> Prefs.camlp5dir:=Some s),
"<dir> Specifies where is the Camlp5 library and tells to use it";
"-arch", arg_string_option Prefs.arch,
"<arch> Specifies the architecture";
- "-opt", Arg.Unit (fun () ->
- prerr_endline "Warning: -opt option is deprecated. Native OCaml executables are detected automatically."),
- " Deprecated: native OCaml executables detected automatically";
"-natdynlink", arg_bool Prefs.natdynlink,
"(yes|no) Use dynamic loading of native code or not";
"-coqide", Arg.String (fun s -> Prefs.coqide := Some (get_ide s)),
@@ -329,33 +319,18 @@ let args_options = Arg.align [
" Do not try to build CoqIDE MacOS integration";
"-browser", arg_string_option Prefs.browser,
"<command> Use <command> to open URL %s";
- "-nodoc", Arg.Unit (fun () ->
- prerr_endline "Warning: -nodoc option is deprecated. Use -with-doc no instead.";
- Prefs.withdoc := false),
- " Deprecated: use -with-doc no instead";
"-with-doc", arg_bool Prefs.withdoc,
"(yes|no) Compile the documentation or not";
"-with-geoproof", arg_bool Prefs.geoproof,
"(yes|no) Use Geoproof binding or not";
"-byte-only", Arg.Set Prefs.byteonly,
" Compiles only bytecode version of Coq";
- "-byteonly", Arg.Unit (fun () ->
- prerr_endline "Warning: -byteonly option is deprecated. Use -byte-only instead.";
- Prefs.byteonly := true),
- " Deprecated: use -byte-only instead";
- "-debug", Arg.Unit (fun () ->
- prerr_endline "Warning: -debug option is deprecated. Coq is compiled in debug mode by default.";
- Prefs.debug := true),
- " Deprecated: Coq is compiled in debug mode by default";
"-nodebug", Arg.Clear Prefs.debug,
" Do not add debugging information in the Coq executables";
"-profile", Arg.Set Prefs.profile,
" Add profiling information in the Coq executables";
"-annotate", Arg.Set Prefs.annotate,
" Dumps ml annotation files while compiling Coq";
- "-makecmd", Arg.String (fun _ ->
- prerr_endline "Warning: -makecmd option is deprecated and doesn't have any effect."),
- "<command> Deprecated";
"-native-compiler", arg_bool Prefs.nativecompiler,
"(yes|no) Compilation to native code for conversion and normalization";
"-coqwebsite", Arg.Set_string Prefs.coqwebsite,
diff --git a/dev/doc/naming-conventions.tex b/dev/doc/naming-conventions.tex
index 349164948..337b9226d 100644
--- a/dev/doc/naming-conventions.tex
+++ b/dev/doc/naming-conventions.tex
@@ -267,7 +267,7 @@ If the conclusion is in the other way than listed below, add suffix
{forall x y:D, op (op' x y) = op' x (op y)}
\itemrule{Idempotency of binary operator {\op} on domain {\D}}{Dop\_idempotent}
-{forall x:D, op x n = x}
+{forall x:D, op x x = x}
\itemrule{Idempotency of unary operator {\op} on domain {\D}}{Dop\_idempotent}
{forall x:D, op (op x) = op x}
diff --git a/dev/top_printers.ml b/dev/top_printers.ml
index 9884a0109..ffa8fffdf 100644
--- a/dev/top_printers.ml
+++ b/dev/top_printers.ml
@@ -229,7 +229,7 @@ let ppenvwithcst e = pp
str "[" ++ pr_rel_context e Evd.empty (rel_context e) ++ str "]" ++ spc() ++
str "{" ++ Cmap_env.fold (fun a _ s -> pr_con a ++ spc () ++ s) (Obj.magic e).Pre_env.env_globals.Pre_env.env_constants (mt ()) ++ str "}")
-let pptac = (fun x -> pp(Ltac_plugin.Pptactic.pr_glob_tactic (Global.env()) x))
+let pptac = (fun x -> pp(Ltac_plugin.Pptactic.pr_glob_tactic (API.Global.env()) x))
let ppobj obj = Format.print_string (Libobject.object_tag obj)
@@ -494,7 +494,6 @@ VERNAC COMMAND EXTEND PrintConstr
END
*)
-open Grammar_API
open Genarg
open Stdarg
open Egramml
diff --git a/doc/stdlib/index-list.html.template b/doc/stdlib/index-list.html.template
index 48f82f2d9..48048b7a0 100644
--- a/doc/stdlib/index-list.html.template
+++ b/doc/stdlib/index-list.html.template
@@ -591,5 +591,6 @@ through the <tt>Require Import</tt> command.</p>
theories/Compat/AdmitAxiom.v
theories/Compat/Coq85.v
theories/Compat/Coq86.v
+ theories/Compat/Coq87.v
</dd>
</dl>
diff --git a/grammar/argextend.mlp b/grammar/argextend.mlp
index cc92680fc..643b6277a 100644
--- a/grammar/argextend.mlp
+++ b/grammar/argextend.mlp
@@ -46,17 +46,17 @@ let make_act loc act pil =
make (List.rev pil)
let make_prod_item = function
- | ExtTerminal s -> <:expr< Grammar_API.Extend.Atoken (Grammar_API.CLexer.terminal $mlexpr_of_string s$) >>
+ | ExtTerminal s -> <:expr< Extend.Atoken (CLexer.terminal $mlexpr_of_string s$) >>
| ExtNonTerminal (g, _) ->
let base s = <:expr< $lid:s$ >> in
mlexpr_of_prod_entry_key base g
let rec make_prod = function
-| [] -> <:expr< Grammar_API.Extend.Stop >>
-| item :: prods -> <:expr< Grammar_API.Extend.Next $make_prod prods$ $make_prod_item item$ >>
+| [] -> <:expr< Extend.Stop >>
+| item :: prods -> <:expr< Extend.Next $make_prod prods$ $make_prod_item item$ >>
let make_rule loc (prods,act) =
- <:expr< Grammar_API.Extend.Rule $make_prod (List.rev prods)$ $make_act loc act prods$ >>
+ <:expr< Extend.Rule $make_prod (List.rev prods)$ $make_act loc act prods$ >>
let is_ident x = function
| <:expr< $lid:s$ >> -> (s : string) = x
@@ -67,7 +67,7 @@ let make_extend loc s cl wit = match cl with
(** Special handling of identity arguments by not redeclaring an entry *)
<:str_item<
value $lid:s$ =
- let () = Grammar_API.Pcoq.register_grammar $wit$ $lid:e$ in
+ let () = Pcoq.register_grammar $wit$ $lid:e$ in
$lid:e$
>>
| _ ->
@@ -75,8 +75,8 @@ let make_extend loc s cl wit = match cl with
let rules = mlexpr_of_list (make_rule loc) (List.rev cl) in
<:str_item<
value $lid:s$ =
- let $lid:s$ = Grammar_API.Pcoq.create_generic_entry Grammar_API.Pcoq.utactic $se$ (Genarg.rawwit $wit$) in
- let () = Grammar_API.Pcoq.grammar_extend $lid:s$ None (None, [(None, None, $rules$)]) in
+ let $lid:s$ = Pcoq.create_generic_entry Pcoq.utactic $se$ (Genarg.rawwit $wit$) in
+ let () = Pcoq.grammar_extend $lid:s$ None (None, [(None, None, $rules$)]) in
$lid:s$ >>
let warning_redundant prefix s =
@@ -127,7 +127,7 @@ let declare_tactic_argument loc s (typ, f, g, h) cl =
begin match globtyp with
| None ->
let typ = match globtyp with None -> ExtraArgType s | Some typ -> typ in
- <:expr< fun ist v -> API.Ftactic.return (API.Geninterp.Val.inject (API.Geninterp.val_tag $make_topwit loc typ$) v) >>
+ <:expr< fun ist v -> Ftactic.return (Geninterp.Val.inject (Geninterp.val_tag $make_topwit loc typ$) v) >>
| Some globtyp ->
<:expr< fun ist x ->
Tacinterp.interp_genarg ist (Genarg.in_gen $make_globwit loc globtyp$ x) >>
@@ -137,10 +137,10 @@ let declare_tactic_argument loc s (typ, f, g, h) cl =
let typ = match globtyp with None -> ExtraArgType s | Some typ -> typ in
<:expr<
let f = $lid:f$ in
- fun ist v -> API.Ftactic.nf_enter (fun gl ->
- let (sigma, v) = API.Tacmach.New.of_old (fun gl -> f ist gl v) gl in
- let v = API.Geninterp.Val.inject (API.Geninterp.val_tag $make_topwit loc typ$) v in
- API.Proofview.tclTHEN (API.Proofview.Unsafe.tclEVARS sigma) (API.Ftactic.return v)
+ fun ist v -> Ftactic.nf_enter (fun gl ->
+ let (sigma, v) = Tacmach.New.of_old (fun gl -> f ist gl v) gl in
+ let v = Geninterp.Val.inject (Geninterp.val_tag $make_topwit loc typ$) v in
+ Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Ftactic.return v)
)
>> in
let subst = match h with
@@ -156,15 +156,15 @@ let declare_tactic_argument loc s (typ, f, g, h) cl =
| Some f -> <:expr< $lid:f$>> in
let dyn = match typ with
| None -> <:expr< None >>
- | Some typ -> <:expr< Some (API.Geninterp.val_tag $make_topwit loc typ$) >>
+ | Some typ -> <:expr< Some (Geninterp.val_tag $make_topwit loc typ$) >>
in
let wit = <:expr< $lid:"wit_"^s$ >> in
declare_str_items loc
[ <:str_item< value ($lid:"wit_"^s$) = Genarg.make0 $se$ >>;
- <:str_item< Grammar_API.Genintern.register_intern0 $wit$ $glob$ >>;
- <:str_item< Grammar_API.Genintern.register_subst0 $wit$ $subst$ >>;
- <:str_item< API.Geninterp.register_interp0 $wit$ $interp$ >>;
- <:str_item< API.Geninterp.register_val0 $wit$ $dyn$ >>;
+ <:str_item< Genintern.register_intern0 $wit$ $glob$ >>;
+ <:str_item< Genintern.register_subst0 $wit$ $subst$ >>;
+ <:str_item< Geninterp.register_interp0 $wit$ $interp$ >>;
+ <:str_item< Geninterp.register_val0 $wit$ $dyn$ >>;
make_extend loc s cl wit;
<:str_item< do {
Pptactic.declare_extra_genarg_pprule
diff --git a/grammar/q_util.mlp b/grammar/q_util.mlp
index 51a4c3573..536ee7ca5 100644
--- a/grammar/q_util.mlp
+++ b/grammar/q_util.mlp
@@ -57,23 +57,23 @@ let mlexpr_of_option f = function
| Some e -> <:expr< Some $f e$ >>
let mlexpr_of_name f = function
- | None -> <:expr< API.Names.Name.Anonymous >>
- | Some e -> <:expr< API.Names.Name.Name $f e$ >>
+ | None -> <:expr< Names.Name.Anonymous >>
+ | Some e -> <:expr< Names.Name.Name $f e$ >>
-let symbol_of_string s = <:expr< Grammar_API.Extend.Atoken (Grammar_API.CLexer.terminal $str:s$) >>
+let symbol_of_string s = <:expr< Extend.Atoken (CLexer.terminal $str:s$) >>
let rec mlexpr_of_prod_entry_key f = function
- | Ulist1 s -> <:expr< Grammar_API.Extend.Alist1 $mlexpr_of_prod_entry_key f s$ >>
- | Ulist1sep (s,sep) -> <:expr< Grammar_API.Extend.Alist1sep $mlexpr_of_prod_entry_key f s$ $symbol_of_string sep$ >>
- | Ulist0 s -> <:expr< Grammar_API.Extend.Alist0 $mlexpr_of_prod_entry_key f s$ >>
- | Ulist0sep (s,sep) -> <:expr< Grammar_API.Extend.Alist0sep $mlexpr_of_prod_entry_key f s$ $symbol_of_string sep$ >>
- | Uopt s -> <:expr< Grammar_API.Extend.Aopt $mlexpr_of_prod_entry_key f s$ >>
- | Uentry e -> <:expr< Grammar_API.Extend.Aentry ($f e$) >>
+ | Ulist1 s -> <:expr< Extend.Alist1 $mlexpr_of_prod_entry_key f s$ >>
+ | Ulist1sep (s,sep) -> <:expr< Extend.Alist1sep $mlexpr_of_prod_entry_key f s$ $symbol_of_string sep$ >>
+ | Ulist0 s -> <:expr< Extend.Alist0 $mlexpr_of_prod_entry_key f s$ >>
+ | Ulist0sep (s,sep) -> <:expr< Extend.Alist0sep $mlexpr_of_prod_entry_key f s$ $symbol_of_string sep$ >>
+ | Uopt s -> <:expr< Extend.Aopt $mlexpr_of_prod_entry_key f s$ >>
+ | Uentry e -> <:expr< Extend.Aentry ($f e$) >>
| Uentryl (e, l) ->
(** Keep in sync with Pcoq! *)
assert (e = "tactic");
- if l = 5 then <:expr< Grammar_API.Extend.Aentry Pltac.binder_tactic >>
- else <:expr< Grammar_API.Extend.Aentryl (Pltac.tactic_expr) $mlexpr_of_int l$ >>
+ if l = 5 then <:expr< Extend.Aentry Pltac.binder_tactic >>
+ else <:expr< Extend.Aentryl (Pltac.tactic_expr) $mlexpr_of_int l$ >>
let rec type_of_user_symbol = function
| Ulist1 s | Ulist1sep (s, _) | Ulist0 s | Ulist0sep (s, _) ->
diff --git a/grammar/tacextend.mlp b/grammar/tacextend.mlp
index 079d2e4e4..0b33dab05 100644
--- a/grammar/tacextend.mlp
+++ b/grammar/tacextend.mlp
@@ -25,7 +25,7 @@ let plugin_name = <:expr< __coq_plugin_name >>
let mlexpr_of_ident id =
(** Workaround for badly-designed generic arguments lacking a closure *)
let id = "$" ^ id in
- <:expr< API.Names.Id.of_string_soft $str:id$ >>
+ <:expr< Names.Id.of_string_soft $str:id$ >>
let rec make_patt = function
| [] -> <:patt< [] >>
@@ -57,18 +57,18 @@ let make_fun_clauses loc s l =
let get_argt e = <:expr< (fun e -> match e with [ Genarg.ExtraArg tag -> tag | _ -> assert False ]) $e$ >>
let rec mlexpr_of_symbol = function
-| Ulist1 s -> <:expr< Grammar_API.Extend.Ulist1 $mlexpr_of_symbol s$ >>
-| Ulist1sep (s,sep) -> <:expr< Grammar_API.Extend.Ulist1sep $mlexpr_of_symbol s$ $str:sep$ >>
-| Ulist0 s -> <:expr< Grammar_API.Extend.Ulist0 $mlexpr_of_symbol s$ >>
-| Ulist0sep (s,sep) -> <:expr< Grammar_API.Extend.Ulist0sep $mlexpr_of_symbol s$ $str:sep$ >>
-| Uopt s -> <:expr< Grammar_API.Extend.Uopt $mlexpr_of_symbol s$ >>
+| Ulist1 s -> <:expr< Extend.Ulist1 $mlexpr_of_symbol s$ >>
+| Ulist1sep (s,sep) -> <:expr< Extend.Ulist1sep $mlexpr_of_symbol s$ $str:sep$ >>
+| Ulist0 s -> <:expr< Extend.Ulist0 $mlexpr_of_symbol s$ >>
+| Ulist0sep (s,sep) -> <:expr< Extend.Ulist0sep $mlexpr_of_symbol s$ $str:sep$ >>
+| Uopt s -> <:expr< Extend.Uopt $mlexpr_of_symbol s$ >>
| Uentry e ->
let arg = get_argt <:expr< $lid:"wit_"^e$ >> in
- <:expr< Grammar_API.Extend.Uentry (Genarg.ArgT.Any $arg$) >>
+ <:expr< Extend.Uentry (Genarg.ArgT.Any $arg$) >>
| Uentryl (e, l) ->
assert (e = "tactic");
let arg = get_argt <:expr< Tacarg.wit_tactic >> in
- <:expr< Grammar_API.Extend.Uentryl (Genarg.ArgT.Any $arg$) $mlexpr_of_int l$>>
+ <:expr< Extend.Uentryl (Genarg.ArgT.Any $arg$) $mlexpr_of_int l$>>
let make_prod_item = function
| ExtTerminal s -> <:expr< Tacentries.TacTerm $str:s$ >>
@@ -113,12 +113,12 @@ let declare_tactic loc tacname ~level classification clause = match clause with
the ML tactic retrieves its arguments in the [ist] environment instead.
This is the rôle of the [lift_constr_tac_to_ml_tac] function. *)
let body = <:expr< Tacexpr.TacFun ($vars$, Tacexpr.TacML (Loc.tag ( $ml$ , []))) >> in
- let name = <:expr< API.Names.Id.of_string $name$ >> in
+ let name = <:expr< Names.Id.of_string $name$ >> in
declare_str_items loc
[ <:str_item< do {
let obj () = Tacenv.register_ltac True False $name$ $body$ in
let () = Tacenv.register_ml_tactic $se$ [|$tac$|] in
- API.Mltop.declare_cache_obj obj $plugin_name$ } >>
+ Mltop.declare_cache_obj obj $plugin_name$ } >>
]
| _ ->
(** Otherwise we add parsing and printing rules to generate a call to a
@@ -131,7 +131,7 @@ let declare_tactic loc tacname ~level classification clause = match clause with
declare_str_items loc
[ <:str_item< do {
Tacenv.register_ml_tactic $se$ (Array.of_list $make_fun_clauses loc tacname clause$);
- Grammar_API.Mltop.declare_cache_obj $obj$ $plugin_name$; } >>
+ Mltop.declare_cache_obj $obj$ $plugin_name$; } >>
]
open Pcaml
diff --git a/grammar/vernacextend.mlp b/grammar/vernacextend.mlp
index af4d0dfe8..a529185dd 100644
--- a/grammar/vernacextend.mlp
+++ b/grammar/vernacextend.mlp
@@ -100,12 +100,12 @@ let make_fun_classifiers loc s c l =
mlexpr_of_list (fun x -> x) cl
let make_prod_item = function
- | ExtTerminal s -> <:expr< Grammar_API.Egramml.GramTerminal $str:s$ >>
+ | ExtTerminal s -> <:expr< Egramml.GramTerminal $str:s$ >>
| ExtNonTerminal (g, ido) ->
let nt = type_of_user_symbol g in
- let base s = <:expr< Grammar_API.Pcoq.genarg_grammar ($mk_extraarg loc s$) >> in
+ let base s = <:expr< Pcoq.genarg_grammar ($mk_extraarg loc s$) >> in
let typ = match ido with None -> None | Some _ -> Some nt in
- <:expr< Grammar_API.Egramml.GramNonTerminal ( Loc.tag ( $mlexpr_of_option (make_rawwit loc) typ$ ,
+ <:expr< Egramml.GramNonTerminal ( Loc.tag ( $mlexpr_of_option (make_rawwit loc) typ$ ,
$mlexpr_of_prod_entry_key base g$ ) ) >>
let mlexpr_of_clause cl =
@@ -122,9 +122,9 @@ let declare_command loc s c nt cl =
let classl = make_fun_classifiers loc s c cl in
declare_str_items loc
[ <:str_item< do {
- CList.iteri (fun i (depr, f) -> Grammar_API.Vernacinterp.vinterp_add depr ($se$, i) f) $funcl$;
- CList.iteri (fun i f -> API.Vernac_classifier.declare_vernac_classifier ($se$, i) f) $classl$;
- CList.iteri (fun i r -> Grammar_API.Egramml.extend_vernac_command_grammar ($se$, i) $nt$ r) $gl$;
+ CList.iteri (fun i (depr, f) -> Vernacinterp.vinterp_add depr ($se$, i) f) $funcl$;
+ CList.iteri (fun i f -> Vernac_classifier.declare_vernac_classifier ($se$, i) f) $classl$;
+ CList.iteri (fun i r -> Egramml.extend_vernac_command_grammar ($se$, i) $nt$ r) $gl$;
} >> ]
open Pcaml
@@ -143,16 +143,16 @@ EXTEND
| "DECLARE"; "PLUGIN"; name = STRING ->
declare_str_items loc [
<:str_item< value __coq_plugin_name = $str:name$ >>;
- <:str_item< value _ = Grammar_API.Mltop.add_known_module __coq_plugin_name >>;
+ <:str_item< value _ = Mltop.add_known_module __coq_plugin_name >>;
]
] ]
;
classification:
[ [ "CLASSIFIED"; "BY"; c = LIDENT -> <:expr< $lid:c$ >>
| "CLASSIFIED"; "AS"; "SIDEFF" ->
- <:expr< fun _ -> API.Vernac_classifier.classify_as_sideeff >>
+ <:expr< fun _ -> Vernac_classifier.classify_as_sideeff >>
| "CLASSIFIED"; "AS"; "QUERY" ->
- <:expr< fun _ -> API.Vernac_classifier.classify_as_query >>
+ <:expr< fun _ -> Vernac_classifier.classify_as_query >>
] ]
;
deprecation:
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index f360fb192..c9fc3aa4f 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -664,11 +664,11 @@ let instantiate_notation_constr loc intern ntnvars subst infos c =
| NProd (Name id, NHole _, c') when option_mem_assoc id binderopt ->
let a,letins = snd (Option.get binderopt) in
let e = make_letins letins (aux subst' infos c') in
- let (loc,(na,bk,t)) = a in
+ let (_loc,(na,bk,t)) = a in
CAst.make ?loc @@ GProd (na,bk,t,e)
| NLambda (Name id,NHole _,c') when option_mem_assoc id binderopt ->
let a,letins = snd (Option.get binderopt) in
- let (loc,(na,bk,t)) = a in
+ let (_loc,(na,bk,t)) = a in
CAst.make ?loc @@ GLambda (na,bk,t,make_letins letins (aux subst' infos c'))
(* Two special cases to keep binder name synchronous with BinderType *)
| NProd (na,NHole(Evar_kinds.BinderType na',naming,arg),c')
diff --git a/kernel/entries.mli b/kernel/entries.ml
index 3fa25c142..3fa25c142 100644
--- a/kernel/entries.mli
+++ b/kernel/entries.ml
diff --git a/kernel/kernel.mllib b/kernel/kernel.mllib
index 0813315b5..994634854 100644
--- a/kernel/kernel.mllib
+++ b/kernel/kernel.mllib
@@ -13,9 +13,11 @@ Mod_subst
Cbytecodes
Copcodes
Cemitcodes
+Opaqueproof
+Declarations
+Entries
Nativevalues
Primitives
-Opaqueproof
Declareops
Retroknowledge
Conv_oracle
@@ -41,5 +43,4 @@ Nativelibrary
Safe_typing
Vm
Csymtable
-Declarations
Vconv
diff --git a/lib/flags.ml b/lib/flags.ml
index 5d9d9dcf5..0bce22f58 100644
--- a/lib/flags.ml
+++ b/lib/flags.ml
@@ -106,7 +106,7 @@ let we_are_parsing = ref false
(* Current means no particular compatibility consideration.
For correct comparisons, this constructor should remain the last one. *)
-type compat_version = VOld | V8_5 | V8_6 | Current
+type compat_version = VOld | V8_5 | V8_6 | V8_7 | Current
let compat_version = ref Current
@@ -120,6 +120,9 @@ let version_compare v1 v2 = match v1, v2 with
| V8_6, V8_6 -> 0
| V8_6, _ -> -1
| _, V8_6 -> 1
+ | V8_7, V8_7 -> 0
+ | V8_7, _ -> -1
+ | _, V8_7 -> 1
| Current, Current -> 0
let version_strictly_greater v = version_compare !compat_version v > 0
@@ -129,6 +132,7 @@ let pr_version = function
| VOld -> "old"
| V8_5 -> "8.5"
| V8_6 -> "8.6"
+ | V8_7 -> "8.7"
| Current -> "current"
(* Translate *)
diff --git a/lib/flags.mli b/lib/flags.mli
index e63f1ec26..eb4c37a54 100644
--- a/lib/flags.mli
+++ b/lib/flags.mli
@@ -77,7 +77,7 @@ val raw_print : bool ref
(* Univ print flag, never set anywere. Maybe should belong to Univ? *)
val univ_print : bool ref
-type compat_version = VOld | V8_5 | V8_6 | Current
+type compat_version = VOld | V8_5 | V8_6 | V8_7 | Current
val compat_version : compat_version ref
val version_compare : compat_version -> compat_version -> int
val version_strictly_greater : compat_version -> bool
diff --git a/lib/profile.ml b/lib/profile.ml
index b66916185..0bc226a45 100644
--- a/lib/profile.ml
+++ b/lib/profile.ml
@@ -85,6 +85,9 @@ let init_alloc = ref 0.0
let reset_profile () = List.iter reset_record !prof_table
let init_profile () =
+ (* We test Flags.profile as a way to support declaring profiled
+ functions in plugins *)
+ if !prof_table <> [] || Flags.profile then begin
let outside = create_record () in
stack := [outside];
last_alloc := get_alloc ();
@@ -92,6 +95,7 @@ let init_profile () =
init_time := get_time ();
outside.tottime <- - !init_time;
outside.owntime <- - !init_time
+ end
let ajoute n o =
o.owntime <- o.owntime + n.owntime;
@@ -317,15 +321,15 @@ let adjust_time ov_bc ov_ad e =
owntime = e.owntime - int_of_float (ad_imm +. bc_imm) }
let close_profile print =
- let dw = spent_alloc () in
- let t = get_time () in
- match !stack with
- | [outside] ->
- outside.tottime <- outside.tottime + t;
- outside.owntime <- outside.owntime + t;
- ajoute_ownalloc outside dw;
- ajoute_totalloc outside dw;
- if !prof_table <> [] then begin
+ if !prof_table <> [] then begin
+ let dw = spent_alloc () in
+ let t = get_time () in
+ match !stack with
+ | [outside] ->
+ outside.tottime <- outside.tottime + t;
+ outside.owntime <- outside.owntime + t;
+ ajoute_ownalloc outside dw;
+ ajoute_totalloc outside dw;
let ov_bc = time_overhead_B_C () (* B+C overhead *) in
let ov_ad = time_overhead_A_D () (* A+D overhead *) in
let adjust (n,e) = (n, adjust_time ov_bc ov_ad e) in
@@ -346,8 +350,8 @@ let close_profile print =
in
if print then format_profile updated_data;
init_profile ()
- end
- | _ -> failwith "Inconsistency"
+ | _ -> failwith "Inconsistency"
+ end
let print_profile () = close_profile true
@@ -358,9 +362,6 @@ let declare_profile name =
prof_table := (name,e)::!prof_table;
e
-(* Default initialization, may be overridden *)
-let _ = init_profile ()
-
(******************************)
(* Entry points for profiling *)
let profile1 e f a =
diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml
index 7d69cbff1..39826d744 100644
--- a/plugins/extraction/extract_env.ml
+++ b/plugins/extraction/extract_env.ml
@@ -175,26 +175,32 @@ let factor_fix env l cb msb =
(hack proposed by Elie)
*)
-let expand_mexpr env mp me =
+let expand_mexpr env mpo me =
let inl = Some (Flags.get_inline_level()) in
- Mod_typing.translate_mse env (Some mp) inl me
+ Mod_typing.translate_mse env mpo inl me
-(** Ad-hoc update of environment, inspired by [Mod_type.check_with_aux_def].
- To check with Elie. *)
-
-let rec mp_of_mexpr = function
- | MEident mp -> mp
- | MEwith (seb,_) -> mp_of_mexpr seb
- | _ -> assert false
+let expand_modtype env mp me =
+ let inl = Some (Flags.get_inline_level()) in
+ Mod_typing.translate_modtype env mp inl ([],me)
let no_delta = Mod_subst.empty_delta_resolver
-let env_for_mtb_with_def env mp me idl =
+let flatten_modtype env mp me_alg struc_opt =
+ match struc_opt with
+ | Some me -> me, no_delta
+ | None ->
+ let mtb = expand_modtype env mp me_alg in
+ mtb.mod_type, mtb.mod_delta
+
+(** Ad-hoc update of environment, inspired by [Mod_typing.check_with_aux_def].
+*)
+
+let env_for_mtb_with_def env mp me reso idl =
let struc = Modops.destr_nofunctor me in
let l = Label.of_id (List.hd idl) in
let spot = function (l',SFBconst _) -> Label.equal l l' | _ -> false in
let before = fst (List.split_when spot struc) in
- Modops.add_structure mp before no_delta env
+ Modops.add_structure mp before reso env
let make_cst resolver mp l =
Mod_subst.constant_of_delta_kn resolver (KerName.make2 mp l)
@@ -234,20 +240,24 @@ let rec extract_structure_spec env mp reso = function
[extract_mexpression_spec] should come from a [mod_type_alg] field.
This way, any encountered [MEident] should be a true module type. *)
-and extract_mexpr_spec env mp1 (me_struct,me_alg) = match me_alg with
+and extract_mexpr_spec env mp1 (me_struct_o,me_alg) = match me_alg with
| MEident mp -> Visit.add_mp_all mp; MTident mp
| MEwith(me',WithDef(idl,(c,ctx)))->
- let env' = env_for_mtb_with_def env (mp_of_mexpr me') me_struct idl in
- let mt = extract_mexpr_spec env mp1 (me_struct,me') in
+ let me_struct,delta = flatten_modtype env mp1 me' me_struct_o in
+ let env' = env_for_mtb_with_def env mp1 me_struct delta idl in
+ let mt = extract_mexpr_spec env mp1 (None,me') in
(match extract_with_type env' c with (* cb may contain some kn *)
| None -> mt
- | Some (vl,typ) -> MTwith(mt,ML_With_type(idl,vl,typ)))
+ | Some (vl,typ) ->
+ type_iter_references Visit.add_ref typ;
+ MTwith(mt,ML_With_type(idl,vl,typ)))
| MEwith(me',WithMod(idl,mp))->
Visit.add_mp_all mp;
- MTwith(extract_mexpr_spec env mp1 (me_struct,me'), ML_With_module(idl,mp))
+ MTwith(extract_mexpr_spec env mp1 (None,me'), ML_With_module(idl,mp))
| MEapply _ ->
(* No higher-order module type in OCaml : we use the expanded version *)
- extract_msignature_spec env mp1 no_delta (*TODO*) me_struct
+ let me_struct,delta = flatten_modtype env mp1 me_alg me_struct_o in
+ extract_msignature_spec env mp1 delta me_struct
and extract_mexpression_spec env mp1 (me_struct,me_alg) = match me_alg with
| MoreFunctor (mbid, mtb, me_alg') ->
@@ -258,8 +268,8 @@ and extract_mexpression_spec env mp1 (me_struct,me_alg) = match me_alg with
let mp = MPbound mbid in
let env' = Modops.add_module_type mp mtb env in
MTfunsig (mbid, extract_mbody_spec env mp mtb,
- extract_mexpression_spec env' mp1 (me_struct',me_alg'))
- | NoFunctor m -> extract_mexpr_spec env mp1 (me_struct,m)
+ extract_mexpression_spec env' mp1 (me_struct',me_alg'))
+ | NoFunctor m -> extract_mexpr_spec env mp1 (Some me_struct,m)
and extract_msignature_spec env mp1 reso = function
| NoFunctor struc ->
@@ -335,7 +345,7 @@ and extract_mexpr env mp = function
(* In Haskell/Scheme, we expand everything.
For now, we also extract everything, dead code will be removed later
(see [Modutil.optimize_struct]. *)
- let sign,_,delta,_ = expand_mexpr env mp me in
+ let sign,_,delta,_ = expand_mexpr env (Some mp) me in
extract_msignature env mp delta ~all:true sign
| MEident mp ->
if is_modfile mp && not (modular ()) then error_MPfile_as_mod mp false;
@@ -685,3 +695,35 @@ let structure_for_compute c =
let struc = optimize_struct (refs,[]) (mono_environment refs []) in
let flatstruc = List.map snd (List.flatten (List.map snd struc)) in
flatstruc, ast, mlt
+
+(* For the test-suite :
+ extraction to a temporary file + run ocamlc on it *)
+
+let compile f =
+ try
+ let args = ["ocamlc";"-I";Filename.dirname f;"-c";f^"i";f] in
+ let res = CUnix.sys_command (Envars.ocamlfind ()) args in
+ match res with
+ | Unix.WEXITED 0 -> ()
+ | Unix.WEXITED n | Unix.WSIGNALED n | Unix.WSTOPPED n ->
+ CErrors.user_err
+ Pp.(str "Compilation of file " ++ str f ++
+ str " failed with exit code " ++ int n)
+ with Unix.Unix_error (e,_,_) ->
+ CErrors.user_err
+ Pp.(str "Compilation of file " ++ str f ++
+ str " failed with error " ++ str (Unix.error_message e))
+
+let remove f =
+ if Sys.file_exists f then Sys.remove f
+
+let extract_and_compile l =
+ if lang () != Ocaml then
+ CErrors.user_err (Pp.str "This command only works with OCaml extraction");
+ let f = Filename.temp_file "testextraction" ".ml" in
+ let () = full_extraction (Some f) l in
+ let () = compile f in
+ let () = remove f; remove (f^"i") in
+ let base = Filename.chop_suffix f ".ml" in
+ let () = remove (base^".cmo"); remove (base^".cmi") in
+ Feedback.msg_notice (str "Extracted code successfully compiled")
diff --git a/plugins/extraction/extract_env.mli b/plugins/extraction/extract_env.mli
index f289b63ad..e10dcd48b 100644
--- a/plugins/extraction/extract_env.mli
+++ b/plugins/extraction/extract_env.mli
@@ -17,6 +17,10 @@ val full_extraction : string option -> reference list -> unit
val separate_extraction : reference list -> unit
val extraction_library : bool -> Id.t -> unit
+(* For the test-suite : extraction to a temporary file + ocamlc on it *)
+
+val extract_and_compile : reference list -> unit
+
(* For debug / external output via coqtop.byte + Drop : *)
val mono_environment :
diff --git a/plugins/extraction/g_extraction.ml4 b/plugins/extraction/g_extraction.ml4
index 4372ea557..23452febd 100644
--- a/plugins/extraction/g_extraction.ml4
+++ b/plugins/extraction/g_extraction.ml4
@@ -8,7 +8,7 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
-open Grammar_API.Pcoq.Prim
+open Pcoq.Prim
DECLARE PLUGIN "extraction_plugin"
@@ -65,6 +65,10 @@ VERNAC COMMAND EXTEND Extraction CLASSIFIED AS QUERY
(* Monolithic extraction to a file *)
| [ "Extraction" string(f) ne_global_list(l) ]
-> [ full_extraction (Some f) l ]
+
+(* Extraction to a temporary file and OCaml compilation *)
+| [ "Extraction" "TestCompile" ne_global_list(l) ]
+ -> [ extract_and_compile l ]
END
VERNAC COMMAND EXTEND SeparateExtraction CLASSIFIED AS QUERY
diff --git a/plugins/extraction/modutil.ml b/plugins/extraction/modutil.ml
index a896a8d03..1e0c33190 100644
--- a/plugins/extraction/modutil.ml
+++ b/plugins/extraction/modutil.ml
@@ -17,10 +17,15 @@ open Mlutil
(*S Functions upon ML modules. *)
+(** Note: a syntax like [(F M) with ...] is actually legal, see for instance
+ bug #4720. Hence the code below tries to handle [MTsig], maybe not in
+ a perfect way, but that should be enough for the use of [se_iter] below. *)
+
let rec msid_of_mt = function
| MTident mp -> mp
+ | MTsig(mp,_) -> mp
| MTwith(mt,_)-> msid_of_mt mt
- | _ -> anomaly ~label:"extraction" (Pp.str "the With operator isn't applied to a name.")
+ | MTfunsig _ -> assert false (* A functor cannot be inside a MTwith *)
(*s Apply some functions upon all [ml_decl] and [ml_spec] found in a
[ml_structure]. *)
@@ -36,7 +41,7 @@ let se_iter do_decl do_spec do_mp =
List.fold_left (fun mp l -> MPdot(mp,Label.of_id l)) mp_mt idl'
in
let r = ConstRef (Constant.make2 mp_w (Label.of_id l')) in
- mt_iter mt; do_decl (Dtype(r,l,t))
+ mt_iter mt; do_spec (Stype(r,l,Some t))
| MTwith (mt,ML_With_module(idl,mp))->
let mp_mt = msid_of_mt mt in
let mp_w =
diff --git a/plugins/extraction/modutil.mli b/plugins/extraction/modutil.mli
index ad60b58d5..17a6e8db6 100644
--- a/plugins/extraction/modutil.mli
+++ b/plugins/extraction/modutil.mli
@@ -17,6 +17,7 @@ val struct_type_search : (ml_type -> bool) -> ml_structure -> bool
type do_ref = global_reference -> unit
+val type_iter_references : do_ref -> ml_type -> unit
val ast_iter_references : do_ref -> do_ref -> do_ref -> ml_ast -> unit
val decl_iter_references : do_ref -> do_ref -> do_ref -> ml_decl -> unit
val spec_iter_references : do_ref -> do_ref -> do_ref -> ml_spec -> unit
diff --git a/plugins/firstorder/g_ground.ml4 b/plugins/firstorder/g_ground.ml4
index c001ee382..1e7da3250 100644
--- a/plugins/firstorder/g_ground.ml4
+++ b/plugins/firstorder/g_ground.ml4
@@ -8,7 +8,6 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
-open Grammar_API
open Ltac_plugin
open Formula
open Sequent
diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4
index c495703ee..16d9f200f 100644
--- a/plugins/funind/g_indfun.ml4
+++ b/plugins/funind/g_indfun.ml4
@@ -6,7 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
(*i camlp4deps: "grammar/grammar.cma" i*)
-open Grammar_API
open Ltac_plugin
open Util
open Pp
diff --git a/plugins/ltac/extraargs.ml4 b/plugins/ltac/extraargs.ml4
index 72c6f9090..609795133 100644
--- a/plugins/ltac/extraargs.ml4
+++ b/plugins/ltac/extraargs.ml4
@@ -8,7 +8,6 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
-open Grammar_API
open Pp
open Genarg
open Stdarg
diff --git a/plugins/ltac/extraargs.mli b/plugins/ltac/extraargs.mli
index 419c5e8c4..acdf67779 100644
--- a/plugins/ltac/extraargs.mli
+++ b/plugins/ltac/extraargs.mli
@@ -6,7 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Grammar_API
open Tacexpr
open Names
open Constrexpr
diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4
index 6d80ab549..f3f2f27e9 100644
--- a/plugins/ltac/extratactics.ml4
+++ b/plugins/ltac/extratactics.ml4
@@ -8,7 +8,6 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
-open Grammar_API
open Pp
open Genarg
open Stdarg
diff --git a/plugins/ltac/g_auto.ml4 b/plugins/ltac/g_auto.ml4
index 4d13d89a4..301943a50 100644
--- a/plugins/ltac/g_auto.ml4
+++ b/plugins/ltac/g_auto.ml4
@@ -8,7 +8,6 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
-open Grammar_API
open Pp
open Genarg
open Stdarg
diff --git a/plugins/ltac/g_ltac.ml4 b/plugins/ltac/g_ltac.ml4
index cc052c8a2..2ea0f60eb 100644
--- a/plugins/ltac/g_ltac.ml4
+++ b/plugins/ltac/g_ltac.ml4
@@ -8,8 +8,6 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
-open Grammar_API
-
DECLARE PLUGIN "ltac_plugin"
open Util
diff --git a/plugins/ltac/g_obligations.ml4 b/plugins/ltac/g_obligations.ml4
index 1935d560a..1a2d89586 100644
--- a/plugins/ltac/g_obligations.ml4
+++ b/plugins/ltac/g_obligations.ml4
@@ -12,7 +12,6 @@
Syntax for the subtac terms and types.
Elaborated from correctness/psyntax.ml4 by Jean-Christophe Filliâtre *)
-open Grammar_API
open Libnames
open Constrexpr
open Constrexpr_ops
diff --git a/plugins/ltac/g_rewrite.ml4 b/plugins/ltac/g_rewrite.ml4
index 3c27b2747..c874f8d5a 100644
--- a/plugins/ltac/g_rewrite.ml4
+++ b/plugins/ltac/g_rewrite.ml4
@@ -10,7 +10,6 @@
(* Syntax for rewriting with strategies *)
-open Grammar_API
open Names
open Misctypes
open Locus
diff --git a/plugins/ltac/g_tactic.ml4 b/plugins/ltac/g_tactic.ml4
index e539b5867..d792d4ff7 100644
--- a/plugins/ltac/g_tactic.ml4
+++ b/plugins/ltac/g_tactic.ml4
@@ -6,7 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Grammar_API
open Pp
open CErrors
open Util
diff --git a/plugins/ltac/pltac.ml b/plugins/ltac/pltac.ml
index 2adcf02e6..2c1b1067e 100644
--- a/plugins/ltac/pltac.ml
+++ b/plugins/ltac/pltac.ml
@@ -6,7 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Grammar_API
open Pcoq
(* Main entry for extensions *)
diff --git a/plugins/ltac/pltac.mli b/plugins/ltac/pltac.mli
index 794cb527f..048dcc8e9 100644
--- a/plugins/ltac/pltac.mli
+++ b/plugins/ltac/pltac.mli
@@ -8,7 +8,6 @@
(** Ltac parsing entries *)
-open Grammar_API
open Loc
open Names
open Pcoq
diff --git a/plugins/ltac/tacentries.ml b/plugins/ltac/tacentries.ml
index 791b7f48d..cf676f598 100644
--- a/plugins/ltac/tacentries.ml
+++ b/plugins/ltac/tacentries.ml
@@ -6,7 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Grammar_API
open Pp
open CErrors
open Util
diff --git a/plugins/ltac/tacentries.mli b/plugins/ltac/tacentries.mli
index ccd44b914..aa8f4efe6 100644
--- a/plugins/ltac/tacentries.mli
+++ b/plugins/ltac/tacentries.mli
@@ -8,7 +8,6 @@
(** Ltac toplevel command entries. *)
-open Grammar_API
open Vernacexpr
open Tacexpr
diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml
index df03c7b47..0554d4364 100644
--- a/plugins/ltac/tacintern.ml
+++ b/plugins/ltac/tacintern.ml
@@ -6,7 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Grammar_API
open Pattern
open Pp
open Genredexpr
diff --git a/plugins/ltac/tacintern.mli b/plugins/ltac/tacintern.mli
index ad2e70908..b262473a9 100644
--- a/plugins/ltac/tacintern.mli
+++ b/plugins/ltac/tacintern.mli
@@ -6,7 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Grammar_API
open Pp
open Names
open Tacexpr
diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml
index 7b054947b..60a8f75ec 100644
--- a/plugins/ltac/tacinterp.ml
+++ b/plugins/ltac/tacinterp.ml
@@ -6,7 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Grammar_API
open Constrintern
open Patternops
open Pp
diff --git a/plugins/ltac/tacsubst.ml b/plugins/ltac/tacsubst.ml
index c1ca85433..180fb2db4 100644
--- a/plugins/ltac/tacsubst.ml
+++ b/plugins/ltac/tacsubst.ml
@@ -6,7 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Grammar_API
open Util
open Tacexpr
open Mod_subst
diff --git a/plugins/setoid_ring/g_newring.ml4 b/plugins/setoid_ring/g_newring.ml4
index 6c82346bc..05ab8ab32 100644
--- a/plugins/setoid_ring/g_newring.ml4
+++ b/plugins/setoid_ring/g_newring.ml4
@@ -8,7 +8,6 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
-open Grammar_API
open Ltac_plugin
open Pp
open Util
diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml
index 608b778e4..799e969ae 100644
--- a/plugins/ssr/ssrcommon.ml
+++ b/plugins/ssr/ssrcommon.ml
@@ -8,7 +8,6 @@
(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
-open Grammar_API
open Util
open Names
open Evd
diff --git a/plugins/ssr/ssrparser.ml4 b/plugins/ssr/ssrparser.ml4
index 228444b82..ce23bb2b3 100644
--- a/plugins/ssr/ssrparser.ml4
+++ b/plugins/ssr/ssrparser.ml4
@@ -8,7 +8,6 @@
(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
-open Grammar_API
open Names
open Pp
open Pcoq
diff --git a/plugins/ssr/ssrparser.mli b/plugins/ssr/ssrparser.mli
index c93e10405..bf6f44f11 100644
--- a/plugins/ssr/ssrparser.mli
+++ b/plugins/ssr/ssrparser.mli
@@ -8,8 +8,6 @@
(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
-open Grammar_API
-
val ssrtacarg : Tacexpr.raw_tactic_expr Pcoq.Gram.entry
val wit_ssrtacarg : (Tacexpr.raw_tactic_expr, Tacexpr.glob_tactic_expr, Geninterp.Val.t) Genarg.genarg_type
val pr_ssrtacarg : 'a -> 'b -> (int * Ppextend.parenRelation -> 'c) -> 'c
diff --git a/plugins/ssr/ssrvernac.ml4 b/plugins/ssr/ssrvernac.ml4
index fbe3cd2b9..9c59d83d4 100644
--- a/plugins/ssr/ssrvernac.ml4
+++ b/plugins/ssr/ssrvernac.ml4
@@ -8,7 +8,6 @@
(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
-open Grammar_API
open Names
open Term
open Termops
diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4
index 74519f6c5..f6300ab7e 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml4
@@ -8,8 +8,6 @@
(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
-open Grammar_API
-
(* Defining grammar rules with "xx" in it automatically declares keywords too,
* we thus save the lexer to restore it at the end of the file *)
let frozen_lexer = CLexer.get_keyword_state () ;;
diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli
index 0c09d7bfb..65ea76d16 100644
--- a/plugins/ssrmatching/ssrmatching.mli
+++ b/plugins/ssrmatching/ssrmatching.mli
@@ -1,7 +1,6 @@
(* (c) Copyright 2006-2015 Microsoft Corporation and Inria. *)
(* Distributed under the terms of CeCILL-B. *)
-open Grammar_API
open Goal
open Genarg
open Tacexpr
diff --git a/proofs/proof_type.mli b/proofs/proof_type.ml
index 11f1a13e6..11f1a13e6 100644
--- a/proofs/proof_type.mli
+++ b/proofs/proof_type.ml
diff --git a/proofs/proofs.mllib b/proofs/proofs.mllib
index 0ea2bd66b..eaf0c693e 100644
--- a/proofs/proofs.mllib
+++ b/proofs/proofs.mllib
@@ -2,6 +2,7 @@ Miscprint
Goal
Evar_refiner
Proof_using
+Proof_type
Logic
Refine
Proof
diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml
index 6cc9d3d55..3fc2fc31b 100644
--- a/tactics/class_tactics.ml
+++ b/tactics/class_tactics.ml
@@ -1143,7 +1143,7 @@ module Search = struct
let res =
if j = 0 then tclUNIT ()
else tclDISPATCH
- (List.init j (fun j' -> (tac_of gls i (Option.default 0 k + j))))
+ (List.init j (fun j' -> (tac_of gls i (Option.default 0 k + j'))))
in
let finish nestedshelf sigma =
let filter ev =
diff --git a/test-suite/bugs/closed/4720.v b/test-suite/bugs/closed/4720.v
new file mode 100644
index 000000000..9265b60c1
--- /dev/null
+++ b/test-suite/bugs/closed/4720.v
@@ -0,0 +1,46 @@
+(** Bug 4720 : extraction and "with" in module type *)
+
+Module Type A.
+ Parameter t : Set.
+End A.
+
+Module A_instance <: A.
+ Definition t := nat.
+End A_instance.
+
+Module A_private : A.
+ Definition t := nat.
+End A_private.
+
+Module Type B.
+End B.
+
+Module Type C (b : B).
+ Declare Module a : A.
+End C.
+
+Module WithMod (a' : A) (b' : B) (c' : C b' with Module a := A_instance).
+End WithMod.
+
+Module WithDef (a' : A) (b' : B) (c' : C b' with Definition a.t := nat).
+End WithDef.
+
+Module WithModPriv (a' : A) (b' : B) (c' : C b' with Module a := A_private).
+End WithModPriv.
+
+(* The initial bug report was concerning the extraction of WithModPriv
+ in Coq 8.4, which was suboptimal: it was compiling, but could have been
+ turned into some faulty code since A_private and c'.a were not seen as
+ identical by the extraction.
+
+ In Coq 8.5 and 8.6, the extractions of WithMod, WithDef, WithModPriv
+ were all causing Anomaly or Assert Failure. This shoud be fixed now.
+*)
+
+Require Extraction.
+
+Recursive Extraction WithMod.
+
+Recursive Extraction WithDef.
+
+Recursive Extraction WithModPriv.
diff --git a/test-suite/bugs/closed/5177.v b/test-suite/bugs/closed/5177.v
new file mode 100644
index 000000000..231d103a5
--- /dev/null
+++ b/test-suite/bugs/closed/5177.v
@@ -0,0 +1,21 @@
+(* Bug 5177 https://coq.inria.fr/bug/5177 :
+ Extraction and module type containing application and "with" *)
+
+Module Type T.
+ Parameter t: Type.
+End T.
+
+Module Type A (MT: T).
+ Parameter t1: Type.
+ Parameter t2: Type.
+ Parameter bar: MT.t -> t1 -> t2.
+End A.
+
+Module MakeA(MT: T): A MT with Definition t1 := nat.
+ Definition t1 := nat.
+ Definition t2 := nat.
+ Definition bar (m: MT.t) (x:t1) := x.
+End MakeA.
+
+Require Extraction.
+Recursive Extraction MakeA.
diff --git a/theories/Compat/Coq86.v b/theories/Compat/Coq86.v
index f46460886..34061ddd6 100644
--- a/theories/Compat/Coq86.v
+++ b/theories/Compat/Coq86.v
@@ -7,5 +7,7 @@
(************************************************************************)
(** Compatibility file for making Coq act similar to Coq v8.6 *)
+Require Export Coq.Compat.Coq87.
+
Require Export Coq.extraction.Extraction.
Require Export Coq.funind.FunInd.
diff --git a/theories/Compat/Coq87.v b/theories/Compat/Coq87.v
new file mode 100644
index 000000000..61e911678
--- /dev/null
+++ b/theories/Compat/Coq87.v
@@ -0,0 +1,9 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** Compatibility file for making Coq act similar to Coq v8.7 *)
diff --git a/theories/FSets/FMapPositive.v b/theories/FSets/FMapPositive.v
index b1c0fdaa2..3e7664929 100644
--- a/theories/FSets/FMapPositive.v
+++ b/theories/FSets/FMapPositive.v
@@ -654,19 +654,19 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
Lemma xelements_bits_lt_1 : forall p p0 q m v,
List.In (p0,v) (xelements m (append p (xO q))) -> E.bits_lt p0 p.
- Proof.
+ Proof using.
intros.
generalize (xelements_complete _ _ _ _ H); clear H; intros.
- revert p0 q m v H.
+ revert p0 H.
induction p; destruct p0; simpl; intros; eauto; try discriminate.
Qed.
Lemma xelements_bits_lt_2 : forall p p0 q m v,
List.In (p0,v) (xelements m (append p (xI q))) -> E.bits_lt p p0.
- Proof.
+ Proof using.
intros.
generalize (xelements_complete _ _ _ _ H); clear H; intros.
- revert p0 q m v H.
+ revert p0 H.
induction p; destruct p0; simpl; intros; eauto; try discriminate.
Qed.
diff --git a/theories/Numbers/Integer/Abstract/ZDivEucl.v b/theories/Numbers/Integer/Abstract/ZDivEucl.v
index ce49877e4..967b68d36 100644
--- a/theories/Numbers/Integer/Abstract/ZDivEucl.v
+++ b/theories/Numbers/Integer/Abstract/ZDivEucl.v
@@ -602,6 +602,14 @@ Proof.
apply div_mod; order.
Qed.
+Lemma mod_div: forall a b, b~=0 ->
+ a mod b / b == 0.
+Proof.
+ intros a b Hb.
+ rewrite div_small_iff by assumption.
+ auto using mod_always_pos.
+Qed.
+
(** A last inequality: *)
Theorem div_mul_le:
diff --git a/theories/Numbers/Integer/Abstract/ZDivFloor.v b/theories/Numbers/Integer/Abstract/ZDivFloor.v
index ca6197002..a9077127e 100644
--- a/theories/Numbers/Integer/Abstract/ZDivFloor.v
+++ b/theories/Numbers/Integer/Abstract/ZDivFloor.v
@@ -650,6 +650,14 @@ Proof.
apply div_mod; order.
Qed.
+Lemma mod_div: forall a b, b~=0 ->
+ a mod b / b == 0.
+Proof.
+ intros a b Hb.
+ rewrite div_small_iff by assumption.
+ auto using mod_bound_or.
+Qed.
+
(** A last inequality: *)
Theorem div_mul_le:
diff --git a/theories/Numbers/Integer/Abstract/ZDivTrunc.v b/theories/Numbers/Integer/Abstract/ZDivTrunc.v
index 46c36015d..bbb8ad5ae 100644
--- a/theories/Numbers/Integer/Abstract/ZDivTrunc.v
+++ b/theories/Numbers/Integer/Abstract/ZDivTrunc.v
@@ -622,6 +622,14 @@ Proof.
apply quot_rem; order.
Qed.
+Lemma rem_quot: forall a b, b~=0 ->
+ a rem b ÷ b == 0.
+Proof.
+ intros a b Hb.
+ rewrite quot_small_iff by assumption.
+ auto using rem_bound_abs.
+Qed.
+
(** A last inequality: *)
Theorem quot_mul_le:
diff --git a/theories/ZArith/Zdiv.v b/theories/ZArith/Zdiv.v
index 55fc90f21..fa1ddf56f 100644
--- a/theories/ZArith/Zdiv.v
+++ b/theories/ZArith/Zdiv.v
@@ -508,6 +508,13 @@ Qed.
(** Unfortunately, the previous result isn't always true on negative numbers.
For instance: 3/(-2)/(-2) = 1 <> 0 = 3 / (-2*-2) *)
+Lemma Zmod_div : forall a b, a mod b / b = 0.
+Proof.
+ intros a b.
+ zero_or_not b.
+ auto using Z.mod_div.
+Qed.
+
(** A last inequality: *)
Theorem Zdiv_mult_le:
diff --git a/toplevel/coqinit.ml b/toplevel/coqinit.ml
index f61416dde..326ef5471 100644
--- a/toplevel/coqinit.ml
+++ b/toplevel/coqinit.ml
@@ -127,7 +127,8 @@ let init_ocaml_path () =
List.iter add_subdir Coq_config.all_src_dirs
let get_compat_version ?(allow_old = true) = function
- | "8.7" -> Flags.Current
+ | "8.8" -> Flags.Current
+ | "8.7" -> Flags.V8_7
| "8.6" -> Flags.V8_6
| "8.5" -> Flags.V8_5
| ("8.4" | "8.3" | "8.2" | "8.1" | "8.0") as s ->
diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml
index 08c235023..515552fe7 100644
--- a/toplevel/coqtop.ml
+++ b/toplevel/coqtop.ml
@@ -209,6 +209,7 @@ let add_compat_require v =
match v with
| Flags.V8_5 -> add_require "Coq.Compat.Coq85"
| Flags.V8_6 -> add_require "Coq.Compat.Coq86"
+ | Flags.V8_7 -> add_require "Coq.Compat.Coq87"
| Flags.VOld | Flags.Current -> ()
let compile_list = ref ([] : (bool * string) list)
@@ -613,6 +614,7 @@ let parse_args arglist =
with any -> fatal_error any
let init_toplevel arglist =
+ Profile.init_profile ();
init_gc ();
Sys.catch_break false; (* Ctrl-C is fatal during the initialisation *)
let init_feeder = Feedback.add_feeder coqtop_init_feed in