author | wenzelm |
Sun, 01 Oct 2006 18:30:04 +0200 | |
changeset 20818 | cb7ec413f95d |
parent 20699 | 0cc77abb185a |
child 20835 | 27d049062b56 |
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 |
20456 | 11 |
include BASIC_CODEGEN_THINGOL; |
20600 | 12 |
val codegen_term: theory -> term -> thm * iterm; |
13 |
val eval_term: theory -> (string (*reference name!*) * 'a option ref) * term -> 'a; |
|
20699 | 14 |
val const_of_idf: theory -> string -> string * typ; |
20600 | 15 |
val get_root_module: theory -> CodegenThingol.module; |
18217 | 16 |
|
19884 | 17 |
type appgen; |
20439 | 18 |
val add_appconst: string * appgen -> theory -> theory; |
18702 | 19 |
val appgen_default: appgen; |
20600 | 20 |
val appgen_numeral: (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
|
21 |
val appgen_char: (term -> int option) -> appgen; |
20105 | 22 |
val appgen_case: (theory -> term |
23 |
-> ((string * typ) list * ((term * typ) * (term * term) list)) option) |
|
24 |
-> appgen; |
|
25 |
val appgen_let: appgen; |
|
18217 | 26 |
|
19008 | 27 |
val print_code: theory -> unit; |
20600 | 28 |
val purge_code: theory -> CodegenThingol.module; |
29 |
structure CodegenPackageData: THEORY_DATA; |
|
30 |
structure Code: CODE_DATA; |
|
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
31 |
end; |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
32 |
|
18217 | 33 |
structure CodegenPackage : CODEGEN_PACKAGE = |
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
34 |
struct |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
35 |
|
18850 | 36 |
open CodegenThingol; |
18217 | 37 |
|
20439 | 38 |
(** preliminaries **) |
39 |
||
20699 | 40 |
val nsp_module = CodegenNames.nsp_module; |
41 |
val nsp_class = CodegenNames.nsp_class; |
|
42 |
val nsp_tyco = CodegenNames.nsp_tyco; |
|
43 |
val nsp_inst = CodegenNames.nsp_inst; |
|
44 |
val nsp_fun = CodegenNames.nsp_fun; |
|
45 |
val nsp_classop = CodegenNames.nsp_classop; |
|
46 |
val nsp_dtco = CodegenNames.nsp_dtco; |
|
47 |
val nsp_eval = CodegenNames.nsp_eval; |
|
18217 | 48 |
|
49 |
||
20699 | 50 |
|
51 |
(** code extraction **) |
|
52 |
||
53 |
(* theory data *) |
|
18217 | 54 |
|
20600 | 55 |
type appgen = theory -> ((sort -> sort) * Sorts.algebra) * Consts.T |
56 |
-> CodegenFuncgr.T |
|
20456 | 57 |
-> bool * string list option -> (string * typ) * term list -> transact -> iterm * transact; |
20439 | 58 |
|
59 |
type appgens = (int * (appgen * stamp)) Symtab.table; |
|
18217 | 60 |
|
20386 | 61 |
fun merge_appgens (x : appgens * appgens) = |
20105 | 62 |
Symtab.merge (fn ((bounds1, (_, stamp1)), (bounds2, (_, stamp2))) => |
20439 | 63 |
bounds1 = bounds2 andalso stamp1 = stamp2) x; |
18217 | 64 |
|
20600 | 65 |
structure Code = CodeDataFun |
20456 | 66 |
(struct |
67 |
val name = "Pure/code"; |
|
68 |
type T = module; |
|
69 |
val empty = empty_module; |
|
70 |
fun merge _ = merge_module; |
|
20600 | 71 |
fun purge _ _ = CodegenThingol.empty_module; |
20456 | 72 |
end); |
73 |
||
20600 | 74 |
structure CodegenPackageData = TheoryDataFun |
18217 | 75 |
(struct |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19884
diff
changeset
|
76 |
val name = "Pure/codegen_package"; |
20699 | 77 |
type T = appgens; |
78 |
val empty = Symtab.empty; |
|
18217 | 79 |
val copy = I; |
80 |
val extend = I; |
|
20699 | 81 |
fun merge _ = merge_appgens; |
20456 | 82 |
fun print _ _ = (); |
18217 | 83 |
end); |
84 |
||
20600 | 85 |
val _ = Context.add_setup (Code.init #> CodegenPackageData.init); |
18708 | 86 |
|
20600 | 87 |
fun print_code thy = |
88 |
let |
|
89 |
val code = Code.get thy; |
|
90 |
in (Pretty.writeln o Pretty.chunks) [pretty_module code, pretty_deps code] end; |
|
20439 | 91 |
|
20600 | 92 |
fun purge_code thy = Code.change thy (K CodegenThingol.empty_module); |
18865 | 93 |
|
94 |
||
20386 | 95 |
(* extraction kernel *) |
18865 | 96 |
|
20699 | 97 |
fun check_strict has_seri x (false, _) = |
19884 | 98 |
false |
20699 | 99 |
| check_strict has_seri x (_, SOME targets) = |
100 |
not (has_seri targets x) |
|
101 |
| check_strict has_seri x (true, _) = |
|
19884 | 102 |
true; |
103 |
||
20439 | 104 |
fun no_strict (_, targets) = (false, targets); |
20386 | 105 |
|
20600 | 106 |
fun ensure_def_class thy algbr funcgr strct cls trns = |
20386 | 107 |
let |
20600 | 108 |
fun defgen_class thy (algbr as ((proj_sort, _), _)) funcgr strct cls trns = |
20699 | 109 |
case CodegenNames.class_rev thy cls |
18865 | 110 |
of SOME cls => |
111 |
let |
|
19283 | 112 |
val (v, cs) = (ClassPackage.the_consts_sign thy) cls; |
20456 | 113 |
val superclasses = (proj_sort o Sign.super_classes thy) cls |
20699 | 114 |
val idfs = map (CodegenNames.const thy o CodegenConsts.norm_of_typ thy) cs; |
18865 | 115 |
in |
116 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
117 |
|> debug_msg (fn _ => "trying defgen class declaration for " ^ quote cls) |
20600 | 118 |
|> fold_map (ensure_def_class thy algbr funcgr strct) superclasses |
119 |
||>> (fold_map (exprgen_type thy algbr funcgr strct) o map snd) cs |
|
20456 | 120 |
|-> (fn (supcls, memtypes) => succeed |
121 |
(Class (supcls, (unprefix "'" v, idfs ~~ memtypes)))) |
|
18865 | 122 |
end |
123 |
| _ => |
|
124 |
trns |
|
20389 | 125 |
|> fail ("No class definition found for " ^ quote cls); |
20699 | 126 |
val cls' = CodegenNames.class thy cls; |
18454 | 127 |
in |
128 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
129 |
|> debug_msg (fn _ => "generating class " ^ quote cls) |
20600 | 130 |
|> ensure_def (defgen_class thy algbr funcgr strct) true ("generating class " ^ quote cls) cls' |
18865 | 131 |
|> pair cls' |
132 |
end |
|
20600 | 133 |
and ensure_def_tyco thy algbr funcgr strct tyco trns = |
18865 | 134 |
let |
20699 | 135 |
val tyco' = CodegenNames.tyco thy tyco; |
136 |
val strict = check_strict (CodegenSerializer.tyco_has_serialization thy) tyco' strct; |
|
20600 | 137 |
fun defgen_datatype thy algbr funcgr strct dtco trns = |
20699 | 138 |
case CodegenNames.tyco_rev thy dtco |
18963 | 139 |
of SOME dtco => |
20600 | 140 |
(case CodegenData.get_datatype thy dtco |
20456 | 141 |
of SOME (vs, cos) => |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
142 |
trns |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
143 |
|> debug_msg (fn _ => "trying defgen datatype for " ^ quote dtco) |
20600 | 144 |
|> fold_map (exprgen_tyvar_sort thy algbr funcgr strct) vs |
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
145 |
||>> fold_map (fn (c, tys) => |
20600 | 146 |
fold_map (exprgen_type thy algbr funcgr strct) tys |
147 |
#-> (fn tys' => |
|
20699 | 148 |
pair ((CodegenNames.const thy o CodegenConsts.norm_of_typ thy) |
20600 | 149 |
(c, tys ---> Type (dtco, map TFree vs)), tys'))) cos |
20456 | 150 |
|-> (fn (vs, cos) => succeed (Datatype (vs, cos))) |
18963 | 151 |
| NONE => |
152 |
trns |
|
20389 | 153 |
|> fail ("No datatype found for " ^ quote dtco)) |
18963 | 154 |
| NONE => |
155 |
trns |
|
20389 | 156 |
|> fail ("Not a type constructor: " ^ quote dtco) |
18865 | 157 |
in |
158 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
159 |
|> debug_msg (fn _ => "generating type constructor " ^ quote tyco) |
20600 | 160 |
|> ensure_def (defgen_datatype thy algbr funcgr strct) strict |
20389 | 161 |
("generating type constructor " ^ quote tyco) tyco' |
18865 | 162 |
|> pair tyco' |
163 |
end |
|
20600 | 164 |
and exprgen_tyvar_sort thy (algbr as ((proj_sort, _), _)) funcgr strct (v, sort) trns = |
18516 | 165 |
trns |
20600 | 166 |
|> fold_map (ensure_def_class thy algbr funcgr strct) (proj_sort sort) |
18865 | 167 |
|-> (fn sort => pair (unprefix "'" v, sort)) |
20600 | 168 |
and exprgen_type thy algbr funcgr strct (TVar _) trns = |
20389 | 169 |
error "TVar encountered in typ during code generation" |
20600 | 170 |
| exprgen_type thy algbr funcgr strct (TFree vs) trns = |
18516 | 171 |
trns |
20600 | 172 |
|> exprgen_tyvar_sort thy algbr funcgr strct vs |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
173 |
|-> (fn (v, sort) => pair (ITyVar v)) |
20600 | 174 |
| exprgen_type thy algbr funcgr strct (Type ("fun", [t1, t2])) trns = |
18516 | 175 |
trns |
20600 | 176 |
|> exprgen_type thy algbr funcgr strct t1 |
177 |
||>> exprgen_type thy algbr funcgr strct t2 |
|
18516 | 178 |
|-> (fn (t1', t2') => pair (t1' `-> t2')) |
20600 | 179 |
| exprgen_type thy algbr funcgr strct (Type (tyco, tys)) trns = |
18516 | 180 |
trns |
20600 | 181 |
|> ensure_def_tyco thy algbr funcgr strct tyco |
182 |
||>> fold_map (exprgen_type thy algbr funcgr strct) tys |
|
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
183 |
|-> (fn (tyco, tys) => pair (tyco `%% tys)); |
18516 | 184 |
|
20600 | 185 |
exception CONSTRAIN of ((string * typ) * typ) * term option; |
186 |
||
187 |
fun exprgen_typinst thy (algbr as ((proj_sort, algebra), consts)) funcgr strct (ty_ctxt, sort_decl) trns = |
|
20456 | 188 |
let |
189 |
val pp = Sign.pp thy; |
|
190 |
datatype inst = |
|
191 |
Inst of (class * string) * inst list list |
|
192 |
| Contxt of (string * sort) * (class list * int); |
|
193 |
fun classrel (l as Contxt (v_sort, (classes, n)), _) class = |
|
194 |
Contxt (v_sort, (class :: classes, n)) |
|
195 |
| classrel (Inst ((_, tyco), lss), _) class = |
|
196 |
Inst ((class, tyco), lss); |
|
197 |
fun constructor tyco iss class = |
|
198 |
Inst ((class, tyco), (map o map) fst iss); |
|
199 |
fun variable (TFree (v, sort)) = |
|
200 |
let |
|
201 |
val sort' = proj_sort sort; |
|
202 |
in map_index (fn (n, class) => (Contxt ((v, sort'), ([], n)), class)) sort' end; |
|
203 |
val insts = Sorts.of_sort_derivation pp algebra |
|
204 |
{classrel = classrel, constructor = constructor, variable = variable} |
|
205 |
(ty_ctxt, proj_sort sort_decl); |
|
206 |
fun mk_dict (Inst (inst, instss)) trns = |
|
207 |
trns |
|
20600 | 208 |
|> ensure_def_inst thy algbr funcgr strct inst |
20456 | 209 |
||>> (fold_map o fold_map) mk_dict instss |
210 |
|-> (fn (inst, instss) => pair (Instance (inst, instss))) |
|
211 |
| mk_dict (Contxt ((v, sort), (classes, k))) trns = |
|
18517 | 212 |
trns |
20600 | 213 |
|> fold_map (ensure_def_class thy algbr funcgr strct) classes |
20456 | 214 |
|-> (fn classes => pair (Context (classes, (unprefix "'" v, |
215 |
if length sort = 1 then ~1 else k)))) |
|
216 |
in |
|
217 |
trns |
|
218 |
|> fold_map mk_dict insts |
|
219 |
end |
|
20600 | 220 |
and exprgen_typinst_const thy (algbr as (_, consts)) funcgr strct (c, ty_ctxt) trns = |
221 |
let |
|
222 |
val c' = CodegenConsts.norm_of_typ thy (c, ty_ctxt) |
|
20699 | 223 |
val idf = CodegenNames.const thy c'; |
20466 | 224 |
val ty_decl = Consts.declaration consts idf; |
225 |
val insts = (op ~~ o apsnd (map (snd o dest_TVar)) oo pairself) |
|
226 |
(curry (Consts.typargs consts) idf) (ty_ctxt, ty_decl); |
|
20600 | 227 |
val _ = if exists not (map (Sign.of_sort thy) insts) |
228 |
then raise CONSTRAIN (((c, ty_decl), ty_ctxt), NONE) else (); |
|
20456 | 229 |
in |
230 |
trns |
|
20600 | 231 |
|> fold_map (exprgen_typinst thy algbr funcgr strct) insts |
20456 | 232 |
end |
20600 | 233 |
and ensure_def_inst thy algbr funcgr strct (cls, tyco) trns = |
20456 | 234 |
let |
20600 | 235 |
fun defgen_inst thy (algbr as ((proj_sort, _), _)) funcgr strct inst trns = |
20699 | 236 |
case CodegenNames.instance_rev thy inst |
19956 | 237 |
of SOME (class, tyco) => |
18865 | 238 |
let |
20456 | 239 |
val (vs, memdefs) = ClassPackage.the_inst_sign thy (class, tyco); |
20389 | 240 |
val (_, members) = ClassPackage.the_consts_sign thy class; |
20456 | 241 |
val arity_typ = Type (tyco, (map TFree vs)); |
242 |
val superclasses = (proj_sort o Sign.super_classes thy) class |
|
18865 | 243 |
fun gen_suparity supclass trns = |
244 |
trns |
|
20600 | 245 |
|> ensure_def_class thy algbr funcgr strct supclass |
246 |
||>> exprgen_typinst thy algbr funcgr strct (arity_typ, [supclass]) |
|
20466 | 247 |
|-> (fn (supclass, [Instance (supints, lss)]) => pair (supclass, (supints, lss))); |
20386 | 248 |
fun gen_membr ((m0, ty0), (m, ty)) trns = |
18865 | 249 |
trns |
20600 | 250 |
|> ensure_def_const thy algbr funcgr strct (CodegenConsts.norm_of_typ thy (m0, ty0)) |
251 |
||>> exprgen_term thy algbr funcgr strct (Const (m, ty)); |
|
18865 | 252 |
in |
253 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
254 |
|> debug_msg (fn _ => "trying defgen class instance for (" ^ quote cls |
18865 | 255 |
^ ", " ^ quote tyco ^ ")") |
20600 | 256 |
|> ensure_def_class thy algbr funcgr strct class |
257 |
||>> ensure_def_tyco thy algbr funcgr strct tyco |
|
258 |
||>> fold_map (exprgen_tyvar_sort thy algbr funcgr strct) vs |
|
20456 | 259 |
||>> fold_map gen_suparity superclasses |
20389 | 260 |
||>> fold_map gen_membr (members ~~ memdefs) |
18885 | 261 |
|-> (fn ((((class, tyco), arity), suparities), memdefs) => |
20389 | 262 |
succeed (Classinst ((class, (tyco, arity)), (suparities, memdefs)))) |
18865 | 263 |
end |
264 |
| _ => |
|
20389 | 265 |
trns |> fail ("No class instance found for " ^ quote inst); |
20699 | 266 |
val inst = CodegenNames.instance thy (cls, tyco); |
18865 | 267 |
in |
268 |
trns |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
269 |
|> debug_msg (fn _ => "generating instance " ^ quote cls ^ " / " ^ quote tyco) |
20600 | 270 |
|> ensure_def (defgen_inst thy algbr funcgr strct) true |
18865 | 271 |
("generating instance " ^ quote cls ^ " / " ^ quote tyco) inst |
272 |
|> pair inst |
|
273 |
end |
|
20600 | 274 |
and ensure_def_const thy algbr funcgr strct (c, tys) trns = |
18865 | 275 |
let |
20600 | 276 |
fun defgen_datatypecons thy algbr funcgr strct co trns = |
20699 | 277 |
case CodegenData.get_datatype_of_constr thy ((the o CodegenNames.const_rev thy) co) |
20386 | 278 |
of SOME tyco => |
18865 | 279 |
trns |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19283
diff
changeset
|
280 |
|> debug_msg (fn _ => "trying defgen datatype constructor for " ^ quote co) |
20600 | 281 |
|> ensure_def_tyco thy algbr funcgr strct tyco |
20386 | 282 |
|-> (fn _ => succeed Bot) |
18865 | 283 |
| _ => |
284 |
trns |
|
20389 | 285 |
|> fail ("Not a datatype constructor: " |
20600 | 286 |
^ (quote o CodegenConsts.string_of_const thy) (c, tys)); |
287 |
fun defgen_clsmem thy algbr funcgr strct m trns = |
|
20699 | 288 |
case AxClass.class_of_param thy ((fst o the o CodegenNames.const_rev thy) m) |
20386 | 289 |
of SOME class => |
290 |
trns |
|
20699 | 291 |
|> debug_msg (fn _ => "trying defgen class operation for " ^ quote m) |
20600 | 292 |
|> ensure_def_class thy algbr funcgr strct class |
20386 | 293 |
|-> (fn _ => succeed Bot) |
294 |
| _ => |
|
20699 | 295 |
trns |> fail ("No class found for " ^ (quote o CodegenConsts.string_of_const thy) (c, tys)) |
20600 | 296 |
fun defgen_funs thy (algbr as (_, consts)) funcgr strct c' trns = |
20699 | 297 |
case CodegenFuncgr.get_funcs funcgr ((the o CodegenNames.const_rev thy) c') |
20456 | 298 |
of eq_thms as eq_thm :: _ => |
299 |
let |
|
300 |
val msg = cat_lines ("generating code for theorems " :: map string_of_thm eq_thms); |
|
20600 | 301 |
val ty = (Logic.unvarifyT o CodegenData.typ_func thy) eq_thm; |
20466 | 302 |
val vs = (map dest_TFree o Consts.typargs consts) (c', ty); |
20456 | 303 |
fun dest_eqthm eq_thm = |
304 |
let |
|
305 |
val ((t, args), rhs) = |
|
20466 | 306 |
(apfst strip_comb o Logic.dest_equals o Logic.unvarify o prop_of) eq_thm; |
20456 | 307 |
in case t |
308 |
of Const (c', _) => if c' = c then (args, rhs) |
|
309 |
else error ("Illegal function equation for " ^ quote c |
|
310 |
^ ", actually defining " ^ quote c') |
|
311 |
| _ => error ("Illegal function equation for " ^ quote c) |
|
312 |
end; |
|
313 |
fun exprgen_eq (args, rhs) trns = |
|
314 |
trns |
|
20600 | 315 |
|> fold_map (exprgen_term thy algbr funcgr strct) args |
316 |
||>> exprgen_term thy algbr funcgr strct rhs; |
|
20456 | 317 |
fun checkvars (args, rhs) = |
318 |
if CodegenThingol.vars_distinct args then (args, rhs) |
|
319 |
else error ("Repeated variables on left hand side of function") |
|
320 |
in |
|
321 |
trns |
|
322 |
|> message msg (fn trns => trns |
|
323 |
|> fold_map (exprgen_eq o dest_eqthm) eq_thms |
|
324 |
|-> (fn eqs => pair (map checkvars eqs)) |
|
20600 | 325 |
||>> fold_map (exprgen_tyvar_sort thy algbr funcgr strct) vs |
326 |
||>> exprgen_type thy algbr funcgr strct ty |
|
20456 | 327 |
|-> (fn ((eqs, vs), ty) => succeed (Fun (eqs, (vs, ty))))) |
328 |
end |
|
329 |
| [] => |
|
330 |
trns |
|
331 |
|> fail ("No defining equations found for " |
|
20600 | 332 |
^ (quote o CodegenConsts.string_of_const thy) (c, tys)); |
333 |
fun get_defgen funcgr strct idf strict = |
|
20699 | 334 |
if CodegenNames.has_nsp nsp_fun idf |
20600 | 335 |
then defgen_funs thy algbr funcgr strct strict |
20699 | 336 |
else if CodegenNames.has_nsp nsp_classop idf |
20600 | 337 |
then defgen_clsmem thy algbr funcgr strct strict |
20699 | 338 |
else if CodegenNames.has_nsp nsp_dtco idf |
20600 | 339 |
then defgen_datatypecons thy algbr funcgr strct strict |
20389 | 340 |
else error ("Illegal shallow name space for constant: " ^ quote idf); |
20699 | 341 |
val idf = CodegenNames.const thy (c, tys); |
342 |
val strict = check_strict (CodegenSerializer.const_has_serialization thy) idf strct; |
|
18865 | 343 |
in |
344 |
trns |
|
20389 | 345 |
|> debug_msg (fn _ => "generating constant " |
20600 | 346 |
^ (quote o CodegenConsts.string_of_const thy) (c, tys)) |
347 |
|> ensure_def (get_defgen funcgr strct idf) strict ("generating constant " |
|
348 |
^ CodegenConsts.string_of_const thy (c, tys)) idf |
|
18963 | 349 |
|> pair idf |
18865 | 350 |
end |
20600 | 351 |
and exprgen_term thy algbr funcgr strct (Const (f, ty)) trns = |
18517 | 352 |
trns |
20600 | 353 |
|> appgen thy algbr funcgr strct ((f, ty), []) |
18516 | 354 |
|-> (fn e => pair e) |
20600 | 355 |
| exprgen_term thy algbr funcgr strct (Var _) trns = |
20389 | 356 |
error "Var encountered in term during code generation" |
20600 | 357 |
| exprgen_term thy algbr funcgr strct (Free (v, ty)) trns = |
18516 | 358 |
trns |
20600 | 359 |
|> exprgen_type thy algbr funcgr strct ty |
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
360 |
|-> (fn ty => pair (IVar v)) |
20600 | 361 |
| exprgen_term thy algbr funcgr strct (Abs (raw_v, ty, raw_t)) trns = |
19136 | 362 |
let |
20386 | 363 |
val (v, t) = Syntax.variant_abs (CodegenNames.purify_var raw_v, ty, raw_t); |
19136 | 364 |
in |
365 |
trns |
|
20600 | 366 |
|> exprgen_type thy algbr funcgr strct ty |
367 |
||>> exprgen_term thy algbr funcgr strct t |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
368 |
|-> (fn (ty, e) => pair ((v, ty) `|-> e)) |
19136 | 369 |
end |
20600 | 370 |
| exprgen_term thy algbr funcgr strct (t as t1 $ t2) trns = |
18516 | 371 |
let |
372 |
val (t', ts) = strip_comb t |
|
373 |
in case t' |
|
374 |
of Const (f, ty) => |
|
375 |
trns |
|
20600 | 376 |
|> appgen thy algbr funcgr strct ((f, ty), ts) |
18516 | 377 |
|-> (fn e => pair e) |
378 |
| _ => |
|
379 |
trns |
|
20600 | 380 |
|> exprgen_term thy algbr funcgr strct t' |
381 |
||>> fold_map (exprgen_term thy algbr funcgr strct) ts |
|
18516 | 382 |
|-> (fn (e, es) => pair (e `$$ es)) |
18865 | 383 |
end |
20600 | 384 |
and appgen_default thy algbr funcgr strct ((c, ty), ts) trns = |
18865 | 385 |
trns |
20600 | 386 |
|> ensure_def_const thy algbr funcgr strct (CodegenConsts.norm_of_typ thy (c, ty)) |
387 |
||>> exprgen_type thy algbr funcgr strct ty |
|
388 |
||>> exprgen_typinst_const thy algbr funcgr strct (c, ty) |
|
389 |
||>> fold_map (exprgen_term thy algbr funcgr strct) ts |
|
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
390 |
|-> (fn (((c, ty), ls), es) => |
19202 | 391 |
pair (IConst (c, (ls, ty)) `$$ es)) |
20600 | 392 |
and appgen thy algbr funcgr strct ((f, ty), ts) trns = |
20699 | 393 |
case Symtab.lookup (CodegenPackageData.get thy) f |
20105 | 394 |
of SOME (i, (ag, _)) => |
395 |
if length ts < i then |
|
18865 | 396 |
let |
20105 | 397 |
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
|
398 |
val vs = Name.names (Name.declare f Name.context) "a" tys; |
18865 | 399 |
in |
400 |
trns |
|
20600 | 401 |
|> fold_map (exprgen_type thy algbr funcgr strct) tys |
402 |
||>> ag thy algbr funcgr strct ((f, ty), ts @ map Free vs) |
|
20105 | 403 |
|-> (fn (tys, e) => pair (map2 (fn (v, _) => pair v) vs tys `|--> e)) |
18865 | 404 |
end |
20105 | 405 |
else if length ts > i then |
18865 | 406 |
trns |
20600 | 407 |
|> ag thy algbr funcgr strct ((f, ty), Library.take (i, ts)) |
408 |
||>> fold_map (exprgen_term thy algbr funcgr strct) (Library.drop (i, ts)) |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
409 |
|-> (fn (e, es) => pair (e `$$ es)) |
18865 | 410 |
else |
411 |
trns |
|
20600 | 412 |
|> ag thy algbr funcgr strct ((f, ty), ts) |
18865 | 413 |
| NONE => |
414 |
trns |
|
20600 | 415 |
|> appgen_default thy algbr funcgr strct ((f, ty), ts); |
416 |
||
417 |
||
418 |
(* entry points into extraction kernel *) |
|
419 |
||
420 |
fun ensure_def_const' thy algbr funcgr strct c trns = |
|
421 |
ensure_def_const thy algbr funcgr strct c trns |
|
422 |
handle CONSTRAIN (((c, ty), ty_decl), NONE) => error ( |
|
423 |
"Constant " ^ c ^ " with most general type\n" |
|
424 |
^ Sign.string_of_typ thy ty |
|
425 |
^ "\noccurs with type\n" |
|
426 |
^ Sign.string_of_typ thy ty_decl) |
|
427 |
handle CONSTRAIN (((c, ty), ty_decl), SOME t) => error ("In term " ^ (quote o Sign.string_of_term thy) t |
|
428 |
^ ",\nconstant " ^ c ^ " with most general type\n" |
|
429 |
^ Sign.string_of_typ thy ty |
|
430 |
^ "\noccurs with type\n" |
|
431 |
^ Sign.string_of_typ thy ty_decl); |
|
432 |
||
433 |
fun exprgen_term' thy algbr funcgr strct t trns = |
|
434 |
exprgen_term thy algbr funcgr strct t trns |
|
435 |
handle CONSTRAIN (((c, ty), ty_decl), _) => error ("In term " ^ (quote o Sign.string_of_term thy) t |
|
436 |
^ ",\nconstant " ^ c ^ " with most general type\n" |
|
437 |
^ Sign.string_of_typ thy ty |
|
438 |
^ "\noccurs with type\n" |
|
439 |
^ Sign.string_of_typ thy ty_decl); |
|
18516 | 440 |
|
18702 | 441 |
|
20439 | 442 |
(* parametrized application generators, for instantiation in object logic *) |
443 |
(* (axiomatic extensions of extraction kernel *) |
|
18217 | 444 |
|
20600 | 445 |
fun appgen_numeral int_of_numeral thy algbr funcgr strct (app as (c, ts)) trns = |
20485 | 446 |
case try (int_of_numeral thy) (list_comb (Const c, ts)) |
20600 | 447 |
of SOME i => (*if i < 0 then (*preprocessor eliminates negative numerals*) |
20353 | 448 |
trns |
20600 | 449 |
|> appgen_default thy algbr funcgr (no_strict strct) app |
450 |
else*) |
|
19884 | 451 |
trns |
20600 | 452 |
|> appgen_default thy algbr funcgr (no_strict strct) app |
20485 | 453 |
|-> (fn e => pair (CodegenThingol.INum (i, e))) |
19884 | 454 |
| NONE => |
455 |
trns |
|
20600 | 456 |
|> appgen_default thy algbr funcgr (no_strict strct) app; |
18217 | 457 |
|
20600 | 458 |
fun appgen_char char_to_index thy algbr funcgr 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
|
459 |
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
|
460 |
of SOME i => |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
461 |
trns |
20600 | 462 |
|> exprgen_type thy algbr funcgr strct ty |
463 |
||>> appgen_default thy algbr funcgr strct app |
|
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
464 |
|-> (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
|
465 |
| NONE => |
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
466 |
trns |
20600 | 467 |
|> appgen_default thy algbr funcgr strct app; |
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
468 |
|
20600 | 469 |
fun appgen_case dest_case_expr thy algbr funcgr strct (app as (c_ty, ts)) trns = |
20105 | 470 |
let |
471 |
val SOME ([], ((st, sty), ds)) = dest_case_expr thy (list_comb (Const c_ty, ts)); |
|
472 |
fun clausegen (dt, bt) trns = |
|
473 |
trns |
|
20600 | 474 |
|> exprgen_term thy algbr funcgr strct dt |
475 |
||>> exprgen_term thy algbr funcgr strct bt; |
|
20105 | 476 |
in |
477 |
trns |
|
20600 | 478 |
|> exprgen_term thy algbr funcgr strct st |
479 |
||>> exprgen_type thy algbr funcgr strct sty |
|
20105 | 480 |
||>> fold_map clausegen ds |
20600 | 481 |
||>> appgen_default thy algbr funcgr strct app |
20105 | 482 |
|-> (fn (((se, sty), ds), e0) => pair (ICase (((se, sty), ds), e0))) |
483 |
end; |
|
484 |
||
20600 | 485 |
fun appgen_let thy algbr funcgr strct (app as (_, [st, ct])) trns = |
20105 | 486 |
trns |
20600 | 487 |
|> exprgen_term thy algbr funcgr strct ct |
488 |
||>> exprgen_term thy algbr funcgr strct st |
|
489 |
||>> appgen_default thy algbr funcgr strct app |
|
20105 | 490 |
|-> (fn (((v, ty) `|-> be, se), e0) => |
491 |
pair (ICase (((se, ty), case be |
|
492 |
of ICase (((IVar w, _), ds), _) => if v = w then ds else [(IVar v, be)] |
|
493 |
| _ => [(IVar v, be)] |
|
494 |
), e0)) |
|
495 |
| (_, e0) => pair e0); |
|
496 |
||
20439 | 497 |
fun add_appconst (c, appgen) thy = |
498 |
let |
|
499 |
val i = (length o fst o strip_type o Sign.the_const_type thy) c |
|
20699 | 500 |
in CodegenPackageData.map (Symtab.update (c, (i, (appgen, stamp ())))) thy end; |
20439 | 501 |
|
18217 | 502 |
|
18516 | 503 |
|
20439 | 504 |
(** code generation interfaces **) |
18516 | 505 |
|
20600 | 506 |
fun generate thy (cs, rhss) targets init gen it = |
20466 | 507 |
let |
20600 | 508 |
val funcgr = CodegenFuncgr.mk_funcgr thy cs rhss; |
20466 | 509 |
val qnaming = NameSpace.qualified_names NameSpace.default_naming |
510 |
val algebr = ClassPackage.operational_algebra thy; |
|
511 |
val consttab = Consts.empty |
|
512 |
|> fold (fn (c, ty) => Consts.declare qnaming |
|
20699 | 513 |
((CodegenNames.const thy c, ty), true)) |
20600 | 514 |
(CodegenFuncgr.get_func_typs funcgr) |
20466 | 515 |
val algbr = (algebr, consttab); |
516 |
in |
|
20600 | 517 |
Code.change_yield thy (start_transact init (gen thy algbr funcgr |
518 |
(true, targets) it)) |
|
519 |
|> (fn (x, modl) => x) |
|
20466 | 520 |
end; |
18516 | 521 |
|
20600 | 522 |
fun codegen_term thy t = |
20353 | 523 |
let |
20600 | 524 |
val ct = Thm.cterm_of thy t; |
525 |
val thm = CodegenData.preprocess_cterm thy ct; |
|
526 |
val t' = (snd o Logic.dest_equals o Drule.plain_prop_of) thm; |
|
20699 | 527 |
val cs_rhss = CodegenConsts.consts_of thy t'; |
20353 | 528 |
in |
20600 | 529 |
(thm, generate thy cs_rhss (SOME []) NONE exprgen_term' t') |
20353 | 530 |
end; |
19136 | 531 |
|
20699 | 532 |
fun const_of_idf thy = |
533 |
CodegenConsts.typ_of_inst thy o the o CodegenNames.const_rev thy; |
|
19150 | 534 |
|
19967 | 535 |
fun get_root_module thy = |
20600 | 536 |
Code.get thy; |
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
537 |
|
20600 | 538 |
fun eval_term thy (ref_spec, t) = |
20213 | 539 |
let |
20401 | 540 |
val _ = Term.fold_atyps (fn _ => |
541 |
error ("Term" ^ Sign.string_of_term thy t ^ "is polymorhpic")) |
|
542 |
(Term.fastype_of t); |
|
20389 | 543 |
fun preprocess_term t = |
544 |
let |
|
545 |
val x = Free (Name.variant (add_term_names (t, [])) "x", fastype_of t); |
|
546 |
(* fake definition *) |
|
547 |
val eq = setmp quick_and_dirty true (SkipProof.make_thm thy) |
|
548 |
(Logic.mk_equals (x, t)); |
|
549 |
fun err () = error "preprocess_term: bad preprocessor" |
|
20600 | 550 |
in case map prop_of (CodegenFuncgr.preprocess thy [eq]) |
20389 | 551 |
of [Const ("==", _) $ x' $ t'] => if x = x' then t' else err () |
552 |
| _ => err () |
|
553 |
end; |
|
20600 | 554 |
val (_, t') = codegen_term thy (preprocess_term t); |
20428 | 555 |
in |
20699 | 556 |
CodegenSerializer.eval_term thy (ref_spec, t') (Code.get thy) |
18217 | 557 |
end; |
558 |
||
559 |
||
18516 | 560 |
|
20439 | 561 |
(** toplevel interface and setup **) |
18756 | 562 |
|
563 |
local |
|
19150 | 564 |
|
20699 | 565 |
structure P = OuterParse |
566 |
and K = OuterKeyword |
|
567 |
||
20439 | 568 |
fun code raw_cs seris thy = |
18217 | 569 |
let |
20600 | 570 |
val cs = map (CodegenConsts.read_const thy) raw_cs; |
20456 | 571 |
val targets = case map fst seris |
572 |
of [] => NONE |
|
573 |
| xs => SOME xs; |
|
20439 | 574 |
val seris' = map_filter (fn (target, SOME seri) => SOME (target, seri) | _ => NONE) seris; |
575 |
fun generate' thy = case cs |
|
20600 | 576 |
of [] => [] |
20439 | 577 |
| _ => |
20600 | 578 |
generate thy (cs, []) targets NONE (fold_map oooo ensure_def_const') cs; |
20699 | 579 |
fun serialize' [] code (_, seri) = |
580 |
seri thy NONE code |
|
581 |
| serialize' cs code (_, seri) = |
|
582 |
seri thy (SOME cs) code; |
|
20600 | 583 |
val cs = generate' thy; |
20699 | 584 |
val code = Code.get thy; |
18217 | 585 |
in |
20699 | 586 |
(map (serialize' cs code) seris'; ()) |
18217 | 587 |
end; |
588 |
||
20699 | 589 |
fun reader_tyco thy rhss target dep = |
590 |
generate thy rhss (SOME [target]) (SOME dep) (fold_map oooo exprgen_type); |
|
18217 | 591 |
|
20699 | 592 |
fun reader_const thy rhss target dep = |
593 |
generate thy rhss (SOME [target]) (SOME dep) (fold_map oooo exprgen_term'); |
|
594 |
||
595 |
val parse_target = P.name >> tap CodegenSerializer.check_serializer; |
|
20456 | 596 |
|
597 |
fun zip_list (x::xs) f g = |
|
598 |
f x #-> (fn y => fold_map (fn x => g |-- f x >> pair x) xs |
|
599 |
#-> (fn xys => pair ((x, y) :: xys))); |
|
600 |
||
601 |
fun parse_multi_syntax parse_thing parse_syntax = |
|
602 |
P.and_list1 parse_thing |
|
603 |
#-> (fn things => Scan.repeat1 (P.$$$ "(" |-- parse_target :-- |
|
604 |
(fn target => zip_list things (parse_syntax target) |
|
605 |
(P.$$$ "and")) --| P.$$$ ")")) |
|
606 |
||
20699 | 607 |
|
18217 | 608 |
in |
609 |
||
20456 | 610 |
val (codeK, |
20600 | 611 |
syntax_classK, syntax_instK, syntax_tycoK, syntax_constK) = |
20456 | 612 |
("code_gen", |
20600 | 613 |
"code_class", "code_instance", "code_type", "code_const"); |
18335 | 614 |
|
20439 | 615 |
val codeP = |
20600 | 616 |
OuterSyntax.command codeK "generate and serialize executable code for constants" K.diag ( |
20439 | 617 |
Scan.repeat P.term |
618 |
-- Scan.repeat (P.$$$ "(" |-- |
|
20699 | 619 |
P.name :-- (fn target => (CodegenSerializer.parse_serializer target >> SOME) || pair NONE) |
20439 | 620 |
--| P.$$$ ")") |
20600 | 621 |
>> (fn (raw_cs, seris) => Toplevel.keep (code raw_cs seris o Toplevel.theory_of)) |
20439 | 622 |
); |
623 |
||
18865 | 624 |
val syntax_classP = |
19884 | 625 |
OuterSyntax.command syntax_classK "define code syntax for class" K.thy_decl ( |
20456 | 626 |
parse_multi_syntax P.xname |
627 |
(fn _ => fn _ => P.string -- Scan.optional (P.$$$ "where" |-- Scan.repeat1 |
|
628 |
(P.term --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.string)) []) |
|
629 |
>> (Toplevel.theory oo fold) (fn (target, syns) => |
|
20699 | 630 |
fold (fn (raw_class, syn) => CodegenSerializer.add_syntax_class target raw_class syn) syns) |
18865 | 631 |
); |
632 |
||
20428 | 633 |
val syntax_instP = |
634 |
OuterSyntax.command syntax_instK "define code syntax for instance" K.thy_decl ( |
|
20456 | 635 |
parse_multi_syntax (P.xname --| P.$$$ "::" -- P.xname) |
636 |
(fn _ => fn _ => P.name #-> |
|
637 |
(fn "-" => Scan.succeed () | _ => Scan.fail_with (fn _ => "\"-\" expected") ())) |
|
638 |
>> (Toplevel.theory oo fold) (fn (target, syns) => |
|
20699 | 639 |
fold (fn (raw_inst, ()) => CodegenSerializer.add_syntax_inst target raw_inst) syns) |
20428 | 640 |
); |
641 |
||
18217 | 642 |
val syntax_tycoP = |
643 |
OuterSyntax.command syntax_tycoK "define code syntax for type constructor" K.thy_decl ( |
|
18702 | 644 |
Scan.repeat1 ( |
20699 | 645 |
parse_multi_syntax P.xname (CodegenSerializer.parse_syntax_tyco reader_tyco) |
18702 | 646 |
) |
20456 | 647 |
>> (Toplevel.theory o (fold o fold) (fold snd o snd)) |
18217 | 648 |
); |
649 |
||
650 |
val syntax_constP = |
|
651 |
OuterSyntax.command syntax_constK "define code syntax for constant" K.thy_decl ( |
|
18702 | 652 |
Scan.repeat1 ( |
20699 | 653 |
parse_multi_syntax P.term (CodegenSerializer.parse_syntax_const reader_const) |
18702 | 654 |
) |
20456 | 655 |
>> (Toplevel.theory o (fold o fold) (fold snd o snd)) |
18217 | 656 |
); |
657 |
||
20456 | 658 |
val _ = OuterSyntax.add_parsers [codeP, |
20600 | 659 |
syntax_classP, syntax_instP, syntax_tycoP, syntax_constP]; |
18217 | 660 |
|
661 |
end; (* local *) |
|
662 |
||
663 |
end; (* struct *) |