| author | wenzelm | 
| Wed, 05 Jul 2023 14:33:32 +0200 | |
| changeset 78254 | 50a949d316d3 | 
| parent 77846 | 5ba68d3bd741 | 
| child 80599 | f8c070249502 | 
| permissions | -rw-r--r-- | 
| 20089 | 1  | 
(* Title: Pure/name.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
| 28965 | 4  | 
Names of basic logical entities (variables etc.).  | 
| 20089 | 5  | 
*)  | 
6  | 
||
7  | 
signature NAME =  | 
|
8  | 
sig  | 
|
| 24849 | 9  | 
val uu: string  | 
| 43683 | 10  | 
val uu_: string  | 
| 24849 | 11  | 
val aT: string  | 
| 20089 | 12  | 
val bound: int -> string  | 
13  | 
val is_bound: string -> bool  | 
|
14  | 
val internal: string -> string  | 
|
15  | 
val dest_internal: string -> string  | 
|
| 55948 | 16  | 
val is_internal: string -> bool  | 
| 55949 | 17  | 
val reject_internal: string * Position.T list -> unit  | 
| 20089 | 18  | 
val skolem: string -> string  | 
19  | 
val dest_skolem: string -> string  | 
|
| 55948 | 20  | 
val is_skolem: string -> bool  | 
| 55949 | 21  | 
val reject_skolem: string * Position.T list -> unit  | 
| 20089 | 22  | 
val clean_index: string * int -> string * int  | 
23  | 
val clean: string -> string  | 
|
24  | 
type context  | 
|
25  | 
val context: context  | 
|
| 74411 | 26  | 
val build_context: (context -> context) -> context  | 
| 20158 | 27  | 
val make_context: string list -> context  | 
| 20089 | 28  | 
val declare: string -> context -> context  | 
| 20158 | 29  | 
val is_declared: context -> string -> bool  | 
| 
43329
 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
30  | 
val invent: context -> string -> int -> string list  | 
| 
 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
31  | 
val invent_names: context -> string -> 'a list -> (string * 'a) list  | 
| 20089 | 32  | 
val invent_list: string list -> string -> int -> string list  | 
| 
43326
 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 
wenzelm 
parents: 
43324 
diff
changeset
 | 
33  | 
val variant: string -> context -> string * context  | 
| 20089 | 34  | 
val variant_list: string list -> string list -> string list  | 
| 56812 | 35  | 
val enforce_case: bool -> string -> string  | 
| 56811 | 36  | 
val desymbolize: bool option -> string -> string  | 
| 20089 | 37  | 
end;  | 
38  | 
||
39  | 
structure Name: NAME =  | 
|
40  | 
struct  | 
|
41  | 
||
| 24849 | 42  | 
(** common defaults **)  | 
43  | 
||
44  | 
val uu = "uu";  | 
|
| 43683 | 45  | 
val uu_ = "uu_";  | 
| 24849 | 46  | 
val aT = "'a";  | 
47  | 
||
48  | 
||
| 20089 | 49  | 
|
50  | 
(** special variable names **)  | 
|
51  | 
||
52  | 
(* encoded bounds *)  | 
|
53  | 
||
54  | 
(*names for numbered variables --  | 
|
55  | 
preserves order wrt. int_ord vs. string_ord, avoids allocating new strings*)  | 
|
56  | 
||
57  | 
val small_int = Vector.tabulate (1000, fn i =>  | 
|
58  | 
let val leading = if i < 10 then "00" else if i < 100 then "0" else ""  | 
|
59  | 
in ":" ^ leading ^ string_of_int i end);  | 
|
60  | 
||
61  | 
fun bound n =  | 
|
| 
77846
 
5ba68d3bd741
more operations, following Isabelle/ML conventions;
 
wenzelm 
parents: 
74411 
diff
changeset
 | 
62  | 
if n < 1000 then Vector.nth small_int n  | 
| 
 
5ba68d3bd741
more operations, following Isabelle/ML conventions;
 
wenzelm 
parents: 
74411 
diff
changeset
 | 
