aboutsummaryrefslogtreecommitdiffhomepage
path: root/library
diff options
context:
space:
mode:
Diffstat (limited to 'library')
-rw-r--r--library/declaremods.ml150
-rw-r--r--library/declaremods.mli2
-rw-r--r--library/global.ml25
-rw-r--r--library/global.mli17
-rw-r--r--library/globnames.ml2
-rw-r--r--library/kindops.ml48
-rw-r--r--library/kindops.mli2
-rw-r--r--library/lib.ml107
-rw-r--r--library/lib.mli11
-rw-r--r--library/libnames.ml19
-rw-r--r--library/libnames.mli20
-rw-r--r--library/library.ml4
-rw-r--r--library/nametab.ml74
-rw-r--r--library/nametab.mli13
-rw-r--r--library/summary.ml203
-rw-r--r--library/summary.mli41
16 files changed, 418 insertions, 320 deletions
diff --git a/library/declaremods.ml b/library/declaremods.ml
index db83dafef..28a04252a 100644
--- a/library/declaremods.ml
+++ b/library/declaremods.ml
@@ -180,16 +180,16 @@ let compute_visibility exists i =
(** Iterate some function [iter_objects] on all components of a module *)
-let do_module exists iter_objects i dir mp sobjs kobjs =
- let prefix = (dir,(mp,DirPath.empty)) in
+let do_module exists iter_objects i obj_dir obj_mp sobjs kobjs =
+ let prefix = { obj_dir ; obj_mp; obj_sec = DirPath.empty } in
let dirinfo = DirModule prefix in
- consistency_checks exists dir dirinfo;
- Nametab.push_dir (compute_visibility exists i) dir dirinfo;
- ModSubstObjs.set mp sobjs;
+ consistency_checks exists obj_dir dirinfo;
+ Nametab.push_dir (compute_visibility exists i) obj_dir dirinfo;
+ ModSubstObjs.set obj_mp sobjs;
(* If we're not a functor, let's iter on the internal components *)
if sobjs_no_functor sobjs then begin
let objs = expand_sobjs sobjs in
- ModObjs.set mp (prefix,objs,kobjs);
+ ModObjs.set obj_mp (prefix,objs,kobjs);
iter_objects (i+1) prefix objs;
iter_objects (i+1) prefix kobjs
end
@@ -222,20 +222,20 @@ let cache_keep _ = anomaly (Pp.str "This module should not be cached!")
let load_keep i ((sp,kn),kobjs) =
(* Invariant : seg isn't empty *)
- let dir = dir_of_sp sp and mp = mp_of_kn kn in
- let prefix = (dir,(mp,DirPath.empty)) in
+ let obj_dir = dir_of_sp sp and obj_mp = mp_of_kn kn in
+ let prefix = { obj_dir ; obj_mp; obj_sec = DirPath.empty } in
let prefix',sobjs,kobjs0 =
- try ModObjs.get mp
+ try ModObjs.get obj_mp
with Not_found -> assert false (* a substobjs should already be loaded *)
in
assert (eq_op prefix' prefix);
assert (List.is_empty kobjs0);
- ModObjs.set mp (prefix,sobjs,kobjs);
+ ModObjs.set obj_mp (prefix,sobjs,kobjs);
Lib.load_objects i prefix kobjs
let open_keep i ((sp,kn),kobjs) =
- let dir = dir_of_sp sp and mp = mp_of_kn kn in
- let prefix = (dir,(mp,DirPath.empty)) in
+ let obj_dir = dir_of_sp sp and obj_mp = mp_of_kn kn in
+ let prefix = { obj_dir; obj_mp; obj_sec = DirPath.empty } in
Lib.open_objects i prefix kobjs
let in_modkeep : Lib.lib_objects -> obj =
@@ -284,9 +284,9 @@ let (in_modtype : substitutive_objects -> obj),
(** {6 Declaration of substitutive objects for Include} *)
let do_include do_load do_open i ((sp,kn),aobjs) =
- let dir = Libnames.dirpath sp in
- let mp = KerName.modpath kn in
- let prefix = (dir,(mp,DirPath.empty)) in
+ let obj_dir = Libnames.dirpath sp in
+ let obj_mp = KerName.modpath kn in
+ let prefix = { obj_dir; obj_mp; obj_sec = DirPath.empty } in
let o = expand_aobjs aobjs in
if do_load then Lib.load_objects i prefix o;
if do_open then Lib.open_objects i prefix o
@@ -442,23 +442,26 @@ let process_module_binding mbid me =
Objects in these parameters are also loaded.
Output is accumulated on top of [acc] (in reverse order). *)
-let intern_arg interp_modast acc (idl,(typ,ann)) =
+let intern_arg interp_modast (acc, cst) (idl,(typ,ann)) =
let inl = inl2intopt ann in
let lib_dir = Lib.library_dp() in
let env = Global.env() in
- let mty,_ = interp_modast env ModType typ in
+ let (mty, _, cst') = interp_modast env ModType typ in
+ let () = Global.push_context_set true cst' in
+ let env = Global.env () in
let sobjs = get_module_sobjs false env inl mty in
let mp0 = get_module_path mty in
- List.fold_left
- (fun acc (_,id) ->
- let dir = DirPath.make [id] in
- let mbid = MBId.make lib_dir id in
- let mp = MPbound mbid in
- let resolver = Global.add_module_parameter mbid mty inl in
- let sobjs = subst_sobjs (map_mp mp0 mp resolver) sobjs in
- do_module false Lib.load_objects 1 dir mp sobjs [];
- (mbid,mty,inl)::acc)
- acc idl
+ let fold acc (_, id) =
+ let dir = DirPath.make [id] in
+ let mbid = MBId.make lib_dir id in
+ let mp = MPbound mbid in
+ let resolver = Global.add_module_parameter mbid mty inl in
+ let sobjs = subst_sobjs (map_mp mp0 mp resolver) sobjs in
+ do_module false Lib.load_objects 1 dir mp sobjs [];
+ (mbid,mty,inl)::acc
+ in
+ let acc = List.fold_left fold acc idl in
+ (acc, Univ.ContextSet.union cst cst')
(** Process a list of declarations of functor parameters
(Id11 .. Id1n : Typ1)..(Idk1 .. Idkm : Typk)
@@ -472,7 +475,7 @@ let intern_arg interp_modast acc (idl,(typ,ann)) =
*)
let intern_args interp_modast params =
- List.fold_left (intern_arg interp_modast) [] params
+ List.fold_left (intern_arg interp_modast) ([], Univ.ContextSet.empty) params
(** {6 Auxiliary functions concerning subtyping checks} *)
@@ -524,13 +527,17 @@ let mk_funct_type env args seb0 =
(** Prepare the module type list for check of subtypes *)
let build_subtypes interp_modast env mp args mtys =
- List.map
- (fun (m,ann) ->
+ let (cst, ans) = List.fold_left_map
+ (fun cst (m,ann) ->
let inl = inl2intopt ann in
- let mte,_ = interp_modast env ModType m in
+ let mte, _, cst' = interp_modast env ModType m in
+ let env = Environ.push_context_set ~strict:true cst' env in
+ let cst = Univ.ContextSet.union cst cst' in
let mtb = Mod_typing.translate_modtype env mp inl ([],mte) in
- { mtb with mod_type = mk_funct_type env args mtb.mod_type })
- mtys
+ cst, { mtb with mod_type = mk_funct_type env args mtb.mod_type })
+ Univ.ContextSet.empty mtys
+ in
+ (ans, cst)
(** {6 Current module information}
@@ -563,21 +570,26 @@ module RawModOps = struct
let start_module interp_modast export id args res fs =
let mp = Global.start_module id in
- let arg_entries_r = intern_args interp_modast args in
+ let arg_entries_r, cst = intern_args interp_modast args in
+ let () = Global.push_context_set true cst in
let env = Global.env () in
- let res_entry_o, subtyps = match res with
+ let res_entry_o, subtyps, cst = match res with
| Enforce (res,ann) ->
let inl = inl2intopt ann in
- let mte,_ = interp_modast env ModType res in
+ let (mte, _, cst) = interp_modast env ModType res in
+ let env = Environ.push_context_set ~strict:true cst env in
(* We check immediately that mte is well-formed *)
- let _ = Mod_typing.translate_mse env None inl mte in
- Some (mte,inl), []
+ let _, _, _, cst' = Mod_typing.translate_mse env None inl mte in
+ let cst = Univ.ContextSet.union cst cst' in
+ Some (mte, inl), [], cst
| Check resl ->
- None, build_subtypes interp_modast env mp arg_entries_r resl
+ let typs, cst = build_subtypes interp_modast env mp arg_entries_r resl in
+ None, typs, cst
in
+ let () = Global.push_context_set true cst in
openmod_info := { cur_typ = res_entry_o; cur_typs = subtyps };
let prefix = Lib.start_module export id mp fs in
- Nametab.push_dir (Nametab.Until 1) (fst prefix) (DirOpenModule prefix);
+ Nametab.push_dir (Nametab.Until 1) (prefix.obj_dir) (DirOpenModule prefix);
mp
let end_module () =
@@ -622,25 +634,33 @@ let declare_module interp_modast id args res mexpr_o fs =
(* We simulate the beginning of an interactive module,
then we adds the module parameters to the global env. *)
let mp = Global.start_module id in
- let arg_entries_r = intern_args interp_modast args in
+ let arg_entries_r, cst = intern_args interp_modast args in
let params = mk_params_entry arg_entries_r in
let env = Global.env () in
- let mty_entry_o, subs, inl_res = match res with
+ let env = Environ.push_context_set ~strict:true cst env in
+ let mty_entry_o, subs, inl_res, cst' = match res with
| Enforce (mty,ann) ->
let inl = inl2intopt ann in
- let mte, _ = interp_modast env ModType mty in
+ let (mte, _, cst) = interp_modast env ModType mty in
+ let env = Environ.push_context_set ~strict:true cst env in
(* We check immediately that mte is well-formed *)
- let _ = Mod_typing.translate_mse env None inl mte in
- Some mte, [], inl
+ let _, _, _, cst' = Mod_typing.translate_mse env None inl mte in
+ let cst = Univ.ContextSet.union cst cst' in
+ Some mte, [], inl, cst
| Check mtys ->
- None, build_subtypes interp_modast env mp arg_entries_r mtys,
- default_inline ()
+ let typs, cst = build_subtypes interp_modast env mp arg_entries_r mtys in
+ None, typs, default_inline (), cst
in
- let mexpr_entry_o, inl_expr = match mexpr_o with
- | None -> None, default_inline ()
+ let env = Environ.push_context_set ~strict:true cst' env in
+ let cst = Univ.ContextSet.union cst cst' in
+ let mexpr_entry_o, inl_expr, cst' = match mexpr_o with
+ | None -> None, default_inline (), Univ.ContextSet.empty
| Some (mexpr,ann) ->
- Some (fst (interp_modast env Module mexpr)), inl2intopt ann
+ let (mte, _, cst) = interp_modast env Module mexpr in
+ Some mte, inl2intopt ann, cst
in
+ let env = Environ.push_context_set ~strict:true cst' env in
+ let cst = Univ.ContextSet.union cst cst' in
let entry = match mexpr_entry_o, mty_entry_o with
| None, None -> assert false (* No body, no type ... *)
| None, Some typ -> MType (params, typ)
@@ -659,6 +679,7 @@ let declare_module interp_modast id args res mexpr_o fs =
| None -> None
| _ -> inl_res
in
+ let () = Global.push_context_set true cst in
let mp_env,resolver = Global.add_module id entry inl in
(* Name consistency check : kernel vs. library *)
@@ -679,12 +700,14 @@ module RawModTypeOps = struct
let start_modtype interp_modast id args mtys fs =
let mp = Global.start_modtype id in
- let arg_entries_r = intern_args interp_modast args in
+ let arg_entries_r, cst = intern_args interp_modast args in
+ let () = Global.push_context_set true cst in
let env = Global.env () in
- let sub_mty_l = build_subtypes interp_modast env mp arg_entries_r mtys in
+ let sub_mty_l, cst = build_subtypes interp_modast env mp arg_entries_r mtys in
+ let () = Global.push_context_set true cst in
openmodtype_info := sub_mty_l;
let prefix = Lib.start_modtype id mp fs in
- Nametab.push_dir (Nametab.Until 1) (fst prefix) (DirOpenModtype prefix);
+ Nametab.push_dir (Nametab.Until 1) (prefix.obj_dir) (DirOpenModtype prefix);
mp
let end_modtype () =
@@ -708,14 +731,21 @@ let declare_modtype interp_modast id args mtys (mty,ann) fs =
(* We simulate the beginning of an interactive module,
then we adds the module parameters to the global env. *)
let mp = Global.start_modtype id in
- let arg_entries_r = intern_args interp_modast args in
+ let arg_entries_r, cst = intern_args interp_modast args in
+ let () = Global.push_context_set true cst in
let params = mk_params_entry arg_entries_r in
let env = Global.env () in
- let mte, _ = interp_modast env ModType mty in
+ let mte, _, cst = interp_modast env ModType mty in
+ let () = Global.push_context_set true cst in
+ let env = Global.env () in
(* We check immediately that mte is well-formed *)
- let _ = Mod_typing.translate_mse env None inl mte in
+ let _, _, _, cst = Mod_typing.translate_mse env None inl mte in
+ let () = Global.push_context_set true cst in
+ let env = Global.env () in
let entry = params, mte in
- let sub_mty_l = build_subtypes interp_modast env mp arg_entries_r mtys in
+ let sub_mty_l, cst = build_subtypes interp_modast env mp arg_entries_r mtys in
+ let () = Global.push_context_set true cst in
+ let env = Global.env () in
let sobjs = get_functor_sobjs false env inl entry in
let subst = map_mp (get_module_path (snd entry)) mp empty_delta_resolver in
let sobjs = subst_sobjs subst sobjs in
@@ -769,7 +799,9 @@ let type_of_incl env is_mod = function
let declare_one_include interp_modast (me_ast,annot) =
let env = Global.env() in
- let me,kind = interp_modast env ModAny me_ast in
+ let me, kind, cst = interp_modast env ModAny me_ast in
+ let () = Global.push_context_set true cst in
+ let env = Global.env () in
let is_mod = (kind == Module) in
let cur_mp = Lib.current_mp () in
let inl = inl2intopt annot in
@@ -947,7 +979,7 @@ let iter_all_segments f =
type 'modast module_interpretor =
Environ.env -> Misctypes.module_kind -> 'modast ->
- Entries.module_struct_entry * Misctypes.module_kind
+ Entries.module_struct_entry * Misctypes.module_kind * Univ.ContextSet.t
type 'modast module_params =
(Id.t Loc.located list * ('modast * inline)) list
diff --git a/library/declaremods.mli b/library/declaremods.mli
index 42e5f4b13..0df55f34d 100644
--- a/library/declaremods.mli
+++ b/library/declaremods.mli
@@ -13,7 +13,7 @@ open Vernacexpr
type 'modast module_interpretor =
Environ.env -> Misctypes.module_kind -> 'modast ->
- Entries.module_struct_entry * Misctypes.module_kind
+ Entries.module_struct_entry * Misctypes.module_kind * Univ.ContextSet.t
type 'modast module_params =
(Id.t Loc.located list * ('modast * inline)) list
diff --git a/library/global.ml b/library/global.ml
index 43097dc5d..ed847b7cd 100644
--- a/library/global.ml
+++ b/library/global.ml
@@ -8,7 +8,6 @@
open Names
open Environ
-open Decl_kinds
(** We introduce here the global environment of the system,
and we declare it as a synchronized table. *)
@@ -21,6 +20,7 @@ module GlobalSafeEnv : sig
val set_safe_env : Safe_typing.safe_environment -> unit
val join_safe_environment : ?except:Future.UUIDSet.t -> unit -> unit
val is_joined_environment : unit -> bool
+ val global_env_summary_tag : Safe_typing.safe_environment Summary.Dyn.tag
end = struct
@@ -31,9 +31,9 @@ let join_safe_environment ?except () =
let is_joined_environment () =
Safe_typing.is_joined_environment !global_env
-
-let () =
- Summary.declare_summary global_env_summary_name
+
+let global_env_summary_tag =
+ Summary.declare_summary_tag global_env_summary_name
{ Summary.freeze_function = (function
| `Yes -> join_safe_environment (); !global_env
| `No -> !global_env
@@ -52,6 +52,8 @@ let set_safe_env e = global_env := e
end
+let global_env_summary_tag = GlobalSafeEnv.global_env_summary_tag
+
let safe_env = GlobalSafeEnv.safe_env
let join_safe_environment ?except () =
GlobalSafeEnv.join_safe_environment ?except ()
@@ -79,7 +81,7 @@ 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 = globalize (Safe_typing.push_named_def d)
+let push_named_def d = globalize0 (Safe_typing.push_named_def d)
let add_constraints c = globalize0 (Safe_typing.add_constraints 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)
@@ -231,18 +233,7 @@ let universes_of_global env r =
let universes_of_global gr =
universes_of_global (env ()) gr
-(** Global universe names *)
-type universe_names =
- (polymorphic * Univ.Level.t) Id.Map.t * Id.t Univ.LMap.t
-
-let global_universes =
- Summary.ref ~name:"Global universe names"
- ((Id.Map.empty, Univ.LMap.empty) : universe_names)
-
-let global_universe_names () = !global_universes
-let set_global_universe_names s = global_universes := s
-
-let is_polymorphic r =
+let is_polymorphic r =
let env = env() in
match r with
| VarRef id -> false
diff --git a/library/global.mli b/library/global.mli
index 51fe53181..03bc945da 100644
--- a/library/global.mli
+++ b/library/global.mli
@@ -32,11 +32,11 @@ val set_typing_flags : Declarations.typing_flags -> unit
(** Variables, Local definitions, constants, inductive types *)
val push_named_assum : (Id.t * Constr.types * bool) Univ.in_universe_context_set -> unit
-val push_named_def : (Id.t * Safe_typing.private_constants Entries.definition_entry) -> Univ.ContextSet.t
+val push_named_def : (Id.t * Entries.section_def_entry) -> unit
val export_private_constants : in_section:bool ->
- Safe_typing.private_constants Entries.constant_entry ->
- unit Entries.constant_entry * Safe_typing.exported_private_constant list
+ Safe_typing.private_constants Entries.definition_entry ->
+ unit Entries.definition_entry * Safe_typing.exported_private_constant list
val add_constant :
DirPath.t -> Id.t -> Safe_typing.global_declaration -> Constant.t
@@ -44,7 +44,7 @@ val add_mind :
DirPath.t -> Id.t -> Entries.mutual_inductive_entry -> MutInd.t
(** Extra universe constraints *)
-val add_constraints : Univ.constraints -> unit
+val add_constraints : Univ.Constraint.t -> unit
val push_context : bool -> Univ.UContext.t -> unit
val push_context_set : bool -> Univ.ContextSet.t -> unit
@@ -102,13 +102,6 @@ val body_of_constant : Constant.t -> (Constr.constr * Univ.AUContext.t) option
val body_of_constant_body : Declarations.constant_body -> (Constr.constr * Univ.AUContext.t) option
(** Same as {!body_of_constant} but on {!Declarations.constant_body}. *)
-(** Global universe name <-> level mapping *)
-type universe_names =
- (Decl_kinds.polymorphic * Univ.Level.t) Id.Map.t * Id.t Univ.LMap.t
-
-val global_universe_names : unit -> universe_names
-val set_global_universe_names : universe_names -> unit
-
(** {6 Compiled libraries } *)
val start_library : DirPath.t -> ModPath.t
@@ -166,4 +159,4 @@ val current_dirpath : unit -> DirPath.t
val with_global : (Environ.env -> DirPath.t -> 'a Univ.in_universe_context_set) -> 'a
-val global_env_summary_name : string
+val global_env_summary_tag : Safe_typing.safe_environment Summary.Dyn.tag
diff --git a/library/globnames.ml b/library/globnames.ml
index 9d7ab2db8..a6e75fdb6 100644
--- a/library/globnames.ml
+++ b/library/globnames.ml
@@ -30,7 +30,7 @@ let eq_gr gr1 gr2 =
| IndRef kn1, IndRef kn2 -> eq_ind kn1 kn2
| ConstructRef kn1, ConstructRef kn2 -> eq_constructor kn1 kn2
| VarRef v1, VarRef v2 -> Id.equal v1 v2
- | _ -> false
+ | (ConstRef _ | IndRef _ | ConstructRef _ | VarRef _), _ -> false
let destVarRef = function VarRef ind -> ind | _ -> failwith "destVarRef"
let destConstRef = function ConstRef ind -> ind | _ -> failwith "destConstRef"
diff --git a/library/kindops.ml b/library/kindops.ml
index 882f62086..83985ce97 100644
--- a/library/kindops.ml
+++ b/library/kindops.ml
@@ -23,45 +23,13 @@ let string_of_theorem_kind = function
| Proposition -> "Proposition"
| Corollary -> "Corollary"
-let string_of_definition_kind def =
- let (locality, poly, kind) = def in
- let error () = CErrors.anomaly (Pp.str "Internal definition kind.") in
- match kind with
- | Definition ->
- begin match locality with
- | Discharge -> "Let"
- | Local -> "Local Definition"
- | Global -> "Definition"
- end
- | Example ->
- begin match locality with
- | Discharge -> error ()
- | Local -> "Local Example"
- | Global -> "Example"
- end
- | Coercion ->
- begin match locality with
- | Discharge -> error ()
- | Local -> "Local Coercion"
- | Global -> "Coercion"
- end
- | SubClass ->
- begin match locality with
- | Discharge -> error ()
- | Local -> "Local SubClass"
- | Global -> "SubClass"
- end
- | CanonicalStructure ->
- begin match locality with
- | Discharge -> error ()
- | Local -> error ()
- | Global -> "Canonical Structure"
- end
- | Instance ->
- begin match locality with
- | Discharge -> error ()
- | Local -> "Instance"
- | Global -> "Global Instance"
- end
+let string_of_definition_object_kind = function
+ | Definition -> "Definition"
+ | Example -> "Example"
+ | Coercion -> "Coercion"
+ | SubClass -> "SubClass"
+ | CanonicalStructure -> "Canonical Structure"
+ | Instance -> "Instance"
+ | Let -> "Let"
| (StructureComponent|Scheme|CoFixpoint|Fixpoint|IdentityCoercion|Method) ->
CErrors.anomaly (Pp.str "Internal definition kind.")
diff --git a/library/kindops.mli b/library/kindops.mli
index 77979c915..06f873e85 100644
--- a/library/kindops.mli
+++ b/library/kindops.mli
@@ -12,4 +12,4 @@ open Decl_kinds
val logical_kind_of_goal_kind : goal_object_kind -> logical_kind
val string_of_theorem_kind : theorem_kind -> string
-val string_of_definition_kind : definition_kind -> string
+val string_of_definition_object_kind : definition_object_kind -> string
diff --git a/library/lib.ml b/library/lib.ml
index 3dbb16c7b..971089c17 100644
--- a/library/lib.ml
+++ b/library/lib.ml
@@ -93,12 +93,16 @@ let segment_of_objects prefix =
sections, but on the contrary there are many constructions of section
paths based on the library path. *)
-let initial_prefix = default_library,(Names.ModPath.initial,Names.DirPath.empty)
+let initial_prefix = {
+ obj_dir = default_library;
+ obj_mp = ModPath.initial;
+ obj_sec = DirPath.empty;
+}
type lib_state = {
- comp_name : Names.DirPath.t option;
+ comp_name : DirPath.t option;
lib_stk : library_segment;
- path_prefix : Names.DirPath.t * (Names.ModPath.t * Names.DirPath.t);
+ path_prefix : object_prefix;
}
let initial_lib_state = {
@@ -115,10 +119,9 @@ let library_dp () =
(* [path_prefix] is a pair of absolute dirpath and a pair of current
module path and relative section path *)
-let cwd () = fst !lib_state.path_prefix
-let current_prefix () = snd !lib_state.path_prefix
-let current_mp () = fst (snd !lib_state.path_prefix)
-let current_sections () = snd (snd !lib_state.path_prefix)
+let cwd () = !lib_state.path_prefix.obj_dir
+let current_mp () = !lib_state.path_prefix.obj_mp
+let current_sections () = !lib_state.path_prefix.obj_sec
let sections_depth () = List.length (Names.DirPath.repr (current_sections ()))
let sections_are_opened () = not (Names.DirPath.is_empty (current_sections ()))
@@ -136,7 +139,7 @@ let make_path_except_section id =
Libnames.make_path (cwd_except_section ()) id
let make_kn id =
- let mp,dir = current_prefix () in
+ let mp, dir = current_mp (), current_sections () in
Names.KerName.make mp dir (Names.Label.of_id id)
let make_oname id = Libnames.make_oname !lib_state.path_prefix id
@@ -152,8 +155,11 @@ let recalc_path_prefix () =
lib_state := { !lib_state with path_prefix = recalc !lib_state.lib_stk }
let pop_path_prefix () =
- let dir,(mp,sec) = !lib_state.path_prefix in
- lib_state := { !lib_state with path_prefix = pop_dirpath dir, (mp, pop_dirpath sec)}
+ let op = !lib_state.path_prefix in
+ lib_state := { !lib_state
+ with path_prefix = { op with obj_dir = pop_dirpath op.obj_dir;
+ obj_sec = pop_dirpath op.obj_sec;
+ } }
let find_entry_p p =
let rec find = function
@@ -226,7 +232,7 @@ let add_anonymous_entry node =
add_entry (make_oname (anonymous_id ())) node
let add_leaf id obj =
- if Names.ModPath.equal (current_mp ()) Names.ModPath.initial then
+ if ModPath.equal (current_mp ()) ModPath.initial then
user_err Pp.(str "No session module started (use -top dir)");
let oname = make_oname id in
cache_object (oname,obj);
@@ -278,8 +284,8 @@ let current_mod_id () =
let start_mod is_type export id mp fs =
- let dir = add_dirpath_suffix (cwd ()) id in
- let prefix = dir,(mp,Names.DirPath.empty) in
+ let dir = add_dirpath_suffix (!lib_state.path_prefix.obj_dir) id in
+ let prefix = { obj_dir = dir; obj_mp = mp; obj_sec = Names.DirPath.empty } in
let exists =
if is_type then Nametab.exists_cci (make_path id)
else Nametab.exists_module dir
@@ -328,10 +334,10 @@ let contents_after sp = let (after,_,_) = split_lib sp in after
let start_compilation s mp =
if !lib_state.comp_name != None then
user_err Pp.(str "compilation unit is already started");
- if not (Names.DirPath.is_empty (current_sections ())) then
+ if not (Names.DirPath.is_empty (!lib_state.path_prefix.obj_sec)) then
user_err Pp.(str "some sections are already opened");
- let prefix = s, (mp, Names.DirPath.empty) in
- let () = add_anonymous_entry (CompilingLibrary prefix) in
+ let prefix = Libnames.{ obj_dir = s; obj_mp = mp; obj_sec = DirPath.empty } in
+ add_anonymous_entry (CompilingLibrary prefix);
lib_state := { !lib_state with comp_name = Some s;
path_prefix = prefix }
@@ -411,8 +417,11 @@ let find_opening_node id =
type variable_info = Context.Named.Declaration.t * Decl_kinds.binding_kind
type variable_context = variable_info list
-type abstr_info = variable_context * Univ.universe_level_subst * Univ.AUContext.t
-
+type abstr_info = {
+ abstr_ctx : variable_context;
+ abstr_subst : Univ.Instance.t;
+ abstr_uctx : Univ.AUContext.t;
+}
type abstr_list = abstr_info Names.Cmap.t * abstr_info Names.Mindmap.t
type secentry =
@@ -477,8 +486,12 @@ let add_section_replacement f g poly hyps =
let inst = Univ.UContext.instance ctx in
let subst, ctx = Univ.abstract_universes ctx in
let args = instance_from_variable_context (List.rev sechyps) in
- sectab := (vars,f (inst,args) exps,
- g (sechyps,subst,ctx) abs)::sl
+ let info = {
+ abstr_ctx = sechyps;
+ abstr_subst = subst;
+ abstr_uctx = ctx;
+ } in
+ sectab := (vars,f (inst,args) exps, g info abs) :: sl
let add_section_kn poly kn =
let f x (l1,l2) = (l1,Names.Mindmap.add kn x l2) in
@@ -496,12 +509,21 @@ let section_segment_of_constant con =
let section_segment_of_mutual_inductive kn =
Names.Mindmap.find kn (snd (pi3 (List.hd !sectab)))
-let variable_section_segment_of_reference = function
- | ConstRef con -> pi1 (section_segment_of_constant con)
- | IndRef (kn,_) | ConstructRef ((kn,_),_) ->
- pi1 (section_segment_of_mutual_inductive kn)
- | _ -> []
-
+let empty_segment = {
+ abstr_ctx = [];
+ abstr_subst = Univ.Instance.empty;
+ abstr_uctx = Univ.AUContext.empty;
+}
+
+let section_segment_of_reference = function
+| ConstRef c -> section_segment_of_constant c
+| IndRef (kn,_) | ConstructRef ((kn,_),_) ->
+ section_segment_of_mutual_inductive kn
+| VarRef _ -> empty_segment
+
+let variable_section_segment_of_reference gr =
+ (section_segment_of_reference gr).abstr_ctx
+
let section_instance = function
| VarRef id ->
let eq = function
@@ -522,15 +544,15 @@ let is_in_section ref =
(*************)
(* Sections. *)
let open_section id =
- let olddir,(mp,oldsec) = !lib_state.path_prefix in
- let dir = add_dirpath_suffix olddir id in
- let prefix = dir, (mp, add_dirpath_suffix oldsec id) in
- if Nametab.exists_section dir then
+ let opp = !lib_state.path_prefix in
+ let obj_dir = add_dirpath_suffix opp.obj_dir id in
+ let prefix = { obj_dir; obj_mp = opp.obj_mp; obj_sec = add_dirpath_suffix opp.obj_sec id } in
+ if Nametab.exists_section obj_dir then
user_err ~hdr:"open_section" (Id.print id ++ str " already exists.");
let fs = Summary.freeze_summaries ~marshallable:`No in
add_entry (make_oname id) (OpenedSection (prefix, fs));
(*Pushed for the lifetime of the section: removed by unfrozing the summary*)
- Nametab.push_dir (Nametab.Until 1) dir (DirOpenSection prefix);
+ Nametab.push_dir (Nametab.Until 1) obj_dir (DirOpenSection prefix);
lib_state := { !lib_state with path_prefix = prefix };
add_section ()
@@ -556,7 +578,7 @@ let close_section () =
in
let (secdecls,mark,before) = split_lib_at_opening oname in
lib_state := { !lib_state with lib_stk = before };
- let full_olddir = fst !lib_state.path_prefix in
+ let full_olddir = !lib_state.path_prefix.obj_dir in
pop_path_prefix ();
add_entry oname (ClosedSection (List.rev (mark::secdecls)));
let newdecls = List.map discharge_item secdecls in
@@ -596,10 +618,10 @@ let init () =
(* Misc *)
let mp_of_global = function
- |VarRef id -> current_mp ()
- |ConstRef cst -> Names.Constant.modpath cst
- |IndRef ind -> Names.ind_modpath ind
- |ConstructRef constr -> Names.constr_modpath constr
+ | VarRef id -> !lib_state.path_prefix.obj_mp
+ | ConstRef cst -> Names.Constant.modpath cst
+ | IndRef ind -> Names.ind_modpath ind
+ | ConstructRef constr -> Names.constr_modpath constr
let rec dp_of_mp = function
|Names.MPfile dp -> dp
@@ -648,15 +670,10 @@ let discharge_con cst =
let discharge_inductive (kn,i) =
(discharge_kn kn,i)
-let discharge_abstract_universe_context (_, subst, abs_ctx) auctx =
+let discharge_abstract_universe_context { abstr_subst = subst; abstr_uctx = abs_ctx } auctx =
let open Univ in
- let len = LMap.cardinal subst in
- let rec gen_subst i acc =
- if i < 0 then acc
- else
- let acc = LMap.add (Level.var i) (Level.var (i + len)) acc in
- gen_subst (pred i) acc
- in
- let subst = gen_subst (AUContext.size auctx - 1) subst in
+ let ainst = make_abstract_instance auctx in
+ let subst = Instance.append subst ainst in
+ let subst = make_instance_subst subst in
let auctx = Univ.subst_univs_level_abstract_universe_context subst auctx in
subst, AUContext.union abs_ctx auctx
diff --git a/library/lib.mli b/library/lib.mli
index 721e2896f..cf75d5f8c 100644
--- a/library/lib.mli
+++ b/library/lib.mli
@@ -153,13 +153,22 @@ val init : unit -> unit
(** {6 Section management for discharge } *)
type variable_info = Context.Named.Declaration.t * Decl_kinds.binding_kind
type variable_context = variable_info list
-type abstr_info = variable_context * Univ.universe_level_subst * Univ.AUContext.t
+type abstr_info = private {
+ abstr_ctx : variable_context;
+ (** Section variables of this prefix *)
+ abstr_subst : Univ.Instance.t;
+ (** Actual names of the abstracted variables *)
+ abstr_uctx : Univ.AUContext.t;
+ (** Universe quantification, same length as the substitution *)
+}
val instance_from_variable_context : variable_context -> Names.Id.t array
val named_of_variable_context : variable_context -> Context.Named.t
val section_segment_of_constant : Names.Constant.t -> abstr_info
val section_segment_of_mutual_inductive: Names.MutInd.t -> abstr_info
+val section_segment_of_reference : Globnames.global_reference -> abstr_info
+
val variable_section_segment_of_reference : Globnames.global_reference -> variable_context
val section_instance : Globnames.global_reference -> Univ.Instance.t * Names.Id.t array
diff --git a/library/libnames.ml b/library/libnames.ml
index 81878e72f..a471d8396 100644
--- a/library/libnames.ml
+++ b/library/libnames.ml
@@ -156,10 +156,15 @@ let qualid_of_dirpath dir =
type object_name = full_path * KerName.t
-type object_prefix = DirPath.t * (ModPath.t * DirPath.t)
+type object_prefix = {
+ obj_dir : DirPath.t;
+ obj_mp : ModPath.t;
+ obj_sec : DirPath.t;
+}
-let make_oname (dirpath,(mp,dir)) id =
- make_path dirpath id, KerName.make mp dir (Label.of_id id)
+(* let make_oname (dirpath,(mp,dir)) id = *)
+let make_oname { obj_dir; obj_mp; obj_sec } id =
+ make_path obj_dir id, KerName.make obj_mp obj_sec (Label.of_id id)
(* to this type are mapped DirPath.t's in the nametab *)
type global_dir_reference =
@@ -170,10 +175,10 @@ type global_dir_reference =
| DirClosedSection of DirPath.t
(* this won't last long I hope! *)
-let eq_op (d1, (mp1, p1)) (d2, (mp2, p2)) =
- DirPath.equal d1 d2 &&
- DirPath.equal p1 p2 &&
- ModPath.equal mp1 mp2
+let eq_op op1 op2 =
+ DirPath.equal op1.obj_dir op2.obj_dir &&
+ DirPath.equal op1.obj_sec op2.obj_sec &&
+ ModPath.equal op1.obj_mp op2.obj_mp
let eq_global_dir_reference r1 r2 = match r1, r2 with
| DirOpenModule op1, DirOpenModule op2 -> eq_op op1 op2
diff --git a/library/libnames.mli b/library/libnames.mli
index ed01163ee..71f542240 100644
--- a/library/libnames.mli
+++ b/library/libnames.mli
@@ -94,7 +94,25 @@ val qualid_of_ident : Id.t -> qualid
type object_name = full_path * KerName.t
-type object_prefix = DirPath.t * (ModPath.t * DirPath.t)
+(** Object prefix morally contains the "prefix" naming of an object to
+ be stored by [library], where [obj_dir] is the "absolute" path,
+ [obj_mp] is the current "module" prefix and [obj_sec] is the
+ "section" prefix.
+
+ Thus, for an object living inside [Module A. Section B.] the
+ prefix would be:
+
+ [ { obj_dir = "A.B"; obj_mp = "A"; obj_sec = "B" } ]
+
+ Note that both [obj_dir] and [obj_sec] are "paths" that is to say,
+ as opposed to [obj_mp] which is a single module name.
+
+ *)
+type object_prefix = {
+ obj_dir : DirPath.t;
+ obj_mp : ModPath.t;
+ obj_sec : DirPath.t;
+}
val eq_op : object_prefix -> object_prefix -> bool
diff --git a/library/library.ml b/library/library.ml
index 88470d121..868e26684 100644
--- a/library/library.ml
+++ b/library/library.ml
@@ -170,7 +170,7 @@ let register_loaded_library m =
let prefix = Nativecode.mod_uid_of_dirpath libname ^ "." in
let f = prefix ^ "cmo" in
let f = Dynlink.adapt_filename f in
- if not Coq_config.no_native_compiler then
+ if Coq_config.native_compiler then
Nativelib.link_library ~prefix ~dirname ~basename:f
in
let rec aux = function
@@ -738,7 +738,7 @@ let save_library_to ?todo dir f otab =
System.marshal_out_segment f' ch (opaque_table : seg_proofs);
close_out ch;
(* Writing native code files *)
- if !Flags.native_compiler then
+ if !Flags.output_native_objects then
let fn = Filename.dirname f'^"/"^Nativecode.mod_uid_of_dirpath dir in
if not (Nativelib.compile_library dir ast fn) then
user_err Pp.(str "Could not compile the library to native code.")
diff --git a/library/nametab.ml b/library/nametab.ml
index 0ec4a37cd..08881d6d7 100644
--- a/library/nametab.ml
+++ b/library/nametab.ml
@@ -81,8 +81,9 @@ struct
Module F (X : S). Module X.
The argument X of the functor F is masked by the inner module X.
*)
- let masking_absolute n =
- Flags.if_verbose Feedback.msg_info (str ("Trying to mask the absolute name \"" ^ U.to_string n ^ "\"!"))
+ let warn_masking_absolute =
+ CWarnings.create ~name:"masking-absolute-name" ~category:"deprecated"
+ (fun n -> str ("Trying to mask the absolute name \"" ^ U.to_string n ^ "\"!"))
type user_name = U.t
@@ -121,7 +122,7 @@ struct
| Absolute (n,_) ->
(* This is an absolute name, we must keep it
otherwise it may become unaccessible forever *)
- masking_absolute n; tree.path
+ warn_masking_absolute n; tree.path
| Nothing
| Relative _ -> Relative (uname,o)
else tree.path
@@ -154,7 +155,7 @@ let rec push_exactly uname o level tree = function
| Absolute (n,_) ->
(* This is an absolute name, we must keep it
otherwise it may become unaccessible forever *)
- masking_absolute n; tree.path
+ warn_masking_absolute n; tree.path
| Nothing
| Relative _ -> Relative (uname,o)
in
@@ -302,6 +303,16 @@ module DirTab = Make(DirPath')(GlobDir)
type dirtab = DirTab.t
let the_dirtab = ref (DirTab.empty : dirtab)
+type universe_id = DirPath.t * int
+
+module UnivIdEqual =
+struct
+ type t = universe_id
+ let equal (d, i) (d', i') = DirPath.equal d d' && Int.equal i i'
+end
+module UnivTab = Make(FullPath)(UnivIdEqual)
+type univtab = UnivTab.t
+let the_univtab = ref (UnivTab.empty : univtab)
(* Reversed name tables ***************************************************)
@@ -318,6 +329,21 @@ let the_modrevtab = ref (MPmap.empty : mprevtab)
type mptrevtab = full_path MPmap.t
let the_modtyperevtab = ref (MPmap.empty : mptrevtab)
+module UnivIdOrdered =
+struct
+ type t = universe_id
+ let hash (d, i) = i + DirPath.hash d
+ let compare (d, i) (d', i') =
+ let c = Int.compare i i' in
+ if Int.equal c 0 then DirPath.compare d d'
+ else c
+end
+
+module UnivIdMap = HMap.Make(UnivIdOrdered)
+
+type univrevtab = full_path UnivIdMap.t
+let the_univrevtab = ref (UnivIdMap.empty : univrevtab)
+
(* Push functions *********************************************************)
(* This is for permanent constructions (never discharged -- but with
@@ -359,9 +385,14 @@ let push_modtype vis sp kn =
let push_dir vis dir dir_ref =
the_dirtab := DirTab.push vis dir dir_ref !the_dirtab;
match dir_ref with
- DirModule (_,(mp,_)) -> the_modrevtab := MPmap.add mp dir !the_modrevtab
- | _ -> ()
+ | DirModule { obj_mp; _ } -> the_modrevtab := MPmap.add obj_mp dir !the_modrevtab
+ | _ -> ()
+
+(* This is for global universe names *)
+let push_universe vis sp univ =
+ the_univtab := UnivTab.push vis sp univ !the_univtab;
+ the_univrevtab := UnivIdMap.add univ sp !the_univrevtab
(* Locate functions *******************************************************)
@@ -382,21 +413,23 @@ let locate_syndef qid = match locate_extended qid with
let locate_modtype qid = MPTab.locate qid !the_modtypetab
let full_name_modtype qid = MPTab.user_name qid !the_modtypetab
+let locate_universe qid = UnivTab.locate qid !the_univtab
+
let locate_dir qid = DirTab.locate qid !the_dirtab
let locate_module qid =
match locate_dir qid with
- | DirModule (_,(mp,_)) -> mp
+ | DirModule { obj_mp ; _} -> obj_mp
| _ -> raise Not_found
let full_name_module qid =
match locate_dir qid with
- | DirModule (dir,_) -> dir
+ | DirModule { obj_dir ; _} -> obj_dir
| _ -> raise Not_found
let locate_section qid =
match locate_dir qid with
- | DirOpenSection (dir, _)
+ | DirOpenSection { obj_dir; _ } -> obj_dir
| DirClosedSection dir -> dir
| _ -> raise Not_found
@@ -447,6 +480,8 @@ let exists_module = exists_dir
let exists_modtype sp = MPTab.exists sp !the_modtypetab
+let exists_universe kn = UnivTab.exists kn !the_univtab
+
(* Reverse locate functions ***********************************************)
let path_of_global ref =
@@ -469,6 +504,9 @@ let dirpath_of_module mp =
let path_of_modtype mp =
MPmap.find mp !the_modtyperevtab
+let path_of_universe mp =
+ UnivIdMap.find mp !the_univrevtab
+
(* Shortest qualid functions **********************************************)
let shortest_qualid_of_global ctx ref =
@@ -490,6 +528,10 @@ let shortest_qualid_of_modtype kn =
let sp = MPmap.find kn !the_modtyperevtab in
MPTab.shortest_qualid Id.Set.empty sp !the_modtypetab
+let shortest_qualid_of_universe kn =
+ let sp = UnivIdMap.find kn !the_univrevtab in
+ UnivTab.shortest_qualid Id.Set.empty sp !the_univtab
+
let pr_global_env env ref =
try pr_qualid (shortest_qualid_of_global env ref)
with Not_found as e ->
@@ -508,24 +550,28 @@ let global_inductive r =
(********************************************************************)
(* Registration of tables as a global table and rollback *)
-type frozen = ccitab * dirtab * mptab
- * globrevtab * mprevtab * mptrevtab
+type frozen = ccitab * dirtab * mptab * univtab
+ * globrevtab * mprevtab * mptrevtab * univrevtab
let freeze _ : frozen =
!the_ccitab,
!the_dirtab,
!the_modtypetab,
+ !the_univtab,
!the_globrevtab,
!the_modrevtab,
- !the_modtyperevtab
+ !the_modtyperevtab,
+ !the_univrevtab
-let unfreeze (ccit,dirt,mtyt,globr,modr,mtyr) =
+let unfreeze (ccit,dirt,mtyt,univt,globr,modr,mtyr,univr) =
the_ccitab := ccit;
the_dirtab := dirt;
the_modtypetab := mtyt;
+ the_univtab := univt;
the_globrevtab := globr;
the_modrevtab := modr;
- the_modtyperevtab := mtyr
+ the_modtyperevtab := mtyr;
+ the_univrevtab := univr
let _ =
Summary.declare_summary "names"
diff --git a/library/nametab.mli b/library/nametab.mli
index c02447a7c..77fafa100 100644
--- a/library/nametab.mli
+++ b/library/nametab.mli
@@ -78,6 +78,12 @@ val push_modtype : visibility -> full_path -> ModPath.t -> unit
val push_dir : visibility -> DirPath.t -> global_dir_reference -> unit
val push_syndef : visibility -> full_path -> syndef_name -> unit
+type universe_id = DirPath.t * int
+
+module UnivIdMap : CMap.ExtS with type key = universe_id
+
+val push_universe : visibility -> full_path -> universe_id -> unit
+
(** {6 The following functions perform globalization of qualified names } *)
(** These functions globalize a (partially) qualified name or fail with
@@ -91,6 +97,7 @@ val locate_modtype : qualid -> ModPath.t
val locate_dir : qualid -> global_dir_reference
val locate_module : qualid -> ModPath.t
val locate_section : qualid -> DirPath.t
+val locate_universe : qualid -> universe_id
(** These functions globalize user-level references into global
references, like [locate] and co, but raise a nice error message
@@ -119,6 +126,7 @@ val exists_modtype : full_path -> bool
val exists_dir : DirPath.t -> bool
val exists_section : DirPath.t -> bool (** deprecated synonym of [exists_dir] *)
val exists_module : DirPath.t -> bool (** deprecated synonym of [exists_dir] *)
+val exists_universe : full_path -> bool
(** {6 These functions locate qualids into full user names } *)
@@ -138,6 +146,10 @@ val path_of_global : global_reference -> full_path
val dirpath_of_module : ModPath.t -> DirPath.t
val path_of_modtype : ModPath.t -> full_path
+(** A universe_id might not be registered with a corresponding user name.
+ @raise Not_found if the universe was not introduced by the user. *)
+val path_of_universe : universe_id -> full_path
+
(** Returns in particular the dirpath or the basename of the full path
associated to global reference *)
@@ -158,6 +170,7 @@ val shortest_qualid_of_global : Id.Set.t -> global_reference -> qualid
val shortest_qualid_of_syndef : Id.Set.t -> syndef_name -> qualid
val shortest_qualid_of_modtype : ModPath.t -> qualid
val shortest_qualid_of_module : ModPath.t -> qualid
+val shortest_qualid_of_universe : universe_id -> qualid
(** Deprecated synonyms *)
diff --git a/library/summary.ml b/library/summary.ml
index 9f49d1f83..6df17476b 100644
--- a/library/summary.ml
+++ b/library/summary.ml
@@ -13,17 +13,22 @@ open Util
module Dyn = Dyn.Make ()
type marshallable = [ `Yes | `No | `Shallow ]
+
type 'a summary_declaration = {
freeze_function : marshallable -> 'a;
unfreeze_function : 'a -> unit;
init_function : unit -> unit }
-let summaries = ref Int.Map.empty
+let sum_mod = ref None
+let sum_map = ref String.Map.empty
let mangle id = id ^ "-SUMMARY"
+let unmangle id = String.(sub id 0 (length id - 8))
+
+let ml_modules = "ML-MODULES"
-let internal_declare_summary hash sumname sdecl =
- let (infun, outfun) = Dyn.Easy.make_dyn (mangle sumname) in
+let internal_declare_summary fadd sumname sdecl =
+ let infun, outfun, tag = Dyn.Easy.make_dyn_tag (mangle sumname) in
let dyn_freeze b = infun (sdecl.freeze_function b)
and dyn_unfreeze sum = sdecl.unfreeze_function (outfun sum)
and dyn_init = sdecl.init_function in
@@ -32,140 +37,116 @@ let internal_declare_summary hash sumname sdecl =
unfreeze_function = dyn_unfreeze;
init_function = dyn_init }
in
- summaries := Int.Map.add hash (sumname, ddecl) !summaries
+ fadd sumname ddecl;
+ tag
-let all_declared_summaries = ref Int.Set.empty
+let declare_ml_modules_summary decl =
+ let ml_add _ ddecl = sum_mod := Some ddecl in
+ internal_declare_summary ml_add ml_modules decl
-let summary_names = ref []
-let name_of_summary name =
- try List.assoc name !summary_names
- with Not_found -> "summary name not found"
+let declare_ml_modules_summary decl =
+ ignore(declare_ml_modules_summary decl)
-let declare_summary sumname decl =
- let hash = String.hash sumname in
- let () = if Int.Map.mem hash !summaries then
- let (name, _) = Int.Map.find hash !summaries in
- anomaly ~label:"Summary.declare_summary"
- (str "Colliding summary names: " ++ str sumname ++ str " vs. " ++ str name ++ str ".")
+let declare_summary_tag sumname decl =
+ let fadd name ddecl = sum_map := String.Map.add name ddecl !sum_map in
+ let () = if String.Map.mem sumname !sum_map then
+ anomaly ~label:"Summary.declare_summary"
+ (str "Colliding summary names: " ++ str sumname ++ str " vs. " ++ str sumname ++ str ".")
in
- all_declared_summaries := Int.Set.add hash !all_declared_summaries;
- summary_names := (hash, sumname) :: !summary_names;
- internal_declare_summary hash sumname decl
+ internal_declare_summary fadd sumname decl
+
+let declare_summary sumname decl =
+ ignore(declare_summary_tag sumname decl)
type frozen = {
- summaries : (int * Dyn.t) list;
+ summaries : Dyn.t String.Map.t;
(** Ordered list w.r.t. the first component. *)
ml_module : Dyn.t option;
(** Special handling of the ml_module summary. *)
}
-let empty_frozen = { summaries = []; ml_module = None; }
-
-let ml_modules = "ML-MODULES"
-let ml_modules_summary = String.hash ml_modules
+let empty_frozen = { summaries = String.Map.empty; ml_module = None }
let freeze_summaries ~marshallable : frozen =
- let fold id (_, decl) accu =
- (* to debug missing Lazy.force
- if marshallable <> `No then begin
- let id, _ = Int.Map.find id !summaries in
- prerr_endline ("begin marshalling " ^ id);
- ignore(Marshal.to_string (decl.freeze_function marshallable) []);
- prerr_endline ("end marshalling " ^ id);
- end;
- /debug *)
- let state = decl.freeze_function marshallable in
- if Int.equal id ml_modules_summary then { accu with ml_module = Some state }
- else { accu with summaries = (id, state) :: accu.summaries }
+ let smap decl = decl.freeze_function marshallable in
+ { summaries = String.Map.map smap !sum_map;
+ ml_module = Option.map (fun decl -> decl.freeze_function marshallable) !sum_mod;
+ }
+
+let unfreeze_single name state =
+ let decl =
+ try String.Map.find name !sum_map
+ with
+ | Not_found ->
+ CErrors.anomaly Pp.(str "trying to unfreeze unregistered summary " ++ str name)
in
- Int.Map.fold_right fold !summaries empty_frozen
-
-let unfreeze_summaries fs =
+ try decl.unfreeze_function state
+ with e when CErrors.noncritical e ->
+ let e = CErrors.push e in
+ Feedback.msg_warning
+ Pp.(seq [str "Error unfreezing summary "; str name; fnl (); CErrors.iprint e]);
+ iraise e
+
+let unfreeze_summaries ?(partial=false) { summaries; ml_module } =
(* The unfreezing of [ml_modules_summary] has to be anticipated since it
- * may modify the content of [summaries] ny loading new ML modules *)
- let (_, decl) =
- try Int.Map.find ml_modules_summary !summaries
- with Not_found -> anomaly (str "Undeclared summary " ++ str ml_modules ++ str ".")
- in
- let () = match fs.ml_module with
+ * may modify the content of [summaries] by loading new ML modules *)
+ begin match !sum_mod with
| None -> anomaly (str "Undeclared summary " ++ str ml_modules ++ str ".")
- | Some state -> decl.unfreeze_function state
- in
- let fold id (_, decl) states =
- if Int.equal id ml_modules_summary then states
- else match states with
- | [] ->
- let () = decl.init_function () in
- []
- | (nid, state) :: rstates ->
- if Int.equal id nid then
- let () = decl.unfreeze_function state in rstates
- else
- let () = decl.init_function () in states
+ | Some decl -> Option.iter (fun state -> decl.unfreeze_function state) ml_module
+ end;
+ (** We must be independent on the order of the map! *)
+ let ufz name decl =
+ try decl.unfreeze_function String.Map.(find name summaries)
+ with Not_found ->
+ if not partial then begin
+ Feedback.msg_warning Pp.(str "Summary was captured out of module scope for entry " ++ str name);
+ decl.init_function ()
+ end;
in
- let fold id decl state =
- try fold id decl state
- with e when CErrors.noncritical e ->
- let e = CErrors.push e in
- Feedback.msg_error
- Pp.(seq [str "Error unfreezing summary %s\n%s\n%!";
- str (name_of_summary id);
- CErrors.iprint e]);
- iraise e
- in
- (** We rely on the order of the frozen list, and the order of folding *)
- ignore (Int.Map.fold_left fold !summaries fs.summaries)
+ (* String.Map.iter unfreeze_single !sum_map *)
+ String.Map.iter ufz !sum_map
let init_summaries () =
- Int.Map.iter (fun _ (_, decl) -> decl.init_function ()) !summaries
+ String.Map.iter (fun _ decl -> decl.init_function ()) !sum_map
(** For global tables registered statically before the end of coqtop
launch, the following empty [init_function] could be used. *)
let nop () = ()
-(** Selective freeze *)
+(** Summary projection *)
+let project_from_summary { summaries } tag =
+ let id = unmangle (Dyn.repr tag) in
+ let state = String.Map.find id summaries in
+ Option.get (Dyn.Easy.prj state tag)
+
+let modify_summary st tag v =
+ let id = unmangle (Dyn.repr tag) in
+ let summaries = String.Map.set id (Dyn.Easy.inj v tag) st.summaries in
+ {st with summaries}
-type frozen_bits = (int * Dyn.t) list
+let remove_from_summary st tag =
+ let id = unmangle (Dyn.repr tag) in
+ let summaries = String.Map.remove id st.summaries in
+ {st with summaries}
+
+(** Selective freeze *)
-let ids_of_string_list complement ids =
- if not complement then List.map String.hash ids
- else
- let fold accu id =
- let id = String.hash id in
- Int.Set.remove id accu
- in
- let ids = List.fold_left fold !all_declared_summaries ids in
- Int.Set.elements ids
+type frozen_bits = Dyn.t String.Map.t
let freeze_summary ~marshallable ?(complement=false) ids =
- let ids = ids_of_string_list complement ids in
- List.map (fun id ->
- let (_, summary) = Int.Map.find id !summaries in
- id, summary.freeze_function marshallable)
- ids
-
-let unfreeze_summary datas =
- List.iter
- (fun (id, data) ->
- let (name, summary) = Int.Map.find id !summaries in
- try summary.unfreeze_function data
- with e ->
- let e = CErrors.push e in
- prerr_endline ("Exception unfreezing " ^ name);
- iraise e)
- datas
+ let sub_map = String.Map.filter (fun id _ -> complement <> List.(mem id ids)) !sum_map in
+ String.Map.map (fun decl -> decl.freeze_function marshallable) sub_map
+
+let unfreeze_summary = String.Map.iter unfreeze_single
let surgery_summary { summaries; ml_module } bits =
- let summaries = List.map (fun (id, _ as orig) ->
- try id, List.assoc id bits
- with Not_found -> orig)
- summaries in
+ let summaries =
+ String.Map.fold (fun hash state sum -> String.Map.set hash state sum ) summaries bits in
{ summaries; ml_module }
let project_summary { summaries; ml_module } ?(complement=false) ids =
- let ids = ids_of_string_list complement ids in
- List.filter (fun (id, _) -> List.mem id ids) summaries
+ String.Map.filter (fun name _ -> complement <> List.(mem name ids)) summaries
let pointer_equal l1 l2 =
let ptr_equal d1 d2 =
@@ -174,19 +155,22 @@ let pointer_equal l1 l2 =
match Dyn.eq t1 t2 with
| None -> false
| Some Refl -> x1 == x2
- in
+ in
+ let l1, l2 = String.Map.bindings l1, String.Map.bindings l2 in
CList.for_all2eq
(fun (id1,v1) (id2,v2) -> id1 = id2 && ptr_equal v1 v2) l1 l2
(** All-in-one reference declaration + registration *)
-let ref ?(freeze=fun _ r -> r) ~name x =
+let ref_tag ?(freeze=fun _ r -> r) ~name x =
let r = ref x in
- declare_summary name
+ let tag = declare_summary_tag name
{ freeze_function = (fun b -> freeze b !r);
unfreeze_function = ((:=) r);
- init_function = (fun () -> r := x) };
- r
+ init_function = (fun () -> r := x) } in
+ r, tag
+
+let ref ?freeze ~name x = fst @@ ref_tag ?freeze ~name x
module Local = struct
@@ -198,8 +182,7 @@ let (!) r =
let key, name = !r in
try CEphemeron.get key
with CEphemeron.InvalidKey ->
- let _, { init_function } =
- Int.Map.find (String.hash (mangle name)) !summaries in
+ let { init_function } = String.Map.find name !sum_map in
init_function ();
CEphemeron.get (fst !r)
diff --git a/library/summary.mli b/library/summary.mli
index d093d95f2..09447199e 100644
--- a/library/summary.mli
+++ b/library/summary.mli
@@ -36,6 +36,12 @@ type 'a summary_declaration = {
val declare_summary : string -> 'a summary_declaration -> unit
+(** We provide safe projection from the summary to the types stored in
+ it.*)
+module Dyn : Dyn.S
+
+val declare_summary_tag : string -> 'a summary_declaration -> 'a Dyn.tag
+
(** All-in-one reference declaration + summary registration.
It behaves just as OCaml's standard [ref] function, except
that a [declare_summary] is done, with [name] as string.
@@ -43,6 +49,7 @@ val declare_summary : string -> 'a summary_declaration -> unit
The [freeze_function] can be overridden *)
val ref : ?freeze:(marshallable -> 'a -> 'a) -> name:string -> 'a -> 'a ref
+val ref_tag : ?freeze:(marshallable -> 'a -> 'a) -> name:string -> 'a -> 'a ref * 'a Dyn.tag
(* As [ref] but the value is local to a process, i.e. not sent to, say, proof
* workers. It is useful to implement a local cache for example. *)
@@ -55,10 +62,11 @@ module Local : sig
end
-(** Special name for the summary of ML modules. This summary entry is
- special because its unfreeze may load ML code and hence add summary
- entries. Thus is has to be recognizable, and handled appropriately *)
-val ml_modules : string
+(** Special summary for ML modules. This summary entry is special
+ because its unfreeze may load ML code and hence add summary
+ entries. Thus is has to be recognizable, and handled properly.
+ *)
+val declare_ml_modules_summary : 'a summary_declaration -> unit
(** For global tables registered statically before the end of coqtop
launch, the following empty [init_function] could be used. *)
@@ -72,19 +80,34 @@ type frozen
val empty_frozen : frozen
val freeze_summaries : marshallable:marshallable -> frozen
-val unfreeze_summaries : frozen -> unit
+val unfreeze_summaries : ?partial:bool -> frozen -> unit
val init_summaries : unit -> unit
-(** The type [frozen_bits] is a snapshot of some of the registered tables *)
+(** Typed projection of the summary. Experimental API, use with CARE *)
+
+val modify_summary : frozen -> 'a Dyn.tag -> 'a -> frozen
+val project_from_summary : frozen -> 'a Dyn.tag -> 'a
+val remove_from_summary : frozen -> 'a Dyn.tag -> frozen
+
+(** The type [frozen_bits] is a snapshot of some of the registered
+ tables. It is DEPRECATED in favor of the typed projection
+ version. *)
+
type frozen_bits
+[@@ocaml.deprecated "Please use the typed version of summary projection"]
-val freeze_summary :
- marshallable:marshallable -> ?complement:bool -> string list -> frozen_bits
+[@@@ocaml.warning "-3"]
+val freeze_summary : marshallable:marshallable -> ?complement:bool -> string list -> frozen_bits
+[@@ocaml.deprecated "Please use the typed version of summary projection"]
val unfreeze_summary : frozen_bits -> unit
+[@@ocaml.deprecated "Please use the typed version of summary projection"]
val surgery_summary : frozen -> frozen_bits -> frozen
+[@@ocaml.deprecated "Please use the typed version of summary projection"]
val project_summary : frozen -> ?complement:bool -> string list -> frozen_bits
+[@@ocaml.deprecated "Please use the typed version of summary projection"]
val pointer_equal : frozen_bits -> frozen_bits -> bool
+[@@ocaml.deprecated "Please use the typed version of summary projection"]
+[@@@ocaml.warning "+3"]
(** {6 Debug} *)
-
val dump : unit -> (int * string) list