| author | wenzelm |
| Thu, 31 Aug 2006 22:55:49 +0200 | |
| changeset 20451 | 27ea2ba48fa3 |
| parent 20439 | 1bf42b262a38 |
| child 20456 | 42be3a46dcd8 |
| 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 |
| 20105 | 11 |
val codegen_term: term -> theory -> CodegenThingol.iterm * theory; |
| 20389 | 12 |
val eval_term: (string (*reference name!*) * 'a ref) * term |
13 |
-> theory -> 'a * theory; |
|
| 19884 | 14 |
val is_dtcon: string -> bool; |
15 |
val consts_of_idfs: theory -> string list -> (string * typ) list; |
|
16 |
val idfs_of_consts: theory -> (string * typ) list -> string list; |
|
| 19967 | 17 |
val get_root_module: theory -> CodegenThingol.module * theory; |
| 19884 | 18 |
val get_ml_fun_datatype: theory -> (string -> string) |
19 |
-> ((string * CodegenThingol.funn) list -> Pretty.T) |
|
20 |
* ((string * CodegenThingol.datatyp) list -> Pretty.T); |
|
| 18702 | 21 |
|
| 20401 | 22 |
val add_pretty_list: string -> string -> string -> (Pretty.T list -> Pretty.T) |
23 |
-> ((string -> string) * (string -> string)) option -> int * string |
|
24 |
-> theory -> theory; |
|
25 |
val add_pretty_ml_string: string -> string -> string -> string |
|
26 |
-> (string -> string) -> (string -> string) -> string -> theory -> theory; |
|
| 20175 | 27 |
val purge_code: theory -> theory; |
| 18217 | 28 |
|
| 19884 | 29 |
type appgen; |
| 20439 | 30 |
val add_appconst: string * appgen -> theory -> theory; |
| 18702 | 31 |
val appgen_default: appgen; |
| 20353 | 32 |
val appgen_rep_bin: (theory -> term -> IntInf.int) -> appgen; |
|
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
33 |
val appgen_char: (term -> int option) -> appgen; |
| 20105 | 34 |
val appgen_case: (theory -> term |
35 |
-> ((string * typ) list * ((term * typ) * (term * term) list)) option) |
|
36 |
-> appgen; |
|
37 |
val appgen_let: appgen; |
|
| 19038 | 38 |
val appgen_wfrec: appgen; |
| 18217 | 39 |
|
| 19008 | 40 |
val print_code: theory -> unit; |
| 18231 | 41 |
structure CodegenData: THEORY_DATA; |
|
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
42 |
end; |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
43 |
|
| 18217 | 44 |
structure CodegenPackage : CODEGEN_PACKAGE = |
|
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
45 |
struct |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
46 |
|
| 18850 | 47 |
open CodegenThingol; |
| 18217 | 48 |
|
| 20439 | 49 |
(** preliminaries **) |
50 |
||
| 18702 | 51 |
(* shallow name spaces *) |
| 18217 | 52 |
|
| 20216 | 53 |
val nsp_module = ""; (*a dummy by convention*) |
| 18217 | 54 |
val nsp_class = "class"; |
| 18454 | 55 |
val nsp_tyco = "tyco"; |
| 18217 | 56 |
val nsp_const = "const"; |
| 18454 | 57 |
val nsp_dtcon = "dtcon"; |
| 18217 | 58 |
val nsp_mem = "mem"; |
59 |
val nsp_inst = "inst"; |
|
| 20216 | 60 |
val nsp_eval = "EVAL"; (*only for evaluation*) |
| 18217 | 61 |
|
| 19038 | 62 |
fun add_nsp shallow name = |
63 |
name |
|
64 |
|> NameSpace.unpack |
|
65 |
|> split_last |
|
66 |
|> apsnd (single #> cons shallow) |
|
67 |
|> (op @) |
|
68 |
|> NameSpace.pack; |
|
69 |
||
70 |
fun dest_nsp nsp idf = |
|
71 |
let |
|
72 |
val idf' = NameSpace.unpack idf; |
|
73 |
val (idf'', idf_base) = split_last idf'; |
|
74 |
val (modl, shallow) = split_last idf''; |
|
75 |
in |
|
76 |
if nsp = shallow |
|
77 |
then (SOME o NameSpace.pack) (modl @ [idf_base]) |
|
78 |
else NONE |
|
79 |
end; |
|
80 |
||
| 20386 | 81 |
fun if_nsp nsp f idf = |
82 |
Option.map f (dest_nsp nsp idf); |
|
| 19956 | 83 |
|
| 18702 | 84 |
val serializers = ref ( |
85 |
Symtab.empty |
|
86 |
|> Symtab.update ( |
|
87 |
#ml CodegenSerializer.serializers |
|
88 |
|> apsnd (fn seri => seri |
|
| 20183 | 89 |
nsp_dtcon |
| 20389 | 90 |
[[nsp_module], [nsp_class, nsp_tyco], |
| 20428 | 91 |
[nsp_const, nsp_dtcon, nsp_class, nsp_mem, nsp_inst]] |
| 18702 | 92 |
) |
93 |
) |
|
94 |
|> Symtab.update ( |
|
95 |
#haskell CodegenSerializer.serializers |
|
96 |
|> apsnd (fn seri => seri |
|
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19884
diff
changeset
|
97 |
(nsp_dtcon, [nsp_module, nsp_class, nsp_tyco, nsp_dtcon]) |
| 20389 | 98 |
[[nsp_module], [nsp_class], [nsp_tyco], [nsp_const, nsp_mem], |
| 20428 | 99 |
[nsp_dtcon], [nsp_inst]] |
| 18702 | 100 |
) |
101 |
) |
|
102 |
); |
|
| 18217 | 103 |
|
104 |
||
| 20439 | 105 |
(* theory data *) |
| 18217 | 106 |
|
| 20439 | 107 |
type appgen = theory -> CodegenTheorems.thmtab -> bool * string list option |
108 |
-> (string * typ) * term list -> transact -> iterm * transact; |
|
109 |
||
110 |
type appgens = (int * (appgen * stamp)) Symtab.table; |
|
| 18217 | 111 |
|
| 20386 | 112 |
fun merge_appgens (x : appgens * appgens) = |
| 20105 | 113 |
Symtab.merge (fn ((bounds1, (_, stamp1)), (bounds2, (_, stamp2))) => |
| 20439 | 114 |
bounds1 = bounds2 andalso stamp1 = stamp2) x; |
| 18217 | 115 |
|
| 18702 | 116 |
type target_data = {
|
| 20428 | 117 |
syntax_class: ((string * (string -> string option)) * stamp) Symtab.table, |
118 |
syntax_inst: unit Symtab.table, |
|
| 18516 | 119 |
syntax_tyco: (itype CodegenSerializer.pretty_syntax * stamp) Symtab.table, |
| 20105 | 120 |
syntax_const: (iterm CodegenSerializer.pretty_syntax * stamp) Symtab.table |
| 18217 | 121 |
}; |
122 |
||
| 18702 | 123 |
fun merge_target_data |
| 20428 | 124 |
({ syntax_class = syntax_class1, syntax_inst = syntax_inst1,
|
125 |
syntax_tyco = syntax_tyco1, syntax_const = syntax_const1 }, |
|
126 |
{ syntax_class = syntax_class2, syntax_inst = syntax_inst2,
|
|
127 |
syntax_tyco = syntax_tyco2, syntax_const = syntax_const2 }) = |
|
128 |
{ syntax_class = Symtab.merge (eq_snd (op =)) (syntax_class1, syntax_class2),
|
|
129 |
syntax_inst = Symtab.merge (op =) (syntax_inst1, syntax_inst2), |
|
| 18865 | 130 |
syntax_tyco = Symtab.merge (eq_snd (op =)) (syntax_tyco1, syntax_tyco2), |
| 18702 | 131 |
syntax_const = Symtab.merge (eq_snd (op =)) (syntax_const1, syntax_const2) } : target_data; |
| 18217 | 132 |
|
133 |
structure CodegenData = TheoryDataFun |
|
134 |
(struct |
|
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19884
diff
changeset
|
135 |
val name = "Pure/codegen_package"; |
| 18217 | 136 |
type T = {
|
137 |
modl: module, |
|
| 20105 | 138 |
appgens: appgens, |
| 18702 | 139 |
target_data: target_data Symtab.table |
| 18217 | 140 |
}; |
141 |
val empty = {
|
|
142 |
modl = empty_module, |
|
| 20105 | 143 |
appgens = Symtab.empty, |
| 18702 | 144 |
target_data = |
| 18217 | 145 |
Symtab.empty |
| 18702 | 146 |
|> Symtab.fold (fn (target, _) => |
| 18865 | 147 |
Symtab.update (target, |
| 20428 | 148 |
{ syntax_class = Symtab.empty, syntax_inst = Symtab.empty,
|
149 |
syntax_tyco = Symtab.empty, syntax_const = Symtab.empty }) |
|
| 18702 | 150 |
) (! serializers) |
| 18217 | 151 |
} : T; |
152 |
val copy = I; |
|
153 |
val extend = I; |
|
154 |
fun merge _ ( |
|
| 20105 | 155 |
{ modl = modl1, appgens = appgens1,
|
| 20386 | 156 |
target_data = target_data1 }, |
| 20105 | 157 |
{ modl = modl2, appgens = appgens2,
|
| 20386 | 158 |
target_data = target_data2 } |
| 18217 | 159 |
) = {
|
160 |
modl = merge_module (modl1, modl2), |
|
| 20105 | 161 |
appgens = merge_appgens (appgens1, appgens2), |
| 19025 | 162 |
target_data = Symtab.join (K merge_target_data) (target_data1, target_data2) |
| 18217 | 163 |
}; |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
164 |
fun print thy (data : T) = |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
165 |
let |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
166 |
val module = #modl data |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
167 |
in |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
168 |
(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
|
169 |
end; |
| 18217 | 170 |
end); |
171 |
||
| 18708 | 172 |
val _ = Context.add_setup CodegenData.init; |
173 |
||
| 18217 | 174 |
fun map_codegen_data f thy = |
175 |
case CodegenData.get thy |
|
| 20386 | 176 |
of { modl, appgens, target_data } =>
|
177 |
let val (modl, appgens, target_data) = |
|
178 |
f (modl, appgens, target_data) |
|
| 20105 | 179 |
in CodegenData.put { modl = modl, appgens = appgens,
|
| 20386 | 180 |
target_data = target_data } thy end; |
| 18217 | 181 |
|
| 20439 | 182 |
fun get_serializer target = |
183 |
case Symtab.lookup (!serializers) target |
|
184 |
of SOME seri => seri |
|
185 |
| NONE => Scan.fail_with (fn _ => "Unknown code target language: " ^ quote target) (); |
|
186 |
||
187 |
fun serialize thy target seri cs = |
|
188 |
let |
|
189 |
val data = CodegenData.get thy |
|
190 |
val code = #modl data; |
|
191 |
val target_data = |
|
192 |
(the oo Symtab.lookup) (#target_data data) target; |
|
193 |
val syntax_class = #syntax_class target_data; |
|
194 |
val syntax_inst = #syntax_inst target_data; |
|
195 |
val syntax_tyco = #syntax_tyco target_data; |
|
196 |
val syntax_const = #syntax_const target_data; |
|
197 |
fun fun_of syntax = (Option.map fst oo Symtab.lookup) syntax; |
|
198 |
in |
|
199 |
seri (fun_of syntax_class, fun_of syntax_tyco, fun_of syntax_const) |
|
200 |
(Symtab.keys syntax_class @ Symtab.keys syntax_inst |
|
201 |
@ Symtab.keys syntax_tyco @ Symtab.keys syntax_const, cs) code : unit |
|
202 |
end; |
|
203 |
||
204 |
fun map_target_data target f = |
|
205 |
let |
|
206 |
val _ = get_serializer target; |
|
207 |
in |
|
208 |
map_codegen_data (fn (modl, appgens, target_data) => |
|
209 |
(modl, appgens, Symtab.map_entry target (fn { syntax_class, syntax_inst, syntax_tyco, syntax_const } =>
|
|
210 |
let |
|
211 |
val (syntax_class, syntax_inst, syntax_tyco, syntax_const) = |
|
212 |
f (syntax_class, syntax_inst, syntax_tyco, syntax_const) |
|
213 |
in {
|
|
214 |
syntax_class = syntax_class, |
|
215 |
syntax_inst = syntax_inst, |
|
216 |
syntax_tyco = syntax_tyco, |
|
217 |
syntax_const = syntax_const } : target_data |
|
218 |
end |
|
219 |
) target_data) |
|
220 |
) |
|
221 |
end; |
|
222 |
||
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
223 |
val print_code = CodegenData.print; |
| 18217 | 224 |
|
| 20439 | 225 |
fun map_module f = |
226 |
map_codegen_data (fn (modl, gens, target_data) => |
|
227 |
(f modl, gens, target_data)); |
|
228 |
||
229 |
val purge_code = map_module (K CodegenThingol.empty_module); |
|
230 |
||
231 |
fun purge_defs NONE = purge_code |
|
232 |
| purge_defs (SOME []) = I |
|
233 |
| purge_defs (SOME cs) = purge_code; |
|
| 18454 | 234 |
|
| 18865 | 235 |
|
| 20386 | 236 |
(* name handling *) |
237 |
||
238 |
fun idf_of_class thy class = |
|
239 |
CodegenNames.class thy class |
|
240 |
|> add_nsp nsp_class; |
|
241 |
||
242 |
fun class_of_idf thy = if_nsp nsp_class (CodegenNames.class_rev thy); |
|
243 |
||
244 |
fun idf_of_tyco thy tyco = |
|
245 |
CodegenNames.tyco thy tyco |
|
246 |
|> add_nsp nsp_tyco; |
|
247 |
||
248 |
fun tyco_of_idf thy = if_nsp nsp_tyco (CodegenNames.tyco_rev thy); |
|
249 |
||
250 |
fun idf_of_inst thy inst = |
|
251 |
CodegenNames.instance thy inst |
|
252 |
|> add_nsp nsp_inst; |
|
253 |
||
254 |
fun inst_of_idf thy = if_nsp nsp_inst (CodegenNames.instance_rev thy); |
|
255 |
||
256 |
fun idf_of_const thy thmtab (c_ty as (c, ty)) = |
|
257 |
if is_some (CodegenTheorems.get_dtyp_of_cons thmtab c_ty) then |
|
258 |
CodegenNames.const thy c_ty |
|
259 |
|> add_nsp nsp_dtcon |
|
260 |
else if (is_some o CodegenConsts.class_of_classop thy o CodegenConsts.typinst_of_typ thy) c_ty then |
|
261 |
CodegenNames.const thy c_ty |
|
262 |
|> add_nsp nsp_mem |
|
263 |
else |
|
264 |
CodegenNames.const thy c_ty |
|
265 |
|> add_nsp nsp_const; |
|
266 |
||
| 20428 | 267 |
fun idf_of_classop thy c_ty = |
268 |
CodegenNames.const thy c_ty |
|
269 |
|> add_nsp nsp_mem; |
|
270 |
||
| 20386 | 271 |
fun const_of_idf thy idf = |
272 |
case dest_nsp nsp_const idf |
|
273 |
of SOME c => CodegenNames.const_rev thy c |> SOME |
|
274 |
| _ => (case dest_nsp nsp_dtcon idf |
|
275 |
of SOME c => CodegenNames.const_rev thy c |> SOME |
|
276 |
| _ => (case dest_nsp nsp_mem idf |
|
277 |
of SOME c => CodegenNames.const_rev thy c |> SOME |
|
278 |
| _ => NONE)); |
|
279 |
||
280 |
||
| 18865 | 281 |
|
| 20439 | 282 |
(** code extraction **) |
| 18865 | 283 |
|
| 20386 | 284 |
(* extraction kernel *) |
| 18865 | 285 |
|
| 20439 | 286 |
fun check_strict thy f x (false, _) = |
| 19884 | 287 |
false |
| 20439 | 288 |
| check_strict thy f x (_, SOME targets) = |
| 19884 | 289 |
exists ( |
| 20389 | 290 |
is_none o (fn tab => Symtab.lookup tab x) o f o the |
291 |
o (Symtab.lookup ((#target_data o CodegenData.get) thy)) |
|
| 19884 | 292 |
) targets |
| 20439 | 293 |
| check_strict thy f x (true, _) = |
| 19884 | 294 |
true; |
295 |
||
| 20439 | 296 |
fun no_strict (_, targets) = (false, targets); |
| 20386 | 297 |
|
| 20439 | 298 |
(*FIXME: provide a unified view on this in codegen_consts.ML*) |
299 |
fun sortlookups_const thy thmtab (c, ty_ctxt) = |
|
| 20386 | 300 |
let |
| 20439 | 301 |
val ty_decl = case CodegenTheorems.get_fun_thms thmtab (c, ty_ctxt) |
| 20386 | 302 |
of thms as thm :: _ => CodegenTheorems.extr_typ thy thm |
303 |
| [] => (case AxClass.class_of_param thy c |
|
304 |
of SOME class => (case ClassPackage.the_consts_sign thy class of (v, cs) => |
|
305 |
(Logic.varifyT o map_type_tfree (fn u as (w, _) => |
|
306 |
if w = v then TFree (v, [class]) else TFree u)) |
|
307 |
((the o AList.lookup (op =) cs) c)) |
|
308 |
| NONE => Sign.the_const_type thy c); |
|
309 |
in |
|
310 |
Vartab.empty |
|
| 20439 | 311 |
|> Sign.typ_match thy (ty_decl, ty_ctxt) |
| 20386 | 312 |
|> Vartab.dest |
| 20389 | 313 |
|> map (fn (_, (sort, ty)) => ClassPackage.sortlookup thy (ty, sort)) |
| 20386 | 314 |
|> filter_out null |
315 |
end; |
|
|
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
316 |
|
| 20439 | 317 |
fun ensure_def_class thy thmtab strct cls trns = |
| 18454 | 318 |
let |
| 20439 | 319 |
fun defgen_class thy thmtab strct cls trns = |
| 20386 | 320 |
case class_of_idf thy cls |
| 18865 | 321 |
of SOME cls => |
322 |
let |
|
| 19283 | 323 |
val (v, cs) = (ClassPackage.the_consts_sign thy) cls; |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19884
diff
changeset
|
324 |
val sortctxts = map (ClassPackage.sortcontext_of_typ thy o snd) cs; |
| 20386 | 325 |
val idfs = map (idf_of_const thy thmtab) cs; |
| 18865 | 326 |
in |
327 |
trns |
|
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
328 |
|> debug_msg (fn _ => "trying defgen class declaration for " ^ quote cls) |
| 20439 | 329 |
|> fold_map (ensure_def_class thy thmtab strct) (ClassPackage.the_superclasses thy cls) |
330 |
||>> (fold_map (exprgen_type thy thmtab strct) o map snd) cs |
|
331 |
||>> (fold_map o fold_map) (exprgen_tyvar_sort thy thmtab strct) sortctxts |
|
| 18865 | 332 |
|-> (fn ((supcls, memtypes), sortctxts) => succeed |
| 19283 | 333 |
(Class (supcls, (unprefix "'" v, idfs ~~ (sortctxts ~~ memtypes))))) |
| 18865 | 334 |
end |
335 |
| _ => |
|
336 |
trns |
|
| 20389 | 337 |
|> fail ("No class definition found for " ^ quote cls);
|
| 20386 | 338 |
val cls' = idf_of_class thy cls; |
| 18454 | 339 |
in |
340 |
trns |
|
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
341 |
|> debug_msg (fn _ => "generating class " ^ quote cls) |
| 20439 | 342 |
|> ensure_def (defgen_class thy thmtab strct) true ("generating class " ^ quote cls) cls'
|
| 18865 | 343 |
|> pair cls' |
344 |
end |
|
| 20439 | 345 |
and ensure_def_tyco thy thmtab strct tyco trns = |
| 18865 | 346 |
let |
| 20386 | 347 |
val tyco' = idf_of_tyco thy tyco; |
| 20439 | 348 |
val strict = check_strict thy #syntax_tyco tyco' strct; |
349 |
fun defgen_datatype thy thmtab strct dtco trns = |
|
| 20386 | 350 |
case tyco_of_idf thy dtco |
| 18963 | 351 |
of SOME dtco => |
| 20353 | 352 |
(case CodegenTheorems.get_dtyp_spec thmtab dtco |
353 |
of SOME (vars, cos) => |
|
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
354 |
trns |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
355 |
|> debug_msg (fn _ => "trying defgen datatype for " ^ quote dtco) |
| 20439 | 356 |
|> fold_map (exprgen_tyvar_sort thy thmtab strct) vars |
|
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
357 |
||>> fold_map (fn (c, tys) => |
| 20439 | 358 |
fold_map (exprgen_type thy thmtab strct) tys |
| 20386 | 359 |
#-> (fn tys' => pair (idf_of_const thy thmtab (c, tys ---> Type (dtco, map TFree vars)), tys'))) cos |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
360 |
|-> (fn (vars, cos) => succeed (Datatype |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
361 |
(vars, cos))) |
| 18963 | 362 |
| NONE => |
363 |
trns |
|
| 20389 | 364 |
|> fail ("No datatype found for " ^ quote dtco))
|
| 18963 | 365 |
| NONE => |
366 |
trns |
|
| 20389 | 367 |
|> fail ("Not a type constructor: " ^ quote dtco)
|
| 18865 | 368 |
in |
369 |
trns |
|
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
370 |
|> debug_msg (fn _ => "generating type constructor " ^ quote tyco) |
| 20439 | 371 |
|> ensure_def (defgen_datatype thy thmtab strct) strict |
| 20389 | 372 |
("generating type constructor " ^ quote tyco) tyco'
|
| 18865 | 373 |
|> pair tyco' |
374 |
end |
|
| 20439 | 375 |
and exprgen_tyvar_sort thy thmtab strct (v, sort) trns = |
| 18516 | 376 |
trns |
| 20439 | 377 |
|> fold_map (ensure_def_class thy thmtab strct) (ClassPackage.operational_sort_of thy sort) |
| 18865 | 378 |
|-> (fn sort => pair (unprefix "'" v, sort)) |
| 20439 | 379 |
and exprgen_type thy thmtab strct (TVar _) trns = |
| 20389 | 380 |
error "TVar encountered in typ during code generation" |
| 20439 | 381 |
| exprgen_type thy thmtab strct (TFree v_s) trns = |
| 18516 | 382 |
trns |
| 20439 | 383 |
|> exprgen_tyvar_sort thy thmtab strct v_s |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
384 |
|-> (fn (v, sort) => pair (ITyVar v)) |
| 20439 | 385 |
| exprgen_type thy thmtab strct (Type ("fun", [t1, t2])) trns =
|
| 18516 | 386 |
trns |
| 20439 | 387 |
|> exprgen_type thy thmtab strct t1 |
388 |
||>> exprgen_type thy thmtab strct t2 |
|
| 18516 | 389 |
|-> (fn (t1', t2') => pair (t1' `-> t2')) |
| 20439 | 390 |
| exprgen_type thy thmtab strct (Type (tyco, tys)) trns = |
| 18516 | 391 |
trns |
| 20439 | 392 |
|> ensure_def_tyco thy thmtab strct tyco |
393 |
||>> fold_map (exprgen_type thy thmtab strct) tys |
|
|
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
394 |
|-> (fn (tyco, tys) => pair (tyco `%% tys)); |
| 18516 | 395 |
|
| 20439 | 396 |
fun exprgen_classlookup thy thmtab strct (ClassPackage.Instance (inst, ls)) trns = |
| 18517 | 397 |
trns |
| 20439 | 398 |
|> ensure_def_inst thy thmtab strct inst |
399 |
||>> (fold_map o fold_map) (exprgen_classlookup thy thmtab strct) ls |
|
| 18885 | 400 |
|-> (fn (inst, ls) => pair (Instance (inst, ls))) |
| 20439 | 401 |
| exprgen_classlookup thy thmtab strct (ClassPackage.Lookup (clss, (v, (i, j)))) trns = |
| 18516 | 402 |
trns |
| 20439 | 403 |
|> fold_map (ensure_def_class thy thmtab strct) clss |
| 19253 | 404 |
|-> (fn clss => pair (Lookup (clss, (v |> unprefix "'", if j = 1 then ~1 else i)))) |
| 20439 | 405 |
and mk_fun thy thmtab strct (c, ty) trns = |
| 20386 | 406 |
case CodegenTheorems.get_fun_thms thmtab (c, ty) |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19884
diff
changeset
|
407 |
of eq_thms as eq_thm :: _ => |
| 18865 | 408 |
let |
| 19884 | 409 |
val msg = cat_lines ("generating code for theorems " :: map string_of_thm eq_thms);
|
| 20386 | 410 |
val ty = (Logic.unvarifyT o CodegenTheorems.extr_typ thy) eq_thm |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19884
diff
changeset
|
411 |
val sortcontext = ClassPackage.sortcontext_of_typ thy ty; |
| 18865 | 412 |
fun dest_eqthm eq_thm = |
413 |
let |
|
414 |
val ((t, args), rhs) = |
|
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19884
diff
changeset
|
415 |
(apfst strip_comb o Logic.dest_equals o Logic.legacy_unvarify o prop_of) eq_thm; |
| 18865 | 416 |
in case t |
417 |
of Const (c', _) => if c' = c then (args, rhs) |
|
| 20389 | 418 |
else error ("Illegal function equation for " ^ quote c
|
| 18865 | 419 |
^ ", actually defining " ^ quote c') |
| 20389 | 420 |
| _ => error ("Illegal function equation for " ^ quote c)
|
| 18865 | 421 |
end; |
|
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
422 |
fun exprgen_eq (args, rhs) trns = |
|
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
423 |
trns |
| 20439 | 424 |
|> fold_map (exprgen_term thy thmtab strct) args |
425 |
||>> exprgen_term thy thmtab strct rhs; |
|
| 20105 | 426 |
fun checkvars (args, rhs) = |
427 |
if CodegenThingol.vars_distinct args then (args, rhs) |
|
| 20389 | 428 |
else error ("Repeated variables on left hand side of function")
|
| 18865 | 429 |
in |
| 18517 | 430 |
trns |
| 19884 | 431 |
|> message msg (fn trns => trns |
|
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
432 |
|> fold_map (exprgen_eq o dest_eqthm) eq_thms |
| 20105 | 433 |
|-> (fn eqs => pair (map checkvars eqs)) |
| 20439 | 434 |
||>> fold_map (exprgen_tyvar_sort thy thmtab strct) sortcontext |
435 |
||>> exprgen_type thy thmtab strct ty |
|
| 20389 | 436 |
|-> (fn ((eqs, sortctxt), ty) => (pair o SOME) ((eqs, (sortctxt, ty)), |
437 |
map snd sortcontext))) |
|
| 18865 | 438 |
end |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19884
diff
changeset
|
439 |
| [] => (NONE, trns) |
| 20439 | 440 |
and ensure_def_inst thy thmtab strct (cls, tyco) trns = |
| 18865 | 441 |
let |
| 20439 | 442 |
fun defgen_inst thy thmtab strct inst trns = |
| 20386 | 443 |
case inst_of_idf thy inst |
| 19956 | 444 |
of SOME (class, tyco) => |
| 18865 | 445 |
let |
| 18885 | 446 |
val (arity, memdefs) = ClassPackage.the_inst_sign thy (class, tyco); |
| 20389 | 447 |
val (_, members) = ClassPackage.the_consts_sign thy class; |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19884
diff
changeset
|
448 |
val arity_typ = Type (tyco, (map TFree arity)); |
| 20389 | 449 |
val operational_arity = map_filter (fn (v, sort) => |
450 |
case ClassPackage.operational_sort_of thy sort |
|
451 |
of [] => NONE |
|
452 |
| sort => SOME (v, sort)) arity; |
|
| 18865 | 453 |
fun gen_suparity supclass trns = |
454 |
trns |
|
| 20439 | 455 |
|> ensure_def_class thy thmtab strct supclass |
456 |
||>> fold_map (exprgen_classlookup thy thmtab strct) |
|
| 20386 | 457 |
(ClassPackage.sortlookup thy (arity_typ, [supclass])); |
458 |
fun gen_membr ((m0, ty0), (m, ty)) trns = |
|
| 18865 | 459 |
trns |
| 20439 | 460 |
|> ensure_def_const thy thmtab strct (m0, ty0) |
461 |
||>> exprgen_term thy thmtab strct (Const (m, ty)); |
|
| 18865 | 462 |
in |
463 |
trns |
|
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
464 |
|> debug_msg (fn _ => "trying defgen class instance for (" ^ quote cls
|
| 18865 | 465 |
^ ", " ^ quote tyco ^ ")") |
| 20439 | 466 |
|> ensure_def_class thy thmtab strct class |
467 |
||>> ensure_def_tyco thy thmtab strct tyco |
|
468 |
||>> fold_map (exprgen_tyvar_sort thy thmtab strct) arity |
|
| 18885 | 469 |
||>> fold_map gen_suparity (ClassPackage.the_superclasses thy class) |
| 20389 | 470 |
||>> fold_map gen_membr (members ~~ memdefs) |
| 18885 | 471 |
|-> (fn ((((class, tyco), arity), suparities), memdefs) => |
| 20389 | 472 |
succeed (Classinst ((class, (tyco, arity)), (suparities, memdefs)))) |
| 18865 | 473 |
end |
474 |
| _ => |
|
| 20389 | 475 |
trns |> fail ("No class instance found for " ^ quote inst);
|
| 20386 | 476 |
val inst = idf_of_inst thy (cls, tyco); |
| 18865 | 477 |
in |
478 |
trns |
|
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
479 |
|> debug_msg (fn _ => "generating instance " ^ quote cls ^ " / " ^ quote tyco) |
| 20439 | 480 |
|> ensure_def (defgen_inst thy thmtab strct) true |
| 18865 | 481 |
("generating instance " ^ quote cls ^ " / " ^ quote tyco) inst
|
482 |
|> pair inst |
|
483 |
end |
|
| 20439 | 484 |
and ensure_def_const thy thmtab strct (c, ty) trns = |
| 18865 | 485 |
let |
| 20439 | 486 |
fun defgen_datatypecons thy thmtab strct co trns = |
| 20386 | 487 |
case CodegenTheorems.get_dtyp_of_cons thmtab ((the o const_of_idf thy) co) |
488 |
of SOME tyco => |
|
| 18865 | 489 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
490 |
|> debug_msg (fn _ => "trying defgen datatype constructor for " ^ quote co) |
| 20439 | 491 |
|> ensure_def_tyco thy thmtab strct tyco |
| 20386 | 492 |
|-> (fn _ => succeed Bot) |
| 18865 | 493 |
| _ => |
494 |
trns |
|
| 20389 | 495 |
|> fail ("Not a datatype constructor: "
|
496 |
^ (quote o CodegenConsts.string_of_const_typ thy) (c, ty)); |
|
| 20439 | 497 |
fun defgen_clsmem thy thmtab strct m trns = |
| 20386 | 498 |
case CodegenConsts.class_of_classop thy |
499 |
((CodegenConsts.typinst_of_typ thy o the o const_of_idf thy) m) |
|
500 |
of SOME class => |
|
501 |
trns |
|
502 |
|> debug_msg (fn _ => "trying defgen class member for " ^ quote m) |
|
| 20439 | 503 |
|> ensure_def_class thy thmtab strct class |
| 20386 | 504 |
|-> (fn _ => succeed Bot) |
505 |
| _ => |
|
| 20389 | 506 |
trns |> fail ("No class found for " ^ (quote o CodegenConsts.string_of_const_typ thy) (c, ty))
|
| 20439 | 507 |
fun defgen_funs thy thmtab strct c' trns = |
| 20386 | 508 |
trns |
| 20439 | 509 |
|> mk_fun thy thmtab strct ((the o const_of_idf thy) c') |
| 20386 | 510 |
|-> (fn SOME (funn, _) => succeed (Fun funn) |
| 20389 | 511 |
| NONE => fail ("No defining equations found for "
|
512 |
^ (quote o CodegenConsts.string_of_const_typ thy) (c, ty))) |
|
| 20439 | 513 |
fun get_defgen thmtab strct idf strict = |
| 20386 | 514 |
if (is_some oo dest_nsp) nsp_const idf |
| 20439 | 515 |
then defgen_funs thy thmtab strct strict |
| 20386 | 516 |
else if (is_some oo dest_nsp) nsp_mem idf |
| 20439 | 517 |
then defgen_clsmem thy thmtab strct strict |
| 20386 | 518 |
else if (is_some oo dest_nsp) nsp_dtcon idf |
| 20439 | 519 |
then defgen_datatypecons thy thmtab strct strict |
| 20389 | 520 |
else error ("Illegal shallow name space for constant: " ^ quote idf);
|
| 20386 | 521 |
val idf = idf_of_const thy thmtab (c, ty); |
| 20439 | 522 |
val strict = check_strict thy #syntax_const idf strct; |
| 18865 | 523 |
in |
524 |
trns |
|
| 20389 | 525 |
|> debug_msg (fn _ => "generating constant " |
526 |
^ (quote o CodegenConsts.string_of_const_typ thy) (c, ty)) |
|
| 20439 | 527 |
|> ensure_def (get_defgen thmtab strct idf) strict ("generating constant "
|
| 20389 | 528 |
^ CodegenConsts.string_of_const_typ thy (c, ty)) idf |
| 18963 | 529 |
|> pair idf |
| 18865 | 530 |
end |
| 20439 | 531 |
and exprgen_term thy thmtab strct (Const (f, ty)) trns = |
| 18517 | 532 |
trns |
| 20439 | 533 |
|> appgen thy thmtab strct ((f, ty), []) |
| 18516 | 534 |
|-> (fn e => pair e) |
| 20439 | 535 |
| exprgen_term thy thmtab strct (Var _) trns = |
| 20389 | 536 |
error "Var encountered in term during code generation" |
| 20439 | 537 |
| exprgen_term thy thmtab strct (Free (v, ty)) trns = |
| 18516 | 538 |
trns |
| 20439 | 539 |
|> exprgen_type thy thmtab strct ty |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
540 |
|-> (fn ty => pair (IVar v)) |
| 20439 | 541 |
| exprgen_term thy thmtab strct (Abs (raw_v, ty, raw_t)) trns = |
| 19136 | 542 |
let |
| 20386 | 543 |
val (v, t) = Syntax.variant_abs (CodegenNames.purify_var raw_v, ty, raw_t); |
| 19136 | 544 |
in |
545 |
trns |
|
| 20439 | 546 |
|> exprgen_type thy thmtab strct ty |
547 |
||>> exprgen_term thy thmtab strct t |
|
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
548 |
|-> (fn (ty, e) => pair ((v, ty) `|-> e)) |
| 19136 | 549 |
end |
| 20439 | 550 |
| exprgen_term thy thmtab strct (t as t1 $ t2) trns = |
| 18516 | 551 |
let |
552 |
val (t', ts) = strip_comb t |
|
553 |
in case t' |
|
554 |
of Const (f, ty) => |
|
555 |
trns |
|
| 20439 | 556 |
|> appgen thy thmtab strct ((f, ty), ts) |
| 18516 | 557 |
|-> (fn e => pair e) |
558 |
| _ => |
|
559 |
trns |
|
| 20439 | 560 |
|> exprgen_term thy thmtab strct t' |
561 |
||>> fold_map (exprgen_term thy thmtab strct) ts |
|
| 18516 | 562 |
|-> (fn (e, es) => pair (e `$$ es)) |
| 18865 | 563 |
end |
| 20439 | 564 |
and appgen_default thy thmtab strct ((c, ty), ts) trns = |
| 18865 | 565 |
trns |
| 20439 | 566 |
|> ensure_def_const thy thmtab strct (c, ty) |
567 |
||>> exprgen_type thy thmtab strct ty |
|
568 |
||>> (fold_map o fold_map) (exprgen_classlookup thy thmtab strct) |
|
| 20386 | 569 |
(sortlookups_const thy thmtab (c, ty)) |
| 20439 | 570 |
||>> fold_map (exprgen_term thy thmtab strct) ts |
|
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
571 |
|-> (fn (((c, ty), ls), es) => |
| 19202 | 572 |
pair (IConst (c, (ls, ty)) `$$ es)) |
| 20439 | 573 |
and appgen thy thmtab strct ((f, ty), ts) trns = |
| 20105 | 574 |
case Symtab.lookup ((#appgens o CodegenData.get) thy) f |
575 |
of SOME (i, (ag, _)) => |
|
576 |
if length ts < i then |
|
| 18865 | 577 |
let |
| 20105 | 578 |
val tys = Library.take (i - length ts, ((fst o strip_type) ty)); |
|
20192
956cd30ef3be
renamed Name.give_names to Name.names and moved Name.alphanum to Symbol.alphanum
haftmann
parents:
20191
diff
changeset
|
579 |
val vs = Name.names (Name.declare f Name.context) "a" tys; |
| 18865 | 580 |
in |
581 |
trns |
|
| 20439 | 582 |
|> fold_map (exprgen_type thy thmtab strct) tys |
583 |
||>> ag thy thmtab strct ((f, ty), ts @ map Free vs) |
|
| 20105 | 584 |
|-> (fn (tys, e) => pair (map2 (fn (v, _) => pair v) vs tys `|--> e)) |
| 18865 | 585 |
end |
| 20105 | 586 |
else if length ts > i then |
| 18865 | 587 |
trns |
| 20439 | 588 |
|> ag thy thmtab strct ((f, ty), Library.take (i, ts)) |
589 |
||>> fold_map (exprgen_term thy thmtab strct) (Library.drop (i, ts)) |
|
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
590 |
|-> (fn (e, es) => pair (e `$$ es)) |
| 18865 | 591 |
else |
592 |
trns |
|
| 20439 | 593 |
|> ag thy thmtab strct ((f, ty), ts) |
| 18865 | 594 |
| NONE => |
595 |
trns |
|
| 20439 | 596 |
|> appgen_default thy thmtab strct ((f, ty), ts); |
| 18516 | 597 |
|
| 18702 | 598 |
|
| 20439 | 599 |
(* parametrized application generators, for instantiation in object logic *) |
600 |
(* (axiomatic extensions of extraction kernel *) |
|
| 18217 | 601 |
|
| 20439 | 602 |
fun appgen_rep_bin int_of_numeral thy thmtab strct (app as (c as (_, ty), [bin])) trns = |
| 20353 | 603 |
case try (int_of_numeral thy) bin |
604 |
of SOME i => if i < 0 then (*preprocessor eliminates negative numerals*) |
|
605 |
trns |
|
| 20439 | 606 |
|> appgen_default thy thmtab (no_strict strct) app |
| 19884 | 607 |
else |
608 |
trns |
|
| 20439 | 609 |
|> exprgen_term thy thmtab (no_strict strct) (Const c) |
610 |
||>> exprgen_term thy thmtab (no_strict strct) bin |
|
| 20353 | 611 |
|-> (fn (e1, e2) => pair (CodegenThingol.INum (i, e1 `$ e2))) |
| 19884 | 612 |
| NONE => |
613 |
trns |
|
| 20439 | 614 |
|> appgen_default thy thmtab strct app; |
| 18217 | 615 |
|
| 20439 | 616 |
fun appgen_char char_to_index thy thmtab strct (app as ((_, ty), _)) trns = |
|
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
617 |
case (char_to_index o list_comb o apfst Const) app |
|
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
618 |
of SOME i => |
|
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
619 |
trns |
| 20439 | 620 |
|> exprgen_type thy thmtab strct ty |
621 |
||>> appgen_default thy thmtab strct app |
|
|
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
622 |
|-> (fn (_, e0) => pair (IChar (chr i, e0))) |
|
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
623 |
| NONE => |
|
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
624 |
trns |
| 20439 | 625 |
|> appgen_default thy thmtab strct app; |
|
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
626 |
|
| 20439 | 627 |
fun appgen_case dest_case_expr thy thmtab strct (app as (c_ty, ts)) trns = |
| 20105 | 628 |
let |
629 |
val SOME ([], ((st, sty), ds)) = dest_case_expr thy (list_comb (Const c_ty, ts)); |
|
630 |
fun clausegen (dt, bt) trns = |
|
631 |
trns |
|
| 20439 | 632 |
|> exprgen_term thy thmtab strct dt |
633 |
||>> exprgen_term thy thmtab strct bt; |
|
| 20105 | 634 |
in |
635 |
trns |
|
| 20439 | 636 |
|> exprgen_term thy thmtab strct st |
637 |
||>> exprgen_type thy thmtab strct sty |
|
| 20105 | 638 |
||>> fold_map clausegen ds |
| 20439 | 639 |
||>> appgen_default thy thmtab strct app |
| 20105 | 640 |
|-> (fn (((se, sty), ds), e0) => pair (ICase (((se, sty), ds), e0))) |
641 |
end; |
|
642 |
||
| 20439 | 643 |
fun appgen_let thy thmtab strct (app as (_, [st, ct])) trns = |
| 20105 | 644 |
trns |
| 20439 | 645 |
|> exprgen_term thy thmtab strct ct |
646 |
||>> exprgen_term thy thmtab strct st |
|
647 |
||>> appgen_default thy thmtab strct app |
|
| 20105 | 648 |
|-> (fn (((v, ty) `|-> be, se), e0) => |
649 |
pair (ICase (((se, ty), case be |
|
650 |
of ICase (((IVar w, _), ds), _) => if v = w then ds else [(IVar v, be)] |
|
651 |
| _ => [(IVar v, be)] |
|
652 |
), e0)) |
|
653 |
| (_, e0) => pair e0); |
|
654 |
||
| 20439 | 655 |
fun appgen_wfrec thy thmtab strct ((c, ty), [_, tf, tx]) trns = |
| 19038 | 656 |
let |
| 20386 | 657 |
val ty_def = (op ---> o apfst tl o strip_type o Logic.unvarifyT o Sign.the_const_type thy) c; |
| 19038 | 658 |
val ty' = (op ---> o apfst tl o strip_type) ty; |
| 20386 | 659 |
val idf = idf_of_const thy thmtab (c, ty); |
| 19038 | 660 |
in |
661 |
trns |
|
| 20386 | 662 |
|> ensure_def ((K o fail) "no extraction for wfrec") false ("generating wfrec") idf
|
| 20439 | 663 |
|> exprgen_type thy thmtab strct ty' |
664 |
||>> exprgen_type thy thmtab strct ty_def |
|
665 |
||>> exprgen_term thy thmtab strct tf |
|
666 |
||>> exprgen_term thy thmtab strct tx |
|
| 20434 | 667 |
|-> (fn (((_, ty), tf), tx) => pair (IConst (idf, ([], ty)) `$ tf `$ tx)) |
| 19038 | 668 |
end; |
669 |
||
| 20439 | 670 |
fun add_appconst (c, appgen) thy = |
671 |
let |
|
672 |
val i = (length o fst o strip_type o Sign.the_const_type thy) c |
|
673 |
in map_codegen_data |
|
674 |
(fn (modl, appgens, target_data) => |
|
675 |
(modl, |
|
676 |
appgens |> Symtab.update (c, (i, (appgen, stamp ()))), |
|
677 |
target_data)) thy |
|
678 |
end; |
|
679 |
||
| 18217 | 680 |
|
| 18516 | 681 |
|
| 20439 | 682 |
(** code generation interfaces **) |
| 18516 | 683 |
|
| 20439 | 684 |
fun generate cs targets init gen it thy = |
| 19597 | 685 |
thy |
686 |
|> CodegenTheorems.notify_dirty |
|
687 |
|> `(#modl o CodegenData.get) |
|
688 |
|> (fn (modl, thy) => |
|
| 20439 | 689 |
(start_transact init (gen thy (CodegenTheorems.mk_thmtab thy cs) |
690 |
(true, targets) it) modl, thy)) |
|
| 19597 | 691 |
|-> (fn (x, modl) => map_module (K modl) #> pair x); |
| 18516 | 692 |
|
| 20353 | 693 |
fun consts_of t = |
694 |
fold_aterms (fn Const c => cons c | _ => I) t []; |
|
695 |
||
| 20105 | 696 |
fun codegen_term t thy = |
| 20353 | 697 |
let |
698 |
val _ = Thm.cterm_of thy t; |
|
699 |
in |
|
700 |
thy |
|
| 20439 | 701 |
|> generate (consts_of t) (SOME []) NONE exprgen_term t |
| 20353 | 702 |
end; |
| 19136 | 703 |
|
704 |
val is_dtcon = has_nsp nsp_dtcon; |
|
705 |
||
706 |
fun consts_of_idfs thy = |
|
| 20386 | 707 |
map (the o const_of_idf thy); |
| 19150 | 708 |
|
| 20353 | 709 |
fun idfs_of_consts thy cs = |
| 20439 | 710 |
map (idf_of_const thy (CodegenTheorems.mk_thmtab thy cs)) cs; |
| 19150 | 711 |
|
| 19967 | 712 |
fun get_root_module thy = |
713 |
thy |
|
714 |
|> CodegenTheorems.notify_dirty |
|
715 |
|> `(#modl o CodegenData.get); |
|
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
716 |
|
| 20213 | 717 |
fun eval_term (ref_spec, t) thy = |
718 |
let |
|
| 20401 | 719 |
val _ = Term.fold_atyps (fn _ => |
720 |
error ("Term" ^ Sign.string_of_term thy t ^ "is polymorhpic"))
|
|
721 |
(Term.fastype_of t); |
|
| 20389 | 722 |
fun preprocess_term t = |
723 |
let |
|
724 |
val x = Free (Name.variant (add_term_names (t, [])) "x", fastype_of t); |
|
725 |
(* fake definition *) |
|
726 |
val eq = setmp quick_and_dirty true (SkipProof.make_thm thy) |
|
727 |
(Logic.mk_equals (x, t)); |
|
728 |
fun err () = error "preprocess_term: bad preprocessor" |
|
729 |
in case map prop_of (CodegenTheorems.preprocess thy [eq]) |
|
730 |
of [Const ("==", _) $ x' $ t'] => if x = x' then t' else err ()
|
|
731 |
| _ => err () |
|
732 |
end; |
|
| 20213 | 733 |
val target_data = |
734 |
((fn data => (the o Symtab.lookup data) "ml") o #target_data o CodegenData.get) thy; |
|
| 20428 | 735 |
val eval = CodegenSerializer.eval_term nsp_eval nsp_dtcon [[nsp_module], [nsp_class, nsp_tyco], [nsp_const, nsp_dtcon, nsp_class, nsp_mem, nsp_inst], [nsp_eval]] |
| 20213 | 736 |
((Option.map fst oo Symtab.lookup) (#syntax_tyco target_data), |
737 |
(Option.map fst oo Symtab.lookup) (#syntax_const target_data)) |
|
738 |
(Symtab.keys (#syntax_tyco target_data) @ Symtab.keys (#syntax_const target_data)) |
|
739 |
in |
|
740 |
thy |
|
| 20389 | 741 |
|> codegen_term (preprocess_term t) |
| 20213 | 742 |
||>> `(#modl o CodegenData.get) |
743 |
|-> (fn (t', modl) => `(fn _ => eval (ref_spec, t') modl)) |
|
744 |
end; |
|
745 |
||
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
746 |
fun get_ml_fun_datatype thy resolv = |
|
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
747 |
let |
| 19150 | 748 |
val target_data = |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
749 |
((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
|
750 |
in |
| 20183 | 751 |
CodegenSerializer.ml_fun_datatype nsp_dtcon |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
752 |
((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
|
753 |
(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
|
754 |
resolv |
|
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
755 |
end; |
| 18516 | 756 |
|
757 |
||
758 |
||
| 20439 | 759 |
(** target syntax **) |
| 18702 | 760 |
|
| 20439 | 761 |
local |
| 18702 | 762 |
|
| 20439 | 763 |
fun gen_add_syntax_class prep_class prep_const raw_class target (syntax, raw_ops) thy = |
| 20428 | 764 |
let |
765 |
val class = (idf_of_class thy o prep_class thy) raw_class; |
|
766 |
val ops = (map o apfst) (idf_of_classop thy o prep_const thy) raw_ops; |
|
767 |
val syntax_ops = AList.lookup (op =) ops; |
|
768 |
in |
|
769 |
thy |
|
| 20439 | 770 |
|> map_target_data target (fn (syntax_class, syntax_inst, syntax_tyco, syntax_const) => |
771 |
(syntax_class |> Symtab.update (class, |
|
772 |
((syntax, syntax_ops), stamp ())), |
|
773 |
syntax_inst, syntax_tyco, syntax_const)) |
|
| 20428 | 774 |
end; |
| 18865 | 775 |
|
| 20428 | 776 |
fun gen_add_syntax_inst prep_class prep_tyco (raw_class, raw_tyco) target thy = |
777 |
let |
|
778 |
val inst = idf_of_inst thy (prep_class thy raw_class, prep_tyco thy raw_tyco); |
|
779 |
in |
|
780 |
thy |
|
| 20439 | 781 |
|> map_target_data target (fn (syntax_class, syntax_inst, syntax_tyco, syntax_const) => |
782 |
(syntax_class, syntax_inst |> Symtab.update (inst, ()), |
|
783 |
syntax_tyco, syntax_const)) |
|
| 20428 | 784 |
end; |
785 |
||
| 20439 | 786 |
fun gen_add_syntax_tyco prep_tyco raw_tyco target syntax thy = |
| 18217 | 787 |
let |
| 20439 | 788 |
val tyco = (idf_of_tyco thy o prep_tyco thy) raw_tyco; |
| 18217 | 789 |
in |
| 20439 | 790 |
thy |
791 |
|> map_target_data target (fn (syntax_class, syntax_inst, syntax_tyco, syntax_const) => |
|
792 |
(syntax_class, syntax_inst, syntax_tyco |
|
793 |
|> Symtab.update (tyco, (syntax, stamp ())), syntax_const)) |
|
794 |
end; |
|
795 |
||
796 |
fun gen_add_syntax_const prep_const raw_c target syntax thy = |
|
797 |
let |
|
798 |
val c_ty = prep_const thy raw_c; |
|
799 |
val c = idf_of_const thy (CodegenTheorems.mk_thmtab thy [c_ty]) c_ty; |
|
800 |
in |
|
801 |
thy |
|
802 |
|> map_target_data target (fn (syntax_class, syntax_inst, syntax_tyco, syntax_const) => |
|
803 |
(syntax_class, syntax_inst, syntax_tyco, syntax_const |
|
804 |
|> Symtab.update (c, (syntax, stamp ())))) |
|
| 18217 | 805 |
end; |
806 |
||
| 20439 | 807 |
fun idfs_of_const_names thy cs = |
808 |
let |
|
809 |
val cs' = AList.make (fn c => Sign.the_const_type thy c) cs |
|
810 |
val thmtab = CodegenTheorems.mk_thmtab thy cs' |
|
811 |
in AList.make (idf_of_const thy thmtab) cs' end; |
|
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
812 |
|
| 20439 | 813 |
fun read_quote reader consts_of target get_init gen raw_it thy = |
| 18217 | 814 |
let |
| 20439 | 815 |
val it = reader thy raw_it; |
816 |
val cs = consts_of it; |
|
| 18217 | 817 |
in |
| 20439 | 818 |
thy |
819 |
|> generate cs (SOME [target]) ((SOME o get_init) thy) |
|
820 |
(fn thy => fn thmtab => fn strct => gen thy thmtab strct) [it] |
|
821 |
|-> (fn [it'] => pair it') |
|
| 18217 | 822 |
end; |
823 |
||
| 20439 | 824 |
fun parse_quote num_of reader consts_of target get_init gen adder = |
825 |
CodegenSerializer.parse_syntax num_of |
|
826 |
(read_quote reader consts_of target get_init gen) |
|
827 |
#-> (fn modifier => pair (modifier #-> adder target)); |
|
828 |
||
829 |
in |
|
830 |
||
831 |
val add_syntax_class = gen_add_syntax_class Sign.intern_class CodegenConsts.read_const_typ; |
|
832 |
val add_syntax_inst = gen_add_syntax_inst Sign.intern_class Sign.intern_type; |
|
833 |
||
834 |
fun parse_syntax_tyco raw_tyco target = |
|
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
835 |
let |
| 20439 | 836 |
fun intern thy = Sign.intern_type thy raw_tyco; |
837 |
fun num_of thy = Sign.arity_number thy (intern thy); |
|
838 |
fun idf_of thy = idf_of_tyco thy (intern thy); |
|
839 |
fun read_typ thy = |
|
840 |
Sign.read_typ (thy, K NONE); |
|
841 |
in |
|
842 |
parse_quote num_of read_typ (K []) target idf_of (fold_map ooo exprgen_type) |
|
843 |
(gen_add_syntax_tyco Sign.intern_type raw_tyco) |
|
844 |
end; |
|
845 |
||
846 |
fun parse_syntax_const raw_const target = |
|
847 |
let |
|
848 |
fun intern thy = CodegenConsts.read_const_typ thy raw_const; |
|
849 |
fun num_of thy = (length o fst o strip_type o Sign.the_const_type thy o fst o intern) thy; |
|
850 |
fun idf_of thy = |
|
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
851 |
let |
| 20439 | 852 |
val c_ty = intern thy; |
853 |
val c = idf_of_const thy (CodegenTheorems.mk_thmtab thy [c_ty]) c_ty; |
|
854 |
in c end; |
|
855 |
in |
|
856 |
parse_quote num_of Sign.read_term consts_of target idf_of (fold_map ooo exprgen_term) |
|
857 |
(gen_add_syntax_const CodegenConsts.read_const_typ raw_const) |
|
858 |
end; |
|
859 |
||
860 |
fun add_pretty_list target nill cons mk_list mk_char_string target_cons thy = |
|
861 |
let |
|
862 |
val [(_, nil''), (cons', cons'')] = idfs_of_const_names thy [nill, cons]; |
|
| 20401 | 863 |
val pr = CodegenSerializer.pretty_list nil'' cons'' mk_list mk_char_string target_cons; |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
864 |
in |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
865 |
thy |
| 20439 | 866 |
|> gen_add_syntax_const (K I) cons' target pr |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
867 |
end; |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
868 |
|
| 20439 | 869 |
fun add_pretty_ml_string target nill cons str mk_char mk_string target_implode thy = |
| 20401 | 870 |
let |
| 20439 | 871 |
val [(_, nil''), (_, cons''), (str', _)] = idfs_of_const_names thy [nill, cons, str]; |
872 |
val pr = CodegenSerializer.pretty_ml_string nil'' cons'' mk_char mk_string target_implode; |
|
| 20401 | 873 |
in |
874 |
thy |
|
| 20439 | 875 |
|> gen_add_syntax_const (K I) str' target pr |
| 20401 | 876 |
end; |
877 |
||
| 20439 | 878 |
end; (*local*) |
879 |
||
| 18217 | 880 |
|
| 18516 | 881 |
|
| 20439 | 882 |
(** toplevel interface and setup **) |
| 18756 | 883 |
|
884 |
local |
|
| 19150 | 885 |
|
| 20439 | 886 |
fun generate_code targets (SOME raw_cs) thy = |
|
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
887 |
let |
| 20439 | 888 |
val cs = map (CodegenConsts.read_const_typ thy) raw_cs; |
| 19884 | 889 |
val _ = case targets of SOME targets => (map get_serializer targets; ()) | _ => (); |
| 18756 | 890 |
in |
891 |
thy |
|
| 20439 | 892 |
|> generate cs targets NONE (fold_map ooo ensure_def_const) cs |
| 18756 | 893 |
|-> (fn cs => pair (SOME cs)) |
894 |
end |
|
|
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
895 |
| generate_code _ NONE thy = |
| 18756 | 896 |
(NONE, thy); |
897 |
||
| 20439 | 898 |
fun serialize_code target seri raw_cs thy = |
899 |
thy |
|
900 |
|> generate_code (SOME [target]) raw_cs |
|
901 |
|-> (fn cs => tap (fn thy => serialize thy target seri cs)); |
|
902 |
||
903 |
fun code raw_cs seris thy = |
|
| 18217 | 904 |
let |
| 20439 | 905 |
val cs = map (CodegenConsts.read_const_typ thy) raw_cs; |
906 |
val targets = map fst seris; |
|
907 |
val seris' = map_filter (fn (target, SOME seri) => SOME (target, seri) | _ => NONE) seris; |
|
908 |
fun generate' thy = case cs |
|
909 |
of [] => ([], thy) |
|
910 |
| _ => |
|
| 18702 | 911 |
thy |
| 20439 | 912 |
|> generate cs (SOME targets) NONE (fold_map ooo ensure_def_const) cs; |
913 |
fun serialize' thy [] (target, seri) = |
|
914 |
serialize thy target seri NONE : unit |
|
915 |
| serialize' thy cs (target, seri) = |
|
916 |
serialize thy target seri (SOME cs) : unit; |
|
| 18217 | 917 |
in |
918 |
thy |
|
| 20439 | 919 |
|> generate' |
920 |
|-> (fn cs => tap (fn thy => map (serialize' thy cs) seris')) |
|
| 18217 | 921 |
end; |
922 |
||
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
923 |
fun purge_consts raw_ts thy = |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
924 |
let |
| 20386 | 925 |
val cs = map (CodegenConsts.read_const_typ thy) raw_ts; |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
926 |
in fold CodegenTheorems.purge_defs cs thy end; |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
927 |
|
| 18217 | 928 |
structure P = OuterParse |
929 |
and K = OuterKeyword |
|
930 |
||
931 |
in |
|
932 |
||
| 20439 | 933 |
val (codeK, generateK, serializeK, |
| 20428 | 934 |
syntax_classK, syntax_instK, syntax_tycoK, syntax_constK, |
| 20386 | 935 |
purgeK) = |
| 20439 | 936 |
("codeK", "code_generate", "code_serialize",
|
| 20428 | 937 |
"code_class", "code_instance", "code_typapp", "code_constapp", |
| 20386 | 938 |
"code_purge"); |
| 18335 | 939 |
|
| 20439 | 940 |
val codeP = |
941 |
OuterSyntax.command codeK "generate and serialize executable code for constants" K.thy_decl ( |
|
942 |
Scan.repeat P.term |
|
943 |
-- Scan.repeat (P.$$$ "(" |--
|
|
944 |
P.name :-- (fn target => (get_serializer target >> SOME) || pair NONE) |
|
945 |
--| P.$$$ ")") |
|
946 |
>> (fn (raw_cs, seris) => Toplevel.theory (code raw_cs seris)) |
|
947 |
); |
|
948 |
||
| 18217 | 949 |
val generateP = |
| 18282 | 950 |
OuterSyntax.command generateK "generate executable code for constants" K.thy_decl ( |
| 19884 | 951 |
(Scan.option (P.$$$ "(" |-- P.list1 P.name --| P.$$$ ")")
|
952 |
>> (fn SOME ["-"] => SOME [] | ts => ts)) |
|
|
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
953 |
-- Scan.repeat1 P.term |
| 19884 | 954 |
>> (fn (targets, raw_consts) => |
955 |
Toplevel.theory (generate_code targets (SOME raw_consts) #> snd)) |
|
| 18217 | 956 |
); |
957 |
||
958 |
val serializeP = |
|
| 18282 | 959 |
OuterSyntax.command serializeK "serialize executable code for constants" K.thy_decl ( |
| 18217 | 960 |
P.name |
| 19136 | 961 |
-- Scan.option (Scan.repeat1 P.term) |
| 18756 | 962 |
#-> (fn (target, raw_consts) => |
| 18850 | 963 |
P.$$$ "("
|
964 |
|-- get_serializer target |
|
965 |
--| P.$$$ ")" |
|
| 18756 | 966 |
>> (fn seri => |
967 |
Toplevel.theory (serialize_code target seri raw_consts) |
|
968 |
)) |
|
| 18217 | 969 |
); |
970 |
||
| 18865 | 971 |
val syntax_classP = |
| 19884 | 972 |
OuterSyntax.command syntax_classK "define code syntax for class" K.thy_decl ( |
| 18865 | 973 |
Scan.repeat1 ( |
974 |
P.xname |
|
975 |
-- Scan.repeat1 ( |
|
| 20428 | 976 |
P.name -- (P.string -- Scan.optional |
977 |
(P.$$$ "(" |-- Scan.repeat1 (P.term -- P.string) --| P.$$$ ")") [])
|
|
| 18865 | 978 |
) |
979 |
) |
|
980 |
>> (Toplevel.theory oo fold) (fn (raw_class, syns) => |
|
981 |
fold (fn (target, p) => add_syntax_class raw_class target p) syns) |
|
982 |
); |
|
983 |
||
| 20428 | 984 |
val syntax_instP = |
985 |
OuterSyntax.command syntax_instK "define code syntax for instance" K.thy_decl ( |
|
986 |
Scan.repeat1 ( |
|
987 |
P.$$$ "(" |-- P.xname --| P.$$$ "::" -- P.xname --| P.$$$ ")"
|
|
988 |
-- Scan.repeat1 P.name |
|
989 |
) |
|
990 |
>> (Toplevel.theory oo fold) (fn (raw_inst, targets) => |
|
991 |
fold (fn target => add_syntax_inst raw_inst target) targets) |
|
992 |
); |
|
993 |
||
| 18217 | 994 |
val syntax_tycoP = |
995 |
OuterSyntax.command syntax_tycoK "define code syntax for type constructor" K.thy_decl ( |
|
| 18702 | 996 |
Scan.repeat1 ( |
997 |
P.xname |
|
| 18963 | 998 |
#-> (fn raw_tyco => Scan.repeat1 ( |
| 20439 | 999 |
P.name #-> parse_syntax_tyco raw_tyco |
| 18963 | 1000 |
)) |
| 18702 | 1001 |
) |
| 18963 | 1002 |
>> (Toplevel.theory oo fold o fold) |
| 20439 | 1003 |
(fn modifier => modifier) |
| 18217 | 1004 |
); |
1005 |
||
1006 |
val syntax_constP = |
|
1007 |
OuterSyntax.command syntax_constK "define code syntax for constant" K.thy_decl ( |
|
| 18702 | 1008 |
Scan.repeat1 ( |
| 19136 | 1009 |
P.term |
| 18963 | 1010 |
#-> (fn raw_const => Scan.repeat1 ( |
| 20439 | 1011 |
P.name #-> parse_syntax_const raw_const |
| 18963 | 1012 |
)) |
| 18702 | 1013 |
) |
| 19008 | 1014 |
>> (Toplevel.theory oo fold o fold) |
| 20439 | 1015 |
(fn modifier => modifier) |
| 18217 | 1016 |
); |
1017 |
||
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
1018 |
val purgeP = |
| 20175 | 1019 |
OuterSyntax.command purgeK "purge all incrementally generated code" K.thy_decl |
1020 |
(Scan.succeed (Toplevel.theory purge_code)); |
|
| 18516 | 1021 |
|
| 20439 | 1022 |
val _ = OuterSyntax.add_parsers [(*codeP, *)generateP, serializeP, |
1023 |
syntax_classP, syntax_instP, syntax_tycoP, syntax_constP, purgeP]; |
|
| 18217 | 1024 |
|
1025 |
end; (* local *) |
|
1026 |
||
| 20439 | 1027 |
(*code basis change notifications*) |
1028 |
val _ = Context.add_setup (CodegenTheorems.add_notify purge_defs); |
|
1029 |
||
| 18217 | 1030 |
end; (* struct *) |