aboutsummaryrefslogtreecommitdiff
path: root/engine/uState.ml
diff options
context:
space:
mode:
authorGaëtan Gilbert2018-10-18 13:08:23 +0200
committerGaëtan Gilbert2018-11-12 18:43:04 +0100
commitccf995fd843f14ae8dfaf18177be6c2494faea35 (patch)
tree25a3ba0347ac8ee654c1a03949fbd5148bea7b20 /engine/uState.ml
parenta44b97e3e7f4a302a5255ca9fc57ea0b81a36b7e (diff)
Automatically generate names for universes.
The names are `uXXX` with `XXX` some number avoiding collision. Note that there may be some collisions with polymorphic binders, eg something like ~~~ Set Universe Polymorphism. Section foo. Universe u0. Definition bar := Type. (* bar@{u0} = Type@{u0} but this isn't the section u0 *) Definition baz := Type@{u0}. (* this one is the section u0 *) Definition foobar := Eval compute in baz -> Type. (* Type@{u0} -> Type@{u0} but these aren't the same u0 *) ~~~ So maybe we should do a nametab lookup too. This is strictly a printing issue (polymorphic binder names have no other use). In the monomorphic case names are qualified by the parent definition so it should be fine (barring module/definition collision but we already have those). Note that there are still unnamed universes as they didn't go through UState (eg schemes).
Diffstat (limited to 'engine/uState.ml')
-rw-r--r--engine/uState.ml20
1 files changed, 19 insertions, 1 deletions
diff --git a/engine/uState.ml b/engine/uState.ml
index 41905feab7..15bd0335f4 100644
--- a/engine/uState.ml
+++ b/engine/uState.ml
@@ -148,7 +148,25 @@ let of_binders b =
in
{ ctx with uctx_names = b, rmap }
-let universe_binders ctx = fst ctx.uctx_names
+let invent_name (named,cnt) u =
+ let rec aux i =
+ let na = Id.of_string ("u"^(string_of_int i)) in
+ if Id.Map.mem na named then aux (i+1)
+ else Id.Map.add na u named, i+1
+ in
+ aux cnt
+
+let universe_binders ctx =
+ let open Univ in
+ let named, rev = ctx.uctx_names in
+ let named, _ = LSet.fold (fun u named ->
+ match LMap.find u rev with
+ | exception Not_found -> (* not sure if possible *) invent_name named u
+ | { uname = None } -> invent_name named u
+ | { uname = Some _ } -> named)
+ (ContextSet.levels ctx.uctx_local) (named, 0)
+ in
+ named
let instantiate_variable l b v =
try v := Univ.LMap.set l (Some b) !v