author | wenzelm |
Tue, 25 Apr 2006 22:23:41 +0200 | |
changeset 19466 | 29bc35832a77 |
parent 19435 | d7c10da57042 |
child 19482 | 9f11af8f7ef9 |
permissions | -rw-r--r-- |
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
1 |
(* Title: Pure/Tools/codegen_package.ML |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
2 |
ID: $Id$ |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
3 |
Author: Florian Haftmann, TU Muenchen |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
4 |
|
18217 | 5 |
Code generator from Isabelle theories to |
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
6 |
intermediate language ("Thin-gol"). |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
7 |
*) |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
8 |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
9 |
signature CODEGEN_PACKAGE = |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
10 |
sig |
18454 | 11 |
type auxtab; |
18702 | 12 |
type appgen = theory -> auxtab |
18756 | 13 |
-> (string * typ) * term list -> CodegenThingol.transact |
14 |
-> CodegenThingol.iexpr * CodegenThingol.transact; |
|
18702 | 15 |
|
18516 | 16 |
val add_appconst: string * ((int * int) * appgen) -> theory -> theory; |
18517 | 17 |
val add_appconst_i: xstring * ((int * int) * appgen) -> theory -> theory; |
18963 | 18 |
val add_prim_class: xstring -> (string * string) |
18517 | 19 |
-> theory -> theory; |
18963 | 20 |
val add_prim_tyco: xstring -> (string * string) |
18517 | 21 |
-> theory -> theory; |
19136 | 22 |
val add_prim_const: xstring -> (string * string) |
18517 | 23 |
-> theory -> theory; |
18963 | 24 |
val add_prim_i: string -> (string * CodegenThingol.prim list) |
18517 | 25 |
-> theory -> theory; |
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
26 |
val add_pretty_list: string -> string -> string * (int * string) |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
27 |
-> theory -> theory; |
18217 | 28 |
val add_alias: string * string -> theory -> theory; |
18455 | 29 |
val set_get_all_datatype_cons : (theory -> (string * string) list) |
18454 | 30 |
-> theory -> theory; |
18963 | 31 |
val set_get_datatype: (theory -> string -> ((string * sort) list * (string * typ list) list) option) |
32 |
-> theory -> theory; |
|
18702 | 33 |
val set_int_tyco: string -> theory -> theory; |
18217 | 34 |
|
19150 | 35 |
val codegen_term: term -> theory -> CodegenThingol.iexpr * theory; |
19136 | 36 |
val is_dtcon: string -> bool; |
19150 | 37 |
val consts_of_idfs: theory -> string list -> (string * typ) list; |
38 |
val idfs_of_consts: theory -> (string * typ) list -> string list; |
|
39 |
val get_root_module: theory -> CodegenThingol.module; |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
40 |
val get_ml_fun_datatype: theory -> (string -> string) |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
41 |
-> ((string * CodegenThingol.funn) list -> Pretty.T) |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
42 |
* ((string * CodegenThingol.datatyp) list -> Pretty.T); |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
43 |
|
18702 | 44 |
val appgen_default: appgen; |
18335 | 45 |
val appgen_let: (int -> term -> term list * term) |
46 |
-> appgen; |
|
47 |
val appgen_split: (int -> term -> term list * term) |
|
48 |
-> appgen; |
|
18915 | 49 |
val appgen_number_of: (term -> term) -> (theory -> term -> IntInf.int) -> string -> string |
18335 | 50 |
-> appgen; |
19038 | 51 |
val appgen_wfrec: appgen; |
18702 | 52 |
val add_case_const: (theory -> string -> (string * int) list option) -> xstring |
18517 | 53 |
-> theory -> theory; |
18702 | 54 |
val add_case_const_i: (theory -> string -> (string * int) list option) -> string |
18517 | 55 |
-> theory -> theory; |
18217 | 56 |
|
19008 | 57 |
val print_code: theory -> unit; |
18454 | 58 |
val rename_inconsistent: theory -> theory; |
18517 | 59 |
val ensure_datatype_case_consts: (theory -> string list) |
60 |
-> (theory -> string -> (string * int) list option) |
|
61 |
-> theory -> theory; |
|
18515 | 62 |
|
63 |
(*debugging purpose only*) |
|
18454 | 64 |
structure InstNameMangler: NAME_MANGLER; |
65 |
structure ConstNameMangler: NAME_MANGLER; |
|
66 |
structure DatatypeconsNameMangler: NAME_MANGLER; |
|
18231 | 67 |
structure CodegenData: THEORY_DATA; |
18454 | 68 |
val mk_tabs: theory -> auxtab; |
69 |
val alias_get: theory -> string -> string; |
|
18515 | 70 |
val idf_of_name: theory -> string -> string -> string; |
71 |
val idf_of_const: theory -> auxtab -> string * typ -> string; |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
72 |
val idf_of_co: theory -> auxtab -> string * string -> string option; |
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
73 |
end; |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
74 |
|
18217 | 75 |
structure CodegenPackage : CODEGEN_PACKAGE = |
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
76 |
struct |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
77 |
|
18850 | 78 |
open CodegenThingol; |
18217 | 79 |
|
18702 | 80 |
(* shallow name spaces *) |
18217 | 81 |
|
19038 | 82 |
val alias_ref = ref (fn thy : theory => fn s : string => s, fn thy : theory => fn s : string => s); |
83 |
fun alias_get name = (fst o !) alias_ref name; |
|
84 |
fun alias_rev name = (snd o !) alias_ref name; |
|
85 |
||
18885 | 86 |
val nsp_module = ""; (* a dummy by convention *) |
18217 | 87 |
val nsp_class = "class"; |
18454 | 88 |
val nsp_tyco = "tyco"; |
18217 | 89 |
val nsp_const = "const"; |
18454 | 90 |
val nsp_overl = "overl"; |
91 |
val nsp_dtcon = "dtcon"; |
|
18217 | 92 |
val nsp_mem = "mem"; |
93 |
val nsp_inst = "inst"; |
|
19213 | 94 |
val nsp_instmem = "instmem"; |
18217 | 95 |
|
19038 | 96 |
fun add_nsp shallow name = |
97 |
name |
|
98 |
|> NameSpace.unpack |
|
99 |
|> split_last |
|
100 |
|> apsnd (single #> cons shallow) |
|
101 |
|> (op @) |
|
102 |
|> NameSpace.pack; |
|
103 |
||
104 |
fun dest_nsp nsp idf = |
|
105 |
let |
|
106 |
val idf' = NameSpace.unpack idf; |
|
107 |
val (idf'', idf_base) = split_last idf'; |
|
108 |
val (modl, shallow) = split_last idf''; |
|
109 |
in |
|
110 |
if nsp = shallow |
|
111 |
then (SOME o NameSpace.pack) (modl @ [idf_base]) |
|
112 |
else NONE |
|
113 |
end; |
|
114 |
||
115 |
fun idf_of_name thy shallow name = |
|
116 |
name |
|
117 |
|> alias_get thy |
|
118 |
|> add_nsp shallow; |
|
119 |
||
120 |
fun name_of_idf thy shallow idf = |
|
121 |
idf |
|
122 |
|> dest_nsp shallow |
|
123 |
|> Option.map (alias_rev thy); |
|
124 |
||
18217 | 125 |
|
18702 | 126 |
(* code generator basics *) |
18454 | 127 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
128 |
type deftab = (typ * thm) list Symtab.table; |
19038 | 129 |
|
19136 | 130 |
fun is_overloaded thy c = case Defs.specifications_of (Theory.defs_of thy) c |
19202 | 131 |
of [] => true |
19435 | 132 |
| [(ty, _)] => not (Sign.typ_equiv thy (ty, Sign.the_const_type thy c)) |
19136 | 133 |
| _ => true; |
134 |
||
18454 | 135 |
structure InstNameMangler = NameManglerFun ( |
136 |
type ctxt = theory; |
|
137 |
type src = string * (class * string); |
|
138 |
val ord = prod_ord string_ord (prod_ord string_ord string_ord); |
|
139 |
fun mk thy ((thyname, (cls, tyco)), i) = |
|
19008 | 140 |
(NameSpace.base o alias_get thy) cls ^ "_" ^ (NameSpace.base o alias_get thy) tyco ^ implode (replicate i "'") |
18454 | 141 |
|> NameSpace.append thyname; |
142 |
fun is_valid _ _ = true; |
|
143 |
fun maybe_unique _ _ = NONE; |
|
144 |
fun re_mangle _ dst = error ("no such instance: " ^ quote dst); |
|
145 |
); |
|
146 |
||
147 |
structure ConstNameMangler = NameManglerFun ( |
|
19136 | 148 |
type ctxt = theory; |
149 |
type src = string * typ; |
|
150 |
val ord = prod_ord string_ord Term.typ_ord; |
|
151 |
fun mk thy ((c, ty), i) = |
|
18454 | 152 |
let |
19038 | 153 |
val c' = idf_of_name thy nsp_overl c; |
19202 | 154 |
val prefix = |
19435 | 155 |
case (AList.lookup (Sign.typ_equiv thy) |
19202 | 156 |
(Defs.specifications_of (Theory.defs_of thy) c)) ty |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
157 |
of SOME (_, thyname) => NameSpace.append thyname nsp_overl |
19202 | 158 |
| NONE => if c = "op =" |
159 |
then |
|
160 |
NameSpace.append |
|
161 |
(((fn s => Codegen.thyname_of_type s thy) o fst o dest_Type o hd o fst o strip_type) ty) |
|
162 |
nsp_overl |
|
163 |
else |
|
164 |
NameSpace.drop_base c'; |
|
19136 | 165 |
val c'' = NameSpace.append prefix (NameSpace.base c'); |
18454 | 166 |
fun mangle (Type (tyco, tys)) = |
19466 | 167 |
(NameSpace.base o alias_get thy) tyco :: flat (List.mapPartial mangle tys) |> SOME |
18454 | 168 |
| mangle _ = |
169 |
NONE |
|
170 |
in |
|
171 |
Vartab.empty |
|
19435 | 172 |
|> Type.raw_match (Sign.the_const_type thy c, ty) |
18454 | 173 |
|> map (snd o snd) o Vartab.dest |
174 |
|> List.mapPartial mangle |
|
19466 | 175 |
|> flat |
18454 | 176 |
|> null ? K ["x"] |
19038 | 177 |
|> cons c'' |
18454 | 178 |
|> space_implode "_" |
179 |
|> curry (op ^ o swap) ((implode oo replicate) i "'") |
|
180 |
end; |
|
181 |
fun is_valid _ _ = true; |
|
19150 | 182 |
fun maybe_unique thy (c, ty) = |
19136 | 183 |
if is_overloaded thy c |
184 |
then NONE |
|
185 |
else (SOME o idf_of_name thy nsp_const) c; |
|
186 |
fun re_mangle thy idf = |
|
187 |
case name_of_idf thy nsp_const idf |
|
188 |
of NONE => error ("no such constant: " ^ quote idf) |
|
189 |
| SOME c => (c, Sign.the_const_type thy c); |
|
18454 | 190 |
); |
191 |
||
192 |
structure DatatypeconsNameMangler = NameManglerFun ( |
|
193 |
type ctxt = theory; |
|
194 |
type src = string * string; |
|
195 |
val ord = prod_ord string_ord string_ord; |
|
18702 | 196 |
fun mk thy ((co, dtco), i) = |
18454 | 197 |
let |
198 |
fun basename 0 = NameSpace.base co |
|
18455 | 199 |
| basename 1 = NameSpace.base dtco ^ "_" ^ NameSpace.base co |
18454 | 200 |
| basename i = NameSpace.base dtco ^ "_" ^ NameSpace.base co ^ "_" ^ (implode oo replicate) (i-1) "'"; |
201 |
fun strip_dtco name = |
|
202 |
case (rev o NameSpace.unpack) name |
|
203 |
of x1::x2::xs => |
|
204 |
if x2 = NameSpace.base dtco |
|
205 |
then NameSpace.pack (x1::xs) |
|
206 |
else name |
|
207 |
| _ => name; |
|
208 |
in |
|
209 |
NameSpace.append (NameSpace.drop_base dtco) (basename i) |
|
210 |
|> strip_dtco |
|
211 |
end; |
|
212 |
fun is_valid _ _ = true; |
|
213 |
fun maybe_unique _ _ = NONE; |
|
214 |
fun re_mangle _ dst = error ("no such datatype constructor: " ^ quote dst); |
|
215 |
); |
|
216 |
||
19213 | 217 |
type auxtab = deftab |
19136 | 218 |
* (InstNameMangler.T * (typ list Symtab.table * ConstNameMangler.T) |
18865 | 219 |
* DatatypeconsNameMangler.T); |
18702 | 220 |
type eqextr = theory -> auxtab |
19008 | 221 |
-> string * typ -> (thm list * typ) option; |
222 |
type eqextr_default = theory -> auxtab |
|
223 |
-> string * typ -> ((thm list * term option) * typ) option; |
|
18865 | 224 |
type appgen = theory -> auxtab |
225 |
-> (string * typ) * term list -> transact -> iexpr * transact; |
|
18217 | 226 |
|
18702 | 227 |
val serializers = ref ( |
228 |
Symtab.empty |
|
229 |
|> Symtab.update ( |
|
230 |
#ml CodegenSerializer.serializers |
|
231 |
|> apsnd (fn seri => seri |
|
18865 | 232 |
(nsp_dtcon, nsp_class, K false) |
19213 | 233 |
[[nsp_module], [nsp_class, nsp_tyco], [nsp_const, nsp_overl, nsp_dtcon, nsp_class, nsp_mem, nsp_inst, nsp_instmem]] |
18702 | 234 |
) |
235 |
) |
|
236 |
|> Symtab.update ( |
|
237 |
#haskell CodegenSerializer.serializers |
|
238 |
|> apsnd (fn seri => seri |
|
18919 | 239 |
[nsp_module, nsp_class, nsp_tyco, nsp_dtcon] |
19213 | 240 |
[[nsp_module], [nsp_class], [nsp_tyco], [nsp_const, nsp_overl, nsp_mem], [nsp_dtcon], [nsp_inst], [nsp_instmem]] |
18702 | 241 |
) |
242 |
) |
|
243 |
); |
|
18217 | 244 |
|
245 |
||
18454 | 246 |
(* theory data for code generator *) |
18217 | 247 |
|
18912 | 248 |
fun merge_opt _ (x1, NONE) = x1 |
249 |
| merge_opt _ (NONE, x2) = x2 |
|
250 |
| merge_opt eq (SOME x1, SOME x2) = |
|
251 |
if eq (x1, x2) then SOME x1 else error ("incompatible options during merge"); |
|
252 |
||
18217 | 253 |
type gens = { |
18516 | 254 |
appconst: ((int * int) * (appgen * stamp)) Symtab.table, |
19008 | 255 |
eqextrs: (string * (eqextr_default * stamp)) list |
18217 | 256 |
}; |
257 |
||
18865 | 258 |
fun map_gens f { appconst, eqextrs } = |
18217 | 259 |
let |
18865 | 260 |
val (appconst, eqextrs) = |
261 |
f (appconst, eqextrs) |
|
262 |
in { appconst = appconst, eqextrs = eqextrs } : gens end; |
|
18217 | 263 |
|
264 |
fun merge_gens |
|
18865 | 265 |
({ appconst = appconst1 , eqextrs = eqextrs1 }, |
266 |
{ appconst = appconst2 , eqextrs = eqextrs2 }) = |
|
18516 | 267 |
{ appconst = Symtab.merge |
18865 | 268 |
(fn ((bounds1, (_, stamp1)), (bounds2, (_, stamp2))) => bounds1 = bounds2 |
269 |
andalso stamp1 = stamp2) |
|
18516 | 270 |
(appconst1, appconst2), |
18865 | 271 |
eqextrs = AList.merge (op =) (eq_snd (op =)) (eqextrs1, eqextrs2) |
18702 | 272 |
} : gens; |
18217 | 273 |
|
274 |
type logic_data = { |
|
18455 | 275 |
get_all_datatype_cons: ((theory -> (string * string) list) * stamp) option, |
18963 | 276 |
get_datatype: ((theory -> string -> ((string * sort) list * (string * typ list) list) option) * stamp) option, |
18217 | 277 |
alias: string Symtab.table * string Symtab.table |
278 |
}; |
|
279 |
||
18963 | 280 |
fun map_logic_data f { get_all_datatype_cons, get_datatype, alias } = |
18217 | 281 |
let |
18963 | 282 |
val ((get_all_datatype_cons, get_datatype), alias) = |
283 |
f ((get_all_datatype_cons, get_datatype), alias) |
|
284 |
in { get_all_datatype_cons = get_all_datatype_cons, |
|
285 |
get_datatype = get_datatype, alias = alias } : logic_data end; |
|
18217 | 286 |
|
287 |
fun merge_logic_data |
|
18963 | 288 |
({ get_all_datatype_cons = get_all_datatype_cons1, |
289 |
get_datatype = get_datatype1, alias = alias1 }, |
|
290 |
{ get_all_datatype_cons = get_all_datatype_cons2, |
|
291 |
get_datatype = get_datatype2, alias = alias2 }) = |
|
18217 | 292 |
let |
293 |
in |
|
18963 | 294 |
{ get_all_datatype_cons = merge_opt (eq_snd (op =)) |
18865 | 295 |
(get_all_datatype_cons1, get_all_datatype_cons2), |
18963 | 296 |
get_datatype = merge_opt (eq_snd (op =)) |
297 |
(get_datatype1, get_datatype2), |
|
18217 | 298 |
alias = (Symtab.merge (op =) (fst alias1, fst alias2), |
18304 | 299 |
Symtab.merge (op =) (snd alias1, snd alias2)) } : logic_data |
18217 | 300 |
end; |
301 |
||
18702 | 302 |
type target_data = { |
18865 | 303 |
syntax_class: string Symtab.table, |
18516 | 304 |
syntax_tyco: (itype CodegenSerializer.pretty_syntax * stamp) Symtab.table, |
305 |
syntax_const: (iexpr CodegenSerializer.pretty_syntax * stamp) Symtab.table |
|
18217 | 306 |
}; |
307 |
||
18865 | 308 |
fun map_target_data f { syntax_class, syntax_tyco, syntax_const } = |
18217 | 309 |
let |
18865 | 310 |
val (syntax_class, syntax_tyco, syntax_const) = |
311 |
f (syntax_class, syntax_tyco, syntax_const) |
|
312 |
in { |
|
313 |
syntax_class = syntax_class, |
|
314 |
syntax_tyco = syntax_tyco, |
|
315 |
syntax_const = syntax_const } : target_data |
|
18454 | 316 |
end; |
18217 | 317 |
|
18702 | 318 |
fun merge_target_data |
18865 | 319 |
({ syntax_class = syntax_class1, syntax_tyco = syntax_tyco1, syntax_const = syntax_const1 }, |
320 |
{ syntax_class = syntax_class2, syntax_tyco = syntax_tyco2, syntax_const = syntax_const2 }) = |
|
321 |
{ syntax_class = Symtab.merge (op =) (syntax_class1, syntax_class2), |
|
322 |
syntax_tyco = Symtab.merge (eq_snd (op =)) (syntax_tyco1, syntax_tyco2), |
|
18702 | 323 |
syntax_const = Symtab.merge (eq_snd (op =)) (syntax_const1, syntax_const2) } : target_data; |
18217 | 324 |
|
325 |
structure CodegenData = TheoryDataFun |
|
326 |
(struct |
|
327 |
val name = "Pure/codegen_package"; |
|
328 |
type T = { |
|
329 |
modl: module, |
|
330 |
gens: gens, |
|
331 |
logic_data: logic_data, |
|
18702 | 332 |
target_data: target_data Symtab.table |
18217 | 333 |
}; |
334 |
val empty = { |
|
335 |
modl = empty_module, |
|
18865 | 336 |
gens = { appconst = Symtab.empty, eqextrs = [] } : gens, |
18963 | 337 |
logic_data = { get_all_datatype_cons = NONE, |
338 |
get_datatype = NONE, |
|
18454 | 339 |
alias = (Symtab.empty, Symtab.empty) } : logic_data, |
18702 | 340 |
target_data = |
18217 | 341 |
Symtab.empty |
18702 | 342 |
|> Symtab.fold (fn (target, _) => |
18865 | 343 |
Symtab.update (target, |
344 |
{ syntax_class = Symtab.empty, syntax_tyco = Symtab.empty, syntax_const = Symtab.empty }) |
|
18702 | 345 |
) (! serializers) |
18217 | 346 |
} : T; |
347 |
val copy = I; |
|
348 |
val extend = I; |
|
349 |
fun merge _ ( |
|
18702 | 350 |
{ modl = modl1, gens = gens1, |
351 |
target_data = target_data1, logic_data = logic_data1 }, |
|
352 |
{ modl = modl2, gens = gens2, |
|
353 |
target_data = target_data2, logic_data = logic_data2 } |
|
18217 | 354 |
) = { |
355 |
modl = merge_module (modl1, modl2), |
|
356 |
gens = merge_gens (gens1, gens2), |
|
357 |
logic_data = merge_logic_data (logic_data1, logic_data2), |
|
19025 | 358 |
target_data = Symtab.join (K merge_target_data) (target_data1, target_data2) |
18217 | 359 |
}; |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
360 |
fun print thy (data : T) = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
361 |
let |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
362 |
val module = #modl data |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
363 |
in |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
364 |
(Pretty.writeln o Pretty.chunks) [pretty_module module, pretty_deps module] |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
365 |
end; |
18217 | 366 |
end); |
367 |
||
18708 | 368 |
val _ = Context.add_setup CodegenData.init; |
369 |
||
18217 | 370 |
fun map_codegen_data f thy = |
371 |
case CodegenData.get thy |
|
18702 | 372 |
of { modl, gens, target_data, logic_data } => |
373 |
let val (modl, gens, target_data, logic_data) = |
|
374 |
f (modl, gens, target_data, logic_data) |
|
375 |
in CodegenData.put { modl = modl, gens = gens, |
|
376 |
target_data = target_data, logic_data = logic_data } thy end; |
|
18217 | 377 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
378 |
val print_code = CodegenData.print; |
18217 | 379 |
|
19038 | 380 |
(* advanced name handling *) |
18217 | 381 |
|
382 |
fun add_alias (src, dst) = |
|
383 |
map_codegen_data |
|
18702 | 384 |
(fn (modl, gens, target_data, logic_data) => |
385 |
(modl, gens, target_data, |
|
18217 | 386 |
logic_data |> map_logic_data |
387 |
(apsnd (fn (tab, tab_rev) => |
|
388 |
(tab |> Symtab.update (src, dst), |
|
389 |
tab_rev |> Symtab.update (dst, src)))))); |
|
390 |
||
19038 | 391 |
val _ = alias_ref := (perhaps o Symtab.lookup o fst o #alias o #logic_data o CodegenData.get, |
392 |
perhaps o Symtab.lookup o snd o #alias o #logic_data o CodegenData.get); |
|
18702 | 393 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
394 |
fun idf_of_co thy (tabs as (_, (_, _, dtcontab))) (co, dtco) = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
395 |
case CodegenTheorems.get_datatypes thy dtco |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
396 |
of SOME ((_, cos), _) => if AList.defined (op =) cos co |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
397 |
then try (DatatypeconsNameMangler.get thy dtcontab) (co, dtco) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
398 |
|> the_default co |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
399 |
|> idf_of_name thy nsp_dtcon |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
400 |
|> SOME |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
401 |
else NONE |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
402 |
| NONE => NONE; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
403 |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
404 |
fun co_of_idf thy (tabs as (_, (_, _, dtcontab))) idf = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
405 |
case name_of_idf thy nsp_dtcon idf |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
406 |
of SOME idf' => let |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
407 |
val (c, dtco) = case try (DatatypeconsNameMangler.rev thy dtcontab) idf' |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
408 |
of SOME c_dtco => c_dtco |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
409 |
| NONE => case (snd o strip_type o Sign.the_const_type thy) idf' |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
410 |
of Type (dtco, _) => (idf', dtco) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
411 |
| _ => (idf', "nat") (*a hack*) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
412 |
in SOME (c, dtco) end |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
413 |
| NONE => NONE; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
414 |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
415 |
fun idf_of_const thy (tabs as (deftab, (_, (overltab1, overltab2), _))) |
18865 | 416 |
(c, ty) = |
18454 | 417 |
let |
18515 | 418 |
fun get_overloaded (c, ty) = |
19202 | 419 |
(case Symtab.lookup overltab1 c |
19136 | 420 |
of SOME tys => |
18515 | 421 |
(case find_first (curry (Sign.typ_instance thy) ty) tys |
19136 | 422 |
of SOME ty' => ConstNameMangler.get thy overltab2 |
423 |
(c, ty') |> SOME |
|
18515 | 424 |
| _ => NONE) |
19202 | 425 |
| _ => NONE) |
18515 | 426 |
fun get_datatypecons (c, ty) = |
427 |
case (snd o strip_type) ty |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
428 |
of Type (tyco, _) => idf_of_co thy tabs (c, tyco) |
18515 | 429 |
| _ => NONE; |
18702 | 430 |
in case get_datatypecons (c, ty) |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
431 |
of SOME idf => idf |
18702 | 432 |
| NONE => case get_overloaded (c, ty) |
18515 | 433 |
of SOME idf => idf |
19213 | 434 |
| NONE => case ClassPackage.lookup_const_class thy c |
18515 | 435 |
of SOME _ => idf_of_name thy nsp_mem c |
436 |
| NONE => idf_of_name thy nsp_const c |
|
437 |
end; |
|
18217 | 438 |
|
19177 | 439 |
fun recconst_of_idf thy (_, (_, (_, overltab2), _)) idf = |
18454 | 440 |
case name_of_idf thy nsp_const idf |
18702 | 441 |
of SOME c => SOME (c, Sign.the_const_type thy c) |
18454 | 442 |
| NONE => ( |
443 |
case dest_nsp nsp_overl idf |
|
444 |
of SOME _ => |
|
445 |
idf |
|
19136 | 446 |
|> ConstNameMangler.rev thy overltab2 |
18454 | 447 |
|> SOME |
448 |
| NONE => NONE |
|
449 |
); |
|
450 |
||
18865 | 451 |
|
452 |
(* further theory data accessors *) |
|
453 |
||
454 |
fun gen_add_appconst prep_const (raw_c, (bounds, ag)) thy = |
|
18454 | 455 |
let |
18865 | 456 |
val c = prep_const thy raw_c; |
457 |
in map_codegen_data |
|
458 |
(fn (modl, gens, target_data, logic_data) => |
|
459 |
(modl, |
|
460 |
gens |> map_gens |
|
461 |
(fn (appconst, eqextrs) => |
|
462 |
(appconst |
|
463 |
|> Symtab.update (c, (bounds, (ag, stamp ()))), |
|
464 |
eqextrs)), target_data, logic_data)) thy |
|
18454 | 465 |
end; |
18217 | 466 |
|
18865 | 467 |
val add_appconst = gen_add_appconst Sign.intern_const; |
468 |
val add_appconst_i = gen_add_appconst (K I); |
|
469 |
||
470 |
fun set_get_all_datatype_cons f = |
|
471 |
map_codegen_data |
|
472 |
(fn (modl, gens, target_data, logic_data) => |
|
473 |
(modl, gens, target_data, |
|
474 |
logic_data |
|
18963 | 475 |
|> map_logic_data ((apfst (fn (get_all_datatype_cons, get_datatype) |
476 |
=> (SOME (f, stamp ()), get_datatype)))))); |
|
18865 | 477 |
|
478 |
fun get_all_datatype_cons thy = |
|
479 |
case (#get_all_datatype_cons o #logic_data o CodegenData.get) thy |
|
480 |
of NONE => [] |
|
481 |
| SOME (f, _) => f thy; |
|
482 |
||
18963 | 483 |
fun set_get_datatype f = |
18865 | 484 |
map_codegen_data |
485 |
(fn (modl, gens, target_data, logic_data) => |
|
486 |
(modl, gens, target_data, |
|
487 |
logic_data |
|
18963 | 488 |
|> map_logic_data ((apfst (fn (get_all_datatype_cons, get_datatype) |
489 |
=> (get_all_datatype_cons, SOME (f, stamp ()))))))); |
|
18865 | 490 |
|
18963 | 491 |
fun get_datatype thy = |
492 |
case (#get_datatype o #logic_data o CodegenData.get) thy |
|
493 |
of NONE => K NONE |
|
494 |
| SOME (f, _) => f thy; |
|
18865 | 495 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
496 |
fun const_of_idf thy (tabs as (_, (_, _, dtcontab))) idf = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
497 |
case recconst_of_idf thy tabs idf |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
498 |
of SOME c_ty => SOME c_ty |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
499 |
| NONE => case dest_nsp nsp_mem idf |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
500 |
of SOME c => SOME (c, Sign.the_const_constraint thy c) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
501 |
| NONE => case co_of_idf thy tabs idf |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
502 |
of SOME (c, dtco) => SOME (c, Type (dtco, (the o AList.lookup (op =) ((snd o fst o the o CodegenTheorems.get_datatypes thy) dtco)) c)) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
503 |
| NONE => NONE; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
504 |
|
18865 | 505 |
fun set_int_tyco tyco thy = |
506 |
(serializers := ( |
|
507 |
! serializers |
|
508 |
|> Symtab.update ( |
|
509 |
#ml CodegenSerializer.serializers |
|
510 |
|> apsnd (fn seri => seri |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
511 |
(nsp_dtcon, nsp_class, fn tyco' => tyco' = idf_of_name thy nsp_tyco tyco) |
19213 | 512 |
[[nsp_module], [nsp_class, nsp_tyco], [nsp_const, nsp_overl, nsp_dtcon, nsp_mem, nsp_inst, nsp_instmem]] |
18865 | 513 |
) |
514 |
) |
|
515 |
); thy); |
|
516 |
||
517 |
||
518 |
(* definition and expression generators *) |
|
519 |
||
520 |
fun ensure_def_class thy tabs cls trns = |
|
18454 | 521 |
let |
18865 | 522 |
fun defgen_class thy (tabs as (_, (insttab, _, _))) cls trns = |
523 |
case name_of_idf thy nsp_class cls |
|
524 |
of SOME cls => |
|
525 |
let |
|
19283 | 526 |
val (v, cs) = (ClassPackage.the_consts_sign thy) cls; |
18865 | 527 |
val sortctxts = map (ClassPackage.extract_sortctxt thy o snd) cs; |
528 |
val idfs = map (idf_of_name thy nsp_mem o fst) cs; |
|
529 |
in |
|
530 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
531 |
|> debug_msg (fn _ => "trying defgen class declaration for " ^ quote cls) |
18865 | 532 |
|> fold_map (ensure_def_class thy tabs) (ClassPackage.the_superclasses thy cls) |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
533 |
||>> (exprsgen_type thy tabs o map snd) cs |
18865 | 534 |
||>> (fold_map o fold_map) (exprgen_tyvar_sort thy tabs) sortctxts |
535 |
|-> (fn ((supcls, memtypes), sortctxts) => succeed |
|
19283 | 536 |
(Class (supcls, (unprefix "'" v, idfs ~~ (sortctxts ~~ memtypes))))) |
18865 | 537 |
end |
538 |
| _ => |
|
539 |
trns |
|
540 |
|> fail ("no class definition found for " ^ quote cls); |
|
541 |
val cls' = idf_of_name thy nsp_class cls; |
|
18454 | 542 |
in |
543 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
544 |
|> debug_msg (fn _ => "generating class " ^ quote cls) |
19136 | 545 |
|> ensure_def [("class", defgen_class thy tabs)] ("generating class " ^ quote cls) cls' |
18865 | 546 |
|> pair cls' |
547 |
end |
|
548 |
and ensure_def_tyco thy tabs tyco trns = |
|
549 |
let |
|
18963 | 550 |
fun defgen_datatype thy (tabs as (_, (_, _, dtcontab))) dtco trns = |
551 |
case name_of_idf thy nsp_tyco dtco |
|
552 |
of SOME dtco => |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
553 |
(case CodegenTheorems.get_datatypes thy dtco |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
554 |
of SOME ((vars, cos), _) => |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
555 |
trns |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
556 |
|> debug_msg (fn _ => "trying defgen datatype for " ^ quote dtco) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
557 |
|> fold_map (exprgen_tyvar_sort thy tabs) vars |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
558 |
||>> fold_map (fn (c, ty) => |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
559 |
exprsgen_type thy tabs ty |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
560 |
#-> (fn ty' => pair ((the o idf_of_co thy tabs) (c, dtco), ty'))) cos |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
561 |
|-> (fn (vars, cos) => succeed (Datatype |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
562 |
(vars, cos))) |
18963 | 563 |
| NONE => |
564 |
trns |
|
565 |
|> fail ("no datatype found for " ^ quote dtco)) |
|
566 |
| NONE => |
|
567 |
trns |
|
568 |
|> fail ("not a type constructor: " ^ quote dtco) |
|
18865 | 569 |
val tyco' = idf_of_name thy nsp_tyco tyco; |
570 |
in |
|
571 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
572 |
|> debug_msg (fn _ => "generating type constructor " ^ quote tyco) |
19136 | 573 |
|> ensure_def [("datatype", defgen_datatype thy tabs)] ("generating type constructor " ^ quote tyco) tyco' |
18865 | 574 |
|> pair tyco' |
575 |
end |
|
576 |
and exprgen_tyvar_sort thy tabs (v, sort) trns = |
|
18516 | 577 |
trns |
18885 | 578 |
|> fold_map (ensure_def_class thy tabs) (ClassPackage.operational_sort_of thy sort) |
18865 | 579 |
|-> (fn sort => pair (unprefix "'" v, sort)) |
580 |
and exprgen_type thy tabs (TVar _) trns = |
|
18516 | 581 |
error "TVar encountered during code generation" |
18702 | 582 |
| exprgen_type thy tabs (TFree v_s) trns = |
18516 | 583 |
trns |
18702 | 584 |
|> exprgen_tyvar_sort thy tabs v_s |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
585 |
|-> (fn (v, sort) => pair (ITyVar v)) |
18516 | 586 |
| exprgen_type thy tabs (Type ("fun", [t1, t2])) trns = |
587 |
trns |
|
588 |
|> exprgen_type thy tabs t1 |
|
589 |
||>> exprgen_type thy tabs t2 |
|
590 |
|-> (fn (t1', t2') => pair (t1' `-> t2')) |
|
591 |
| exprgen_type thy tabs (Type (tyco, tys)) trns = |
|
592 |
trns |
|
593 |
|> ensure_def_tyco thy tabs tyco |
|
594 |
||>> fold_map (exprgen_type thy tabs) tys |
|
18912 | 595 |
|-> (fn (tyco, tys) => pair (tyco `%% tys)) |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
596 |
and exprsgen_type thy tabs = |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
597 |
fold_map (exprgen_type thy tabs); |
18516 | 598 |
|
18885 | 599 |
fun exprgen_classlookup thy tabs (ClassPackage.Instance (inst, ls)) trns = |
18517 | 600 |
trns |
18885 | 601 |
|> ensure_def_inst thy tabs inst |
602 |
||>> (fold_map o fold_map) (exprgen_classlookup thy tabs) ls |
|
603 |
|-> (fn (inst, ls) => pair (Instance (inst, ls))) |
|
19253 | 604 |
| exprgen_classlookup thy tabs (ClassPackage.Lookup (clss, (v, (i, j)))) trns = |
18516 | 605 |
trns |
18517 | 606 |
|> fold_map (ensure_def_class thy tabs) clss |
19253 | 607 |
|-> (fn clss => pair (Lookup (clss, (v |> unprefix "'", if j = 1 then ~1 else i)))) |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
608 |
and mk_fun thy tabs (c, ty) trns = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
609 |
case CodegenTheorems.get_funs thy (c, Type.varifyT ty) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
610 |
of SOME (ty, eq_thms) => |
18865 | 611 |
let |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
612 |
val sortctxt = ClassPackage.extract_sortctxt thy ty; |
18865 | 613 |
fun dest_eqthm eq_thm = |
614 |
let |
|
615 |
val ((t, args), rhs) = |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
616 |
(apfst strip_comb o Logic.dest_equals |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
617 |
o prop_of) eq_thm; |
18865 | 618 |
in case t |
619 |
of Const (c', _) => if c' = c then (args, rhs) |
|
620 |
else error ("illegal function equation for " ^ quote c |
|
621 |
^ ", actually defining " ^ quote c') |
|
622 |
| _ => error ("illegal function equation for " ^ quote c) |
|
623 |
end; |
|
624 |
in |
|
18517 | 625 |
trns |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
626 |
|> (exprsgen_eqs thy tabs o map dest_eqthm) eq_thms |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
627 |
||>> exprsgen_type thy tabs [ty] |
18865 | 628 |
||>> fold_map (exprgen_tyvar_sort thy tabs) sortctxt |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
629 |
|-> (fn ((eqs, [ity]), sortctxt) => (pair o SOME) ((eqs, (sortctxt, ity)), ty)) |
18865 | 630 |
end |
631 |
| NONE => (NONE, trns) |
|
632 |
and ensure_def_inst thy (tabs as (_, (insttab, _, _))) (cls, tyco) trns = |
|
633 |
let |
|
634 |
fun defgen_inst thy (tabs as (_, (insttab, _, _))) inst trns = |
|
635 |
case Option.map (InstNameMangler.rev thy insttab) (name_of_idf thy nsp_inst inst) |
|
18885 | 636 |
of SOME (_, (class, tyco)) => |
18865 | 637 |
let |
18885 | 638 |
val (arity, memdefs) = ClassPackage.the_inst_sign thy (class, tyco); |
18865 | 639 |
fun gen_suparity supclass trns = |
640 |
trns |
|
19136 | 641 |
|> ensure_def_class thy tabs supclass |
18885 | 642 |
||>> ensure_def_inst thy tabs (supclass, tyco) |
19136 | 643 |
||>> (fold_map o fold_map) (exprgen_classlookup thy tabs) |
19213 | 644 |
(ClassPackage.extract_classlookup_inst thy (class, tyco) supclass) |
19136 | 645 |
|-> (fn ((supclass, inst), lss) => pair (supclass, (inst, lss))); |
18865 | 646 |
fun gen_membr (m, ty) trns = |
647 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
648 |
|> mk_fun thy tabs (m, ty) |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
649 |
|-> (fn NONE => error ("could not derive definition for member " |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
650 |
^ quote m ^ " :: " ^ Sign.string_of_typ thy ty) |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
651 |
| SOME (funn, ty_decl) => (fold_map o fold_map) (exprgen_classlookup thy tabs) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
652 |
(ClassPackage.extract_classlookup_member thy (ty_decl, ty)) |
19253 | 653 |
#-> (fn lss => |
654 |
pair (idf_of_name thy nsp_mem m, ((idf_of_name thy nsp_instmem m, funn), lss)))); |
|
18865 | 655 |
in |
656 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
657 |
|> debug_msg (fn _ => "trying defgen class instance for (" ^ quote cls |
18865 | 658 |
^ ", " ^ quote tyco ^ ")") |
18885 | 659 |
|> ensure_def_class thy tabs class |
18865 | 660 |
||>> ensure_def_tyco thy tabs tyco |
661 |
||>> fold_map (exprgen_tyvar_sort thy tabs) arity |
|
18885 | 662 |
||>> fold_map gen_suparity (ClassPackage.the_superclasses thy class) |
18865 | 663 |
||>> fold_map gen_membr memdefs |
18885 | 664 |
|-> (fn ((((class, tyco), arity), suparities), memdefs) => |
665 |
succeed (Classinst (((class, (tyco, arity)), suparities), memdefs))) |
|
18865 | 666 |
end |
667 |
| _ => |
|
668 |
trns |> fail ("no class instance found for " ^ quote inst); |
|
19213 | 669 |
val (_, thyname) = (the o AList.lookup (op =) (ClassPackage.the_instances thy cls)) tyco; |
18865 | 670 |
val inst = idf_of_name thy nsp_inst (InstNameMangler.get thy insttab (thyname, (cls, tyco))); |
671 |
in |
|
672 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
673 |
|> debug_msg (fn _ => "generating instance " ^ quote cls ^ " / " ^ quote tyco) |
19136 | 674 |
|> ensure_def [("instance", defgen_inst thy tabs)] |
18865 | 675 |
("generating instance " ^ quote cls ^ " / " ^ quote tyco) inst |
676 |
|> pair inst |
|
677 |
end |
|
19213 | 678 |
and ensure_def_const thy (tabs as (_, (_, overltab, dtcontab))) (c, ty) trns = |
18865 | 679 |
let |
680 |
fun defgen_funs thy tabs c trns = |
|
681 |
case recconst_of_idf thy tabs c |
|
682 |
of SOME (c, ty) => |
|
683 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
684 |
|> mk_fun thy tabs (c, ty) |
19253 | 685 |
|-> (fn SOME (funn, _) => succeed (Fun funn) |
18865 | 686 |
| NONE => fail ("no defining equations found for " ^ quote c)) |
687 |
| NONE => |
|
688 |
trns |
|
689 |
|> fail ("not a constant: " ^ quote c); |
|
18963 | 690 |
fun defgen_clsmem thy tabs m trns = |
691 |
case name_of_idf thy nsp_mem m |
|
692 |
of SOME m => |
|
693 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
694 |
|> debug_msg (fn _ => "trying defgen class member for " ^ quote m) |
18963 | 695 |
|> ensure_def_class thy tabs ((the o ClassPackage.lookup_const_class thy) m) |
696 |
|-> (fn cls => succeed Undef) |
|
697 |
| _ => |
|
698 |
trns |> fail ("no class member found for " ^ quote m) |
|
18865 | 699 |
fun defgen_datatypecons thy (tabs as (_, (_, _, dtcontab))) co trns = |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
700 |
case co_of_idf thy tabs co |
18865 | 701 |
of SOME (co, dtco) => |
702 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
703 |
|> debug_msg (fn _ => "trying defgen datatype constructor for " ^ quote co) |
18865 | 704 |
|> ensure_def_tyco thy tabs dtco |
705 |
|-> (fn dtco => succeed Undef) |
|
706 |
| _ => |
|
707 |
trns |
|
708 |
|> fail ("not a datatype constructor: " ^ quote co); |
|
18963 | 709 |
fun get_defgen idf = |
710 |
if (is_some oo name_of_idf thy) nsp_const idf |
|
711 |
orelse (is_some oo name_of_idf thy) nsp_overl idf |
|
712 |
then ("funs", defgen_funs thy tabs) |
|
713 |
else if (is_some oo name_of_idf thy) nsp_mem idf |
|
714 |
then ("clsmem", defgen_clsmem thy tabs) |
|
715 |
else if (is_some oo name_of_idf thy) nsp_dtcon idf |
|
716 |
then ("datatypecons", defgen_datatypecons thy tabs) |
|
717 |
else error ("illegal shallow name space for constant: " ^ quote idf); |
|
718 |
val idf = idf_of_const thy tabs (c, ty); |
|
18865 | 719 |
in |
720 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
721 |
|> debug_msg (fn _ => "generating constant " ^ quote c) |
19136 | 722 |
|> ensure_def ((single o get_defgen) idf) ("generating constant " ^ quote c) idf |
18963 | 723 |
|> pair idf |
18865 | 724 |
end |
18517 | 725 |
and exprgen_term thy tabs (Const (f, ty)) trns = |
726 |
trns |
|
727 |
|> appgen thy tabs ((f, ty), []) |
|
18516 | 728 |
|-> (fn e => pair e) |
18912 | 729 |
| exprgen_term thy tabs (Var _) trns = |
730 |
error "Var encountered during code generation" |
|
18516 | 731 |
| exprgen_term thy tabs (Free (v, ty)) trns = |
732 |
trns |
|
18912 | 733 |
|> exprgen_type thy tabs ty |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
734 |
|-> (fn ty => pair (IVar v)) |
19136 | 735 |
| exprgen_term thy tabs (Abs (abs as (_, ty, _))) trns = |
736 |
let |
|
737 |
val (v, t) = Term.variant_abs abs |
|
738 |
in |
|
739 |
trns |
|
740 |
|> exprgen_type thy tabs ty |
|
741 |
||>> exprgen_term thy tabs t |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
742 |
|-> (fn (ty, e) => pair ((v, ty) `|-> e)) |
19136 | 743 |
end |
18516 | 744 |
| exprgen_term thy tabs (t as t1 $ t2) trns = |
745 |
let |
|
746 |
val (t', ts) = strip_comb t |
|
747 |
in case t' |
|
748 |
of Const (f, ty) => |
|
749 |
trns |
|
18517 | 750 |
|> appgen thy tabs ((f, ty), ts) |
18516 | 751 |
|-> (fn e => pair e) |
752 |
| _ => |
|
753 |
trns |
|
754 |
|> exprgen_term thy tabs t' |
|
755 |
||>> fold_map (exprgen_term thy tabs) ts |
|
756 |
|-> (fn (e, es) => pair (e `$$ es)) |
|
18865 | 757 |
end |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
758 |
and exprsgen_term thy tabs = |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
759 |
fold_map (exprgen_term thy tabs) |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
760 |
and exprsgen_eqs thy tabs = |
18912 | 761 |
apfst (map (fn (rhs::args) => (args, rhs))) |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
762 |
oo fold_burrow (exprsgen_term thy tabs) |
18912 | 763 |
o map (fn (args, rhs) => (rhs :: args)) |
18865 | 764 |
and appgen_default thy tabs ((c, ty), ts) trns = |
765 |
trns |
|
766 |
|> ensure_def_const thy tabs (c, ty) |
|
19202 | 767 |
||>> exprsgen_type thy tabs [ty] |
18885 | 768 |
||>> (fold_map o fold_map) (exprgen_classlookup thy tabs) |
769 |
(ClassPackage.extract_classlookup thy (c, ty)) |
|
18912 | 770 |
||>> fold_map (exprgen_term thy tabs) ts |
19202 | 771 |
|-> (fn (((c, [ty]), ls), es) => |
772 |
pair (IConst (c, (ls, ty)) `$$ es)) |
|
18865 | 773 |
and appgen thy tabs ((f, ty), ts) trns = |
774 |
case Symtab.lookup ((#appconst o #gens o CodegenData.get) thy) f |
|
775 |
of SOME ((imin, imax), (ag, _)) => |
|
776 |
if length ts < imin then |
|
777 |
let |
|
778 |
val d = imin - length ts; |
|
779 |
val vs = Term.invent_names (add_term_names (Const (f, ty), [])) "x" d; |
|
780 |
val tys = Library.take (d, ((fst o strip_type) ty)); |
|
781 |
in |
|
782 |
trns |
|
18912 | 783 |
|> fold_map (exprgen_type thy tabs) tys |
18865 | 784 |
||>> ag thy tabs ((f, ty), ts @ map2 (curry Free) vs tys) |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
785 |
|-> (fn (tys, e) => pair (vs ~~ tys `|--> e)) |
18865 | 786 |
end |
787 |
else if length ts > imax then |
|
788 |
trns |
|
789 |
|> ag thy tabs ((f, ty), Library.take (imax, ts)) |
|
790 |
||>> fold_map (exprgen_term thy tabs) (Library.drop (imax, ts)) |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
791 |
|-> (fn (e, es) => pair (e `$$ es)) |
18865 | 792 |
else |
793 |
trns |
|
794 |
|> ag thy tabs ((f, ty), ts) |
|
795 |
| NONE => |
|
796 |
trns |
|
797 |
|> appgen_default thy tabs ((f, ty), ts); |
|
18516 | 798 |
|
18702 | 799 |
|
18217 | 800 |
(* parametrized generators, for instantiation in HOL *) |
801 |
||
19213 | 802 |
fun appgen_split strip_abs thy tabs (app as (c_ty, [t])) trns = |
803 |
case strip_abs 1 (Const c_ty $ t) |
|
19202 | 804 |
of ([vt], bt) => |
19150 | 805 |
trns |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
806 |
|> exprgen_term thy tabs vt |
19202 | 807 |
||>> exprgen_type thy tabs (type_of vt) |
808 |
||>> exprgen_term thy tabs bt |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
809 |
||>> appgen_default thy tabs app |
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
810 |
|-> (fn (((ve, vty), be), e0) => pair (IAbs (((ve, vty), be), e0))) |
19150 | 811 |
| _ => |
812 |
trns |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
813 |
|> appgen_default thy tabs app; |
18335 | 814 |
|
19202 | 815 |
fun appgen_let strip_abs thy tabs (app as ((c, ty), [dt, ct])) trns = |
816 |
case strip_abs 1 ct |
|
817 |
of ([st], bt) => |
|
818 |
trns |
|
819 |
|> exprgen_term thy tabs dt |
|
820 |
||>> exprgen_type thy tabs (type_of dt) |
|
821 |
||>> exprgen_term thy tabs st |
|
822 |
||>> exprgen_term thy tabs bt |
|
823 |
||>> appgen_default thy tabs app |
|
824 |
|-> (fn ((((de, dty), se), be), e0) => pair (ICase (((de, dty), [(se, be)]), e0))) |
|
825 |
| _ => |
|
826 |
trns |
|
827 |
|> appgen_default thy tabs app; |
|
828 |
||
18702 | 829 |
fun appgen_number_of mk_int_to_nat bin_to_int tyco_int tyco_nat thy tabs ((_, |
830 |
Type (_, [_, ty as Type (tyco, [])])), [bin]) trns = |
|
19202 | 831 |
if tyco = tyco_nat then |
18702 | 832 |
trns |
19202 | 833 |
|> exprgen_term thy tabs (mk_int_to_nat bin) (*should be a preprocessor's work*) |
834 |
else |
|
835 |
let |
|
836 |
val i = bin_to_int thy bin; |
|
837 |
val _ = if i < 0 then error ("negative numeral: " ^ IntInf.toString i) else (); |
|
838 |
(*should be a preprocessor's work*) |
|
839 |
in |
|
840 |
trns |
|
841 |
|> exprgen_type thy tabs ty |
|
842 |
|-> (fn ty => pair (CodegenThingol.INum ((i, ty), ()))) |
|
843 |
end; |
|
18217 | 844 |
|
19038 | 845 |
fun appgen_wfrec thy tabs ((c, ty), [_, tf, tx]) trns = |
846 |
let |
|
847 |
val ty_def = (op ---> o apfst tl o strip_type o Sign.the_const_type thy) c; |
|
848 |
val ty' = (op ---> o apfst tl o strip_type) ty; |
|
849 |
val idf = idf_of_const thy tabs (c, ty); |
|
850 |
in |
|
851 |
trns |
|
19136 | 852 |
|> ensure_def [("wfrec", (K o succeed) Undef)] ("generating wfrec") idf |
853 |
|> exprgen_type thy tabs ty' |
|
854 |
||>> (fold_map o fold_map) (exprgen_classlookup thy tabs) |
|
855 |
(ClassPackage.extract_classlookup thy (c, ty)) |
|
856 |
||>> exprsgen_type thy tabs [ty_def] |
|
857 |
||>> exprgen_term thy tabs tf |
|
858 |
||>> exprgen_term thy tabs tx |
|
19202 | 859 |
|-> (fn ((((_, ls), [ty]), tf), tx) => pair (IConst (idf, (ls, ty)) `$ tf `$ tx)) |
19038 | 860 |
end; |
861 |
||
19008 | 862 |
fun eqextr_eq f fals thy tabs ("op =", ty) = |
863 |
(case ty |
|
864 |
of Type ("fun", [Type (dtco, _), _]) => |
|
865 |
(case f thy dtco |
|
866 |
of [] => NONE |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
867 |
| [eq] => SOME ((the o CodegenTheorems.preprocess_fun thy) [eq], NONE) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
868 |
| eqs => SOME ((the o CodegenTheorems.preprocess_fun thy) eqs, SOME fals)) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
869 |
| _ => NONE) |> Option.map (fn ((ty, eqs), default) => ((eqs, default), ty)) |
19008 | 870 |
| eqextr_eq f fals thy tabs _ = |
871 |
NONE; |
|
872 |
||
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
873 |
fun appgen_datatype_case cos thy tabs (app as ((_, ty), ts)) trns = |
18247
b17724cae935
code generator: case expressions, improved name resolving
haftmann
parents:
18231
diff
changeset
|
874 |
let |
18517 | 875 |
val (ts', t) = split_last ts; |
876 |
val (tys, dty) = (split_last o fst o strip_type) ty; |
|
877 |
fun gen_names i = |
|
878 |
variantlist (replicate i "x", foldr add_term_names |
|
879 |
(map (fst o fst o dest_Var) (foldr add_term_vars [] ts)) ts); |
|
880 |
fun cg_case_d (((cname, i), ty), t) trns = |
|
18247
b17724cae935
code generator: case expressions, improved name resolving
haftmann
parents:
18231
diff
changeset
|
881 |
let |
b17724cae935
code generator: case expressions, improved name resolving
haftmann
parents:
18231
diff
changeset
|
882 |
val vs = gen_names i; |
b17724cae935
code generator: case expressions, improved name resolving
haftmann
parents:
18231
diff
changeset
|
883 |
val tys = Library.take (i, (fst o strip_type) ty); |
18330 | 884 |
val frees = map2 (curry Free) vs tys; |
18247
b17724cae935
code generator: case expressions, improved name resolving
haftmann
parents:
18231
diff
changeset
|
885 |
val t' = Envir.beta_norm (list_comb (t, frees)); |
b17724cae935
code generator: case expressions, improved name resolving
haftmann
parents:
18231
diff
changeset
|
886 |
in |
b17724cae935
code generator: case expressions, improved name resolving
haftmann
parents:
18231
diff
changeset
|
887 |
trns |
18516 | 888 |
|> exprgen_term thy tabs (list_comb (Const (cname, tys ---> dty), frees)) |
889 |
||>> exprgen_term thy tabs t' |
|
18865 | 890 |
|-> (fn (ep, e) => pair (ep, e)) |
18247
b17724cae935
code generator: case expressions, improved name resolving
haftmann
parents:
18231
diff
changeset
|
891 |
end; |
18517 | 892 |
in |
893 |
trns |
|
894 |
|> exprgen_term thy tabs t |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
895 |
||>> exprgen_type thy tabs dty |
18517 | 896 |
||>> fold_map cg_case_d ((cos ~~ tys) ~~ ts') |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
897 |
||>> appgen_default thy tabs app |
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
898 |
|-> (fn (((de, dty), bses), e0) => pair (ICase (((de, dty), bses), e0))) |
18247
b17724cae935
code generator: case expressions, improved name resolving
haftmann
parents:
18231
diff
changeset
|
899 |
end; |
b17724cae935
code generator: case expressions, improved name resolving
haftmann
parents:
18231
diff
changeset
|
900 |
|
18702 | 901 |
fun gen_add_case_const prep_c get_case_const_data raw_c thy = |
18517 | 902 |
let |
903 |
val c = prep_c thy raw_c; |
|
18702 | 904 |
val (tys, dty) = (split_last o fst o strip_type o Sign.the_const_type thy) c; |
18517 | 905 |
val cos = (the o get_case_const_data thy) c; |
906 |
val n_eta = length cos + 1; |
|
907 |
in |
|
908 |
thy |
|
909 |
|> add_appconst_i (c, ((n_eta, n_eta), appgen_datatype_case cos)) |
|
910 |
end; |
|
911 |
||
18702 | 912 |
val add_case_const = gen_add_case_const Sign.intern_const; |
913 |
val add_case_const_i = gen_add_case_const (K I); |
|
18517 | 914 |
|
18217 | 915 |
|
18516 | 916 |
|
917 |
(** theory interface **) |
|
18217 | 918 |
|
18454 | 919 |
fun mk_tabs thy = |
18217 | 920 |
let |
18454 | 921 |
fun mk_insttab thy = |
922 |
InstNameMangler.empty |
|
923 |
|> Symtab.fold_map |
|
19213 | 924 |
(fn (cls, clsinsts) => fold_map |
18454 | 925 |
(fn (tyco, thyname) => InstNameMangler.declare thy (thyname, (cls, tyco))) clsinsts) |
926 |
(ClassPackage.get_classtab thy) |
|
927 |
|-> (fn _ => I); |
|
19136 | 928 |
fun mk_overltabs thy = |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
929 |
(Symtab.empty, ConstNameMangler.empty) |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
930 |
|> Symtab.fold |
19136 | 931 |
(fn (c, _) => |
932 |
let |
|
933 |
val deftab = Defs.specifications_of (Theory.defs_of thy) c |
|
934 |
val is_overl = (is_none o ClassPackage.lookup_const_class thy) c |
|
935 |
andalso case deftab |
|
936 |
of [] => false |
|
19435 | 937 |
| [(ty, _)] => not (Sign.typ_equiv thy (ty, Sign.the_const_type thy c)) |
19136 | 938 |
| _ => true; |
939 |
in if is_overl then (fn (overltab1, overltab2) => ( |
|
940 |
overltab1 |
|
941 |
|> Symtab.update_new (c, map fst deftab), |
|
942 |
overltab2 |
|
943 |
|> fold_map (fn (ty, _) => ConstNameMangler.declare thy (c, ty)) deftab |
|
944 |
|-> (fn _ => I))) else I |
|
945 |
end) ((#2 o #constants o Consts.dest o #consts o Sign.rep_sg) thy) |
|
946 |
|> (fn (overltab1, overltab2) => |
|
947 |
let |
|
948 |
val c = "op ="; |
|
949 |
val ty = Sign.the_const_type thy c; |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
950 |
fun inst tyco = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
951 |
let |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
952 |
val ty_inst = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
953 |
tyco |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
954 |
|> Symtab.lookup ((snd o #types o Type.rep_tsig o Sign.tsig_of) thy) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
955 |
|> (fn SOME (Type.LogicalType i, _) => i) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
956 |
|> Term.invent_names [] "'a" |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
957 |
|> map (fn v => (TVar ((v, 0), Sign.defaultS thy))) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
958 |
|> (fn tys => Type (tyco, tys)) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
959 |
in map_atyps (fn _ => ty_inst) ty end; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
960 |
val tys = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
961 |
(Type.logical_types o Sign.tsig_of) thy |
19466 | 962 |
|> filter (fn tyco => can (Sign.arity_sorts thy tyco) (Sign.defaultS thy)) |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
963 |
|> map inst |
19136 | 964 |
in |
965 |
(overltab1 |
|
966 |
|> Symtab.update_new (c, tys), |
|
967 |
overltab2 |
|
968 |
|> fold (fn ty => ConstNameMangler.declare thy (c, ty) #> snd) tys) |
|
969 |
end); |
|
18454 | 970 |
fun mk_dtcontab thy = |
971 |
DatatypeconsNameMangler.empty |
|
972 |
|> fold_map |
|
18455 | 973 |
(fn (_, co_dtco) => DatatypeconsNameMangler.declare_multi thy co_dtco) |
974 |
(fold (fn (co, dtco) => |
|
975 |
let |
|
19008 | 976 |
val key = ((NameSpace.drop_base o NameSpace.drop_base) co, NameSpace.base co); |
18455 | 977 |
in AList.default (op =) (key, []) #> AList.map_entry (op =) key (cons (co, dtco)) end |
978 |
) (get_all_datatype_cons thy) []) |
|
18454 | 979 |
|-> (fn _ => I); |
980 |
val insttab = mk_insttab thy; |
|
19136 | 981 |
val overltabs = mk_overltabs thy; |
18454 | 982 |
val dtcontab = mk_dtcontab thy; |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
983 |
in (Symtab.empty, (insttab, overltabs, dtcontab)) end; |
18217 | 984 |
|
18756 | 985 |
fun get_serializer target = |
986 |
case Symtab.lookup (!serializers) target |
|
987 |
of SOME seri => seri |
|
19008 | 988 |
| NONE => Scan.fail_with (fn _ => "unknown code target language: " ^ quote target) (); |
18335 | 989 |
|
18516 | 990 |
fun map_module f = |
18702 | 991 |
map_codegen_data (fn (modl, gens, target_data, logic_data) => |
992 |
(f modl, gens, target_data, logic_data)); |
|
18516 | 993 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
994 |
(*fun delete_defs NONE thy = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
995 |
map_module (K CodegenThingol.empty_module) thy |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
996 |
| delete_defs (SOME c) thy = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
997 |
let |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
998 |
val tabs = mk_tabs thy; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
999 |
in |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1000 |
map_module (CodegenThingol.purge_module []) thy |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1001 |
end; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1002 |
does not make sense: |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1003 |
* primitve definitions have to be kept |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1004 |
* *all* overloaded defintitions for a constant have to be purged |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1005 |
* precondition: improved representation of definitions hidden by customary serializations |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1006 |
*) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1007 |
|
19136 | 1008 |
fun expand_module init gen arg thy = |
18516 | 1009 |
(#modl o CodegenData.get) thy |
19136 | 1010 |
|> start_transact init (gen thy (mk_tabs thy) arg) |
18516 | 1011 |
|-> (fn x:'a => fn modl => (x, map_module (K modl) thy)); |
1012 |
||
1013 |
fun rename_inconsistent thy = |
|
18217 | 1014 |
let |
18516 | 1015 |
fun get_inconsistent thyname = |
1016 |
let |
|
1017 |
val thy = theory thyname; |
|
1018 |
fun own_tables get = |
|
1019 |
(get thy) |
|
1020 |
|> fold (Symtab.fold (Symtab.remove (K true)) o get) (Theory.parents_of thy) |
|
1021 |
|> Symtab.keys; |
|
1022 |
val names = own_tables (#2 o #types o Type.rep_tsig o Sign.tsig_of) |
|
18960 | 1023 |
@ own_tables (#2 o #constants o Consts.dest o #consts o Sign.rep_sg); |
18516 | 1024 |
fun diff names = |
1025 |
fold (fn name => |
|
1026 |
if is_prefix (op =) (NameSpace.unpack thyname) (NameSpace.unpack name) |
|
1027 |
then I |
|
1028 |
else cons (name, NameSpace.append thyname (NameSpace.base name))) names []; |
|
1029 |
in diff names end; |
|
19466 | 1030 |
val inconsistent = map get_inconsistent (ThyInfo.names ()) |> flat; |
18516 | 1031 |
fun add (src, dst) thy = |
1032 |
if (is_some oo Symtab.lookup o fst o #alias o #logic_data o CodegenData.get) thy src |
|
1033 |
then (warning ("code generator alias already defined for " ^ quote src ^ ", will not overwrite"); thy) |
|
1034 |
else add_alias (src, dst) thy |
|
1035 |
in fold add inconsistent thy end; |
|
1036 |
||
18517 | 1037 |
fun ensure_datatype_case_consts get_datatype_case_consts get_case_const_data thy = |
1038 |
let |
|
1039 |
fun ensure case_c thy = |
|
1040 |
if |
|
1041 |
Symtab.defined ((#appconst o #gens o CodegenData.get) thy) case_c |
|
1042 |
then |
|
1043 |
(warning ("case constant " ^ quote case_c ^ " already present in application table, will not overwrite"); thy) |
|
1044 |
else |
|
18702 | 1045 |
add_case_const_i get_case_const_data case_c thy; |
18517 | 1046 |
in |
1047 |
fold ensure (get_datatype_case_consts thy) thy |
|
1048 |
end; |
|
1049 |
||
19150 | 1050 |
fun codegen_term t thy = |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1051 |
thy |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1052 |
|> expand_module NONE exprsgen_term [CodegenTheorems.preprocess_term thy t] |
19150 | 1053 |
|-> (fn [t] => pair t); |
19136 | 1054 |
|
1055 |
val is_dtcon = has_nsp nsp_dtcon; |
|
1056 |
||
1057 |
fun consts_of_idfs thy = |
|
19177 | 1058 |
map (the o const_of_idf thy (mk_tabs thy)); |
19150 | 1059 |
|
1060 |
fun idfs_of_consts thy = |
|
1061 |
map (idf_of_const thy (mk_tabs thy)); |
|
1062 |
||
1063 |
val get_root_module = (#modl o CodegenData.get); |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1064 |
|
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1065 |
fun get_ml_fun_datatype thy resolv = |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1066 |
let |
19150 | 1067 |
val target_data = |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1068 |
((fn data => (the o Symtab.lookup data) "ml") o #target_data o CodegenData.get) thy; |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1069 |
in |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1070 |
CodegenSerializer.ml_fun_datatype (nsp_dtcon, nsp_class, K false) |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1071 |
((Option.map fst oo Symtab.lookup o #syntax_tyco) target_data, |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1072 |
(Option.map fst oo Symtab.lookup o #syntax_const) target_data) |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1073 |
resolv |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1074 |
end; |
18516 | 1075 |
|
1076 |
||
1077 |
(** target languages **) |
|
1078 |
||
1079 |
(* primitive definitions *) |
|
1080 |
||
18702 | 1081 |
fun read_typ thy = |
1082 |
Sign.read_typ (thy, K NONE); |
|
1083 |
||
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1084 |
fun read_const thy raw_t = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1085 |
let |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1086 |
val t = Sign.read_term thy raw_t |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1087 |
in case try dest_Const t |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1088 |
of SOME c => c |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1089 |
| NONE => error ("not a constant: " ^ Sign.string_of_term thy t) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1090 |
end; |
18517 | 1091 |
|
18963 | 1092 |
fun read_quote get reader gen raw thy = |
18912 | 1093 |
thy |
18963 | 1094 |
|> expand_module ((SOME o get) thy) |
19136 | 1095 |
(fn thy => fn tabs => gen thy tabs o single o reader thy) raw |
18912 | 1096 |
|-> (fn [x] => pair x); |
18702 | 1097 |
|
18963 | 1098 |
fun gen_add_prim prep_name prep_primdef raw_name (target, raw_primdef) thy = |
18516 | 1099 |
let |
18702 | 1100 |
val _ = if Symtab.defined ((#target_data o CodegenData.get) thy) target |
18516 | 1101 |
then () else error ("unknown target language: " ^ quote target); |
1102 |
val tabs = mk_tabs thy; |
|
1103 |
val name = prep_name thy tabs raw_name; |
|
1104 |
val primdef = prep_primdef raw_primdef; |
|
18217 | 1105 |
in |
18516 | 1106 |
thy |
18963 | 1107 |
|> map_module (CodegenThingol.add_prim name (target, primdef)) |
18217 | 1108 |
end; |
1109 |
||
18516 | 1110 |
val add_prim_i = gen_add_prim ((K o K) I) I; |
1111 |
val add_prim_class = gen_add_prim |
|
1112 |
(fn thy => K (idf_of_name thy nsp_class o Sign.intern_class thy)) |
|
18963 | 1113 |
CodegenSerializer.parse_targetdef; |
18516 | 1114 |
val add_prim_tyco = gen_add_prim |
1115 |
(fn thy => K (idf_of_name thy nsp_tyco o Sign.intern_type thy)) |
|
18963 | 1116 |
CodegenSerializer.parse_targetdef; |
18516 | 1117 |
val add_prim_const = gen_add_prim |
18517 | 1118 |
(fn thy => fn tabs => idf_of_const thy tabs o read_const thy) |
18963 | 1119 |
CodegenSerializer.parse_targetdef; |
18516 | 1120 |
|
19008 | 1121 |
val ensure_prim = map_module oo CodegenThingol.ensure_prim; |
18217 | 1122 |
|
18517 | 1123 |
|
18217 | 1124 |
(* syntax *) |
1125 |
||
18865 | 1126 |
fun gen_add_syntax_class prep_class class target pretty thy = |
1127 |
thy |
|
1128 |
|> map_codegen_data |
|
1129 |
(fn (modl, gens, target_data, logic_data) => |
|
1130 |
(modl, gens, |
|
1131 |
target_data |> Symtab.map_entry target |
|
1132 |
(map_target_data |
|
1133 |
(fn (syntax_class, syntax_tyco, syntax_const) => |
|
1134 |
(syntax_class |
|
1135 |
|> Symtab.update (prep_class thy class, pretty), syntax_tyco, syntax_const))), |
|
1136 |
logic_data)); |
|
1137 |
||
1138 |
val add_syntax_class = gen_add_syntax_class Sign.intern_class; |
|
1139 |
||
18963 | 1140 |
fun parse_syntax_tyco raw_tyco = |
18217 | 1141 |
let |
18963 | 1142 |
fun check_tyco thy tyco = |
1143 |
if Sign.declared_tyname thy tyco |
|
1144 |
then tyco |
|
1145 |
else error ("no such type constructor: " ^ quote tyco); |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1146 |
fun prep_tyco thy raw_tyco = |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1147 |
raw_tyco |
18963 | 1148 |
|> Sign.intern_type thy |
1149 |
|> check_tyco thy |
|
1150 |
|> idf_of_name thy nsp_tyco; |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1151 |
fun no_args_tyco thy raw_tyco = |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1152 |
AList.lookup (op =) ((NameSpace.dest_table o #types o Type.rep_tsig o Sign.tsig_of) thy) |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1153 |
(Sign.intern_type thy raw_tyco) |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1154 |
|> (fn SOME ((Type.LogicalType i), _) => i); |
18963 | 1155 |
fun mk reader target thy = |
18702 | 1156 |
let |
18756 | 1157 |
val _ = get_serializer target; |
18702 | 1158 |
val tyco = prep_tyco thy raw_tyco; |
1159 |
in |
|
1160 |
thy |
|
1161 |
|> ensure_prim tyco target |
|
1162 |
|> reader |
|
1163 |
|-> (fn pretty => map_codegen_data |
|
1164 |
(fn (modl, gens, target_data, logic_data) => |
|
1165 |
(modl, gens, |
|
1166 |
target_data |> Symtab.map_entry target |
|
1167 |
(map_target_data |
|
18865 | 1168 |
(fn (syntax_class, syntax_tyco, syntax_const) => |
1169 |
(syntax_class, syntax_tyco |> Symtab.update |
|
18702 | 1170 |
(tyco, (pretty, stamp ())), |
1171 |
syntax_const))), |
|
1172 |
logic_data))) |
|
1173 |
end; |
|
18217 | 1174 |
in |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1175 |
CodegenSerializer.parse_syntax (fn thy => no_args_tyco thy raw_tyco) |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1176 |
(read_quote (fn thy => prep_tyco thy raw_tyco) read_typ exprsgen_type) |
18702 | 1177 |
#-> (fn reader => pair (mk reader)) |
18217 | 1178 |
end; |
1179 |
||
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1180 |
fun add_pretty_syntax_const c target pretty = |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1181 |
map_codegen_data |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1182 |
(fn (modl, gens, target_data, logic_data) => |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1183 |
(modl, gens, |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1184 |
target_data |> Symtab.map_entry target |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1185 |
(map_target_data |
18865 | 1186 |
(fn (syntax_class, syntax_tyco, syntax_const) => |
1187 |
(syntax_class, syntax_tyco, |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1188 |
syntax_const |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1189 |
|> Symtab.update |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1190 |
(c, (pretty, stamp ()))))), |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1191 |
logic_data)); |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1192 |
|
18963 | 1193 |
fun parse_syntax_const raw_const = |
18217 | 1194 |
let |
18963 | 1195 |
fun prep_const thy raw_const = |
1196 |
idf_of_const thy (mk_tabs thy) (read_const thy raw_const); |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1197 |
fun no_args_const thy raw_const = |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1198 |
(length o fst o strip_type o snd o read_const thy) raw_const; |
18963 | 1199 |
fun mk reader target thy = |
18702 | 1200 |
let |
18756 | 1201 |
val _ = get_serializer target; |
18963 | 1202 |
val c = prep_const thy raw_const; |
18702 | 1203 |
in |
1204 |
thy |
|
1205 |
|> ensure_prim c target |
|
1206 |
|> reader |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1207 |
|-> (fn pretty => add_pretty_syntax_const c target pretty) |
18702 | 1208 |
end; |
18217 | 1209 |
in |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1210 |
CodegenSerializer.parse_syntax (fn thy => no_args_const thy raw_const) |
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1211 |
(read_quote (fn thy => prep_const thy raw_const) Sign.read_term exprsgen_term) |
18702 | 1212 |
#-> (fn reader => pair (mk reader)) |
18217 | 1213 |
end; |
1214 |
||
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1215 |
fun add_pretty_list raw_nil raw_cons (target, seri) thy = |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1216 |
let |
18756 | 1217 |
val _ = get_serializer target; |
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1218 |
val tabs = mk_tabs thy; |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1219 |
fun mk_const raw_name = |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1220 |
let |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1221 |
val name = Sign.intern_const thy raw_name; |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1222 |
in idf_of_const thy tabs (name, Sign.the_const_type thy name) end; |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1223 |
val nil' = mk_const raw_nil; |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1224 |
val cons' = mk_const raw_cons; |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1225 |
val pr' = CodegenSerializer.pretty_list nil' cons' seri; |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1226 |
in |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1227 |
thy |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1228 |
|> ensure_prim cons' target |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1229 |
|> add_pretty_syntax_const cons' target pr' |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1230 |
end; |
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1231 |
|
18217 | 1232 |
|
18516 | 1233 |
|
18756 | 1234 |
(** toplevel interface **) |
1235 |
||
1236 |
local |
|
19150 | 1237 |
|
18756 | 1238 |
fun generate_code (SOME raw_consts) thy = |
19150 | 1239 |
let |
18756 | 1240 |
val consts = map (read_const thy) raw_consts; |
1241 |
in |
|
1242 |
thy |
|
19136 | 1243 |
|> expand_module NONE (fold_map oo ensure_def_const) consts |
18756 | 1244 |
|-> (fn cs => pair (SOME cs)) |
1245 |
end |
|
1246 |
| generate_code NONE thy = |
|
1247 |
(NONE, thy); |
|
1248 |
||
1249 |
fun serialize_code target seri raw_consts thy = |
|
18217 | 1250 |
let |
18756 | 1251 |
fun serialize cs thy = |
18702 | 1252 |
let |
18756 | 1253 |
val module = (#modl o CodegenData.get) thy; |
18702 | 1254 |
val target_data = |
1255 |
thy |
|
1256 |
|> CodegenData.get |
|
1257 |
|> #target_data |
|
1258 |
|> (fn data => (the oo Symtab.lookup) data target); |
|
18756 | 1259 |
in (seri ( |
18865 | 1260 |
(Symtab.lookup o #syntax_class) target_data, |
18702 | 1261 |
(Option.map fst oo Symtab.lookup o #syntax_tyco) target_data, |
1262 |
(Option.map fst oo Symtab.lookup o #syntax_const) target_data |
|
18850 | 1263 |
) cs module : unit; thy) end; |
18217 | 1264 |
in |
1265 |
thy |
|
18756 | 1266 |
|> generate_code raw_consts |
1267 |
|-> (fn cs => serialize cs) |
|
18217 | 1268 |
end; |
1269 |
||
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1270 |
fun purge_consts raw_ts thy = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1271 |
let |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1272 |
val cs = map (read_const thy) raw_ts; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1273 |
in fold CodegenTheorems.purge_defs cs thy end; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1274 |
|
18217 | 1275 |
structure P = OuterParse |
1276 |
and K = OuterKeyword |
|
1277 |
||
1278 |
in |
|
1279 |
||
18850 | 1280 |
val (generateK, serializeK, |
18517 | 1281 |
primclassK, primtycoK, primconstK, |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1282 |
syntax_classK, syntax_tycoK, syntax_constK, |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1283 |
purgeK, aliasK) = |
18850 | 1284 |
("code_generate", "code_serialize", |
18517 | 1285 |
"code_primclass", "code_primtyco", "code_primconst", |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1286 |
"code_syntax_class", "code_syntax_tyco", "code_syntax_const", |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1287 |
"code_purge", "code_alias"); |
18335 | 1288 |
|
18217 | 1289 |
val generateP = |
18282 | 1290 |
OuterSyntax.command generateK "generate executable code for constants" K.thy_decl ( |
19136 | 1291 |
Scan.repeat1 P.term |
18756 | 1292 |
>> (fn raw_consts => |
1293 |
Toplevel.theory (generate_code (SOME raw_consts) #> snd)) |
|
18217 | 1294 |
); |
1295 |
||
1296 |
val serializeP = |
|
18282 | 1297 |
OuterSyntax.command serializeK "serialize executable code for constants" K.thy_decl ( |
18217 | 1298 |
P.name |
19136 | 1299 |
-- Scan.option (Scan.repeat1 P.term) |
18756 | 1300 |
#-> (fn (target, raw_consts) => |
18850 | 1301 |
P.$$$ "(" |
1302 |
|-- get_serializer target |
|
1303 |
--| P.$$$ ")" |
|
18756 | 1304 |
>> (fn seri => |
1305 |
Toplevel.theory (serialize_code target seri raw_consts) |
|
1306 |
)) |
|
18217 | 1307 |
); |
1308 |
||
18517 | 1309 |
val primclassP = |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1310 |
OuterSyntax.command primclassK "define target-language specific class" K.thy_decl ( |
18517 | 1311 |
P.xname |
1312 |
-- Scan.repeat1 (P.name -- P.text) |
|
18963 | 1313 |
>> (fn (raw_class, primdefs) => |
1314 |
(Toplevel.theory oo fold) (add_prim_class raw_class) primdefs) |
|
18517 | 1315 |
); |
1316 |
||
1317 |
val primtycoP = |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1318 |
OuterSyntax.command primtycoK "define target-language specific type constructor" K.thy_decl ( |
18517 | 1319 |
P.xname |
1320 |
-- Scan.repeat1 (P.name -- P.text) |
|
18963 | 1321 |
>> (fn (raw_tyco, primdefs) => |
1322 |
(Toplevel.theory oo fold) (add_prim_tyco raw_tyco) primdefs) |
|
18517 | 1323 |
); |
1324 |
||
1325 |
val primconstP = |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1326 |
OuterSyntax.command primconstK "define target-language specific constant" K.thy_decl ( |
19136 | 1327 |
P.term |
18517 | 1328 |
-- Scan.repeat1 (P.name -- P.text) |
18963 | 1329 |
>> (fn (raw_const, primdefs) => |
1330 |
(Toplevel.theory oo fold) (add_prim_const raw_const) primdefs) |
|
18517 | 1331 |
); |
1332 |
||
18865 | 1333 |
val syntax_classP = |
1334 |
OuterSyntax.command syntax_tycoK "define code syntax for class" K.thy_decl ( |
|
1335 |
Scan.repeat1 ( |
|
1336 |
P.xname |
|
1337 |
-- Scan.repeat1 ( |
|
1338 |
P.name -- P.string |
|
1339 |
) |
|
1340 |
) |
|
1341 |
>> (Toplevel.theory oo fold) (fn (raw_class, syns) => |
|
1342 |
fold (fn (target, p) => add_syntax_class raw_class target p) syns) |
|
1343 |
); |
|
1344 |
||
18217 | 1345 |
val syntax_tycoP = |
1346 |
OuterSyntax.command syntax_tycoK "define code syntax for type constructor" K.thy_decl ( |
|
18702 | 1347 |
Scan.repeat1 ( |
1348 |
P.xname |
|
18963 | 1349 |
#-> (fn raw_tyco => Scan.repeat1 ( |
1350 |
P.name -- parse_syntax_tyco raw_tyco |
|
1351 |
)) |
|
18702 | 1352 |
) |
18963 | 1353 |
>> (Toplevel.theory oo fold o fold) |
1354 |
(fn (target, modifier) => modifier target) |
|
18217 | 1355 |
); |
1356 |
||
1357 |
val syntax_constP = |
|
1358 |
OuterSyntax.command syntax_constK "define code syntax for constant" K.thy_decl ( |
|
18702 | 1359 |
Scan.repeat1 ( |
19136 | 1360 |
P.term |
18963 | 1361 |
#-> (fn raw_const => Scan.repeat1 ( |
1362 |
P.name -- parse_syntax_const raw_const |
|
1363 |
)) |
|
18702 | 1364 |
) |
19008 | 1365 |
>> (Toplevel.theory oo fold o fold) |
18963 | 1366 |
(fn (target, modifier) => modifier target) |
18217 | 1367 |
); |
1368 |
||
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1369 |
val purgeP = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1370 |
OuterSyntax.command purgeK "purge all defintions for constant" K.thy_decl ( |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1371 |
Scan.repeat1 P.term |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1372 |
>> (Toplevel.theory o purge_consts) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1373 |
); |
18516 | 1374 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1375 |
val aliasP = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1376 |
OuterSyntax.command aliasK "declare an alias for a theory identifier" K.thy_decl ( |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1377 |
Scan.repeat1 (P.name -- P.name) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1378 |
>> (Toplevel.theory oo fold) add_alias |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1379 |
); |
18516 | 1380 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1381 |
val _ = OuterSyntax.add_parsers [generateP, serializeP, |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1382 |
primclassP, primtycoP, primconstP, syntax_tycoP, syntax_constP, |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1383 |
purgeP, aliasP]; |
18217 | 1384 |
|
1385 |
end; (* local *) |
|
1386 |
||
1387 |
end; (* struct *) |