aboutsummaryrefslogtreecommitdiff
path: root/engine
diff options
context:
space:
mode:
Diffstat (limited to 'engine')
-rw-r--r--engine/uState.ml19
-rw-r--r--engine/univNames.ml17
-rw-r--r--engine/univNames.mli2
3 files changed, 28 insertions, 10 deletions
diff --git a/engine/uState.ml b/engine/uState.ml
index aa7ec63a6f..41905feab7 100644
--- a/engine/uState.ml
+++ b/engine/uState.ml
@@ -101,13 +101,21 @@ let context ctx = Univ.ContextSet.to_context ctx.uctx_local
let const_univ_entry ~poly uctx =
let open Entries in
- if poly then Polymorphic_const_entry (context uctx)
+ if poly then
+ let (binders, _) = uctx.uctx_names in
+ let uctx = context uctx in
+ let nas = UnivNames.compute_instance_binders (Univ.UContext.instance uctx) binders in
+ Polymorphic_const_entry (nas, uctx)
else Monomorphic_const_entry (context_set uctx)
(* does not support cumulativity since you need more info *)
let ind_univ_entry ~poly uctx =
let open Entries in
- if poly then Polymorphic_ind_entry (context uctx)
+ if poly then
+ let (binders, _) = uctx.uctx_names in
+ let uctx = context uctx in
+ let nas = UnivNames.compute_instance_binders (Univ.UContext.instance uctx) binders in
+ Polymorphic_ind_entry (nas, uctx)
else Monomorphic_ind_entry (context_set uctx)
let of_context_set ctx = { empty with uctx_local = ctx }
@@ -394,8 +402,11 @@ let check_univ_decl ~poly uctx decl =
let ctx =
let names = decl.univdecl_instance in
let extensible = decl.univdecl_extensible_instance in
- if poly
- then Entries.Polymorphic_const_entry (universe_context ~names ~extensible uctx)
+ if poly then
+ let (binders, _) = uctx.uctx_names in
+ let uctx = universe_context ~names ~extensible uctx in
+ let nas = UnivNames.compute_instance_binders (Univ.UContext.instance uctx) binders in
+ Entries.Polymorphic_const_entry (nas, uctx)
else
let () = check_universe_context_set ~names ~extensible uctx in
Entries.Monomorphic_const_entry uctx.uctx_local
diff --git a/engine/univNames.ml b/engine/univNames.ml
index a71f9c5736..a037e577c4 100644
--- a/engine/univNames.ml
+++ b/engine/univNames.ml
@@ -81,18 +81,23 @@ let ubinder_obj : GlobRef.t * Id.t list -> Libobject.obj =
discharge_function = discharge_ubinder;
rebuild_function = (fun x -> x); }
+let compute_instance_binders inst ubinders =
+ let revmap = Id.Map.fold (fun id lvl accu -> LMap.add lvl id accu) ubinders LMap.empty in
+ let map lvl =
+ try Name (LMap.find lvl revmap)
+ with Not_found -> Name (name_universe lvl)
+ in
+ Array.map_to_list map (Instance.to_array inst)
+
let register_universe_binders ref ubinders =
(** TODO: change the API to register a [Name.t list] instead. This is the last
part of the code that depends on the internal representation of names in
abstract contexts, but removing it requires quite a rework of the
callers. *)
let univs = AUContext.instance (Environ.universes_of_global (Global.env()) ref) in
- let revmap = Id.Map.fold (fun id lvl accu -> LMap.add lvl id accu) ubinders LMap.empty in
- let map lvl =
- try LMap.find lvl revmap
- with Not_found -> name_universe lvl
- in
- let ubinders = Array.map_to_list map (Instance.to_array univs) in
+ let ubinders = compute_instance_binders univs ubinders in
+ (** FIXME: the function above always generate names but this may change *)
+ let ubinders = List.map (function Name id -> id | Anonymous -> assert false) ubinders in
if not (List.is_empty ubinders) then Lib.add_anonymous_leaf (ubinder_obj (ref, ubinders))
type univ_name_list = Names.lname list
diff --git a/engine/univNames.mli b/engine/univNames.mli
index bd4062ade4..634db9581c 100644
--- a/engine/univNames.mli
+++ b/engine/univNames.mli
@@ -19,6 +19,8 @@ type universe_binders = Univ.Level.t Names.Id.Map.t
val empty_binders : universe_binders
+val compute_instance_binders : Instance.t -> universe_binders -> Names.Name.t list
+
val register_universe_binders : Names.GlobRef.t -> universe_binders -> unit
type univ_name_list = Names.lname list