| author | desharna | 
| Fri, 29 Sep 2023 15:27:43 +0200 | |
| changeset 81528 | e4b4141e6954 | 
| parent 77967 | 6bb2f9b32804 | 
| permissions | -rw-r--r-- | 
| 52265 | 1  | 
(* Title: Tools/Code/code_symbol.ML  | 
| 52136 | 2  | 
Author: Florian Haftmann, TU Muenchen  | 
3  | 
||
4  | 
Data related to symbols in programs: constants, type constructors, classes,  | 
|
5  | 
class relations, class instances, modules.  | 
|
6  | 
*)  | 
|
7  | 
||
| 55150 | 8  | 
signature BASIC_CODE_SYMBOL =  | 
| 52136 | 9  | 
sig  | 
10  | 
  datatype ('a, 'b, 'c, 'd, 'e, 'f) attr =
 | 
|
11  | 
Constant of 'a | Type_Constructor of 'b | Type_Class of 'c |  | 
|
12  | 
Class_Relation of 'd | Class_Instance of 'e | Module of 'f  | 
|
| 55150 | 13  | 
end  | 
14  | 
||
15  | 
signature CODE_SYMBOL =  | 
|
16  | 
sig  | 
|
17  | 
include BASIC_CODE_SYMBOL  | 
|
18  | 
type T = (string, string, class, class * class, string * class, string) attr  | 
|
19  | 
structure Table: TABLE  | 
|
20  | 
structure Graph: GRAPH  | 
|
21  | 
val default_base: T -> string  | 
|
22  | 
val default_prefix: Proof.context -> T -> string  | 
|
23  | 
val quote: Proof.context -> T -> string  | 
|
24  | 
val marker: T -> string  | 
|
25  | 
val value: T  | 
|
26  | 
val is_value: T -> bool  | 
|
| 52136 | 27  | 
  val map_attr: ('a -> 'g) -> ('b -> 'h) -> ('c -> 'i) -> ('d -> 'j) -> ('e -> 'k) -> ('f -> 'l)
 | 
28  | 
    -> ('a, 'b, 'c, 'd, 'e, 'f) attr -> ('g, 'h, 'i, 'j, 'k, 'l) attr
 | 
|
29  | 
  val maps_attr: ('a -> 'g list) -> ('b -> 'h list)
 | 
|
30  | 
    -> ('c -> 'i list) -> ('d -> 'j list) -> ('e -> 'k list) -> ('f -> 'l list)
 | 
|
31  | 
    -> ('a, 'b, 'c, 'd, 'e, 'f) attr -> ('g, 'h, 'i, 'j, 'k, 'l) attr list
 | 
