aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Matthieu Sozeau <matthieu.sozeau@inria.fr>2015-09-23 16:00:49 +0200
committerGravatar Matthieu Sozeau <mattam@mattam.org>2015-10-02 15:54:10 +0200
commitcc69a4697633e14fc00c9bd0858b38120645464b (patch)
tree6a589f6916cf85bbc008ef3db788db8b90056278
parentff4d0d98ab0e0e81bd1acf9a7bf4b64913834911 (diff)
Univs: module constraints move to universe contexts as they might
declare new universes (e.g. with).
-rw-r--r--kernel/declarations.mli4
-rw-r--r--kernel/environ.ml11
-rw-r--r--kernel/mod_typing.ml35
-rw-r--r--kernel/mod_typing.mli2
-rw-r--r--kernel/safe_typing.ml55
5 files changed, 56 insertions, 51 deletions
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 c433c0789..429aba4f7 100644
--- a/kernel/environ.ml
+++ b/kernel/environ.ml
@@ -197,8 +197,10 @@ let push_constraints_to_env (_,univs) env =
add_constraints univs env
let add_universes strict ctx g =
- let g = Array.fold_left (fun g v -> Univ.add_universe v strict g)
- g (Univ.Instance.to_array (Univ.UContext.instance ctx))
+ 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
@@ -206,8 +208,9 @@ 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 -> Univ.add_universe v strict g)
- (Univ.ContextSet.levels 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 =
diff --git a/kernel/mod_typing.ml b/kernel/mod_typing.ml
index 4f20e5f62..7da0958ea 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
@@ -75,30 +75,30 @@ let rec check_with_def env struc (idl,(c,ctx)) mp equiv =
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 ctxs = Univ.ContextSet.of_context ctx in
let env' = Environ.add_constraints cst env' in
- let c',cst = match cb.const_body with
+ let c',ctx' = 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
+ j.uj_val, Univ.ContextSet.add_constraints cst' ctxs
| 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
+ (* if cb.const_polymorphic then *)Univ.ContextSet.add_constraints cst' ctxs
+ (* else cst' +++ cst *)
in
c, cst
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 = Univ.ContextSet.to_context ctx' }
in
before@(lab,SFBconst(cb'))::after, c', ctx'
else
@@ -145,8 +145,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 +193,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 +206,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 +237,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 +255,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 +300,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 +308,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 +323,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/safe_typing.ml b/kernel/safe_typing.ml
index 55e767321..9417aa080 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,7 +221,7 @@ 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 Univ.ContextSet.t | Later of Univ.ContextSet.t Future.computation
let add_constraints cst senv =
match cst with
@@ -229,14 +229,14 @@ let add_constraints cst senv =
{senv with future_cst = fc :: senv.future_cst}
| Now cst ->
{ senv with
- env = Environ.add_constraints cst senv.env;
- univ = Univ.Constraint.union cst senv.univ }
+ env = Environ.push_context_set ~strict:true cst senv.env;
+ univ = Univ.ContextSet.union cst senv.univ }
let add_constraints_list cst senv =
List.fold_right add_constraints cst senv
-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 ctx = add_constraints (Now ctx)
+let push_context ctx = add_constraints (Now (Univ.ContextSet.of_context ctx))
let is_curmod_library senv =
match senv.modvariant with LIBRARY -> true | _ -> false
@@ -373,9 +373,9 @@ let labels_of_mib mib =
let globalize_constant_universes env cb =
if cb.const_polymorphic then
- [Now Univ.Constraint.empty]
+ [Now Univ.ContextSet.empty]
else
- let cstrs = Univ.UContext.constraints cb.const_universes in
+ let cstrs = Univ.ContextSet.of_context cb.const_universes in
Now cstrs ::
(match cb.const_body with
| (Undef _ | Def _) -> []
@@ -384,16 +384,14 @@ 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 c])
let globalize_mind_universes mb =
if mb.mind_polymorphic then
- [Now Univ.Constraint.empty]
+ [Now Univ.ContextSet.empty]
else
- [Now (Univ.UContext.constraints mb.mind_universes)]
+ [Now (Univ.ContextSet.of_context mb.mind_universes)]
let constraints_of_sfb env sfb =
match sfb with
@@ -617,8 +615,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 +639,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 +665,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
@@ -696,7 +694,8 @@ let add_include me is_module inl 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 (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 +706,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 +800,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 +856,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 (Univ.ContextSet.add_constraints c Univ.ContextSet.empty))
(* NB: The next old comment probably refers to [propagate_loads] above.