63  | 
else ":" ^ bound (n div 1000) ^ Vector.nth small_int (n mod 1000);  | 
| 20089 | 64  | 
|
65  | 
val is_bound = String.isPrefix ":";  | 
|
66  | 
||
67  | 
||
| 55949 | 68  | 
(* internal names -- NB: internal subsumes skolem *)  | 
| 20089 | 69  | 
|
70  | 
val internal = suffix "_";  | 
|
71  | 
val dest_internal = unsuffix "_";  | 
|
| 55948 | 72  | 
val is_internal = String.isSuffix "_";  | 
| 55949 | 73  | 
fun reject_internal (x, ps) =  | 
74  | 
  if is_internal x then error ("Bad name: " ^ quote x ^ Position.here_list ps) else ();
 | 
|
| 20089 | 75  | 
|
76  | 
val skolem = suffix "__";  | 
|
77  | 
val dest_skolem = unsuffix "__";  | 
|
| 55948 | 78  | 
val is_skolem = String.isSuffix "__";  | 
| 55949 | 79  | 
fun reject_skolem (x, ps) =  | 
80  | 
  if is_skolem x then error ("Bad name: " ^ quote x ^ Position.here_list ps) else ();
 | 
|
| 20089 | 81  | 
|
| 20099 | 82  | 
fun clean_index (x, i) =  | 
83  | 
(case try dest_internal x of  | 
|
84  | 
NONE => (x, i)  | 
|
85  | 
| SOME x' => clean_index (x', i + 1));  | 
|
| 20089 | 86  | 
|
87  | 
fun clean x = #1 (clean_index (x, 0));  | 
|
88  | 
||
89  | 
||
90  | 
||
91  | 
(** generating fresh names **)  | 
|
92  | 
||
93  | 
(* context *)  | 
|
94  | 
||
95  | 
datatype context =  | 
|
96  | 
Context of string option Symtab.table; (*declared names with latest renaming*)  | 
|
97  | 
||
| 20099 | 98  | 
fun declare x (Context tab) =  | 
99  | 
Context (Symtab.default (clean x, NONE) tab);  | 
|
100  | 
||
101  | 
fun declare_renaming (x, x') (Context tab) =  | 
|
102  | 
Context (Symtab.update (clean x, SOME (clean x')) tab);  | 
|
| 20089 | 103  | 
|
| 20158 | 104  | 
fun is_declared (Context tab) = Symtab.defined tab;  | 
| 20089 | 105  | 
fun declared (Context tab) = Symtab.lookup tab;  | 
106  | 
||
| 20121 | 107  | 
val context = Context Symtab.empty |> fold declare ["", "'"];  | 
| 74411 | 108  | 
|
109  | 
fun build_context (f: context -> context) = f context;  | 
|
110  | 
||
111  | 
val make_context = build_context o fold declare;  | 
|
| 20089 | 112  | 
|
113  | 
||
| 
43329
 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
114  | 
(* invent names *)  | 
| 20089 | 115  | 
|
| 
43329
 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
116  | 
fun invent ctxt =  | 
| 20089 | 117  | 
let  | 
118  | 
fun invs _ 0 = []  | 
|
119  | 
| invs x n =  | 
|
| 
43329
 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
120  | 
let val x' = Symbol.bump_string x  | 
| 
 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
121  | 
in if is_declared ctxt x then invs x' n else x :: invs x' (n - 1) end;  | 
| 20099 | 122  | 
in invs o clean end;  | 
| 20089 | 123  | 
|
| 
43329
 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
124  | 
fun invent_names ctxt x xs = invent ctxt x (length xs) ~~ xs;  | 
| 20198 | 125  | 
|
| 
43329
 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
126  | 
val invent_list = invent o make_context;  | 
| 20089 | 127  | 
|
128  | 
||
129  | 
(* variants *)  | 
|
130  | 
||
131  | 
(*makes a variant of a name distinct from already used names in a  | 
|
132  | 
context; preserves a suffix of underscores "_"*)  | 
|
| 
43326
 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 
wenzelm 
parents: 
43324 
diff
changeset
 | 
133  | 
fun variant name ctxt =  | 
| 20089 | 134  | 
let  | 
135  | 
fun vary x =  | 
|
136  | 
(case declared ctxt x of  | 
|
137  | 
NONE => x  | 
|
138  | 
| SOME x' => vary (Symbol.bump_string (the_default x x')));  | 
|
139  | 
||
| 20121 | 140  | 
val (x, n) = clean_index (name, 0);  | 
| 20089 | 141  | 
val (x', ctxt') =  | 
| 20158 | 142  | 
if not (is_declared ctxt x) then (x, declare x ctxt)  | 
| 20089 | 143  | 
else  | 
144  | 
let  | 
|
145  | 
val x0 = Symbol.bump_init x;  | 
|
146  | 
val x' = vary x0;  | 
|
147  | 
val ctxt' = ctxt  | 
|
| 21565 | 148  | 
|> x0 <> x' ? declare_renaming (x0, x')  | 
| 20089 | 149  | 
|> declare x';  | 
150  | 
in (x', ctxt') end;  | 
|
| 
43326
 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 
wenzelm 
parents: 
43324 
diff
changeset
 | 
151  | 
in (x' ^ replicate_string n "_", ctxt') end;  | 
| 20089 | 152  | 
|
| 
43326
 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 
wenzelm 
parents: 
43324 
diff
changeset
 | 
153  | 
fun variant_list used names = #1 (make_context used |> fold_map variant names);  | 
| 20089 | 154  | 
|
| 31013 | 155  | 
|
| 31035 | 156  | 
(* names conforming to typical requirements of identifiers in the world outside *)  | 
| 31013 | 157  | 
|
| 
69137
 
90fce429e1bc
Jenkins: run ocaml_setup
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
56812 
diff
changeset
 | 
158  | 
fun enforce_case' false cs =  | 
| 56812 | 159  | 
(if forall Symbol.is_ascii_upper cs then map else nth_map 0) Symbol.to_ascii_lower cs  | 
| 
69137
 
90fce429e1bc
Jenkins: run ocaml_setup
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
56812 
diff
changeset
 | 
160  | 
| enforce_case' true cs =  | 
| 56812 | 161  | 
nth_map 0 Symbol.to_ascii_upper cs;  | 
162  | 
||
163  | 
fun enforce_case upper = implode o enforce_case' upper o raw_explode;  | 
|
164  | 
||
| 56811 | 165  | 
fun desymbolize perhaps_upper "" =  | 
166  | 
if the_default false perhaps_upper then "X" else "x"  | 
|
167  | 
| desymbolize perhaps_upper s =  | 
|
| 31013 | 168  | 
let  | 
| 31035 | 169  | 
val xs as (x :: _) = Symbol.explode s;  | 
| 33547 | 170  | 
val ys =  | 
171  | 
if Symbol.is_ascii_letter x orelse Symbol.is_symbolic x then xs  | 
|
| 31013 | 172  | 
else "x" :: xs;  | 
173  | 
fun is_valid x =  | 
|
| 34307 | 174  | 
Symbol.is_ascii_letter x orelse Symbol.is_ascii_digit x;  | 
| 31013 | 175  | 
fun sep [] = []  | 
176  | 
| sep (xs as "_" :: _) = xs  | 
|
177  | 
| sep xs = "_" :: xs;  | 
|
| 31035 | 178  | 
        fun desep ("_" :: xs) = xs
 | 
179  | 
| desep xs = xs;  | 
|
| 33547 | 180  | 
fun desymb x xs =  | 
181  | 
if is_valid x then x :: xs  | 
|
| 31013 | 182  | 
else  | 
| 33547 | 183  | 
(case Symbol.decode x of  | 
| 
40627
 
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
 
wenzelm 
parents: 
34307 
diff
changeset
 | 
184  | 
Symbol.Sym name => "_" :: raw_explode name @ sep xs  | 
| 33547 | 185  | 
| _ => sep xs);  | 
| 56812 | 186  | 
val upper_lower = Option.map enforce_case' perhaps_upper |> the_default I;  | 
| 31035 | 187  | 
in fold_rev desymb ys [] |> desep |> upper_lower |> implode end;  | 
| 31013 | 188  | 
|
| 20089 | 189  | 
end;  |