aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--checker/univ.ml8
-rw-r--r--checker/univ.mli2
-rw-r--r--engine/evd.ml2
-rw-r--r--engine/uState.ml2
-rw-r--r--engine/universes.ml78
-rw-r--r--engine/universes.mli5
-rw-r--r--kernel/univ.ml32
-rw-r--r--kernel/univ.mli9
-rw-r--r--kernel/vars.ml43
-rw-r--r--kernel/vars.mli6
-rw-r--r--pretyping/unification.ml6
-rw-r--r--tactics/ind_tables.ml3
12 files changed, 89 insertions, 107 deletions
diff --git a/checker/univ.ml b/checker/univ.ml
index 4f3131813..7d01657df 100644
--- a/checker/univ.ml
+++ b/checker/univ.ml
@@ -881,14 +881,6 @@ type universe_level_subst = universe_level universe_map
(** A full substitution might involve algebraic universes *)
type universe_subst = universe universe_map
-let level_subst_of f =
- fun l ->
- try let u = f l in
- match Universe.level u with
- | None -> l
- | Some l -> l
- with Not_found -> l
-
module Instance : sig
type t = Level.t array
diff --git a/checker/univ.mli b/checker/univ.mli
index 0eadc6801..21c94d952 100644
--- a/checker/univ.mli
+++ b/checker/univ.mli
@@ -150,8 +150,6 @@ type universe_level_subst_fn = universe_level -> universe_level
type universe_subst = universe universe_map
type universe_level_subst = universe_level universe_map
-val level_subst_of : universe_subst_fn -> universe_level_subst_fn
-
(** {6 Universe instances} *)
module Instance :
diff --git a/engine/evd.ml b/engine/evd.ml
index e33c851f6..0e9472158 100644
--- a/engine/evd.ml
+++ b/engine/evd.ml
@@ -855,7 +855,7 @@ let normalize_universe evd =
let normalize_universe_instance evd l =
let vars = ref (UState.subst evd.universes) in
- let normalize = Univ.level_subst_of (Universes.normalize_univ_variable_opt_subst vars) in
+ let normalize = Universes.level_subst_of (Universes.normalize_univ_variable_opt_subst vars) in
Univ.Instance.subst_fn normalize l
let normalize_sort evars s =
diff --git a/engine/uState.ml b/engine/uState.ml
index 6131f4c03..6f2b3c4b2 100644
--- a/engine/uState.ml
+++ b/engine/uState.ml
@@ -516,7 +516,7 @@ let is_sort_variable uctx s =
| _ -> None
let subst_univs_context_with_def def usubst (ctx, cst) =
- (Univ.LSet.diff ctx def, Univ.subst_univs_constraints usubst cst)
+ (Univ.LSet.diff ctx def, Universes.subst_univs_constraints usubst cst)
let normalize_variables uctx =
let normalized_variables, undef, def, subst =
diff --git a/engine/universes.ml b/engine/universes.ml
index 30490ec56..eaddf98a8 100644
--- a/engine/universes.ml
+++ b/engine/universes.ml
@@ -181,6 +181,30 @@ let enforce_eq_instances_univs strict x y c =
(fun x y -> Constraints.add (Universe.make x, d, Universe.make y))
ax ay c
+let enforce_univ_constraint (u,d,v) =
+ match d with
+ | Eq -> enforce_eq u v
+ | Le -> enforce_leq u v
+ | Lt -> enforce_leq (super u) v
+
+let subst_univs_level fn l =
+ try Some (fn l)
+ with Not_found -> None
+
+let subst_univs_constraint fn (u,d,v as c) cstrs =
+ let u' = subst_univs_level fn u in
+ let v' = subst_univs_level fn v in
+ match u', v' with
+ | None, None -> Constraint.add c cstrs
+ | Some u, None -> enforce_univ_constraint (u,d,Universe.make v) cstrs
+ | None, Some v -> enforce_univ_constraint (Universe.make u,d,v) cstrs
+ | Some u, Some v -> enforce_univ_constraint (u,d,v) cstrs
+
+let subst_univs_constraints subst csts =
+ Constraint.fold
+ (fun c cstrs -> subst_univs_constraint subst c cstrs)
+ csts Constraint.empty
+
let subst_univs_universe_constraint fn (u,d,v) =
let u' = subst_univs_universe fn u and v' = subst_univs_universe fn v in
if Universe.equal u' v' then None
@@ -519,13 +543,60 @@ let choose_canonical ctx flexible algs s =
let canon = LSet.choose algs in
canon, (global, rigid, LSet.remove canon flexible)
+let level_subst_of f =
+ fun l ->
+ try let u = f l in
+ match Universe.level u with
+ | None -> l
+ | Some l -> l
+ with Not_found -> l
+
+let subst_univs_fn_constr f c =
+ let changed = ref false in
+ let fu = Univ.subst_univs_universe f in
+ let fi = Univ.Instance.subst_fn (level_subst_of f) in
+ let rec aux t =
+ match kind t with
+ | Sort (Sorts.Type u) ->
+ let u' = fu u in
+ if u' == u then t else
+ (changed := true; mkSort (Sorts.sort_of_univ u'))
+ | Const (c, u) ->
+ let u' = fi u in
+ if u' == u then t
+ else (changed := true; mkConstU (c, u'))
+ | Ind (i, u) ->
+ let u' = fi u in
+ if u' == u then t
+ else (changed := true; mkIndU (i, u'))
+ | Construct (c, u) ->
+ let u' = fi u in
+ if u' == u then t
+ else (changed := true; mkConstructU (c, u'))
+ | _ -> map aux t
+ in
+ let c' = aux c in
+ if !changed then c' else c
+
+let subst_univs_constr subst c =
+ if Univ.is_empty_subst subst then c
+ else
+ let f = Univ.make_subst subst in
+ subst_univs_fn_constr f c
+
+let subst_univs_constr =
+ if Flags.profile then
+ let subst_univs_constr_key = CProfile.declare_profile "subst_univs_constr" in
+ 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 = Univ.level_subst_of subst in
+ let lsubst = level_subst_of subst in
let rec aux c =
match kind c with
| Evar (evk, args) ->
@@ -605,7 +676,7 @@ let normalize_opt_subst ctx =
in !ectx
type universe_opt_subst = Universe.t option universe_map
-
+
let make_opt_subst s =
fun x ->
(match Univ.LMap.find x s with
@@ -614,8 +685,7 @@ let make_opt_subst s =
let subst_opt_univs_constr s =
let f = make_opt_subst s in
- Vars.subst_univs_fn_constr f
-
+ subst_univs_fn_constr f
let normalize_univ_variables ctx =
let ctx = normalize_opt_subst ctx in
diff --git a/engine/universes.mli b/engine/universes.mli
index 1a98d969b..130dcf8bb 100644
--- a/engine/universes.mli
+++ b/engine/universes.mli
@@ -154,6 +154,11 @@ val extend_context : 'a in_universe_context_set -> ContextSet.t ->
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
+
+val subst_univs_constr : universe_subst -> constr -> constr
+
type universe_opt_subst = Universe.t option universe_map
val make_opt_subst : universe_opt_subst -> universe_subst_fn
diff --git a/kernel/univ.ml b/kernel/univ.ml
index fee431ff4..f72f6f26a 100644
--- a/kernel/univ.ml
+++ b/kernel/univ.ml
@@ -686,12 +686,6 @@ let enforce_leq u v c =
let enforce_leq_level u v c =
if Level.equal u v then c else Constraint.add (u,Le,v) c
-let enforce_univ_constraint (u,d,v) =
- match d with
- | Eq -> enforce_eq u v
- | Le -> enforce_leq u v
- | Lt -> enforce_leq (super u) v
-
(* Miscellaneous functions to remove or test local univ assumed to
occur in a universe *)
@@ -718,14 +712,6 @@ type universe_level_subst = universe_level universe_map
(** A full substitution might involve algebraic universes *)
type universe_subst = universe universe_map
-let level_subst_of f =
- fun l ->
- try let u = f l in
- match Universe.level u with
- | None -> l
- | Some l -> l
- with Not_found -> l
-
module Instance : sig
type t = Level.t array
@@ -1128,24 +1114,6 @@ let subst_univs_universe fn ul =
List.fold_left (fun acc u -> Universe.merge_univs acc (Universe.tip u))
substs nosubst
-let subst_univs_level fn l =
- try Some (fn l)
- with Not_found -> None
-
-let subst_univs_constraint fn (u,d,v as c) cstrs =
- let u' = subst_univs_level fn u in
- let v' = subst_univs_level fn v in
- match u', v' with
- | None, None -> Constraint.add c cstrs
- | Some u, None -> enforce_univ_constraint (u,d,make v) cstrs
- | None, Some v -> enforce_univ_constraint (make u,d,v) cstrs
- | Some u, Some v -> enforce_univ_constraint (u,d,v) cstrs
-
-let subst_univs_constraints subst csts =
- Constraint.fold
- (fun c cstrs -> subst_univs_constraint subst c cstrs)
- csts Constraint.empty
-
let make_instance_subst i =
let arr = Instance.to_array i in
Array.fold_left_i (fun i acc l ->
diff --git a/kernel/univ.mli b/kernel/univ.mli
index 324167890..63bef1b81 100644
--- a/kernel/univ.mli
+++ b/kernel/univ.mli
@@ -238,8 +238,6 @@ type universe_level_subst_fn = Level.t -> Level.t
type universe_subst = Universe.t universe_map
type universe_level_subst = Level.t universe_map
-val level_subst_of : universe_subst_fn -> universe_level_subst_fn
-
(** {6 Universe instances} *)
module Instance :
@@ -461,18 +459,21 @@ val is_empty_subst : universe_subst -> bool
val make_subst : universe_subst -> universe_subst_fn
val subst_univs_universe : universe_subst_fn -> Universe.t -> Universe.t
-val subst_univs_constraints : universe_subst_fn -> Constraint.t -> Constraint.t
+(** Only user in the kernel is template polymorphism. Ideally we get rid of
+ this code if it goes away. *)
(** Substitution of instances *)
val subst_instance_instance : Instance.t -> Instance.t -> Instance.t
val subst_instance_universe : Instance.t -> Universe.t -> Universe.t
val make_instance_subst : Instance.t -> universe_level_subst
+(** Creates [u(0) ↦ 0; ...; u(n-1) ↦ n - 1] out of [u(0); ...; u(n - 1)] *)
+
val make_inverse_instance_subst : Instance.t -> universe_level_subst
val abstract_universes : UContext.t -> Instance.t * AUContext.t
-
val abstract_cumulativity_info : CumulativityInfo.t -> Instance.t * ACumulativityInfo.t
+(** TODO: move universe abstraction out of the kernel *)
val make_abstract_instance : AUContext.t -> Instance.t
diff --git a/kernel/vars.ml b/kernel/vars.ml
index eae917b5a..b3b3eff62 100644
--- a/kernel/vars.ml
+++ b/kernel/vars.ml
@@ -235,49 +235,6 @@ let subst_vars subst c = substn_vars 1 subst c
(** Universe substitutions *)
open Constr
-let subst_univs_fn_puniverses fn =
- let f = Univ.Instance.subst_fn fn in
- fun ((c, u) as x) -> let u' = f u in if u' == u then x else (c, u')
-
-let subst_univs_fn_constr f c =
- let changed = ref false in
- let fu = Univ.subst_univs_universe f in
- let fi = Univ.Instance.subst_fn (Univ.level_subst_of f) in
- let rec aux t =
- match kind t with
- | Sort (Sorts.Type u) ->
- let u' = fu u in
- if u' == u then t else
- (changed := true; mkSort (Sorts.sort_of_univ u'))
- | Const (c, u) ->
- let u' = fi u in
- if u' == u then t
- else (changed := true; mkConstU (c, u'))
- | Ind (i, u) ->
- let u' = fi u in
- if u' == u then t
- else (changed := true; mkIndU (i, u'))
- | Construct (c, u) ->
- let u' = fi u in
- if u' == u then t
- else (changed := true; mkConstructU (c, u'))
- | _ -> map aux t
- in
- let c' = aux c in
- if !changed then c' else c
-
-let subst_univs_constr subst c =
- if Univ.is_empty_subst subst then c
- else
- let f = Univ.make_subst subst in
- subst_univs_fn_constr f c
-
-let subst_univs_constr =
- if Flags.profile then
- let subst_univs_constr_key = CProfile.declare_profile "subst_univs_constr" in
- CProfile.profile2 subst_univs_constr_key subst_univs_constr
- else subst_univs_constr
-
let subst_univs_level_constr subst c =
if Univ.is_empty_level_subst subst then c
else
diff --git a/kernel/vars.mli b/kernel/vars.mli
index 964de4e95..b74d25260 100644
--- a/kernel/vars.mli
+++ b/kernel/vars.mli
@@ -129,12 +129,6 @@ val subst_var : Id.t -> constr -> constr
open Univ
-val subst_univs_fn_constr : universe_subst_fn -> constr -> constr
-val subst_univs_fn_puniverses : universe_level_subst_fn ->
- 'a puniverses -> 'a puniverses
-
-val subst_univs_constr : universe_subst -> constr -> constr
-
(** Level substitutions for polymorphism. *)
val subst_univs_level_constr : universe_level_subst -> constr -> constr
diff --git a/pretyping/unification.ml b/pretyping/unification.ml
index b41fb4e4d..8df8f8474 100644
--- a/pretyping/unification.ml
+++ b/pretyping/unification.ml
@@ -6,8 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-module CVars = Vars
-
open CErrors
open Pp
open Util
@@ -1527,7 +1525,7 @@ let indirectly_dependent sigma c d decls =
let finish_evar_resolution ?(flags=Pretyping.all_and_fail_flags) env current_sigma (pending,c) =
let sigma = Pretyping.solve_remaining_evars flags env current_sigma pending in
let sigma, subst = nf_univ_variables sigma in
- (sigma, EConstr.of_constr (CVars.subst_univs_constr subst (EConstr.Unsafe.to_constr (nf_evar sigma c))))
+ (sigma, EConstr.of_constr (Universes.subst_univs_constr subst (EConstr.Unsafe.to_constr (nf_evar sigma c))))
let default_matching_core_flags sigma =
let ts = Names.full_transparent_state in {
@@ -1617,7 +1615,7 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) =
| Some (sigma,_,l) ->
let c = applist (nf_evar sigma (local_strong whd_meta sigma c), l) in
let univs, subst = nf_univ_variables sigma in
- Some (sigma,EConstr.of_constr (CVars.subst_univs_constr subst (EConstr.Unsafe.to_constr c))))
+ Some (sigma,EConstr.of_constr (Universes.subst_univs_constr subst (EConstr.Unsafe.to_constr c))))
let make_eq_test env evd c =
let out cstr =
diff --git a/tactics/ind_tables.ml b/tactics/ind_tables.ml
index e1bf32f3c..bc2fea2bd 100644
--- a/tactics/ind_tables.ml
+++ b/tactics/ind_tables.ml
@@ -121,8 +121,7 @@ let define internal id c p univs =
let fd = declare_constant ~internal in
let id = compute_name internal id in
let ctx = Evd.normalize_evar_universe_context univs in
- let c = Vars.subst_univs_fn_constr
- (Universes.make_opt_subst (Evd.evar_universe_context_subst ctx)) c in
+ let c = Universes.subst_opt_univs_constr (Evd.evar_universe_context_subst ctx) c in
let univs =
if p then Polymorphic_const_entry (UState.context ctx)
else Monomorphic_const_entry (UState.context_set ctx)