aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--API/API.ml4
-rw-r--r--API/API.mli380
-rw-r--r--CONTRIBUTING.md4
-rw-r--r--dev/ci/user-overlays/06197-ejgallego-plugins+remove_locality_hack.sh4
-rw-r--r--dev/top_printers.ml6
-rw-r--r--engine/eConstr.ml13
-rw-r--r--engine/engine.mllib7
-rw-r--r--engine/evarutil.ml4
-rw-r--r--engine/evarutil.mli5
-rw-r--r--engine/evd.ml6
-rw-r--r--engine/evd.mli13
-rw-r--r--engine/nameops.ml (renamed from library/nameops.ml)15
-rw-r--r--engine/nameops.mli (renamed from library/nameops.mli)18
-rw-r--r--engine/proofview.ml2
-rw-r--r--engine/uState.ml3
-rw-r--r--engine/uState.mli5
-rw-r--r--engine/universes.mli1
-rw-r--r--engine/univops.ml (renamed from library/univops.ml)4
-rw-r--r--engine/univops.mli (renamed from library/univops.mli)0
-rw-r--r--grammar/vernacextend.mlp27
-rw-r--r--interp/impargs.ml1
-rw-r--r--interp/reserve.ml2
-rw-r--r--kernel/constr.ml163
-rw-r--r--kernel/constr.mli106
-rw-r--r--kernel/indtypes.ml16
-rw-r--r--kernel/inductive.ml26
-rw-r--r--kernel/modops.ml4
-rw-r--r--kernel/names.ml2
-rw-r--r--kernel/names.mli1
-rw-r--r--kernel/term.ml204
-rw-r--r--kernel/term.mli54
-rw-r--r--library/coqlib.ml10
-rw-r--r--library/declaremods.ml4
-rw-r--r--library/heads.ml1
-rw-r--r--library/lib.ml5
-rw-r--r--library/libnames.ml10
-rw-r--r--library/libnames.mli20
-rw-r--r--library/library.ml21
-rw-r--r--library/library.mllib2
-rw-r--r--library/loadpath.ml6
-rw-r--r--parsing/egramcoq.ml8
-rw-r--r--plugins/cc/ccalgo.ml16
-rw-r--r--plugins/cc/cctac.ml1
-rw-r--r--plugins/extraction/table.ml2
-rw-r--r--plugins/firstorder/g_ground.ml411
-rw-r--r--plugins/firstorder/ground.ml2
-rw-r--r--plugins/firstorder/rules.ml4
-rw-r--r--plugins/funind/functional_principles_proofs.ml2
-rw-r--r--plugins/funind/functional_principles_types.ml3
-rw-r--r--plugins/funind/glob_term_to_relation.ml1
-rw-r--r--plugins/funind/indfun.ml3
-rw-r--r--plugins/funind/recdef.ml1
-rw-r--r--plugins/ltac/extratactics.ml438
-rw-r--r--plugins/ltac/g_auto.ml415
-rw-r--r--plugins/ltac/g_ltac.ml418
-rw-r--r--plugins/ltac/g_obligations.ml412
-rw-r--r--plugins/ltac/g_rewrite.ml427
-rw-r--r--plugins/ltac/rewrite.ml4
-rw-r--r--plugins/ltac/rewrite.mli2
-rw-r--r--plugins/quote/quote.ml2
-rw-r--r--plugins/romega/const_omega.ml1
-rw-r--r--plugins/ssr/ssrequality.ml2
-rw-r--r--plugins/ssr/ssrfwd.ml7
-rw-r--r--plugins/ssr/ssrvernac.ml411
-rw-r--r--pretyping/cases.mli4
-rw-r--r--pretyping/evarconv.ml3
-rw-r--r--pretyping/evardefine.ml2
-rw-r--r--pretyping/evarsolve.ml2
-rw-r--r--pretyping/geninterp.ml (renamed from engine/geninterp.ml)0
-rw-r--r--pretyping/geninterp.mli (renamed from engine/geninterp.mli)0
-rw-r--r--pretyping/pretyping.mli2
-rw-r--r--pretyping/pretyping.mllib1
-rw-r--r--pretyping/recordops.ml2
-rw-r--r--printing/ppvernac.ml6
-rw-r--r--printing/prettyp.ml9
-rw-r--r--printing/printer.ml1
-rw-r--r--proofs/logic.ml1
-rw-r--r--tactics/autorewrite.ml1
-rw-r--r--tactics/hipattern.ml2
-rw-r--r--test-suite/bugs/closed/5215.v286
-rw-r--r--test-suite/bugs/closed/5215_2.v8
-rw-r--r--test-suite/success/Notations2.v9
-rw-r--r--toplevel/coqinit.ml6
-rw-r--r--vernac/command.ml6
-rw-r--r--vernac/himsg.ml4
-rw-r--r--vernac/locality.ml20
-rw-r--r--vernac/locality.mli8
-rw-r--r--vernac/obligations.ml4
-rw-r--r--vernac/record.ml3
-rw-r--r--vernac/vernacentries.ml20
-rw-r--r--vernac/vernacinterp.ml18
-rw-r--r--vernac/vernacinterp.mli13
-rw-r--r--vernac/vernacstate.ml2
-rw-r--r--vernac/vernacstate.mli2
94 files changed, 1269 insertions, 548 deletions
diff --git a/API/API.ml b/API/API.ml
index 78d9c0c26..378c03ee4 100644
--- a/API/API.ml
+++ b/API/API.ml
@@ -20,10 +20,6 @@
(******************************************************************************)
module Coq_config = Coq_config
-(* Reexporting deprecated symbols throu module aliases triggers a
- warning in 4.06.0 *)
-[@@@ocaml.warning "-3"]
-
(******************************************************************************)
(* Kernel *)
(******************************************************************************)
diff --git a/API/API.mli b/API/API.mli
index a043a802a..1f1b05ead 100644
--- a/API/API.mli
+++ b/API/API.mli
@@ -20,10 +20,6 @@
See below in the file for their concrete position.
*)
-(* Reexporting deprecated symbols throu module aliases triggers a
- warning in 4.06.0 *)
-[@@@ocaml.warning "-3"]
-
(************************************************************************)
(* Modules from config/ *)
(************************************************************************)
@@ -87,6 +83,7 @@ sig
val repr : t -> Id.t list
val equal : t -> t -> bool
val to_string : t -> string
+ val print : t -> Pp.t
end
module MBId : sig
@@ -327,7 +324,7 @@ sig
type identifier = Id.t
[@@ocaml.deprecated "Alias of Names"]
- module Idset : Set.S with type elt = identifier and type t = Id.Set.t
+ module Idset : Set.S with type elt = Id.t and type t = Id.Set.t
[@@ocaml.deprecated "Alias of Id.Set.t"]
end
@@ -347,7 +344,7 @@ sig
module LSet :
sig
- include CSig.SetS with type elt = universe_level
+ include CSig.SetS with type elt = Level.t
val pr : (Level.t -> Pp.t) -> t -> Pp.t
end
@@ -375,7 +372,7 @@ sig
type constraint_type = Lt | Le | Eq
- type univ_constraint = universe_level * constraint_type * universe_level
+ type univ_constraint = Level.t * constraint_type * Level.t
module Constraint : sig
include Set.S with type elt = univ_constraint
@@ -437,7 +434,7 @@ sig
module LMap :
sig
- include CMap.ExtS with type key = universe_level and module Set := LSet
+ include CMap.ExtS with type key = Level.t and module Set := LSet
val union : 'a t -> 'a t -> 'a t
val diff : 'a t -> 'a t -> 'a t
@@ -446,8 +443,8 @@ sig
end
type 'a universe_map = 'a LMap.t
- type universe_subst = universe universe_map
- type universe_level_subst = universe_level universe_map
+ type universe_subst = Universe.t universe_map
+ type universe_level_subst = Level.t universe_map
val enforce_leq : Universe.t constraint_function
val pr_uni : Universe.t -> Pp.t
@@ -481,6 +478,7 @@ sig
type family = InProp | InSet | InType
val family : t -> family
+ val univ_of_sort : t -> Univ.Universe.t
end
module Evar :
@@ -501,6 +499,7 @@ end
module Constr :
sig
+
open Names
type t
@@ -578,13 +577,13 @@ sig
val kind : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term
val of_kind : (constr, types, Sorts.t, Univ.Instance.t) kind_of_term -> constr
-val map_with_binders :
- ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr
-val map : (constr -> constr) -> constr -> constr
+ val map_with_binders :
+ ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr
+ val map : (constr -> constr) -> constr -> constr
-val fold : ('a -> constr -> 'a) -> 'a -> constr -> 'a
-val iter : (constr -> unit) -> constr -> unit
-val compare_head : (constr -> constr -> bool) -> constr -> constr -> bool
+ val fold : ('a -> constr -> 'a) -> 'a -> constr -> 'a
+ val iter : (constr -> unit) -> constr -> unit
+ val compare_head : (constr -> constr -> bool) -> constr -> constr -> bool
val equal : t -> t -> bool
val eq_constr_nounivs : t -> t -> bool
@@ -626,6 +625,109 @@ val compare_head : (constr -> constr -> bool) -> constr -> constr -> bool
val mkCase : case_info * t * t * t array -> t
+ (** {6 Simple case analysis} *)
+ val isRel : constr -> bool
+ val isRelN : int -> constr -> bool
+ val isVar : constr -> bool
+ val isVarId : Id.t -> constr -> bool
+ val isInd : constr -> bool
+ val isEvar : constr -> bool
+ val isMeta : constr -> bool
+ val isEvar_or_Meta : constr -> bool
+ val isSort : constr -> bool
+ val isCast : constr -> bool
+ val isApp : constr -> bool
+ val isLambda : constr -> bool
+ val isLetIn : constr -> bool
+ val isProd : constr -> bool
+ val isConst : constr -> bool
+ val isConstruct : constr -> bool
+ val isFix : constr -> bool
+ val isCoFix : constr -> bool
+ val isCase : constr -> bool
+ val isProj : constr -> bool
+
+ val is_Prop : constr -> bool
+ val is_Set : constr -> bool
+ val isprop : constr -> bool
+ val is_Type : constr -> bool
+ val iskind : constr -> bool
+ val is_small : Sorts.t -> bool
+
+ (** {6 Term destructors } *)
+ (** Destructor operations are partial functions and
+ @raise DestKO if the term has not the expected form. *)
+
+ exception DestKO
+
+ (** Destructs a de Bruijn index *)
+ val destRel : constr -> int
+
+ (** Destructs an existential variable *)
+ val destMeta : constr -> metavariable
+
+ (** Destructs a variable *)
+ val destVar : constr -> Id.t
+
+ (** Destructs a sort. [is_Prop] recognizes the sort {% \textsf{%}Prop{% }%}, whether
+ [isprop] recognizes both {% \textsf{%}Prop{% }%} and {% \textsf{%}Set{% }%}. *)
+ val destSort : constr -> Sorts.t
+
+ (** Destructs a casted term *)
+ val destCast : constr -> constr * cast_kind * constr
+
+ (** Destructs the product {% $ %}(x:t_1)t_2{% $ %} *)
+ val destProd : types -> Name.t * types * types
+
+ (** Destructs the abstraction {% $ %}[x:t_1]t_2{% $ %} *)
+ val destLambda : constr -> Name.t * types * constr
+
+ (** Destructs the let {% $ %}[x:=b:t_1]t_2{% $ %} *)
+ val destLetIn : constr -> Name.t * constr * types * constr
+
+ (** Destructs an application *)
+ val destApp : constr -> constr * constr array
+
+ (** Decompose any term as an applicative term; the list of args can be empty *)
+ val decompose_app : constr -> constr * constr list
+
+ (** Same as [decompose_app], but returns an array. *)
+ val decompose_appvect : constr -> constr * constr array
+
+ (** Destructs a constant *)
+ val destConst : constr -> Constant.t puniverses
+
+ (** Destructs an existential variable *)
+ val destEvar : constr -> existential
+
+ (** Destructs a (co)inductive type *)
+ val destInd : constr -> inductive puniverses
+
+ (** Destructs a constructor *)
+ val destConstruct : constr -> constructor puniverses
+
+ (** Destructs a [match c as x in I args return P with ... |
+ Ci(...yij...) => ti | ... end] (or [let (..y1i..) := c as x in I args
+ return P in t1], or [if c then t1 else t2])
+ @return [(info,c,fun args x => P,[|...|fun yij => ti| ...|])]
+ where [info] is pretty-printing information *)
+ val destCase : constr -> case_info * constr * constr * constr array
+
+ (** Destructs a projection *)
+ val destProj : constr -> Projection.t * constr
+
+ (** Destructs the {% $ %}i{% $ %}th function of the block
+ [Fixpoint f{_ 1} ctx{_ 1} = b{_ 1}
+ with f{_ 2} ctx{_ 2} = b{_ 2}
+ ...
+ with f{_ n} ctx{_ n} = b{_ n}],
+ where the length of the {% $ %}j{% $ %}th context is {% $ %}ij{% $ %}.
+ *)
+ val destFix : constr -> fixpoint
+
+ type cofixpoint = int * rec_declaration
+ val destCoFix : constr -> cofixpoint
+
end
module Context :
@@ -856,6 +958,7 @@ end
module Term :
sig
+ open Constr
type sorts_family = Sorts.family = InProp | InSet | InType
[@@ocaml.deprecated "Alias of Sorts.family"]
@@ -863,15 +966,10 @@ sig
[@@ocaml.deprecated "Alias of Sorts.contents"]
type sorts = Sorts.t =
- | Prop of contents
+ | Prop of Sorts.contents
| Type of Univ.Universe.t
[@@ocaml.deprecated "alias of API.Sorts.t"]
- type constr = Constr.t
- [@@ocaml.deprecated "Alias of Constr.t"]
- type types = Constr.t
- [@@ocaml.deprecated "Alias of Constr.types"]
-
type metavariable = int
[@@ocaml.deprecated "Alias of Constr.metavariable"]
@@ -890,11 +988,11 @@ sig
type 'a puniverses = 'a Univ.puniverses
[@@ocaml.deprecated "Alias of Constr.puniverses"]
- type pconstant = Names.Constant.t puniverses
+ type pconstant = Names.Constant.t Constr.puniverses
[@@ocaml.deprecated "Alias of Constr.pconstant"]
- type pinductive = Names.inductive puniverses
+ type pinductive = Names.inductive Constr.puniverses
[@@ocaml.deprecated "Alias of Constr.pinductive"]
- type pconstructor = Names.constructor puniverses
+ type pconstructor = Names.constructor Constr.puniverses
[@@ocaml.deprecated "Alias of Constr.pconstructor"]
type case_style = Constr.case_style =
| LetStyle
@@ -907,7 +1005,7 @@ sig
type case_printing = Constr.case_printing =
{ ind_tags : bool list;
cstr_tags : bool list array;
- style : case_style
+ style : Constr.case_style
}
[@@ocaml.deprecated "Alias of Constr.case_printing"]
@@ -916,25 +1014,25 @@ sig
ci_npar : int;
ci_cstr_ndecls: int array;
ci_cstr_nargs : int array;
- ci_pp_info : case_printing
+ ci_pp_info : Constr.case_printing
}
[@@ocaml.deprecated "Alias of Constr.case_info"]
type ('constr, 'types) pfixpoint =
- (int array * int) * ('constr, 'types) prec_declaration
+ (int array * int) * ('constr, 'types) Constr.prec_declaration
[@@ocaml.deprecated "Alias of Constr.pfixpoint"]
type ('constr, 'types) pcofixpoint =
- int * ('constr, 'types) prec_declaration
+ int * ('constr, 'types) Constr.prec_declaration
[@@ocaml.deprecated "Alias of Constr.pcofixpoint"]
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 Constr.metavariable
- | Evar of 'constr pexistential
+ | Evar of 'constr Constr.pexistential
| Sort of 'sort
- | Cast of 'constr * cast_kind * 'types
+ | Cast of 'constr * 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
@@ -942,22 +1040,18 @@ sig
| Const of (Names.Constant.t * 'univs)
| Ind of (Names.inductive * 'univs)
| Construct of (Names.constructor * 'univs)
- | Case of case_info * 'constr * 'constr * 'constr array
- | Fix of ('constr, 'types) pfixpoint
- | CoFix of ('constr, 'types) pcofixpoint
+ | Case of Constr.case_info * 'constr * 'constr * 'constr array
+ | Fix of ('constr, 'types) Constr.pfixpoint
+ | CoFix of ('constr, 'types) Constr.pcofixpoint
| Proj of Names.Projection.t * 'constr
[@@ocaml.deprecated "Alias of Constr.kind_of_term"]
- type existential = Constr.existential_key * constr array
+ type existential = Constr.existential_key * Constr.constr array
[@@ocaml.deprecated "Alias of Constr.existential"]
- type rec_declaration = Names.Name.t array * constr array * constr array
+ type rec_declaration = Names.Name.t array * Constr.constr array * Constr.constr array
[@@ocaml.deprecated "Alias of Constr.rec_declaration"]
- type fixpoint = (int array * int) * rec_declaration
- [@@ocaml.deprecated "Alias of Constr.fixpoint"]
- type cofixpoint = int * rec_declaration
- [@@ocaml.deprecated "Alias of Constr.cofixpoint"]
- val kind_of_term : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term
+ val kind_of_term : Constr.constr -> (Constr.constr, Constr.types, Sorts.t, Univ.Instance.t) Constr.kind_of_term
[@@ocaml.deprecated "Alias of Constr.kind"]
- val applistc : constr -> constr list -> constr
+ val applistc : Constr.constr -> Constr.constr list -> Constr.constr
val applist : constr * constr list -> constr
[@@ocaml.deprecated "(sort of an) alias of API.Term.applistc"]
@@ -971,7 +1065,7 @@ sig
val mkMeta : Constr.metavariable -> constr
[@@ocaml.deprecated "Alias of similarly named Constr function"]
- val mkEvar : existential -> constr
+ val mkEvar : Constr.existential -> constr
[@@ocaml.deprecated "Alias of similarly named Constr function"]
val mkSort : Sorts.t -> types
[@@ocaml.deprecated "Alias of similarly named Constr function"]
@@ -981,7 +1075,7 @@ sig
[@@ocaml.deprecated "Alias of similarly named Constr function"]
val mkType : Univ.Universe.t -> types
[@@ocaml.deprecated "Alias of similarly named Constr function"]
- val mkCast : constr * cast_kind * constr -> constr
+ val mkCast : constr * Constr.cast_kind * constr -> constr
[@@ocaml.deprecated "Alias of similarly named Constr function"]
val mkProd : Names.Name.t * types * types -> types
[@@ocaml.deprecated "Alias of similarly named Constr function"]
@@ -999,11 +1093,11 @@ sig
[@@ocaml.deprecated "Alias of similarly named Constr function"]
val mkConstruct : Names.constructor -> constr
[@@ocaml.deprecated "Alias of similarly named Constr function"]
- val mkConstructU : Names.constructor puniverses -> constr
+ val mkConstructU : Names.constructor Constr.puniverses -> constr
[@@ocaml.deprecated "Alias of similarly named Constr function"]
- val mkConstructUi : (pinductive * int) -> constr
+ val mkConstructUi : (Constr.pinductive * int) -> constr
[@@ocaml.deprecated "Alias of similarly named Constr function"]
- val mkCase : case_info * constr * constr * constr array -> constr
+ val mkCase : Constr.case_info * constr * constr * constr array -> constr
[@@ocaml.deprecated "Alias of similarly named Constr function"]
val mkFix : fixpoint -> constr
[@@ocaml.deprecated "Alias of similarly named Constr function"]
@@ -1015,6 +1109,8 @@ sig
val mkNamedProd : Names.Id.t -> types -> types -> types
val decompose_app : constr -> constr * constr list
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
+
val decompose_prod : constr -> (Names.Name.t*constr) list * constr
val decompose_prod_n : int -> constr -> (Names.Name.t * constr) list * constr
val decompose_prod_assum : types -> Context.Rel.t * types
@@ -1026,26 +1122,46 @@ sig
val compose_lam : (Names.Name.t * constr) list -> constr -> constr
val destSort : constr -> Sorts.t
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val destVar : constr -> Names.Id.t
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val destApp : constr -> constr * constr array
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val destProd : types -> Names.Name.t * types * types
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val destLetIn : constr -> Names.Name.t * constr * types * constr
- val destEvar : constr -> existential
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
+ val destEvar : constr -> Constr.existential
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val destRel : constr -> int
- val destConst : constr -> Names.Constant.t puniverses
- val destCast : constr -> constr * cast_kind * constr
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
+ val destConst : constr -> Names.Constant.t Constr.puniverses
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
+ val destCast : constr -> constr * Constr.cast_kind * constr
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val destLambda : constr -> Names.Name.t * types * constr
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val isRel : constr -> bool
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val isVar : constr -> bool
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val isEvar : constr -> bool
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val isLetIn : constr -> bool
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val isLambda : constr -> bool
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val isConst : constr -> bool
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val isEvar_or_Meta : constr -> bool
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val isCast : constr -> bool
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val isMeta : constr -> bool
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val isApp : constr -> bool
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val fold_constr : ('a -> constr -> 'a) -> 'a -> constr -> 'a
[@@ocaml.deprecated "Alias of Constr.fold"]
@@ -1059,13 +1175,13 @@ sig
val it_mkLambda_or_LetIn : constr -> Context.Rel.t -> constr
val it_mkProd_or_LetIn : types -> Context.Rel.t -> types
val prod_applist : constr -> constr list -> constr
- exception DestKO
+
val map_constr : (constr -> constr) -> constr -> constr
[@@ocaml.deprecated "Alias of Constr.map"]
- val mkIndU : pinductive -> constr
+ val mkIndU : Constr.pinductive -> constr
[@@ocaml.deprecated "Alias of Constr.mkIndU"]
- val mkConstU : pconstant -> constr
+ val mkConstU : Constr.pconstant -> constr
[@@ocaml.deprecated "Alias of Constr.mkConstU"]
val map_constr_with_binders :
('a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr
@@ -1104,18 +1220,31 @@ sig
val constr_ord : constr -> constr -> int
[@@ocaml.deprecated "alias of Term.compare"]
- val destInd : constr -> Names.inductive puniverses
+ val destInd : constr -> Names.inductive Constr.puniverses
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val univ_of_sort : Sorts.t -> Univ.Universe.t
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val strip_lam : constr -> constr
val strip_prod_assum : types -> types
val decompose_lam_assum : constr -> Context.Rel.t * constr
val destFix : constr -> fixpoint
+ [@@ocaml.deprecated "Alias for the function in [Constr]"]
val compare_constr : (constr -> constr -> bool) -> constr -> constr -> bool
[@@ocaml.deprecated "Alias of Constr.compare_head."]
+ type constr = Constr.t
+ [@@ocaml.deprecated "Alias of Constr.t"]
+ type types = Constr.t
+ [@@ocaml.deprecated "Alias of Constr.types"]
+
+ type fixpoint = (int array * int) * Constr.rec_declaration
+ [@@ocaml.deprecated "Alias of Constr.Constr.fixpoint"]
+ type cofixpoint = int * Constr.rec_declaration
+ [@@ocaml.deprecated "Alias of Constr.cofixpoint"]
+
end
module Mod_subst :
@@ -1288,8 +1417,8 @@ sig
| TemplateArity of 'b
type constant_universes =
- | Monomorphic_const of Univ.universe_context
- | Polymorphic_const of Univ.abstract_universe_context
+ | Monomorphic_const of Univ.UContext.t
+ | Polymorphic_const of Univ.AUContext.t
type projection_body = {
proj_ind : Names.MutInd.t;
@@ -1308,7 +1437,7 @@ sig
type constant_body = {
const_hyps : Context.Named.t;
const_body : constant_def;
- const_type : Term.types;
+ const_type : Constr.types;
const_body_code : Cemitcodes.to_patch_substituted option;
const_universes : constant_universes;
const_proj : projection_body option;
@@ -1355,12 +1484,12 @@ sig
| MEwith of module_alg_expr * with_declaration
type abstract_inductive_universes =
- | Monomorphic_ind of Univ.universe_context
- | Polymorphic_ind of Univ.abstract_universe_context
- | Cumulative_ind of Univ.abstract_cumulativity_info
+ | Monomorphic_ind of Univ.UContext.t
+ | Polymorphic_ind of Univ.AUContext.t
+ | Cumulative_ind of Univ.ACumulativityInfo.t
type record_body = (Id.t * Constant.t array * projection_body array) option
-
+
type mutual_inductive_body = {
mind_packets : one_inductive_body array;
mind_record : record_body option;
@@ -1422,9 +1551,9 @@ sig
| 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
+ | Monomorphic_ind_entry of Univ.UContext.t
+ | Polymorphic_ind_entry of Univ.UContext.t
+ | Cumulative_ind_entry of Univ.CumulativityInfo.t
type one_inductive_entry = {
mind_entry_typename : Id.t;
@@ -1451,8 +1580,8 @@ sig
type 'a proof_output = Constr.t Univ.in_universe_context_set * 'a
type 'a const_entry_body = 'a proof_output Future.computation
type constant_universes_entry =
- | Monomorphic_const_entry of Univ.universe_context
- | Polymorphic_const_entry of Univ.universe_context
+ | Monomorphic_const_entry of Univ.UContext.t
+ | Polymorphic_const_entry of Univ.UContext.t
type 'a definition_entry =
{ const_entry_body : 'a const_entry_body;
(* List of section variables *)
@@ -1493,12 +1622,12 @@ sig
utj_val : 'types;
utj_type : Sorts.t }
- type unsafe_type_judgment = Term.types punsafe_type_judgment
+ type unsafe_type_judgment = Constr.types punsafe_type_judgment
val empty_env : env
val lookup_mind : Names.MutInd.t -> env -> Declarations.mutual_inductive_body
val push_rel : Context.Rel.Declaration.t -> env -> env
val push_rel_context : Context.Rel.t -> env -> env
- val push_rec_types : Term.rec_declaration -> env -> env
+ val push_rec_types : Constr.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 -> named_context_val -> Context.Named.Declaration.t
@@ -1538,13 +1667,13 @@ sig
| 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 *)
+ | FFix of Constr.fixpoint * fconstr Esubst.subs
+ | FCoFix of Constr.cofixpoint * fconstr Esubst.subs
+ | FCaseT of Constr.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
+ | FEvar of Constr.existential * fconstr Esubst.subs
| FLIFT of int * fconstr
| FCLOS of Constr.t * fconstr Esubst.subs
| FLOCKED
@@ -1580,7 +1709,7 @@ sig
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 create_clos_infos : ?evars:(Constr.existential -> Constr.t option) -> RedFlags.reds -> Environ.env -> clos_infos
val whd_val : clos_infos -> fconstr -> Constr.t
@@ -1601,13 +1730,13 @@ sig
val whd_betaiotazeta : Environ.env -> Constr.t -> Constr.t
- val is_arity : Environ.env -> Term.types -> bool
+ val is_arity : Environ.env -> Constr.types -> bool
- val dest_prod : Environ.env -> Term.types -> Context.Rel.t * Term.types
+ val dest_prod : Environ.env -> Constr.types -> Context.Rel.t * Constr.types
type 'a extended_conversion_function =
?l2r:bool -> ?reds:Names.transparent_state -> Environ.env ->
- ?evars:((Term.existential->Constr.t option) * UGraph.t) ->
+ ?evars:((Constr.existential->Constr.t option) * UGraph.t) ->
'a -> 'a -> unit
val conv : Constr.t extended_conversion_function
end
@@ -1616,7 +1745,7 @@ module Type_errors :
sig
open Names
- open Term
+ open Constr
open Environ
type 'constr pguard_error =
@@ -1648,9 +1777,9 @@ sig
| 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
+ | ReferenceVariables of Id.t * '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
@@ -1682,16 +1811,16 @@ end
module Inductive :
sig
type mind_specif = Declarations.mutual_inductive_body * Declarations.one_inductive_body
- val type_of_inductive : Environ.env -> mind_specif Univ.puniverses -> Term.types
+ val type_of_inductive : Environ.env -> mind_specif Univ.puniverses -> Constr.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 * Constr.t list
+ val find_inductive : Environ.env -> Constr.types -> Constr.pinductive * Constr.t list
end
module Typeops :
sig
- val infer_type : Environ.env -> Term.types -> Environ.unsafe_type_judgment
- val type_of_constant_in : Environ.env -> Term.pconstant -> Term.types
+ val infer_type : Environ.env -> Constr.types -> Environ.unsafe_type_judgment
+ val type_of_constant_in : Environ.env -> Constr.pconstant -> Constr.types
end
module Mod_typing :
@@ -1756,7 +1885,7 @@ sig
type glob_constraint = glob_level * Univ.constraint_type * glob_level
- type case_style = Term.case_style =
+ type case_style = Constr.case_style =
| LetStyle
| IfStyle
| LetPatternStyle
@@ -1857,8 +1986,8 @@ end
module Univops :
sig
- val universes_of_constr : Term.constr -> Univ.universe_set
- val restrict_universe_context : Univ.universe_context_set -> Univ.universe_set -> Univ.universe_context_set
+ val universes_of_constr : Constr.constr -> Univ.LSet.t
+ val restrict_universe_context : Univ.ContextSet.t -> Univ.LSet.t -> Univ.ContextSet.t
end
module Nameops :
@@ -1934,8 +2063,10 @@ sig
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.t
+ [@@ocaml.deprecated "Alias for DirPath.print"]
val string_of_path : full_path -> string
+
val basename : full_path -> Names.Id.t
type object_name = full_path * Names.KerName.t
@@ -2006,7 +2137,7 @@ module Pattern :
sig
type case_info_pattern =
- { cip_style : Misctypes.case_style;
+ { cip_style : Constr.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 _ => _ ? *) }
@@ -2027,8 +2158,8 @@ sig
| 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
+ | PFix of Constr.fixpoint
+ | PCoFix of Constr.cofixpoint
end
@@ -2079,7 +2210,7 @@ sig
| GLambda of Names.Name.t * Decl_kinds.binding_kind * 'a glob_constr_g * 'a glob_constr_g
| GProd of Names.Name.t * Decl_kinds.binding_kind * 'a glob_constr_g * 'a glob_constr_g
| GLetIn of Names.Name.t * 'a glob_constr_g * 'a glob_constr_g option * 'a glob_constr_g
- | GCases of Term.case_style * 'a glob_constr_g option * 'a tomatch_tuples_g * 'a cases_clauses_g
+ | GCases of Constr.case_style * 'a glob_constr_g option * 'a tomatch_tuples_g * 'a cases_clauses_g
| GLetTuple of Names.Name.t list * (Names.Name.t * 'a glob_constr_g option) * 'a glob_constr_g * 'a glob_constr_g
| GIf of 'a glob_constr_g * (Names.Name.t * 'a glob_constr_g option) * 'a glob_constr_g * 'a glob_constr_g
| GRec of 'a fix_kind_g * Names.Id.t array * 'a glob_decl_g list array *
@@ -2142,7 +2273,7 @@ sig
| 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 *
+ | NCases of Constr.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) *
@@ -2214,7 +2345,7 @@ sig
| 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
+ | CCases of Constr.case_style
* constr_expr option
* case_expr list
* branch_expr list
@@ -2606,9 +2737,9 @@ 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 fresh_inductive_instance : Environ.env -> Names.inductive -> Constr.pinductive Univ.in_universe_context_set
+ val new_Type : Names.DirPath.t -> Constr.types
+ val type_of_global : Globnames.global_reference -> Constr.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
@@ -2733,7 +2864,7 @@ sig
val create_evar_defs : evar_map -> evar_map
- val meta_declare : Constr.metavariable -> Term.types -> ?name:Names.Name.t -> evar_map -> evar_map
+ val meta_declare : Constr.metavariable -> Constr.types -> ?name:Names.Name.t -> evar_map -> evar_map
val clear_metas : evar_map -> evar_map
@@ -2744,7 +2875,7 @@ sig
val fresh_global : ?loc:Loc.t -> ?rigid:rigid -> ?names:Univ.Instance.t -> Environ.env ->
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 fresh_inductive_instance : ?loc:Loc.t -> Environ.env -> evar_map -> Names.inductive -> evar_map * Constr.pinductive
val fold_undefined : (Evar.t -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'a
val universe_context_set : evar_map -> Univ.ContextSet.t
@@ -2801,8 +2932,8 @@ sig
type evar_universe_context = UState.t
[@@ocaml.deprecated "alias of API.UState.t"]
- val existential_opt_value : evar_map -> Term.existential -> Constr.t option
- val existential_value : evar_map -> Term.existential -> Constr.t
+ val existential_opt_value : evar_map -> Constr.existential -> Constr.t option
+ val existential_value : evar_map -> Constr.existential -> Constr.t
exception NotInstantiatedEvar
@@ -3033,7 +3164,7 @@ sig
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. *)
+ (** Remove the outer-most {!Constr.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].
@@ -3044,7 +3175,7 @@ sig
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
+ val push_rels_assum : (Names.Name.t * Constr.types) list -> Environ.env -> Environ.env
type meta_value_map = (Constr.metavariable * Constr.t) list
@@ -3146,7 +3277,7 @@ sig
?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 safe_evar_value : Evd.evar_map -> Constr.existential -> Constr.t option
val evd_comb1 : (Evd.evar_map -> 'b -> Evd.evar_map * 'a) -> Evd.evar_map ref -> 'b -> 'a
end
@@ -3514,14 +3645,14 @@ sig
| IndType of inductive_family * EConstr.constr list
type constructor_summary =
{
- cs_cstr : Term.pconstructor;
+ cs_cstr : Constr.pconstructor;
cs_params : Constr.t list;
cs_nargs : int;
cs_args : Context.Rel.t;
cs_concl_realargs : Constr.t array;
}
- val arities_of_constructors : Environ.env -> Term.pinductive -> Term.types array
+ val arities_of_constructors : Environ.env -> Constr.pinductive -> Constr.types array
val constructors_nrealargs_env : Environ.env -> Names.inductive -> int array
val constructor_nallargs_env : Environ.env -> Names.constructor -> int
@@ -3529,16 +3660,16 @@ sig
val inductive_nparamdecls : Names.inductive -> int
- val type_of_constructors : Environ.env -> Term.pinductive -> Term.types array
+ val type_of_constructors : Environ.env -> Constr.pinductive -> Constr.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 dest_ind_family : inductive_family -> Names.inductive Constr.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
+ val type_of_inductive : Environ.env -> Constr.pinductive -> Constr.types
end
module Impargs :
@@ -4190,12 +4321,12 @@ 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 ->
+ val build_case_analysis_scheme : Environ.env -> Evd.evar_map -> Constr.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 ->
+ Environ.env -> Evd.evar_map -> (Constr.pinductive * dep_flag * Sorts.family) list -> Evd.evar_map * Constr.t list
+ val build_case_analysis_scheme_default : Environ.env -> Evd.evar_map -> Constr.pinductive ->
Sorts.family -> Evd.evar_map * Constr.t
end
@@ -4490,13 +4621,13 @@ sig
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 ->
- Constrexpr.constr_expr -> Term.types Evd.in_evar_universe_context
+ Constrexpr.constr_expr -> Constr.types Evd.in_evar_universe_context
val interp_context_evars :
?global_level:bool -> ?impl_env:internalization_env -> ?shift:int ->
Environ.env -> Evd.evar_map ref -> Constrexpr.local_binder_expr list ->
internalization_env * ((Environ.env * EConstr.rel_context) * Impargs.manual_implicits)
val compute_internalization_data : Environ.env -> var_internalization_type ->
- Term.types -> Impargs.manual_explicitation list -> var_internalization_data
+ Constr.types -> Impargs.manual_explicitation list -> var_internalization_data
val empty_internalization_env : internalization_env
val global_reference : Names.Id.t -> Globnames.global_reference
end
@@ -4525,7 +4656,7 @@ sig
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
+ | SectionLocalAssum of Constr.types Univ.in_universe_context_set * Decl_kinds.polymorphic * bool
type variable_declaration = Names.DirPath.t * section_variable_entry * Decl_kinds.logical_kind
@@ -4539,7 +4670,7 @@ sig
?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 ->
+ ?opaque:bool -> ?inline:bool -> ?types:Constr.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
@@ -5188,9 +5319,8 @@ sig
val pr_closed_glob_env : Environ.env -> Evd.evar_map -> Ltac_pretype.closed_glob_constr -> Pp.t
val pr_rel_context_of : Environ.env -> Evd.evar_map -> Pp.t
val pr_named_context_of : Environ.env -> Evd.evar_map -> Pp.t
- val pr_ltype : Term.types -> Pp.t
+ val pr_ltype : Constr.types -> Pp.t
[@@ocaml.deprecated "The global printing API is deprecated, please use the _env functions"]
-
val pr_ljudge : EConstr.unsafe_judgment -> Pp.t * Pp.t
[@@ocaml.deprecated "The global printing API is deprecated, please use the _env functions"]
@@ -5622,7 +5752,7 @@ end
module Hints :
sig
- type raw_hint = EConstr.t * EConstr.types * Univ.universe_context_set
+ type raw_hint = EConstr.t * EConstr.types * Univ.ContextSet.t
type 'a hint_ast =
| Res_pf of 'a (* Hint Apply *)
@@ -5772,7 +5902,7 @@ end
module Autorewrite :
sig
type rew_rule = { rew_lemma: Constr.t;
- rew_type: Term.types;
+ rew_type: Constr.types;
rew_pat: Constr.t;
rew_ctx: Univ.ContextSet.t;
rew_l2r: bool;
@@ -5836,9 +5966,6 @@ end
module Locality :
sig
val make_section_locality : bool option -> bool
- module LocalityFixme : sig
- val consume : unit -> bool option
- end
val make_module_locality : bool option -> bool
end
@@ -5981,8 +6108,15 @@ sig
type deprecation = bool
+ type atts = {
+ loc : Loc.t option;
+ locality : bool option;
+ }
+
type vernac_command =
- Genarg.raw_generic_argument list -> Loc.t option -> Vernacstate.t -> Vernacstate.t
+ Genarg.raw_generic_argument list ->
+ atts:atts -> st:Vernacstate.t ->
+ Vernacstate.t
val vinterp_add : deprecation -> Vernacexpr.extend_name -> vernac_command -> unit
diff --git a/CONTRIBUTING.md b/CONTRIBUTING.md
index db02f7834..b4e6a1418 100644
--- a/CONTRIBUTING.md
+++ b/CONTRIBUTING.md
@@ -36,6 +36,10 @@ Here are a few tags Coq developers may add to your PR and what they mean. In gen
- [needs: fixing](https://github.com/coq/coq/pulls?q=is%3Aopen+is%3Apr+label%3A%22needs%3A+fixing%22) indicates the PR needs a fix, as discussed in the comments.
- [needs: testing](https://github.com/coq/coq/pulls?q=is%3Aopen+is%3Apr+label%3A%22needs%3A+testing%22) indicates the PR needs testing. This is often used when testing beyond what the test suite can handle is required. For example, performance benchmarking is currently performed with a different infrastructure. Unless some followup is specifically requested you aren't expected to do this additional testing.
+The release manager uses the following filter to know which PRs seem ready for merge. If you are waiting for a PR to be merged, make sure it appears in this list:
+
+- [Pull requests ready for merge](https://github.com/coq/coq/pulls?utf8=%E2%9C%93&q=is%3Apr%20is%3Aopen%20-label%3A%22needs%3A%20discussion%22%20-label%3A%22needs%3A%20testing%22%20-label%3A%22needs%3A%20fixing%22%20-label%3A%22needs%3A%20progress%22%20-label%3A%22needs%3A%20rebase%22%20-label%3A%22needs%3A%20review%22%20-label%3A%22needs%3A%20help%22%20-label%3A%22needs%3A%20independent%20fix%22%20-label%3A%22needs%3A%20feedback%22%20-label%3A%22help%20wanted%22%20-review%3Achanges_requested%20-status%3Apending%20base%3Amaster%20sort%3Aupdated-asc%20-label%3A%22needs%3A%20squashing%22%20)
+
## Documentation
Currently the process for contributing to the documentation is the same as for changing anything else in Coq, so please submit a pull request as described above.
diff --git a/dev/ci/user-overlays/06197-ejgallego-plugins+remove_locality_hack.sh b/dev/ci/user-overlays/06197-ejgallego-plugins+remove_locality_hack.sh
new file mode 100644
index 000000000..c9f1272be
--- /dev/null
+++ b/dev/ci/user-overlays/06197-ejgallego-plugins+remove_locality_hack.sh
@@ -0,0 +1,4 @@
+if [ "$TRAVIS_PULL_REQUEST" = "6197" ] || [ "$TRAVIS_BRANCH" = "plugins+remove_locality_hack" ]; then
+ ltac2_CI_BRANCH=localityfixyou
+ ltac2_CI_GITURL=https://github.com/ejgallego/ltac2.git
+fi
diff --git a/dev/top_printers.ml b/dev/top_printers.ml
index 0f496d3b9..4e7b94e41 100644
--- a/dev/top_printers.ml
+++ b/dev/top_printers.ml
@@ -39,7 +39,7 @@ let ppfuture kx = pp (Future.print (fun _ -> str "_") kx)
let ppid id = pp (Id.print id)
let pplab l = pp (Label.print l)
let ppmbid mbid = pp (str (MBId.debug_to_string mbid))
-let ppdir dir = pp (pr_dirpath dir)
+let ppdir dir = pp (DirPath.print dir)
let ppmp mp = pp(str (ModPath.debug_to_string mp))
let ppcon con = pp(Constant.debug_print con)
let ppproj con = pp(Constant.debug_print (Projection.constant con))
@@ -509,7 +509,7 @@ let _ =
(function
[c] when genarg_tag c = unquote (topwit wit_constr) && true ->
let c = out_gen (rawwit wit_constr) c in
- (fun _ st -> in_current_context constr_display c; st)
+ (fun ~atts ~st -> in_current_context constr_display c; st)
| _ -> failwith "Vernac extension: cannot occur")
with
e -> pp (CErrors.print e)
@@ -525,7 +525,7 @@ let _ =
(function
[c] when genarg_tag c = unquote (topwit wit_constr) && true ->
let c = out_gen (rawwit wit_constr) c in
- (fun _ st -> in_current_context print_pure_constr c; st)
+ (fun ~atts ~st -> in_current_context print_pure_constr c; st)
| _ -> failwith "Vernac extension: cannot occur")
with
e -> pp (CErrors.print e)
diff --git a/engine/eConstr.ml b/engine/eConstr.ml
index bcfbc8081..afdceae06 100644
--- a/engine/eConstr.ml
+++ b/engine/eConstr.ml
@@ -9,7 +9,6 @@
open CErrors
open Util
open Names
-open Term
open Constr
open Context
open Evd
@@ -55,7 +54,7 @@ struct
type t = Sorts.t
let make s = s
let kind sigma = function
- | Type u -> sort_of_univ (Evd.normalize_universe sigma u)
+ | Sorts.Type u -> Sorts.sort_of_univ (Evd.normalize_universe sigma u)
| s -> s
let unsafe_to_sorts s = s
end
@@ -85,16 +84,16 @@ let rec whd_evar sigma c =
| Some c -> whd_evar sigma c
| None -> c
end
- | App (f, args) when Term.isEvar f ->
+ | App (f, args) when isEvar f ->
(** Enforce smart constructor invariant on applications *)
- let ev = Term.destEvar f in
+ let ev = destEvar f in
begin match safe_evar_value sigma ev with
| None -> c
| Some f -> whd_evar sigma (mkApp (f, args))
end
- | Cast (c0, k, t) when Term.isEvar c0 ->
+ | Cast (c0, k, t) when isEvar c0 ->
(** Enforce smart constructor invariant on casts. *)
- let ev = Term.destEvar c0 in
+ let ev = destEvar c0 in
begin match safe_evar_value sigma ev with
| None -> c
| Some c -> whd_evar sigma (mkCast (c, k, t))
@@ -115,7 +114,7 @@ let rec to_constr sigma c = match Constr.kind c with
| Some c -> to_constr sigma c
| None -> Constr.map (fun c -> to_constr sigma c) c
end
-| Sort (Type u) ->
+| Sort (Sorts.Type u) ->
let u' = Evd.normalize_universe sigma u in
if u' == u then c else mkSort (Sorts.sort_of_univ u')
| Const (c', u) when not (Univ.Instance.is_empty u) ->
diff --git a/engine/engine.mllib b/engine/engine.mllib
index afc02d7f6..a3614f6c4 100644
--- a/engine/engine.mllib
+++ b/engine/engine.mllib
@@ -1,12 +1,13 @@
-Logic_monad
Universes
+Univops
UState
+Nameops
Evd
EConstr
Namegen
Termops
-Proofview_monad
Evarutil
+Logic_monad
+Proofview_monad
Proofview
Ftactic
-Geninterp
diff --git a/engine/evarutil.ml b/engine/evarutil.ml
index df4ef2ce7..14d07ccae 100644
--- a/engine/evarutil.ml
+++ b/engine/evarutil.ml
@@ -11,11 +11,11 @@ open Util
open Names
open Term
open Constr
-open Termops
-open Namegen
open Pre_env
open Environ
open Evd
+open Termops
+open Namegen
module RelDecl = Context.Rel.Declaration
module NamedDecl = Context.Named.Declaration
diff --git a/engine/evarutil.mli b/engine/evarutil.mli
index 62288ced4..42f2d5f25 100644
--- a/engine/evarutil.mli
+++ b/engine/evarutil.mli
@@ -238,7 +238,8 @@ val subterm_source : existential_key -> Evar_kinds.t Loc.located ->
val meta_counter_summary_name : string
-(** Deprecater *)
-
+(** Deprecated *)
type type_constraint = types option
+[@@ocaml.deprecated "use the version in Evardefine"]
type val_constraint = constr option
+[@@ocaml.deprecated "use the version in Evardefine"]
diff --git a/engine/evd.ml b/engine/evd.ml
index 8d465384b..60bd6de2a 100644
--- a/engine/evd.ml
+++ b/engine/evd.ml
@@ -243,7 +243,7 @@ let evar_instance_array test_id info args =
instrec filter (evar_context info) 0
let make_evar_instance_array info args =
- evar_instance_array (NamedDecl.get_id %> Term.isVarId) info args
+ evar_instance_array (NamedDecl.get_id %> isVarId) info args
let instantiate_evar_array info c args =
let inst = make_evar_instance_array info args in
@@ -707,10 +707,10 @@ let extract_all_conv_pbs evd =
extract_conv_pbs evd (fun _ -> true)
let loc_of_conv_pb evd (pbty,env,t1,t2) =
- match kind (fst (Term.decompose_app t1)) with
+ match kind (fst (decompose_app t1)) with
| Evar (evk1,_) -> fst (evar_source evk1 evd)
| _ ->
- match kind (fst (Term.decompose_app t2)) with
+ match kind (fst (decompose_app t2)) with
| Evar (evk2,_) -> fst (evar_source evk2 evd)
| _ -> None
diff --git a/engine/evd.mli b/engine/evd.mli
index af5373582..17fa15045 100644
--- a/engine/evd.mli
+++ b/engine/evd.mli
@@ -599,11 +599,16 @@ type open_constr = evar_map * constr (* Special case when before is empty *)
type unsolvability_explanation = SeveralInstancesFound of int
(** Failure explanation. *)
+(** {5 Summary names} *)
+
+(* This stuff is internal and should not be used. Currently a hack in
+ the STM relies on it. *)
+val evar_counter_summary_name : string
+
(** {5 Deprecated functions} *)
+val create_evar_defs : evar_map -> evar_map
+(* XXX: This is supposed to be deprecated by used by ssrmatching, what
+ should the replacement be? *)
-val create_evar_defs : evar_map -> evar_map
(** Create an [evar_map] with empty meta map: *)
-(** {5 Summary names} *)
-
-val evar_counter_summary_name : string
diff --git a/library/nameops.ml b/engine/nameops.ml
index d598a63b8..5105d7bec 100644
--- a/library/nameops.ml
+++ b/engine/nameops.ml
@@ -203,13 +203,14 @@ let pr_name = print
let pr_lab l = Label.print l
-let default_library = Names.DirPath.initial (* = ["Top"] *)
-
-(*s Roots of the space of absolute names *)
-let coq_string = "Coq"
-let coq_root = Id.of_string coq_string
-let default_root_prefix = DirPath.empty
-
(* Metavariables *)
let pr_meta = Pp.int
let string_of_meta = string_of_int
+
+(* Deprecated *)
+open Libnames
+let default_library = default_library
+let coq_string = coq_string
+let coq_root = coq_root
+let default_root_prefix = default_root_prefix
+
diff --git a/library/nameops.mli b/engine/nameops.mli
index 60e5a90bb..0fec8a925 100644
--- a/library/nameops.mli
+++ b/engine/nameops.mli
@@ -89,6 +89,10 @@ module Name : sig
end
+(** Metavariables *)
+val pr_meta : Constr.metavariable -> Pp.t
+val string_of_meta : Constr.metavariable -> string
+
val out_name : Name.t -> Id.t
[@@ocaml.deprecated "Same as [Name.get_id]"]
@@ -119,18 +123,16 @@ val pr_id : Id.t -> Pp.t
val pr_lab : Label.t -> Pp.t
[@@ocaml.deprecated "Same as [Names.Label.print]"]
-(** some preset paths *)
-
+(** Deprecated stuff to libnames *)
val default_library : DirPath.t
+[@@ocaml.deprecated "Same as [Libnames.default_library]"]
-(** This is the root of the standard library of Coq *)
val coq_root : module_ident (** "Coq" *)
+[@@ocaml.deprecated "Same as [Libnames.coq_root]"]
+
val coq_string : string (** "Coq" *)
+[@@ocaml.deprecated "Same as [Libnames.coq_string]"]
-(** This is the default root prefix for developments which doesn't
- mention a root *)
val default_root_prefix : DirPath.t
+[@@ocaml.deprecated "Same as [Libnames.default_root_prefix]"]
-(** Metavariables *)
-val pr_meta : Constr.metavariable -> Pp.t
-val string_of_meta : Constr.metavariable -> string
diff --git a/engine/proofview.ml b/engine/proofview.ml
index 598358c47..3b945c87f 100644
--- a/engine/proofview.ml
+++ b/engine/proofview.ml
@@ -1200,7 +1200,7 @@ module V82 = struct
{ Evd.it = comb ; sigma = solution }
let top_goals initial { solution=solution; } =
- let goals = CList.map (fun (t,_) -> fst (Term.destEvar (EConstr.Unsafe.to_constr t))) initial in
+ let goals = CList.map (fun (t,_) -> fst (Constr.destEvar (EConstr.Unsafe.to_constr t))) initial in
{ Evd.it = goals ; sigma=solution; }
let top_evars initial =
diff --git a/engine/uState.ml b/engine/uState.ml
index dfea25dd0..01a479821 100644
--- a/engine/uState.ml
+++ b/engine/uState.ml
@@ -437,6 +437,9 @@ let make_flexible_variable ctx ~algebraic u =
{ctx with uctx_univ_variables = uvars';
uctx_univ_algebraic = avars'}
+let make_flexible_nonalgebraic ctx =
+ {ctx with uctx_univ_algebraic = Univ.LSet.empty}
+
let is_sort_variable uctx s =
match s with
| Sorts.Type u ->
diff --git a/engine/uState.mli b/engine/uState.mli
index b31e94b28..1c906fcb2 100644
--- a/engine/uState.mli
+++ b/engine/uState.mli
@@ -104,6 +104,11 @@ val add_global_univ : t -> Univ.Level.t -> t
universe. Otherwise the variable is just made flexible. *)
val make_flexible_variable : t -> algebraic:bool -> Univ.Level.t -> t
+(** Turn all undefined flexible algebraic variables into simply flexible
+ ones. Can be used in case the variables might appear in universe instances
+ (typically for polymorphic program obligations). *)
+val make_flexible_nonalgebraic : t -> t
+
val is_sort_variable : t -> Sorts.t -> Univ.Level.t option
val normalize_variables : t -> Univ.universe_subst * t
diff --git a/engine/universes.mli b/engine/universes.mli
index 24613c4b9..a960099ed 100644
--- a/engine/universes.mli
+++ b/engine/universes.mli
@@ -169,6 +169,7 @@ val constr_of_global : Globnames.global_reference -> constr
(** ** DEPRECATED ** synonym of [constr_of_global] *)
val constr_of_reference : Globnames.global_reference -> constr
+[@@ocaml.deprecated "synonym of [constr_of_global]"]
(** Returns the type of the global reference, by creating a fresh instance of polymorphic
references and computing their instantiated universe context. (side-effect on the
diff --git a/library/univops.ml b/engine/univops.ml
index 9dc138eb8..d498b2e0d 100644
--- a/library/univops.ml
+++ b/engine/univops.ml
@@ -6,8 +6,8 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Constr
open Univ
+open Constr
let universes_of_constr c =
let rec aux s c =
@@ -15,7 +15,7 @@ let universes_of_constr c =
| Const (_, u) | Ind (_, u) | Construct (_, u) ->
LSet.fold LSet.add (Instance.levels u) s
| Sort u when not (Sorts.is_small u) ->
- let u = Term.univ_of_sort u in
+ let u = Sorts.univ_of_sort u in
LSet.fold LSet.add (Universe.levels u) s
| _ -> Constr.fold aux s c
in aux LSet.empty c
diff --git a/library/univops.mli b/engine/univops.mli
index 9af568bcb..9af568bcb 100644
--- a/library/univops.mli
+++ b/engine/univops.mli
diff --git a/grammar/vernacextend.mlp b/grammar/vernacextend.mlp
index 12308bede..5bc8f1504 100644
--- a/grammar/vernacextend.mlp
+++ b/grammar/vernacextend.mlp
@@ -136,6 +136,10 @@ EXTEND
OPT "|"; l = LIST1 rule SEP "|";
"END" ->
declare_command loc s c <:expr<None>> l
+ | "VERNAC"; "COMMAND"; "FUNCTIONAL"; "EXTEND"; s = UIDENT; c = OPT classification;
+ OPT "|"; l = LIST1 fun_rule SEP "|";
+ "END" ->
+ declare_command loc s c <:expr<None>> l
| "VERNAC"; nt = LIDENT ; "EXTEND"; s = UIDENT; c = OPT classification;
OPT "|"; l = LIST1 rule SEP "|";
"END" ->
@@ -162,22 +166,27 @@ EXTEND
(which otherwise could have been another argument) is not passed
to the VernacExtend interpreter function to discriminate between
the clauses. *)
-
- (* ejga: Due to the LocalityFixme abomination we cannot eta-expand
- [e] as we'd like to, so we need to use the below mess with [fun
- st -> st].
-
- At some point We should solve the mess and extend
- vernacextend.mlp with locality info. *)
rule:
[ [ "["; s = STRING; l = LIST0 args; "]";
d = OPT deprecation; c = OPT classifier; "->"; "["; e = Pcaml.expr; "]" ->
let () = if s = "" then failwith "Command name is empty." in
- let b = <:expr< fun loc -> ( let () = $e$ in fun st -> st ) >> in
+ let b = <:expr< fun ~atts ~st -> ( let () = $e$ in st ) >> in
+ { r_head = Some s; r_patt = l; r_class = c; r_branch = b; r_depr = d; }
+ | "[" ; "-" ; l = LIST1 args ; "]" ;
+ d = OPT deprecation; c = OPT classifier; "->"; "["; e = Pcaml.expr; "]" ->
+ let b = <:expr< fun ~atts ~st -> ( let () = $e$ in st ) >> in
+ { r_head = None; r_patt = l; r_class = c; r_branch = b; r_depr = d; }
+ ] ]
+ ;
+ fun_rule:
+ [ [ "["; s = STRING; l = LIST0 args; "]";
+ d = OPT deprecation; c = OPT classifier; "->"; "["; e = Pcaml.expr; "]" ->
+ let () = if s = "" then failwith "Command name is empty." in
+ let b = <:expr< $e$ >> in
{ r_head = Some s; r_patt = l; r_class = c; r_branch = b; r_depr = d; }
| "[" ; "-" ; l = LIST1 args ; "]" ;
d = OPT deprecation; c = OPT classifier; "->"; "["; e = Pcaml.expr; "]" ->
- let b = <:expr< fun loc -> ( let () = $e$ in fun st -> st ) >> in
+ let b = <:expr< $e$ >> in
{ r_head = None; r_patt = l; r_class = c; r_branch = b; r_depr = d; }
] ]
;
diff --git a/interp/impargs.ml b/interp/impargs.ml
index f70154e61..3105214d5 100644
--- a/interp/impargs.ml
+++ b/interp/impargs.ml
@@ -10,7 +10,6 @@ open CErrors
open Util
open Names
open Globnames
-open Term
open Constr
open Reduction
open Declarations
diff --git a/interp/reserve.ml b/interp/reserve.ml
index 6fd1d0b58..22c5a2f5e 100644
--- a/interp/reserve.ml
+++ b/interp/reserve.ml
@@ -102,7 +102,7 @@ let declare_reserved_type idl t =
let find_reserved_type id = Id.Map.find (root_of_id id) !reserve_table
let constr_key c =
- try RefKey (canonical_gr (global_of_constr (fst (Term.decompose_app c))))
+ try RefKey (canonical_gr (global_of_constr (fst (Constr.decompose_app c))))
with Not_found -> Oth
let revert_reserved_type t =
diff --git a/kernel/constr.ml b/kernel/constr.ml
index cec00c04b..be59f9341 100644
--- a/kernel/constr.ml
+++ b/kernel/constr.ml
@@ -233,7 +233,6 @@ let mkMeta n = Meta n
(* Constructs a Variable named id *)
let mkVar id = Var id
-
(************************************************************************)
(* kind_of_term = constructions as seen by the user *)
(************************************************************************)
@@ -250,6 +249,168 @@ let of_kind = function
| Cast (c, knd, t) -> mkCast (c, knd, t)
| k -> k
+(**********************************************************************)
+(* Non primitive term destructors *)
+(**********************************************************************)
+
+(* Destructor operations : partial functions
+ Raise [DestKO] if the const has not the expected form *)
+
+exception DestKO
+
+let isMeta c = match kind c with Meta _ -> true | _ -> false
+
+(* Destructs a type *)
+let isSort c = match kind c with
+ | Sort _ -> true
+ | _ -> false
+
+let rec isprop c = match kind c with
+ | Sort (Sorts.Prop _) -> true
+ | Cast (c,_,_) -> isprop c
+ | _ -> false
+
+let rec is_Prop c = match kind c with
+ | Sort (Sorts.Prop Sorts.Null) -> true
+ | Cast (c,_,_) -> is_Prop c
+ | _ -> false
+
+let rec is_Set c = match kind c with
+ | Sort (Sorts.Prop Sorts.Pos) -> true
+ | Cast (c,_,_) -> is_Set c
+ | _ -> false
+
+let rec is_Type c = match kind c with
+ | Sort (Sorts.Type _) -> true
+ | Cast (c,_,_) -> is_Type c
+ | _ -> false
+
+let is_small = Sorts.is_small
+let iskind c = isprop c || is_Type c
+
+(* Tests if an evar *)
+let isEvar c = match kind c with Evar _ -> true | _ -> false
+let isEvar_or_Meta c = match kind c with
+ | Evar _ | Meta _ -> true
+ | _ -> false
+
+let isCast c = match kind c with Cast _ -> true | _ -> false
+(* Tests if a de Bruijn index *)
+let isRel c = match kind c with Rel _ -> true | _ -> false
+let isRelN n c =
+ match kind c with Rel n' -> Int.equal n n' | _ -> false
+(* Tests if a variable *)
+let isVar c = match kind c with Var _ -> true | _ -> false
+let isVarId id c = match kind c with Var id' -> Id.equal id id' | _ -> false
+(* Tests if an inductive *)
+let isInd c = match kind c with Ind _ -> true | _ -> false
+let isProd c = match kind c with | Prod _ -> true | _ -> false
+let isLambda c = match kind c with | Lambda _ -> true | _ -> false
+let isLetIn c = match kind c with LetIn _ -> true | _ -> false
+let isApp c = match kind c with App _ -> true | _ -> false
+let isConst c = match kind c with Const _ -> true | _ -> false
+let isConstruct c = match kind c with Construct _ -> true | _ -> false
+let isCase c = match kind c with Case _ -> true | _ -> false
+let isProj c = match kind c with Proj _ -> true | _ -> false
+let isFix c = match kind c with Fix _ -> true | _ -> false
+let isCoFix c = match kind c with CoFix _ -> true | _ -> false
+
+(* Destructs a de Bruijn index *)
+let destRel c = match kind c with
+ | Rel n -> n
+ | _ -> raise DestKO
+
+(* Destructs an existential variable *)
+let destMeta c = match kind c with
+ | Meta n -> n
+ | _ -> raise DestKO
+
+(* Destructs a variable *)
+let destVar c = match kind c with
+ | Var id -> id
+ | _ -> raise DestKO
+
+let destSort c = match kind c with
+ | Sort s -> s
+ | _ -> raise DestKO
+
+(* Destructs a casted term *)
+let destCast c = match kind c with
+ | Cast (t1,k,t2) -> (t1,k,t2)
+ | _ -> raise DestKO
+
+(* Destructs the product (x:t1)t2 *)
+let destProd c = match kind c with
+ | Prod (x,t1,t2) -> (x,t1,t2)
+ | _ -> raise DestKO
+
+(* Destructs the abstraction [x:t1]t2 *)
+let destLambda c = match kind c with
+ | Lambda (x,t1,t2) -> (x,t1,t2)
+ | _ -> raise DestKO
+
+(* Destructs the let [x:=b:t1]t2 *)
+let destLetIn c = match kind c with
+ | LetIn (x,b,t1,t2) -> (x,b,t1,t2)
+ | _ -> raise DestKO
+
+(* Destructs an application *)
+let destApp c = match kind c with
+ | App (f,a) -> (f, a)
+ | _ -> raise DestKO
+
+(* Destructs a constant *)
+let destConst c = match kind c with
+ | Const kn -> kn
+ | _ -> raise DestKO
+
+(* Destructs an existential variable *)
+let destEvar c = match kind c with
+ | Evar (kn, a as r) -> r
+ | _ -> raise DestKO
+
+(* Destructs a (co)inductive type named kn *)
+let destInd c = match kind c with
+ | Ind (kn, a as r) -> r
+ | _ -> raise DestKO
+
+(* Destructs a constructor *)
+let destConstruct c = match kind c with
+ | Construct (kn, a as r) -> r
+ | _ -> raise DestKO
+
+(* Destructs a term <p>Case c of lc1 | lc2 .. | lcn end *)
+let destCase c = match kind c with
+ | Case (ci,p,c,v) -> (ci,p,c,v)
+ | _ -> raise DestKO
+
+let destProj c = match kind c with
+ | Proj (p, c) -> (p, c)
+ | _ -> raise DestKO
+
+let destFix c = match kind c with
+ | Fix fix -> fix
+ | _ -> raise DestKO
+
+let destCoFix c = match kind c with
+ | CoFix cofix -> cofix
+ | _ -> raise DestKO
+
+
+(******************************************************************)
+(* Flattening and unflattening of embedded applications and casts *)
+(******************************************************************)
+
+let decompose_app c =
+ match kind c with
+ | App (f,cl) -> (f, Array.to_list cl)
+ | _ -> (c,[])
+
+let decompose_appvect c =
+ match kind c with
+ | App (f,cl) -> (f, cl)
+ | _ -> (c,[||])
+
(****************************************************************************)
(* Functions to recur through subterms *)
(****************************************************************************)
diff --git a/kernel/constr.mli b/kernel/constr.mli
index 474ab3884..4c5ea9e95 100644
--- a/kernel/constr.mli
+++ b/kernel/constr.mli
@@ -225,6 +225,110 @@ type ('constr, 'types, 'sort, 'univs) kind_of_term =
val kind : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term
val of_kind : (constr, types, Sorts.t, Univ.Instance.t) kind_of_term -> constr
+(** {6 Simple case analysis} *)
+val isRel : constr -> bool
+val isRelN : int -> constr -> bool
+val isVar : constr -> bool
+val isVarId : Id.t -> constr -> bool
+val isInd : constr -> bool
+val isEvar : constr -> bool
+val isMeta : constr -> bool
+val isEvar_or_Meta : constr -> bool
+val isSort : constr -> bool
+val isCast : constr -> bool
+val isApp : constr -> bool
+val isLambda : constr -> bool
+val isLetIn : constr -> bool
+val isProd : constr -> bool
+val isConst : constr -> bool
+val isConstruct : constr -> bool
+val isFix : constr -> bool
+val isCoFix : constr -> bool
+val isCase : constr -> bool
+val isProj : constr -> bool
+
+val is_Prop : constr -> bool
+val is_Set : constr -> bool
+val isprop : constr -> bool
+val is_Type : constr -> bool
+val iskind : constr -> bool
+val is_small : Sorts.t -> bool
+
+(** {6 Term destructors } *)
+(** Destructor operations are partial functions and
+ @raise DestKO if the term has not the expected form. *)
+
+exception DestKO
+
+(** Destructs a de Bruijn index *)
+val destRel : constr -> int
+
+(** Destructs an existential variable *)
+val destMeta : constr -> metavariable
+
+(** Destructs a variable *)
+val destVar : constr -> Id.t
+
+(** Destructs a sort. [is_Prop] recognizes the sort {% \textsf{%}Prop{% }%}, whether
+ [isprop] recognizes both {% \textsf{%}Prop{% }%} and {% \textsf{%}Set{% }%}. *)
+val destSort : constr -> Sorts.t
+
+(** Destructs a casted term *)
+val destCast : constr -> constr * cast_kind * constr
+
+(** Destructs the product {% $ %}(x:t_1)t_2{% $ %} *)
+val destProd : types -> Name.t * types * types
+
+(** Destructs the abstraction {% $ %}[x:t_1]t_2{% $ %} *)
+val destLambda : constr -> Name.t * types * constr
+
+(** Destructs the let {% $ %}[x:=b:t_1]t_2{% $ %} *)
+val destLetIn : constr -> Name.t * constr * types * constr
+
+(** Destructs an application *)
+val destApp : constr -> constr * constr array
+
+(** Decompose any term as an applicative term; the list of args can be empty *)
+val decompose_app : constr -> constr * constr list
+
+(** Same as [decompose_app], but returns an array. *)
+val decompose_appvect : constr -> constr * constr array
+
+(** Destructs a constant *)
+val destConst : constr -> Constant.t puniverses
+
+(** Destructs an existential variable *)
+val destEvar : constr -> existential
+
+(** Destructs a (co)inductive type *)
+val destInd : constr -> inductive puniverses
+
+(** Destructs a constructor *)
+val destConstruct : constr -> constructor puniverses
+
+(** Destructs a [match c as x in I args return P with ... |
+Ci(...yij...) => ti | ... end] (or [let (..y1i..) := c as x in I args
+return P in t1], or [if c then t1 else t2])
+@return [(info,c,fun args x => P,[|...|fun yij => ti| ...|])]
+where [info] is pretty-printing information *)
+val destCase : constr -> case_info * constr * constr * constr array
+
+(** Destructs a projection *)
+val destProj : constr -> projection * constr
+
+(** Destructs the {% $ %}i{% $ %}th function of the block
+ [Fixpoint f{_ 1} ctx{_ 1} = b{_ 1}
+ with f{_ 2} ctx{_ 2} = b{_ 2}
+ ...
+ with f{_ n} ctx{_ n} = b{_ n}],
+ where the length of the {% $ %}j{% $ %}th context is {% $ %}ij{% $ %}.
+*)
+val destFix : constr -> fixpoint
+
+val destCoFix : constr -> cofixpoint
+
+(** {6 Equality} *)
+
(** [equal a b] is true if [a] equals [b] modulo alpha, casts,
and application grouping *)
val equal : constr -> constr -> bool
@@ -344,7 +448,7 @@ val compare_head_gen_leq : (bool -> Univ.Instance.t -> Univ.Instance.t -> bool)
(constr -> constr -> bool) ->
(constr -> constr -> bool) ->
constr -> constr -> bool
-
+
(** {6 Hashconsing} *)
val hash : constr -> int
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index f4e611c19..083b0ae40 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -56,7 +56,7 @@ let weaker_noccur_between env x nvars t =
else None
let is_constructor_head t =
- Term.isRel(fst(Term.decompose_app t))
+ isRel(fst(decompose_app t))
(************************************************************************)
(* Various well-formedness check for inductive declarations *)
@@ -135,7 +135,7 @@ let infos_and_sort env t =
| Prod (name,c1,c2) ->
let varj = infer_type env c1 in
let env1 = Environ.push_rel (LocalAssum (name,varj.utj_val)) env in
- let max = Universe.sup max (Term.univ_of_sort varj.utj_type) in
+ let max = Universe.sup max (Sorts.univ_of_sort varj.utj_type) in
aux env1 c2 max
| _ when is_constructor_head t -> max
| _ -> (* don't fail if not positive, it is tested later *) max
@@ -184,7 +184,7 @@ let cumulate_arity_large_levels env sign =
match d with
| LocalAssum (_,t) ->
let tj = infer_type env t in
- let u = Term.univ_of_sort tj.utj_type in
+ let u = Sorts.univ_of_sort tj.utj_type in
(Universe.sup u lev, push_rel d env)
| LocalDef _ ->
lev, push_rel d env)
@@ -351,7 +351,7 @@ let typecheck_inductive env mie =
| None -> clev
in
let full_polymorphic () =
- let defu = Term.univ_of_sort def_level in
+ let defu = Sorts.univ_of_sort def_level in
let is_natural =
type_in_type env || (UGraph.check_leq (universes env') infu defu)
in
@@ -555,7 +555,7 @@ let check_positivity_one ~chkpos recursive (env,_,ntypes,_ as ienv) paramsctxt (
constructor [cn] has a type of the shape [… -> c … -> P], where,
more generally, the arrows may be dependent). *)
let rec check_pos (env, n, ntypes, ra_env as ienv) nmr c =
- let x,largs = Term.decompose_app (whd_all env c) in
+ let x,largs = decompose_app (whd_all env c) in
match kind x with
| Prod (na,b,d) ->
let () = assert (List.is_empty largs) in
@@ -663,7 +663,7 @@ let check_positivity_one ~chkpos recursive (env,_,ntypes,_ as ienv) paramsctxt (
inductive type. *)
and check_constructors ienv check_head nmr c =
let rec check_constr_rec (env,n,ntypes,ra_env as ienv) nmr lrec c =
- let x,largs = Term.decompose_app (whd_all env c) in
+ let x,largs = decompose_app (whd_all env c) in
match kind x with
| Prod (na,b,d) ->
@@ -916,11 +916,11 @@ let build_inductive env prv iu env_ar paramsctxt kn isrecord isfinite inds nmr r
let ar = {template_param_levels = paramlevs; template_level = lev} in
TemplateArity ar, all_sorts
| RegularArity (info,ar,defs) ->
- let s = sort_of_univ defs in
+ let s = Sorts.sort_of_univ defs in
let kelim = allowed_sorts info s in
let ar = RegularArity
{ mind_user_arity = Vars.subst_univs_level_constr substunivs ar;
- mind_sort = sort_of_univ (Univ.subst_univs_level_universe substunivs defs); } in
+ mind_sort = Sorts.sort_of_univ (Univ.subst_univs_level_universe substunivs defs); } in
ar, kelim in
(* Assigning VM tags to constructors *)
let nconst, nblock = ref 0, ref 0 in
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index cb03a4d6b..0782ea820 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -29,20 +29,20 @@ let lookup_mind_specif env (kn,tyi) =
(mib, mib.mind_packets.(tyi))
let find_rectype env c =
- let (t, l) = Term.decompose_app (whd_all env c) in
+ let (t, l) = decompose_app (whd_all env c) in
match kind t with
| Ind ind -> (ind, l)
| _ -> raise Not_found
let find_inductive env c =
- let (t, l) = Term.decompose_app (whd_all env c) in
+ let (t, l) = decompose_app (whd_all env c) in
match kind t with
| Ind ind
when (fst (lookup_mind_specif env (out_punivs ind))).mind_finite <> Decl_kinds.CoFinite -> (ind, l)
| _ -> raise Not_found
let find_coinductive env c =
- let (t, l) = Term.decompose_app (whd_all env c) in
+ let (t, l) = decompose_app (whd_all env c) in
match kind t with
| Ind ind
when (fst (lookup_mind_specif env (out_punivs ind))).mind_finite == Decl_kinds.CoFinite -> (ind, l)
@@ -354,7 +354,7 @@ let build_branches_type (ind,u) (_,mip as specif) params p =
let typi = full_constructor_instantiate (ind,u,specif,params) cty in
let (cstrsign,ccl) = Term.decompose_prod_assum typi in
let nargs = Context.Rel.length cstrsign in
- let (_,allargs) = Term.decompose_app ccl in
+ let (_,allargs) = decompose_app ccl in
let (lparams,vargs) = List.chop (inductive_params specif) allargs in
let cargs =
let cstr = ith_constructor_of_inductive ind (i+1) in
@@ -566,8 +566,8 @@ let check_inductive_codomain env p =
let env = push_rel_context absctx env in
let arctx, s = dest_prod_assum env ar in
let env = push_rel_context arctx env in
- let i,l' = Term.decompose_app (whd_all env s) in
- Term.isInd i
+ let i,l' = decompose_app (whd_all env s) in
+ isInd i
(* The following functions are almost duplicated from indtypes.ml, except
that they carry here a poorer environment (containing less information). *)
@@ -621,7 +621,7 @@ close to check_positive in indtypes.ml, but does no positivity check and does no
compute the number of recursive arguments. *)
let get_recargs_approx env tree ind args =
let rec build_recargs (env, ra_env as ienv) tree c =
- let x,largs = Term.decompose_app (whd_all env c) in
+ let x,largs = decompose_app (whd_all env c) in
match kind x with
| Prod (na,b,d) ->
assert (List.is_empty largs);
@@ -680,7 +680,7 @@ let get_recargs_approx env tree ind args =
and build_recargs_constructors ienv trees c =
let rec recargs_constr_rec (env,ra_env as ienv) trees lrec c =
- let x,largs = Term.decompose_app (whd_all env c) in
+ let x,largs = decompose_app (whd_all env c) in
match kind x with
| Prod (na,b,d) ->
@@ -709,7 +709,7 @@ let restrict_spec env spec p =
let env = push_rel_context absctx env in
let arctx, s = dest_prod_assum env ar in
let env = push_rel_context arctx env in
- let i,args = Term.decompose_app (whd_all env s) in
+ let i,args = decompose_app (whd_all env s) in
match kind i with
| Ind i ->
begin match spec with
@@ -730,7 +730,7 @@ let restrict_spec env spec p =
let rec subterm_specif renv stack t =
(* maybe reduction is not always necessary! *)
- let f,l = Term.decompose_app (whd_all renv.env t) in
+ let f,l = decompose_app (whd_all renv.env t) in
match kind f with
| Rel k -> subterm_var k renv
| Case (ci,p,c,lbr) ->
@@ -863,7 +863,7 @@ let filter_stack_domain env ci p stack =
let d = LocalAssum (n,a) in
let ctx, a = dest_prod_assum env a in
let env = push_rel_context ctx env in
- let ty, args = Term.decompose_app (whd_all env a) in
+ let ty, args = decompose_app (whd_all env a) in
let elt = match kind ty with
| Ind ind ->
let spec' = stack_element_specif elt in
@@ -894,7 +894,7 @@ let check_one_fix renv recpos trees def =
(* if [t] does not make recursive calls, it is guarded: *)
if noccur_with_meta renv.rel_min nfi t then ()
else
- let (f,l) = Term.decompose_app (whd_betaiotazeta renv.env t) in
+ let (f,l) = decompose_app (whd_betaiotazeta renv.env t) in
match kind f with
| Rel p ->
(* Test if [p] is a fixpoint (recursive call) *)
@@ -1120,7 +1120,7 @@ let rec codomain_is_coind env c =
let check_one_cofix env nbfix def deftype =
let rec check_rec_call env alreadygrd n tree vlra t =
if not (noccur_with_meta n nbfix t) then
- let c,args = Term.decompose_app (whd_all env t) in
+ let c,args = decompose_app (whd_all env t) in
match kind c with
| Rel p when n <= p && p < n+nbfix ->
(* recursive call: must be guarded and no nested recursive
diff --git a/kernel/modops.ml b/kernel/modops.ml
index b1df1a187..11e6be659 100644
--- a/kernel/modops.ml
+++ b/kernel/modops.ml
@@ -266,9 +266,9 @@ let subst_structure subst = subst_structure subst do_delta_codom
(* lclrk : retroknowledge_action list, rkaction : retroknowledge action *)
let add_retroknowledge mp =
let perform rkaction env = match rkaction with
- |Retroknowledge.RKRegister (f, e) when (Term.isConst e || Term.isInd e) ->
+ | Retroknowledge.RKRegister (f, e) when (isConst e || isInd e) ->
Environ.register env f e
- |_ ->
+ | _ ->
CErrors.anomaly ~label:"Modops.add_retroknowledge"
(Pp.str "had to import an unsupported kind of term.")
in
diff --git a/kernel/names.ml b/kernel/names.ml
index cb27104d1..b02c0b840 100644
--- a/kernel/names.ml
+++ b/kernel/names.ml
@@ -179,6 +179,8 @@ struct
| [] -> "<>"
| sl -> String.concat "." (List.rev_map Id.to_string sl)
+ let print dp = str (to_string dp)
+
let initial = [default_module_name]
module Hdir = Hashcons.Hlist(Id)
diff --git a/kernel/names.mli b/kernel/names.mli
index ba0637c8a..709ebeb7f 100644
--- a/kernel/names.mli
+++ b/kernel/names.mli
@@ -159,6 +159,7 @@ sig
val hcons : t -> t
(** Hashconsing of directory paths. *)
+ val print : t -> Pp.t
end
(** {6 Names of structure elements } *)
diff --git a/kernel/term.ml b/kernel/term.ml
index 1c970867a..4217cfac7 100644
--- a/kernel/term.ml
+++ b/kernel/term.ml
@@ -11,6 +11,7 @@ open Pp
open CErrors
open Names
open Vars
+open Constr
(**********************************************************************)
(** Redeclaration of types from module Constr *)
@@ -165,167 +166,52 @@ let hcons_types = Constr.hcons
(* Non primitive term destructors *)
(**********************************************************************)
-(* Destructor operations : partial functions
- Raise [DestKO] if the const has not the expected form *)
-
-exception DestKO
-
+exception DestKO = DestKO
(* Destructs a de Bruijn index *)
-let destRel c = match kind_of_term c with
- | Rel n -> n
- | _ -> raise DestKO
-
-(* Destructs an existential variable *)
-let destMeta c = match kind_of_term c with
- | Meta n -> n
- | _ -> raise DestKO
-
-let isMeta c = match kind_of_term c with Meta _ -> true | _ -> false
-
-(* Destructs a variable *)
-let destVar c = match kind_of_term c with
- | Var id -> id
- | _ -> raise DestKO
-
-(* Destructs a type *)
-let isSort c = match kind_of_term c with
- | Sort _ -> true
- | _ -> false
-
-let destSort c = match kind_of_term c with
- | Sort s -> s
- | _ -> raise DestKO
-
-let rec isprop c = match kind_of_term c with
- | Sort (Prop _) -> true
- | Cast (c,_,_) -> isprop c
- | _ -> false
-
-let rec is_Prop c = match kind_of_term c with
- | Sort (Prop Null) -> true
- | Cast (c,_,_) -> is_Prop c
- | _ -> false
-
-let rec is_Set c = match kind_of_term c with
- | Sort (Prop Pos) -> true
- | Cast (c,_,_) -> is_Set c
- | _ -> false
-
-let rec is_Type c = match kind_of_term c with
- | Sort (Type _) -> true
- | Cast (c,_,_) -> is_Type c
- | _ -> false
-
-let is_small = Sorts.is_small
-
-let iskind c = isprop c || is_Type c
-
-(* Tests if an evar *)
-let isEvar c = match kind_of_term c with Evar _ -> true | _ -> false
-
-let isEvar_or_Meta c = match kind_of_term c with
- | Evar _ | Meta _ -> true
- | _ -> false
-
-(* Destructs a casted term *)
-let destCast c = match kind_of_term c with
- | Cast (t1,k,t2) -> (t1,k,t2)
- | _ -> raise DestKO
-
-let isCast c = match kind_of_term c with Cast _ -> true | _ -> false
-
-
-(* Tests if a de Bruijn index *)
-let isRel c = match kind_of_term c with Rel _ -> true | _ -> false
-let isRelN n c =
- match kind_of_term c with Rel n' -> Int.equal n n' | _ -> false
-
-(* Tests if a variable *)
-let isVar c = match kind_of_term c with Var _ -> true | _ -> false
-let isVarId id c =
- match kind_of_term c with Var id' -> Id.equal id id' | _ -> false
-
-(* Tests if an inductive *)
-let isInd c = match kind_of_term c with Ind _ -> true | _ -> false
-
-(* Destructs the product (x:t1)t2 *)
-let destProd c = match kind_of_term c with
- | Prod (x,t1,t2) -> (x,t1,t2)
- | _ -> raise DestKO
-
-let isProd c = match kind_of_term c with | Prod _ -> true | _ -> false
-
-(* Destructs the abstraction [x:t1]t2 *)
-let destLambda c = match kind_of_term c with
- | Lambda (x,t1,t2) -> (x,t1,t2)
- | _ -> raise DestKO
-
-let isLambda c = match kind_of_term c with | Lambda _ -> true | _ -> false
-
-(* Destructs the let [x:=b:t1]t2 *)
-let destLetIn c = match kind_of_term c with
- | LetIn (x,b,t1,t2) -> (x,b,t1,t2)
- | _ -> raise DestKO
-
-let isLetIn c = match kind_of_term c with LetIn _ -> true | _ -> false
-
-(* Destructs an application *)
-let destApp c = match kind_of_term c with
- | App (f,a) -> (f, a)
- | _ -> raise DestKO
-
+let destRel = destRel
+let destMeta = destRel
+let isMeta = isMeta
+let destVar = destVar
+let isSort = isSort
+let destSort = destSort
+let isprop = isprop
+let is_Prop = is_Prop
+let is_Set = is_Set
+let is_Type = is_Type
+let is_small = is_small
+let iskind = iskind
+let isEvar = isEvar
+let isEvar_or_Meta = isEvar_or_Meta
+let destCast = destCast
+let isCast = isCast
+let isRel = isRel
+let isRelN = isRelN
+let isVar = isVar
+let isVarId = isVarId
+let isInd = isInd
+let destProd = destProd
+let isProd = isProd
+let destLambda = destLambda
+let isLambda = isLambda
+let destLetIn = destLetIn
+let isLetIn = isLetIn
+let destApp = destApp
let destApplication = destApp
-
-let isApp c = match kind_of_term c with App _ -> true | _ -> false
-
-(* Destructs a constant *)
-let destConst c = match kind_of_term c with
- | Const kn -> kn
- | _ -> raise DestKO
-
-let isConst c = match kind_of_term c with Const _ -> true | _ -> false
-
-(* Destructs an existential variable *)
-let destEvar c = match kind_of_term c with
- | Evar (kn, a as r) -> r
- | _ -> raise DestKO
-
-(* Destructs a (co)inductive type named kn *)
-let destInd c = match kind_of_term c with
- | Ind (kn, a as r) -> r
- | _ -> raise DestKO
-
-(* Destructs a constructor *)
-let destConstruct c = match kind_of_term c with
- | Construct (kn, a as r) -> r
- | _ -> raise DestKO
-
-let isConstruct c = match kind_of_term c with Construct _ -> true | _ -> false
-
-(* Destructs a term <p>Case c of lc1 | lc2 .. | lcn end *)
-let destCase c = match kind_of_term c with
- | Case (ci,p,c,v) -> (ci,p,c,v)
- | _ -> raise DestKO
-
-let isCase c = match kind_of_term c with Case _ -> true | _ -> false
-
-let isProj c = match kind_of_term c with Proj _ -> true | _ -> false
-
-let destProj c = match kind_of_term c with
- | Proj (p, c) -> (p, c)
- | _ -> raise DestKO
-
-let destFix c = match kind_of_term c with
- | Fix fix -> fix
- | _ -> raise DestKO
-
-let isFix c = match kind_of_term c with Fix _ -> true | _ -> false
-
-let destCoFix c = match kind_of_term c with
- | CoFix cofix -> cofix
- | _ -> raise DestKO
-
-let isCoFix c = match kind_of_term c with CoFix _ -> true | _ -> false
+let isApp = isApp
+let destConst = destConst
+let isConst = isConst
+let destEvar = destEvar
+let destInd = destInd
+let destConstruct = destConstruct
+let isConstruct = isConstruct
+let destCase = destCase
+let isCase = isCase
+let isProj = isProj
+let destProj = destProj
+let destFix = destFix
+let isFix = isFix
+let destCoFix = destCoFix
+let isCoFix = isCoFix
(******************************************************************)
(* Flattening and unflattening of embedded applications and casts *)
diff --git a/kernel/term.mli b/kernel/term.mli
index 33c6b0b08..4efb582d0 100644
--- a/kernel/term.mli
+++ b/kernel/term.mli
@@ -16,90 +16,133 @@ open Constr
*)
+exception DestKO
+[@@ocaml.deprecated "Alias for [Constr.DestKO]"]
+
(** {5 Simple term case analysis. } *)
val isRel : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isRel]"]
val isRelN : int -> constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isRelN]"]
val isVar : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isVar]"]
val isVarId : Id.t -> constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isVarId]"]
val isInd : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isInd]"]
val isEvar : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isEvar]"]
val isMeta : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isMeta]"]
val isEvar_or_Meta : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isEvar_or_Meta]"]
val isSort : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isSort]"]
val isCast : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isCast]"]
val isApp : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isApp]"]
val isLambda : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isLambda]"]
val isLetIn : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isletIn]"]
val isProd : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isProp]"]
val isConst : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isConst]"]
val isConstruct : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isConstruct]"]
val isFix : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isFix]"]
val isCoFix : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isCoFix]"]
val isCase : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isCase]"]
val isProj : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isProj]"]
val is_Prop : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.is_Prop]"]
val is_Set : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.is_Set]"]
val isprop : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.isprop]"]
val is_Type : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.is_Type]"]
val iskind : constr -> bool
+[@@ocaml.deprecated "Alias for [Constr.is_kind]"]
val is_small : Sorts.t -> bool
+[@@ocaml.deprecated "Alias for [Constr.is_small]"]
(** {5 Term destructors } *)
(** Destructor operations are partial functions and
@raise DestKO if the term has not the expected form. *)
-exception DestKO
-
(** Destructs a de Bruijn index *)
val destRel : constr -> int
+[@@ocaml.deprecated "Alias for [Constr.destRel]"]
(** Destructs an existential variable *)
val destMeta : constr -> metavariable
+[@@ocaml.deprecated "Alias for [Constr.destMeta]"]
(** Destructs a variable *)
val destVar : constr -> Id.t
+[@@ocaml.deprecated "Alias for [Constr.destVar]"]
(** Destructs a sort. [is_Prop] recognizes the sort {% \textsf{%}Prop{% }%}, whether
[isprop] recognizes both {% \textsf{%}Prop{% }%} and {% \textsf{%}Set{% }%}. *)
val destSort : constr -> Sorts.t
+[@@ocaml.deprecated "Alias for [Constr.destSort]"]
(** Destructs a casted term *)
val destCast : constr -> constr * cast_kind * constr
+[@@ocaml.deprecated "Alias for [Constr.destCast]"]
(** Destructs the product {% $ %}(x:t_1)t_2{% $ %} *)
val destProd : types -> Name.t * types * types
+[@@ocaml.deprecated "Alias for [Constr.destProd]"]
(** Destructs the abstraction {% $ %}[x:t_1]t_2{% $ %} *)
val destLambda : constr -> Name.t * types * constr
+[@@ocaml.deprecated "Alias for [Constr.destLambda]"]
(** Destructs the let {% $ %}[x:=b:t_1]t_2{% $ %} *)
val destLetIn : constr -> Name.t * constr * types * constr
+[@@ocaml.deprecated "Alias for [Constr.destLetIn]"]
(** Destructs an application *)
val destApp : constr -> constr * constr array
+[@@ocaml.deprecated "Alias for [Constr.destApp]"]
(** Obsolete synonym of destApp *)
val destApplication : constr -> constr * constr array
+[@@ocaml.deprecated "Alias for [Constr.destApplication]"]
(** Decompose any term as an applicative term; the list of args can be empty *)
val decompose_app : constr -> constr * constr list
+[@@ocaml.deprecated "Alias for [Constr.decompose_app]"]
(** Same as [decompose_app], but returns an array. *)
val decompose_appvect : constr -> constr * constr array
+[@@ocaml.deprecated "Alias for [Constr.decompose_appvect]"]
(** Destructs a constant *)
val destConst : constr -> Constant.t puniverses
+[@@ocaml.deprecated "Alias for [Constr.destConst]"]
(** Destructs an existential variable *)
val destEvar : constr -> existential
+[@@ocaml.deprecated "Alias for [Constr.destEvar]"]
(** Destructs a (co)inductive type *)
val destInd : constr -> inductive puniverses
+[@@ocaml.deprecated "Alias for [Constr.destInd]"]
(** Destructs a constructor *)
val destConstruct : constr -> constructor puniverses
+[@@ocaml.deprecated "Alias for [Constr.destConstruct]"]
(** Destructs a [match c as x in I args return P with ... |
Ci(...yij...) => ti | ... end] (or [let (..y1i..) := c as x in I args
@@ -107,9 +150,11 @@ return P in t1], or [if c then t1 else t2])
@return [(info,c,fun args x => P,[|...|fun yij => ti| ...|])]
where [info] is pretty-printing information *)
val destCase : constr -> case_info * constr * constr * constr array
+[@@ocaml.deprecated "Alias for [Constr.destCase]"]
(** Destructs a projection *)
val destProj : constr -> projection * constr
+[@@ocaml.deprecated "Alias for [Constr.destProj]"]
(** Destructs the {% $ %}i{% $ %}th function of the block
[Fixpoint f{_ 1} ctx{_ 1} = b{_ 1}
@@ -119,8 +164,10 @@ val destProj : constr -> projection * constr
where the length of the {% $ %}j{% $ %}th context is {% $ %}ij{% $ %}.
*)
val destFix : constr -> fixpoint
+[@@ocaml.deprecated "Alias for [Constr.destFix]"]
val destCoFix : constr -> cofixpoint
+[@@ocaml.deprecated "Alias for [Constr.destCoFix]"]
(** {5 Derived constructors} *)
@@ -415,8 +462,11 @@ val map_constr_with_binders :
[@@ocaml.deprecated "Alias for [Constr.map_with_binders]"]
val map_puniverses : ('a -> 'b) -> 'a puniverses -> 'b puniverses
+[@@ocaml.deprecated "Alias for [Constr.map_puniverses]"]
val univ_of_sort : Sorts.t -> Univ.Universe.t
+[@@ocaml.deprecated "Alias for [Sorts.univ_of_sort]"]
val sort_of_univ : Univ.Universe.t -> Sorts.t
+[@@ocaml.deprecated "Alias for [Sorts.sort_of_univ]"]
val iter_constr : (constr -> unit) -> constr -> unit
[@@ocaml.deprecated "Alias for [Constr.iter]"]
diff --git a/library/coqlib.ml b/library/coqlib.ml
index 141fff033..4a2390985 100644
--- a/library/coqlib.ml
+++ b/library/coqlib.ml
@@ -14,7 +14,7 @@ open Libnames
open Globnames
open Nametab
-let coq = Nameops.coq_string (* "Coq" *)
+let coq = Libnames.coq_string (* "Coq" *)
(************************************************************************)
(* Generic functions to find Coq objects *)
@@ -32,7 +32,7 @@ let find_reference locstr dir s =
of not found errors here *)
user_err ~hdr:locstr
Pp.(str "cannot find " ++ Libnames.pr_path sp ++
- str "; maybe library " ++ Libnames.pr_dirpath dp ++
+ str "; maybe library " ++ DirPath.print dp ++
str " has to be required first.")
let coq_reference locstr dir s = find_reference locstr (coq::dir) s
@@ -52,14 +52,14 @@ let gen_reference_in_modules locstr dirs s =
| [] ->
anomaly ~label:locstr (str "cannot find " ++ str s ++
str " in module" ++ str (if List.length dirs > 1 then "s " else " ") ++
- prlist_with_sep pr_comma pr_dirpath dirs ++ str ".")
+ prlist_with_sep pr_comma DirPath.print dirs ++ str ".")
| l ->
anomaly ~label:locstr
(str "ambiguous name " ++ str s ++ str " can represent " ++
prlist_with_sep pr_comma
(fun x -> Libnames.pr_path (Nametab.path_of_global x)) l ++
str " in module" ++ str (if List.length dirs > 1 then "s " else " ") ++
- prlist_with_sep pr_comma pr_dirpath dirs ++ str ".")
+ prlist_with_sep pr_comma DirPath.print dirs ++ str ".")
(* For tactics/commands requiring vernacular libraries *)
@@ -79,7 +79,7 @@ let check_required_library d =
*)
(* or failing ...*)
user_err ~hdr:"Coqlib.check_required_library"
- (str "Library " ++ pr_dirpath dir ++ str " has to be required first.")
+ (str "Library " ++ DirPath.print dir ++ str " has to be required first.")
(************************************************************************)
(* Specific Coq objects *)
diff --git a/library/declaremods.ml b/library/declaremods.ml
index cda40f49f..db83dafef 100644
--- a/library/declaremods.ml
+++ b/library/declaremods.ml
@@ -167,13 +167,13 @@ let consistency_checks exists dir dirinfo =
try Nametab.locate_dir (qualid_of_dirpath dir)
with Not_found ->
user_err ~hdr:"consistency_checks"
- (pr_dirpath dir ++ str " should already exist!")
+ (DirPath.print dir ++ str " should already exist!")
in
assert (eq_global_dir_reference globref dirinfo)
else
if Nametab.exists_dir dir then
user_err ~hdr:"consistency_checks"
- (pr_dirpath dir ++ str " already exists")
+ (DirPath.print dir ++ str " already exists")
let compute_visibility exists i =
if exists then Nametab.Exactly i else Nametab.Until i
diff --git a/library/heads.ml b/library/heads.ml
index 8b8e407f7..ee3bfe1bd 100644
--- a/library/heads.ml
+++ b/library/heads.ml
@@ -8,7 +8,6 @@
open Util
open Names
-open Term
open Constr
open Vars
open Mod_subst
diff --git a/library/lib.ml b/library/lib.ml
index 36292d367..3dbb16c7b 100644
--- a/library/lib.ml
+++ b/library/lib.ml
@@ -12,7 +12,6 @@ open Util
open Names
open Libnames
open Globnames
-open Nameops
open Libobject
open Context.Named.Declaration
@@ -361,8 +360,8 @@ let end_compilation_checks dir =
| None -> anomaly (Pp.str "There should be a module name...")
| Some m ->
if not (Names.DirPath.equal m dir) then anomaly
- (str "The current open module has name" ++ spc () ++ pr_dirpath m ++
- spc () ++ str "and not" ++ spc () ++ pr_dirpath m ++ str ".");
+ (str "The current open module has name" ++ spc () ++ DirPath.print m ++
+ spc () ++ str "and not" ++ spc () ++ DirPath.print m ++ str ".");
in
oname
diff --git a/library/libnames.ml b/library/libnames.ml
index efb1348ab..81878e72f 100644
--- a/library/libnames.ml
+++ b/library/libnames.ml
@@ -13,7 +13,7 @@ open Names
(**********************************************)
-let pr_dirpath sl = str (DirPath.to_string sl)
+let pr_dirpath sl = DirPath.print sl
(*s Operations on dirpaths *)
@@ -232,6 +232,14 @@ let join_reference ns r =
Qualid (loc, make_qualid
(dirpath_of_string (Names.Id.to_string id1)) id2)
+(* Default paths *)
+let default_library = Names.DirPath.initial (* = ["Top"] *)
+
+(*s Roots of the space of absolute names *)
+let coq_string = "Coq"
+let coq_root = Id.of_string coq_string
+let default_root_prefix = DirPath.empty
+
(* Deprecated synonyms *)
let make_short_qualid = qualid_of_ident
diff --git a/library/libnames.mli b/library/libnames.mli
index ab2585334..ed01163ee 100644
--- a/library/libnames.mli
+++ b/library/libnames.mli
@@ -11,12 +11,13 @@ open Loc
open Names
(** {6 Dirpaths } *)
-(** FIXME: ought to be in Names.dir_path *)
+val dirpath_of_string : string -> DirPath.t
val pr_dirpath : DirPath.t -> Pp.t
+[@@ocaml.deprecated "Alias for DirPath.print"]
-val dirpath_of_string : string -> DirPath.t
val string_of_dirpath : DirPath.t -> string
+[@@ocaml.deprecated "Alias for DirPath.to_string"]
(** Pop the suffix of a [DirPath.t]. Raises a [Failure] for an empty path *)
val pop_dirpath : DirPath.t -> DirPath.t
@@ -127,7 +128,20 @@ val pr_reference : reference -> Pp.t
val loc_of_reference : reference -> Loc.t option
val join_reference : reference -> reference -> reference
-(** Deprecated synonyms *)
+(** some preset paths *)
+val default_library : DirPath.t
+
+(** This is the root of the standard library of Coq *)
+val coq_root : module_ident (** "Coq" *)
+val coq_string : string (** "Coq" *)
+(** This is the default root prefix for developments which doesn't
+ mention a root *)
+val default_root_prefix : DirPath.t
+
+(** Deprecated synonyms *)
val make_short_qualid : Id.t -> qualid (** = qualid_of_ident *)
+[@@ocaml.deprecated "Alias for qualid_of_ident"]
+
val qualid_of_sp : full_path -> qualid (** = qualid_of_path *)
+[@@ocaml.deprecated "Alias for qualid_of_sp"]
diff --git a/library/library.ml b/library/library.ml
index 99ef66699..88470d121 100644
--- a/library/library.ml
+++ b/library/library.ml
@@ -12,9 +12,8 @@ open Util
open Names
open Libnames
-open Nameops
-open Libobject
open Lib
+open Libobject
(************************************************************************)
(*s Low-level interning/externing of libraries to files *)
@@ -132,7 +131,7 @@ let try_find_library dir =
try find_library dir
with Not_found ->
user_err ~hdr:"Library.find_library"
- (str "Unknown library " ++ pr_dirpath dir)
+ (str "Unknown library " ++ DirPath.print dir)
let register_library_filename dir f =
(* Not synchronized: overwrite the previous binding if one existed *)
@@ -331,7 +330,7 @@ let error_unmapped_dir qid =
let prefix, _ = repr_qualid qid in
user_err ~hdr:"load_absolute_library_from"
(str "Cannot load " ++ pr_qualid qid ++ str ":" ++ spc () ++
- str "no physical path bound to" ++ spc () ++ pr_dirpath prefix ++ fnl ())
+ str "no physical path bound to" ++ spc () ++ DirPath.print prefix ++ fnl ())
let error_lib_not_found qid =
user_err ~hdr:"load_absolute_library_from"
@@ -465,8 +464,8 @@ let rec intern_library (needed, contents) (dir, f) from =
if not (DirPath.equal dir m.library_name) then
user_err ~hdr:"load_physical_library"
(str "The file " ++ str f ++ str " contains library" ++ spc () ++
- pr_dirpath m.library_name ++ spc () ++ str "and not library" ++
- spc() ++ pr_dirpath dir);
+ DirPath.print m.library_name ++ spc () ++ str "and not library" ++
+ spc() ++ DirPath.print dir);
Feedback.feedback (Feedback.FileLoaded(DirPath.to_string dir, f));
m.library_digests, intern_library_deps (needed, contents) dir m f
@@ -477,9 +476,9 @@ and intern_library_deps libs dir m from =
and intern_mandatory_library caller from libs (dir,d) =
let digest, libs = intern_library libs (dir, None) (Some from) in
if not (Safe_typing.digest_match ~actual:digest ~required:d) then
- user_err (str "Compiled library " ++ pr_dirpath caller ++
+ user_err (str "Compiled library " ++ DirPath.print caller ++
str " (in file " ++ str from ++ str ") makes inconsistent assumptions \
- over library " ++ pr_dirpath dir);
+ over library " ++ DirPath.print dir);
libs
let rec_intern_library libs (dir, f) =
@@ -617,7 +616,7 @@ let check_coq_overwriting p id =
let is_empty = match l with [] -> true | _ -> false in
if not !Flags.boot && not is_empty && Id.equal (List.last l) coq_root then
user_err
- (str "Cannot build module " ++ pr_dirpath p ++ str "." ++ Id.print id ++ str "." ++ spc () ++
+ (str "Cannot build module " ++ DirPath.print p ++ str "." ++ Id.print id ++ str "." ++ spc () ++
str "it starts with prefix \"Coq\" which is reserved for the Coq library.")
let start_library fo =
@@ -625,7 +624,7 @@ let start_library fo =
try
let lp = Loadpath.find_load_path (Filename.dirname fo) in
Loadpath.logical lp
- with Not_found -> Nameops.default_root_prefix
+ with Not_found -> Libnames.default_root_prefix
in
let file = Filename.chop_extension (Filename.basename fo) in
let id = Id.of_string file in
@@ -665,7 +664,7 @@ let current_reexports () = !libraries_exports_list
let error_recursively_dependent_library dir =
user_err
- (strbrk "Unable to use logical name " ++ pr_dirpath dir ++
+ (strbrk "Unable to use logical name " ++ DirPath.print dir ++
strbrk " to save current library because" ++
strbrk " it already depends on a library of this name.")
diff --git a/library/library.mllib b/library/library.mllib
index d94fc2291..e43bfb5a1 100644
--- a/library/library.mllib
+++ b/library/library.mllib
@@ -1,5 +1,3 @@
-Univops
-Nameops
Libnames
Globnames
Libobject
diff --git a/library/loadpath.ml b/library/loadpath.ml
index 757e972b1..eb6dae84a 100644
--- a/library/loadpath.ml
+++ b/library/loadpath.ml
@@ -54,8 +54,8 @@ let warn_overriding_logical_loadpath =
CWarnings.create ~name:"overriding-logical-loadpath" ~category:"loadpath"
(fun (phys_path, old_path, coq_path) ->
str phys_path ++ strbrk " was previously bound to " ++
- pr_dirpath old_path ++ strbrk "; it is remapped to " ++
- pr_dirpath coq_path)
+ DirPath.print old_path ++ strbrk "; it is remapped to " ++
+ DirPath.print coq_path)
let add_load_path phys_path coq_path ~implicit =
let phys_path = CUnix.canonical_path_name phys_path in
@@ -75,7 +75,7 @@ let add_load_path phys_path coq_path ~implicit =
else
let () =
(* Do not warn when overriding the default "-I ." path *)
- if not (DirPath.equal old_path Nameops.default_root_prefix) then
+ if not (DirPath.equal old_path Libnames.default_root_prefix) then
warn_overriding_logical_loadpath (phys_path, old_path, coq_path)
in
true in
diff --git a/parsing/egramcoq.ml b/parsing/egramcoq.ml
index 7f50fd22a..2cb7da569 100644
--- a/parsing/egramcoq.ml
+++ b/parsing/egramcoq.ml
@@ -259,9 +259,11 @@ let is_binder_level from e = match e with
| (NumLevel 200, (BorderProd (Right, _) | InternalProd)) -> from = 200
| _ -> false
-let make_sep_rules tkl =
- let rec mkrule : Tok.t list -> unit rules = function
- | [] -> Rules ({ norec_rule = Stop }, ignore)
+let make_sep_rules = function
+ | [tk] -> Atoken tk
+ | tkl ->
+ let rec mkrule : Tok.t list -> string rules = function
+ | [] -> Rules ({ norec_rule = Stop }, fun _ -> (* dropped anyway: *) "")
| tkn :: rem ->
let Rules ({ norec_rule = r }, f) = mkrule rem in
let r = { norec_rule = Next (r, Atoken tkn) } in
diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml
index faabd7c14..ccef9ab96 100644
--- a/plugins/cc/ccalgo.ml
+++ b/plugins/cc/ccalgo.ml
@@ -12,13 +12,13 @@
open CErrors
open Pp
-open Goptions
open Names
-open Term
+open Sorts
open Constr
open Vars
-open Tacmach
open Evd
+open Goptions
+open Tacmach
open Util
let init_size=5
@@ -437,7 +437,7 @@ and make_app l=function
and applist_proj c l =
match c with
| Symb s -> applist_projection s l
- | _ -> applistc (constr_of_term c) l
+ | _ -> Term.applistc (constr_of_term c) l
and applist_projection c l =
match Constr.kind c with
| Const c when Environ.is_projection (fst c) (Global.env()) ->
@@ -447,10 +447,10 @@ and applist_projection c l =
let ty = Typeops.type_of_constant_in (Global.env ()) c in (* FIXME constraints *)
let pb = Environ.lookup_projection p (Global.env()) in
let ctx,_ = Term.decompose_prod_n_assum (pb.Declarations.proj_npars + 1) ty in
- it_mkLambda_or_LetIn (mkProj(p,mkRel 1)) ctx
+ Term.it_mkLambda_or_LetIn (mkProj(p,mkRel 1)) ctx
| hd :: tl ->
- applistc (mkProj (p, hd)) tl)
- | _ -> applistc c l
+ Term.applistc (mkProj (p, hd)) tl)
+ | _ -> Term.applistc c l
let rec canonize_name sigma c =
let c = EConstr.Unsafe.to_constr c in
@@ -838,7 +838,7 @@ let complete_one_class state i=
let _args =
List.map (fun i -> constr_of_term (term state.uf i))
pac.args in
- let typ = prod_applist _c (List.rev _args) in
+ let typ = Term.prod_applist _c (List.rev _args) in
let ct = app (term state.uf i) typ pac.arity in
state.uf.epsilons <- pac :: state.uf.epsilons;
ignore (add_term state ct)
diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml
index 7dec34d4d..8642df684 100644
--- a/plugins/cc/cctac.ml
+++ b/plugins/cc/cctac.ml
@@ -12,7 +12,6 @@ open Evd
open Names
open Inductiveops
open Declarations
-open Term
open Constr
open EConstr
open Vars
diff --git a/plugins/extraction/table.ml b/plugins/extraction/table.ml
index 995d5fd19..5903733a6 100644
--- a/plugins/extraction/table.ml
+++ b/plugins/extraction/table.ml
@@ -486,7 +486,7 @@ let check_loaded_modfile mp = match base_mp mp with
if not (Library.library_is_loaded dp) then begin
match base_mp (Lib.current_mp ()) with
| MPfile dp' when not (DirPath.equal dp dp') ->
- err (str "Please load library " ++ pr_dirpath dp ++ str " first.")
+ err (str "Please load library " ++ DirPath.print dp ++ str " first.")
| _ -> ()
end
| _ -> ()
diff --git a/plugins/firstorder/g_ground.ml4 b/plugins/firstorder/g_ground.ml4
index 1e7da3250..938bec25b 100644
--- a/plugins/firstorder/g_ground.ml4
+++ b/plugins/firstorder/g_ground.ml4
@@ -65,11 +65,14 @@ let default_intuition_tac =
let (set_default_solver, default_solver, print_default_solver) =
Tactic_option.declare_tactic_option ~default:default_intuition_tac "Firstorder default solver"
-VERNAC COMMAND EXTEND Firstorder_Set_Solver CLASSIFIED AS SIDEFF
+VERNAC COMMAND FUNCTIONAL EXTEND Firstorder_Set_Solver CLASSIFIED AS SIDEFF
| [ "Set" "Firstorder" "Solver" tactic(t) ] -> [
- set_default_solver
- (Locality.make_section_locality (Locality.LocalityFixme.consume ()))
- (Tacintern.glob_tactic t) ]
+ fun ~atts ~st -> let open Vernacinterp in
+ set_default_solver
+ (Locality.make_section_locality atts.locality)
+ (Tacintern.glob_tactic t);
+ st
+ ]
END
VERNAC COMMAND EXTEND Firstorder_Print_Solver CLASSIFIED AS QUERY
diff --git a/plugins/firstorder/ground.ml b/plugins/firstorder/ground.ml
index f660ba734..d46201335 100644
--- a/plugins/firstorder/ground.ml
+++ b/plugins/firstorder/ground.ml
@@ -11,7 +11,7 @@ open Formula
open Sequent
open Rules
open Instances
-open Term
+open Constr
open Tacmach.New
open Tacticals.New
diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml
index d6309b057..1a6eba8c6 100644
--- a/plugins/firstorder/rules.ml
+++ b/plugins/firstorder/rules.ml
@@ -235,8 +235,8 @@ let constant str = Universes.constr_of_global
@@ Coqlib.coq_reference "User" ["Init";"Logic"] str
let defined_connectives=lazy
- [AllOccurrences,EvalConstRef (fst (Term.destConst (constant "not")));
- AllOccurrences,EvalConstRef (fst (Term.destConst (constant "iff")))]
+ [AllOccurrences,EvalConstRef (fst (Constr.destConst (constant "not")));
+ AllOccurrences,EvalConstRef (fst (Constr.destConst (constant "iff")))]
let normalize_evaluables=
Proofview.Goal.enter begin fun gl ->
diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml
index 29e824f44..62ca70626 100644
--- a/plugins/funind/functional_principles_proofs.ml
+++ b/plugins/funind/functional_principles_proofs.ml
@@ -1,7 +1,7 @@
open Printer
open CErrors
open Util
-open Term
+open Constr
open EConstr
open Vars
open Namegen
diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml
index 3899bc709..996e2b6af 100644
--- a/plugins/funind/functional_principles_types.ml
+++ b/plugins/funind/functional_principles_types.ml
@@ -1,7 +1,8 @@
open Printer
open CErrors
-open Util
open Term
+open Sorts
+open Util
open Constr
open Vars
open Namegen
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index d04b7a33d..fa4353630 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -1,7 +1,6 @@
open Printer
open Pp
open Names
-open Term
open Constr
open Vars
open Glob_term
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index f01b6669d..9e22ad306 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -1,7 +1,8 @@
open CErrors
+open Sorts
open Util
open Names
-open Term
+open Constr
open EConstr
open Pp
open Indfun_common
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index 04d729b10..3089ec470 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -9,7 +9,6 @@
module CVars = Vars
-open Term
open Constr
open EConstr
open Vars
diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4
index 71db919ef..d6cfa3cf9 100644
--- a/plugins/ltac/extratactics.ml4
+++ b/plugins/ltac/extratactics.ml4
@@ -320,24 +320,44 @@ let project_hint pri l2r r =
let info = {Vernacexpr.hint_priority = pri; hint_pattern = None} in
(info,false,true,Hints.PathAny, Hints.IsGlobRef (Globnames.ConstRef c))
-let add_hints_iff l2r lc n bl =
- let l = Locality.LocalityFixme.consume () in
- Hints.add_hints (Locality.make_module_locality l) bl
+let add_hints_iff ?locality l2r lc n bl =
+ Hints.add_hints (Locality.make_module_locality locality) bl
(Hints.HintsResolveEntry (List.map (project_hint n l2r) lc))
-VERNAC COMMAND EXTEND HintResolveIffLR CLASSIFIED AS SIDEFF
+VERNAC COMMAND FUNCTIONAL EXTEND HintResolveIffLR CLASSIFIED AS SIDEFF
[ "Hint" "Resolve" "->" ne_global_list(lc) natural_opt(n)
":" preident_list(bl) ] ->
- [ add_hints_iff true lc n bl ]
+ [ fun ~atts ~st -> begin
+ let open Vernacinterp in
+ add_hints_iff ?locality:atts.locality true lc n bl;
+ st
+ end
+ ]
| [ "Hint" "Resolve" "->" ne_global_list(lc) natural_opt(n) ] ->
- [ add_hints_iff true lc n ["core"] ]
+ [ fun ~atts ~st -> begin
+ let open Vernacinterp in
+ add_hints_iff ?locality:atts.locality true lc n ["core"];
+ st
+ end
+ ]
END
-VERNAC COMMAND EXTEND HintResolveIffRL CLASSIFIED AS SIDEFF
+
+VERNAC COMMAND FUNCTIONAL EXTEND HintResolveIffRL CLASSIFIED AS SIDEFF
[ "Hint" "Resolve" "<-" ne_global_list(lc) natural_opt(n)
":" preident_list(bl) ] ->
- [ add_hints_iff false lc n bl ]
+ [ fun ~atts ~st -> begin
+ let open Vernacinterp in
+ add_hints_iff ?locality:atts.locality false lc n bl;
+ st
+ end
+ ]
| [ "Hint" "Resolve" "<-" ne_global_list(lc) natural_opt(n) ] ->
- [ add_hints_iff false lc n ["core"] ]
+ [ fun ~atts ~st -> begin
+ let open Vernacinterp in
+ add_hints_iff ?locality:atts.locality false lc n ["core"];
+ st
+ end
+ ]
END
(**********************************************************************)
diff --git a/plugins/ltac/g_auto.ml4 b/plugins/ltac/g_auto.ml4
index 84e79d8ab..90a44708f 100644
--- a/plugins/ltac/g_auto.ml4
+++ b/plugins/ltac/g_auto.ml4
@@ -190,7 +190,7 @@ END
let pr_hints_path prc prx pry c = Hints.pp_hints_path c
let pr_pre_hints_path prc prx pry c = Hints.pp_hints_path_gen Libnames.pr_reference c
let glob_hints_path ist = Hints.glob_hints_path
-
+
ARGUMENT EXTEND hints_path
PRINTED BY pr_hints_path
@@ -214,10 +214,15 @@ ARGUMENT EXTEND opthints
| [ ] -> [ None ]
END
-VERNAC COMMAND EXTEND HintCut CLASSIFIED AS SIDEFF
+VERNAC COMMAND FUNCTIONAL EXTEND HintCut CLASSIFIED AS SIDEFF
| [ "Hint" "Cut" "[" hints_path(p) "]" opthints(dbnames) ] -> [
- let entry = Hints.HintsCutEntry (Hints.glob_hints_path p) in
- Hints.add_hints (Locality.make_section_locality (Locality.LocalityFixme.consume ()))
- (match dbnames with None -> ["core"] | Some l -> l) entry ]
+ fun ~atts ~st -> begin
+ let open Vernacinterp in
+ let entry = Hints.HintsCutEntry (Hints.glob_hints_path p) in
+ Hints.add_hints (Locality.make_section_locality atts.locality)
+ (match dbnames with None -> ["core"] | Some l -> l) entry;
+ st
+ end
+ ]
END
diff --git a/plugins/ltac/g_ltac.ml4 b/plugins/ltac/g_ltac.ml4
index 116152568..34fea6175 100644
--- a/plugins/ltac/g_ltac.ml4
+++ b/plugins/ltac/g_ltac.ml4
@@ -469,13 +469,13 @@ VERNAC ARGUMENT EXTEND ltac_production_item PRINTED BY pr_ltac_production_item
[ Tacentries.TacNonTerm (Loc.tag ~loc ((Id.to_string nt, None), None)) ]
END
-VERNAC COMMAND EXTEND VernacTacticNotation
+VERNAC COMMAND FUNCTIONAL EXTEND VernacTacticNotation
| [ "Tactic" "Notation" ltac_tactic_level_opt(n) ne_ltac_production_item_list(r) ":=" tactic(e) ] =>
[ VtUnknown, VtNow ] ->
- [
- let l = Locality.LocalityFixme.consume () in
- let n = Option.default 0 n in
- Tacentries.add_tactic_notation (Locality.make_module_locality l) n r e
+ [ fun ~atts ~st -> let open Vernacinterp in
+ let n = Option.default 0 n in
+ Tacentries.add_tactic_notation (Locality.make_module_locality atts.locality) n r e;
+ st
]
END
@@ -512,15 +512,15 @@ PRINTED BY pr_tacdef_body
| [ tacdef_body(t) ] -> [ t ]
END
-VERNAC COMMAND EXTEND VernacDeclareTacticDefinition
+VERNAC COMMAND FUNCTIONAL EXTEND VernacDeclareTacticDefinition
| [ "Ltac" ne_ltac_tacdef_body_list_sep(l, "with") ] => [
VtSideff (List.map (function
| TacticDefinition ((_,r),_) -> r
| TacticRedefinition (Ident (_,r),_) -> r
| TacticRedefinition (Qualid (_,q),_) -> snd(repr_qualid q)) l), VtLater
- ] -> [
- let lc = Locality.LocalityFixme.consume () in
- Tacentries.register_ltac (Locality.make_module_locality lc) l
+ ] -> [ fun ~atts ~st -> let open Vernacinterp in
+ Tacentries.register_ltac (Locality.make_module_locality atts.locality) l;
+ st
]
END
diff --git a/plugins/ltac/g_obligations.ml4 b/plugins/ltac/g_obligations.ml4
index fea9e837b..f6cc3833a 100644
--- a/plugins/ltac/g_obligations.ml4
+++ b/plugins/ltac/g_obligations.ml4
@@ -123,11 +123,15 @@ VERNAC COMMAND EXTEND Admit_Obligations CLASSIFIED AS SIDEFF
| [ "Admit" "Obligations" ] -> [ admit_obligations None ]
END
-VERNAC COMMAND EXTEND Set_Solver CLASSIFIED AS SIDEFF
+VERNAC COMMAND FUNCTIONAL EXTEND Set_Solver CLASSIFIED AS SIDEFF
| [ "Obligation" "Tactic" ":=" tactic(t) ] -> [
- set_default_tactic
- (Locality.make_section_locality (Locality.LocalityFixme.consume ()))
- (Tacintern.glob_tactic t) ]
+ fun ~atts ~st -> begin
+ let open Vernacinterp in
+ set_default_tactic
+ (Locality.make_section_locality atts.locality)
+ (Tacintern.glob_tactic t);
+ st
+ end]
END
open Pp
diff --git a/plugins/ltac/g_rewrite.ml4 b/plugins/ltac/g_rewrite.ml4
index 91abe1019..ea1808a25 100644
--- a/plugins/ltac/g_rewrite.ml4
+++ b/plugins/ltac/g_rewrite.ml4
@@ -243,22 +243,37 @@ VERNAC COMMAND EXTEND AddParametricRelation3 CLASSIFIED AS SIDEFF
[ declare_relation ~binders:b a aeq n None None (Some lemma3) ]
END
-VERNAC COMMAND EXTEND AddSetoid1 CLASSIFIED AS SIDEFF
+VERNAC COMMAND FUNCTIONAL EXTEND AddSetoid1 CLASSIFIED AS SIDEFF
[ "Add" "Setoid" constr(a) constr(aeq) constr(t) "as" ident(n) ] ->
- [ add_setoid (not (Locality.make_section_locality (Locality.LocalityFixme.consume ()))) [] a aeq t n ]
+ [ fun ~atts ~st -> let open Vernacinterp in
+ add_setoid (not (Locality.make_section_locality atts.locality)) [] a aeq t n;
+ st
+ ]
| [ "Add" "Parametric" "Setoid" binders(binders) ":" constr(a) constr(aeq) constr(t) "as" ident(n) ] ->
- [ add_setoid (not (Locality.make_section_locality (Locality.LocalityFixme.consume ()))) binders a aeq t n ]
+ [ fun ~atts ~st -> let open Vernacinterp in
+ add_setoid (not (Locality.make_section_locality atts.locality)) binders a aeq t n;
+ st
+ ]
| [ "Add" "Morphism" constr(m) ":" ident(n) ]
(* This command may or may not open a goal *)
=> [ Vernacexpr.VtUnknown, Vernacexpr.VtNow ]
- -> [ add_morphism_infer (not (Locality.make_section_locality (Locality.LocalityFixme.consume ()))) m n ]
+ -> [ fun ~atts ~st -> let open Vernacinterp in
+ add_morphism_infer (not (Locality.make_section_locality atts.locality)) m n;
+ st
+ ]
| [ "Add" "Morphism" constr(m) "with" "signature" lconstr(s) "as" ident(n) ]
=> [ Vernacexpr.(VtStartProof("Classic",GuaranteesOpacity,[n]), VtLater) ]
- -> [ add_morphism (not (Locality.make_section_locality (Locality.LocalityFixme.consume ()))) [] m s n ]
+ -> [ fun ~atts ~st -> let open Vernacinterp in
+ add_morphism (not (Locality.make_section_locality atts.locality)) [] m s n;
+ st
+ ]
| [ "Add" "Parametric" "Morphism" binders(binders) ":" constr(m)
"with" "signature" lconstr(s) "as" ident(n) ]
=> [ Vernacexpr.(VtStartProof("Classic",GuaranteesOpacity,[n]), VtLater) ]
- -> [ add_morphism (not (Locality.make_section_locality (Locality.LocalityFixme.consume ()))) binders m s n ]
+ -> [ fun ~atts ~st -> let open Vernacinterp in
+ add_morphism (not (Locality.make_section_locality atts.locality)) binders m s n;
+ st
+ ]
END
TACTIC EXTEND setoid_symmetry
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index c63492d1b..14b0742a7 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -1800,9 +1800,9 @@ let declare_instance_trans global binders a aeq n lemma =
in anew_instance global binders instance
[(Ident (Loc.tag @@ Id.of_string "transitivity"),lemma)]
-let declare_relation ?(binders=[]) a aeq n refl symm trans =
+let declare_relation ?locality ?(binders=[]) a aeq n refl symm trans =
init_setoid ();
- let global = not (Locality.make_section_locality (Locality.LocalityFixme.consume ())) in
+ let global = not (Locality.make_section_locality locality) in
let instance = declare_instance a aeq (add_suffix n "_relation") "Coq.Classes.RelationClasses.RewriteRelation"
in ignore(anew_instance global binders instance []);
match (refl,symm,trans) with
diff --git a/plugins/ltac/rewrite.mli b/plugins/ltac/rewrite.mli
index 1306c590b..17e7244b3 100644
--- a/plugins/ltac/rewrite.mli
+++ b/plugins/ltac/rewrite.mli
@@ -75,7 +75,7 @@ val cl_rewrite_clause :
val is_applied_rewrite_relation :
env -> evar_map -> rel_context -> constr -> types option
-val declare_relation :
+val declare_relation : ?locality:bool ->
?binders:local_binder_expr list -> constr_expr -> constr_expr -> Id.t ->
constr_expr option -> constr_expr option -> constr_expr option -> unit
diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml
index 96bf31b11..0ea8904f2 100644
--- a/plugins/quote/quote.ml
+++ b/plugins/quote/quote.ml
@@ -104,7 +104,7 @@
open CErrors
open Util
open Names
-open Term
+open Constr
open EConstr
open Pattern
open Patternops
diff --git a/plugins/romega/const_omega.ml b/plugins/romega/const_omega.ml
index 7956be58b..337510ef1 100644
--- a/plugins/romega/const_omega.ml
+++ b/plugins/romega/const_omega.ml
@@ -7,7 +7,6 @@
*************************************************************************)
open Names
-open Term
open Constr
let module_refl_name = "ReflOmegaCore"
diff --git a/plugins/ssr/ssrequality.ml b/plugins/ssr/ssrequality.ml
index 274c7110c..bd9633afb 100644
--- a/plugins/ssr/ssrequality.ml
+++ b/plugins/ssr/ssrequality.ml
@@ -342,7 +342,7 @@ let pirrel_rewrite pred rdx rdx_ty new_rdx dir (sigma, c) c_ty gl =
let sort = elimination_sort_of_goal gl in
let elim, gl = pf_fresh_global (Indrec.lookup_eliminator ind sort) gl in
if dir = R2L then elim, gl else (* taken from Coq's rewrite *)
- let elim, _ = Term.destConst elim in
+ let elim, _ = destConst elim in
let mp,dp,l = Constant.repr3 (Constant.make1 (Constant.canonical elim)) in
let l' = Label.of_id (Nameops.add_suffix (Label.to_id l) "_r") in
let c1' = Global.constant_of_delta_kn (Constant.canonical (Constant.make3 mp dp l')) in
diff --git a/plugins/ssr/ssrfwd.ml b/plugins/ssr/ssrfwd.ml
index a707226cd..5c1b399a8 100644
--- a/plugins/ssr/ssrfwd.ml
+++ b/plugins/ssr/ssrfwd.ml
@@ -8,11 +8,12 @@
(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+open Pp
open Names
+open Constr
open Tacmach
open Ssrmatching_plugin.Ssrmatching
-
open Ssrprinters
open Ssrcommon
open Ssrtacticals
@@ -30,10 +31,6 @@ let ssrposetac ist (id, (_, t)) gl =
let sigma, t, ucst, _ = pf_abs_ssrterm ist gl t in
posetac id t (pf_merge_uc ucst gl)
-open Pp
-open Term
-open Constr
-
let ssrsettac ist id ((_, (pat, pty)), (_, occ)) gl =
let pat = interp_cpattern ist gl pat (Option.map snd pty) in
let cl, sigma, env = pf_concl gl, project gl, pf_env gl in
diff --git a/plugins/ssr/ssrvernac.ml4 b/plugins/ssr/ssrvernac.ml4
index cd614fee9..7385ed84c 100644
--- a/plugins/ssr/ssrvernac.ml4
+++ b/plugins/ssr/ssrvernac.ml4
@@ -158,11 +158,14 @@ let declare_one_prenex_implicit locality f =
| impls ->
Impargs.declare_manual_implicits locality fref ~enriching:false [impls]
-VERNAC COMMAND EXTEND Ssrpreneximplicits CLASSIFIED AS SIDEFF
+VERNAC COMMAND FUNCTIONAL EXTEND Ssrpreneximplicits CLASSIFIED AS SIDEFF
| [ "Prenex" "Implicits" ne_global_list(fl) ]
- -> [ let locality =
- Locality.make_section_locality (Locality.LocalityFixme.consume ()) in
- List.iter (declare_one_prenex_implicit locality) fl ]
+ -> [ fun ~atts ~st ->
+ let open Vernacinterp in
+ let locality = Locality.make_section_locality atts.locality in
+ List.iter (declare_one_prenex_implicit locality) fl;
+ st
+ ]
END
(* Vernac grammar visibility patch *)
diff --git a/pretyping/cases.mli b/pretyping/cases.mli
index 3a139b7b0..43dbc3105 100644
--- a/pretyping/cases.mli
+++ b/pretyping/cases.mli
@@ -13,8 +13,8 @@ open Environ
open EConstr
open Inductiveops
open Glob_term
-open Evarutil
open Ltac_pretype
+open Evardefine
(** {5 Compilation of pattern-matching } *)
@@ -116,7 +116,7 @@ type 'a pattern_matching_problem =
val compile : 'a pattern_matching_problem -> unsafe_judgment
val prepare_predicate : ?loc:Loc.t ->
- (Evarutil.type_constraint ->
+ (type_constraint ->
Environ.env -> Evd.evar_map ref -> ltac_var_map -> glob_constr -> unsafe_judgment) ->
Environ.env ->
Evd.evar_map ->
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index 681eb17d3..18e0c31dd 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -9,7 +9,6 @@
open CErrors
open Util
open Names
-open Term
open Constr
open Termops
open Environ
@@ -49,7 +48,7 @@ let _ = Goptions.declare_bool_option {
"data.id.type" etc... *)
let impossible_default_case () =
let c, ctx = Universes.fresh_global_instance (Global.env()) (Globnames.ConstRef Coqlib.id) in
- let (_, u) = Term.destConst c in
+ let (_, u) = Constr.destConst c in
Some (c, Constr.mkConstU (Coqlib.type_of_id, u), ctx)
let coq_unit_judge =
diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml
index 18dbbea1b..b646a37f8 100644
--- a/pretyping/evardefine.ml
+++ b/pretyping/evardefine.ml
@@ -6,10 +6,10 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Sorts
open Util
open Pp
open Names
-open Term
open Constr
open Termops
open EConstr
diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml
index fba154291..e6d1e59b3 100644
--- a/pretyping/evarsolve.ml
+++ b/pretyping/evarsolve.ml
@@ -6,10 +6,10 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Sorts
open Util
open CErrors
open Names
-open Term
open Constr
open Environ
open Termops
diff --git a/engine/geninterp.ml b/pretyping/geninterp.ml
index 768ef3cfd..768ef3cfd 100644
--- a/engine/geninterp.ml
+++ b/pretyping/geninterp.ml
diff --git a/engine/geninterp.mli b/pretyping/geninterp.mli
index ae0b26e59..ae0b26e59 100644
--- a/engine/geninterp.mli
+++ b/pretyping/geninterp.mli
diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli
index eb2b435bf..b2735ee22 100644
--- a/pretyping/pretyping.mli
+++ b/pretyping/pretyping.mli
@@ -17,8 +17,8 @@ open Environ
open Evd
open EConstr
open Glob_term
-open Evarutil
open Ltac_pretype
+open Evardefine
(** An auxiliary function for searching for fixpoint guard indexes *)
diff --git a/pretyping/pretyping.mllib b/pretyping/pretyping.mllib
index 9904b7354..1da5b4567 100644
--- a/pretyping/pretyping.mllib
+++ b/pretyping/pretyping.mllib
@@ -1,3 +1,4 @@
+Geninterp
Ltac_pretype
Locusops
Pretype_errors
diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml
index e6d8a0af2..9ff9a75b3 100644
--- a/pretyping/recordops.ml
+++ b/pretyping/recordops.ml
@@ -213,7 +213,7 @@ let compute_canonical_projections warn (con,ind) =
let sign = List.map (on_snd EConstr.Unsafe.to_constr) sign in
let t = EConstr.Unsafe.to_constr t in
let lt = List.rev_map snd sign in
- let args = snd (Term.decompose_app t) in
+ let args = snd (decompose_app t) in
let { s_EXPECTEDPARAM = p; s_PROJ = lpj; s_PROJKIND = kl } =
lookup_structure ind in
let params, projs = List.chop p args in
diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml
index 143f9ddcc..e897b1938 100644
--- a/printing/ppvernac.ml
+++ b/printing/ppvernac.ml
@@ -447,7 +447,7 @@ open Decl_kinds
| PrintGrammar ent ->
keyword "Print Grammar" ++ spc() ++ str ent
| PrintLoadPath dir ->
- keyword "Print LoadPath" ++ pr_opt pr_dirpath dir
+ keyword "Print LoadPath" ++ pr_opt DirPath.print dir
| PrintModules ->
keyword "Print Modules"
| PrintMLLoadPath ->
@@ -518,7 +518,7 @@ open Decl_kinds
in
keyword cmd ++ spc() ++ pr_smart_global qid
| PrintNamespace dp ->
- keyword "Print Namespace" ++ pr_dirpath dp
+ keyword "Print Namespace" ++ DirPath.print dp
| PrintStrategy None ->
keyword "Print Strategies"
| PrintStrategy (Some qid) ->
@@ -964,7 +964,7 @@ open Decl_kinds
keyword "LoadPath" ++ spc() ++ qs s ++
(match d with
| None -> mt()
- | Some dir -> spc() ++ keyword "as" ++ spc() ++ pr_dirpath dir))
+ | Some dir -> spc() ++ keyword "as" ++ spc() ++ DirPath.print dir))
)
| VernacRemoveLoadPath s ->
return (keyword "Remove LoadPath" ++ qs s)
diff --git a/printing/prettyp.ml b/printing/prettyp.ml
index e2d23ce7b..8fc00ed96 100644
--- a/printing/prettyp.ml
+++ b/printing/prettyp.ml
@@ -15,7 +15,6 @@ open CErrors
open Util
open Names
open Nameops
-open Term
open Termops
open Declarations
open Environ
@@ -139,7 +138,7 @@ let print_renames_list prefix l =
let need_expansion impl ref =
let typ, _ = Global.type_of_global_in_context (Global.env ()) ref in
- let ctx = prod_assum typ in
+ let ctx = Term.prod_assum typ in
let nprods = List.count is_local_assum ctx in
not (List.is_empty impl) && List.length impl >= nprods &&
let _,lastimpl = List.chop nprods impl in
@@ -366,7 +365,7 @@ let pr_located_qualid = function
| DirModule (dir,_) -> "Module", dir
| DirClosedSection dir -> "Closed Section", dir
in
- str s ++ spc () ++ pr_dirpath dir
+ str s ++ spc () ++ DirPath.print dir
| ModuleType mp ->
str "Module Type" ++ spc () ++ pr_path (Nametab.path_of_modtype mp)
| Other (obj, info) -> info.name obj
@@ -490,7 +489,7 @@ let gallina_print_typed_value_in_env env sigma (trm,typ) =
let print_named_def env sigma name body typ =
let pbody = pr_lconstr_env env sigma body in
let ptyp = pr_ltype_env env sigma typ in
- let pbody = if isCast body then surround pbody else pbody in
+ let pbody = if Constr.isCast body then surround pbody else pbody in
(str "*** [" ++ str name ++ str " " ++
hov 0 (str ":=" ++ brk (1,2) ++ pbody ++ spc () ++
str ":" ++ brk (1,2) ++ ptyp) ++
@@ -647,7 +646,7 @@ let gallina_print_library_entry env sigma with_values ent =
| (oname,Lib.ClosedSection _) ->
Some (str " >>>>>>> Closed Section " ++ pr_name oname)
| (_,Lib.CompilingLibrary (dir,_)) ->
- Some (str " >>>>>>> Library " ++ pr_dirpath dir)
+ Some (str " >>>>>>> Library " ++ DirPath.print dir)
| (oname,Lib.OpenedModule _) ->
Some (str " >>>>>>> Module " ++ pr_name oname)
| (oname,Lib.ClosedModule _) ->
diff --git a/printing/printer.ml b/printing/printer.ml
index 377a6b4e1..d7bb0460d 100644
--- a/printing/printer.ml
+++ b/printing/printer.ml
@@ -10,7 +10,6 @@ open Pp
open CErrors
open Util
open Names
-open Term
open Constr
open Environ
open Globnames
diff --git a/proofs/logic.ml b/proofs/logic.ml
index 13a4e4ce3..a9ad606a0 100644
--- a/proofs/logic.ml
+++ b/proofs/logic.ml
@@ -11,7 +11,6 @@ open CErrors
open Util
open Names
open Nameops
-open Term
open Constr
open Vars
open Termops
diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml
index 6a9cd7e20..de98f6382 100644
--- a/tactics/autorewrite.ml
+++ b/tactics/autorewrite.ml
@@ -9,7 +9,6 @@
open Equality
open Names
open Pp
-open Term
open Constr
open Termops
open CErrors
diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml
index 75fae6647..8e851375a 100644
--- a/tactics/hipattern.ml
+++ b/tactics/hipattern.ml
@@ -160,7 +160,7 @@ let test_strict_disjunction n lc =
let open Term in
Array.for_all_i (fun i c ->
match (prod_assum (snd (decompose_prod_n_assum n c))) with
- | [LocalAssum (_,c)] -> isRel c && Int.equal (destRel c) (n - i)
+ | [LocalAssum (_,c)] -> Constr.isRel c && Int.equal (Constr.destRel c) (n - i)
| _ -> false) 0 lc
let match_with_disjunction ?(strict=false) ?(onlybinary=false) sigma t =
diff --git a/test-suite/bugs/closed/5215.v b/test-suite/bugs/closed/5215.v
new file mode 100644
index 000000000..ecf529159
--- /dev/null
+++ b/test-suite/bugs/closed/5215.v
@@ -0,0 +1,286 @@
+Require Import Coq.Logic.FunctionalExtensionality.
+Require Import Coq.Program.Tactics.
+
+Global Set Primitive Projections.
+
+Global Set Universe Polymorphism.
+
+Global Unset Universe Minimization ToSet.
+
+Class Category : Type :=
+{
+ Obj : Type;
+ Hom : Obj -> Obj -> Type;
+ compose : forall {a b c : Obj}, (Hom a b) -> (Hom b c) -> (Hom a c);
+ id : forall {a : Obj}, Hom a a;
+}.
+
+Arguments Obj {_}, _.
+Arguments id {_ _}, {_} _, _ _.
+Arguments Hom {_} _ _, _ _ _.
+Arguments compose {_} {_ _ _} _ _, _ {_ _ _} _ _, _ _ _ _ _ _.
+
+Coercion Obj : Category >-> Sortclass.
+
+Definition Opposite (C : Category) : Category :=
+{|
+
+ Obj := Obj C;
+ Hom := fun a b => Hom b a;
+ compose :=
+ fun a b c (f : Hom b a) (g : Hom c b) => compose C c b a g f;
+ id := fun c => id C c;
+|}.
+
+Record Functor (C C' : Category) : Type :=
+{
+ FO : C -> C';
+ FA : forall {a b}, Hom a b -> Hom (FO a) (FO b);
+}.
+
+Arguments FO {_ _} _ _.
+Arguments FA {_ _} _ {_ _} _, {_ _} _ _ _ _.
+
+Section Opposite_Functor.
+ Context {C D : Category} (F : Functor C D).
+
+ Program Definition Opposite_Functor : (Functor (Opposite C) (Opposite D)) :=
+ {|
+ FO := FO F;
+ FA := fun _ _ h => FA F h;
+ |}.
+
+End Opposite_Functor.
+
+Section Functor_Compose.
+ Context {C C' C'' : Category} (F : Functor C C') (F' : Functor C' C'').
+
+ Program Definition Functor_compose : Functor C C'' :=
+ {|
+ FO := fun c => FO F' (FO F c);
+ FA := fun c d f => FA F' (FA F f)
+ |}.
+
+End Functor_Compose.
+
+Section Algebras.
+ Context {C : Category} (T : Functor C C).
+ Record Algebra : Type :=
+ {
+ Alg_Carrier : C;
+ Constructors : Hom (FO T Alg_Carrier) Alg_Carrier
+ }.
+
+ Record Algebra_Hom (alg alg' : Algebra) : Type :=
+ {
+ Alg_map : Hom (Alg_Carrier alg) (Alg_Carrier alg');
+
+ Alg_map_com : compose (FA T Alg_map) (Constructors alg')
+ = compose (Constructors alg) Alg_map
+ }.
+
+ Arguments Alg_map {_ _} _.
+ Arguments Alg_map_com {_ _} _.
+ Program Definition Algebra_Hom_compose
+ {alg alg' alg'' : Algebra}
+ (h : Algebra_Hom alg alg')
+ (h' : Algebra_Hom alg' alg'')
+ : Algebra_Hom alg alg''
+ :=
+ {|
+ Alg_map := compose (Alg_map h) (Alg_map h')
+ |}.
+
+ Next Obligation. Proof. Admitted.
+
+ Lemma Algebra_Hom_eq_simplify (alg alg' : Algebra)
+ (ah ah' : Algebra_Hom alg alg')
+ : (Alg_map ah) = (Alg_map ah') -> ah = ah'.
+ Proof. Admitted.
+
+ Program Definition Algebra_Hom_id (alg : Algebra) : Algebra_Hom alg alg :=
+ {|
+ Alg_map := id
+ |}.
+
+ Next Obligation. Admitted.
+
+ Definition Algebra_Cat : Category :=
+ {|
+ Obj := Algebra;
+ Hom := Algebra_Hom;
+ compose := @Algebra_Hom_compose;
+ id := Algebra_Hom_id;
+ |}.
+
+End Algebras.
+
+Arguments Alg_Carrier {_ _} _.
+Arguments Constructors {_ _} _.
+Arguments Algebra_Hom {_ _} _ _.
+Arguments Alg_map {_ _ _ _} _.
+Arguments Alg_map_com {_ _ _ _} _.
+Arguments Algebra_Hom_id {_ _} _.
+
+Section CoAlgebras.
+ Context {C : Category}.
+
+ Definition CoAlgebra (T : Functor C C) :=
+ @Algebra (Opposite C) (Opposite_Functor T).
+
+ Definition CoAlgebra_Hom {T : Functor C C} :=
+ @Algebra_Hom (Opposite C) (Opposite_Functor T).
+
+ Definition CoAlgebra_Hom_id {T : Functor C C} :=
+ @Algebra_Hom_id (Opposite C) (Opposite_Functor T).
+
+ Definition CoAlgebra_Cat (T : Functor C C) :=
+ @Algebra_Cat (Opposite C) (Opposite_Functor T).
+
+End CoAlgebras.
+
+Program Definition Type_Cat : Category :=
+{|
+ Obj := Type;
+ Hom := (fun A B => A -> B);
+ compose := fun A B C (g : A -> B) (h : B -> C) => fun (x : A) => h (g x);
+ id := fun A => fun x => x
+|}.
+
+Local Obligation Tactic := idtac.
+
+Program Definition Prod_Cat (C C' : Category) : Category :=
+{|
+ Obj := C * C';
+ Hom :=
+ fun a b =>
+ ((Hom (fst a) (fst b)) * (Hom (snd a) (snd b)))%type;
+ compose :=
+ fun a b c f g =>
+ ((compose (fst f) (fst g)), (compose (snd f)(snd g)));
+ id := fun c => (id, id)
+|}.
+
+Class Terminal (C : Category) : Type :=
+{
+ terminal : C;
+ t_morph : forall (d : Obj), Hom d terminal;
+ t_morph_unique : forall (d : Obj) (f g : (Hom d terminal)), f = g
+}.
+
+Arguments terminal {_} _.
+Arguments t_morph {_} _ _.
+Arguments t_morph_unique {_} _ _ _ _.
+
+Coercion terminal : Terminal >-> Obj.
+
+Definition Initial (C : Category) := Terminal (Opposite C).
+Existing Class Initial.
+
+Record Product {C : Category} (c d : C) : Type :=
+{
+ product : C;
+ Pi_1 : Hom product c;
+ Pi_2 : Hom product d;
+ Prod_morph_ex : forall (p' : Obj) (r1 : Hom p' c) (r2 : Hom p' d), (Hom p' product);
+}.
+
+Arguments Product _ _ _, {_} _ _.
+
+Arguments Pi_1 {_ _ _ _}, {_ _ _} _.
+Arguments Pi_2 {_ _ _ _}, {_ _ _} _.
+Arguments Prod_morph_ex {_ _ _} _ _ _ _.
+
+Coercion product : Product >-> Obj.
+
+Definition Has_Products (C : Category) : Type := forall a b, Product a b.
+
+Existing Class Has_Products.
+
+Program Definition Prod_Func (C : Category) {HP : Has_Products C}
+ : Functor (Prod_Cat C C) C :=
+{|
+ FO := fun x => HP (fst x) (snd x);
+ FA := fun a b f => Prod_morph_ex _ _ (compose Pi_1 (fst f)) (compose Pi_2 (snd f))
+|}.
+
+Arguments Prod_Func _ _, _ {_}.
+
+Definition Sum (C : Category) := @Product (Opposite C).
+
+Arguments Sum _ _ _, {_} _ _.
+
+Definition Has_Sums (C : Category) : Type := forall (a b : C), (Sum a b).
+
+Existing Class Has_Sums.
+
+Program Definition sum_Sum (A B : Type) : (@Sum Type_Cat A B) :=
+{|
+ product := (A + B)%type;
+ Prod_morph_ex :=
+ fun (p' : Type)
+ (r1 : A -> p')
+ (r2 : B -> p')
+ (X : A + B) =>
+ match X return p' with
+ | inl a => r1 a
+ | inr b => r2 b
+ end
+|}.
+Next Obligation. simpl; auto. Defined.
+Next Obligation. simpl; auto. Defined.
+
+Program Instance Type_Cat_Has_Sums : Has_Sums Type_Cat := sum_Sum.
+
+Definition Sum_Func {C : Category} {HS : Has_Sums C} :
+ Functor (Prod_Cat C C) C := Opposite_Functor (Prod_Func (Opposite C) HS).
+
+Arguments Sum_Func _ _, _ {_}.
+
+Program Instance unit_Type_term : Terminal Type_Cat :=
+{
+ terminal := unit;
+ t_morph := fun _ _=> tt
+}.
+
+Next Obligation. Proof. Admitted.
+
+Program Definition term_id : Functor Type_Cat (Prod_Cat Type_Cat Type_Cat) :=
+{|
+ FO := fun a => (@terminal Type_Cat _, a);
+ FA := fun a b f => (@id _ (@terminal Type_Cat _), f)
+|}.
+
+Definition S_nat_func : Functor Type_Cat Type_Cat :=
+ Functor_compose term_id (Sum_Func Type_Cat _).
+
+Definition S_nat_alg_cat := Algebra_Cat S_nat_func.
+
+CoInductive CoNat : Set :=
+ | CoO : CoNat
+ | CoS : CoNat -> CoNat
+.
+
+Definition S_nat_coalg_cat := @CoAlgebra_Cat Type_Cat S_nat_func.
+
+Set Printing Universes.
+Program Definition CoNat_alg_term : Initial S_nat_coalg_cat :=
+{|
+ terminal := _;
+ t_morph := _
+|}.
+
+Next Obligation. Admitted.
+Next Obligation. Admitted.
+
+Axiom Admit : False.
+
+Next Obligation.
+Proof.
+ intros d f g.
+ assert(H1 := (@Alg_map_com _ _ _ _ f)). clear.
+ assert (inl tt = inr tt) by (exfalso; apply Admit).
+ discriminate.
+ all: exfalso; apply Admit.
+ Show Universes.
+Qed.
diff --git a/test-suite/bugs/closed/5215_2.v b/test-suite/bugs/closed/5215_2.v
new file mode 100644
index 000000000..399947f00
--- /dev/null
+++ b/test-suite/bugs/closed/5215_2.v
@@ -0,0 +1,8 @@
+Require Import Coq.Program.Tactics.
+Set Universe Polymorphism.
+Set Printing Universes.
+Definition typ := Type.
+
+Program Definition foo : typ := _ -> _.
+Next Obligation. Admitted.
+Next Obligation. exact typ. Show Proof. Show Universes. Defined.
diff --git a/test-suite/success/Notations2.v b/test-suite/success/Notations2.v
index e86b3edb8..2655b651a 100644
--- a/test-suite/success/Notations2.v
+++ b/test-suite/success/Notations2.v
@@ -96,3 +96,12 @@ Check fun A (x :prod' bool A) => match x with ##### 0 _ y 0%bool => 2 | _ => 1 e
Notation "'FUNNAT' i => t" := (fun i : nat => i = t) (at level 200).
Notation "'Funnat' i => t" := (FUNNAT i => t + i%nat) (at level 200).
+
+(* 11. Notations with needed factorization of a recursive pattern *)
+(* See https://github.com/coq/coq/issues/6078#issuecomment-342287412 *)
+Module A.
+Notation "[:: x1 ; .. ; xn & s ]" := (cons x1 .. (cons xn s) ..).
+Notation "[:: x1 ; .. ; xn ]" := (cons x1 .. (cons xn nil) ..).
+Check [:: 1 ; 2 ; 3 ].
+Check [:: 1 ; 2 ; 3 & nil ]. (* was failing *)
+End A.
diff --git a/toplevel/coqinit.ml b/toplevel/coqinit.ml
index c80899288..3a195c1df 100644
--- a/toplevel/coqinit.ml
+++ b/toplevel/coqinit.ml
@@ -65,7 +65,7 @@ let add_stdlib_path ~load_init ~unix_path ~coq_root ~with_ml =
let add_userlib_path ~unix_path =
Mltop.add_rec_path Mltop.AddRecML ~unix_path
- ~coq_root:Nameops.default_root_prefix ~implicit:false
+ ~coq_root:Libnames.default_root_prefix ~implicit:false
(* Options -I, -I-as, and -R of the command line *)
let includes = ref []
@@ -80,7 +80,7 @@ let init_load_path ~load_init =
let user_contrib = coqlib/"user-contrib" in
let xdg_dirs = Envars.xdg_dirs ~warn:(fun x -> Feedback.msg_warning (str x)) in
let coqpath = Envars.coqpath in
- let coq_root = Names.DirPath.make [Nameops.coq_root] in
+ let coq_root = Names.DirPath.make [Libnames.coq_root] in
(* NOTE: These directories are searched from last to first *)
(* first, developer specific directory to open *)
if Coq_config.local then
@@ -105,7 +105,7 @@ let init_load_path ~load_init =
List.iter (fun s -> add_userlib_path ~unix_path:s) coqpath;
(* then current directory (not recursively!) *)
Mltop.add_ml_dir ".";
- Loadpath.add_load_path "." Nameops.default_root_prefix ~implicit:false;
+ Loadpath.add_load_path "." Libnames.default_root_prefix ~implicit:false;
(* additional loadpath, given with options -Q and -R *)
List.iter
(fun (unix_path, coq_root, implicit) ->
diff --git a/vernac/command.ml b/vernac/command.ml
index fd0027c40..257c003b5 100644
--- a/vernac/command.ml
+++ b/vernac/command.ml
@@ -8,8 +8,8 @@
open Pp
open CErrors
+open Sorts
open Util
-open Term
open Constr
open Vars
open Termops
@@ -376,8 +376,8 @@ let rec check_anonymous_type ind =
| _ -> false
let make_conclusion_flexible evdref ty poly =
- if poly && isArity ty then
- let _, concl = destArity ty in
+ if poly && Term.isArity ty then
+ let _, concl = Term.destArity ty in
match concl with
| Type u ->
(match Univ.universe_level u with
diff --git a/vernac/himsg.ml b/vernac/himsg.ml
index 839064aa0..e8c5aeedd 100644
--- a/vernac/himsg.ml
+++ b/vernac/himsg.ml
@@ -11,7 +11,7 @@ open Util
open Names
open Nameops
open Namegen
-open Term
+open Constr
open Termops
open Indtypes
open Environ
@@ -405,7 +405,7 @@ let explain_not_product env sigma c =
let pr = pr_lconstr_env env sigma c in
str "The type of this term is a product" ++ spc () ++
str "while it is expected to be" ++
- (if Term.is_Type c then str " a sort" else (brk(1,1) ++ pr)) ++ str "."
+ (if Constr.is_Type c then str " a sort" else (brk(1,1) ++ pr)) ++ str "."
(* TODO: use the names *)
(* (co)fixpoints *)
diff --git a/vernac/locality.ml b/vernac/locality.ml
index 054a451a4..681b1ab20 100644
--- a/vernac/locality.ml
+++ b/vernac/locality.ml
@@ -6,22 +6,12 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Pp
-
(** * Managing locality *)
let local_of_bool = function
| true -> Decl_kinds.Local
| false -> Decl_kinds.Global
-let check_locality locality_flag =
- match locality_flag with
- | Some b ->
- let s = if b then "Local" else "Global" in
- CErrors.user_err ~hdr:"Locality.check_locality"
- (str "This command does not support the \"" ++ str s ++ str "\" prefix.")
- | None -> ()
-
(** Extracting the locality flag *)
(* Commands which supported an inlined Local flag *)
@@ -95,13 +85,3 @@ let make_module_locality = function
let enforce_module_locality locality_flag local =
make_module_locality (enforce_locality_full locality_flag local)
-
-module LocalityFixme = struct
- let locality = ref None
- let set l = locality := l
- let consume () =
- let l = !locality in
- locality := None;
- l
- let assert_consumed () = check_locality !locality
-end
diff --git a/vernac/locality.mli b/vernac/locality.mli
index c1c45d6b0..bef66d8bc 100644
--- a/vernac/locality.mli
+++ b/vernac/locality.mli
@@ -41,11 +41,3 @@ val enforce_section_locality : bool option -> bool -> bool
val make_module_locality : bool option -> bool
val enforce_module_locality : bool option -> bool -> bool
-
-(* This is the old imperative interface that is still used for
- * VernacExtend vernaculars. Time permitting this could be trashed too *)
-module LocalityFixme : sig
- val set : bool option -> unit
- val consume : unit -> bool option
- val assert_consumed : unit -> unit
-end
diff --git a/vernac/obligations.ml b/vernac/obligations.ml
index ed4d8b888..a44de66e9 100644
--- a/vernac/obligations.ml
+++ b/vernac/obligations.ml
@@ -678,6 +678,7 @@ let init_prog_info ?(opaque = false) sign n udecl b t ctx deps fixkind
obl_deps = d; obl_tac = tac })
obls, b
in
+ let ctx = UState.make_flexible_nonalgebraic ctx in
{ prg_name = n ; prg_body = b; prg_type = reduce t;
prg_ctx = ctx; prg_univdecl = udecl;
prg_obligations = (obls', Array.length obls');
@@ -841,6 +842,9 @@ let obligation_terminator name num guard hook auto pf =
Inductiveops.control_only_guard (Global.env ()) body;
(** Declare the obligation ourselves and drop the hook *)
let prg = get_info (ProgMap.find name !from_prg) in
+ (** Ensure universes are substituted properly in body and type *)
+ let body = EConstr.to_constr sigma (EConstr.of_constr body) in
+ let ty = Option.map (fun x -> EConstr.to_constr sigma (EConstr.of_constr x)) ty in
let ctx = Evd.evar_universe_context sigma in
let prg = { prg with prg_ctx = ctx } in
let obls, rem = prg.prg_obligations in
diff --git a/vernac/record.ml b/vernac/record.ml
index 1bd47a556..f09b57048 100644
--- a/vernac/record.ml
+++ b/vernac/record.ml
@@ -8,11 +8,12 @@
open Pp
open CErrors
+open Term
+open Sorts
open Util
open Names
open Globnames
open Nameops
-open Term
open Constr
open Vars
open Environ
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index 10c139e5a..62c7edb19 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -152,7 +152,7 @@ let show_match id =
(* "Print" commands *)
let print_path_entry p =
- let dir = pr_dirpath (Loadpath.logical p) in
+ let dir = DirPath.print (Loadpath.logical p) in
let path = str (Loadpath.physical p) in
Pp.hov 2 (dir ++ spc () ++ path)
@@ -175,9 +175,9 @@ let print_modules () =
let loaded_opened = List.intersect DirPath.equal opened loaded
and only_loaded = List.subtract DirPath.equal loaded opened in
str"Loaded and imported library files: " ++
- pr_vertical_list pr_dirpath loaded_opened ++ fnl () ++
+ pr_vertical_list DirPath.print loaded_opened ++ fnl () ++
str"Loaded and not imported library files: " ++
- pr_vertical_list pr_dirpath only_loaded
+ pr_vertical_list DirPath.print only_loaded
let print_module r =
@@ -361,29 +361,29 @@ let locate_file f =
let msg_found_library = function
| Library.LibLoaded, fulldir, file ->
Feedback.msg_info (hov 0
- (pr_dirpath fulldir ++ strbrk " has been loaded from file " ++
+ (DirPath.print fulldir ++ strbrk " has been loaded from file " ++
str file))
| Library.LibInPath, fulldir, file ->
Feedback.msg_info (hov 0
- (pr_dirpath fulldir ++ strbrk " is bound to file " ++ str file))
+ (DirPath.print fulldir ++ strbrk " is bound to file " ++ str file))
let err_unmapped_library ?loc ?from qid =
let dir = fst (repr_qualid qid) in
let prefix = match from with
| None -> str "."
| Some from ->
- str " and prefix " ++ pr_dirpath from ++ str "."
+ str " and prefix " ++ DirPath.print from ++ str "."
in
user_err ?loc
~hdr:"locate_library"
(strbrk "Cannot find a physical path bound to logical path matching suffix " ++
- pr_dirpath dir ++ prefix)
+ DirPath.print dir ++ prefix)
let err_notfound_library ?loc ?from qid =
let prefix = match from with
| None -> str "."
| Some from ->
- str " with prefix " ++ pr_dirpath from ++ str "."
+ str " with prefix " ++ DirPath.print from ++ str "."
in
user_err ?loc ~hdr:"locate_library"
(strbrk "Unable to locate library " ++ pr_qualid qid ++ prefix)
@@ -893,7 +893,7 @@ let expand filename =
let vernac_add_loadpath implicit pdir ldiropt =
let pdir = expand pdir in
- let alias = Option.default Nameops.default_root_prefix ldiropt in
+ let alias = Option.default Libnames.default_root_prefix ldiropt in
Mltop.add_rec_path Mltop.AddTopML ~unix_path:pdir ~coq_root:alias ~implicit
let vernac_remove_loadpath path =
@@ -2078,7 +2078,7 @@ let interp ?proof ?loc locality poly st c =
(* Extensions *)
| VernacExtend (opn,args) ->
(* XXX: Here we are returning the state! :) *)
- let _st : Vernacstate.t = Vernacinterp.call ?locality ?loc (opn,args) st in
+ let _st : Vernacstate.t = Vernacinterp.call ?locality ?loc (opn,args) ~st in
()
(* Vernaculars that take a locality flag *)
diff --git a/vernac/vernacinterp.ml b/vernac/vernacinterp.ml
index 1d024386e..47dec1958 100644
--- a/vernac/vernacinterp.ml
+++ b/vernac/vernacinterp.ml
@@ -11,8 +11,16 @@ open Pp
open CErrors
type deprecation = bool
-type vernac_command = Genarg.raw_generic_argument list -> Loc.t option ->
- Vernacstate.t -> Vernacstate.t
+
+type atts = {
+ loc : Loc.t option;
+ locality : bool option;
+}
+
+type vernac_command =
+ Genarg.raw_generic_argument list ->
+ atts:atts -> st:Vernacstate.t ->
+ Vernacstate.t
(* Table of vernac entries *)
let vernac_tab =
@@ -66,10 +74,8 @@ let call ?locality ?loc (opn,converted_args) =
phase := "Checking arguments";
let hunk = callback converted_args in
phase := "Executing command";
- Locality.LocalityFixme.set locality;
- let res = hunk loc in
- Locality.LocalityFixme.assert_consumed ();
- res
+ let atts = { loc; locality } in
+ hunk ~atts
with
| Drop -> raise Drop
| reraise ->
diff --git a/vernac/vernacinterp.mli b/vernac/vernacinterp.mli
index 1c66b1c04..602ccba15 100644
--- a/vernac/vernacinterp.mli
+++ b/vernac/vernacinterp.mli
@@ -10,8 +10,15 @@
type deprecation = bool
-type vernac_command = Genarg.raw_generic_argument list -> Loc.t option ->
- Vernacstate.t -> Vernacstate.t
+type atts = {
+ loc : Loc.t option;
+ locality : bool option;
+}
+
+type vernac_command =
+ Genarg.raw_generic_argument list ->
+ atts:atts -> st:Vernacstate.t ->
+ Vernacstate.t
val vinterp_add : deprecation -> Vernacexpr.extend_name -> vernac_command -> unit
@@ -21,4 +28,4 @@ val vinterp_init : unit -> unit
val call : ?locality:bool -> ?loc:Loc.t ->
Vernacexpr.extend_name * Genarg.raw_generic_argument list ->
- Vernacstate.t -> Vernacstate.t
+ st:Vernacstate.t -> Vernacstate.t
diff --git a/vernac/vernacstate.ml b/vernac/vernacstate.ml
index 9802a03ca..eb1359d52 100644
--- a/vernac/vernacstate.ml
+++ b/vernac/vernacstate.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-type t = { (* TODO: inline records in OCaml 4.03 *)
+type t = {
system : States.state; (* summary + libstack *)
proof : Proof_global.state; (* proof state *)
shallow : bool (* is the state trimmed down (libstack) *)
diff --git a/vernac/vernacstate.mli b/vernac/vernacstate.mli
index 63a5b3b1e..bcfa49aa3 100644
--- a/vernac/vernacstate.mli
+++ b/vernac/vernacstate.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-type t = { (* TODO: inline records in OCaml 4.03 *)
+type t = {
system : States.state; (* summary + libstack *)
proof : Proof_global.state; (* proof state *)
shallow : bool (* is the state trimmed down (libstack) *)