Skip to content

Commit 303d2c0

Browse files
committed
[nametab] Introduce type of imperative nametabs, unify API
This is both a cleanup and a step towards pushing the state upwards, and handling the nametabs functionally. Related to rocq-prover#16746 and rocq-prover/rfcs#65
1 parent de8bbed commit 303d2c0

Some content is hidden

Large Commits have some content hidden by default. Use the searchbox below for content that may be hidden.

43 files changed

+524
-477
lines changed

dev/base_include

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -172,8 +172,13 @@ let constr_of_string s =
172172
open Declarations;;
173173
open Declareops;;
174174

175+
let locate_constant qid =
176+
match Nametab.GlobRef.locate qid with
177+
| GlobRef.ConstRef c -> c
178+
| _ -> raise Not_found
179+
175180
let constbody_of_string s =
176-
let b = Global.lookup_constant (Nametab.locate_constant (qualid_of_string s)) in
181+
let b = Global.lookup_constant (locate_constant (qualid_of_string s)) in
177182
Option.get (Global.body_of_constant_body Library.indirect_accessor b);;
178183

179184
(* Get the current goal *)

engine/namegen.ml

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -81,14 +81,14 @@ let is_imported_ref = let open GlobRef in function
8181

8282
let is_global id =
8383
try
84-
let ref = Nametab.locate (qualid_of_ident id) in
84+
let ref = Nametab.GlobRef.locate (qualid_of_ident id) in
8585
not (is_imported_ref ref)
8686
with Not_found ->
8787
false
8888

8989
let is_constructor id =
9090
try
91-
match Nametab.locate (qualid_of_ident id) with
91+
match Nametab.GlobRef.locate (qualid_of_ident id) with
9292
| GlobRef.ConstructRef _ -> true
9393
| _ -> false
9494
with Not_found ->
@@ -284,7 +284,7 @@ let visible_ids sigma (nenv, c) =
284284
let g = global_of_constr c in
285285
if not (GlobRef.Set_env.mem g gseen) then
286286
let gseen = GlobRef.Set_env.add g gseen in
287-
let ids = match Nametab.shortest_qualid_of_global Id.Set.empty g with
287+
let ids = match Nametab.GlobRef.shortest_qualid Id.Set.empty g with
288288
| short ->
289289
let dir, id = repr_qualid short in
290290
if DirPath.is_empty dir then Id.Set.add id ids else ids

engine/termops.ml

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -889,7 +889,7 @@ let free_rels_and_unqualified_refs sigma t =
889889
if not (GlobRef.Set_env.mem g gseen) then begin
890890
try
891891
let gseen = GlobRef.Set_env.add g gseen in
892-
let short = Nametab.shortest_qualid_of_global Id.Set.empty g in
892+
let short = Nametab.GlobRef.shortest_qualid Id.Set.empty g in
893893
let dir, id = Libnames.repr_qualid short in
894894
let ids = if DirPath.is_empty dir then Id.Set.add id ids else ids in
895895
(gseen, vseen, ids)