|
32  | 
  val maps_attr': ('a -> ('m * 'g) list) -> ('b -> ('m * 'h) list)
 | 
|
33  | 
    -> ('c -> ('m * 'i) list) -> ('d -> ('m * 'j) list) -> ('e -> ('m * 'k) list) -> ('f -> ('m * 'l) list)
 | 
|
34  | 
    -> ('a, 'b, 'c, 'd, 'e, 'f) attr -> ('m * ('g, 'h, 'i, 'j, 'k, 'l) attr) list
 | 
|
35  | 
  type ('a, 'b, 'c, 'd, 'e, 'f) data
 | 
|
36  | 
  val empty_data: ('a, 'b, 'c, 'd, 'e, 'f) data
 | 
|
37  | 
  val merge_data: ('a, 'b, 'c, 'd, 'e, 'f) data * ('a, 'b, 'c, 'd, 'e, 'f) data
 | 
|
38  | 
    -> ('a, 'b, 'c, 'd, 'e, 'f) data
 | 
|
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
39  | 
val set_data: (string * 'a option, string * 'b option, string * 'c option,  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
40  | 
(class * class) * 'd option, (string * class) * 'e option, string * 'f option) attr  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
41  | 
    -> ('a, 'b, 'c, 'd, 'e, 'f) data -> ('a, 'b, 'c, 'd, 'e, 'f) data
 | 
| 52136 | 42  | 
  val lookup_constant_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> string -> 'a option
 | 
43  | 
  val lookup_type_constructor_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> string -> 'b option
 | 
|
44  | 
  val lookup_type_class_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> class -> 'c option
 | 
|
45  | 
  val lookup_class_relation_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> class * class -> 'd option
 | 
|
46  | 
  val lookup_class_instance_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> string * class -> 'e option
 | 
|
47  | 
  val lookup_module_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> string -> 'f option
 | 
|
| 55150 | 48  | 
  val lookup: ('a, 'a, 'a, 'a, 'a, 'a) data -> T -> 'a option
 | 
49  | 
  val symbols_of: ('a, 'b, 'c, 'd, 'e, 'f) data -> T list
 | 
|
| 52136 | 50  | 
  val dest_module_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> (string * 'f) list
 | 
51  | 
end;  | 
|
52  | 
||
53  | 
structure Code_Symbol : CODE_SYMBOL =  | 
|
54  | 
struct  | 
|
55  | 
||
56  | 
(* data for symbols *)  | 
|
57  | 
||
58  | 
datatype ('a, 'b, 'c, 'd, 'e, 'f) attr =
 | 
|
59  | 
Constant of 'a | Type_Constructor of 'b | Type_Class of 'c | Class_Relation of 'd | Class_Instance of 'e | Module of 'f;  | 
|
60  | 
||
| 55150 | 61  | 
type T = (string, string, class, string * class, class * class, string) attr;  | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
62  | 
|
| 54933 | 63  | 
fun symbol_ord (Constant c1, Constant c2) = fast_string_ord (c1, c2)  | 
64  | 
| symbol_ord (Constant _, _) = GREATER  | 
|
65  | 
| symbol_ord (_, Constant _) = LESS  | 
|
66  | 
| symbol_ord (Type_Constructor tyco1, Type_Constructor tyco2) = fast_string_ord (tyco1, tyco2)  | 
|
67  | 
| symbol_ord (Type_Constructor _, _) = GREATER  | 
|
68  | 
| symbol_ord (_, Type_Constructor _) = LESS  | 
|
69  | 
| symbol_ord (Type_Class class1, Type_Class class2) = fast_string_ord (class1, class2)  | 
|
70  | 
| symbol_ord (Type_Class _, _) = GREATER  | 
|
71  | 
| symbol_ord (_, Type_Class _) = LESS  | 
|
72  | 
| symbol_ord (Class_Relation classrel1, Class_Relation classrel2) =  | 
|
73  | 
prod_ord fast_string_ord fast_string_ord (classrel1, classrel2)  | 
|
74  | 
| symbol_ord (Class_Relation _, _) = GREATER  | 
|
75  | 
| symbol_ord (_, Class_Relation _) = LESS  | 
|
76  | 
| symbol_ord (Class_Instance inst1, Class_Instance inst2) =  | 
|
77  | 
prod_ord fast_string_ord fast_string_ord (inst1, inst2)  | 
|
78  | 
| symbol_ord (Class_Instance _, _) = GREATER  | 
|
79  | 
| symbol_ord (_, Class_Instance _) = LESS  | 
|
80  | 
| symbol_ord (Module name1, Module name2) = fast_string_ord (name1, name2);  | 
|
81  | 
||
| 55150 | 82  | 
structure Table = Table(type key = T val ord = symbol_ord);  | 
83  | 
structure Graph = Graph(type key = T val ord = symbol_ord);  | 
|
84  | 
||
85  | 
local  | 
|
| 
56826
 
ba18bd41e510
enforce case of identifiers only to accomodate strict language requirements (or clear separation of constructors from variables in the case of SML)
 
haftmann 
parents: 
56811 
diff
changeset
 | 
86  | 
val base = Name.desymbolize NONE o Long_Name.base_name;  | 
| 55150 | 87  | 
fun base_rel (x, y) = base y ^ "_" ^ base x;  | 
88  | 
in  | 
|
89  | 
||
90  | 
fun default_base (Constant "") = "value"  | 
|
91  | 
| default_base (Constant c) = base c  | 
|
92  | 
| default_base (Type_Constructor tyco) = base tyco  | 
|
93  | 
| default_base (Type_Class class) = base class  | 
|
94  | 
| default_base (Class_Relation classrel) = base_rel classrel  | 
|
95  | 
| default_base (Class_Instance inst) = base_rel inst;  | 
|
96  | 
||
97  | 
end;  | 
|
| 54933 | 98  | 
|
99  | 
local  | 
|
| 77967 | 100  | 
  val thyname_of = Name_Space.theory_name {long = false};
 | 
101  | 
val thyname_of_type = thyname_of o Sign.type_space;  | 
|
102  | 
val thyname_of_class = thyname_of o Sign.class_space;  | 
|
| 77891 | 103  | 
  fun thyname_of_instance thy inst = case Thm.theory_names_of_arity {long = false} thy inst
 | 
| 
70475
 
98b6da301e13
backed out changeset 1b8858f4c393: odd problems e.g. in CAVA_LTL_Modelchecker;
 
wenzelm 
parents: 
70469 
diff
changeset
 | 
104  | 
   of [] => error ("No such instance: " ^ quote (fst inst ^ " :: " ^  snd inst))
 | 
| 
 
98b6da301e13
backed out changeset 1b8858f4c393: odd problems e.g. in CAVA_LTL_Modelchecker;
 
wenzelm 
parents: 
70469 
diff
changeset
 | 
105  | 
| thyname :: _ => thyname;  | 
| 54933 | 106  | 
fun thyname_of_const thy c = case Axclass.class_of_param thy c  | 
107  | 
of SOME class => thyname_of_class thy class  | 
|
108  | 
| NONE => (case Code.get_type_of_constr_or_abstr thy c  | 
|
109  | 
of SOME (tyco, _) => thyname_of_type thy tyco  | 
|
| 77967 | 110  | 
| NONE => thyname_of (Sign.const_space thy) c);  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
111  | 
fun prefix thy (Constant "") = "Code"  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
112  | 
| prefix thy (Constant c) = thyname_of_const thy c  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
113  | 
| prefix thy (Type_Constructor tyco) = thyname_of_type thy tyco  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
114  | 
| prefix thy (Type_Class class) = thyname_of_class thy class  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
115  | 
| prefix thy (Class_Relation (class, _)) = thyname_of_class thy class  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
116  | 
| prefix thy (Class_Instance inst) = thyname_of_instance thy inst;  | 
| 54933 | 117  | 
in  | 
118  | 
||
| 55150 | 119  | 
fun default_prefix ctxt = prefix (Proof_Context.theory_of ctxt);  | 
| 54933 | 120  | 
|
| 55150 | 121  | 
end;  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
122  | 
|
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
123  | 
val value = Constant "";  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
124  | 
fun is_value (Constant "") = true  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
125  | 
| is_value _ = false;  | 
| 54933 | 126  | 
|
| 55304 | 127  | 
fun quote ctxt (Constant c) =  | 
128  | 
Library.quote (Code.string_of_const (Proof_Context.theory_of ctxt) c)  | 
|
129  | 
| quote ctxt (Type_Constructor tyco) =  | 
|
130  | 
"type " ^ Library.quote (Proof_Context.markup_type ctxt tyco)  | 
|
131  | 
| quote ctxt (Type_Class class) =  | 
|
132  | 
"class " ^ Library.quote (Proof_Context.markup_class ctxt class)  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
133  | 
| quote ctxt (Class_Relation (sub, super)) =  | 
| 55304 | 134  | 
Library.quote (Proof_Context.markup_class ctxt sub) ^ " < " ^  | 
135  | 
Library.quote (Proof_Context.markup_class ctxt super)  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
136  | 
| quote ctxt (Class_Instance (tyco, class)) =  | 
| 55304 | 137  | 
Library.quote (Proof_Context.markup_type ctxt tyco) ^ " :: " ^  | 
138  | 
Library.quote (Proof_Context.markup_class ctxt class);  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
139  | 
|
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
140  | 
fun marker (Constant c) = "CONST " ^ c  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
141  | 
| marker (Type_Constructor tyco) = "TYPE " ^ tyco  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
142  | 
| marker (Type_Class class) = "CLASS " ^ class  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
143  | 
| marker (Class_Relation (sub, super)) = "CLASSREL " ^ sub ^ " < " ^ super  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
144  | 
| marker (Class_Instance (tyco, class)) = "INSTANCE " ^ tyco ^ " :: " ^ class;  | 
| 54933 | 145  | 
|
| 52136 | 146  | 
fun map_attr const tyco class classrel inst module (Constant x) = Constant (const x)  | 
147  | 
| map_attr const tyco class classrel inst module (Type_Constructor x) = Type_Constructor (tyco x)  | 
|
148  | 
| map_attr const tyco class classrel inst module (Type_Class x) = Type_Class (class x)  | 
|
149  | 
| map_attr const tyco class classrel inst module (Class_Relation x) = Class_Relation (classrel x)  | 
|
150  | 
| map_attr const tyco class classrel inst module (Class_Instance x) = Class_Instance (inst x)  | 
|
151  | 
| map_attr const tyco class classrel inst module (Module x) = Module (module x);  | 
|
152  | 
||
153  | 
fun maps_attr const tyco class classrel inst module (Constant x) = map Constant (const x)  | 
|
154  | 
| maps_attr const tyco class classrel inst module (Type_Constructor x) = map Type_Constructor (tyco x)  | 
|
155  | 
| maps_attr const tyco class classrel inst module (Type_Class x) = map Type_Class (class x)  | 
|
156  | 
| maps_attr const tyco class classrel inst module (Class_Relation x) = map Class_Relation (classrel x)  | 
|
157  | 
| maps_attr const tyco class classrel inst module (Class_Instance x) = map Class_Instance (inst x)  | 
|
158  | 
| maps_attr const tyco class classrel inst module (Module x) = map Module (module x);  | 
|
159  | 
||
160  | 
fun maps_attr' const tyco class classrel inst module (Constant x) = (map o apsnd) Constant (const x)  | 
|
161  | 
| maps_attr' const tyco class classrel inst module (Type_Constructor x) = (map o apsnd) Type_Constructor (tyco x)  | 
|
162  | 
| maps_attr' const tyco class classrel inst module (Type_Class x) = (map o apsnd) Type_Class (class x)  | 
|
163  | 
| maps_attr' const tyco class classrel inst module (Class_Relation x) = (map o apsnd) Class_Relation (classrel x)  | 
|
164  | 
| maps_attr' const tyco class classrel inst module (Class_Instance x) = (map o apsnd) Class_Instance (inst x)  | 
|
165  | 
| maps_attr' const tyco class classrel inst module (Module x) = (map o apsnd) Module (module x);  | 
|
166  | 
||
167  | 
datatype ('a, 'b, 'c, 'd, 'e, 'f) data =
 | 
|
168  | 
  Data of { constant: 'a Symtab.table, type_constructor: 'b Symtab.table,
 | 
|
169  | 
type_class: 'c Symtab.table, class_relation: 'd Symreltab.table, class_instance: 'e Symreltab.table,  | 
|
170  | 
module: 'f Symtab.table };  | 
|
171  | 
||
172  | 
fun make_data constant type_constructor type_class class_relation class_instance module =  | 
|
173  | 
  Data { constant = constant, type_constructor = type_constructor,
 | 
|
174  | 
type_class = type_class, class_relation = class_relation, class_instance = class_instance, module = module };  | 
|
175  | 
fun dest_data (Data x) = x;  | 
|
176  | 
fun map_data map_constant map_type_constructor map_type_class map_class_relation map_class_instance map_module  | 
|
177  | 
  (Data { constant, type_constructor, type_class, class_relation, class_instance, module }) =
 | 
|
178  | 
    Data { constant = map_constant constant, type_constructor = map_type_constructor type_constructor,
 | 
|
179  | 
type_class = map_type_class type_class, class_relation = map_class_relation class_relation,  | 
|
180  | 
class_instance = map_class_instance class_instance, module = map_module module };  | 
|
181  | 
||
182  | 
val empty_data = Data { constant = Symtab.empty, type_constructor = Symtab.empty,
 | 
|
183  | 
type_class = Symtab.empty, class_relation = Symreltab.empty, class_instance = Symreltab.empty, module = Symtab.empty };  | 
|
184  | 
fun merge_data (Data { constant = constant1, type_constructor = type_constructor1,
 | 
|
185  | 
type_class = type_class1, class_instance = class_instance1, class_relation = class_relation1, module = module1 },  | 
|
186  | 
  Data { constant = constant2, type_constructor = type_constructor2,
 | 
|
187  | 
type_class = type_class2, class_instance = class_instance2, class_relation = class_relation2, module = module2 }) =  | 
|
188  | 
make_data (Symtab.join (K snd) (constant1, constant2))  | 
|
189  | 
(Symtab.join (K snd) (type_constructor1, type_constructor2))  | 
|
190  | 
(Symtab.join (K snd) (type_class1, type_class2))  | 
|
191  | 
(Symreltab.join (K snd) (class_relation1, class_relation2))  | 
|
192  | 
(Symreltab.join (K snd) (class_instance1, class_instance2))  | 
|
193  | 
(Symtab.join (K snd) (module1, module2)); (*prefer later entries: K snd*)  | 
|
194  | 
||
195  | 
fun set_sym (sym, NONE) = Symtab.delete_safe sym  | 
|
196  | 
| set_sym (sym, SOME y) = Symtab.update (sym, y);  | 
|
197  | 
fun set_symrel (symrel, NONE) = Symreltab.delete_safe symrel  | 
|
198  | 
| set_symrel (symrel, SOME y) = Symreltab.update (symrel, y);  | 
|
199  | 
||
200  | 
fun set_data (Constant x) = map_data (set_sym x) I I I I I  | 
|
201  | 
| set_data (Type_Constructor x) = map_data I (set_sym x) I I I I  | 
|
202  | 
| set_data (Type_Class x) = map_data I I (set_sym x) I I I  | 
|
203  | 
| set_data (Class_Relation x) = map_data I I I (set_symrel x) I I  | 
|
204  | 
| set_data (Class_Instance x) = map_data I I I I (set_symrel x) I  | 
|
205  | 
| set_data (Module x) = map_data I I I I I (set_sym x);  | 
|
206  | 
||
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
207  | 
fun lookup_constant_data x = (Symtab.lookup o #constant o dest_data) x;  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
208  | 
fun lookup_type_constructor_data x = (Symtab.lookup o #type_constructor o dest_data) x;  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
209  | 
fun lookup_type_class_data x = (Symtab.lookup o #type_class o dest_data) x;  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
210  | 
fun lookup_class_relation_data x = (Symreltab.lookup o #class_relation o dest_data) x;  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
211  | 
fun lookup_class_instance_data x = (Symreltab.lookup o #class_instance o dest_data) x;  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
212  | 
fun lookup_module_data x = (Symtab.lookup o #module o dest_data) x;  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
213  | 
|
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
214  | 
fun lookup data (Constant x) = lookup_constant_data data x  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
215  | 
| lookup data (Type_Constructor x) = lookup_type_constructor_data data x  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
216  | 
| lookup data (Type_Class x) = lookup_type_class_data data x  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
217  | 
| lookup data (Class_Relation x) = lookup_class_relation_data data x  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
218  | 
| lookup data (Class_Instance x) = lookup_class_instance_data data x  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
219  | 
| lookup data (Module x) = lookup_module_data data x;  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
220  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
221  | 
fun symbols_of x = (map Constant o Symtab.keys o #constant o dest_data) x  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
222  | 
@ (map Type_Constructor o Symtab.keys o #type_constructor o dest_data) x  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
223  | 
@ (map Type_Class o Symtab.keys o #type_class o dest_data) x  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
224  | 
@ (map Class_Relation o Symreltab.keys o #class_relation o dest_data) x  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
225  | 
@ (map Class_Instance o Symreltab.keys o #class_instance o dest_data) x  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
226  | 
@ (map Module o Symtab.keys o #module o dest_data) x;  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55042 
diff
changeset
 | 
227  | 
|
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
228  | 
fun dest_module_data x = (Symtab.dest o #module o dest_data) x;  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52136 
diff
changeset
 | 
229  | 
|
| 52136 | 230  | 
end;  | 
| 55150 | 231  | 
|
232  | 
||
233  | 
structure Basic_Code_Symbol: BASIC_CODE_SYMBOL = Code_Symbol;  |