aboutsummaryrefslogtreecommitdiffhomepage
path: root/engine
diff options
context:
space:
mode:
Diffstat (limited to 'engine')
-rw-r--r--engine/eConstr.ml1
-rw-r--r--engine/eConstr.mli4
-rw-r--r--engine/evar_kinds.ml3
-rw-r--r--engine/evarutil.ml2
-rw-r--r--engine/evarutil.mli9
-rw-r--r--engine/evd.ml32
-rw-r--r--engine/evd.mli2
-rw-r--r--engine/termops.mli8
-rw-r--r--engine/universes.ml317
-rw-r--r--engine/universes.mli20
10 files changed, 181 insertions, 217 deletions
diff --git a/engine/eConstr.ml b/engine/eConstr.ml
index a72bdee12..d30cb74d7 100644
--- a/engine/eConstr.ml
+++ b/engine/eConstr.ml
@@ -500,6 +500,7 @@ let eq_universes env sigma cstrs cv_pb ref nargs l l' =
let l = Evd.normalize_universe_instance sigma l
and l' = Evd.normalize_universe_instance sigma l' in
let open Universes in
+ let open GlobRef in
match ref with
| VarRef _ -> assert false (* variables don't have instances *)
| ConstRef _ ->
diff --git a/engine/eConstr.mli b/engine/eConstr.mli
index 9a5b5ec3a..743888a9b 100644
--- a/engine/eConstr.mli
+++ b/engine/eConstr.mli
@@ -284,9 +284,9 @@ val map_rel_context_in_env :
(* XXX Missing Sigma proxy *)
val fresh_global :
?loc:Loc.t -> ?rigid:Evd.rigid -> ?names:Univ.Instance.t -> Environ.env ->
- Evd.evar_map -> Globnames.global_reference -> Evd.evar_map * t
+ Evd.evar_map -> GlobRef.t -> Evd.evar_map * t
-val is_global : Evd.evar_map -> Globnames.global_reference -> t -> bool
+val is_global : Evd.evar_map -> GlobRef.t -> t -> bool
(** {5 Extra} *)
diff --git a/engine/evar_kinds.ml b/engine/evar_kinds.ml
index c964ecf1f..6e123d642 100644
--- a/engine/evar_kinds.ml
+++ b/engine/evar_kinds.ml
@@ -9,7 +9,6 @@
(************************************************************************)
open Names
-open Globnames
open Misctypes
(** The kinds of existential variable *)
@@ -24,7 +23,7 @@ type matching_var_kind = FirstOrderPatVar of patvar | SecondOrderPatVar of patva
type subevar_kind = Domain | Codomain | Body
type t =
- | ImplicitArg of global_reference * (int * Id.t option)
+ | ImplicitArg of GlobRef.t * (int * Id.t option)
* bool (** Force inference *)
| BinderType of Name.t
| NamedHole of Id.t (* coming from some ?[id] syntax *)
diff --git a/engine/evarutil.ml b/engine/evarutil.ml
index 710491f84..6c27d5937 100644
--- a/engine/evarutil.ml
+++ b/engine/evarutil.ml
@@ -534,7 +534,7 @@ type clear_dependency_error =
| OccurHypInSimpleClause of Id.t option
| EvarTypingBreak of existential
-exception ClearDependencyError of Id.t * clear_dependency_error * Globnames.global_reference option
+exception ClearDependencyError of Id.t * clear_dependency_error * GlobRef.t option
exception Depends of Id.t
diff --git a/engine/evarutil.mli b/engine/evarutil.mli
index d3937f28e..3bbd2923c 100644
--- a/engine/evarutil.mli
+++ b/engine/evarutil.mli
@@ -78,8 +78,8 @@ val restrict_evar : evar_map -> Evar.t -> Filter.t ->
(** Polymorphic constants *)
-val new_global : evar_map -> Globnames.global_reference -> evar_map * constr
-val e_new_global : evar_map ref -> Globnames.global_reference -> constr
+val new_global : evar_map -> GlobRef.t -> evar_map * constr
+val e_new_global : evar_map ref -> GlobRef.t -> constr
(** Create a fresh evar in a context different from its definition context:
[new_evar_instance sign evd ty inst] creates a new evar of context
@@ -186,11 +186,14 @@ val nf_evar_map_undefined : evar_map -> evar_map
val nf_evars_universes : evar_map -> Constr.constr -> Constr.constr
val nf_evars_and_universes : evar_map -> evar_map * (Constr.constr -> Constr.constr)
+[@@ocaml.deprecated "Use Evd.minimize_universes and nf_evars_universes"]
val e_nf_evars_and_universes : evar_map ref -> (Constr.constr -> Constr.constr) * Universes.universe_opt_subst
+[@@ocaml.deprecated "Use Evd.minimize_universes and nf_evars_universes"]
(** Normalize the evar map w.r.t. universes, after simplification of constraints.
Return the substitution function for constrs as well. *)
val nf_evar_map_universes : evar_map -> evar_map * (Constr.constr -> Constr.constr)
+[@@ocaml.deprecated "Use Evd.minimize_universes and nf_evar_map and nf_evars_universes"]
(** Replacing all evars, possibly raising [Uninstantiated_evar] *)
exception Uninstantiated_evar of Evar.t
@@ -232,7 +235,7 @@ type clear_dependency_error =
| OccurHypInSimpleClause of Id.t option
| EvarTypingBreak of Constr.existential
-exception ClearDependencyError of Id.t * clear_dependency_error * Globnames.global_reference option
+exception ClearDependencyError of Id.t * clear_dependency_error * GlobRef.t option
val clear_hyps_in_evi : env -> evar_map ref -> named_context_val -> types ->
Id.Set.t -> named_context_val * types
diff --git a/engine/evd.ml b/engine/evd.ml
index 6dcec2760..af22de5cd 100644
--- a/engine/evd.ml
+++ b/engine/evd.ml
@@ -1290,30 +1290,14 @@ module MiniEConstr = struct
let unsafe_eq = Refl
let to_constr ?(abort_on_undefined_evars=true) sigma c =
- let rec to_constr c = match Constr.kind c with
- | Evar ev ->
- begin match safe_evar_value sigma ev with
- | Some c -> to_constr c
- | None ->
- if abort_on_undefined_evars then
- anomaly ~label:"econstr" Pp.(str "grounding a non evar-free term")
- else
- Constr.map (fun c -> to_constr c) c
- end
- | Sort (Sorts.Type u) ->
- let u' = normalize_universe sigma u in
- if u' == u then c else mkSort (Sorts.sort_of_univ u')
- | Const (c', u) when not (Univ.Instance.is_empty u) ->
- let u' = normalize_universe_instance sigma u in
- if u' == u then c else mkConstU (c', u')
- | Ind (i, u) when not (Univ.Instance.is_empty u) ->
- let u' = normalize_universe_instance sigma u in
- if u' == u then c else mkIndU (i, u')
- | Construct (co, u) when not (Univ.Instance.is_empty u) ->
- let u' = normalize_universe_instance sigma u in
- if u' == u then c else mkConstructU (co, u')
- | _ -> Constr.map (fun c -> to_constr c) c
- in to_constr c
+ let evar_value =
+ if not abort_on_undefined_evars then fun ev -> safe_evar_value sigma ev
+ else fun ev ->
+ match safe_evar_value sigma ev with
+ | Some _ as v -> v
+ | None -> anomaly ~label:"econstr" Pp.(str "grounding a non evar-free term")
+ in
+ Universes.nf_evars_and_universes_opt_subst evar_value (universe_subst sigma) c
let of_named_decl d = d
let unsafe_to_named_decl d = d
diff --git a/engine/evd.mli b/engine/evd.mli
index 5ce16459c..509db525d 100644
--- a/engine/evd.mli
+++ b/engine/evd.mli
@@ -649,7 +649,7 @@ val fresh_inductive_instance : ?loc:Loc.t -> env -> evar_map -> inductive -> eva
val fresh_constructor_instance : ?loc:Loc.t -> env -> evar_map -> constructor -> evar_map * pconstructor
val fresh_global : ?loc:Loc.t -> ?rigid:rigid -> ?names:Univ.Instance.t -> env ->
- evar_map -> Globnames.global_reference -> evar_map * econstr
+ evar_map -> GlobRef.t -> evar_map * econstr
(********************************************************************)
(* constr with holes and pending resolution of classes, conversion *)
diff --git a/engine/termops.mli b/engine/termops.mli
index 3b0c4bba6..e2ddcd36e 100644
--- a/engine/termops.mli
+++ b/engine/termops.mli
@@ -112,7 +112,7 @@ val dependent_in_decl : Evd.evar_map -> constr -> named_declaration -> bool
val count_occurrences : Evd.evar_map -> constr -> constr -> int
val collect_metas : Evd.evar_map -> constr -> int list
val collect_vars : Evd.evar_map -> constr -> Id.Set.t (** for visible vars only *)
-val vars_of_global_reference : env -> Globnames.global_reference -> Id.Set.t
+val vars_of_global_reference : env -> GlobRef.t -> Id.Set.t
val occur_term : Evd.evar_map -> constr -> constr -> bool (** Synonymous of dependent *)
[@@ocaml.deprecated "alias of Termops.dependent"]
@@ -261,7 +261,7 @@ val clear_named_body : Id.t -> env -> env
val global_vars : env -> Evd.evar_map -> constr -> Id.t list
val global_vars_set : env -> Evd.evar_map -> constr -> Id.Set.t
val global_vars_set_of_decl : env -> Evd.evar_map -> named_declaration -> Id.Set.t
-val global_app_of_constr : Evd.evar_map -> constr -> (Globnames.global_reference * EInstance.t) * constr option
+val global_app_of_constr : Evd.evar_map -> constr -> (GlobRef.t * EInstance.t) * constr option
(** Gives an ordered list of hypotheses, closed by dependencies,
containing a given set *)
@@ -270,9 +270,9 @@ val dependency_closure : env -> Evd.evar_map -> named_context -> Id.Set.t -> Id.
(** Test if an identifier is the basename of a global reference *)
val is_section_variable : Id.t -> bool
-val global_of_constr : Evd.evar_map -> constr -> Globnames.global_reference * EInstance.t
+val global_of_constr : Evd.evar_map -> constr -> GlobRef.t * EInstance.t
-val is_global : Evd.evar_map -> Globnames.global_reference -> constr -> bool
+val is_global : Evd.evar_map -> GlobRef.t -> constr -> bool
val isGlobalRef : Evd.evar_map -> constr -> bool
diff --git a/engine/universes.ml b/engine/universes.ml
index e5f9212a7..938f5ad9c 100644
--- a/engine/universes.ml
+++ b/engine/universes.ml
@@ -74,7 +74,7 @@ let subst_ubinder (subst,(ref,l as orig)) =
let discharge_ubinder (_,(ref,l)) =
Some (Lib.discharge_global ref, l)
-let ubinder_obj : Globnames.global_reference * universe_binders -> Libobject.obj =
+let ubinder_obj : GlobRef.t * universe_binders -> Libobject.obj =
let open Libobject in
declare_object { (default_object "universe binder") with
cache_function = cache_ubinder;
@@ -524,8 +524,6 @@ let new_global_univ () =
(** Simplification *)
-module LevelUnionFind = Unionfind.Make (Univ.LSet) (Univ.LMap)
-
let add_list_map u t map =
try
let l = LMap.find u map in
@@ -533,8 +531,6 @@ let add_list_map u t map =
with Not_found ->
LMap.add u [t] map
-module UF = LevelUnionFind
-
(** Precondition: flexible <= ctx *)
let choose_canonical ctx flexible algs s =
let global = LSet.diff s ctx in
@@ -596,35 +592,6 @@ let subst_univs_constr =
CProfile.profile2 subst_univs_constr_key subst_univs_constr
else subst_univs_constr
-let subst_univs_fn_puniverses lsubst (c, u as cu) =
- let u' = Instance.subst_fn lsubst u in
- if u' == u then cu else (c, u')
-
-let nf_evars_and_universes_opt_subst f subst =
- let subst = fun l -> match LMap.find l subst with None -> raise Not_found | Some l' -> l' in
- let lsubst = level_subst_of subst in
- let rec aux c =
- match kind c with
- | Evar (evk, args) ->
- let args = Array.map aux args in
- (match try f (evk, args) with Not_found -> None with
- | None -> c
- | Some c -> aux c)
- | Const pu ->
- let pu' = subst_univs_fn_puniverses lsubst pu in
- if pu' == pu then c else mkConstU pu'
- | Ind pu ->
- let pu' = subst_univs_fn_puniverses lsubst pu in
- if pu' == pu then c else mkIndU pu'
- | Construct pu ->
- let pu' = subst_univs_fn_puniverses lsubst pu in
- if pu' == pu then c else mkConstructU pu'
- | Sort (Type u) ->
- let u' = Univ.subst_univs_universe subst u in
- if u' == u then c else mkSort (sort_of_univ u')
- | _ -> Constr.map aux c
- in aux
-
let fresh_universe_context_set_instance ctx =
if ContextSet.is_empty ctx then LMap.empty, ctx
else
@@ -683,6 +650,35 @@ let normalize_opt_subst ctx =
type universe_opt_subst = Universe.t option universe_map
+let subst_univs_fn_puniverses f (c, u as cu) =
+ let u' = Instance.subst_fn f u in
+ if u' == u then cu else (c, u')
+
+let nf_evars_and_universes_opt_subst f subst =
+ let subst = normalize_univ_variable_opt_subst (ref subst) in
+ let lsubst = level_subst_of subst in
+ let rec aux c =
+ match kind c with
+ | Evar (evk, args) ->
+ let args = Array.map aux args in
+ (match try f (evk, args) with Not_found -> None with
+ | None -> mkEvar (evk, args)
+ | Some c -> aux c)
+ | Const pu ->
+ let pu' = subst_univs_fn_puniverses lsubst pu in
+ if pu' == pu then c else mkConstU pu'
+ | Ind pu ->
+ let pu' = subst_univs_fn_puniverses lsubst pu in
+ if pu' == pu then c else mkIndU pu'
+ | Construct pu ->
+ let pu' = subst_univs_fn_puniverses lsubst pu in
+ if pu' == pu then c else mkConstructU pu'
+ | Sort (Type u) ->
+ let u' = Univ.subst_univs_universe subst u in
+ if u' == u then c else mkSort (sort_of_univ u')
+ | _ -> Constr.map aux c
+ in aux
+
let make_opt_subst s =
fun x ->
(match Univ.LMap.find x s with
@@ -709,6 +705,7 @@ let pr_universe_body = function
let pr_universe_opt_subst = Univ.LMap.pr pr_universe_body
+(* Eq < Le < Lt *)
let compare_constraint_type d d' =
match d, d' with
| Eq, Eq -> 0
@@ -742,10 +739,12 @@ let lower_add l c m =
let lower_of_list l =
List.fold_left (fun acc (d,l) -> LMap.add l d acc) LMap.empty l
+type lbound = { enforce : bool; alg : bool; lbound: Universe.t; lower : lowermap }
+
exception Found of Level.t * lowermap
let find_inst insts v =
- try LMap.iter (fun k (enf,alg,v',lower) ->
- if not alg && enf && Universe.equal v' v then raise (Found (k, lower)))
+ try LMap.iter (fun k {enforce;alg;lbound=v';lower} ->
+ if not alg && enforce && Universe.equal v' v then raise (Found (k, lower)))
insts; raise Not_found
with Found (f,l) -> (f,l)
@@ -765,18 +764,18 @@ let compute_lbound left =
sup (Universe.super l) lbound
else None))
None left
-
-let instantiate_with_lbound u lbound lower alg enforce (ctx, us, algs, insts, cstrs) =
+
+let instantiate_with_lbound u lbound lower ~alg ~enforce (ctx, us, algs, insts, cstrs) =
if enforce then
let inst = Universe.make u in
let cstrs' = enforce_leq lbound inst cstrs in
(ctx, us, LSet.remove u algs,
- LMap.add u (enforce,alg,lbound,lower) insts, cstrs'),
- (enforce, alg, inst, lower)
+ LMap.add u {enforce;alg;lbound;lower} insts, cstrs'),
+ {enforce; alg; lbound=inst; lower}
else (* Actually instantiate *)
(Univ.LSet.remove u ctx, Univ.LMap.add u (Some lbound) us, algs,
- LMap.add u (enforce,alg,lbound,lower) insts, cstrs),
- (enforce, alg, lbound, lower)
+ LMap.add u {enforce;alg;lbound;lower} insts, cstrs),
+ {enforce; alg; lbound; lower}
type constraints_map = (Univ.constraint_type * Univ.LMap.key) list Univ.LMap.t
@@ -790,73 +789,82 @@ let _pr_constraints_map (cmap:constraints_map) =
let remove_alg l (ctx, us, algs, insts, cstrs) =
(ctx, us, LSet.remove l algs, insts, cstrs)
-let remove_lower u lower =
- let levels = Universe.levels u in
- LSet.fold (fun l acc -> LMap.remove l acc) levels lower
-
+let not_lower lower (d,l) =
+ (* We're checking if (d,l) is already implied by the lower
+ constraints on some level u. If it represents l < u (d is Lt
+ or d is Le and i > 0, the i < 0 case is impossible due to
+ invariants of Univ), and the lower constraints only have l <=
+ u then it is not implied. *)
+ Univ.Universe.exists
+ (fun (l,i) ->
+ let d =
+ if i == 0 then d
+ else match d with
+ | Le -> Lt
+ | d -> d
+ in
+ try let d' = LMap.find l lower in
+ (* If d is stronger than the already implied lower
+ * constraints we must keep it. *)
+ compare_constraint_type d d' > 0
+ with Not_found ->
+ (** No constraint existing on l *) true) l
+
+exception UpperBoundedAlg
+(** [enforce_uppers upper lbound cstrs] interprets [upper] as upper
+ constraints to [lbound], adding them to [cstrs].
+
+ @raise UpperBoundedAlg if any [upper] constraints are strict and
+ [lbound] algebraic. *)
+let enforce_uppers upper lbound cstrs =
+ List.fold_left (fun cstrs (d, r) ->
+ if d == Univ.Le then
+ enforce_leq lbound (Universe.make r) cstrs
+ else
+ match Universe.level lbound with
+ | Some lev -> Constraint.add (lev, d, r) cstrs
+ | None -> raise UpperBoundedAlg)
+ cstrs upper
+
let minimize_univ_variables ctx us algs left right cstrs =
let left, lbounds =
Univ.LMap.fold (fun r lower (left, lbounds as acc) ->
if Univ.LMap.mem r us || not (Univ.LSet.mem r ctx) then acc
else (* Fixed universe, just compute its glb for sharing *)
- let lbounds' =
+ let lbounds =
match compute_lbound (List.map (fun (d,l) -> d, Universe.make l) lower) with
| None -> lbounds
- | Some lbound -> LMap.add r (true, false, lbound, lower_of_list lower)
+ | Some lbound -> LMap.add r {enforce=true; alg=false; lbound; lower=lower_of_list lower}
lbounds
- in (Univ.LMap.remove r left, lbounds'))
+ in (Univ.LMap.remove r left, lbounds))
left (left, Univ.LMap.empty)
in
- let rec instance (ctx', us, algs, insts, cstrs as acc) u =
+ let rec instance (ctx, us, algs, insts, cstrs as acc) u =
let acc, left, lower =
- try
- let l = LMap.find u left in
+ match LMap.find u left with
+ | exception Not_found -> acc, [], LMap.empty
+ | l ->
let acc, left, newlow, lower =
List.fold_left
- (fun (acc, left', newlow, lower') (d, l) ->
- let acc', (enf,alg,l',lower) = aux acc l in
+ (fun (acc, left, newlow, lower') (d, l) ->
+ let acc', {enforce=enf;alg;lbound=l';lower} = aux acc l in
let l' =
if enf then Universe.make l
else l'
- in acc', (d, l') :: left',
+ in acc', (d, l') :: left,
lower_add l d newlow, lower_union lower lower')
(acc, [], LMap.empty, LMap.empty) l
in
- let not_lower (d,l) =
- (* We're checking if (d,l) is already implied by the lower
- constraints on some level u. If it represents l < u (d is Lt
- or d is Le and i > 0, the i < 0 case is impossible due to
- invariants of Univ), and the lower constraints only have l <=
- u then it is not implied. *)
- Univ.Universe.exists
- (fun (l,i) ->
- let d =
- if i == 0 then d
- else match d with
- | Le -> Lt
- | d -> d
- in
- try let d' = LMap.find l lower in
- (* If d is stronger than the already implied lower
- * constraints we must keep it. *)
- compare_constraint_type d d' > 0
- with Not_found ->
- (** No constraint existing on l *) true) l
- in
- let left = List.uniquize (List.filter not_lower left) in
+ let left = List.uniquize (List.filter (not_lower lower) left) in
(acc, left, LMap.union newlow lower)
- with Not_found -> acc, [], LMap.empty
- and right =
- try Some (LMap.find u right)
- with Not_found -> None
in
let instantiate_lbound lbound =
let alg = LSet.mem u algs in
if alg then
(* u is algebraic: we instantiate it with its lower bound, if any,
or enforce the constraints if it is bounded from the top. *)
- let lower = remove_lower lbound lower in
- instantiate_with_lbound u lbound lower true false acc
+ let lower = LSet.fold LMap.remove (Universe.levels lbound) lower in
+ instantiate_with_lbound u lbound lower ~alg:true ~enforce:false acc
else (* u is non algebraic *)
match Universe.level lbound with
| Some l -> (* The lowerbound is directly a level *)
@@ -867,125 +875,96 @@ let minimize_univ_variables ctx us algs left right cstrs =
if not (Level.equal l u) then
(* Should check that u does not
have upper constraints that are not already in right *)
- let acc' = remove_alg l acc in
- instantiate_with_lbound u lbound lower false false acc'
- else acc, (true, false, lbound, lower)
+ let acc = remove_alg l acc in
+ instantiate_with_lbound u lbound lower ~alg:false ~enforce:false acc
+ else acc, {enforce=true; alg=false; lbound; lower}
| None ->
- try
- (* Another universe represents the same lower bound,
- we can share them with no harm. *)
- let can, lower = find_inst insts lbound in
- let lower = LMap.remove can lower in
- instantiate_with_lbound u (Universe.make can) lower false false acc
- with Not_found ->
- (* We set u as the canonical universe representing lbound *)
- instantiate_with_lbound u lbound lower false true acc
+ begin match find_inst insts lbound with
+ | can, lower ->
+ (* Another universe represents the same lower bound,
+ we can share them with no harm. *)
+ let lower = LMap.remove can lower in
+ instantiate_with_lbound u (Universe.make can) lower ~alg:false ~enforce:false acc
+ | exception Not_found ->
+ (* We set u as the canonical universe representing lbound *)
+ instantiate_with_lbound u lbound lower ~alg:false ~enforce:true acc
+ end
in
- let acc' acc =
- match right with
- | None -> acc
- | Some cstrs ->
- let dangling = List.filter (fun (d, r) -> not (LMap.mem r us)) cstrs in
- if List.is_empty dangling then acc
- else
- let ((ctx', us, algs, insts, cstrs), (enf,_,inst,lower as b)) = acc in
- let cstrs' = List.fold_left (fun cstrs (d, r) ->
- if d == Univ.Le then
- enforce_leq inst (Universe.make r) cstrs
- else
- try let lev = Option.get (Universe.level inst) in
- Constraint.add (lev, d, r) cstrs
- with Option.IsNone -> failwith "")
- cstrs dangling
- in
- (ctx', us, algs, insts, cstrs'), b
+ let enforce_uppers ((ctx,us,algs,insts,cstrs), b as acc) =
+ match LMap.find u right with
+ | exception Not_found -> acc
+ | upper ->
+ let upper = List.filter (fun (d, r) -> not (LMap.mem r us)) upper in
+ let cstrs = enforce_uppers upper b.lbound cstrs in
+ (ctx, us, algs, insts, cstrs), b
in
- if not (LSet.mem u ctx) then acc' (acc, (true, false, Universe.make u, lower))
- else
- let lbound = compute_lbound left in
- match lbound with
- | None -> (* Nothing to do *)
- acc' (acc, (true, false, Universe.make u, lower))
- | Some lbound ->
- try acc' (instantiate_lbound lbound)
- with Failure _ -> acc' (acc, (true, false, Universe.make u, lower))
- and aux (ctx', us, algs, seen, cstrs as acc) u =
+ if not (LSet.mem u ctx)
+ then enforce_uppers (acc, {enforce=true; alg=false; lbound=Universe.make u; lower})
+ else
+ let lbound = compute_lbound left in
+ match lbound with
+ | None -> (* Nothing to do *)
+ enforce_uppers (acc, {enforce=true;alg=false;lbound=Universe.make u; lower})
+ | Some lbound ->
+ try enforce_uppers (instantiate_lbound lbound)
+ with UpperBoundedAlg ->
+ enforce_uppers (acc, {enforce=true; alg=false; lbound=Universe.make u; lower})
+ and aux (ctx, us, algs, seen, cstrs as acc) u =
try acc, LMap.find u seen
with Not_found -> instance acc u
in
- LMap.fold (fun u v (ctx', us, algs, seen, cstrs as acc) ->
+ LMap.fold (fun u v (ctx, us, algs, seen, cstrs as acc) ->
if v == None then fst (aux acc u)
- else LSet.remove u ctx', us, LSet.remove u algs, seen, cstrs)
+ else LSet.remove u ctx, us, LSet.remove u algs, seen, cstrs)
us (ctx, us, algs, lbounds, cstrs)
let normalize_context_set g ctx us algs weak =
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
- if is_set_minimization () && LSet.mem r ctx then
- (Constraint.add cstr smallles, noneqs)
- else (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)
+ Constraint.partition (fun (l,d,r) -> d == Le && Level.is_small l) csts
+ in
+ let smallles = if is_set_minimization ()
+ then Constraint.filter (fun (l,d,r) -> LSet.mem r ctx) smallles
+ else Constraint.empty
in
- let csts =
+ let csts, partition =
(* We first put constraints in a normal-form: all self-loops are collapsed
to equalities. *)
- let g = Univ.LSet.fold (fun v g -> UGraph.add_universe v false g)
+ let g = LSet.fold (fun v g -> UGraph.add_universe v false g)
ctx UGraph.initial_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 UGraph.add_universe l false g
- with UGraph.AlreadyDeclared -> g
- else g
- in
- let g =
- if not (Level.is_small r || LSet.mem r ctx) then
- try UGraph.add_universe r false g
- with UGraph.AlreadyDeclared -> g
- else g
- in g) csts g
+ let add_soft u g =
+ if not (Level.is_small u || LSet.mem u ctx)
+ then try UGraph.add_universe u false g with UGraph.AlreadyDeclared -> g
+ else g
+ in
+ let g = Constraint.fold
+ (fun (l, d, r) g -> add_soft r (add_soft l g))
+ csts g
in
- let g = Univ.Constraint.fold UGraph.enforce_constraint csts g in
+ let g = UGraph.merge_constraints csts g in
UGraph.constraints_of_universes g
in
+ (* We ignore the trivial Prop/Set <= i constraints. *)
let noneqs =
- 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 if Univ.Level.is_prop l && d == Lt && Univ.Level.is_set r
- then noneqs
- else Constraint.add cstr noneqs)
- csts Constraint.empty
+ Constraint.filter
+ (fun (l,d,r) -> not ((d == Le && Level.is_small l) ||
+ (Level.is_prop l && d == Lt && Level.is_set r)))
+ csts
in
let noneqs = Constraint.union noneqs smallles in
- let partition = UF.partition uf in
let flex x = LMap.mem x us in
let ctx, us, eqs = List.fold_left (fun (ctx, us, cstrs) s ->
let canon, (global, rigid, flexible) = choose_canonical ctx flex algs s in
(* Add equalities for globals which can't be merged anymore. *)
let cstrs = LSet.fold (fun g cst ->
- Constraint.add (canon, Univ.Eq, g) cst) global
+ Constraint.add (canon, Eq, g) cst) global
cstrs
in
(* Also add equalities for rigid variables *)
let cstrs = LSet.fold (fun g cst ->
- Constraint.add (canon, Univ.Eq, g) cst) rigid
+ Constraint.add (canon, Eq, g) cst) rigid
cstrs
in
let canonu = Some (Universe.make canon) in
diff --git a/engine/universes.mli b/engine/universes.mli
index 4823c5746..e6bee42bb 100644
--- a/engine/universes.mli
+++ b/engine/universes.mli
@@ -37,8 +37,8 @@ type universe_binders = Univ.Level.t Names.Id.Map.t
val empty_binders : universe_binders
-val register_universe_binders : Globnames.global_reference -> universe_binders -> unit
-val universe_binders_of_global : Globnames.global_reference -> universe_binders
+val register_universe_binders : GlobRef.t -> universe_binders -> unit
+val universe_binders_of_global : GlobRef.t -> universe_binders
type univ_name_list = Misctypes.lname list
@@ -48,7 +48,7 @@ type univ_name_list = Misctypes.lname list
May error if the lengths mismatch.
Otherwise return [universe_binders_of_global ref]. *)
-val universe_binders_with_opt_names : Globnames.global_reference ->
+val universe_binders_with_opt_names : Names.GlobRef.t ->
Univ.Level.t list -> univ_name_list option -> universe_binders
(** The global universe counter *)
@@ -132,7 +132,7 @@ val fresh_inductive_instance : env -> inductive ->
val fresh_constructor_instance : env -> constructor ->
pconstructor in_universe_context_set
-val fresh_global_instance : ?names:Univ.Instance.t -> env -> Globnames.global_reference ->
+val fresh_global_instance : ?names:Univ.Instance.t -> env -> GlobRef.t ->
constr in_universe_context_set
val fresh_global_or_constr_instance : env -> Globnames.global_reference_or_constr ->
@@ -144,9 +144,9 @@ val fresh_universe_context_set_instance : ContextSet.t ->
universe_level_subst * ContextSet.t
(** Raises [Not_found] if not a global reference. *)
-val global_of_constr : constr -> Globnames.global_reference puniverses
+val global_of_constr : constr -> GlobRef.t puniverses
-val constr_of_global_univ : Globnames.global_reference puniverses -> constr
+val constr_of_global_univ : GlobRef.t puniverses -> constr
val extend_context : 'a in_universe_context_set -> ContextSet.t ->
'a in_universe_context_set
@@ -162,8 +162,6 @@ val extend_context : 'a in_universe_context_set -> ContextSet.t ->
(a global one if there is one) and transitively saturate
the constraints w.r.t to the equalities. *)
-module UF : Unionfind.PartitionSig with type elt = Level.t
-
val level_subst_of : universe_subst_fn -> universe_level_subst_fn
val subst_univs_constraints : universe_subst_fn -> Constraint.t -> Constraint.t
@@ -206,16 +204,16 @@ val normalize_universe_subst : universe_subst ref ->
the constraints should be properly added to an evd.
See Evd.fresh_global, Evarutil.new_global, and pf_constr_of_global for
the proper way to get a fresh copy of a global reference. *)
-val constr_of_global : Globnames.global_reference -> constr
+val constr_of_global : GlobRef.t -> constr
(** ** DEPRECATED ** synonym of [constr_of_global] *)
-val constr_of_reference : Globnames.global_reference -> constr
+val constr_of_reference : GlobRef.t -> constr
[@@ocaml.deprecated "synonym of [constr_of_global]"]
(** Returns the type of the global reference, by creating a fresh instance of polymorphic
references and computing their instantiated universe context. (side-effect on the
universe counter, use with care). *)
-val type_of_global : Globnames.global_reference -> types in_universe_context_set
+val type_of_global : GlobRef.t -> types in_universe_context_set
(** Full universes substitutions into terms *)