engine/univNames.ml

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -15,7 +15,8 @@ open Univ
1515
let qualid_of_level ctx l =
1616
match Level.name l with
1717
| Some qid ->
18-
(try Some (Nametab.shortest_qualid_of_universe ctx qid)
18+
let ctx = Id.Map.domain ctx in
19+
(try Some (Nametab.Universe.shortest_qualid ctx qid)
1920
with Not_found -> None)
2021
| None -> None
2122

interp/abbreviation.ml

Lines changed: 8 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -44,11 +44,11 @@ let toggle_abbreviation ~on ~use kn =
4444
| OnlyParsing | ParsingAndPrinting ->
4545
if on then
4646
begin
47-
Nametab.push_abbreviation (Nametab.Until 1) sp kn;
48-
Nametab.push_abbreviation (Nametab.Exactly 1) sp kn
47+
Nametab.Abbrev.push (Nametab.Until 1) sp kn;
48+
Nametab.Abbrev.push (Nametab.Exactly 1) sp kn
4949
end
5050
else
51-
Nametab.remove_abbreviation sp kn
51+
Nametab.Abbrev.remove sp kn
5252
end
5353

5454
let toggle_abbreviations ~on ~use filter =
@@ -57,23 +57,23 @@ let toggle_abbreviations ~on ~use filter =
5757
!abbrev_table ()
5858

5959
let load_abbreviation i ((sp,kn),(_local,abbrev)) =
60-
if Nametab.exists_cci sp then
60+
if Nametab.GlobRef.exists sp then
6161
user_err
6262
(Id.print (basename sp) ++ str " already exists.");
6363
add_abbreviation kn sp abbrev;
64-
Nametab.push_abbreviation (Nametab.Until i) sp kn
64+
Nametab.Abbrev.push (Nametab.Until i) sp kn
6565

6666
let is_alias_of_already_visible_name sp = function
6767
| _,NRef (ref,_) ->
68-
let (dir,id) = repr_qualid (Nametab.shortest_qualid_of_global Id.Set.empty ref) in
68+
let (dir,id) = repr_qualid (Nametab.GlobRef.shortest_qualid Id.Set.empty ref) in
6969
DirPath.is_empty dir && Id.equal id (basename sp)
7070
| _ ->
7171
false
7272

7373
let open_abbreviation i ((sp,kn),(_local,abbrev)) =
7474
let pat = abbrev.abbrev_pattern in
7575
if not (Int.equal i 1 && is_alias_of_already_visible_name sp pat) then begin
76-
Nametab.push_abbreviation (Nametab.Exactly i) sp kn;
76+
Nametab.Abbrev.push (Nametab.Exactly i) sp kn;
7777
if not abbrev.abbrev_onlyparsing then
7878
(* Redeclare it to be used as (short) name in case an other (distfix)
7979
notation was declared in between *)
@@ -114,7 +114,7 @@ let declare_abbreviation ~local ?(also_in_cases_pattern=true) deprecation id ~on
114114
in
115115
add_leaf (inAbbreviation id (local,abbrev))
116116

117-
let pr_abbreviation kn = pr_qualid (Nametab.shortest_qualid_of_abbreviation Id.Set.empty kn)
117+
let pr_abbreviation kn = pr_qualid (Nametab.Abbrev.shortest_qualid Id.Set.empty kn)
118118

119119
let warn_deprecated_abbreviation =
120120
Deprecation.create_warning ~object_name:"Notation" ~warning_name:"deprecated-syntactic-definition"

interp/constrextern.ml

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -200,7 +200,7 @@ let path_of_global = function
200200
| GlobRef.ConstructRef ((ind,n),p) -> Libnames.make_path (dirpath_of_modpath (MutInd.modpath ind)) (Id.of_string_soft ("<constructor:" ^ Label.to_string (MutInd.label ind) ^ ":" ^ string_of_int n ^ ":" ^ string_of_int (p+1) ^ ">"))
201201

202202
let default_extern_reference ?loc vars r =
203-
try Nametab.shortest_qualid_of_global ?loc vars r
203+
try Nametab.GlobRef.shortest_qualid ?loc vars r
204204
with Not_found when GlobRef.is_bound r -> qualid_of_path (path_of_global r)
205205

206206
let my_extern_reference = ref default_extern_reference
@@ -430,7 +430,7 @@ and apply_notation_to_pattern ?loc gr ((subst,substlist),(no_implicit,nb_to_drop
430430
match availability_of_entry_coercion custom InConstrEntrySomeLevel with
431431
| None -> raise No_match
432432
| Some coercion ->
433-
let qid = Nametab.shortest_qualid_of_abbreviation ?loc vars kn in
433+
let qid = Nametab.Abbrev.shortest_qualid ?loc vars kn in
434434
let l1 =
435435
List.rev_map (fun (c,(subentry,(scopt,scl))) ->
436436
extern_cases_pattern_in_scope (subentry,(scopt,scl@scopes)) vars c)
@@ -795,7 +795,7 @@ let extended_glob_local_binder_of_decl ?loc u = DAst.make ?loc (extended_glob_lo
795795

796796
(* this helper function is copied from notation.ml as it's not exported *)
797797
let qualid_of_ref n =
798-
n |> Coqlib.lib_ref |> Nametab.shortest_qualid_of_global Id.Set.empty
798+
n |> Coqlib.lib_ref |> Nametab.GlobRef.shortest_qualid Id.Set.empty
799799

800800
let q_infinity () = qualid_of_ref "num.float.infinity"
801801
let q_neg_infinity () = qualid_of_ref "num.float.neg_infinity"
@@ -1273,7 +1273,7 @@ and extern_notation inctx (custom,scopes as allscopes) vars t rules =
12731273
extern true (subentry,(scopt,scl@snd scopes)) (vars,uvars) c)
12741274
terms
12751275
in
1276-
let cf = Nametab.shortest_qualid_of_abbreviation ?loc vars kn in
1276+
let cf = Nametab.Abbrev.shortest_qualid ?loc vars kn in
12771277
let a = CRef (cf,None) in
12781278
let args = fill_arg_scopes args argsscopes allscopes in
12791279
let args = extern_args (extern true) (vars,uvars) args in

interp/constrintern.ml

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -130,7 +130,7 @@ let explain_bad_patterns_number n1 n2 =
130130

131131
let inductive_of_record s =
132132
let inductive = GlobRef.IndRef (s.Structure.name) in
133-
Nametab.shortest_qualid_of_global Id.Set.empty inductive
133+
Nametab.GlobRef.shortest_qualid Id.Set.empty inductive
134134

135135
let explain_field_not_a_projection field_id =
136136
pr_qualid field_id ++ str ": Not a projection"
@@ -797,7 +797,7 @@ let terms_of_binders bl =
797797
| PatVar (Name id) -> CRef (qualid_of_ident id, None)
798798
| PatVar (Anonymous) -> error_cannot_coerce_wildcard_term ?loc ()
799799
| PatCstr (c,l,_) ->
800-
let qid = qualid_of_path ?loc (Nametab.path_of_global (GlobRef.ConstructRef c)) in
800+
let qid = qualid_of_path ?loc (Nametab.GlobRef.path (GlobRef.ConstructRef c)) in
801801
let hole = CAst.make ?loc @@ CHole (None,IntroAnonymous,None) in
802802
let params = List.make (Inductiveops.inductive_nparams (Global.env()) (fst c)) hole in
803803
CAppExpl ((qid,None),params @ List.map term_of_pat l)) pt in
@@ -1164,7 +1164,7 @@ let intern_sort_name ~local_univs = function
11641164
match local with
11651165
| Some u -> GUniv u
11661166
| None ->
1167-
try GUniv (Univ.Level.make (Nametab.locate_universe qid))
1167+
try GUniv (Univ.Level.make (Nametab.Universe.locate qid))
11681168
with Not_found ->
11691169
if is_id && local_univs.unb_univs
11701170
then GLocalUniv (CAst.make ?loc:qid.loc (qualid_basename qid))
@@ -2815,7 +2815,7 @@ let interp_named_context_evars ?(program_mode=false) ?(impl_env=empty_internaliz
28152815
let known_universe_level_name evd lid =
28162816
try Evd.universe_of_name evd lid.CAst.v
28172817
with Not_found ->
2818-
let u = Nametab.locate_universe (Libnames.qualid_of_lident lid) in
2818+
let u = Nametab.Universe.locate (Libnames.qualid_of_lident lid) in
28192819
Univ.Level.make u
28202820

28212821
let known_glob_level evd = function

interp/dumpglob.ml

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -248,7 +248,7 @@ let add_glob_gen ?loc sp lib_dp ty =
248248

249249
let add_glob ?loc ref =
250250
if dump () then
251-
let sp = Nametab.path_of_global ref in
251+
let sp = Nametab.GlobRef.path ref in
252252
let lib_dp = Lib.library_part ref in
253253
let ty = type_of_global_ref ref in
254254
add_glob_gen ?loc sp lib_dp ty
@@ -259,7 +259,7 @@ let mp_of_kn kn =
259259

260260
let add_glob_kn ?loc kn =
261261
if dump () then
262-
let sp = Nametab.path_of_abbreviation kn in
262+
let sp = Nametab.Abbrev.path kn in
263263
let lib_dp = Names.ModPath.dp (mp_of_kn kn) in
264264
add_glob_gen ?loc sp lib_dp "abbrev"
265265

interp/implicit_quantifiers.ml

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -188,7 +188,7 @@ let implicit_application env ty =
188188
let ({CAst.v=(qid, _, _)} as clapp) = destClassAppExpl ty in
189189
if Libnames.idset_mem_qualid qid env then None
190190
else
191-
let gr = Nametab.locate qid in
191+
let gr = Nametab.GlobRef.locate qid in
192192
if Typeclasses.is_class gr then Some (clapp, gr) else None
193193
with Not_found -> None
194194
in

interp/modintern.ml

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -51,12 +51,12 @@ let lookup_module_or_modtype kind qid =
5151
let loc = qid.CAst.loc in
5252
try
5353
if kind == ModType then raise Not_found;
54-
let mp = Nametab.locate_module qid in
54+
let mp = Nametab.Module.locate qid in
5555
Dumpglob.dump_modref ?loc mp "modtype"; (mp,Module)
5656
with Not_found ->
5757
try
5858
if kind == Module then raise Not_found;
59-
let mp = Nametab.locate_modtype qid in
59+
let mp = Nametab.ModType.locate qid in
6060
Dumpglob.dump_modref ?loc mp "mod"; (mp,ModType)
6161
with Not_found as exn ->
6262
let _, info = Exninfo.capture exn in

0 commit comments

Comments
 (0)