aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--Makefile.build2
-rw-r--r--checker/cic.mli2
-rw-r--r--checker/environ.ml19
-rw-r--r--checker/environ.mli2
-rw-r--r--checker/mod_checking.ml22
-rw-r--r--checker/modops.ml7
-rw-r--r--checker/reduction.ml2
-rw-r--r--checker/safe_typing.ml4
-rw-r--r--checker/univ.ml87
-rw-r--r--checker/univ.mli20
-rw-r--r--checker/values.ml6
-rw-r--r--dev/doc/univpoly.txt48
-rw-r--r--engine/evd.ml187
-rw-r--r--engine/evd.mli8
-rw-r--r--kernel/declarations.mli4
-rw-r--r--kernel/environ.ml42
-rw-r--r--kernel/environ.mli4
-rw-r--r--kernel/indtypes.ml5
-rw-r--r--kernel/inductive.ml5
-rw-r--r--kernel/mod_typing.ml97
-rw-r--r--kernel/mod_typing.mli2
-rw-r--r--kernel/modops.ml6
-rw-r--r--kernel/safe_typing.ml97
-rw-r--r--kernel/safe_typing.mli12
-rw-r--r--kernel/subtyping.ml13
-rw-r--r--kernel/term_typing.ml46
-rw-r--r--kernel/univ.ml127
-rw-r--r--kernel/univ.mli17
-rw-r--r--library/declare.ml25
-rw-r--r--library/global.ml8
-rw-r--r--library/global.mli8
-rw-r--r--library/lib.ml1
-rw-r--r--library/universes.ml68
-rw-r--r--plugins/cc/ccalgo.ml16
-rw-r--r--plugins/cc/ccalgo.mli2
-rw-r--r--plugins/cc/cctac.ml2
-rw-r--r--plugins/funind/functional_principles_types.ml20
-rw-r--r--plugins/funind/g_indfun.ml42
-rw-r--r--plugins/funind/glob_term_to_relation.ml95
-rw-r--r--plugins/funind/indfun.ml34
-rw-r--r--plugins/funind/indfun_common.ml7
-rw-r--r--plugins/funind/invfun.ml6
-rw-r--r--plugins/funind/recdef.ml7
-rw-r--r--plugins/setoid_ring/newring.ml11
-rw-r--r--pretyping/evarsolve.ml43
-rw-r--r--pretyping/pretyping.ml31
-rw-r--r--pretyping/typeclasses.ml2
-rw-r--r--proofs/pfedit.ml4
-rw-r--r--proofs/proof_global.ml6
-rw-r--r--stm/lemmas.ml11
-rw-r--r--tactics/autorewrite.ml5
-rw-r--r--tactics/equality.ml6
-rw-r--r--tactics/extratactics.ml415
-rw-r--r--tactics/hints.ml7
-rw-r--r--tactics/rewrite.ml24
-rw-r--r--tactics/tacticals.ml10
-rw-r--r--test-suite/Makefile2
-rw-r--r--test-suite/bugs/closed/2016.v62
-rw-r--r--test-suite/bugs/closed/2584.v89
-rw-r--r--test-suite/bugs/closed/3309.v334
-rw-r--r--test-suite/bugs/closed/3314.v8
-rw-r--r--test-suite/bugs/closed/3685.v (renamed from test-suite/bugs/opened/3685.v)2
-rw-r--r--test-suite/bugs/closed/3777.v16
-rw-r--r--test-suite/bugs/closed/4251.v17
-rw-r--r--test-suite/bugs/closed/4287.v124
-rw-r--r--test-suite/bugs/closed/4294.v31
-rw-r--r--test-suite/bugs/closed/4298.v7
-rw-r--r--test-suite/bugs/closed/4301.v12
-rw-r--r--test-suite/bugs/closed/4328.v6
-rw-r--r--test-suite/bugs/closed/HoTT_coq_053.v2
-rw-r--r--test-suite/bugs/closed/HoTT_coq_093.v2
-rw-r--r--test-suite/bugs/closed/HoTT_coq_108.v2
-rw-r--r--test-suite/bugs/closed/HoTT_coq_120.v (renamed from test-suite/bugs/opened/HoTT_coq_120.v)5
-rw-r--r--test-suite/failure/guard-cofix.v2
-rw-r--r--theories/Classes/CMorphisms.v3
-rw-r--r--toplevel/auto_ind_decl.ml10
-rw-r--r--toplevel/classes.ml15
-rw-r--r--toplevel/command.ml34
-rw-r--r--toplevel/indschemes.ml2
-rw-r--r--toplevel/obligations.ml25
-rw-r--r--toplevel/record.ml92
-rw-r--r--toplevel/vernacentries.ml8
82 files changed, 1352 insertions, 929 deletions
diff --git a/Makefile.build b/Makefile.build
index c554adadb..957592640 100644
--- a/Makefile.build
+++ b/Makefile.build
@@ -69,7 +69,7 @@ TIMED= # non-empty will activate a default time command
TIMECMD= # if you prefer a specific time command instead of $(STDTIME)
# e.g. "'time -p'"
-
+CAMLFLAGS:=${CAMLFLAGS} -w -3
# NB: if you want to collect compilation timings of .v and import them
# in a spreadsheet, I suggest something like:
# make TIMED=1 2> timings.csv
diff --git a/checker/cic.mli b/checker/cic.mli
index 881d3ca79..bd75111a2 100644
--- a/checker/cic.mli
+++ b/checker/cic.mli
@@ -380,7 +380,7 @@ and module_body =
(** algebraic type, kept if it's relevant for extraction *)
mod_type_alg : module_expression option;
(** set of all constraints in the module *)
- mod_constraints : Univ.constraints;
+ mod_constraints : Univ.ContextSet.t;
(** quotiented set of equivalent constants and inductive names *)
mod_delta : delta_resolver;
mod_retroknowledge : action list }
diff --git a/checker/environ.ml b/checker/environ.ml
index 6dbc44d6b..f8f5c29b7 100644
--- a/checker/environ.ml
+++ b/checker/environ.ml
@@ -84,13 +84,20 @@ let push_rec_types (lna,typarray,_) env =
Array.fold_left (fun e assum -> push_rel assum e) env ctxt
(* Universe constraints *)
-let add_constraints c env =
- if c == Univ.Constraint.empty then
- env
- else
- let s = env.env_stratification in
+let map_universes f env =
+ let s = env.env_stratification in
{ env with env_stratification =
- { s with env_universes = Univ.merge_constraints c s.env_universes } }
+ { s with env_universes = f s.env_universes } }
+
+let add_constraints c env =
+ if c == Univ.Constraint.empty then env
+ else map_universes (Univ.merge_constraints c) env
+
+let push_context ?(strict=false) ctx env =
+ map_universes (Univ.merge_context strict ctx) env
+
+let push_context_set ?(strict=false) ctx env =
+ map_universes (Univ.merge_context_set strict ctx) env
let check_constraints cst env =
Univ.check_constraints cst env.env_stratification.env_universes
diff --git a/checker/environ.mli b/checker/environ.mli
index f3b2dd839..87f143d1b 100644
--- a/checker/environ.mli
+++ b/checker/environ.mli
@@ -39,6 +39,8 @@ val push_rec_types : name array * constr array * 'a -> env -> env
(* Universes *)
val universes : env -> Univ.universes
val add_constraints : Univ.constraints -> env -> env
+val push_context : ?strict:bool -> Univ.universe_context -> env -> env
+val push_context_set : ?strict:bool -> Univ.universe_context_set -> env -> env
val check_constraints : Univ.constraints -> env -> bool
(* Constants *)
diff --git a/checker/mod_checking.ml b/checker/mod_checking.ml
index 998e23c6e..3ea5ed0d3 100644
--- a/checker/mod_checking.ml
+++ b/checker/mod_checking.ml
@@ -18,19 +18,27 @@ let refresh_arity ar =
let ctxt, hd = decompose_prod_assum ar in
match hd with
Sort (Type u) when not (Univ.is_univ_variable u) ->
- let u' = Univ.Universe.make (Univ.Level.make empty_dirpath 1) in
- mkArity (ctxt,Prop Null),
- Univ.enforce_leq u u' Univ.empty_constraint
- | _ -> ar, Univ.empty_constraint
+ let ul = Univ.Level.make empty_dirpath 1 in
+ let u' = Univ.Universe.make ul in
+ let cst = Univ.enforce_leq u u' Univ.empty_constraint in
+ let ctx = Univ.ContextSet.make (Univ.LSet.singleton ul) cst in
+ mkArity (ctxt,Prop Null), ctx
+ | _ -> ar, Univ.ContextSet.empty
let check_constant_declaration env kn cb =
Flags.if_verbose ppnl (str " checking cst: " ++ prcon kn); pp_flush ();
- let env' = add_constraints (Univ.UContext.constraints cb.const_universes) env in
+ let env' =
+ if cb.const_polymorphic then
+ let inst = Univ.make_abstract_instance cb.const_universes in
+ let ctx = Univ.UContext.make (inst, Univ.UContext.constraints cb.const_universes) in
+ push_context ~strict:false ctx env
+ else push_context ~strict:true cb.const_universes env
+ in
let envty, ty =
match cb.const_type with
RegularArity ty ->
let ty', cu = refresh_arity ty in
- let envty = add_constraints cu env' in
+ let envty = push_context_set cu env' in
let _ = infer_type envty ty' in envty, ty
| TemplateArity(ctxt,par) ->
let _ = check_ctxt env' ctxt in
@@ -69,7 +77,7 @@ let mk_mtb mp sign delta =
mod_expr = Abstract;
mod_type = sign;
mod_type_alg = None;
- mod_constraints = Univ.Constraint.empty;
+ mod_constraints = Univ.ContextSet.empty;
mod_delta = delta;
mod_retroknowledge = []; }
diff --git a/checker/modops.ml b/checker/modops.ml
index 8ccf118d3..7f07f8bf8 100644
--- a/checker/modops.ml
+++ b/checker/modops.ml
@@ -83,12 +83,13 @@ let strengthen_const mp_from l cb resolver =
| Def _ -> cb
| _ ->
let con = Constant.make2 mp_from l in
- (* let con = constant_of_delta resolver con in*)
let u =
- if cb.const_polymorphic then Univ.UContext.instance cb.const_universes
+ if cb.const_polymorphic then
+ Univ.make_abstract_instance cb.const_universes
else Univ.Instance.empty
in
- { cb with const_body = Def (Declarations.from_val (Const (con,u))) }
+ { cb with
+ const_body = Def (Declarations.from_val (Const (con,u))) }
let rec strengthen_mod mp_from mp_to mb =
if Declarations.mp_in_delta mb.mod_mp mb.mod_delta then mb
diff --git a/checker/reduction.ml b/checker/reduction.ml
index 8ddeea2a2..384d883ea 100644
--- a/checker/reduction.ml
+++ b/checker/reduction.ml
@@ -175,7 +175,7 @@ let sort_cmp env univ pb s0 s1 =
then begin
if !Flags.debug then begin
let op = match pb with CONV -> "=" | CUMUL -> "<=" in
- Printf.eprintf "cort_cmp: %s\n%!" Pp.(string_of_ppcmds
+ Printf.eprintf "sort_cmp: %s\n%!" Pp.(string_of_ppcmds
(str"Error: " ++ Univ.pr_uni u1 ++ str op ++ Univ.pr_uni u2 ++ str ":" ++ cut()
++ Univ.pr_universes univ))
end;
diff --git a/checker/safe_typing.ml b/checker/safe_typing.ml
index dd9482313..d3bc8373a 100644
--- a/checker/safe_typing.ml
+++ b/checker/safe_typing.ml
@@ -27,7 +27,7 @@ let set_engagement c =
(* full_add_module adds module with universes and constraints *)
let full_add_module dp mb univs digest =
let env = !genv in
- let env = add_constraints mb.mod_constraints env in
+ let env = push_context_set ~strict:true mb.mod_constraints env in
let env = add_constraints univs env in
let env = Modops.add_module mb env in
genv := add_digest env dp digest
@@ -84,7 +84,7 @@ let import file clib univs digest =
let mb = clib.comp_mod in
Mod_checking.check_module
(add_constraints univs
- (add_constraints mb.mod_constraints env)) mb.mod_mp mb;
+ (push_context_set ~strict:true mb.mod_constraints env)) mb.mod_mp mb;
stamp_library file digest;
full_add_module clib.comp_name mb univs digest
diff --git a/checker/univ.ml b/checker/univ.ml
index 3bcb3bc95..648e47817 100644
--- a/checker/univ.ml
+++ b/checker/univ.ml
@@ -244,7 +244,8 @@ module Level = struct
let set = make Set
let prop = make Prop
-
+ let var i = make (Var i)
+
let is_small x =
match data x with
| Level _ -> false
@@ -281,8 +282,8 @@ module Level = struct
end
(** Level sets and maps *)
-module LSet = Set.Make (Level)
-module LMap = Map.Make (Level)
+module LMap = HMap.Make (Level)
+module LSet = LMap.Set
type 'a universe_map = 'a LMap.t
@@ -559,20 +560,26 @@ let repr g u =
in
repr_rec u
-(* [safe_repr] also search for the canonical representative, but
- if the graph doesn't contain the searched universe, we add it. *)
-
-let safe_repr g u =
- let rec safe_repr_rec u =
- match UMap.find u g with
- | Equiv v -> safe_repr_rec v
- | Canonical arc -> arc
- in
- try g, safe_repr_rec u
- with Not_found ->
- let can = terminal u in
- enter_arc can g, can
+let get_set_arc g = repr g Level.set
+exception AlreadyDeclared
+
+let add_universe vlev strict g =
+ try
+ let _arcv = UMap.find vlev g in
+ raise AlreadyDeclared
+ with Not_found ->
+ let v = terminal vlev in
+ let arc =
+ let arc = get_set_arc g in
+ if strict then
+ { arc with lt=vlev::arc.lt}
+ else
+ { arc with le=vlev::arc.le}
+ in
+ let g = enter_arc arc g in
+ enter_arc v g
+
(* reprleq : canonical_arc -> canonical_arc list *)
(* All canonical arcv such that arcu<=arcv with arcv#arcu *)
let reprleq g arcu =
@@ -739,8 +746,8 @@ let is_lt g arcu arcv =
(** First, checks on universe levels *)
let check_equal g u v =
- let g, arcu = safe_repr g u in
- let _, arcv = safe_repr g v in
+ let arcu = repr g u in
+ let arcv = repr g v in
arcu == arcv
let check_eq_level g u v = u == v || check_equal g u v
@@ -749,8 +756,8 @@ let is_set_arc u = Level.is_set u.univ
let is_prop_arc u = Level.is_prop u.univ
let check_smaller g strict u v =
- let g, arcu = safe_repr g u in
- let g, arcv = safe_repr g v in
+ let arcu = repr g u in
+ let arcv = repr g v in
if strict then
is_lt g arcu arcv
else
@@ -900,8 +907,8 @@ let error_inconsistency o u v =
(* enforc_univ_eq u v will force u=v if possible, will fail otherwise *)
let enforce_univ_eq u v g =
- let g,arcu = safe_repr g u in
- let g,arcv = safe_repr g v in
+ let arcu = repr g u in
+ let arcv = repr g v in
match fast_compare g arcu arcv with
| FastEQ -> g
| FastLT -> error_inconsistency Eq v u
@@ -916,8 +923,8 @@ let enforce_univ_eq u v g =
(* enforce_univ_leq : Level.t -> Level.t -> unit *)
(* enforce_univ_leq u v will force u<=v if possible, will fail otherwise *)
let enforce_univ_leq u v g =
- let g,arcu = safe_repr g u in
- let g,arcv = safe_repr g v in
+ let arcu = repr g u in
+ let arcv = repr g v in
if is_leq g arcu arcv then g
else
match fast_compare g arcv arcu with
@@ -928,8 +935,8 @@ let enforce_univ_leq u v g =
(* enforce_univ_lt u v will force u<v if possible, will fail otherwise *)
let enforce_univ_lt u v g =
- let g,arcu = safe_repr g u in
- let g,arcv = safe_repr g v in
+ let arcu = repr g u in
+ let arcv = repr g v in
match fast_compare g arcu arcv with
| FastLT -> g
| FastLE -> fst (setlt g arcu arcv)
@@ -941,7 +948,10 @@ let enforce_univ_lt u v g =
| FastLE | FastLT -> error_inconsistency Lt u v
(* Prop = Set is forbidden here. *)
-let initial_universes = enforce_univ_lt Level.prop Level.set UMap.empty
+let initial_universes =
+ let g = enter_arc (terminal Level.set) UMap.empty in
+ let g = enter_arc (terminal Level.prop) g in
+ enforce_univ_lt Level.prop Level.set g
(* Constraints and sets of constraints. *)
@@ -970,7 +980,7 @@ module Constraint = Set.Make(UConstraintOrd)
let empty_constraint = Constraint.empty
let merge_constraints c g =
Constraint.fold enforce_constraint c g
-
+
type constraints = Constraint.t
(** A value with universe constraints. *)
@@ -1146,7 +1156,7 @@ struct
(** Universe contexts (variables as a list) *)
let empty = (Instance.empty, Constraint.empty)
-
+ let make x = x
let instance (univs, cst) = univs
let constraints (univs, cst) = cst
end
@@ -1158,6 +1168,8 @@ struct
type t = LSet.t constrained
let empty = LSet.empty, Constraint.empty
let constraints (_, cst) = cst
+ let levels (ctx, _) = ctx
+ let make ctx cst = (ctx, cst)
end
type universe_context_set = ContextSet.t
@@ -1207,6 +1219,9 @@ let subst_instance_constraints s csts =
(fun c csts -> Constraint.add (subst_instance_constraint s c) csts)
csts Constraint.empty
+let make_abstract_instance (ctx, _) =
+ Array.mapi (fun i l -> Level.var i) ctx
+
(** Substitute instance inst for ctx in csts *)
let instantiate_univ_context (ctx, csts) =
(ctx, subst_instance_constraints ctx csts)
@@ -1238,6 +1253,20 @@ let subst_univs_universe fn ul =
List.fold_left (fun acc u -> Universe.merge_univs acc (Universe.Huniv.tip u))
substs nosubst
+let merge_context strict ctx g =
+ let g = Array.fold_left
+ (* Be lenient, module typing reintroduces universes and
+ constraints due to includes *)
+ (fun g v -> try add_universe v strict g with AlreadyDeclared -> g)
+ g (UContext.instance ctx)
+ in merge_constraints (UContext.constraints ctx) g
+
+let merge_context_set strict ctx g =
+ let g = LSet.fold
+ (fun v g -> try add_universe v strict g with AlreadyDeclared -> g)
+ (ContextSet.levels ctx) g
+ in merge_constraints (ContextSet.constraints ctx) g
+
(** Pretty-printing *)
let pr_arc = function
diff --git a/checker/univ.mli b/checker/univ.mli
index 742ef91ae..02c1bbdb9 100644
--- a/checker/univ.mli
+++ b/checker/univ.mli
@@ -74,6 +74,13 @@ val check_eq : universe check_function
(** The initial graph of universes: Prop < Set *)
val initial_universes : universes
+(** Adds a universe to the graph, ensuring it is >= or > Set.
+ @raises AlreadyDeclared if the level is already declared in the graph. *)
+
+exception AlreadyDeclared
+
+val add_universe : universe_level -> bool -> universes -> universes
+
(** {6 Constraints. } *)
type constraint_type = Lt | Le | Eq
@@ -117,14 +124,14 @@ type univ_inconsistency = constraint_type * universe * universe
exception UniverseInconsistency of univ_inconsistency
val merge_constraints : constraints -> universes -> universes
-
+
val check_constraints : constraints -> universes -> bool
(** {6 Support for universe polymorphism } *)
(** Polymorphic maps from universe levels to 'a *)
module LMap : Map.S with type key = universe_level
-
+module LSet : CSig.SetS with type elt = universe_level
type 'a universe_map = 'a LMap.t
(** {6 Substitution} *)
@@ -177,7 +184,7 @@ sig
type t
val empty : t
-
+ val make : universe_instance constrained -> t
val instance : t -> Instance.t
val constraints : t -> constraints
@@ -186,6 +193,7 @@ end
module ContextSet :
sig
type t
+ val make : LSet.t -> constraints -> t
val empty : t
val constraints : t -> constraints
end
@@ -193,6 +201,9 @@ module ContextSet :
type universe_context = UContext.t
type universe_context_set = ContextSet.t
+val merge_context : bool -> universe_context -> universes -> universes
+val merge_context_set : bool -> universe_context_set -> universes -> universes
+
val empty_level_subst : universe_level_subst
val is_empty_level_subst : universe_level_subst -> bool
@@ -219,6 +230,9 @@ val subst_instance_constraints : universe_instance -> constraints -> constraints
val instantiate_univ_context : universe_context -> universe_context
val instantiate_univ_constraints : universe_instance -> universe_context -> constraints
+(** Build the relative instance corresponding to the context *)
+val make_abstract_instance : universe_context -> universe_instance
+
(** {6 Pretty-printing of universes. } *)
val pr_universes : universes -> Pp.std_ppcmds
diff --git a/checker/values.ml b/checker/values.ml
index 45220bd05..34de511c8 100644
--- a/checker/values.ml
+++ b/checker/values.ml
@@ -13,7 +13,7 @@
To ensure this file is up-to-date, 'make' now compares the md5 of cic.mli
with a copy we maintain here:
-MD5 8b7e75b4b94a2d8506a62508e0374c0a checker/cic.mli
+MD5 76312d06933f47498a1981a6261c9f75 checker/cic.mli
*)
@@ -307,10 +307,10 @@ and v_impl =
and v_noimpl = v_enum "no_impl" 1 (* Abstract is mandatory for mtb *)
and v_module =
Tuple ("module_body",
- [|v_mp;v_impl;v_sign;Opt v_mexpr;v_cstrs;v_resolver;Any|])
+ [|v_mp;v_impl;v_sign;Opt v_mexpr;v_context_set;v_resolver;Any|])
and v_modtype =
Tuple ("module_type_body",
- [|v_mp;v_noimpl;v_sign;Opt v_mexpr;v_cstrs;v_resolver;Any|])
+ [|v_mp;v_noimpl;v_sign;Opt v_mexpr;v_context_set;v_resolver;Any|])
(** kernel/safe_typing *)
diff --git a/dev/doc/univpoly.txt b/dev/doc/univpoly.txt
index 4c89af01d..bad2ae36e 100644
--- a/dev/doc/univpoly.txt
+++ b/dev/doc/univpoly.txt
@@ -1,5 +1,5 @@
-Notes on universe polymorphism and primitive projections, M. Sozeau - WIP
-=========================================================================
+Notes on universe polymorphism and primitive projections, M. Sozeau
+===================================================================
The new implementation of universe polymorphism and primitive
projections introduces a few changes to the API of Coq. First and
@@ -46,15 +46,16 @@ universes and constraints to the global universe context when it is put
in the environment. No other universes than the global ones and the
declared local ones are needed to check a declaration, hence the kernel
does not produce any constraints anymore, apart from module
-subtyping.... There are hance two conversion functions now: check_conv
-and infer_conv: the former just checks the definition in the current env
+subtyping.... There are hence two conversion functions now: [check_conv]
+and [infer_conv]: the former just checks the definition in the current env
(in which we usually push_universe_context of the associated context),
-and infer_conv which produces constraints that were not implied by the
+and [infer_conv] which produces constraints that were not implied by the
ambient constraints. Ideally, that one could be put out of the kernel,
-but again, module subtyping needs it.
+but currently module subtyping needs it.
Inference of universes is now done during refinement, and the evar_map
-carries the incrementally built universe context. [Evd.conversion] is a
+carries the incrementally built universe context, starting from the
+global universe constraints (see [Evd.from_env]). [Evd.conversion] is a
wrapper around [infer_conv] that will do the bookkeeping for you, it
uses [evar_conv_x]. There is a universe substitution being built
incrementally according to the constraints, so one should normalize at
@@ -72,7 +73,7 @@ val pf_constr_of_global : Globnames.global_reference -> (constr -> tactic) -> ta
Is the way to make a constr out of a global reference in the new API.
If they constr is polymorphic, it will add the necessary constraints to
the evar_map. Even if a constr is not polymorphic, we have to take care
-of keeping track of it's universes. Typically, using:
+of keeping track of its universes. Typically, using:
mkApp (coq_id_function, [| A; a |])
@@ -84,8 +85,8 @@ produce the right constraints and put them in the evar_map. Of course in
some cases you might now from an invariant that no new constraint would
be produced and get rid of it. Anyway the kernel will tell you if you
forgot some. As a temporary way out, [Universes.constr_of_global] allows
-you to make a constr from any non-polymorphic constant, but it might
-forget constraints.
+you to make a constr from any non-polymorphic constant, but it will fail
+on polymorphic ones.
Other than that, unification (w_unify and evarconv) now take account of universes and
produce only well-typed evar_maps.
@@ -157,6 +158,30 @@ this is the only solution I found. In the case of global_references
only, it's just a matter of using [Evd.fresh_global] /
[pf_constr_of_global] to let the system take care of universes.
+
+The universe graph
+==================
+
+To accomodate universe polymorphic definitions, the graph structure in
+kernel/univ.ml was modified. The new API forces every universe to be
+declared before it is mentionned in any constraint. This forces to
+declare every universe to be >= Set or > Set. Every universe variable
+introduced during elaboration is >= Set. Every _global_ universe is now
+declared explicitely > Set, _after_ typechecking the definition. In
+polymorphic definitions Type@{i} ranges over Set and any other universe
+j. However, at instantiation time for polymorphic references, one can
+try to instantiate a universe parameter with Prop as well, if the
+instantiated constraints allow it. The graph invariants ensure that
+no universe i can be set lower than Set, so the chain of universes
+always bottoms down at Prop < Set.
+
+Modules
+=======
+
+One has to think of universes in modules as being globally declared, so
+when including a module (type) which declares a type i (e.g. through a
+parameter), we get back a copy of i and not some fresh universe.
+
Projections
===========
@@ -208,8 +233,7 @@ constants left (the most common case). E.g. Ring with Set Universe
Polymorphism and Set Primitive Projections work (at least it did at some
point, I didn't recheck yet).
-- [native_compute] is untested: it should deal with primitive
-projections right but not universes.
+- [native_compute] works with universes and projections.
Incompatibilities
diff --git a/engine/evd.ml b/engine/evd.ml
index 2fdf3699f..574be9844 100644
--- a/engine/evd.ml
+++ b/engine/evd.ml
@@ -310,6 +310,9 @@ let union_evar_universe_context ctx ctx' =
else
let local = Univ.ContextSet.union ctx.uctx_local ctx'.uctx_local in
let names = UNameMap.union (fst ctx.uctx_names) (fst ctx'.uctx_names) in
+ let newus = Univ.LSet.diff (Univ.ContextSet.levels ctx'.uctx_local)
+ (Univ.ContextSet.levels ctx.uctx_local) in
+ let declarenew g = Univ.LSet.fold (fun u g -> Univ.add_universe u false g) newus g in
let names_rev = Univ.LMap.union (snd ctx.uctx_names) (snd ctx'.uctx_names) in
{ uctx_names = (names, names_rev);
uctx_local = local;
@@ -317,12 +320,12 @@ let union_evar_universe_context ctx ctx' =
Univ.LMap.subst_union ctx.uctx_univ_variables ctx'.uctx_univ_variables;
uctx_univ_algebraic =
Univ.LSet.union ctx.uctx_univ_algebraic ctx'.uctx_univ_algebraic;
- uctx_initial_universes = ctx.uctx_initial_universes;
+ uctx_initial_universes = declarenew ctx.uctx_initial_universes;
uctx_universes =
if local == ctx.uctx_local then ctx.uctx_universes
else
let cstrsr = Univ.ContextSet.constraints ctx'.uctx_local in
- Univ.merge_constraints cstrsr ctx.uctx_universes }
+ Univ.merge_constraints cstrsr (declarenew ctx.uctx_universes) }
(* let union_evar_universe_context_key = Profile.declare_profile "union_evar_universe_context";; *)
(* let union_evar_universe_context = *)
@@ -385,7 +388,7 @@ let process_universe_constraints univs vars alg cstrs =
let levels = Univ.Universe.levels l in
Univ.LSet.fold (fun l local ->
if Univ.Level.is_small l || Univ.LMap.mem l !vars then
- Univ.enforce_eq (Univ.Universe.make l) r local
+ unify_universes fo (Univ.Universe.make l) Universes.UEq r local
else raise (Univ.UniverseInconsistency (Univ.Le, Univ.Universe.make l, r, None)))
levels local
else
@@ -453,7 +456,7 @@ let add_constraints_context ctx cstrs =
in
{ ctx with uctx_local = (univs, Univ.Constraint.union local local');
uctx_univ_variables = vars;
- uctx_universes = Univ.merge_constraints cstrs ctx.uctx_universes }
+ uctx_universes = Univ.merge_constraints local' ctx.uctx_universes }
(* let addconstrkey = Profile.declare_profile "add_constraints_context";; *)
(* let add_constraints_context = Profile.profile2 addconstrkey add_constraints_context;; *)
@@ -929,38 +932,6 @@ let evars_of_filtered_evar_info evi =
(evars_of_named_context (evar_filtered_context evi)))
(**********************************************************)
-(* Side effects *)
-
-let emit_side_effects eff evd =
- { evd with effects = Declareops.union_side_effects eff evd.effects; }
-
-let drop_side_effects evd =
- { evd with effects = Declareops.no_seff; }
-
-let eval_side_effects evd = evd.effects
-
-(* Future goals *)
-let declare_future_goal evk evd =
- { evd with future_goals = evk::evd.future_goals }
-
-let declare_principal_goal evk evd =
- match evd.principal_future_goal with
- | None -> { evd with
- future_goals = evk::evd.future_goals;
- principal_future_goal=Some evk; }
- | Some _ -> Errors.error "Only one main subgoal per instantiation."
-
-let future_goals evd = evd.future_goals
-
-let principal_future_goal evd = evd.principal_future_goal
-
-let reset_future_goals evd =
- { evd with future_goals = [] ; principal_future_goal=None }
-
-let restore_future_goals evd gls pgl =
- { evd with future_goals = gls ; principal_future_goal = pgl }
-
-(**********************************************************)
(* Sort variables *)
type rigid =
@@ -1015,13 +986,13 @@ let restrict_universe_context evd vars =
let universe_subst evd =
evd.universes.uctx_univ_variables
-let merge_uctx rigid uctx ctx' =
+let merge_uctx sideff rigid uctx ctx' =
let open Univ in
- let uctx =
+ let levels = ContextSet.levels ctx' in
+ let uctx = if sideff then uctx else
match rigid with
| UnivRigid -> uctx
| UnivFlexible b ->
- let levels = ContextSet.levels ctx' in
let fold u accu =
if LMap.mem u accu then accu
else LMap.add u None accu
@@ -1032,12 +1003,23 @@ let merge_uctx rigid uctx ctx' =
uctx_univ_algebraic = LSet.union uctx.uctx_univ_algebraic levels }
else { uctx with uctx_univ_variables = uvars' }
in
- let uctx_local = ContextSet.append ctx' uctx.uctx_local in
- let uctx_universes = merge_constraints (ContextSet.constraints ctx') uctx.uctx_universes in
- { uctx with uctx_local; uctx_universes }
+ let uctx_local =
+ if sideff then uctx.uctx_local
+ else ContextSet.append ctx' uctx.uctx_local
+ in
+ let declare g =
+ LSet.fold (fun u g ->
+ try Univ.add_universe u false g
+ with Univ.AlreadyDeclared when sideff -> g)
+ levels g
+ in
+ let initial = declare uctx.uctx_initial_universes in
+ let univs = declare uctx.uctx_universes in
+ let uctx_universes = merge_constraints (ContextSet.constraints ctx') univs in
+ { uctx with uctx_local; uctx_universes; uctx_initial_universes = initial }
let merge_context_set rigid evd ctx' =
- {evd with universes = merge_uctx rigid evd.universes ctx'}
+ {evd with universes = merge_uctx false rigid evd.universes ctx'}
let merge_uctx_subst uctx s =
{ uctx with uctx_univ_variables = Univ.LMap.subst_union uctx.uctx_univ_variables s }
@@ -1048,41 +1030,84 @@ let merge_universe_subst evd subst =
let with_context_set rigid d (a, ctx) =
(merge_context_set rigid d ctx, a)
+let emit_universe_side_effects eff u =
+ Declareops.fold_side_effects
+ (fun acc eff ->
+ match eff with
+ | Declarations.SEscheme (l,s) ->
+ List.fold_left
+ (fun acc (_,_,cb,c) ->
+ let acc = match c with
+ | `Nothing -> acc
+ | `Opaque (s, ctx) -> merge_uctx true univ_rigid acc ctx
+ in if cb.Declarations.const_polymorphic then acc
+ else
+ merge_uctx true univ_rigid acc
+ (Univ.ContextSet.of_context cb.Declarations.const_universes))
+ acc l
+ | Declarations.SEsubproof _ -> acc)
+ u eff
+
let add_uctx_names s l (names, names_rev) =
(UNameMap.add s l names, Univ.LMap.add l s names_rev)
-let uctx_new_univ_variable rigid name
+let uctx_new_univ_variable rigid name predicative
({ uctx_local = ctx; uctx_univ_variables = uvars; uctx_univ_algebraic = avars} as uctx) =
let u = Universes.new_univ_level (Global.current_dirpath ()) in
let ctx' = Univ.ContextSet.add_universe u ctx in
- let uctx' =
+ let uctx', pred =
match rigid with
- | UnivRigid -> uctx
+ | UnivRigid -> uctx, true
| UnivFlexible b ->
let uvars' = Univ.LMap.add u None uvars in
if b then {uctx with uctx_univ_variables = uvars';
- uctx_univ_algebraic = Univ.LSet.add u avars}
- else {uctx with uctx_univ_variables = Univ.LMap.add u None uvars} in
+ uctx_univ_algebraic = Univ.LSet.add u avars}, false
+ else {uctx with uctx_univ_variables = uvars'}, false
+ in
+ (* let ctx' = *)
+ (* if pred then *)
+ (* Univ.ContextSet.add_constraints *)
+ (* (Univ.Constraint.singleton (Univ.Level.set, Univ.Le, u)) ctx' *)
+ (* else ctx' *)
+ (* in *)
let names =
match name with
| Some n -> add_uctx_names n u uctx.uctx_names
| None -> uctx.uctx_names
in
+ let initial =
+ Univ.add_universe u false uctx.uctx_initial_universes
+ in
+ let uctx' =
{uctx' with uctx_names = names; uctx_local = ctx';
- uctx_universes = Univ.add_universe u uctx.uctx_universes}, u
-
-let new_univ_level_variable ?name rigid evd =
- let uctx', u = uctx_new_univ_variable rigid name evd.universes in
+ uctx_universes = Univ.add_universe u false uctx.uctx_universes;
+ uctx_initial_universes = initial}
+ in uctx', u
+
+let new_univ_level_variable ?name ?(predicative=true) rigid evd =
+ let uctx', u = uctx_new_univ_variable rigid name predicative evd.universes in
({evd with universes = uctx'}, u)
-let new_univ_variable ?name rigid evd =
- let uctx', u = uctx_new_univ_variable rigid name evd.universes in
+let new_univ_variable ?name ?(predicative=true) rigid evd =
+ let uctx', u = uctx_new_univ_variable rigid name predicative evd.universes in
({evd with universes = uctx'}, Univ.Universe.make u)
-let new_sort_variable ?name rigid d =
- let (d', u) = new_univ_variable rigid ?name d in
+let new_sort_variable ?name ?(predicative=true) rigid d =
+ let (d', u) = new_univ_variable rigid ?name ~predicative d in
(d', Type u)
+let add_global_univ d u =
+ let uctx = d.universes in
+ let initial =
+ Univ.add_universe u true uctx.uctx_initial_universes
+ in
+ let univs =
+ Univ.add_universe u true uctx.uctx_universes
+ in
+ { d with universes = { uctx with uctx_local = Univ.ContextSet.add_universe u uctx.uctx_local;
+ uctx_initial_universes = initial;
+ uctx_universes = univs } }
+
let make_flexible_variable evd b u =
let {uctx_univ_variables = uvars; uctx_univ_algebraic = avars} as ctx = evd.universes in
let uvars' = Univ.LMap.add u None uvars in
@@ -1104,9 +1129,10 @@ let make_evar_universe_context e l =
match l with
| None -> uctx
| Some us ->
- List.fold_left (fun uctx (loc,id) ->
- fst (uctx_new_univ_variable univ_rigid (Some (Id.to_string id)) uctx))
- uctx us
+ List.fold_left
+ (fun uctx (loc,id) ->
+ fst (uctx_new_univ_variable univ_rigid (Some (Id.to_string id)) true uctx))
+ uctx us
(****************************************)
(* Operations on constants *)
@@ -1271,12 +1297,16 @@ let refresh_undefined_univ_variables uctx =
Univ.LMap.add (Univ.subst_univs_level_level subst u)
(Option.map (Univ.subst_univs_level_universe subst) v) acc)
uctx.uctx_univ_variables Univ.LMap.empty
- in
+ in
+ let declare g = Univ.LSet.fold (fun u g -> Univ.add_universe u false g)
+ (Univ.ContextSet.levels ctx') g in
+ let initial = declare uctx.uctx_initial_universes in
+ let univs = declare Univ.initial_universes in
let uctx' = {uctx_names = uctx.uctx_names;
uctx_local = ctx';
uctx_univ_variables = vars; uctx_univ_algebraic = alg;
- uctx_universes = Univ.initial_universes;
- uctx_initial_universes = uctx.uctx_initial_universes } in
+ uctx_universes = univs;
+ uctx_initial_universes = initial } in
uctx', subst
let refresh_undefined_universes evd =
@@ -1363,6 +1393,39 @@ let e_eq_constr_univs evdref t u =
evdref := evd; b
(**********************************************************)
+(* Side effects *)
+
+let emit_side_effects eff evd =
+ { evd with effects = Declareops.union_side_effects eff evd.effects;
+ universes = emit_universe_side_effects eff evd.universes }
+
+let drop_side_effects evd =
+ { evd with effects = Declareops.no_seff; }
+
+let eval_side_effects evd = evd.effects
+
+(* Future goals *)
+let declare_future_goal evk evd =
+ { evd with future_goals = evk::evd.future_goals }
+
+let declare_principal_goal evk evd =
+ match evd.principal_future_goal with
+ | None -> { evd with
+ future_goals = evk::evd.future_goals;
+ principal_future_goal=Some evk; }
+ | Some _ -> Errors.error "Only one main subgoal per instantiation."
+
+let future_goals evd = evd.future_goals
+
+let principal_future_goal evd = evd.principal_future_goal
+
+let reset_future_goals evd =
+ { evd with future_goals = [] ; principal_future_goal=None }
+
+let restore_future_goals evd gls pgl =
+ { evd with future_goals = gls ; principal_future_goal = pgl }
+
+(**********************************************************)
(* Accessing metas *)
(** We use this function to overcome OCaml compiler limitations and to prevent
diff --git a/engine/evd.mli b/engine/evd.mli
index bc81bd818..db60f5ff4 100644
--- a/engine/evd.mli
+++ b/engine/evd.mli
@@ -501,9 +501,11 @@ val normalize_evar_universe_context_variables : evar_universe_context ->
val normalize_evar_universe_context : evar_universe_context ->
evar_universe_context
-val new_univ_level_variable : ?name:string -> rigid -> evar_map -> evar_map * Univ.universe_level
-val new_univ_variable : ?name:string -> rigid -> evar_map -> evar_map * Univ.universe
-val new_sort_variable : ?name:string -> rigid -> evar_map -> evar_map * sorts
+val new_univ_level_variable : ?name:string -> ?predicative:bool -> rigid -> evar_map -> evar_map * Univ.universe_level
+val new_univ_variable : ?name:string -> ?predicative:bool -> rigid -> evar_map -> evar_map * Univ.universe
+val new_sort_variable : ?name:string -> ?predicative:bool -> rigid -> evar_map -> evar_map * sorts
+val add_global_univ : evar_map -> Univ.Level.t -> evar_map
+
val make_flexible_variable : evar_map -> bool -> Univ.universe_level -> evar_map
val is_sort_variable : evar_map -> sorts -> Univ.universe_level option
(** [is_sort_variable evm s] returns [Some u] or [None] if [s] is
diff --git a/kernel/declarations.mli b/kernel/declarations.mli
index 561c66b42..7def963e7 100644
--- a/kernel/declarations.mli
+++ b/kernel/declarations.mli
@@ -246,8 +246,8 @@ and module_body =
mod_type : module_signature; (** expanded type *)
(** algebraic type, kept if it's relevant for extraction *)
mod_type_alg : module_expression option;
- (** set of all constraints in the module *)
- mod_constraints : Univ.constraints;
+ (** set of all universes constraints in the module *)
+ mod_constraints : Univ.ContextSet.t;
(** quotiented set of equivalent constants and inductive names *)
mod_delta : Mod_subst.delta_resolver;
mod_retroknowledge : Retroknowledge.action list }
diff --git a/kernel/environ.ml b/kernel/environ.ml
index bf12d6c6d..1cc07c0ab 100644
--- a/kernel/environ.ml
+++ b/kernel/environ.ml
@@ -181,26 +181,44 @@ let fold_named_context_reverse f ~init env =
(* Universe constraints *)
-let add_constraints c env =
- if Univ.Constraint.is_empty c then
- env
- else
- let s = env.env_stratification in
+let map_universes f env =
+ let s = env.env_stratification in
{ env with env_stratification =
- { s with env_universes = Univ.merge_constraints c s.env_universes } }
+ { s with env_universes = f s.env_universes } }
+
+let add_constraints c env =
+ if Univ.Constraint.is_empty c then env
+ else map_universes (Univ.merge_constraints c) env
let check_constraints c env =
Univ.check_constraints c env.env_stratification.env_universes
-let set_engagement c env = (* Unsafe *)
- { env with env_stratification =
- { env.env_stratification with env_engagement = c } }
-
let push_constraints_to_env (_,univs) env =
add_constraints univs env
-let push_context ctx env = add_constraints (Univ.UContext.constraints ctx) env
-let push_context_set ctx env = add_constraints (Univ.ContextSet.constraints ctx) env
+let add_universes strict ctx g =
+ let g = Array.fold_left
+ (* Be lenient, module typing reintroduces universes and constraints due to includes *)
+ (fun g v -> try Univ.add_universe v strict g with Univ.AlreadyDeclared -> g)
+ g (Univ.Instance.to_array (Univ.UContext.instance ctx))
+ in
+ Univ.merge_constraints (Univ.UContext.constraints ctx) g
+
+let push_context ?(strict=false) ctx env =
+ map_universes (add_universes strict ctx) env
+
+let add_universes_set strict ctx g =
+ let g = Univ.LSet.fold
+ (fun v g -> try Univ.add_universe v strict g with Univ.AlreadyDeclared -> g)
+ (Univ.ContextSet.levels ctx) g
+ in Univ.merge_constraints (Univ.ContextSet.constraints ctx) g
+
+let push_context_set ?(strict=false) ctx env =
+ map_universes (add_universes_set strict ctx) env
+
+let set_engagement c env = (* Unsafe *)
+ { env with env_stratification =
+ { env.env_stratification with env_engagement = c } }
(* Global constants *)
diff --git a/kernel/environ.mli b/kernel/environ.mli
index 4ad0327fc..9f6ea522a 100644
--- a/kernel/environ.mli
+++ b/kernel/environ.mli
@@ -208,8 +208,8 @@ val add_constraints : Univ.constraints -> env -> env
(** Check constraints are satifiable in the environment. *)
val check_constraints : Univ.constraints -> env -> bool
-val push_context : Univ.universe_context -> env -> env
-val push_context_set : Univ.universe_context_set -> env -> env
+val push_context : ?strict:bool -> Univ.universe_context -> env -> env
+val push_context_set : ?strict:bool -> Univ.universe_context_set -> env -> env
val push_constraints_to_env : 'a Univ.constrained -> env -> env
val set_engagement : engagement -> env -> env
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index e3457006d..5a234d09b 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -291,7 +291,10 @@ let typecheck_inductive env mie =
let defu = Term.univ_of_sort def_level in
let is_natural =
type_in_type env || (check_leq (universes env') infu defu &&
- not (is_type0m_univ defu && not is_unit))
+ not (is_type0m_univ defu && not is_unit)
+ (* (~ is_type0m_univ defu \/ is_unit) (\* infu <= defu && not prop or unital *\) *)
+
+ )
in
let _ =
(** Impredicative sort, always allow *)
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index 87c139f48..a02d5e205 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -165,7 +165,10 @@ let rec make_subst env =
(* to be greater than the level of the argument; this is probably *)
(* a useless extra constraint *)
let s = sort_as_univ (snd (dest_arity env (Lazy.force a))) in
- make (cons_subst u s subst) (sign, exp, args)
+ if Univ.Universe.is_levels s then
+ make (cons_subst u s subst) (sign, exp, args)
+ else (* Cannot handle substitution by i+n universes. *)
+ make subst (sign, exp, args)
| (na,None,t)::sign, Some u::exp, [] ->
(* No more argument here: we add the remaining universes to the *)
(* substitution (when [u] is distinct from all other universes in the *)
diff --git a/kernel/mod_typing.ml b/kernel/mod_typing.ml
index 4f20e5f62..922652287 100644
--- a/kernel/mod_typing.ml
+++ b/kernel/mod_typing.ml
@@ -21,7 +21,7 @@ open Modops
open Mod_subst
type 'alg translation =
- module_signature * 'alg option * delta_resolver * Univ.constraints
+ module_signature * 'alg option * delta_resolver * Univ.ContextSet.t
let rec mp_from_mexpr = function
| MEident mp -> mp
@@ -52,7 +52,7 @@ let rec rebuild_mp mp l =
| []-> mp
| i::r -> rebuild_mp (MPdot(mp,Label.of_id i)) r
-let (+++) = Univ.Constraint.union
+let (+++) = Univ.ContextSet.union
let rec check_with_def env struc (idl,(c,ctx)) mp equiv =
let lab,idl = match idl with
@@ -72,33 +72,65 @@ let rec check_with_def env struc (idl,(c,ctx)) mp equiv =
(* In the spirit of subtyping.check_constant, we accept
any implementations of parameters and opaques terms,
as long as they have the right type *)
- let ccst = Declareops.constraints_of_constant (opaque_tables env) cb in
- let env' = Environ.add_constraints ccst env' in
- let newus, cst = Univ.UContext.dest ctx in
- let env' = Environ.add_constraints cst env' in
- let c',cst = match cb.const_body with
- | Undef _ | OpaqueDef _ ->
- let j = Typeops.infer env' c in
- let typ = Typeops.type_of_constant_type env' cb.const_type in
- let cst' = Reduction.infer_conv_leq env' (Environ.universes env')
- j.uj_type typ in
- j.uj_val,cst' +++ cst
- | Def cs ->
- let cst' = Reduction.infer_conv env' (Environ.universes env') c
- (Mod_subst.force_constr cs) in
- let cst = (*FIXME MS: what to check here? subtyping of polymorphic constants... *)
- if cb.const_polymorphic then cst' +++ cst
- else cst' +++ cst
+ let uctx = Declareops.universes_of_constant (opaque_tables env) cb in
+ let uctx = (* Context of the spec *)
+ if cb.const_polymorphic then
+ Univ.instantiate_univ_context uctx
+ else uctx
+ in
+ let c', univs, ctx' =
+ if not cb.const_polymorphic then
+ let env' = Environ.push_context ~strict:true uctx env' in
+ let env' = Environ.push_context ~strict:true ctx env' in
+ let c',cst = match cb.const_body with
+ | Undef _ | OpaqueDef _ ->
+ let j = Typeops.infer env' c in
+ let typ = Typeops.type_of_constant_type env' cb.const_type in
+ let cst' = Reduction.infer_conv_leq env' (Environ.universes env')
+ j.uj_type typ in
+ j.uj_val, cst'
+ | Def cs ->
+ let c' = Mod_subst.force_constr cs in
+ c, Reduction.infer_conv env' (Environ.universes env') c c'
+ in c', ctx, Univ.ContextSet.add_constraints cst (Univ.ContextSet.of_context ctx)
+ else
+ let cus, ccst = Univ.UContext.dest uctx in
+ let newus, cst = Univ.UContext.dest ctx in
+ let () =
+ if not (Univ.Instance.length cus == Univ.Instance.length newus) then
+ error_incorrect_with_constraint lab
+ in
+ let inst = Univ.Instance.append cus newus in
+ let csti = Univ.enforce_eq_instances cus newus cst in
+ let csta = Univ.Constraint.union csti ccst in
+ let env' = Environ.push_context ~strict:false (Univ.UContext.make (inst, csta)) env in
+ let () = if not (Univ.check_constraints cst (Environ.universes env')) then
+ error_incorrect_with_constraint lab
+ in
+ let cst = match cb.const_body with
+ | Undef _ | OpaqueDef _ ->
+ let j = Typeops.infer env' c in
+ let typ = Typeops.type_of_constant_type env' cb.const_type in
+ let typ = Vars.subst_instance_constr cus typ in
+ let cst' = Reduction.infer_conv_leq env' (Environ.universes env')
+ j.uj_type typ in
+ cst'
+ | Def cs ->
+ let c' = Vars.subst_instance_constr cus (Mod_subst.force_constr cs) in
+ let cst' = Reduction.infer_conv env' (Environ.universes env') c c' in
+ cst'
in
- c, cst
+ if not (Univ.Constraint.is_empty cst) then
+ error_incorrect_with_constraint lab;
+ let subst, ctx = Univ.abstract_universes true ctx in
+ Vars.subst_univs_level_constr subst c, ctx, Univ.ContextSet.empty
in
let def = Def (Mod_subst.from_val c') in
- let ctx' = Univ.UContext.make (newus, cst) in
let cb' =
{ cb with
const_body = def;
- const_body_code = Option.map Cemitcodes.from_val (compile_constant_body env' def);
- const_universes = ctx' }
+ const_universes = univs;
+ const_body_code = Option.map Cemitcodes.from_val (compile_constant_body env' def) }
in
before@(lab,SFBconst(cb'))::after, c', ctx'
else
@@ -145,8 +177,7 @@ let rec check_with_mod env struc (idl,mp1) mp equiv =
begin
try
let mtb_old = module_type_of_module old in
- Subtyping.check_subtypes env' mtb_mp1 mtb_old
- +++ old.mod_constraints
+ Univ.ContextSet.add_constraints (Subtyping.check_subtypes env' mtb_mp1 mtb_old) old.mod_constraints
with Failure _ -> error_incorrect_with_constraint lab
end
| Algebraic (NoFunctor (MEident(mp'))) ->
@@ -194,7 +225,7 @@ let rec check_with_mod env struc (idl,mp1) mp equiv =
| Algebraic (NoFunctor (MEident mp0)) ->
let mpnew = rebuild_mp mp0 idl in
check_modpath_equiv env' mpnew mp;
- before@(lab,spec)::after, equiv, Univ.Constraint.empty
+ before@(lab,spec)::after, equiv, Univ.ContextSet.empty
| _ -> error_generative_module_expected lab
end
with
@@ -207,8 +238,8 @@ let check_with env mp (sign,alg,reso,cst) = function
|WithDef(idl,c) ->
let struc = destr_nofunctor sign in
let struc',c',cst' = check_with_def env struc (idl,c) mp reso in
- let alg' = mk_alg_with alg (WithDef (idl,(c',cst'))) in
- (NoFunctor struc'),alg',reso, cst+++(Univ.UContext.constraints cst')
+ let alg' = mk_alg_with alg (WithDef (idl,(c',Univ.ContextSet.to_context cst'))) in
+ (NoFunctor struc'),alg',reso, cst+++cst'
|WithMod(idl,mp1) as wd ->
let struc = destr_nofunctor sign in
let struc',reso',cst' = check_with_mod env struc (idl,mp1) mp reso in
@@ -238,7 +269,7 @@ let rec translate_mse env mpo inl = function
let mtb = lookup_modtype mp1 env in
mtb.mod_type, mtb.mod_delta
in
- sign,Some (MEident mp1),reso,Univ.Constraint.empty
+ sign,Some (MEident mp1),reso,Univ.ContextSet.empty
|MEapply (fe,mp1) ->
translate_apply env inl (translate_mse env mpo inl fe) mp1 (mk_alg_app mpo)
|MEwith(me, with_decl) ->
@@ -256,7 +287,7 @@ and translate_apply env inl (sign,alg,reso,cst1) mp1 mkalg =
let body = subst_signature subst fbody_b in
let alg' = mkalg alg mp1 in
let reso' = subst_codom_delta_resolver subst reso in
- body,alg',reso', cst1 +++ cst2
+ body,alg',reso', Univ.ContextSet.add_constraints cst2 cst1
let mk_alg_funct mpo mbid mtb alg = match mpo, alg with
| Some _, Some alg -> Some (MoreFunctor (mbid,mtb,alg))
@@ -301,7 +332,7 @@ let finalize_module env mp (sign,alg,reso,cst) restype = match restype with
mk_mod mp impl sign None cst reso
|Some (params_mte,inl) ->
let res_mtb = translate_modtype env mp inl params_mte in
- let auto_mtb = mk_modtype mp sign Univ.Constraint.empty reso in
+ let auto_mtb = mk_modtype mp sign Univ.ContextSet.empty reso in
let cst' = Subtyping.check_subtypes env auto_mtb res_mtb in
let impl = match alg with Some e -> Algebraic e | None -> Struct sign in
{ res_mtb with
@@ -309,7 +340,7 @@ let finalize_module env mp (sign,alg,reso,cst) restype = match restype with
mod_expr = impl;
(** cst from module body typing, cst' from subtyping,
and constraints from module type. *)
- mod_constraints = cst +++ cst' +++ res_mtb.mod_constraints }
+ mod_constraints = Univ.ContextSet.add_constraints cst' (cst +++ res_mtb.mod_constraints) }
let translate_module env mp inl = function
|MType (params,ty) ->
@@ -324,7 +355,7 @@ let rec translate_mse_incl env mp inl = function
|MEident mp1 ->
let mb = strengthen_and_subst_mb (lookup_module mp1 env) mp true in
let sign = clean_bounded_mod_expr mb.mod_type in
- sign,None,mb.mod_delta,Univ.Constraint.empty
+ sign,None,mb.mod_delta,Univ.ContextSet.empty
|MEapply (fe,arg) ->
let ftrans = translate_mse_incl env mp inl fe in
translate_apply env inl ftrans arg (fun _ _ -> None)
diff --git a/kernel/mod_typing.mli b/kernel/mod_typing.mli
index b39e82125..80db12b0d 100644
--- a/kernel/mod_typing.mli
+++ b/kernel/mod_typing.mli
@@ -30,7 +30,7 @@ val translate_modtype :
*)
type 'alg translation =
- module_signature * 'alg option * delta_resolver * Univ.constraints
+ module_signature * 'alg option * delta_resolver * Univ.ContextSet.t
val translate_mse :
env -> module_path option -> inline -> module_struct_entry ->
diff --git a/kernel/modops.ml b/kernel/modops.ml
index d52fe611c..8733ca8c2 100644
--- a/kernel/modops.ml
+++ b/kernel/modops.ml
@@ -331,8 +331,10 @@ let strengthen_const mp_from l cb resolver =
let kn = KerName.make2 mp_from l in
let con = constant_of_delta_kn resolver kn in
let u =
- if cb.const_polymorphic then
- Univ.UContext.instance cb.const_universes
+ if cb.const_polymorphic then
+ let u = Univ.UContext.instance cb.const_universes in
+ let s = Univ.make_instance_subst u in
+ Univ.subst_univs_level_instance s u
else Univ.Instance.empty
in
{ cb with
diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml
index 55e767321..9329b1686 100644
--- a/kernel/safe_typing.ml
+++ b/kernel/safe_typing.ml
@@ -118,8 +118,8 @@ type safe_environment =
revstruct : structure_body;
modlabels : Label.Set.t;
objlabels : Label.Set.t;
- univ : Univ.constraints;
- future_cst : Univ.constraints Future.computation list;
+ univ : Univ.ContextSet.t;
+ future_cst : Univ.ContextSet.t Future.computation list;
engagement : engagement option;
required : vodigest DPMap.t;
loads : (module_path * module_body) list;
@@ -148,7 +148,7 @@ let empty_environment =
modlabels = Label.Set.empty;
objlabels = Label.Set.empty;
future_cst = [];
- univ = Univ.Constraint.empty;
+ univ = Univ.ContextSet.empty;
engagement = None;
required = DPMap.empty;
loads = [];
@@ -221,22 +221,23 @@ let env_of_safe_env senv = senv.env
let env_of_senv = env_of_safe_env
type constraints_addition =
- Now of Univ.constraints | Later of Univ.constraints Future.computation
+ | Now of bool * Univ.ContextSet.t
+ | Later of Univ.ContextSet.t Future.computation
let add_constraints cst senv =
match cst with
| Later fc ->
{senv with future_cst = fc :: senv.future_cst}
- | Now cst ->
+ | Now (poly,cst) ->
{ senv with
- env = Environ.add_constraints cst senv.env;
- univ = Univ.Constraint.union cst senv.univ }
+ env = Environ.push_context_set ~strict:(not poly) cst senv.env;
+ univ = Univ.ContextSet.union cst senv.univ }
let add_constraints_list cst senv =
- List.fold_right add_constraints cst senv
+ List.fold_left (fun acc c -> add_constraints c acc) senv cst
-let push_context_set ctx = add_constraints (Now (Univ.ContextSet.constraints ctx))
-let push_context ctx = add_constraints (Now (Univ.UContext.constraints ctx))
+let push_context_set poly ctx = add_constraints (Now (poly,ctx))
+let push_context poly ctx = add_constraints (Now (poly,Univ.ContextSet.of_context ctx))
let is_curmod_library senv =
match senv.modvariant with LIBRARY -> true | _ -> false
@@ -246,7 +247,7 @@ let join_safe_environment ?(except=Future.UUIDSet.empty) e =
List.fold_left
(fun e fc ->
if Future.UUIDSet.mem (Future.uuid fc) except then e
- else add_constraints (Now (Future.join fc)) e)
+ else add_constraints (Now (false, Future.join fc)) e)
{e with future_cst = []} e.future_cst
let is_joined_environment e = List.is_empty e.future_cst
@@ -337,20 +338,21 @@ let safe_push_named (id,_,_ as d) env =
let push_named_def (id,de) senv =
let c,typ,univs = Term_typing.translate_local_def senv.env id de in
- let senv' = push_context univs senv in
- let c, senv' = match c with
- | Def c -> Mod_subst.force_constr c, senv'
+ let poly = de.Entries.const_entry_polymorphic in
+ let univs = Univ.ContextSet.of_context univs in
+ let c, univs = match c with
+ | Def c -> Mod_subst.force_constr c, univs
| OpaqueDef o ->
- Opaqueproof.force_proof (Environ.opaque_tables senv'.env) o,
- push_context_set
- (Opaqueproof.force_constraints (Environ.opaque_tables senv'.env) o)
- senv'
+ Opaqueproof.force_proof (Environ.opaque_tables senv.env) o,
+ Univ.ContextSet.union univs
+ (Opaqueproof.force_constraints (Environ.opaque_tables senv.env) o)
| _ -> assert false in
+ let senv' = push_context_set poly univs senv in
let env'' = safe_push_named (id,Some c,typ) senv'.env in
- {senv' with env=env''}
+ univs, {senv' with env=env''}
-let push_named_assum ((id,t),ctx) senv =
- let senv' = push_context_set ctx senv in
+let push_named_assum ((id,t,poly),ctx) senv =
+ let senv' = push_context_set poly ctx senv in
let t = Term_typing.translate_local_assum senv'.env t in
let env'' = safe_push_named (id,None,t) senv'.env in
{senv' with env=env''}
@@ -373,10 +375,10 @@ let labels_of_mib mib =
let globalize_constant_universes env cb =
if cb.const_polymorphic then
- [Now Univ.Constraint.empty]
+ [Now (true, Univ.ContextSet.empty)]
else
- let cstrs = Univ.UContext.constraints cb.const_universes in
- Now cstrs ::
+ let cstrs = Univ.ContextSet.of_context cb.const_universes in
+ Now (false, cstrs) ::
(match cb.const_body with
| (Undef _ | Def _) -> []
| OpaqueDef lc ->
@@ -384,23 +386,21 @@ let globalize_constant_universes env cb =
| None -> []
| Some fc ->
match Future.peek_val fc with
- | None -> [Later (Future.chain
- ~greedy:(not (Future.is_exn fc))
- ~pure:true fc Univ.ContextSet.constraints)]
- | Some c -> [Now (Univ.ContextSet.constraints c)])
+ | None -> [Later fc]
+ | Some c -> [Now (false, c)])
let globalize_mind_universes mb =
if mb.mind_polymorphic then
- [Now Univ.Constraint.empty]
+ [Now (true, Univ.ContextSet.empty)]
else
- [Now (Univ.UContext.constraints mb.mind_universes)]
+ [Now (false, Univ.ContextSet.of_context mb.mind_universes)]
let constraints_of_sfb env sfb =
match sfb with
| SFBconst cb -> globalize_constant_universes env cb
| SFBmind mib -> globalize_mind_universes mib
- | SFBmodtype mtb -> [Now mtb.mod_constraints]
- | SFBmodule mb -> [Now mb.mod_constraints]
+ | SFBmodtype mtb -> [Now (false, mtb.mod_constraints)]
+ | SFBmodule mb -> [Now (false, mb.mod_constraints)]
(** A generic function for adding a new field in a same environment.
It also performs the corresponding [add_constraints]. *)
@@ -503,13 +503,13 @@ let add_modtype l params_mte inl senv =
(** full_add_module adds module with universes and constraints *)
let full_add_module mb senv =
- let senv = add_constraints (Now mb.mod_constraints) senv in
+ let senv = add_constraints (Now (false, mb.mod_constraints)) senv in
let dp = ModPath.dp mb.mod_mp in
let linkinfo = Nativecode.link_info_of_dirpath dp in
{ senv with env = Modops.add_linked_module mb linkinfo senv.env }
let full_add_module_type mp mt senv =
- let senv = add_constraints (Now mt.mod_constraints) senv in
+ let senv = add_constraints (Now (false, mt.mod_constraints)) senv in
{ senv with env = Modops.add_module_type mp mt senv.env }
(** Insertion of modules *)
@@ -617,8 +617,8 @@ let propagate_senv newdef newenv newresolver senv oldsenv =
modlabels = Label.Set.add (fst newdef) oldsenv.modlabels;
univ =
List.fold_left (fun acc cst ->
- Univ.Constraint.union acc (Future.force cst))
- (Univ.Constraint.union senv.univ oldsenv.univ)
+ Univ.ContextSet.union acc (Future.force cst))
+ (Univ.ContextSet.union senv.univ oldsenv.univ)
now_cst;
future_cst = later_cst @ oldsenv.future_cst;
(* engagement is propagated to the upper level *)
@@ -641,8 +641,8 @@ let end_module l restype senv =
let senv'=
propagate_loads { senv with
env = newenv;
- univ = Univ.Constraint.union senv.univ mb.mod_constraints} in
- let newenv = Environ.add_constraints mb.mod_constraints senv'.env in
+ univ = Univ.ContextSet.union senv.univ mb.mod_constraints} in
+ let newenv = Environ.push_context_set ~strict:true mb.mod_constraints senv'.env in
let newenv = Modops.add_module mb newenv in
let newresolver =
if Modops.is_functor mb.mod_type then oldsenv.modresolver
@@ -667,7 +667,7 @@ let end_modtype l senv =
let () = check_empty_context senv in
let mbids = List.rev_map fst params in
let newenv = Environ.set_opaque_tables oldsenv.env (Environ.opaque_tables senv.env) in
- let newenv = Environ.add_constraints senv.univ newenv in
+ let newenv = Environ.push_context_set ~strict:true senv.univ newenv in
let newenv = set_engagement_opt newenv senv.engagement in
let senv' = propagate_loads {senv with env=newenv} in
let auto_tb = functorize params (NoFunctor (List.rev senv.revstruct)) in
@@ -690,13 +690,16 @@ let add_include me is_module inl senv =
let mtb = translate_modtype senv.env mp_sup inl ([],me) in
mtb.mod_type,mtb.mod_constraints,mtb.mod_delta
in
- let senv = add_constraints (Now cst) senv in
+ let senv = add_constraints (Now (false, cst)) senv in
(* Include Self support *)
let rec compute_sign sign mb resolver senv =
match sign with
| MoreFunctor(mbid,mtb,str) ->
let cst_sub = Subtyping.check_subtypes senv.env mb mtb in
- let senv = add_constraints (Now cst_sub) senv in
+ let senv =
+ add_constraints
+ (Now (false, Univ.ContextSet.add_constraints cst_sub Univ.ContextSet.empty))
+ senv in
let mpsup_delta =
Modops.inline_delta_resolver senv.env inl mp_sup mbid mtb mb.mod_delta
in
@@ -707,7 +710,7 @@ let add_include me is_module inl senv =
in
let resolver,sign,senv =
let struc = NoFunctor (List.rev senv.revstruct) in
- let mtb = build_mtb mp_sup struc Univ.Constraint.empty senv.modresolver in
+ let mtb = build_mtb mp_sup struc Univ.ContextSet.empty senv.modresolver in
compute_sign sign mtb resolver senv
in
let str = match sign with
@@ -801,8 +804,10 @@ let import lib cst vodigest senv =
check_engagement senv.env lib.comp_enga;
let mp = MPfile lib.comp_name in
let mb = lib.comp_mod in
- let env = Environ.add_constraints mb.mod_constraints senv.env in
- let env = Environ.push_context_set cst env in
+ let env = Environ.push_context_set ~strict:true
+ (Univ.ContextSet.union mb.mod_constraints cst)
+ senv.env
+ in
mp,
{ senv with
env =
@@ -855,7 +860,9 @@ let register_inline kn senv =
let env = { env with env_globals = new_globals } in
{ senv with env = env_of_pre_env env }
-let add_constraints c = add_constraints (Now c)
+let add_constraints c =
+ add_constraints
+ (Now (false, Univ.ContextSet.add_constraints c Univ.ContextSet.empty))
(* NB: The next old comment probably refers to [propagate_loads] above.
diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli
index 2b4324b96..eac08eb83 100644
--- a/kernel/safe_typing.mli
+++ b/kernel/safe_typing.mli
@@ -57,9 +57,13 @@ val is_joined_environment : safe_environment -> bool
(** Insertion of local declarations (Local or Variables) *)
val push_named_assum :
- (Id.t * Term.types) Univ.in_universe_context_set -> safe_transformer0
+ (Id.t * Term.types * bool (* polymorphic *))
+ Univ.in_universe_context_set -> safe_transformer0
+
+(** Returns the full universe context necessary to typecheck the definition
+ (futures are forced) *)
val push_named_def :
- Id.t * Entries.definition_entry -> safe_transformer0
+ Id.t * Entries.definition_entry -> Univ.universe_context_set safe_transformer
(** Insertion of global axioms or definitions *)
@@ -88,10 +92,10 @@ val add_modtype :
(** Adding universe constraints *)
val push_context_set :
- Univ.universe_context_set -> safe_transformer0
+ bool -> Univ.universe_context_set -> safe_transformer0
val push_context :
- Univ.universe_context -> safe_transformer0
+ bool -> Univ.universe_context -> safe_transformer0
val add_constraints :
Univ.constraints -> safe_transformer0
diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml
index db155e6c8..58f3bcdf0 100644
--- a/kernel/subtyping.ml
+++ b/kernel/subtyping.ml
@@ -311,15 +311,19 @@ let check_constant cst env mp1 l info1 cb2 spec2 subst1 subst2 =
try
(* The environment with the expected universes plus equality
of the body instances with the expected instance *)
- let env = Environ.add_constraints cstrs env in
- (* Check that the given definition does not add any constraint over
- the expected ones, so that it can be used in place of the original. *)
+ let ctxi = Univ.Instance.append inst1 inst2 in
+ let ctx = Univ.UContext.make (ctxi, cstrs) in
+ let env = Environ.push_context ctx env in
+ (* Check that the given definition does not add any constraint over
+ the expected ones, so that it can be used in place of
+ the original. *)
if Univ.check_constraints ctx1 (Environ.universes env) then
cstrs, env, inst2
else error (IncompatibleConstraints ctx1)
with Univ.UniverseInconsistency incon ->
error (IncompatibleUniverses incon)
- else cst, env, Univ.Instance.empty
+ else
+ cst, env, Univ.Instance.empty
in
(* Now check types *)
let typ1 = Typeops.type_of_constant_type env' cb1.const_type in
@@ -456,6 +460,7 @@ and check_modtypes cst env mtb1 mtb2 subst1 subst2 equiv =
let check_subtypes env sup super =
let env = add_module_type sup.mod_mp sup env in
+ let env = Environ.push_context_set ~strict:true super.mod_constraints env in
check_modtypes Univ.Constraint.empty env
(strengthen sup sup.mod_mp) super empty_subst
(map_mp super.mod_mp sup.mod_mp sup.mod_delta) false
diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml
index 83e566041..926b38794 100644
--- a/kernel/term_typing.ml
+++ b/kernel/term_typing.ml
@@ -45,8 +45,8 @@ let map_option_typ = function None -> `None | Some x -> `Some x
let mk_pure_proof c = (c, Univ.ContextSet.empty), Declareops.no_seff
-let handle_side_effects env body side_eff =
- let handle_sideff t se =
+let handle_side_effects env body ctx side_eff =
+ let handle_sideff (t,ctx) se =
let cbl = match se with
| SEsubproof (c,cb,b) -> [c,cb,b]
| SEscheme (cl,_) -> List.map (fun (_,c,cb,b) -> c,cb,b) cl in
@@ -66,7 +66,7 @@ let handle_side_effects env body side_eff =
| Const (c',u') when eq_constant c c' ->
Vars.subst_instance_constr u' b
| _ -> map_constr_with_binders ((+) 1) (fun i x -> sub_body c u b i x) i x in
- let fix_body (c,cb,b) t =
+ let fix_body (c,cb,b) (t,ctx) =
match cb.const_body, b with
| Def b, _ ->
let b = Mod_subst.force_constr b in
@@ -74,25 +74,29 @@ let handle_side_effects env body side_eff =
if not poly then
let b_ty = Typeops.type_of_constant_type env cb.const_type in
let t = sub c 1 (Vars.lift 1 t) in
- mkLetIn (cname c, b, b_ty, t)
+ mkLetIn (cname c, b, b_ty, t),
+ Univ.ContextSet.union ctx
+ (Univ.ContextSet.of_context cb.const_universes)
else
let univs = cb.const_universes in
- sub_body c (Univ.UContext.instance univs) b 1 (Vars.lift 1 t)
+ sub_body c (Univ.UContext.instance univs) b 1 (Vars.lift 1 t), ctx
| OpaqueDef _, `Opaque (b,_) ->
let poly = cb.const_polymorphic in
if not poly then
let b_ty = Typeops.type_of_constant_type env cb.const_type in
let t = sub c 1 (Vars.lift 1 t) in
- mkApp (mkLambda (cname c, b_ty, t), [|b|])
+ mkApp (mkLambda (cname c, b_ty, t), [|b|]),
+ Univ.ContextSet.union ctx
+ (Univ.ContextSet.of_context cb.const_universes)
else
let univs = cb.const_universes in
- sub_body c (Univ.UContext.instance univs) b 1 (Vars.lift 1 t)
+ sub_body c (Univ.UContext.instance univs) b 1 (Vars.lift 1 t), ctx
| _ -> assert false
in
- List.fold_right fix_body cbl t
+ List.fold_right fix_body cbl (t,ctx)
in
(* CAVEAT: we assure a proper order *)
- Declareops.fold_side_effects handle_sideff body
+ Declareops.fold_side_effects handle_sideff (body,ctx)
(Declareops.uniquize_side_effects side_eff)
let hcons_j j =
@@ -100,11 +104,11 @@ let hcons_j j =
let feedback_completion_typecheck =
Option.iter (fun state_id -> Pp.feedback ~state_id Feedback.Complete)
-
+
let infer_declaration env kn dcl =
match dcl with
| ParameterEntry (ctx,poly,(t,uctx),nl) ->
- let env = push_context uctx env in
+ let env = push_context ~strict:(not poly) uctx env in
let j = infer env t in
let abstract = poly && not (Option.is_empty kn) in
let usubst, univs = Univ.abstract_universes abstract uctx in
@@ -115,12 +119,12 @@ let infer_declaration env kn dcl =
| DefinitionEntry ({ const_entry_type = Some typ;
const_entry_opaque = true;
const_entry_polymorphic = false} as c) ->
- let env = push_context c.const_entry_universes env in
+ let env = push_context ~strict:true c.const_entry_universes env in
let { const_entry_body = body; const_entry_feedback = feedback_id } = c in
let tyj = infer_type env typ in
let proofterm =
Future.chain ~greedy:true ~pure:true body (fun ((body, ctx),side_eff) ->
- let body = handle_side_effects env body side_eff in
+ let body,ctx = handle_side_effects env body ctx side_eff in
let env' = push_context_set ctx env in
let j = infer env' body in
let j = hcons_j j in
@@ -135,14 +139,16 @@ let infer_declaration env kn dcl =
c.const_entry_inline_code, c.const_entry_secctx
| DefinitionEntry c ->
- let env = push_context c.const_entry_universes env in
let { const_entry_type = typ; const_entry_opaque = opaque } = c in
let { const_entry_body = body; const_entry_feedback = feedback_id } = c in
let (body, ctx), side_eff = Future.join body in
- assert(Univ.ContextSet.is_empty ctx);
- let body = handle_side_effects env body side_eff in
+ let univsctx = Univ.ContextSet.of_context c.const_entry_universes in
+ let body, ctx = handle_side_effects env body
+ (Univ.ContextSet.union univsctx ctx) side_eff in
+ let env = push_context_set ~strict:(not c.const_entry_polymorphic) ctx env in
let abstract = c.const_entry_polymorphic && not (Option.is_empty kn) in
- let usubst, univs = Univ.abstract_universes abstract c.const_entry_universes in
+ let usubst, univs =
+ Univ.abstract_universes abstract (Univ.ContextSet.to_context ctx) in
let j = infer env body in
let typ = constrain_type env j c.const_entry_polymorphic usubst (map_option_typ typ) in
let def = hcons_constr (Vars.subst_univs_level_constr usubst j.uj_val) in
@@ -306,5 +312,9 @@ let translate_mind env kn mie = Indtypes.check_inductive env kn mie
let handle_entry_side_effects env ce = { ce with
const_entry_body = Future.chain ~greedy:true ~pure:true
ce.const_entry_body (fun ((body, ctx), side_eff) ->
- (handle_side_effects env body side_eff, ctx), Declareops.no_seff);
+ let body, ctx' = handle_side_effects env body ctx side_eff in
+ (body, ctx'), Declareops.no_seff);
}
+
+let handle_side_effects env body side_eff =
+ fst (handle_side_effects env body Univ.ContextSet.empty side_eff)
diff --git a/kernel/univ.ml b/kernel/univ.ml
index 336cdb653..73d323426 100644
--- a/kernel/univ.ml
+++ b/kernel/univ.ml
@@ -270,8 +270,10 @@ module Level = struct
let is_small x =
match data x with
| Level _ -> false
- | _ -> true
-
+ | Var _ -> false
+ | Prop -> true
+ | Set -> true
+
let is_prop x =
match data x with
| Prop -> true
@@ -551,6 +553,10 @@ struct
| Cons (l, _, Nil) -> Expr.is_level l
| _ -> false
+ let rec is_levels l = match l with
+ | Cons (l, _, r) -> Expr.is_level l && is_levels r
+ | Nil -> true
+
let level l = match l with
| Cons (l, _, Nil) -> Expr.level l
| _ -> None
@@ -655,7 +661,6 @@ type canonical_arc =
lt: Level.t list;
le: Level.t list;
rank : int;
- predicative : bool;
mutable status : status;
(** Guaranteed to be unset out of the [compare_neq] functions. It is used
to do an imperative traversal of the graph, ensuring a O(1) check that
@@ -670,7 +675,7 @@ let arc_is_lt arc = match arc.status with
| Unset | SetLe -> false
| SetLt -> true
-let terminal u = {univ=u; lt=[]; le=[]; rank=0; predicative=false; status = Unset}
+let terminal u = {univ=u; lt=[]; le=[]; rank=0; status = Unset}
module UMap :
sig
@@ -720,6 +725,16 @@ let enter_arc ca g =
(* Every Level.t has a unique canonical arc representative *)
+(** The graph always contains nodes for Prop and Set. *)
+
+let terminal_lt u v =
+ {(terminal u) with lt=[v]}
+
+let empty_universes =
+ let g = enter_arc (terminal Level.set) UMap.empty in
+ let g = enter_arc (terminal_lt Level.prop Level.set) g in
+ g
+
(* repr : universes -> Level.t -> canonical_arc *)
(* canonical representative : we follow the Equiv links *)
@@ -733,19 +748,28 @@ let rec repr g u =
| Equiv v -> repr g v
| Canonical arc -> arc
-(* [safe_repr] also search for the canonical representative, but
- if the graph doesn't contain the searched universe, we add it. *)
+let get_prop_arc g = repr g Level.prop
+let get_set_arc g = repr g Level.set
+let is_set_arc u = Level.is_set u.univ
+let is_prop_arc u = Level.is_prop u.univ
-let safe_repr g u =
- let rec safe_repr_rec g u =
- match UMap.find u g with
- | Equiv v -> safe_repr_rec g v
- | Canonical arc -> arc
- in
- try g, safe_repr_rec g u
- with Not_found ->
- let can = terminal u in
- enter_arc can g, can
+exception AlreadyDeclared
+
+let add_universe vlev strict g =
+ try
+ let _arcv = UMap.find vlev g in
+ raise AlreadyDeclared
+ with Not_found ->
+ let v = terminal vlev in
+ let arc =
+ let arc = get_set_arc g in
+ if strict then
+ { arc with lt=vlev::arc.lt}
+ else
+ { arc with le=vlev::arc.le}
+ in
+ let g = enter_arc arc g in
+ enter_arc v g
(* reprleq : canonical_arc -> canonical_arc list *)
(* All canonical arcv such that arcu<=arcv with arcv#arcu *)
@@ -789,7 +813,6 @@ let between g arcu arcv =
in
let good,_,_ = explore ([arcv],[],false) arcu in
good
-
(* We assume compare(u,v) = LE with v canonical (see compare below).
In this case List.hd(between g u v) = repr u
Otherwise, between g u v = []
@@ -900,8 +923,9 @@ let get_explanation strict g arcu arcv =
in
find [] arc.lt
in
+ let start = (* if is_prop_arc arcu then [Le, make arcv.univ] else *) [] in
try
- let (to_revert, c) = cmp [] [] [] [(arcu, [])] in
+ let (to_revert, c) = cmp start [] [] [(arcu, [])] in
(** Reset all the touched arcs. *)
let () = List.iter (fun arc -> arc.status <- Unset) to_revert in
List.rev c
@@ -972,7 +996,6 @@ let fast_compare_neq strict g arcu arcv =
else process_le c to_revert (arc :: lt_todo) le_todo lt le
in
-
try
let (to_revert, c) = cmp FastNLE [] [] [arcu] in
(** Reset all the touched arcs. *)
@@ -1015,24 +1038,18 @@ let is_lt g arcu arcv =
(** First, checks on universe levels *)
let check_equal g u v =
- let g, arcu = safe_repr g u in
- let _, arcv = safe_repr g v in
- arcu == arcv
+ let arcu = repr g u and arcv = repr g v in
+ arcu == arcv
let check_eq_level g u v = u == v || check_equal g u v
-let is_set_arc u = Level.is_set u.univ
-let is_prop_arc u = Level.is_prop u.univ
-let get_prop_arc g = snd (safe_repr g Level.prop)
-
let check_smaller g strict u v =
- let g, arcu = safe_repr g u in
- let g, arcv = safe_repr g v in
+ let arcu = repr g u and arcv = repr g v in
if strict then
is_lt g arcu arcv
else
is_prop_arc arcu
- || (is_set_arc arcu && arcv.predicative)
+ || (is_set_arc arcu && not (is_prop_arc arcv))
|| is_leq g arcu arcv
(** Then, checks on universes *)
@@ -1074,19 +1091,11 @@ let check_leq g u v =
(** Enforcing new constraints : [setlt], [setleq], [merge], [merge_disc] *)
-(** To speed up tests of Set </<= i *)
-let set_predicative g arcv =
- enter_arc {arcv with predicative = true} g
-
(* setlt : Level.t -> Level.t -> reason -> unit *)
(* forces u > v *)
(* this is normally an update of u in g rather than a creation. *)
let setlt g arcu arcv =
let arcu' = {arcu with lt=arcv.univ::arcu.lt} in
- let g =
- if is_set_arc arcu then set_predicative g arcv
- else g
- in
enter_arc arcu' g, arcu'
(* checks that non-redundant *)
@@ -1100,11 +1109,6 @@ let setlt_if (g,arcu) v =
(* this is normally an update of u in g rather than a creation. *)
let setleq g arcu arcv =
let arcu' = {arcu with le=arcv.univ::arcu.le} in
- let g =
- if is_set_arc arcu' then
- set_predicative g arcv
- else g
- in
enter_arc arcu' g, arcu'
(* checks that non-redundant *)
@@ -1120,7 +1124,8 @@ let merge g arcu arcv =
(* we find the arc with the biggest rank, and we redirect all others to it *)
let arcu, g, v =
let best_ranked (max_rank, old_max_rank, best_arc, rest) arc =
- if Level.is_small arc.univ || arc.rank >= max_rank
+ if Level.is_small arc.univ ||
+ (arc.rank >= max_rank && not (Level.is_small best_arc.univ))
then (arc.rank, max_rank, arc, best_arc::rest)
else (max_rank, old_max_rank, best_arc, arc::rest)
in
@@ -1150,7 +1155,7 @@ let merge g arcu arcv =
(* we assume compare(u,v) = compare(v,u) = NLE *)
(* merge_disc u v forces u ~ v with repr u as canonical repr *)
let merge_disc g arc1 arc2 =
- let arcu, arcv = if arc1.rank < arc2.rank then arc2, arc1 else arc1, arc2 in
+ let arcu, arcv = if Level.is_small arc2.univ || arc1.rank < arc2.rank then arc2, arc1 else arc1, arc2 in
let arcu, g =
if not (Int.equal arc1.rank arc2.rank) then arcu, g
else
@@ -1173,12 +1178,11 @@ exception UniverseInconsistency of univ_inconsistency
let error_inconsistency o u v (p:explanation option) =
raise (UniverseInconsistency (o,make u,make v,p))
-(* enforc_univ_eq : Level.t -> Level.t -> unit *)
-(* enforc_univ_eq u v will force u=v if possible, will fail otherwise *)
+(* enforce_univ_eq : Level.t -> Level.t -> unit *)
+(* enforce_univ_eq u v will force u=v if possible, will fail otherwise *)
let enforce_univ_eq u v g =
- let g,arcu = safe_repr g u in
- let g,arcv = safe_repr g v in
+ let arcu = repr g u and arcv = repr g v in
match fast_compare g arcu arcv with
| FastEQ -> g
| FastLT ->
@@ -1197,8 +1201,7 @@ let enforce_univ_eq u v g =
(* enforce_univ_leq : Level.t -> Level.t -> unit *)
(* enforce_univ_leq u v will force u<=v if possible, will fail otherwise *)
let enforce_univ_leq u v g =
- let g,arcu = safe_repr g u in
- let g,arcv = safe_repr g v in
+ let arcu = repr g u and arcv = repr g v in
if is_leq g arcu arcv then g
else
match fast_compare g arcv arcu with
@@ -1211,8 +1214,7 @@ let enforce_univ_leq u v g =
(* enforce_univ_lt u v will force u<v if possible, will fail otherwise *)
let enforce_univ_lt u v g =
- let g,arcu = safe_repr g u in
- let g,arcv = safe_repr g v in
+ let arcu = repr g u and arcv = repr g v in
match fast_compare g arcu arcv with
| FastLT -> g
| FastLE -> fst (setlt g arcu arcv)
@@ -1225,18 +1227,10 @@ let enforce_univ_lt u v g =
let p = get_explanation false g arcv arcu in
error_inconsistency Lt u v p
-let empty_universes = UMap.empty
-
(* Prop = Set is forbidden here. *)
-let initial_universes = enforce_univ_lt Level.prop Level.set UMap.empty
+let initial_universes = empty_universes
let is_initial_universes g = UMap.equal (==) g initial_universes
-
-let add_universe vlev g =
- let v = terminal vlev in
- let proparc = get_prop_arc g in
- enter_arc {proparc with le=vlev::proparc.le}
- (enter_arc v g)
(* Constraints and sets of constraints. *)
@@ -1370,10 +1364,12 @@ let check_univ_leq u v =
let enforce_leq u v c =
let open Universe.Huniv in
+ let rec aux acc v =
match v with
- | Cons (v, _, Nil) ->
- fold (fun u -> constraint_add_leq u v) u c
- | _ -> anomaly (Pp.str"A universe bound can only be a variable")
+ | Cons (v, _, l) ->
+ aux (fold (fun u -> constraint_add_leq u v) u c) l
+ | Nil -> acc
+ in aux c v
let enforce_leq u v c =
if check_univ_leq u v then c
@@ -1444,7 +1440,6 @@ let normalize_universes g =
lt = LSet.elements lt;
le = LSet.elements le;
rank = rank;
- predicative = false;
status = Unset;
}
in
@@ -1589,7 +1584,7 @@ let sort_universes orig =
let fold i accu u =
if 0 < i then
let pred = types.(i - 1) in
- let arc = {univ = u; lt = [pred]; le = []; rank = 0; predicative = false; status = Unset; } in
+ let arc = {univ = u; lt = [pred]; le = []; rank = 0; status = Unset; } in
UMap.add u (Canonical arc) accu
else accu
in
diff --git a/kernel/univ.mli b/kernel/univ.mli
index 7aaf2ffe6..4cc8a2528 100644
--- a/kernel/univ.mli
+++ b/kernel/univ.mli
@@ -20,7 +20,11 @@ sig
val is_small : t -> bool
(** Is the universe set or prop? *)
-
+
+ val is_prop : t -> bool
+ val is_set : t -> bool
+ (** Is it specifically Prop or Set *)
+
val compare : t -> t -> int
(** Comparison function *)
@@ -87,6 +91,9 @@ sig
val is_level : t -> bool
(** Test if the universe is a level or an algebraic universe. *)
+ val is_levels : t -> bool
+ (** Test if the universe is a lub of levels or contains +n's. *)
+
val level : t -> Level.t option
(** Try to get a level out of a universe, returns [None] if it
is an algebraic universe. *)
@@ -159,8 +166,12 @@ val is_initial_universes : universes -> bool
val sort_universes : universes -> universes
-(** Adds a universe to the graph, ensuring it is >= Prop. *)
-val add_universe : universe_level -> universes -> universes
+(** Adds a universe to the graph, ensuring it is >= or > Set.
+ @raises AlreadyDeclared if the level is already declared in the graph. *)
+
+exception AlreadyDeclared
+
+val add_universe : universe_level -> bool -> universes -> universes
(** {6 Constraints. } *)
diff --git a/library/declare.ml b/library/declare.ml
index 8438380c9..16803b3bf 100644
--- a/library/declare.ml
+++ b/library/declare.ml
@@ -42,7 +42,7 @@ type variable_declaration = DirPath.t * section_variable_entry * logical_kind
let cache_variable ((sp,_),o) =
match o with
- | Inl ctx -> Global.push_context_set ctx
+ | Inl ctx -> Global.push_context_set false ctx
| Inr (id,(p,d,mk)) ->
(* Constr raisonne sur les noms courts *)
if variable_exists id then
@@ -50,13 +50,13 @@ let cache_variable ((sp,_),o) =
let impl,opaq,poly,ctx = match d with (* Fails if not well-typed *)
| SectionLocalAssum ((ty,ctx),poly,impl) ->
- let () = Global.push_named_assum ((id,ty),ctx) in
+ let () = Global.push_named_assum ((id,ty,poly),ctx) in
let impl = if impl then Implicit else Explicit in
impl, true, poly, ctx
| SectionLocalDef (de) ->
- let () = Global.push_named_def (id,de) in
- Explicit, de.const_entry_opaque, de.const_entry_polymorphic,
- (Univ.ContextSet.of_context de.const_entry_universes) in
+ let univs = Global.push_named_def (id,de) in
+ Explicit, de.const_entry_opaque,
+ de.const_entry_polymorphic, univs in
Nametab.push (Nametab.Until 1) (restrict_path 0 sp) (VarRef id);
add_section_variable id impl poly ctx;
Dischargedhypsmap.set_discharged_hyps sp [];
@@ -116,8 +116,9 @@ let open_constant i ((sp,kn), obj) =
match (Global.lookup_constant con).const_body with
| (Def _ | Undef _) -> ()
| OpaqueDef lc ->
- match Opaqueproof.get_constraints (Global.opaque_tables ())lc with
- | Some f when Future.is_val f -> Global.push_context_set (Future.force f)
+ match Opaqueproof.get_constraints (Global.opaque_tables ()) lc with
+ | Some f when Future.is_val f ->
+ Global.push_context_set false (Future.force f)
| _ -> ()
let exists_name id =
@@ -455,12 +456,14 @@ let input_universes : universe_names -> Libobject.obj =
let do_universe l =
let glob = Universes.global_universe_names () in
- let glob' =
- List.fold_left (fun (idl,lid) (l, id) ->
+ let glob', ctx =
+ List.fold_left (fun ((idl,lid),ctx) (l, id) ->
let lev = Universes.new_univ_level (Global.current_dirpath ()) in
- (Idmap.add id lev idl, Univ.LMap.add lev id lid))
- glob l
+ ((Idmap.add id lev idl, Univ.LMap.add lev id lid),
+ Univ.ContextSet.add_universe lev ctx))
+ (glob, Univ.ContextSet.empty) l
in
+ Global.push_context_set false ctx;
Lib.add_anonymous_leaf (input_universes glob')
diff --git a/library/global.ml b/library/global.ml
index 0419799b6..6002382c1 100644
--- a/library/global.ml
+++ b/library/global.ml
@@ -78,10 +78,10 @@ let globalize_with_summary fs f =
let i2l = Label.of_id
let push_named_assum a = globalize0 (Safe_typing.push_named_assum a)
-let push_named_def d = globalize0 (Safe_typing.push_named_def d)
+let push_named_def d = globalize (Safe_typing.push_named_def d)
let add_constraints c = globalize0 (Safe_typing.add_constraints c)
-let push_context_set c = globalize0 (Safe_typing.push_context_set c)
-let push_context c = globalize0 (Safe_typing.push_context c)
+let push_context_set b c = globalize0 (Safe_typing.push_context_set b c)
+let push_context b c = globalize0 (Safe_typing.push_context b c)
let set_engagement c = globalize0 (Safe_typing.set_engagement c)
let add_constant dir id d = globalize (Safe_typing.add_constant dir (i2l id) d)
@@ -249,7 +249,7 @@ let current_dirpath () =
let with_global f =
let (a, ctx) = f (env ()) (current_dirpath ()) in
- push_context_set ctx; a
+ push_context_set false ctx; a
(* spiwack: register/unregister functions for retroknowledge *)
let register field value by_clause =
diff --git a/library/global.mli b/library/global.mli
index 363bb5789..ac231f7fd 100644
--- a/library/global.mli
+++ b/library/global.mli
@@ -30,8 +30,8 @@ val set_engagement : Declarations.engagement -> unit
(** Variables, Local definitions, constants, inductive types *)
-val push_named_assum : (Id.t * Constr.types) Univ.in_universe_context_set -> unit
-val push_named_def : (Id.t * Entries.definition_entry) -> unit
+val push_named_assum : (Id.t * Constr.types * bool) Univ.in_universe_context_set -> unit
+val push_named_def : (Id.t * Entries.definition_entry) -> Univ.universe_context_set
val add_constant :
DirPath.t -> Id.t -> Safe_typing.global_declaration -> constant
@@ -41,8 +41,8 @@ val add_mind :
(** Extra universe constraints *)
val add_constraints : Univ.constraints -> unit
-val push_context : Univ.universe_context -> unit
-val push_context_set : Univ.universe_context_set -> unit
+val push_context : bool -> Univ.universe_context -> unit
+val push_context_set : bool -> Univ.universe_context_set -> unit
(** Non-interactive modules and module types *)
diff --git a/library/lib.ml b/library/lib.ml
index 81db547ef..f4f52db53 100644
--- a/library/lib.ml
+++ b/library/lib.ml
@@ -420,6 +420,7 @@ let extract_hyps (secs,ohyps) =
in aux (secs,ohyps)
let instance_from_variable_context sign =
+
let rec inst_rec = function
| (id,b,None,_) :: sign -> id :: inst_rec sign
| _ :: sign -> inst_rec sign
diff --git a/library/universes.ml b/library/universes.ml
index 1c8a5ad77..bc42cc044 100644
--- a/library/universes.ml
+++ b/library/universes.ml
@@ -182,10 +182,13 @@ let leq_constr_univs_infer univs m n =
if Sorts.equal s1 s2 then true
else
let u1 = Sorts.univ_of_sort s1 and u2 = Sorts.univ_of_sort s2 in
- if Univ.check_leq univs u1 u2 then true
- else
- (cstrs := Constraints.add (u1, ULe, u2) !cstrs;
- true)
+ if Univ.check_leq univs u1 u2 then
+ ((if Univ.is_small_univ u1 then
+ cstrs := Constraints.add (u1, ULe, u2) !cstrs);
+ true)
+ else
+ (cstrs := Constraints.add (u1, ULe, u2) !cstrs;
+ true)
in
let rec eq_constr' m n =
m == n || Constr.compare_head_gen eq_universes eq_sorts eq_constr' m n
@@ -820,22 +823,59 @@ let minimize_univ_variables ctx us algs left right cstrs =
if v == None then fst (aux acc u)
else LSet.remove u ctx', us, LSet.remove u algs, seen, cstrs)
us (ctx, us, algs, lbounds, cstrs)
-
+
let normalize_context_set ctx us algs =
let (ctx, csts) = ContextSet.levels ctx, ContextSet.constraints ctx in
let uf = UF.create () in
+ (** Keep the Prop/Set <= i constraints separate for minimization *)
+ let smallles, csts =
+ Constraint.fold (fun (l,d,r as cstr) (smallles, noneqs) ->
+ if d == Le then
+ if Univ.Level.is_small l then
+ (Constraint.add cstr smallles, noneqs)
+ else if Level.is_small r then
+ if Level.is_prop r then
+ raise (Univ.UniverseInconsistency
+ (Le,Universe.make l,Universe.make r,None))
+ else (smallles, Constraint.add (l,Eq,r) noneqs)
+ else (smallles, Constraint.add cstr noneqs)
+ else (smallles, Constraint.add cstr noneqs))
+ csts (Constraint.empty, Constraint.empty)
+ in
let csts =
(* We first put constraints in a normal-form: all self-loops are collapsed
to equalities. *)
- let g = Univ.merge_constraints csts Univ.empty_universes in
+ let g = Univ.LSet.fold (fun v g -> Univ.add_universe v false g)
+ ctx Univ.empty_universes
+ in
+ let g =
+ Univ.Constraint.fold
+ (fun (l, d, r) g ->
+ let g =
+ if not (Level.is_small l || LSet.mem l ctx) then
+ try Univ.add_universe l false g
+ with Univ.AlreadyDeclared -> g
+ else g
+ in
+ let g =
+ if not (Level.is_small r || LSet.mem r ctx) then
+ try Univ.add_universe r false g
+ with Univ.AlreadyDeclared -> g
+ else g
+ in g) csts g
+ in
+ let g = Univ.Constraint.fold Univ.enforce_constraint csts g in
Univ.constraints_of_universes g
in
let noneqs =
- Constraint.fold (fun (l,d,r) noneqs ->
- if d == Eq then (UF.union l r uf; noneqs)
- else Constraint.add (l,d,r) noneqs)
- csts Constraint.empty
+ Constraint.fold (fun (l,d,r as cstr) noneqs ->
+ if d == Eq then (UF.union l r uf; noneqs)
+ else (* We ignore the trivial Prop/Set <= i constraints. *)
+ if d == Le && Univ.Level.is_small l then noneqs
+ else Constraint.add cstr noneqs)
+ csts Constraint.empty
in
+ let noneqs = Constraint.union noneqs smallles in
let partition = UF.partition uf in
let flex x = LMap.mem x us in
let ctx, subst, eqs = List.fold_left (fun (ctx, subst, cstrs) s ->
@@ -941,12 +981,12 @@ let simplify_universe_context (univs,csts) =
let csts' = subst_univs_level_constraints subst csts' in
(univs', csts'), subst
-let is_small_leq (l,d,r) =
- Level.is_small l && d == Univ.Le
+let is_trivial_leq (l,d,r) =
+ Univ.Level.is_prop l && (d == Univ.Le || (d == Univ.Lt && Univ.Level.is_set r))
(* Prop < i <-> Set+1 <= i <-> Set < i *)
let translate_cstr (l,d,r as cstr) =
- if Level.equal Level.prop l && d == Univ.Lt then
+ if Level.equal Level.prop l && d == Univ.Lt && not (Level.equal Level.set r) then
(Level.set, d, r)
else cstr
@@ -954,7 +994,7 @@ let refresh_constraints univs (ctx, cstrs) =
let cstrs', univs' =
Univ.Constraint.fold (fun c (cstrs', univs as acc) ->
let c = translate_cstr c in
- if Univ.check_constraint univs c && not (is_small_leq c) then acc
+ if is_trivial_leq c then acc
else (Univ.Constraint.add c cstrs', Univ.enforce_constraint c univs))
cstrs (Univ.Constraint.empty, univs)
in ((ctx, cstrs'), univs')
diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml
index d5d6bdf74..97ea5fdc5 100644
--- a/plugins/cc/ccalgo.ml
+++ b/plugins/cc/ccalgo.ml
@@ -129,14 +129,14 @@ type cinfo=
ci_nhyps: int} (* # projectable args *)
let family_eq f1 f2 = match f1, f2 with
-| InProp, InProp
-| InSet, InSet
-| InType, InType -> true
-| _ -> false
+ | Prop Pos, Prop Pos
+ | Prop Null, Prop Null
+ | Type _, Type _ -> true
+ | _ -> false
type term=
Symb of constr
- | Product of sorts_family * sorts_family
+ | Product of sorts * sorts
| Eps of Id.t
| Appli of term*term
| Constructor of cinfo (* constructor arity + nhyps *)
@@ -161,7 +161,7 @@ let hash_sorts_family = function
let rec hash_term = function
| Symb c -> combine 1 (hash_constr c)
- | Product (s1, s2) -> combine3 2 (hash_sorts_family s1) (hash_sorts_family s2)
+ | Product (s1, s2) -> combine3 2 (Sorts.hash s1) (Sorts.hash s2)
| Eps i -> combine 3 (Id.hash i)
| Appli (t1, t2) -> combine3 4 (hash_term t1) (hash_term t2)
| Constructor {ci_constr=(c,u); ci_arity=i; ci_nhyps=j} -> combine4 5 (constructor_hash c) i j
@@ -425,8 +425,8 @@ let _B_ = Name (Id.of_string "A")
let _body_ = mkProd(Anonymous,mkRel 2,mkRel 2)
let cc_product s1 s2 =
- mkLambda(_A_,mkSort(Universes.new_sort_in_family s1),
- mkLambda(_B_,mkSort(Universes.new_sort_in_family s2),_body_))
+ mkLambda(_A_,mkSort(s1),
+ mkLambda(_B_,mkSort(s2),_body_))
let rec constr_of_term = function
Symb s-> applist_projection s []
diff --git a/plugins/cc/ccalgo.mli b/plugins/cc/ccalgo.mli
index c72843d55..0dcf3a870 100644
--- a/plugins/cc/ccalgo.mli
+++ b/plugins/cc/ccalgo.mli
@@ -30,7 +30,7 @@ type cinfo =
type term =
Symb of constr
- | Product of sorts_family * sorts_family
+ | Product of sorts * sorts
| Eps of Id.t
| Appli of term*term
| Constructor of cinfo (* constructor arity + nhyps *)
diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml
index 9c3a0f729..6439f58d2 100644
--- a/plugins/cc/cctac.ml
+++ b/plugins/cc/cctac.ml
@@ -46,7 +46,7 @@ let whd_delta env=
(* decompose member of equality in an applicative format *)
(** FIXME: evar leak *)
-let sf_of env sigma c = family_of_sort (sort_of env (ref sigma) c)
+let sf_of env sigma c = sort_of env (ref sigma) c
let rec decompose_term env sigma t=
match kind_of_term (whd env t) with
diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml
index 3edc590cc..64284c6fe 100644
--- a/plugins/funind/functional_principles_types.ml
+++ b/plugins/funind/functional_principles_types.ml
@@ -303,7 +303,8 @@ let generate_functional_principle (evd: Evd.evar_map ref)
try
let f = funs.(i) in
- let type_sort = Universes.new_sort_in_family InType in
+ let env = Global.env () in
+ let type_sort = Evarutil.evd_comb1 (Evd.fresh_sort_in_family env) evd InType in
let new_sorts =
match sorts with
| None -> Array.make (Array.length funs) (type_sort)
@@ -317,14 +318,14 @@ let generate_functional_principle (evd: Evd.evar_map ref)
id_of_f,Indrec.make_elimination_ident id_of_f (family_of_sort type_sort)
in
let names = ref [new_princ_name] in
- let evd' = !evd in
let hook =
fun new_principle_type _ _ ->
if Option.is_empty sorts
then
(* let id_of_f = Label.to_id (con_label f) in *)
let register_with_sort fam_sort =
- let s = Universes.new_sort_in_family fam_sort in
+ let evd' = Evd.from_env (Global.env ()) in
+ let evd',s = Evd.fresh_sort_in_family env evd' fam_sort in
let name = Indrec.make_elimination_ident base_new_princ_name fam_sort in
let evd',value = change_property_sort evd' s new_principle_type new_princ_name in
let evd' = fst (Typing.type_of ~refresh:true (Global.env ()) evd' value) in
@@ -394,7 +395,7 @@ let get_funs_constant mp dp =
let body = Tacred.cbv_norm_flags
(Closure.RedFlags.mkflags [Closure.RedFlags.fZETA])
(Global.env ())
- (Evd.empty)
+ (Evd.from_env (Global.env ()))
body
in
body
@@ -483,11 +484,10 @@ let make_scheme evd (fas : (pconstant*glob_sort) list) : Entries.definition_entr
let i = ref (-1) in
let sorts =
List.rev_map (fun (_,x) ->
- Universes.new_sort_in_family (Pretyping.interp_elimination_sort x)
+ Evarutil.evd_comb1 (Evd.fresh_sort_in_family env) evd (Pretyping.interp_elimination_sort x)
)
fas
in
- evd:=sigma;
(* We create the first priciple by tactic *)
let first_type,other_princ_types =
match l_schemes with
@@ -597,7 +597,7 @@ let make_scheme evd (fas : (pconstant*glob_sort) list) : Entries.definition_entr
let build_scheme fas =
Dumpglob.pause ();
- let evd = (ref Evd.empty) in
+ let evd = (ref (Evd.from_env (Global.env ()))) in
let pconstants = (List.map
(fun (_,f,sort) ->
let f_as_constant =
@@ -633,7 +633,7 @@ let build_scheme fas =
let build_case_scheme fa =
let env = Global.env ()
- and sigma = Evd.empty in
+ and sigma = (Evd.from_env (Global.env ())) in
(* let id_to_constr id = *)
(* Constrintern.global_reference id *)
(* in *)
@@ -673,14 +673,14 @@ let build_case_scheme fa =
);
*)
generate_functional_principle
- (ref Evd.empty)
+ (ref (Evd.from_env (Global.env ())))
false
scheme_type
(Some ([|sorts|]))
(Some princ_name)
this_block_funs
0
- (prove_princ_for_struct (ref Evd.empty) false 0 [|fst (destConst funs)|])
+ (prove_princ_for_struct (ref (Evd.from_env (Global.env ()))) false 0 [|fst (destConst funs)|])
in
()
diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4
index 61f03d6f2..bc7e6f8b0 100644
--- a/plugins/funind/g_indfun.ml4
+++ b/plugins/funind/g_indfun.ml4
@@ -388,7 +388,7 @@ let finduction (oid:Id.t option) (heuristic: fapp_info list -> fapp_info list)
| Some id ->
let idref = const_of_id id in
(* JF : FIXME : we probably need to keep trace of evd in presence of universe polymorphism *)
- let idconstr = snd (Evd.fresh_global (Global.env ()) Evd.empty idref) in
+ let idconstr = snd (Evd.fresh_global (Global.env ()) (Evd.from_env (Global.env ())) idref) in
(fun u -> constr_head_match u idconstr) (* select only id *)
| None -> (fun u -> isApp u) in (* select calls to any function *)
let info_list = find_fapp test g in
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index 07efaae27..1b12cd42c 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -333,20 +333,20 @@ let raw_push_named (na,raw_value,raw_typ) env =
match na with
| Anonymous -> env
| Name id ->
- let value = Option.map (fun x-> fst (Pretyping.understand env Evd.empty x)) raw_value in
- let typ,ctx = Pretyping.understand env Evd.empty ~expected_type:Pretyping.IsType raw_typ in
+ let value = Option.map (fun x-> fst (Pretyping.understand env (Evd.from_env env) x)) raw_value in
+ let typ,ctx = Pretyping.understand env (Evd.from_env env) ~expected_type:Pretyping.IsType raw_typ in
Environ.push_named (id,value,typ) env
let add_pat_variables pat typ env : Environ.env =
let rec add_pat_variables env pat typ : Environ.env =
- observe (str "new rel env := " ++ Printer.pr_rel_context_of env Evd.empty);
+ observe (str "new rel env := " ++ Printer.pr_rel_context_of env (Evd.from_env env));
match pat with
| PatVar(_,na) -> Environ.push_rel (na,None,typ) env
| PatCstr(_,c,patl,na) ->
let Inductiveops.IndType(indf,indargs) =
- try Inductiveops.find_rectype env Evd.empty typ
+ try Inductiveops.find_rectype env (Evd.from_env env) typ
with Not_found -> assert false
in
let constructors = Inductiveops.get_constructors env indf in
@@ -376,7 +376,7 @@ let add_pat_variables pat typ env : Environ.env =
~init:(env,[])
)
in
- observe (str "new var env := " ++ Printer.pr_named_context_of res Evd.empty);
+ observe (str "new var env := " ++ Printer.pr_named_context_of res (Evd.from_env env));
res
@@ -393,7 +393,7 @@ let rec pattern_to_term_and_type env typ = function
constr
in
let Inductiveops.IndType(indf,indargs) =
- try Inductiveops.find_rectype env Evd.empty typ
+ try Inductiveops.find_rectype env (Evd.from_env env) typ
with Not_found -> assert false
in
let constructors = Inductiveops.get_constructors env indf in
@@ -405,7 +405,7 @@ let rec pattern_to_term_and_type env typ = function
Array.to_list
(Array.init
(cst_narg - List.length patternl)
- (fun i -> Detyping.detype false [] env Evd.empty csta.(i))
+ (fun i -> Detyping.detype false [] env (Evd.from_env env) csta.(i))
)
in
let patl_as_term =
@@ -486,9 +486,9 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
a pseudo value "v1 ... vn".
The "value" of this branch is then simply [res]
*)
- let rt_as_constr,ctx = Pretyping.understand env Evd.empty rt in
- let rt_typ = Typing.unsafe_type_of env Evd.empty rt_as_constr in
- let res_raw_type = Detyping.detype false [] env Evd.empty rt_typ in
+ let rt_as_constr,ctx = Pretyping.understand env (Evd.from_env env) rt in
+ let rt_typ = Typing.unsafe_type_of env (Evd.from_env env) rt_as_constr in
+ let res_raw_type = Detyping.detype false [] env (Evd.from_env env) rt_typ in
let res = fresh_id args_res.to_avoid "_res" in
let new_avoid = res::args_res.to_avoid in
let res_rt = mkGVar res in
@@ -594,8 +594,8 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
and combine the two result
*)
let v_res = build_entry_lc env funnames avoid v in
- let v_as_constr,ctx = Pretyping.understand env Evd.empty v in
- let v_type = Typing.unsafe_type_of env Evd.empty v_as_constr in
+ let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in
+ let v_type = Typing.unsafe_type_of env (Evd.from_env env) v_as_constr in
let new_env =
match n with
Anonymous -> env
@@ -610,10 +610,10 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
let make_discr = make_discr_match brl in
build_entry_lc_from_case env funnames make_discr el brl avoid
| GIf(_,b,(na,e_option),lhs,rhs) ->
- let b_as_constr,ctx = Pretyping.understand env Evd.empty b in
- let b_typ = Typing.unsafe_type_of env Evd.empty b_as_constr in
+ let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in
+ let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in
let (ind,_) =
- try Inductiveops.find_inductive env Evd.empty b_typ
+ try Inductiveops.find_inductive env (Evd.from_env env) b_typ
with Not_found ->
errorlabstrm "" (str "Cannot find the inductive associated to " ++
Printer.pr_glob_constr b ++ str " in " ++
@@ -642,10 +642,10 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
)
nal
in
- let b_as_constr,ctx = Pretyping.understand env Evd.empty b in
- let b_typ = Typing.unsafe_type_of env Evd.empty b_as_constr in
+ let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in
+ let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in
let (ind,_) =
- try Inductiveops.find_inductive env Evd.empty b_typ
+ try Inductiveops.find_inductive env (Evd.from_env env) b_typ
with Not_found ->
errorlabstrm "" (str "Cannot find the inductive associated to " ++
Printer.pr_glob_constr b ++ str " in " ++
@@ -689,8 +689,8 @@ and build_entry_lc_from_case env funname make_discr
in
let types =
List.map (fun (case_arg,_) ->
- let case_arg_as_constr,ctx = Pretyping.understand env Evd.empty case_arg in
- Typing.unsafe_type_of env Evd.empty case_arg_as_constr
+ let case_arg_as_constr,ctx = Pretyping.understand env (Evd.from_env env) case_arg in
+ Typing.unsafe_type_of env (Evd.from_env env) case_arg_as_constr
) el
in
(****** The next works only if the match is not dependent ****)
@@ -737,11 +737,11 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve
List.fold_right
(fun id acc ->
let typ_of_id =
- Typing.unsafe_type_of env_with_pat_ids Evd.empty (mkVar id)
+ Typing.unsafe_type_of env_with_pat_ids (Evd.from_env env) (mkVar id)
in
let raw_typ_of_id =
Detyping.detype false []
- env_with_pat_ids Evd.empty typ_of_id
+ env_with_pat_ids (Evd.from_env env) typ_of_id
in
mkGProd (Name id,raw_typ_of_id,acc))
pat_ids
@@ -785,15 +785,15 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve
List.map3
(fun pat e typ_as_constr ->
let this_pat_ids = ids_of_pat pat in
- let typ = Detyping.detype false [] new_env Evd.empty typ_as_constr in
+ let typ = Detyping.detype false [] new_env (Evd.from_env env) typ_as_constr in
let pat_as_term = pattern_to_term pat in
List.fold_right
(fun id acc ->
if Id.Set.mem id this_pat_ids
then (Prod (Name id),
- let typ_of_id = Typing.unsafe_type_of new_env Evd.empty (mkVar id) in
+ let typ_of_id = Typing.unsafe_type_of new_env (Evd.from_env env) (mkVar id) in
let raw_typ_of_id =
- Detyping.detype false [] new_env Evd.empty typ_of_id
+ Detyping.detype false [] new_env (Evd.from_env env) typ_of_id
in
raw_typ_of_id
)::acc
@@ -894,7 +894,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let new_t =
mkGApp(mkGVar(mk_rel_id this_relname),args'@[res_rt])
in
- let t',ctx = Pretyping.understand env Evd.empty new_t in
+ let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in
let new_env = Environ.push_rel (n,None,t') env in
let new_b,id_to_exclude =
rebuild_cons new_env
@@ -914,7 +914,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
try
observe (str "computing new type for eq : " ++ pr_glob_constr rt);
let t' =
- try fst (Pretyping.understand env Evd.empty t)(*FIXME*)
+ try fst (Pretyping.understand env (Evd.from_env env) t)(*FIXME*)
with e when Errors.noncritical e -> raise Continue
in
let is_in_b = is_free_in id b in
@@ -937,7 +937,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
mkGProd(n,t,new_b),id_to_exclude
with Continue ->
let jmeq = Globnames.IndRef (fst (destInd (jmeq ()))) in
- let ty',ctx = Pretyping.understand env Evd.empty ty in
+ let ty',ctx = Pretyping.understand env (Evd.from_env env) ty in
let ind,args' = Inductive.find_inductive env ty' in
let mib,_ = Global.lookup_inductive (fst ind) in
let nparam = mib.Declarations.mind_nparams in
@@ -949,7 +949,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
GRef (Loc.ghost,Globnames.IndRef (fst ind),None),
(List.map
(fun p -> Detyping.detype false []
- env Evd.empty
+ env (Evd.from_env env)
p) params)@(Array.to_list
(Array.make
(List.length args' - nparam)
@@ -959,7 +959,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
GApp(loc1,GRef(loc2,jmeq,None),[ty;GVar(loc3,id);rt_typ;rt])
in
observe (str "computing new type for jmeq : " ++ pr_glob_constr eq');
- let eq'_as_constr,ctx = Pretyping.understand env Evd.empty eq' in
+ let eq'_as_constr,ctx = Pretyping.understand env (Evd.from_env env) eq' in
observe (str " computing new type for jmeq : done") ;
let new_args =
match kind_of_term eq'_as_constr with
@@ -978,12 +978,12 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
| Name id' ->
(id',Detyping.detype false []
env
- Evd.empty
+ (Evd.from_env env)
arg)::acc
else if isVar var_as_constr
then (destVar var_as_constr,Detyping.detype false []
env
- Evd.empty
+ (Evd.from_env env)
arg)::acc
else acc
)
@@ -1009,7 +1009,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
if is_in_b then b else replace_var_by_term id rt b
in
let new_env =
- let t',ctx = Pretyping.understand env Evd.empty eq' in
+ let t',ctx = Pretyping.understand env (Evd.from_env env) eq' in
Environ.push_rel (n,None,t') env
in
let new_b,id_to_exclude =
@@ -1047,7 +1047,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
else raise Continue
with Continue ->
observe (str "computing new type for prod : " ++ pr_glob_constr rt);
- let t',ctx = Pretyping.understand env Evd.empty t in
+ let t',ctx = Pretyping.understand env (Evd.from_env env) t in
let new_env = Environ.push_rel (n,None,t') env in
let new_b,id_to_exclude =
rebuild_cons new_env
@@ -1063,7 +1063,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
end
| _ ->
observe (str "computing new type for prod : " ++ pr_glob_constr rt);
- let t',ctx = Pretyping.understand env Evd.empty t in
+ let t',ctx = Pretyping.understand env (Evd.from_env env) t in
let new_env = Environ.push_rel (n,None,t') env in
let new_b,id_to_exclude =
rebuild_cons new_env
@@ -1082,7 +1082,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let not_free_in_t id = not (is_free_in id t) in
let new_crossed_types = t :: crossed_types in
observe (str "computing new type for lambda : " ++ pr_glob_constr rt);
- let t',ctx = Pretyping.understand env Evd.empty t in
+ let t',ctx = Pretyping.understand env (Evd.from_env env) t in
match n with
| Name id ->
let new_env = Environ.push_rel (n,None,t') env in
@@ -1104,8 +1104,10 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
| GLetIn(_,n,t,b) ->
begin
let not_free_in_t id = not (is_free_in id t) in
- let t',ctx = Pretyping.understand env Evd.empty t in
- let type_t' = Typing.unsafe_type_of env Evd.empty t' in
+ let evd = (Evd.from_env env) in
+ let t',ctx = Pretyping.understand env evd t in
+ let evd = Evd.from_ctx ctx in
+ let type_t' = Typing.unsafe_type_of env evd t' in
let new_env = Environ.push_rel (n,Some t',type_t') env in
let new_b,id_to_exclude =
rebuild_cons new_env
@@ -1129,7 +1131,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
args (crossed_types)
depth t
in
- let t',ctx = Pretyping.understand env Evd.empty new_t in
+ let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in
let new_env = Environ.push_rel (na,None,t') env in
let new_b,id_to_exclude =
rebuild_cons new_env
@@ -1297,7 +1299,7 @@ let do_build_inductive
let rel_arities = Array.mapi rel_arity funsargs in
Util.Array.fold_left2 (fun env rel_name rel_ar ->
Environ.push_named (rel_name,None,
- fst (with_full_print (Constrintern.interp_constr env Evd.empty) rel_ar)) env) env relnames rel_arities
+ fst (with_full_print (Constrintern.interp_constr env evd) rel_ar)) env) env relnames rel_arities
in
(* and of the real constructors*)
let constr i res =
@@ -1460,8 +1462,17 @@ let do_build_inductive
let build_inductive evd funconstants funsargs returned_types rtl =
+ let pu = !Detyping.print_universes in
+ let cu = !Constrextern.print_universes in
try
- do_build_inductive evd funconstants funsargs returned_types rtl
- with e when Errors.noncritical e -> raise (Building_graph e)
+ Detyping.print_universes := true;
+ Constrextern.print_universes := true;
+ do_build_inductive evd funconstants funsargs returned_types rtl;
+ Detyping.print_universes := pu;
+ Constrextern.print_universes := cu
+ with e when Errors.noncritical e ->
+ Detyping.print_universes := pu;
+ Constrextern.print_universes := cu;
+ raise (Building_graph e)
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index d9d059f8f..eadeebd38 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -145,15 +145,14 @@ let interp_casted_constr_with_implicits env sigma impls c =
let build_newrecursive
lnameargsardef =
- let env0 = Global.env()
- and sigma = Evd.empty
- in
+ let env0 = Global.env() in
+ let sigma = Evd.from_env env0 in
let (rec_sign,rec_impls) =
List.fold_left
(fun (env,impls) (((_,recname),_),bl,arityc,_) ->
let arityc = Constrexpr_ops.prod_constr_expr arityc bl in
let arity,ctx = Constrintern.interp_type env0 sigma arityc in
- let evdref = ref (Evd.from_env env0) in
+ let evdref = ref (Evd.from_env env0) in
let _, (_, impls') = Constrintern.interp_context_evars env evdref bl in
let impl = Constrintern.compute_internalization_data env0 Constrintern.Recursive arity impls' in
(Environ.push_named (recname,None,arity) env, Id.Map.add recname impl impls))
@@ -228,7 +227,7 @@ let process_vernac_interp_error e =
let derive_inversion fix_names =
try
- let evd' = Evd.empty in
+ let evd' = Evd.from_env (Global.env ()) in
(* we first transform the fix_names identifier into their corresponding constant *)
let evd',fix_names_as_constant =
List.fold_right
@@ -355,9 +354,11 @@ let generate_principle (evd:Evd.evar_map ref) pconstants on_error
List.map_i
(fun i x ->
let princ = Indrec.lookup_eliminator (ind_kn,i) (InProp) in
- let evd',uprinc = Evd.fresh_global (Global.env ()) !evd princ in
+ let env = Global.env () in
+ let evd = ref (Evd.from_env env) in
+ let evd',uprinc = Evd.fresh_global env !evd princ in
let _ = evd := evd' in
- let princ_type = Typing.e_type_of ~refresh:true (Global.env ()) evd uprinc in
+ let princ_type = Typing.e_type_of ~refresh:true env evd uprinc in
Functional_principles_types.generate_functional_principle
evd
interactive_proof
@@ -394,7 +395,7 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp
(Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname)) in
evd,((destConst c)::l)
)
- (Evd.empty,[])
+ (Evd.from_env (Global.env ()),[])
fixpoint_exprl
in
evd,List.rev rev_pconstants
@@ -408,7 +409,7 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp
(Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname)) in
evd,((destConst c)::l)
)
- (Evd.empty,[])
+ (Evd.from_env (Global.env ()),[])
fixpoint_exprl
in
evd,List.rev rev_pconstants
@@ -594,9 +595,9 @@ let rebuild_bl (aux,assoc) bl typ = rebuild_bl (aux,assoc) bl typ
let recompute_binder_list (fixpoint_exprl : (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) =
let fixl,ntns = Command.extract_fixpoint_components false fixpoint_exprl in
- let ((_,_,typel),_,_) = Command.interp_fixpoint fixl ntns in
+ let ((_,_,typel),ctx,_) = Command.interp_fixpoint fixl ntns in
let constr_expr_typel =
- with_full_print (List.map (Constrextern.extern_constr false (Global.env ()) Evd.empty)) typel in
+ with_full_print (List.map (Constrextern.extern_constr false (Global.env ()) (Evd.from_ctx ctx))) typel in
let fixpoint_exprl_with_new_bl =
List.map2 (fun ((lna,(rec_arg_opt,rec_order),bl,ret_typ,opt_body),notation_list) fix_typ ->
@@ -625,7 +626,7 @@ let do_generate_principle pconstants on_error register_built interactive_proof
let using_lemmas = [] in
let pre_hook pconstants =
generate_principle
- (ref (Evd.empty))
+ (ref (Evd.from_env (Global.env ())))
pconstants
on_error
true
@@ -649,7 +650,7 @@ let do_generate_principle pconstants on_error register_built interactive_proof
let body = match body with | Some body -> body | None -> user_err_loc (Loc.ghost,"Function",str "Body of Function must be given") in
let pre_hook pconstants =
generate_principle
- (ref Evd.empty)
+ (ref (Evd.from_env (Global.env ())))
pconstants
on_error
true
@@ -680,7 +681,7 @@ let do_generate_principle pconstants on_error register_built interactive_proof
let evd,pconstants =
if register_built
then register_struct is_rec fixpoint_exprl
- else (Evd.empty,pconstants)
+ else (Evd.from_env (Global.env ()),pconstants)
in
let evd = ref evd in
generate_principle
@@ -835,10 +836,11 @@ let make_graph (f_ref:global_reference) =
| None -> error "Cannot build a graph over an axiom !"
| Some body ->
let env = Global.env () in
+ let sigma = Evd.from_env env in
let extern_body,extern_type =
with_full_print (fun () ->
- (Constrextern.extern_constr false env Evd.empty body,
- Constrextern.extern_type false env Evd.empty
+ (Constrextern.extern_constr false env sigma body,
+ Constrextern.extern_type false env sigma
((*FIXME*) Typeops.type_of_constant_type env c_body.const_type)
)
)
diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml
index 1c409500e..35bd1c36d 100644
--- a/plugins/funind/indfun_common.ml
+++ b/plugins/funind/indfun_common.ml
@@ -180,9 +180,10 @@ let get_proof_clean do_reduce =
let with_full_print f a =
let old_implicit_args = Impargs.is_implicit_args ()
and old_strict_implicit_args = Impargs.is_strict_implicit_args ()
- and old_contextual_implicit_args = Impargs.is_contextual_implicit_args ()
- in
+ and old_contextual_implicit_args = Impargs.is_contextual_implicit_args () in
let old_rawprint = !Flags.raw_print in
+ let old_printuniverses = !Constrextern.print_universes in
+ Constrextern.print_universes := true;
Flags.raw_print := true;
Impargs.make_implicit_args false;
Impargs.make_strict_implicit_args false;
@@ -195,6 +196,7 @@ let with_full_print f a =
Impargs.make_strict_implicit_args old_strict_implicit_args;
Impargs.make_contextual_implicit_args old_contextual_implicit_args;
Flags.raw_print := old_rawprint;
+ Constrextern.print_universes := old_printuniverses;
Dumpglob.continue ();
res
with
@@ -203,6 +205,7 @@ let with_full_print f a =
Impargs.make_strict_implicit_args old_strict_implicit_args;
Impargs.make_contextual_implicit_args old_contextual_implicit_args;
Flags.raw_print := old_rawprint;
+ Constrextern.print_universes := old_printuniverses;
Dumpglob.continue ();
raise reraise
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index 89ceb751a..d97940142 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -760,7 +760,8 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
let funs_constr = Array.map mkConstU funs in
States.with_state_protection_on_exception
(fun () ->
- let evd = ref Evd.empty in
+ let env = Global.env () in
+ let evd = ref (Evd.from_env env) in
let graphs_constr = Array.map mkInd graphs in
let lemmas_types_infos =
Util.Array.map2_i
@@ -829,7 +830,6 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
)
funs;
- (* let evd = ref Evd.empty in *)
let lemmas_types_infos =
Util.Array.map2_i
(fun i f_constr graph ->
@@ -875,7 +875,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
i*)
let lem_id = mk_complete_id f_id in
Lemmas.start_proof lem_id
- (Decl_kinds.Global,Flags.is_universe_polymorphism (),(Decl_kinds.Proof Decl_kinds.Theorem)) !evd
+ (Decl_kinds.Global,Flags.is_universe_polymorphism (),(Decl_kinds.Proof Decl_kinds.Theorem)) sigma
(fst lemmas_types_infos.(i))
(Lemmas.mk_hook (fun _ _ -> ()));
ignore (Pfedit.by
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index 9c09a4a0c..ac7140b9b 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -194,7 +194,7 @@ let (value_f:constr list -> global_reference -> constr) =
Anonymous)],
GVar(d0,v_id)])
in
- let body = fst (understand env Evd.empty glob_body)(*FIXME*) in
+ let body = fst (understand env (Evd.from_env env) glob_body)(*FIXME*) in
it_mkLambda_or_LetIn body context
let (declare_f : Id.t -> logical_kind -> constr list -> global_reference -> global_reference) =
@@ -1293,8 +1293,9 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
ref_ := Some lemma ;
let lid = ref [] in
let h_num = ref (-1) in
+ let env = Global.env () in
Proof_global.discard_all ();
- build_proof Evd.empty
+ build_proof (Evd.from_env env)
( fun gls ->
let hid = next_ident_away_in_goal h_id (pf_ids_of_hyps gls) in
observe_tclTHENLIST (str "")
@@ -1513,7 +1514,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
let relation =
fst (*FIXME*)(interp_constr
env_with_pre_rec_args
- Evd.empty
+ (Evd.from_env env_with_pre_rec_args)
r)
in
let tcc_lemma_name = add_suffix function_name "_tcc" in
diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml
index f73a15dba..942ca15a5 100644
--- a/plugins/setoid_ring/newring.ml
+++ b/plugins/setoid_ring/newring.ml
@@ -145,14 +145,19 @@ let _ = add_tacdef false ((Loc.ghost,Id.of_string"ring_closed_term"
(****************************************************************************)
let ic c =
- let env = Global.env() and sigma = Evd.empty in
+ let env = Global.env() in
+ let sigma = Evd.from_env env in
Constrintern.interp_open_constr env sigma c
let ic_unsafe c = (*FIXME remove *)
- let env = Global.env() and sigma = Evd.empty in
+ let env = Global.env() in
+ let sigma = Evd.from_env env in
fst (Constrintern.interp_constr env sigma c)
-let ty c = Typing.unsafe_type_of (Global.env()) Evd.empty c
+let ty c =
+ let env = Global.env() in
+ let sigma = Evd.from_env env in
+ Typing.unsafe_type_of env sigma c
let decl_constant na ctx c =
let vars = Universes.universes_of_constr c in
diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml
index ac1692f45..a2189d5e4 100644
--- a/pretyping/evarsolve.ml
+++ b/pretyping/evarsolve.ml
@@ -107,6 +107,7 @@ let refresh_universes ?(inferred=false) ?(onlyalg=false) pbty env evd t =
let get_type_of_refresh ?(polyprop=true) ?(lax=false) env sigma c =
let ty = Retyping.get_type_of ~polyprop ~lax env sigma c in
refresh_universes (Some false) env sigma ty
+
(************************)
(* Unification results *)
@@ -127,6 +128,32 @@ let add_conv_oriented_pb ?(tail=true) (pbty,env,t1,t2) evd =
| Some false -> add_conv_pb ~tail (Reduction.CUMUL,env,t2,t1) evd
| None -> add_conv_pb ~tail (Reduction.CONV,env,t1,t2) evd
+(* We retype applications to ensure the universe constraints are collected *)
+
+let recheck_applications conv_algo env evdref t =
+ let rec aux env t =
+ match kind_of_term t with
+ | App (f, args) ->
+ let () = aux env f in
+ let fty = Retyping.get_type_of env !evdref f in
+ let argsty = Array.map (fun x -> aux env x; Retyping.get_type_of env !evdref x) args in
+ let rec aux i ty =
+ if i < Array.length argsty then
+ match kind_of_term (whd_betadeltaiota env !evdref ty) with
+ | Prod (na, dom, codom) ->
+ (match conv_algo env !evdref Reduction.CUMUL argsty.(i) dom with
+ | Success evd -> evdref := evd;
+ aux (succ i) (subst1 args.(i) codom)
+ | UnifFailure (evd, reason) ->
+ Pretype_errors.error_cannot_unify env evd ~reason (argsty.(i), dom))
+ | _ -> assert false
+ else ()
+ in aux 0 fty
+ | _ ->
+ iter_constr_with_full_binders (fun d env -> push_rel d env) aux env t
+ in aux env t
+
+
(*------------------------------------*
* Restricting existing evars *
*------------------------------------*)
@@ -1404,10 +1431,10 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
evdref := restrict_evar evd (fst ev'') None (UpdateWith candidates);
evar'')
| None ->
- (* Evar/Rigid problem (or assimilated if not normal): we "imitate" *)
- map_constr_with_full_binders (fun d (env,k) -> push_rel d env, k+1)
- imitate envk t in
-
+ (* Evar/Rigid problem (or assimilated if not normal): we "imitate" *)
+ map_constr_with_full_binders (fun d (env,k) -> push_rel d env, k+1)
+ imitate envk t
+ in
let rhs = whd_beta evd rhs (* heuristic *) in
let fast rhs =
let filter_ctxt = evar_filtered_context evi in
@@ -1426,8 +1453,12 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
in
let body =
if fast rhs then nf_evar evd rhs
- else imitate (env,0) rhs
- in (!evdref,body)
+ else
+ let t' = imitate (env,0) rhs in
+ if !progress then
+ (recheck_applications conv_algo (evar_env evi) evdref t'; t')
+ else t'
+ in (!evdref,body)
(* [define] tries to solve the problem "?ev[args] = rhs" when "?ev" is
* an (uninstantiated) evar such that "hyps |- ?ev : typ". Otherwise said,
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 2858a5c1f..f18657da8 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -102,14 +102,27 @@ let ((constr_in : constr -> Dyn.t),
(** Miscellaneous interpretation functions *)
let interp_universe_level_name evd s =
let names, _ = Universes.global_universe_names () in
- try
- let id = try Id.of_string s with _ -> raise Not_found in
- evd, Idmap.find id names
- with Not_found ->
- try let level = Evd.universe_of_name evd s in
- evd, level
- with Not_found ->
- new_univ_level_variable ~name:s univ_rigid evd
+ if CString.string_contains s "." then
+ match List.rev (CString.split '.' s) with
+ | [] -> anomaly (str"Invalid universe name " ++ str s)
+ | n :: dp ->
+ let num = int_of_string n in
+ let dp = DirPath.make (List.map Id.of_string dp) in
+ let level = Univ.Level.make dp num in
+ let evd =
+ try Evd.add_global_univ evd level
+ with Univ.AlreadyDeclared -> evd
+ in evd, level
+ else
+ try
+ let id =
+ try Id.of_string s with _ -> raise Not_found in
+ evd, Idmap.find id names
+ with Not_found ->
+ try let level = Evd.universe_of_name evd s in
+ evd, level
+ with Not_found ->
+ new_univ_level_variable ~name:s univ_rigid evd
let interp_universe evd = function
| [] -> let evd, l = new_univ_level_variable univ_rigid evd in
@@ -632,7 +645,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) env evdref (lvar : ltac_
match evar_kind_of_term !evdref resj.uj_val with
| App (f,args) ->
let f = whd_evar !evdref f in
- if isInd f && is_template_polymorphic env f then
+ if is_template_polymorphic env f then
(* Special case for inductive type applications that must be
refreshed right away. *)
let sigma = !evdref in
diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml
index 18e83056b..2ef289650 100644
--- a/pretyping/typeclasses.ml
+++ b/pretyping/typeclasses.ml
@@ -370,7 +370,7 @@ let add_instance check inst =
List.iter (fun (path, pri, c) -> add_instance_hint (IsConstr c) path
(is_local inst) pri poly)
(build_subclasses ~check:(check && not (isVarRef inst.is_impl))
- (Global.env ()) Evd.empty inst.is_impl inst.is_pri)
+ (Global.env ()) (Evd.from_env (Global.env ())) inst.is_impl inst.is_pri)
let rebuild_instance (action, inst) =
let () = match action with
diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml
index b88913dec..2ab3dc67a 100644
--- a/proofs/pfedit.ml
+++ b/proofs/pfedit.ml
@@ -153,9 +153,9 @@ let build_by_tactic ?(side_eff=true) env ctx ?(poly=false) typ tac =
let ce, status, univs = build_constant_by_tactic id ctx sign ~goal_kind:gk typ tac in
let ce = if side_eff then Term_typing.handle_entry_side_effects env ce else { ce with const_entry_body = Future.chain ~pure:true ce.const_entry_body (fun (pt, _) -> pt, Declareops.no_seff) } in
let (cb, ctx), se = Future.force ce.const_entry_body in
+ let univs' = Evd.merge_context_set Evd.univ_rigid (Evd.from_ctx univs) ctx in
assert(Declareops.side_effects_is_empty se);
- assert(Univ.ContextSet.is_empty ctx);
- cb, status, univs
+ cb, status, Evd.evar_universe_context univs'
let refine_by_tactic env sigma ty tac =
(** Save the initial side-effects to restore them afterwards. We set the
diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml
index 21009d120..b5e25cc7c 100644
--- a/proofs/proof_global.ml
+++ b/proofs/proof_global.ml
@@ -293,16 +293,14 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now fpl =
let body = c and typ = nf t in
let used_univs_body = Universes.universes_of_constr body in
let used_univs_typ = Universes.universes_of_constr typ in
- if keep_body_ucst_separate then
+ if keep_body_ucst_separate || not (Declareops.side_effects_is_empty eff) then
let initunivs = Evd.evar_context_universe_context initial_euctx in
let ctx = Evd.evar_universe_context_set initunivs universes in
(* For vi2vo compilation proofs are computed now but we need to
* complement the univ constraints of the typ with the ones of
* the body. So we keep the two sets distinct. *)
let ctx_body = restrict_universe_context ctx used_univs_body in
- let ctx_typ = restrict_universe_context ctx used_univs_typ in
- let univs_typ = Univ.ContextSet.to_context ctx_typ in
- (univs_typ, typ), ((body, ctx_body), eff)
+ (initunivs, typ), ((body, ctx_body), eff)
else
let initunivs = Univ.UContext.empty in
let ctx = Evd.evar_universe_context_set initunivs universes in
diff --git a/stm/lemmas.ml b/stm/lemmas.ml
index 934642f12..33db349c8 100644
--- a/stm/lemmas.ml
+++ b/stm/lemmas.ml
@@ -457,8 +457,12 @@ let start_proof_com kind thms hook =
let recguard,thms,snl = look_for_possibly_mutual_statements thms in
let evd, nf = Evarutil.nf_evars_and_universes !evdref in
let thms = List.map (fun (n, (t, info)) -> (n, (nf t, info))) thms in
- start_proof_with_initialization kind evd
- recguard thms snl hook
+ let evd =
+ if pi2 kind then evd
+ else (* We fix the variables to ensure they won't be lowered to Set *)
+ Evd.fix_undefined_variables evd
+ in
+ start_proof_with_initialization kind evd recguard thms snl hook
(* Saving a proof *)
@@ -511,4 +515,5 @@ let save_proof ?proof = function
let get_current_context () =
try Pfedit.get_current_goal_context ()
with e when Logic.catchable_exception e ->
- (Evd.empty, Global.env())
+ let env = Global.env () in
+ (Evd.from_env env, env)
diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml
index 2b3fadf7f..3a9d40de0 100644
--- a/tactics/autorewrite.ml
+++ b/tactics/autorewrite.ml
@@ -292,10 +292,13 @@ let find_applied_relation metas loc env sigma c left2right =
(* To add rewriting rules to a base *)
let add_rew_rules base lrul =
let counter = ref 0 in
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
let lrul =
List.fold_left
(fun dn (loc,(c,ctx),b,t) ->
- let info = find_applied_relation false loc (Global.env ()) Evd.empty c b in
+ let sigma = Evd.merge_context_set Evd.univ_rigid sigma ctx in
+ let info = find_applied_relation false loc env sigma c b in
let pat = if b then info.hyp_left else info.hyp_right in
let rul = { rew_lemma = c; rew_type = info.hyp_ty;
rew_pat = pat; rew_ctx = ctx; rew_l2r = b;
diff --git a/tactics/equality.ml b/tactics/equality.ml
index a10d8a074..4c48003b9 100644
--- a/tactics/equality.ml
+++ b/tactics/equality.ml
@@ -335,7 +335,9 @@ let find_elim hdcncl lft2rgt dep cls ot gl =
| Ind (ind,u) ->
let c, eff = find_scheme scheme_name ind in
(* MS: cannot use pf_constr_of_global as the eliminator might be generated by side-effect *)
- let sigma, elim = Evd.fresh_global (Global.env ()) (Proofview.Goal.sigma gl) (ConstRef c) in
+ let sigma, elim =
+ Evd.fresh_global (Global.env ()) (Proofview.Goal.sigma gl) (ConstRef c)
+ in
sigma, elim, eff
| _ -> assert false
@@ -899,7 +901,7 @@ let discrimination_pf env sigma e (t,t1,t2) discriminator lbeq =
let i = build_coq_I () in
let absurd_term = build_coq_False () in
let eq_elim, eff = ind_scheme_of_eq lbeq in
- let sigma, eq_elim = Evd.fresh_global env sigma eq_elim in
+ let sigma, eq_elim = Evd.fresh_global (Global.env ()) sigma eq_elim in
sigma, (applist (eq_elim, [t;t1;mkNamedLambda e t discriminator;i;t2]), absurd_term),
eff
diff --git a/tactics/extratactics.ml4 b/tactics/extratactics.ml4
index af0870bc9..cab74968d 100644
--- a/tactics/extratactics.ml4
+++ b/tactics/extratactics.ml4
@@ -262,16 +262,15 @@ TACTIC EXTEND rewrite_star
(* Hint Rewrite *)
let add_rewrite_hint bases ort t lcsr =
- let env = Global.env() and sigma = Evd.empty in
+ let env = Global.env() in
+ let sigma = Evd.from_env env in
let poly = Flags.is_universe_polymorphism () in
let f ce =
let c, ctx = Constrintern.interp_constr env sigma ce in
let ctx =
- if poly then
- Evd.evar_universe_context_set Univ.UContext.empty ctx
- else
- let cstrs = Evd.evar_universe_context_constraints ctx in
- (Global.add_constraints cstrs; Univ.ContextSet.empty)
+ let ctx = Evd.evar_universe_context_set Univ.UContext.empty ctx in
+ if poly then ctx
+ else (Global.push_context_set false ctx; Univ.ContextSet.empty)
in
Constrexpr_ops.constr_loc ce, (c, ctx), ort, t in
let eqs = List.map f lcsr in
@@ -490,7 +489,9 @@ let inTransitivity : bool * constr -> obj =
(* Main entry points *)
let add_transitivity_lemma left lem =
- let lem',ctx (*FIXME*) = Constrintern.interp_constr (Global.env ()) Evd.empty lem in
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
+ let lem',ctx (*FIXME*) = Constrintern.interp_constr env sigma lem in
add_anonymous_leaf (inTransitivity (left,lem'))
(* Vernacular syntax *)
diff --git a/tactics/hints.ml b/tactics/hints.ml
index 0df1a35c6..a7eae667d 100644
--- a/tactics/hints.ml
+++ b/tactics/hints.ml
@@ -1085,8 +1085,10 @@ let prepare_hint check env init (sigma,c) =
let interp_hints poly =
fun h ->
+ let env = (Global.env()) in
+ let sigma = Evd.from_env env in
let f c =
- let evd,c = Constrintern.interp_open_constr (Global.env()) Evd.empty c in
+ let evd,c = Constrintern.interp_open_constr env sigma c in
prepare_hint true (Global.env()) Evd.empty (evd,c) in
let fref r =
let gr = global_with_alias r in
@@ -1135,7 +1137,8 @@ let add_hints local dbnames0 h =
if String.List.mem "nocore" dbnames0 then
error "The hint database \"nocore\" is meant to stay empty.";
let dbnames = if List.is_empty dbnames0 then ["core"] else dbnames0 in
- let env = Global.env() and sigma = Evd.empty in
+ let env = Global.env() in
+ let sigma = Evd.from_env env in
match h with
| HintsResolveEntry lhints -> add_resolves env sigma lhints local dbnames
| HintsImmediateEntry lhints -> add_trivials env sigma lhints local dbnames
diff --git a/tactics/rewrite.ml b/tactics/rewrite.ml
index c64a1226a..937ad2b9d 100644
--- a/tactics/rewrite.ml
+++ b/tactics/rewrite.ml
@@ -1797,11 +1797,13 @@ let proper_projection r ty =
it_mkLambda_or_LetIn app ctx
let declare_projection n instance_id r =
- let c,uctx = Universes.fresh_global_instance (Global.env()) r in
let poly = Global.is_polymorphic r in
- let ty = Retyping.get_type_of (Global.env ()) Evd.empty c in
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
+ let evd,c = Evd.fresh_global env sigma r in
+ let ty = Retyping.get_type_of env sigma c in
let term = proper_projection c ty in
- let typ = Typing.unsafe_type_of (Global.env ()) Evd.empty term in
+ let typ = Typing.unsafe_type_of env sigma term in
let ctx, typ = decompose_prod_assum typ in
let typ =
let n =
@@ -1824,15 +1826,16 @@ let declare_projection n instance_id r =
in
let typ = it_mkProd_or_LetIn typ ctx in
let cst =
- Declare.definition_entry ~types:typ ~poly ~univs:(Univ.ContextSet.to_context uctx)
- term
+ Declare.definition_entry ~types:typ ~poly
+ ~univs:(Evd.universe_context sigma) term
in
ignore(Declare.declare_constant n
(Entries.DefinitionEntry cst, Decl_kinds.IsDefinition Decl_kinds.Definition))
let build_morphism_signature m =
let env = Global.env () in
- let m,ctx = Constrintern.interp_constr env (Evd.from_env env) m in
+ let sigma = Evd.from_env env in
+ let m,ctx = Constrintern.interp_constr env sigma m in
let sigma = Evd.from_ctx ctx in
let t = Typing.unsafe_type_of env sigma m in
let cstrs =
@@ -1844,7 +1847,7 @@ let build_morphism_signature m =
in aux t
in
let evars, t', sig_, cstrs =
- PropGlobal.build_signature (Evd.empty, Evar.Set.empty) env t cstrs None in
+ PropGlobal.build_signature (sigma, Evar.Set.empty) env t cstrs None in
let evd = ref evars in
let _ = List.iter
(fun (ty, rel) ->
@@ -1861,9 +1864,10 @@ let build_morphism_signature m =
let default_morphism sign m =
let env = Global.env () in
- let t = Typing.unsafe_type_of env Evd.empty m in
+ let sigma = Evd.from_env env in
+ let t = Typing.unsafe_type_of env sigma m in
let evars, _, sign, cstrs =
- PropGlobal.build_signature (Evd.empty, Evar.Set.empty) env t (fst sign) (snd sign)
+ PropGlobal.build_signature (sigma, Evar.Set.empty) env t (fst sign) (snd sign)
in
let evars, morph = app_poly_check env evars PropGlobal.proper_type [| t; sign; m |] in
let evars, mor = resolve_one_typeclass env (goalevars evars) morph in
@@ -1894,7 +1898,7 @@ let add_morphism_infer glob m n =
let poly = Flags.is_universe_polymorphism () in
let instance_id = add_suffix n "_Proper" in
let instance = build_morphism_signature m in
- let evd = Evd.empty (*FIXME *) in
+ let evd = Evd.from_env (Global.env ()) in
if Lib.is_modtype () then
let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest instance_id
(Entries.ParameterEntry
diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml
index 7d1cc3341..bc82e9ef4 100644
--- a/tactics/tacticals.ml
+++ b/tactics/tacticals.ml
@@ -593,10 +593,12 @@ module New = struct
(* c should be of type A1->.. An->B with B an inductive definition *)
let general_elim_then_using mk_elim
isrec allnames tac predicate ind (c, t) =
- Proofview.Goal.nf_enter begin fun gl ->
- let indclause = Tacmach.New.of_old (fun gl -> mk_clenv_from gl (c, t)) gl in
- (** FIXME: evar leak. *)
+ Proofview.Goal.nf_enter
+ begin fun gl ->
let sigma, elim = Tacmach.New.of_old (mk_elim ind) gl in
+ Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma)
+ (Proofview.Goal.nf_enter begin fun gl ->
+ let indclause = Tacmach.New.of_old (fun gl -> mk_clenv_from gl (c, t)) gl in
(* applying elimination_scheme just a little modified *)
let elimclause = Tacmach.New.of_old (fun gls -> mk_clenv_from gls (elim,Tacmach.New.pf_unsafe_type_of gl elim)) gl in
let indmv =
@@ -647,7 +649,7 @@ module New = struct
Proofview.tclTHEN
(Clenvtac.clenv_refine false clenv')
(Proofview.tclEXTEND [] tclIDTAC branchtacs)
- end
+ end) end
let elimination_then tac c =
Proofview.Goal.nf_enter begin fun gl ->
diff --git a/test-suite/Makefile b/test-suite/Makefile
index 39c36d541..31b212900 100644
--- a/test-suite/Makefile
+++ b/test-suite/Makefile
@@ -388,7 +388,7 @@ misc/deps-order.log:
} > "$@"
# Sort universes for the whole standard library
-EXPECTED_UNIVERSES := 3
+EXPECTED_UNIVERSES := 5
universes: misc/universes.log
misc/universes.log: misc/universes/all_stdlib.v
@echo "TEST misc/universes"
diff --git a/test-suite/bugs/closed/2016.v b/test-suite/bugs/closed/2016.v
new file mode 100644
index 000000000..13ec5bea9
--- /dev/null
+++ b/test-suite/bugs/closed/2016.v
@@ -0,0 +1,62 @@
+(* Coq 8.2beta4 *)
+Require Import Classical_Prop.
+
+Record coreSemantics : Type := CoreSemantics {
+ core: Type;
+ corestep: core -> core -> Prop;
+ corestep_fun: forall q q1 q2, corestep q q1 -> corestep q q2 -> q1 = q2
+}.
+
+Definition state : Type := {sem: coreSemantics & sem.(core)}.
+
+Inductive step: state -> state -> Prop :=
+ | step_core: forall sem st st'
+ (Hcs: sem.(corestep) st st'),
+ step (existT _ sem st) (existT _ sem st').
+
+Lemma step_fun: forall st1 st2 st2', step st1 st2 -> step st1 st2' -> st2 = st2'.
+Proof.
+intros.
+inversion H; clear H; subst. inversion H0; clear H0; subst; auto.
+generalize (inj_pairT2 _ _ _ _ _ H2); clear H2; intro; subst.
+rewrite (corestep_fun _ _ _ _ Hcs Hcs0); auto.
+Qed.
+
+Record oe_core := oe_Core {
+ in_core: Type;
+ in_corestep: in_core -> in_core -> Prop;
+ in_corestep_fun: forall q q1 q2, in_corestep q q1 -> in_corestep q q2 -> q1 = q2;
+ in_q: in_core
+}.
+
+Definition oe2coreSem (oec : oe_core) : coreSemantics :=
+ CoreSemantics oec.(in_core) oec.(in_corestep) oec.(in_corestep_fun).
+
+Definition oe_corestep (q q': oe_core) :=
+ step (existT _ (oe2coreSem q) q.(in_q)) (existT _ (oe2coreSem q') q'.(in_q)).
+
+Lemma inj_pairT1: forall (U: Type) (P: U -> Type) (p1 p2: U) x y,
+ existT P p1 x = existT P p2 y -> p1=p2.
+Proof. intros; injection H; auto.
+Qed.
+
+Definition f := CoreSemantics oe_core.
+
+Lemma oe_corestep_fun: forall q q1 q2,
+ oe_corestep q q1 -> oe_corestep q q2 -> q1 = q2.
+Proof.
+unfold oe_corestep; intros.
+assert (HH:= step_fun _ _ _ H H0); clear H H0.
+destruct q1; destruct q2; unfold oe2coreSem; simpl in *.
+generalize (inj_pairT1 _ _ _ _ _ _ HH); clear HH; intros.
+injection H; clear H; intros.
+revert in_q1 in_corestep1 in_corestep_fun1
+ H.
+pattern in_core1.
+apply eq_ind_r with (x := in_core0).
+admit.
+apply sym_eq.
+(** good to here **)
+Show Universes.
+Print Universes.
+Fail apply H0. \ No newline at end of file
diff --git a/test-suite/bugs/closed/2584.v b/test-suite/bugs/closed/2584.v
new file mode 100644
index 000000000..a5f4ae64a
--- /dev/null
+++ b/test-suite/bugs/closed/2584.v
@@ -0,0 +1,89 @@
+Require Import List.
+
+Set Implicit Arguments.
+
+Definition err : Type := unit.
+
+Inductive res (A: Type) : Type :=
+| OK: A -> res A
+| Error: err -> res A.
+
+Implicit Arguments Error [A].
+
+Set Printing Universes.
+
+Section FOO.
+
+Inductive ftyp : Type :=
+ | Funit : ftyp
+ | Ffun : list ftyp -> ftyp
+ | Fref : area -> ftyp
+with area : Type :=
+ | Stored : ftyp -> area
+.
+
+Print ftyp.
+(* yields:
+Inductive ftyp : Type (* Top.27429 *) :=
+ Funit : ftyp | Ffun : list ftyp -> ftyp | Fref : area -> ftyp
+ with area : Type (* Set *) := Stored : ftyp -> area
+*)
+
+Fixpoint tc_wf_type (ftype: ftyp) {struct ftype}: res unit :=
+ match ftype with
+ | Funit => OK tt
+ | Ffun args =>
+ ((fix tc_wf_types (ftypes: list ftyp){struct ftypes}: res unit :=
+ match ftypes with
+ | nil => OK tt
+ | t::ts =>
+ match tc_wf_type t with
+ | OK tt => tc_wf_types ts
+ | Error m => Error m
+ end
+ end) args)
+ | Fref a => tc_wf_area a
+ end
+with tc_wf_area (ar:area): res unit :=
+ match ar with
+ | Stored c => tc_wf_type c
+ end.
+
+End FOO.
+
+Print ftyp.
+(* yields:
+Inductive ftyp : Type (* Top.27465 *) :=
+ Funit : ftyp | Ffun : list ftyp -> ftyp | Fref : area -> ftyp
+ with area : Set := Stored : ftyp -> area
+*)
+
+Fixpoint tc_wf_type' (ftype: ftyp) {struct ftype}: res unit :=
+ match ftype with
+ | Funit => OK tt
+ | Ffun args =>
+ ((fix tc_wf_types (ftypes: list ftyp){struct ftypes}: res unit :=
+ match ftypes with
+ | nil => OK tt
+ | t::ts =>
+ match tc_wf_type' t with
+ | OK tt => tc_wf_types ts
+ | Error m => Error m
+ end
+ end) args)
+ | Fref a => tc_wf_area' a
+ end
+with tc_wf_area' (ar:area): res unit :=
+ match ar with
+ | Stored c => tc_wf_type' c
+ end.
+
+(* yields:
+Error:
+Incorrect elimination of "ar" in the inductive type "area":
+the return type has sort "Type (* max(Set, Top.27424) *)" while it
+should be "Prop" or "Set".
+Elimination of an inductive object of sort Set
+is not allowed on a predicate in sort Type
+because strong elimination on non-small inductive types leads to paradoxes.
+*) \ No newline at end of file
diff --git a/test-suite/bugs/closed/3309.v b/test-suite/bugs/closed/3309.v
deleted file mode 100644
index 980431576..000000000
--- a/test-suite/bugs/closed/3309.v
+++ /dev/null
@@ -1,334 +0,0 @@
-Require Import TestSuite.admit.
-(* -*- coq-prog-args: ("-emacs" "-impredicative-set") -*- *)
-(* File reduced by coq-bug-finder from original input, then from 5968 lines to 11933 lines, then from 11239 lines to 11231 lines, then from 10365 lines to 446 lines, then from 456 lines to 379 lines, then from 391 lines to 373 lines, then from 369 lines to 351 lines, then from 350 lines to 340 lines, then from 348 lines to 320 lines, then from 328 lines to 302 lines *)
-Set Universe Polymorphism.
-Record sigT' {A} (P : A -> Type) := existT' { projT1' : A; projT2' : P projT1' }.
-Notation "{ x : A &' P }" := (sigT' (A := A) (fun x => P)) : type_scope.
-Arguments existT' {A} P _ _.
-Axiom admit : forall {T}, T.
-Notation paths := identity .
-
-Unset Automatic Introduction.
-
-Definition UU := Set.
-
-Definition dirprod ( X Y : UU ) := sigT' ( fun x : X => Y ) .
-Definition dirprodpair { X Y : UU } := existT' ( fun x : X => Y ) .
-
-Definition ddualand { X Y P : UU } (xp : ( X -> P ) -> P ) ( yp : ( Y -> P ) -> P ) : ( dirprod X Y -> P ) -> P.
-Proof.
- intros X Y P xp yp X0 .
- set ( int1 := fun ypp : ( ( Y -> P ) -> P ) => fun x : X => yp ( fun y : Y => X0 ( dirprodpair x y) ) ) .
- apply ( xp ( int1 yp ) ) .
-Defined .
-Definition weq ( X Y : UU ) : UU .
-intros; exact ( sigT' (fun f:X->Y => admit) ).
-Defined.
-Definition pr1weq ( X Y : UU):= @projT1' _ _ : weq X Y -> (X -> Y).
-Coercion pr1weq : weq >-> Funclass.
-
-Definition invweq { X Y : UU } ( w : weq X Y ) : weq Y X .
-admit.
-Defined.
-
-Definition hProp := sigT' (fun X : Type => admit).
-
-Definition hProppair ( X : UU ) ( is : admit ) : hProp@{i j Set k}.
-intros; exact (existT' (fun X : UU => admit ) X is ).
-Defined.
-Definition hProptoType := @projT1' _ _ : hProp -> Type .
-Coercion hProptoType: hProp >-> Sortclass.
-
-Definition ishinh_UU ( X : UU ) : UU := forall P: Set, ( ( X -> P ) -> P ).
-
-Definition ishinh ( X : UU ) : hProp := hProppair ( ishinh_UU X ) admit.
-
-Definition hinhfun { X Y : UU } ( f : X -> Y ) : ishinh_UU X -> ishinh_UU Y.
-intros X Y f; exact ( fun isx : ishinh X => fun P : _ => fun yp : Y -> P => isx P ( fun x : X => yp ( f x ) ) ).
-Defined.
-
-Definition hinhuniv { X : UU } { P : hProp } ( f : X -> P ) ( wit : ishinh_UU X ) : P.
-intros; exact ( wit P f ).
-Defined.
-
-Definition hinhand { X Y : UU } ( inx1 : ishinh_UU X ) ( iny1 : ishinh_UU Y) : ishinh ( dirprod X Y ).
-intros; exact ( fun P:_ => ddualand (inx1 P) (iny1 P)) .
-Defined.
-
-Definition UU' := Type.
-Definition hSet:= sigT' (fun X : UU' => admit) .
-Definition hSetpair := existT' (fun X : UU' => admit).
-Definition pr1hSet:= @projT1' UU (fun X : UU' => admit) : hSet -> Type.
-Coercion pr1hSet: hSet >-> Sortclass.
-
-Definition hPropset : hSet := existT' _ hProp admit .
-
-Definition hsubtypes ( X : UU ) : Type.
-intros; exact (X -> hProp ).
-Defined.
-Definition carrier { X : UU } ( A : hsubtypes X ) : Type.
-intros; exact (sigT' A).
-Defined.
-Coercion carrier : hsubtypes >-> Sortclass.
-
-Definition subtypesdirprod { X Y : UU } ( A : hsubtypes X ) ( B : hsubtypes Y ) : hsubtypes ( dirprod X Y ).
-admit.
-Defined.
-
-Lemma weqsubtypesdirprod { X Y : UU } ( A : hsubtypes X ) ( B : hsubtypes Y ) : weq ( subtypesdirprod A B ) ( dirprod A B ) .
- admit.
-Defined.
-
-Lemma ishinhsubtypesdirprod { X Y : UU } ( A : hsubtypes X ) ( B : hsubtypes Y ) ( isa : ishinh A ) ( isb : ishinh B ) : ishinh ( subtypesdirprod A B ) .
-Proof .
- intros .
- apply ( hinhfun ( invweq ( weqsubtypesdirprod A B ) ) ) .
- apply hinhand .
- apply isa .
- apply isb .
-Defined .
-
-Definition hrel ( X : UU ) : Type.
-intros; exact ( X -> X -> hProp).
-Defined.
-
-Definition iseqrel { X : UU } ( R : hrel X ) : Type.
-admit.
-Defined.
-
-Definition eqrel ( X : UU ) : Type.
-intros; exact ( sigT' ( fun R : hrel X => iseqrel R ) ).
-Defined.
-Definition pr1eqrel ( X : UU ) : eqrel X -> ( X -> ( X -> hProp ) ) := @projT1' _ _ .
-Coercion pr1eqrel : eqrel >-> Funclass .
-
-Definition hreldirprod { X Y : UU } ( RX : hrel X ) ( RY : hrel Y ) : hrel ( dirprod X Y ) .
-admit.
-Defined.
-Set Printing Universes.
-Print hProp.
-Print ishinh_UU.
-Print hProppair.
-Definition iseqclass { X : UU } ( R : hrel X ) ( A : hsubtypes X ) : Type.
-intros; exact ( dirprod ( ishinh ( carrier A ) ) ( dirprod ( forall x1 x2 : X , R x1 x2 -> A x1 -> A x2 ) ( forall x1 x2 : X, A x1 -> A x2 -> R x1 x2 ) )) .
-Defined.
-Definition iseqclassconstr { X : UU } ( R : hrel X ) { A : hsubtypes X } ( ax0 : ishinh ( carrier A ) ) ( ax1 : forall x1 x2 : X , R x1 x2 -> A x1 -> A x2 ) ( ax2 : forall x1 x2 : X, A x1 -> A x2 -> R x1 x2 ) : iseqclass R A.
-intros. hnf. apply dirprodpair. exact ax0. apply dirprodpair. exact ax1. exact ax2.
-Defined.
-
-Definition eqax0 { X : UU } { R : hrel X } { A : hsubtypes X } : iseqclass R A -> ishinh ( carrier A ) .
-intros X R A; exact ( fun is : iseqclass R A => projT1' _ is ).
-Defined.
-
-Lemma iseqclassdirprod { X Y : UU } { R : hrel X } { Q : hrel Y } { A : hsubtypes X } { B : hsubtypes Y } ( isa : iseqclass R A ) ( isb : iseqclass Q B ) : iseqclass ( hreldirprod R Q ) ( subtypesdirprod A B ) .
-Proof .
- intros .
- set ( XY := dirprod X Y ) .
- set ( AB := subtypesdirprod A B ) .
- set ( RQ := hreldirprod R Q ) .
- set ( ax0 := ishinhsubtypesdirprod A B ( eqax0 isa ) admit ) .
- apply ( iseqclassconstr _ ax0 admit admit ) .
-Defined .
-
-Definition image { X Y : UU } ( f : X -> Y ) : Type.
-intros; exact ( sigT' ( fun y : Y => admit ) ).
-Defined.
-Definition pr1image { X Y : UU } ( f : X -> Y ) : image f -> Y.
-intros X Y f; exact ( @projT1' _ ( fun y : Y => admit ) ).
-Defined.
-
-Definition prtoimage { X Y : UU } (f : X -> Y) : X -> image f.
- admit.
-Defined.
-
-Definition setquot { X : UU } ( R : hrel X ) : Type.
-intros; exact ( sigT' ( fun A : _ => iseqclass R A ) ).
-Defined.
-Definition setquotpair { X : UU } ( R : hrel X ) ( A : hsubtypes X ) ( is : iseqclass R A ) : setquot R.
-intros; exact (existT' _ A is ).
-Defined.
-Definition pr1setquot { X : UU } ( R : hrel X ) : setquot R -> ( hsubtypes X ).
-intros X R.
-exact ( @projT1' _ ( fun A : _ => iseqclass R A ) ).
-Defined.
-Coercion pr1setquot : setquot >-> hsubtypes .
-
-Definition setquotinset { X : UU } ( R : hrel X ) : hSet.
-intros; exact ( hSetpair (setquot R) admit) .
-Defined.
-
-Definition dirprodtosetquot { X Y : UU } ( RX : hrel X ) ( RY : hrel Y ) (cd : dirprod ( setquot RX ) ( setquot RY ) ) : setquot ( hreldirprod RX RY ).
-intros; exact ( setquotpair _ _ ( iseqclassdirprod ( projT2' _ ( projT1' _ cd ) ) ( projT2' _ ( projT2' _ cd ) ) ) ).
-Defined.
-
-Definition iscomprelfun2 { X Y : UU } ( R : hrel X ) ( f : X -> X -> Y ) := forall x x' x0 x0' : X , R x x' -> R x0 x0' -> paths ( f x x0 ) ( f x' x0' ) .
-
-Definition binop ( X : UU ) : Type.
-intros; exact ( X -> X -> X ).
-Defined.
-
-Definition setwithbinop : Type.
-exact (sigT' ( fun X : hSet => binop X ) ).
-Defined.
-Definition pr1setwithbinop : setwithbinop -> hSet@{j k Set l}.
-unfold setwithbinop.
-exact ( @projT1' _ ( fun X : hSet@{j k Set l} => binop@{Set} X ) ).
-Defined.
-Coercion pr1setwithbinop : setwithbinop >-> hSet .
-
-Definition op { X : setwithbinop } : binop X.
-intros; exact ( projT2' _ X ).
-Defined.
-
-Definition subsetswithbinop { X : setwithbinop } : Type.
-admit.
-Defined.
-
-Definition carrierofasubsetwithbinop { X : setwithbinop } ( A : @subsetswithbinop X ) : setwithbinop .
-admit.
-Defined.
-
-Coercion carrierofasubsetwithbinop : subsetswithbinop >-> setwithbinop .
-
-Definition binopeqrel { X : setwithbinop } : Type.
-intros; exact (sigT' ( fun R : eqrel X => admit ) ).
-Defined.
-Definition binopeqrelpair { X : setwithbinop } := existT' ( fun R : eqrel X => admit ).
-Definition pr1binopeqrel ( X : setwithbinop ) : @binopeqrel X -> eqrel X.
-intros X; exact ( @projT1' _ ( fun R : eqrel X => admit ) ) .
-Defined.
-Coercion pr1binopeqrel : binopeqrel >-> eqrel .
-
-Definition setwithbinopdirprod ( X Y : setwithbinop ) : setwithbinop .
-admit.
-Defined.
-
-Definition monoid : Type.
-exact ( sigT' ( fun X : setwithbinop => admit ) ).
-Defined.
-Definition monoidpair := existT' ( fun X : setwithbinop => admit ) .
-Definition pr1monoid : monoid -> setwithbinop := @projT1' _ _ .
-Coercion pr1monoid : monoid >-> setwithbinop .
-
-Notation "x + y" := ( op x y ) : addmonoid_scope .
-
-Definition submonoids { X : monoid } : Type.
-admit.
-Defined.
-
-Definition submonoidstosubsetswithbinop ( X : monoid ) : @submonoids X -> @subsetswithbinop X.
-admit.
-Defined.
-Coercion submonoidstosubsetswithbinop : submonoids >-> subsetswithbinop .
-
-Definition abmonoid : Type.
-exact (sigT' ( fun X : setwithbinop => admit ) ).
-Defined.
-
-Definition abmonoidtomonoid : abmonoid -> monoid.
-exact (fun X : _ => monoidpair ( projT1' _ X ) admit ).
-Defined.
-Coercion abmonoidtomonoid : abmonoid >-> monoid .
-
-Definition subabmonoids { X : abmonoid } := @submonoids X .
-
-Definition carrierofsubabmonoid { X : abmonoid } ( A : @subabmonoids X ) : abmonoid .
-Proof .
- intros .
- unfold subabmonoids in A .
- split with A .
- admit.
-Defined .
-
-Coercion carrierofsubabmonoid : subabmonoids >-> abmonoid .
-
-Definition abmonoiddirprod ( X Y : abmonoid ) : abmonoid .
-Proof .
- intros .
- split with ( setwithbinopdirprod X Y ) .
- admit.
-Defined .
-
-Open Scope addmonoid_scope .
-
-Definition eqrelabmonoidfrac ( X : abmonoid ) ( A : @submonoids X ) : eqrel ( setwithbinopdirprod X A ).
-admit.
-Defined.
-
-Definition binopeqrelabmonoidfrac ( X : abmonoid ) ( A : @subabmonoids X ) : @binopeqrel ( abmonoiddirprod X A ).
-intros; exact ( @binopeqrelpair ( setwithbinopdirprod X A ) ( eqrelabmonoidfrac X A ) admit ).
-Defined.
-
-Theorem setquotuniv { X : UU } ( R : hrel X ) ( Y : hSet ) ( f : X -> Y ) ( c : setquot R ) : Y .
-Proof.
- intros.
- apply ( pr1image ( fun x : c => f ( projT1' _ x ) ) ) .
- apply ( @hinhuniv ( projT1' _ c ) ( hProppair _ admit ) ( prtoimage ( fun x : c => f ( projT1' _ x ) ) ) ) .
- pose ( eqax0 ( projT2' _ c ) ) as h.
- simpl in *.
- Set Printing Universes.
- exact h.
-Defined .
-
-Definition setquotuniv2 { X : UU } ( R : hrel X ) ( Y : hSet ) ( f : X -> X -> Y ) ( is : iscomprelfun2 R f ) ( c c0 : setquot R ) : Y .
-Proof.
- intros .
- set ( RR := hreldirprod R R ) .
- apply (setquotuniv RR Y admit).
- apply dirprodtosetquot.
- apply dirprodpair.
- exact c.
- exact c0.
-Defined .
-
-Definition setquotfun2 { X Y : UU } ( RX : hrel X ) ( RY : eqrel Y ) ( f : X -> X -> Y ) ( cx cx0 : setquot RX ) : setquot RY .
-Proof .
- intros .
- apply ( setquotuniv2 RX ( setquotinset RY ) admit admit admit admit ) .
-Defined .
-
-Definition quotrel { X : UU } { R : hrel X } : hrel ( setquot R ).
-intros; exact ( setquotuniv2 R hPropset admit admit ).
-Defined.
-
-Definition setwithbinopquot { X : setwithbinop } ( R : @binopeqrel X ) : setwithbinop .
-Proof .
- intros .
- split with ( setquotinset R ) .
- set ( qtmlt := setquotfun2 R R op ) .
- simpl .
- unfold binop .
- apply qtmlt .
-Defined .
-
-Definition abmonoidquot { X : abmonoid } ( R : @binopeqrel X ) : abmonoid .
-Proof .
- intros .
- split with ( setwithbinopquot R ) .
- admit.
-Defined .
-
-Definition abmonoidfrac ( X : abmonoid ) ( A : @submonoids X ) : abmonoid.
-intros; exact ( @abmonoidquot (abmonoiddirprod X (@carrierofsubabmonoid X A)) ( binopeqrelabmonoidfrac X A ) ).
-Defined.
-
-Definition abmonoidfracrel ( X : abmonoid ) ( A : @submonoids X ) : hrel (@setquot (setwithbinopdirprod X A) (eqrelabmonoidfrac X A)).
-intros; exact (@quotrel _ _).
-Defined.
-
-Fail Timeout 1 Axiom ispartlbinopabmonoidfracrel : forall ( X : abmonoid ) ( A : @subabmonoids X ) { L : hrel X } ( z : abmonoidfrac X A ) , @abmonoidfracrel X A ( ( admit + z ) )admit.
-
-Definition ispartlbinopabmonoidfracrel_type : Type :=
- forall ( X : abmonoid ) ( A : @subabmonoids X ) { L : hrel X } ( z : abmonoidfrac X A ),
- @abmonoidfracrel X A ( ( admit + z ) )admit.
-
-Fail Timeout 1 Axiom ispartlbinopabmonoidfracrel' : $(let t:= eval unfold ispartlbinopabmonoidfracrel_type in
- ispartlbinopabmonoidfracrel_type in exact t)$.
-
-Unset Kernel Term Sharing.
-
-Axiom ispartlbinopabmonoidfracrel : forall ( X : abmonoid ) ( A : @subabmonoids X ) { L : hrel X } ( z : abmonoidfrac X A ) , @abmonoidfracrel X A ( ( admit + z ) )admit.
-
-Axiom ispartlbinopabmonoidfracrel' : $(let t:= eval unfold ispartlbinopabmonoidfracrel_type in
- ispartlbinopabmonoidfracrel_type in exact t)$.
-
diff --git a/test-suite/bugs/closed/3314.v b/test-suite/bugs/closed/3314.v
index e63c46da0..fb3791af5 100644
--- a/test-suite/bugs/closed/3314.v
+++ b/test-suite/bugs/closed/3314.v
@@ -122,12 +122,12 @@ Definition depsort (T : Type) (x : bool) : informative x :=
end.
(** This definition should fail *)
-Definition Box (T : Type1) : Prop := Lift T.
+Fail Definition Box (T : Type1) : Prop := Lift T.
-Definition prop {T : Type1} (t : Box T) : T := t.
-Definition wrap {T : Type1} (t : T) : Box T := t.
+Fail Definition prop {T : Type1} (t : Box T) : T := t.
+Fail Definition wrap {T : Type1} (t : T) : Box T := t.
-Definition down (x : Type1) : Prop := Box x.
+Fail Definition down (x : Type1) : Prop := Box x.
Definition up (x : Prop) : Type1 := x.
Fail Definition back A : up (down A) -> A := @prop A.
diff --git a/test-suite/bugs/opened/3685.v b/test-suite/bugs/closed/3685.v
index b2b5db6be..a5bea34a9 100644
--- a/test-suite/bugs/opened/3685.v
+++ b/test-suite/bugs/closed/3685.v
@@ -63,7 +63,7 @@ Module Success.
End Success.
Module Bad.
Include PointwiseCore.
- Fail Definition functor_uncurried `{Funext} (P : PreCategory -> Type)
+ Definition functor_uncurried `{Funext} (P : PreCategory -> Type)
(has_functor_categories : forall C D : sub_pre_cat P, P (C -> D))
: object (((sub_pre_cat P)^op * (sub_pre_cat P)) -> (sub_pre_cat P))
:= Eval cbv zeta in
diff --git a/test-suite/bugs/closed/3777.v b/test-suite/bugs/closed/3777.v
new file mode 100644
index 000000000..b9b2dd6b3
--- /dev/null
+++ b/test-suite/bugs/closed/3777.v
@@ -0,0 +1,16 @@
+Module WithoutPoly.
+ Unset Universe Polymorphism.
+ Definition foo (A : Type@{i}) (B : Type@{i}) := A -> B.
+ Set Printing Universes.
+ Definition bla := ((@foo : Set -> _ -> _) : _ -> Type -> _).
+ (* ((fun A : Set => foo A):Set -> Type@{Top.55} -> Type@{Top.55})
+:Set -> Type@{Top.55} -> Type@{Top.55}
+ : Set -> Type@{Top.55} -> Type@{Top.55}
+(* |= Set <= Top.55
+ *) *)
+End WithoutPoly.
+Module WithPoly.
+ Set Universe Polymorphism.
+ Definition foo (A : Type@{i}) (B : Type@{i}) := A -> B.
+ Set Printing Universes.
+ Fail Check ((@foo : Set -> _ -> _) : _ -> Type -> _).
diff --git a/test-suite/bugs/closed/4251.v b/test-suite/bugs/closed/4251.v
new file mode 100644
index 000000000..66343d667
--- /dev/null
+++ b/test-suite/bugs/closed/4251.v
@@ -0,0 +1,17 @@
+
+Inductive array : Type -> Type :=
+| carray : forall A, array A.
+
+Inductive Mtac : Type -> Prop :=
+| bind : forall {A B}, Mtac A -> (A -> Mtac B) -> Mtac B
+| array_make : forall {A}, A -> Mtac (array A).
+
+Definition Ref := array.
+
+Definition ref : forall {A}, A -> Mtac (Ref A) :=
+ fun A x=> array_make x.
+Check array Type.
+Check fun A : Type => Ref A.
+
+Definition abs_val (a : Type) :=
+ bind (ref a) (fun r : array Type => array_make tt). \ No newline at end of file
diff --git a/test-suite/bugs/closed/4287.v b/test-suite/bugs/closed/4287.v
new file mode 100644
index 000000000..e139c5b6c
--- /dev/null
+++ b/test-suite/bugs/closed/4287.v
@@ -0,0 +1,124 @@
+
+Universe b.
+
+Universe c.
+
+Definition U : Type@{b} := Type@{c}.
+
+Module Type MT.
+
+Definition T := Prop.
+End MT.
+
+Module M : MT.
+ Definition T := Type@{b}.
+
+Print Universes.
+Fail End M.
+
+Set Universe Polymorphism.
+
+(* This is a modified version of Hurkens with all universes floating *)
+Section Hurkens.
+
+Variable down : Type -> Type.
+Variable up : Type -> Type.
+
+Hypothesis back : forall A, up (down A) -> A.
+
+Hypothesis forth : forall A, A -> up (down A).
+
+Hypothesis backforth : forall (A:Type) (P:A->Type) (a:A),
+ P (back A (forth A a)) -> P a.
+
+Hypothesis backforth_r : forall (A:Type) (P:A->Type) (a:A),
+ P a -> P (back A (forth A a)).
+
+(** Proof *)
+Definition V : Type := forall A:Type, ((up A -> Type) -> up A -> Type) -> up A -> Type.
+Definition U : Type := V -> Type.
+
+Definition sb (z:V) : V := fun A r a => r (z A r) a.
+Definition le (i:U -> Type) (x:U) : Type := x (fun A r a => i (fun v => sb v A r a)).
+Definition le' (i:up (down U) -> Type) (x:up (down U)) : Type := le (fun a:U => i (forth _ a)) (back _ x).
+Definition induct (i:U -> Type) : Type := forall x:U, up (le i x) -> up (i x).
+Definition WF : U := fun z => down (induct (fun a => z (down U) le' (forth _ a))).
+Definition I (x:U) : Type :=
+ (forall i:U -> Type, up (le i x) -> up (i (fun v => sb v (down U) le' (forth _ x)))) -> False.
+
+Lemma Omega : forall i:U -> Type, induct i -> up (i WF).
+Proof.
+intros i y.
+apply y.
+unfold le, WF, induct.
+apply forth.
+intros x H0.
+apply y.
+unfold sb, le', le.
+compute.
+apply backforth_r.
+exact H0.
+Qed.
+
+Lemma lemma1 : induct (fun u => down (I u)).
+Proof.
+unfold induct.
+intros x p.
+apply forth.
+intro q.
+generalize (q (fun u => down (I u)) p).
+intro r.
+apply back in r.
+apply r.
+intros i j.
+unfold le, sb, le', le in j |-.
+apply backforth in j.
+specialize q with (i := fun y => i (fun v:V => sb v (down U) le' (forth _ y))).
+apply q.
+exact j.
+Qed.
+
+Lemma lemma2 : (forall i:U -> Type, induct i -> up (i WF)) -> False.
+Proof.
+intro x.
+generalize (x (fun u => down (I u)) lemma1).
+intro r; apply back in r.
+apply r.
+intros i H0.
+apply (x (fun y => i (fun v => sb v (down U) le' (forth _ y)))).
+unfold le, WF in H0.
+apply back in H0.
+exact H0.
+Qed.
+
+Theorem paradox : False.
+Proof.
+exact (lemma2 Omega).
+Qed.
+
+End Hurkens.
+
+Polymorphic Record box (T : Type) := wrap {unwrap : T}.
+
+(* Here we instantiate to Set *)
+
+Fail Definition down (x : Type) : Prop := box x.
+Definition up (x : Prop) : Type := x.
+
+Fail Definition back A : up (down A) -> A := unwrap A.
+
+Fail Definition forth A : A -> up (down A) := wrap A.
+
+Definition id {A : Type} (a : A) := a.
+Definition setlt (A : Type@{i}) :=
+ let foo := Type@{i} : Type@{j} in True.
+
+Definition setle (B : Type@{i}) :=
+ let foo (A : Type@{j}) := A in foo B.
+
+Fail Check @setlt@{j Prop}.
+Check @setlt@{Prop j}.
+Check @setle@{Prop j}.
+
+Fail Definition foo := @setle@{j Prop}.
+Definition foo := @setle@{Prop j}.
diff --git a/test-suite/bugs/closed/4294.v b/test-suite/bugs/closed/4294.v
new file mode 100644
index 000000000..1d5e3c71b
--- /dev/null
+++ b/test-suite/bugs/closed/4294.v
@@ -0,0 +1,31 @@
+Require Import Hurkens.
+
+Module NonPoly.
+Module Type Foo.
+ Definition U := Type.
+ Parameter eq : Type = U.
+End Foo.
+
+Module M : Foo with Definition U := Type.
+ Definition U := Type.
+ Definition eq : Type = U := eq_refl.
+End M.
+
+Print Universes.
+Fail Definition bad : False := TypeNeqSmallType.paradox M.U M.eq.
+End NonPoly.
+
+Set Universe Polymorphism.
+
+Module Type Foo.
+ Definition U := Type.
+ Monomorphic Parameter eq : Type = U.
+End Foo.
+
+Module M : Foo with Definition U := Type.
+ Definition U := Type.
+ Monomorphic Definition eq : Type = U := eq_refl.
+End M.
+
+Fail Definition bad : False := TypeNeqSmallType.paradox Type M.eq.
+(* Print Assumptions bad. *)
diff --git a/test-suite/bugs/closed/4298.v b/test-suite/bugs/closed/4298.v
new file mode 100644
index 000000000..875612ddf
--- /dev/null
+++ b/test-suite/bugs/closed/4298.v
@@ -0,0 +1,7 @@
+Set Universe Polymorphism.
+
+Module Type Foo.
+ Definition U := Type.
+End Foo.
+
+Fail Module M : Foo with Definition U := Prop.
diff --git a/test-suite/bugs/closed/4301.v b/test-suite/bugs/closed/4301.v
new file mode 100644
index 000000000..3b00efb21
--- /dev/null
+++ b/test-suite/bugs/closed/4301.v
@@ -0,0 +1,12 @@
+Set Universe Polymorphism.
+
+Module Type Foo.
+ Parameter U : Type.
+End Foo.
+
+Module Lower (X : Foo with Definition U := True : Type).
+End Lower.
+
+Module M : Foo.
+ Definition U := nat : Type@{i}.
+End M.
diff --git a/test-suite/bugs/closed/4328.v b/test-suite/bugs/closed/4328.v
new file mode 100644
index 000000000..8e1bb3100
--- /dev/null
+++ b/test-suite/bugs/closed/4328.v
@@ -0,0 +1,6 @@
+Inductive M (A:Type) : Type := M'.
+Axiom pi : forall (P : Prop) (p : P), Prop.
+Definition test1 A (x : _) := pi A x. (* success *)
+Fail Definition test2 A (x : A) := pi A x. (* failure ??? *)
+Fail Definition test3 A (x : A) (_ : M A) := pi A x. (* failure *)
+Fail Definition test4 A (_ : M A) (x : A) := pi A x. (* success ??? *) \ No newline at end of file
diff --git a/test-suite/bugs/closed/HoTT_coq_053.v b/test-suite/bugs/closed/HoTT_coq_053.v
index a14fb6aa5..e2bf1dbed 100644
--- a/test-suite/bugs/closed/HoTT_coq_053.v
+++ b/test-suite/bugs/closed/HoTT_coq_053.v
@@ -39,7 +39,7 @@ Definition NatCategory (n : nat) :=
Definition NatCategory' (n : nat) :=
match n with
| 0 => (fun X => @Build_PreCategory X
- (fun _ _ => Unit : Prop)) Unit
+ (fun _ _ => Unit : Set)) Unit
| _ => DiscreteCategory Bool
end.
diff --git a/test-suite/bugs/closed/HoTT_coq_093.v b/test-suite/bugs/closed/HoTT_coq_093.v
index 38943ab35..f382dac97 100644
--- a/test-suite/bugs/closed/HoTT_coq_093.v
+++ b/test-suite/bugs/closed/HoTT_coq_093.v
@@ -21,7 +21,7 @@ Section lift.
Definition Lift (A : Type@{i}) : Type@{j} := A.
End lift.
-Goal forall (A : Type@{i}) (x y : A), @paths@{i} A x y -> @paths@{j} A x y.
+Goal forall (A : Type@{i}) (x y : A), @paths@{i j} A x y -> @paths@{j k} A x y.
intros A x y p.
compute in *. destruct p. exact idpath.
Defined.
diff --git a/test-suite/bugs/closed/HoTT_coq_108.v b/test-suite/bugs/closed/HoTT_coq_108.v
index 4f5ef9974..b6c0da76b 100644
--- a/test-suite/bugs/closed/HoTT_coq_108.v
+++ b/test-suite/bugs/closed/HoTT_coq_108.v
@@ -107,7 +107,7 @@ Section path_functor.
Variable D : PreCategory.
Local Notation path_functor'_T F G
:= { HO : object_of F = object_of G
- | transport (fun GO => forall s d, morphism C s d -> morphism D (GO s) (GO d))
+ & transport (fun GO => forall s d, morphism C s d -> morphism D (GO s) (GO d))
HO
(morphism_of F)
= morphism_of G }
diff --git a/test-suite/bugs/opened/HoTT_coq_120.v b/test-suite/bugs/closed/HoTT_coq_120.v
index 05ee6c7b6..e46ea58bb 100644
--- a/test-suite/bugs/opened/HoTT_coq_120.v
+++ b/test-suite/bugs/closed/HoTT_coq_120.v
@@ -116,7 +116,8 @@ Section fully_faithful_helpers.
Variables x y : hSet.
Variable m : x -> y.
- Let isequiv_isepi_ismono_helper ua := (@isequiv_isepi_ismono ua fs0 x y m : isepi m -> ismono m -> IsEquiv m).
+ Fail Let isequiv_isepi_ismono_helper ua :=
+ (@isequiv_isepi_ismono ua fs0 x y m : isepi m -> ismono m -> IsEquiv m).
Goal True.
Fail set (isequiv_isepimorphism_ismonomorphism
@@ -126,7 +127,7 @@ Section fully_faithful_helpers.
=> (@isequiv_isepi_ismono_helper _ Hepi Hmono : @IsEquiv _ _ m)).
admit.
Undo.
- Fail set (isequiv_isepimorphism_ismonomorphism'
+ Fail set (isequiv_isepimorphism_ismonomorphism
:= fun `{Univalence}
(Hepi : IsEpimorphism (m : morphism set_cat x y))
(Hmono : IsMonomorphism (m : morphism set_cat x y))
diff --git a/test-suite/failure/guard-cofix.v b/test-suite/failure/guard-cofix.v
index 64faa0ce0..eda4a1867 100644
--- a/test-suite/failure/guard-cofix.v
+++ b/test-suite/failure/guard-cofix.v
@@ -25,7 +25,7 @@ Fail Definition ff : False := match loop with CF _ t => t end.
(* Second example *)
-Inductive omega := Omega : omega -> omega.
+Inductive omega : Prop := Omega : omega -> omega.
Lemma H : omega = CoFalse.
Proof.
diff --git a/theories/Classes/CMorphisms.v b/theories/Classes/CMorphisms.v
index 048faa916..9d3952e64 100644
--- a/theories/Classes/CMorphisms.v
+++ b/theories/Classes/CMorphisms.v
@@ -267,10 +267,9 @@ Section GenericInstances.
Qed.
(** The complement of a crelation conserves its proper elements. *)
-
Program Definition complement_proper
`(mR : Proper (A -> A -> Prop) (RA ==> RA ==> iff) R) :
- Proper (RA ==> RA ==> iff) (complement R) := _.
+ Proper (RA ==> RA ==> iff) (complement@{i j Prop} R) := _.
Next Obligation.
Proof.
diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml
index 4122487e2..8ac273c84 100644
--- a/toplevel/auto_ind_decl.ml
+++ b/toplevel/auto_ind_decl.ml
@@ -304,7 +304,7 @@ let build_beq_scheme mode kn =
raise (NonSingletonProp (kn,i));
let fix = mkFix (((Array.make nb_ind 0),i),(names,types,cores)) in
create_input fix),
- Evd.empty_evar_universe_context (* FIXME *)),
+ Evd.make_evar_universe_context (Global.env ()) None),
!eff
let beq_scheme_kind = declare_mutual_scheme_object "_beq" build_beq_scheme
@@ -641,7 +641,7 @@ let make_bl_scheme mode mind =
let lnonparrec,lnamesparrec = (* TODO subst *)
context_chop (nparams-nparrec) mib.mind_params_ctxt in
let bl_goal, eff = compute_bl_goal ind lnamesparrec nparrec in
- let ctx = Evd.empty_evar_universe_context (*FIXME univs *) in
+ let ctx = Evd.make_evar_universe_context (Global.env ()) None in
let side_eff = side_effect_of_mode mode in
let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx bl_goal
(compute_bl_tact mode (!bl_scheme_kind_aux()) (ind, Univ.Instance.empty) lnamesparrec nparrec)
@@ -764,12 +764,12 @@ let make_lb_scheme mode mind =
let lnonparrec,lnamesparrec =
context_chop (nparams-nparrec) mib.mind_params_ctxt in
let lb_goal, eff = compute_lb_goal ind lnamesparrec nparrec in
- let ctx = Evd.empty_evar_universe_context in
+ let ctx = Evd.make_evar_universe_context (Global.env ()) None in
let side_eff = side_effect_of_mode mode in
let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx lb_goal
(compute_lb_tact mode (!lb_scheme_kind_aux()) ind lnamesparrec nparrec)
in
- ([|ans|], ctx (* FIXME *)), eff
+ ([|ans|], ctx), eff
let lb_scheme_kind = declare_mutual_scheme_object "_dec_lb" make_lb_scheme
@@ -934,7 +934,7 @@ let make_eq_decidability mode mind =
let nparams = mib.mind_nparams in
let nparrec = mib.mind_nparams_rec in
let u = Univ.Instance.empty in
- let ctx = Evd.empty_evar_universe_context (* FIXME *)in
+ let ctx = Evd.make_evar_universe_context (Global.env ()) None in
let lnonparrec,lnamesparrec =
context_chop (nparams-nparrec) mib.mind_params_ctxt in
let side_eff = side_effect_of_mode mode in
diff --git a/toplevel/classes.ml b/toplevel/classes.ml
index 7fe79d948..805a29e39 100644
--- a/toplevel/classes.ml
+++ b/toplevel/classes.ml
@@ -347,7 +347,7 @@ let named_of_rel_context l =
let context poly l =
let env = Global.env() in
- let evars = ref Evd.empty in
+ let evars = ref (Evd.from_env env) in
let _, ((env', fullctx), impls) = interp_context_evars env evars l in
let subst = Evarutil.evd_comb0 Evarutil.nf_evars_and_universes evars in
let fullctx = Context.map_rel_context subst fullctx in
@@ -358,11 +358,13 @@ let context poly l =
with e when Errors.noncritical e ->
error "Anonymous variables not allowed in contexts."
in
- let uctx = Evd.universe_context_set !evars in
+ let uctx = ref (Evd.universe_context_set !evars) in
let fn status (id, b, t) =
if Lib.is_modtype () && not (Lib.sections_are_opened ()) then
- let uctx = Univ.ContextSet.to_context uctx in
- let decl = (ParameterEntry (None,poly,(t,uctx),None), IsAssumption Logical) in
+ let ctx = Univ.ContextSet.to_context !uctx in
+ (* Declare the universe context once *)
+ let () = uctx := Univ.ContextSet.empty in
+ let decl = (ParameterEntry (None,poly,(t,ctx),None), IsAssumption Logical) in
let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id decl in
match class_of_constr t with
| Some (rels, ((tc,_), args) as _cl) ->
@@ -379,8 +381,9 @@ let context poly l =
let impl = List.exists test impls in
let decl = (Discharge, poly, Definitional) in
let nstatus =
- pi3 (Command.declare_assumption false decl (t, uctx) [] impl
+ pi3 (Command.declare_assumption false decl (t, !uctx) [] impl
Vernacexpr.NoInline (Loc.ghost, id))
in
- status && nstatus
+ let () = uctx := Univ.ContextSet.empty in
+ status && nstatus
in List.fold_left fn true (List.rev ctx)
diff --git a/toplevel/command.ml b/toplevel/command.ml
index e2e5d8704..b65ff73fe 100644
--- a/toplevel/command.ml
+++ b/toplevel/command.ml
@@ -241,11 +241,14 @@ let interp_assumption evdref env impls bl c =
let ctx = Evd.universe_context_set evd in
((nf ty, ctx), impls)
-let declare_assumptions idl is_coe k c imps impl_is_on nl =
- let refs, status =
- List.fold_left (fun (refs,status) id ->
- let ref',u',status' = declare_assumption is_coe k c imps impl_is_on nl id in
- (ref',u')::refs, status' && status) ([],true) idl in
+let declare_assumptions idl is_coe k (c,ctx) imps impl_is_on nl =
+ let refs, status, _ =
+ List.fold_left (fun (refs,status,ctx) id ->
+ let ref',u',status' =
+ declare_assumption is_coe k (c,ctx) imps impl_is_on nl id in
+ (ref',u')::refs, status' && status, Univ.ContextSet.empty)
+ ([],true,ctx) idl
+ in
List.rev refs, status
let do_assumptions (_, poly, _ as kind) nl l =
@@ -467,16 +470,17 @@ let inductive_levels env evdref poly arities inds =
Evd.set_leq_sort env evd (Prop Pos) du
else evd
in
- (* let arity = it_mkProd_or_LetIn (mkType cu) ctx in *\) *)
- let duu = Sorts.univ_of_sort du in
- let evd =
- if not (Univ.is_small_univ duu) && Evd.check_eq evd cu duu then
- if is_flexible_sort evd duu then
- Evd.set_eq_sort env evd (Prop Null) du
- else evd
- else Evd.set_eq_sort env evd (Type cu) du
- in
- (evd, arity :: arities))
+ let duu = Sorts.univ_of_sort du in
+ let evd =
+ if not (Univ.is_small_univ duu) && Evd.check_eq evd cu duu then
+ if is_flexible_sort evd duu then
+ if Evd.check_leq evd Univ.type0_univ duu then
+ evd
+ else Evd.set_eq_sort env evd (Prop Null) du
+ else evd
+ else Evd.set_eq_sort env evd (Type cu) du
+ in
+ (evd, arity :: arities))
(!evdref,[]) (Array.to_list levels') destarities sizes
in evdref := evd; List.rev arities
diff --git a/toplevel/indschemes.ml b/toplevel/indschemes.ml
index 452d5fbe5..ae8ee7670 100644
--- a/toplevel/indschemes.ml
+++ b/toplevel/indschemes.ml
@@ -423,7 +423,7 @@ let fold_left' f = function
let build_combined_scheme env schemes =
let defs = List.map (fun cst -> (* FIXME *)
- let evd, c = Evd.fresh_constant_instance env Evd.empty cst in
+ let evd, c = Evd.fresh_constant_instance env (Evd.from_env env) cst in
(c, Typeops.type_of_constant_in env c)) schemes in
(* let nschemes = List.length schemes in *)
let find_inductive ty =
diff --git a/toplevel/obligations.ml b/toplevel/obligations.ml
index 0b15ecf33..d2bd2c07b 100644
--- a/toplevel/obligations.ml
+++ b/toplevel/obligations.ml
@@ -625,7 +625,7 @@ let declare_obligation prg obl body ty uctx =
let body = prg.prg_reduce body in
let ty = Option.map prg.prg_reduce ty in
match obl.obl_status with
- | Evar_kinds.Expand -> { obl with obl_body = Some (TermObl body) }
+ | Evar_kinds.Expand -> false, { obl with obl_body = Some (TermObl body) }
| Evar_kinds.Define opaque ->
let opaque = if get_proofs_transparency () then false else opaque in
let poly = pi2 prg.prg_kind in
@@ -649,7 +649,7 @@ let declare_obligation prg obl body ty uctx =
in
if not opaque then add_hint false prg constant;
definition_message obl.obl_name;
- { obl with obl_body =
+ true, { obl with obl_body =
if poly then
Some (DefinedObl constant)
else
@@ -798,9 +798,9 @@ let solve_by_tac name evi t poly ctx =
let entry = Term_typing.handle_entry_side_effects env entry in
let body, eff = Future.force entry.Entries.const_entry_body in
assert(Declareops.side_effects_is_empty eff);
- assert(Univ.ContextSet.is_empty (snd body));
+ let ctx' = Evd.merge_context_set Evd.univ_rigid (Evd.from_ctx ctx') (snd body) in
Inductiveops.control_only_guard (Global.env ()) (fst body);
- (fst body), entry.Entries.const_entry_type, ctx'
+ (fst body), entry.Entries.const_entry_type, Evd.evar_universe_context ctx'
let obligation_terminator name num guard hook pf =
let open Proof_global in
@@ -824,7 +824,7 @@ let obligation_terminator name num guard hook pf =
let obls, rem = prg.prg_obligations in
let obl = obls.(num) in
let ctx = Evd.evar_context_universe_context uctx in
- let obl = declare_obligation prg obl body ty ctx in
+ let (_, obl) = declare_obligation prg obl body ty ctx in
let obls = Array.copy obls in
let _ = obls.(num) <- obl in
try ignore (update_obls prg obls (pred rem))
@@ -847,9 +847,9 @@ let obligation_hook prg obl num auto ctx' _ gr =
let ctx' = match ctx' with None -> prg.prg_ctx | Some ctx' -> ctx' in
let ctx' =
if not (pi2 prg.prg_kind) (* Not polymorphic *) then
- (* This context is already declared globally, we cannot
- instantiate the rigid variables anymore *)
- Evd.abstract_undefined_variables ctx'
+ (* The universe context was declared globally, we continue
+ from the new global environment. *)
+ Evd.evar_universe_context (Evd.from_env (Global.env ()))
else ctx'
in
let prg = { prg with prg_ctx = ctx' } in
@@ -922,8 +922,13 @@ and solve_obligation_by_tac prg obls i tac =
(pi2 !prg.prg_kind) !prg.prg_ctx
in
let uctx = Evd.evar_context_universe_context ctx in
- prg := {!prg with prg_ctx = ctx};
- obls.(i) <- declare_obligation !prg obl t ty uctx;
+ let () = prg := {!prg with prg_ctx = ctx} in
+ let def, obl' = declare_obligation !prg obl t ty uctx in
+ obls.(i) <- obl';
+ if def && not (pi2 !prg.prg_kind) then (
+ (* Declare the term constraints with the first obligation only *)
+ let ctx' = Evd.evar_universe_context (Evd.from_env (Global.env ())) in
+ prg := {!prg with prg_ctx = ctx'});
true
else false
with e when Errors.noncritical e ->
diff --git a/toplevel/record.ml b/toplevel/record.ml
index e214f9ca7..60fe76bb8 100644
--- a/toplevel/record.ml
+++ b/toplevel/record.ml
@@ -131,14 +131,21 @@ let typecheck_params_and_fields def id pl t ps nots fs =
Pretyping.solve_remaining_evars Pretyping.all_and_fail_flags env_ar !evars (Evd.empty,!evars) in
let evars, nf = Evarutil.nf_evars_and_universes sigma in
let arity = nf t' in
- let evars =
+ let arity, evars =
let _, univ = compute_constructor_level evars env_ar newfs in
let ctx, aritysort = Reduction.dest_arity env0 arity in
assert(List.is_empty ctx); (* Ensured by above analysis *)
if Sorts.is_prop aritysort ||
(Sorts.is_set aritysort && is_impredicative_set env0) then
- evars
- else Evd.set_leq_sort env_ar evars (Type univ) aritysort
+ arity, evars
+ else
+ let evars = Evd.set_leq_sort env_ar evars (Type univ) aritysort in
+ if Univ.is_small_univ univ then
+ (* We can assume that the level aritysort is not constrained
+ and clear it. *)
+ mkArity (ctx, Sorts.sort_of_univ univ),
+ Evd.set_eq_sort env_ar evars (Prop Pos) aritysort
+ else arity, evars
in
let evars, nf = Evarutil.nf_evars_and_universes evars in
let newps = map_rel_context nf newps in
@@ -233,7 +240,8 @@ let declare_projections indsp ?(kind=StructureComponent) binder_name coers field
let (mib,mip) = Global.lookup_inductive indsp in
let u = Declareops.inductive_instance mib in
let paramdecls = Inductive.inductive_paramdecls (mib, u) in
- let poly = mib.mind_polymorphic and ctx = Univ.instantiate_univ_context mib.mind_universes in
+ let poly = mib.mind_polymorphic in
+ let ctx = Univ.instantiate_univ_context mib.mind_universes in
let indu = indsp, u in
let r = mkIndU (indsp,u) in
let rp = applist (r, Termops.extended_rel_list 0 paramdecls) in
@@ -293,7 +301,8 @@ let declare_projections indsp ?(kind=StructureComponent) binder_name coers field
const_entry_secctx = None;
const_entry_type = Some projtyp;
const_entry_polymorphic = poly;
- const_entry_universes = ctx;
+ const_entry_universes =
+ if poly then ctx else Univ.UContext.empty;
const_entry_opaque = false;
const_entry_inline_code = false;
const_entry_feedback = None } in
@@ -397,44 +406,49 @@ let declare_class finite def poly ctx id idbuild paramimpls params arity
let impl, projs =
match fields with
| [(Name proj_name, _, field)] when def ->
- let class_body = it_mkLambda_or_LetIn field params in
- let _class_type = it_mkProd_or_LetIn arity params in
- let class_entry =
- Declare.definition_entry (* ?types:class_type *) ~poly ~univs:ctx class_body in
- let cst = Declare.declare_constant (snd id)
- (DefinitionEntry class_entry, IsDefinition Definition)
- in
- let cstu = (cst, if poly then Univ.UContext.instance ctx else Univ.Instance.empty) in
- let inst_type = appvectc (mkConstU cstu) (Termops.rel_vect 0 (List.length params)) in
- let proj_type =
- it_mkProd_or_LetIn (mkProd(Name binder_name, inst_type, lift 1 field)) params in
- let proj_body =
- it_mkLambda_or_LetIn (mkLambda (Name binder_name, inst_type, mkRel 1)) params in
- let proj_entry = Declare.definition_entry ~types:proj_type ~poly ~univs:ctx proj_body in
- let proj_cst = Declare.declare_constant proj_name
- (DefinitionEntry proj_entry, IsDefinition Definition)
- in
- let cref = ConstRef cst in
- Impargs.declare_manual_implicits false cref [paramimpls];
- Impargs.declare_manual_implicits false (ConstRef proj_cst) [List.hd fieldimpls];
- Classes.set_typeclass_transparency (EvalConstRef cst) false false;
- let sub = match List.hd coers with
- | Some b -> Some ((if b then Backward else Forward), List.hd priorities)
- | None -> None
- in
- cref, [Name proj_name, sub, Some proj_cst]
+ let class_body = it_mkLambda_or_LetIn field params in
+ let _class_type = it_mkProd_or_LetIn arity params in
+ let class_entry =
+ Declare.definition_entry (* ?types:class_type *) ~poly ~univs:ctx class_body in
+ let cst = Declare.declare_constant (snd id)
+ (DefinitionEntry class_entry, IsDefinition Definition)
+ in
+ let cstu = (cst, if poly then Univ.UContext.instance ctx else Univ.Instance.empty) in
+ let inst_type = appvectc (mkConstU cstu)
+ (Termops.rel_vect 0 (List.length params)) in
+ let proj_type =
+ it_mkProd_or_LetIn (mkProd(Name binder_name, inst_type, lift 1 field)) params in
+ let proj_body =
+ it_mkLambda_or_LetIn (mkLambda (Name binder_name, inst_type, mkRel 1)) params in
+ let proj_entry =
+ Declare.definition_entry ~types:proj_type ~poly
+ ~univs:(if poly then ctx else Univ.UContext.empty) proj_body
+ in
+ let proj_cst = Declare.declare_constant proj_name
+ (DefinitionEntry proj_entry, IsDefinition Definition)
+ in
+ let cref = ConstRef cst in
+ Impargs.declare_manual_implicits false cref [paramimpls];
+ Impargs.declare_manual_implicits false (ConstRef proj_cst) [List.hd fieldimpls];
+ Classes.set_typeclass_transparency (EvalConstRef cst) false false;
+ let sub = match List.hd coers with
+ | Some b -> Some ((if b then Backward else Forward), List.hd priorities)
+ | None -> None
+ in
+ cref, [Name proj_name, sub, Some proj_cst]
| _ ->
- let ind = declare_structure BiFinite poly ctx (snd id) idbuild paramimpls
+ let ind = declare_structure BiFinite poly ctx (snd id) idbuild paramimpls
params arity template fieldimpls fields
~kind:Method ~name:binder_name false (List.map (fun _ -> false) fields) sign
- in
- let coers = List.map2 (fun coe pri ->
- Option.map (fun b ->
- if b then Backward, pri else Forward, pri) coe)
+ in
+ let coers = List.map2 (fun coe pri ->
+ Option.map (fun b ->
+ if b then Backward, pri else Forward, pri) coe)
coers priorities
- in
- IndRef ind, (List.map3 (fun (id, _, _) b y -> (id, b, y))
- (List.rev fields) coers (Recordops.lookup_projections ind))
+ in
+ let l = List.map3 (fun (id, _, _) b y -> (id, b, y))
+ (List.rev fields) coers (Recordops.lookup_projections ind)
+ in IndRef ind, l
in
let ctx_context =
List.map (fun (na, b, t) ->
diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml
index 8ae6ac2bc..c07c756c0 100644
--- a/toplevel/vernacentries.ml
+++ b/toplevel/vernacentries.ml
@@ -76,9 +76,8 @@ let show_universes () =
let gls = Proof.V82.subgoals pfts in
let sigma = gls.Evd.sigma in
let ctx = Evd.universe_context_set (Evd.nf_constraints sigma) in
- let cstrs = Univ.merge_constraints (Univ.ContextSet.constraints ctx) Univ.empty_universes in
msg_notice (Evd.pr_evar_universe_context (Evd.evar_universe_context sigma));
- msg_notice (str"Normalized constraints: " ++ Univ.pr_universes (Evd.pr_evd_level sigma) cstrs)
+ msg_notice (str"Normalized constraints: " ++ Univ.pr_universe_context_set (Evd.pr_evd_level sigma) ctx)
let show_prooftree () =
(* Spiwack: proof tree is currently not working *)
@@ -1185,8 +1184,9 @@ let default_env () = {
let vernac_reserve bl =
let sb_decl = (fun (idl,c) ->
let env = Global.env() in
- let t,ctx = Constrintern.interp_type env Evd.empty c in
- let t = Detyping.detype false [] env Evd.empty t in
+ let sigma = Evd.from_env env in
+ let t,ctx = Constrintern.interp_type env sigma c in
+ let t = Detyping.detype false [] env (Evd.from_ctx ctx) t in
let t = Notation_ops.notation_constr_of_glob_constr (default_env ()) t in
Reserve.declare_reserved_type idl t)
in List.iter sb_decl bl