| author | wenzelm |
| Wed, 11 Oct 2006 22:59:36 +0200 | |
| changeset 20988 | 0887d0dd3210 |
| parent 20976 | e324808e9f1f |
| child 21012 | f08574148b7a |
| 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 |
|
| 20855 | 5 |
Code generator extraction kernel. Code generator Isar setup. |
|
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
6 |
*) |
|
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 |
signature CODEGEN_PACKAGE = |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
9 |
sig |
| 20456 | 10 |
include BASIC_CODEGEN_THINGOL; |
| 20600 | 11 |
val codegen_term: theory -> term -> thm * iterm; |
12 |
val eval_term: theory -> (string (*reference name!*) * 'a option ref) * term -> 'a; |
|
| 20699 | 13 |
val const_of_idf: theory -> string -> string * typ; |
| 20896 | 14 |
val get_root_module: theory -> CodegenThingol.code; |
| 18217 | 15 |
|
| 19884 | 16 |
type appgen; |
| 20439 | 17 |
val add_appconst: string * appgen -> theory -> theory; |
| 20600 | 18 |
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
|
19 |
val appgen_char: (term -> int option) -> appgen; |
| 20105 | 20 |
val appgen_case: (theory -> term |
21 |
-> ((string * typ) list * ((term * typ) * (term * term) list)) option) |
|
22 |
-> appgen; |
|
23 |
val appgen_let: appgen; |
|
|
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
24 |
end; |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
25 |
|
| 18217 | 26 |
structure CodegenPackage : CODEGEN_PACKAGE = |
|
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
27 |
struct |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
28 |
|
| 20896 | 29 |
open BasicCodegenThingol; |
30 |
val tracing = CodegenThingol.tracing; |
|
31 |
val succeed = CodegenThingol.succeed; |
|
32 |
val fail = CodegenThingol.fail; |
|
| 20699 | 33 |
|
34 |
(** code extraction **) |
|
35 |
||
36 |
(* theory data *) |
|
| 18217 | 37 |
|
| 20931 | 38 |
structure Code = CodeDataFun |
39 |
(struct |
|
40 |
val name = "Pure/code"; |
|
41 |
type T = CodegenThingol.code; |
|
42 |
val empty = CodegenThingol.empty_code; |
|
43 |
fun merge _ = CodegenThingol.merge_code; |
|
44 |
fun purge _ NONE _ = CodegenThingol.empty_code |
|
45 |
| purge NONE _ _ = CodegenThingol.empty_code |
|
46 |
| purge (SOME thy) (SOME cs) code = |
|
47 |
let |
|
48 |
val cs_exisiting = |
|
49 |
map_filter (CodegenNames.const_rev thy) (Graph.keys code); |
|
50 |
in |
|
51 |
Graph.del_nodes |
|
52 |
((Graph.all_succs code |
|
53 |
o map (CodegenNames.const thy) |
|
54 |
o filter (member CodegenConsts.eq_const cs_exisiting) |
|
55 |
) cs) |
|
56 |
code |
|
57 |
end; |
|
58 |
end); |
|
59 |
||
| 20600 | 60 |
type appgen = theory -> ((sort -> sort) * Sorts.algebra) * Consts.T |
61 |
-> CodegenFuncgr.T |
|
| 20896 | 62 |
-> bool * string list option |
63 |
-> (string * typ) * term list -> CodegenThingol.transact -> iterm * CodegenThingol.transact; |
|
| 20439 | 64 |
|
65 |
type appgens = (int * (appgen * stamp)) Symtab.table; |
|
| 20931 | 66 |
val merge_appgens : appgens * appgens -> appgens = |
| 20105 | 67 |
Symtab.merge (fn ((bounds1, (_, stamp1)), (bounds2, (_, stamp2))) => |
| 20931 | 68 |
bounds1 = bounds2 andalso stamp1 = stamp2); |
| 18217 | 69 |
|
| 20931 | 70 |
structure Consttab = CodegenConsts.Consttab; |
71 |
type abstypes = typ Symtab.table * CodegenConsts.const Consttab.table; |
|
72 |
fun merge_abstypes ((typs1, consts1) : abstypes, (typs2, consts2) : abstypes) = |
|
73 |
(Symtab.merge (Type.eq_type Vartab.empty) (typs1, typs2), |
|
74 |
Consttab.merge CodegenConsts.eq_const (consts1, consts2)); |
|
| 20456 | 75 |
|
| 20600 | 76 |
structure CodegenPackageData = TheoryDataFun |
| 18217 | 77 |
(struct |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19884
diff
changeset
|
78 |
val name = "Pure/codegen_package"; |
| 20931 | 79 |
type T = appgens * abstypes; |
80 |
val empty = (Symtab.empty, (Symtab.empty, Consttab.empty)); |
|
| 18217 | 81 |
val copy = I; |
82 |
val extend = I; |
|
| 20931 | 83 |
fun merge _ ((appgens1, abstypes1), (appgens2, abstypes2)) = |
84 |
(merge_appgens (appgens1, appgens2), merge_abstypes (abstypes1, abstypes2)); |
|
| 20456 | 85 |
fun print _ _ = (); |
| 18217 | 86 |
end); |
87 |
||
| 20600 | 88 |
val _ = Context.add_setup (Code.init #> CodegenPackageData.init); |
| 18708 | 89 |
|
| 18865 | 90 |
|
| 20386 | 91 |
(* extraction kernel *) |
| 18865 | 92 |
|
| 20931 | 93 |
fun check_strict (false, _) has_seri x = |
| 19884 | 94 |
false |
| 20931 | 95 |
| check_strict (_, SOME targets) has_seri x = |
| 20699 | 96 |
not (has_seri targets x) |
| 20931 | 97 |
| check_strict (true, _) has_seri x = |
| 19884 | 98 |
true; |
99 |
||
| 20931 | 100 |
fun get_abstype thy (tyco, tys) = case Symtab.lookup ((fst o snd o CodegenPackageData.get) thy) tyco |
101 |
of SOME ty => SOME ((map_atyps (fn TFree (n, _) => nth tys (the (Int.fromString n)))) ty) |
|
102 |
| NONE => NONE; |
|
103 |
||
| 20896 | 104 |
fun ensure_def_class thy (algbr as ((proj_sort, _), _)) funcgr strct class trns = |
| 20386 | 105 |
let |
| 20896 | 106 |
val (v, cs) = (ClassPackage.the_consts_sign thy) class; |
107 |
val superclasses = (proj_sort o Sign.super_classes thy) class; |
|
108 |
val classops' = map (CodegenNames.const thy o CodegenConsts.norm_of_typ thy) cs; |
|
109 |
val class' = CodegenNames.class thy class; |
|
110 |
fun defgen_class trns = |
|
111 |
trns |
|
112 |
|> fold_map (ensure_def_class thy algbr funcgr strct) superclasses |
|
113 |
||>> (fold_map (exprgen_type thy algbr funcgr strct) o map snd) cs |
|
114 |
|-> (fn (superclasses, classoptyps) => succeed |
|
115 |
(CodegenThingol.Class (superclasses, |
|
116 |
(unprefix "'" v, classops' ~~ classoptyps)))) |
|
| 18454 | 117 |
in |
118 |
trns |
|
| 20896 | 119 |
|> tracing (fn _ => "generating class " ^ quote class) |
120 |
|> CodegenThingol.ensure_def defgen_class true |
|
121 |
("generating class " ^ quote class) class'
|
|
122 |
|> pair class' |
|
| 18865 | 123 |
end |
| 20600 | 124 |
and ensure_def_tyco thy algbr funcgr strct tyco trns = |
| 18865 | 125 |
let |
| 20896 | 126 |
fun defgen_datatype trns = |
127 |
case CodegenData.get_datatype thy tyco |
|
128 |
of SOME (vs, cos) => |
|
129 |
trns |
|
130 |
|> fold_map (exprgen_tyvar_sort thy algbr funcgr strct) vs |
|
131 |
||>> fold_map (fn (c, tys) => |
|
132 |
fold_map (exprgen_type thy algbr funcgr strct) tys |
|
133 |
#-> (fn tys' => |
|
134 |
pair ((CodegenNames.const thy o CodegenConsts.norm_of_typ thy) |
|
135 |
(c, tys ---> Type (tyco, map TFree vs)), tys'))) cos |
|
136 |
|-> (fn (vs, cos) => succeed (CodegenThingol.Datatype (vs, cos))) |
|
137 |
| NONE => |
|
138 |
trns |
|
139 |
|> fail ("No such datatype: " ^ quote tyco)
|
|
| 20699 | 140 |
val tyco' = CodegenNames.tyco thy tyco; |
| 20931 | 141 |
val strict = check_strict strct (CodegenSerializer.tyco_has_serialization thy) tyco'; |
| 18865 | 142 |
in |
143 |
trns |
|
| 20835 | 144 |
|> tracing (fn _ => "generating type constructor " ^ quote tyco) |
| 20896 | 145 |
|> CodegenThingol.ensure_def defgen_datatype strict |
| 20389 | 146 |
("generating type constructor " ^ quote tyco) tyco'
|
| 18865 | 147 |
|> pair tyco' |
148 |
end |
|
| 20600 | 149 |
and exprgen_tyvar_sort thy (algbr as ((proj_sort, _), _)) funcgr strct (v, sort) trns = |
| 18516 | 150 |
trns |
| 20600 | 151 |
|> fold_map (ensure_def_class thy algbr funcgr strct) (proj_sort sort) |
| 18865 | 152 |
|-> (fn sort => pair (unprefix "'" v, sort)) |
| 20600 | 153 |
and exprgen_type thy algbr funcgr strct (TVar _) trns = |
| 20389 | 154 |
error "TVar encountered in typ during code generation" |
| 20600 | 155 |
| exprgen_type thy algbr funcgr strct (TFree vs) trns = |
| 18516 | 156 |
trns |
| 20600 | 157 |
|> exprgen_tyvar_sort thy algbr funcgr strct vs |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
158 |
|-> (fn (v, sort) => pair (ITyVar v)) |
| 20600 | 159 |
| exprgen_type thy algbr funcgr strct (Type ("fun", [t1, t2])) trns =
|
| 18516 | 160 |
trns |
| 20600 | 161 |
|> exprgen_type thy algbr funcgr strct t1 |
162 |
||>> exprgen_type thy algbr funcgr strct t2 |
|
| 18516 | 163 |
|-> (fn (t1', t2') => pair (t1' `-> t2')) |
| 20600 | 164 |
| exprgen_type thy algbr funcgr strct (Type (tyco, tys)) trns = |
| 20931 | 165 |
case get_abstype thy (tyco, tys) |
166 |
of SOME ty => |
|
167 |
trns |
|
168 |
|> exprgen_type thy algbr funcgr strct ty |
|
169 |
| NONE => |
|
170 |
trns |
|
171 |
|> ensure_def_tyco thy algbr funcgr strct tyco |
|
172 |
||>> fold_map (exprgen_type thy algbr funcgr strct) tys |
|
173 |
|-> (fn (tyco, tys) => pair (tyco `%% tys)); |
|
| 18516 | 174 |
|
| 20835 | 175 |
exception CONSTRAIN of (string * typ) * typ; |
| 20600 | 176 |
|
177 |
fun exprgen_typinst thy (algbr as ((proj_sort, algebra), consts)) funcgr strct (ty_ctxt, sort_decl) trns = |
|
| 20456 | 178 |
let |
179 |
val pp = Sign.pp thy; |
|
180 |
datatype inst = |
|
181 |
Inst of (class * string) * inst list list |
|
182 |
| Contxt of (string * sort) * (class list * int); |
|
183 |
fun classrel (l as Contxt (v_sort, (classes, n)), _) class = |
|
184 |
Contxt (v_sort, (class :: classes, n)) |
|
185 |
| classrel (Inst ((_, tyco), lss), _) class = |
|
186 |
Inst ((class, tyco), lss); |
|
187 |
fun constructor tyco iss class = |
|
188 |
Inst ((class, tyco), (map o map) fst iss); |
|
189 |
fun variable (TFree (v, sort)) = |
|
190 |
let |
|
191 |
val sort' = proj_sort sort; |
|
192 |
in map_index (fn (n, class) => (Contxt ((v, sort'), ([], n)), class)) sort' end; |
|
193 |
val insts = Sorts.of_sort_derivation pp algebra |
|
194 |
{classrel = classrel, constructor = constructor, variable = variable}
|
|
195 |
(ty_ctxt, proj_sort sort_decl); |
|
196 |
fun mk_dict (Inst (inst, instss)) trns = |
|
197 |
trns |
|
| 20600 | 198 |
|> ensure_def_inst thy algbr funcgr strct inst |
| 20456 | 199 |
||>> (fold_map o fold_map) mk_dict instss |
200 |
|-> (fn (inst, instss) => pair (Instance (inst, instss))) |
|
201 |
| mk_dict (Contxt ((v, sort), (classes, k))) trns = |
|
| 18517 | 202 |
trns |
| 20600 | 203 |
|> fold_map (ensure_def_class thy algbr funcgr strct) classes |
| 20456 | 204 |
|-> (fn classes => pair (Context (classes, (unprefix "'" v, |
205 |
if length sort = 1 then ~1 else k)))) |
|
206 |
in |
|
207 |
trns |
|
208 |
|> fold_map mk_dict insts |
|
209 |
end |
|
| 20600 | 210 |
and exprgen_typinst_const thy (algbr as (_, consts)) funcgr strct (c, ty_ctxt) trns = |
211 |
let |
|
212 |
val c' = CodegenConsts.norm_of_typ thy (c, ty_ctxt) |
|
| 20699 | 213 |
val idf = CodegenNames.const thy c'; |
| 20466 | 214 |
val ty_decl = Consts.declaration consts idf; |
215 |
val insts = (op ~~ o apsnd (map (snd o dest_TVar)) oo pairself) |
|
216 |
(curry (Consts.typargs consts) idf) (ty_ctxt, ty_decl); |
|
| 20600 | 217 |
val _ = if exists not (map (Sign.of_sort thy) insts) |
| 20835 | 218 |
then raise CONSTRAIN ((c, ty_decl), ty_ctxt) else (); |
| 20456 | 219 |
in |
220 |
trns |
|
| 20600 | 221 |
|> fold_map (exprgen_typinst thy algbr funcgr strct) insts |
| 20456 | 222 |
end |
| 20896 | 223 |
and ensure_def_inst thy (algbr as ((proj_sort, _), _)) funcgr strct (class, tyco) trns = |
| 20456 | 224 |
let |
| 20896 | 225 |
val (vs, classop_defs) = ((apsnd o map) Const o ClassPackage.the_inst_sign thy) |
226 |
(class, tyco); |
|
227 |
val classops = (map (CodegenConsts.norm_of_typ thy) o snd |
|
228 |
o ClassPackage.the_consts_sign thy) class; |
|
229 |
val arity_typ = Type (tyco, (map TFree vs)); |
|
230 |
val superclasses = (proj_sort o Sign.super_classes thy) class |
|
231 |
fun gen_superarity superclass trns = |
|
232 |
trns |
|
233 |
|> ensure_def_class thy algbr funcgr strct superclass |
|
234 |
||>> exprgen_typinst thy algbr funcgr strct (arity_typ, [superclass]) |
|
235 |
|-> (fn (superclass, [Instance (inst, iss)]) => pair (superclass, (inst, iss))); |
|
236 |
fun gen_classop_def (classop, t) trns = |
|
237 |
trns |
|
238 |
|> ensure_def_const thy algbr funcgr strct classop |
|
239 |
||>> exprgen_term thy algbr funcgr strct t; |
|
240 |
fun defgen_inst trns = |
|
241 |
trns |
|
242 |
|> ensure_def_class thy algbr funcgr strct class |
|
243 |
||>> ensure_def_tyco thy algbr funcgr strct tyco |
|
244 |
||>> fold_map (exprgen_tyvar_sort thy algbr funcgr strct) vs |
|
245 |
||>> fold_map gen_superarity superclasses |
|
246 |
||>> fold_map gen_classop_def (classops ~~ classop_defs) |
|
247 |
|-> (fn ((((class, tyco), arity), superarities), classops) => |
|
248 |
succeed (CodegenThingol.Classinst ((class, (tyco, arity)), (superarities, classops)))); |
|
249 |
val inst = CodegenNames.instance thy (class, tyco); |
|
| 18865 | 250 |
in |
251 |
trns |
|
| 20896 | 252 |
|> tracing (fn _ => "generating instance " ^ quote class ^ " / " ^ quote tyco) |
253 |
|> CodegenThingol.ensure_def defgen_inst true |
|
254 |
("generating instance " ^ quote class ^ " / " ^ quote tyco) inst
|
|
| 18865 | 255 |
|> pair inst |
256 |
end |
|
| 20896 | 257 |
and ensure_def_const thy (algbr as (_, consts)) funcgr strct (c, tys) trns = |
| 18865 | 258 |
let |
| 20896 | 259 |
val c' = CodegenNames.const thy (c, tys); |
260 |
fun defgen_datatypecons trns = |
|
261 |
trns |
|
262 |
|> ensure_def_tyco thy algbr funcgr strct |
|
263 |
((the o CodegenData.get_datatype_of_constr thy) (c, tys)) |
|
264 |
|-> (fn _ => succeed CodegenThingol.Bot); |
|
265 |
fun defgen_classop trns = |
|
266 |
trns |
|
267 |
|> ensure_def_class thy algbr funcgr strct ((the o AxClass.class_of_param thy) c) |
|
268 |
|-> (fn _ => succeed CodegenThingol.Bot); |
|
269 |
fun defgen_fun trns = |
|
| 20931 | 270 |
case CodegenFuncgr.get_funcs funcgr |
271 |
(perhaps (Consttab.lookup ((snd o snd o CodegenPackageData.get) thy)) (c, tys)) |
|
| 20456 | 272 |
of eq_thms as eq_thm :: _ => |
273 |
let |
|
274 |
val msg = cat_lines ("generating code for theorems " :: map string_of_thm eq_thms);
|
|
| 20600 | 275 |
val ty = (Logic.unvarifyT o CodegenData.typ_func thy) eq_thm; |
| 20466 | 276 |
val vs = (map dest_TFree o Consts.typargs consts) (c', ty); |
| 20835 | 277 |
val dest_eqthm = |
278 |
apfst (snd o strip_comb) o Logic.dest_equals o Logic.unvarify o prop_of; |
|
| 20456 | 279 |
fun exprgen_eq (args, rhs) trns = |
280 |
trns |
|
| 20600 | 281 |
|> fold_map (exprgen_term thy algbr funcgr strct) args |
282 |
||>> exprgen_term thy algbr funcgr strct rhs; |
|
| 20456 | 283 |
in |
284 |
trns |
|
| 20896 | 285 |
|> CodegenThingol.message msg (fn trns => trns |
| 20456 | 286 |
|> fold_map (exprgen_eq o dest_eqthm) eq_thms |
| 20600 | 287 |
||>> fold_map (exprgen_tyvar_sort thy algbr funcgr strct) vs |
288 |
||>> exprgen_type thy algbr funcgr strct ty |
|
| 20896 | 289 |
|-> (fn ((eqs, vs), ty) => succeed (CodegenThingol.Fun (eqs, (vs, ty))))) |
| 20456 | 290 |
end |
291 |
| [] => |
|
292 |
trns |
|
293 |
|> fail ("No defining equations found for "
|
|
| 20600 | 294 |
^ (quote o CodegenConsts.string_of_const thy) (c, tys)); |
| 20896 | 295 |
val defgen = |
296 |
if CodegenNames.has_nsp CodegenNames.nsp_fun c' |
|
297 |
then defgen_fun |
|
298 |
else if CodegenNames.has_nsp CodegenNames.nsp_classop c' |
|
299 |
then defgen_classop |
|
300 |
else if CodegenNames.has_nsp CodegenNames.nsp_dtco c' |
|
301 |
then defgen_datatypecons |
|
302 |
else error ("Illegal shallow name space for constant: " ^ quote c');
|
|
| 20931 | 303 |
val strict = check_strict strct (CodegenSerializer.const_has_serialization thy) c'; |
| 18865 | 304 |
in |
305 |
trns |
|
| 20835 | 306 |
|> tracing (fn _ => "generating constant " |
| 20600 | 307 |
^ (quote o CodegenConsts.string_of_const thy) (c, tys)) |
| 20896 | 308 |
|> CodegenThingol.ensure_def defgen strict ("generating constant "
|
309 |
^ CodegenConsts.string_of_const thy (c, tys)) c' |
|
310 |
|> pair c' |
|
| 18865 | 311 |
end |
| 20896 | 312 |
and exprgen_term thy algbr funcgr strct (Const (c, ty)) trns = |
| 18517 | 313 |
trns |
| 20896 | 314 |
|> select_appgen thy algbr funcgr strct ((c, ty), []) |
| 20600 | 315 |
| exprgen_term thy algbr funcgr strct (Var _) trns = |
| 20389 | 316 |
error "Var encountered in term during code generation" |
| 20600 | 317 |
| exprgen_term thy algbr funcgr strct (Free (v, ty)) trns = |
| 18516 | 318 |
trns |
| 20600 | 319 |
|> exprgen_type thy algbr funcgr strct ty |
| 20896 | 320 |
|-> (fn _ => pair (IVar v)) |
| 20600 | 321 |
| exprgen_term thy algbr funcgr strct (Abs (raw_v, ty, raw_t)) trns = |
| 19136 | 322 |
let |
| 20386 | 323 |
val (v, t) = Syntax.variant_abs (CodegenNames.purify_var raw_v, ty, raw_t); |
| 19136 | 324 |
in |
325 |
trns |
|
| 20600 | 326 |
|> exprgen_type thy algbr funcgr strct ty |
327 |
||>> exprgen_term thy algbr funcgr strct t |
|
| 20896 | 328 |
|-> (fn (ty, t) => pair ((v, ty) `|-> t)) |
| 19136 | 329 |
end |
| 20896 | 330 |
| exprgen_term thy algbr funcgr strct (t as _ $ _) trns = |
331 |
case strip_comb t |
|
332 |
of (Const (c, ty), ts) => |
|
| 18516 | 333 |
trns |
| 20896 | 334 |
|> select_appgen thy algbr funcgr strct ((c, ty), ts) |
335 |
|-> (fn t => pair t) |
|
336 |
| (t', ts) => |
|
| 18516 | 337 |
trns |
| 20600 | 338 |
|> exprgen_term thy algbr funcgr strct t' |
339 |
||>> fold_map (exprgen_term thy algbr funcgr strct) ts |
|
| 20896 | 340 |
|-> (fn (t, ts) => pair (t `$$ ts)) |
| 20600 | 341 |
and appgen_default thy algbr funcgr strct ((c, ty), ts) trns = |
| 18865 | 342 |
trns |
| 20600 | 343 |
|> ensure_def_const thy algbr funcgr strct (CodegenConsts.norm_of_typ thy (c, ty)) |
344 |
||>> exprgen_type thy algbr funcgr strct ty |
|
345 |
||>> exprgen_typinst_const thy algbr funcgr strct (c, ty) |
|
346 |
||>> fold_map (exprgen_term thy algbr funcgr strct) ts |
|
| 20896 | 347 |
|-> (fn (((c, ty), iss), ts) => |
348 |
pair (IConst (c, (iss, ty)) `$$ ts)) |
|
349 |
and select_appgen thy algbr funcgr strct ((c, ty), ts) trns = |
|
| 20931 | 350 |
case Symtab.lookup (fst (CodegenPackageData.get thy)) c |
| 20896 | 351 |
of SOME (i, (appgen, _)) => |
| 20105 | 352 |
if length ts < i then |
| 18865 | 353 |
let |
| 20105 | 354 |
val tys = Library.take (i - length ts, ((fst o strip_type) ty)); |
| 20896 | 355 |
val vs = Name.names (Name.declare c Name.context) "a" tys; |
| 18865 | 356 |
in |
357 |
trns |
|
| 20600 | 358 |
|> fold_map (exprgen_type thy algbr funcgr strct) tys |
| 20896 | 359 |
||>> appgen thy algbr funcgr strct ((c, ty), ts @ map Free vs) |
360 |
|-> (fn (tys, t) => pair (map2 (fn (v, _) => pair v) vs tys `|--> t)) |
|
| 18865 | 361 |
end |
| 20105 | 362 |
else if length ts > i then |
| 18865 | 363 |
trns |
| 20896 | 364 |
|> appgen thy algbr funcgr strct ((c, ty), Library.take (i, ts)) |
| 20600 | 365 |
||>> fold_map (exprgen_term thy algbr funcgr strct) (Library.drop (i, ts)) |
| 20896 | 366 |
|-> (fn (t, ts) => pair (t `$$ ts)) |
| 18865 | 367 |
else |
368 |
trns |
|
| 20896 | 369 |
|> appgen thy algbr funcgr strct ((c, ty), ts) |
| 18865 | 370 |
| NONE => |
371 |
trns |
|
| 20896 | 372 |
|> appgen_default thy algbr funcgr strct ((c, ty), ts); |
| 20600 | 373 |
|
374 |
||
| 20835 | 375 |
(* entrance points into extraction kernel *) |
| 20600 | 376 |
|
377 |
fun ensure_def_const' thy algbr funcgr strct c trns = |
|
378 |
ensure_def_const thy algbr funcgr strct c trns |
|
| 20835 | 379 |
handle CONSTRAIN ((c, ty), ty_decl) => error ( |
| 20600 | 380 |
"Constant " ^ c ^ " with most general type\n" |
381 |
^ Sign.string_of_typ thy ty |
|
382 |
^ "\noccurs with type\n" |
|
383 |
^ Sign.string_of_typ thy ty_decl); |
|
384 |
fun exprgen_term' thy algbr funcgr strct t trns = |
|
385 |
exprgen_term thy algbr funcgr strct t trns |
|
| 20835 | 386 |
handle CONSTRAIN ((c, ty), ty_decl) => error ("In term " ^ (quote o Sign.string_of_term thy) t
|
| 20600 | 387 |
^ ",\nconstant " ^ c ^ " with most general type\n" |
388 |
^ Sign.string_of_typ thy ty |
|
389 |
^ "\noccurs with type\n" |
|
390 |
^ Sign.string_of_typ thy ty_decl); |
|
| 18516 | 391 |
|
| 18702 | 392 |
|
| 20439 | 393 |
(* parametrized application generators, for instantiation in object logic *) |
394 |
(* (axiomatic extensions of extraction kernel *) |
|
| 18217 | 395 |
|
| 20600 | 396 |
fun appgen_numeral int_of_numeral thy algbr funcgr strct (app as (c, ts)) trns = |
| 20485 | 397 |
case try (int_of_numeral thy) (list_comb (Const c, ts)) |
| 20835 | 398 |
of SOME i => |
| 20353 | 399 |
trns |
| 20835 | 400 |
|> appgen_default thy algbr funcgr strct app |
| 20896 | 401 |
|-> (fn t => pair (CodegenThingol.INum (i, t))) |
| 19884 | 402 |
| NONE => |
403 |
trns |
|
| 20835 | 404 |
|> appgen_default thy algbr funcgr strct app; |
| 18217 | 405 |
|
| 20600 | 406 |
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
|
407 |
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
|
408 |
of SOME i => |
|
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
409 |
trns |
| 20600 | 410 |
|> exprgen_type thy algbr funcgr strct ty |
411 |
||>> appgen_default thy algbr funcgr strct app |
|
| 20896 | 412 |
|-> (fn (_, t0) => pair (IChar (chr i, t0))) |
|
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
413 |
| NONE => |
|
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
414 |
trns |
| 20600 | 415 |
|> 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
|
416 |
|
| 20600 | 417 |
fun appgen_case dest_case_expr thy algbr funcgr strct (app as (c_ty, ts)) trns = |
| 20105 | 418 |
let |
419 |
val SOME ([], ((st, sty), ds)) = dest_case_expr thy (list_comb (Const c_ty, ts)); |
|
| 20976 | 420 |
fun fold_map_aterms f (t $ u) s = |
421 |
s |
|
422 |
|> fold_map_aterms f t |
|
423 |
||>> fold_map_aterms f u |
|
424 |
|-> (fn (t1, t2) => pair (t1 $ t2)) |
|
425 |
| fold_map_aterms f (Abs (v, ty, t)) s = |
|
426 |
s |
|
427 |
|> fold_map_aterms f t |
|
428 |
|-> (fn t' => pair (Abs (v, ty, t'))) |
|
429 |
| fold_map_aterms f a s = f a s; |
|
430 |
fun purify_term_frees (Free (v, ty)) ctxt = |
|
431 |
let |
|
432 |
val ([v'], ctxt') = Name.variants [CodegenNames.purify_var v] ctxt; |
|
433 |
in (Free (v, ty), ctxt') end |
|
434 |
| purify_term_frees t ctxt = (t, ctxt); |
|
435 |
fun clausegen (raw_dt, raw_bt) trns = |
|
436 |
let |
|
437 |
val ((dt, bt), _) = Name.context |
|
438 |
|> fold_map_aterms purify_term_frees raw_dt |
|
439 |
||>> fold_map_aterms purify_term_frees raw_bt; |
|
440 |
in |
|
441 |
trns |
|
442 |
|> exprgen_term thy algbr funcgr strct dt |
|
443 |
||>> exprgen_term thy algbr funcgr strct bt |
|
444 |
end; |
|
| 20105 | 445 |
in |
446 |
trns |
|
| 20600 | 447 |
|> exprgen_term thy algbr funcgr strct st |
448 |
||>> exprgen_type thy algbr funcgr strct sty |
|
| 20105 | 449 |
||>> fold_map clausegen ds |
| 20600 | 450 |
||>> appgen_default thy algbr funcgr strct app |
| 20105 | 451 |
|-> (fn (((se, sty), ds), e0) => pair (ICase (((se, sty), ds), e0))) |
452 |
end; |
|
453 |
||
| 20600 | 454 |
fun appgen_let thy algbr funcgr strct (app as (_, [st, ct])) trns = |
| 20105 | 455 |
trns |
| 20600 | 456 |
|> exprgen_term thy algbr funcgr strct ct |
457 |
||>> exprgen_term thy algbr funcgr strct st |
|
458 |
||>> appgen_default thy algbr funcgr strct app |
|
| 20105 | 459 |
|-> (fn (((v, ty) `|-> be, se), e0) => |
460 |
pair (ICase (((se, ty), case be |
|
461 |
of ICase (((IVar w, _), ds), _) => if v = w then ds else [(IVar v, be)] |
|
462 |
| _ => [(IVar v, be)] |
|
463 |
), e0)) |
|
464 |
| (_, e0) => pair e0); |
|
465 |
||
| 20439 | 466 |
fun add_appconst (c, appgen) thy = |
467 |
let |
|
| 20931 | 468 |
val i = (length o fst o strip_type o Sign.the_const_type thy) c; |
469 |
val _ = Code.change thy (K CodegenThingol.empty_code); |
|
470 |
in |
|
471 |
(CodegenPackageData.map o apfst) |
|
472 |
(Symtab.update (c, (i, (appgen, stamp ())))) thy |
|
473 |
end; |
|
474 |
||
475 |
||
476 |
||
477 |
(** abstype and constsubst interface **) |
|
478 |
||
479 |
fun gen_abstyp prep_const prep_typ (raw_abstyp, raw_substtyp) raw_absconsts thy = |
|
480 |
let |
|
481 |
val abstyp = Type.no_tvars (prep_typ thy raw_abstyp); |
|
482 |
val substtyp = Type.no_tvars (prep_typ thy raw_substtyp); |
|
483 |
val absconsts = (map o pairself) (prep_const thy) raw_absconsts; |
|
484 |
val Type (abstyco, tys) = abstyp handle BIND => error ("bad type: " ^ Sign.string_of_typ thy abstyp);
|
|
485 |
val typarms = map (fst o dest_TFree) tys handle MATCH => error ("bad type: " ^ Sign.string_of_typ thy abstyp);
|
|
486 |
fun mk_index v = |
|
487 |
let |
|
488 |
val k = find_index (fn w => v = w) typarms; |
|
489 |
in if k = ~1 |
|
490 |
then error ("free type variable: " ^ quote v)
|
|
491 |
else TFree (string_of_int k, []) |
|
492 |
end; |
|
493 |
val typpat = map_atyps (fn TFree (v, _) => mk_index v) substtyp; |
|
494 |
fun apply_typpat (Type (tyco, tys)) = |
|
495 |
let |
|
496 |
val tys' = map apply_typpat tys; |
|
497 |
in if tyco = abstyco then |
|
498 |
(map_atyps (fn TFree (n, _) => nth tys' (the (Int.fromString n)))) typpat |
|
499 |
else |
|
500 |
Type (tyco, tys') |
|
501 |
end |
|
502 |
| apply_typpat ty = ty; |
|
503 |
val string_of_typ = setmp show_sorts true (Sign.string_of_typ thy); |
|
504 |
fun add_consts (c1, c2) = |
|
505 |
let |
|
506 |
val _ = if CodegenNames.has_nsp CodegenNames.nsp_fun (CodegenNames.const thy c2) |
|
507 |
then () |
|
508 |
else error ("not a function: " ^ CodegenConsts.string_of_const thy c2);
|
|
509 |
val funcgr = CodegenFuncgr.mk_funcgr thy [c1, c2] []; |
|
510 |
val ty_map = CodegenFuncgr.get_func_typs funcgr; |
|
511 |
val ty1 = (apply_typpat o the o AList.lookup CodegenConsts.eq_const ty_map) c1; |
|
512 |
val ty2 = (the o AList.lookup CodegenConsts.eq_const ty_map) c2; |
|
513 |
val _ = if Sign.typ_equiv thy (ty1, ty2) then () else |
|
514 |
error ("Incompatiable type signatures of " ^ CodegenConsts.string_of_const thy c1
|
|
515 |
^ " and " ^ CodegenConsts.string_of_const thy c2 ^ ":\n" |
|
516 |
^ string_of_typ ty1 ^ "\n" ^ string_of_typ ty2); |
|
517 |
in Consttab.update (c1, c2) end; |
|
518 |
val _ = Code.change thy (K CodegenThingol.empty_code); |
|
519 |
in |
|
520 |
thy |
|
521 |
|> (CodegenPackageData.map o apsnd) (fn (abstypes, abscs) => |
|
522 |
(abstypes |
|
523 |
|> Symtab.update (abstyco, typpat), |
|
524 |
abscs |
|
525 |
|> fold add_consts absconsts) |
|
526 |
) |
|
527 |
end; |
|
528 |
||
529 |
fun gen_constsubst prep_const raw_constsubsts thy = |
|
530 |
let |
|
531 |
val constsubsts = (map o pairself) (prep_const thy) raw_constsubsts; |
|
532 |
val string_of_typ = setmp show_sorts true (Sign.string_of_typ thy); |
|
533 |
fun add_consts (c1, c2) = |
|
534 |
let |
|
535 |
val _ = if CodegenNames.has_nsp CodegenNames.nsp_fun (CodegenNames.const thy c2) |
|
536 |
then () |
|
537 |
else error ("not a function: " ^ CodegenConsts.string_of_const thy c2);
|
|
538 |
val funcgr = CodegenFuncgr.mk_funcgr thy [c1, c2] []; |
|
539 |
val ty_map = CodegenFuncgr.get_func_typs funcgr; |
|
540 |
val ty1 = (the o AList.lookup CodegenConsts.eq_const ty_map) c1; |
|
541 |
val ty2 = (the o AList.lookup CodegenConsts.eq_const ty_map) c2; |
|
542 |
val _ = if Sign.typ_equiv thy (ty1, ty2) then () else |
|
543 |
error ("Incompatiable type signatures of " ^ CodegenConsts.string_of_const thy c1
|
|
544 |
^ " and " ^ CodegenConsts.string_of_const thy c2 ^ ":\n" |
|
545 |
^ string_of_typ ty1 ^ "\n" ^ string_of_typ ty2); |
|
546 |
in Consttab.update (c1, c2) end; |
|
547 |
val _ = Code.change thy (K CodegenThingol.empty_code); |
|
548 |
in |
|
549 |
thy |
|
550 |
|> (CodegenPackageData.map o apsnd o apsnd) (fold add_consts constsubsts) |
|
551 |
end; |
|
552 |
||
553 |
val abstyp = gen_abstyp CodegenConsts.norm Sign.certify_typ; |
|
554 |
val abstyp_e = gen_abstyp CodegenConsts.read_const (fn thy => Sign.read_typ (thy, K NONE)); |
|
555 |
||
556 |
val constsubst = gen_constsubst CodegenConsts.norm; |
|
557 |
val constsubst_e = gen_constsubst CodegenConsts.read_const; |
|
| 20439 | 558 |
|
| 18217 | 559 |
|
| 18516 | 560 |
|
| 20439 | 561 |
(** code generation interfaces **) |
| 18516 | 562 |
|
| 20600 | 563 |
fun generate thy (cs, rhss) targets init gen it = |
| 20466 | 564 |
let |
| 20931 | 565 |
val raw_funcgr = CodegenFuncgr.mk_funcgr thy cs rhss; |
566 |
val cs' = map_filter (Consttab.lookup ((snd o snd o CodegenPackageData.get) thy)) |
|
567 |
(CodegenFuncgr.Constgraph.keys raw_funcgr); |
|
568 |
val funcgr = CodegenFuncgr.mk_funcgr thy cs' []; |
|
569 |
val qnaming = NameSpace.qualified_names NameSpace.default_naming; |
|
| 20466 | 570 |
val algebr = ClassPackage.operational_algebra thy; |
571 |
val consttab = Consts.empty |
|
572 |
|> fold (fn (c, ty) => Consts.declare qnaming |
|
| 20699 | 573 |
((CodegenNames.const thy c, ty), true)) |
| 20931 | 574 |
(CodegenFuncgr.get_func_typs funcgr); |
| 20466 | 575 |
val algbr = (algebr, consttab); |
576 |
in |
|
| 20896 | 577 |
Code.change_yield thy (CodegenThingol.start_transact init (gen thy algbr funcgr |
| 20600 | 578 |
(true, targets) it)) |
579 |
|> (fn (x, modl) => x) |
|
| 20466 | 580 |
end; |
| 18516 | 581 |
|
| 20600 | 582 |
fun codegen_term thy t = |
| 20353 | 583 |
let |
| 20835 | 584 |
fun const_typ (c, ty) = |
585 |
let |
|
586 |
val const = CodegenConsts.norm_of_typ thy (c, ty); |
|
587 |
val funcgr = CodegenFuncgr.mk_funcgr thy [const] []; |
|
588 |
in case CodegenFuncgr.get_funcs funcgr const |
|
589 |
of (thm :: _) => CodegenData.typ_func thy thm |
|
590 |
| [] => Sign.the_const_type thy c |
|
591 |
end; |
|
| 20600 | 592 |
val ct = Thm.cterm_of thy t; |
| 20835 | 593 |
val (thm, ct') = CodegenData.preprocess_cterm thy |
594 |
(const_typ) ct; |
|
595 |
val t' = Thm.term_of ct'; |
|
| 20699 | 596 |
val cs_rhss = CodegenConsts.consts_of thy t'; |
| 20353 | 597 |
in |
| 20600 | 598 |
(thm, generate thy cs_rhss (SOME []) NONE exprgen_term' t') |
| 20353 | 599 |
end; |
| 19136 | 600 |
|
| 20699 | 601 |
fun const_of_idf thy = |
602 |
CodegenConsts.typ_of_inst thy o the o CodegenNames.const_rev thy; |
|
| 19150 | 603 |
|
| 19967 | 604 |
fun get_root_module thy = |
| 20600 | 605 |
Code.get thy; |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
606 |
|
| 20600 | 607 |
fun eval_term thy (ref_spec, t) = |
| 20213 | 608 |
let |
| 20401 | 609 |
val _ = Term.fold_atyps (fn _ => |
610 |
error ("Term" ^ Sign.string_of_term thy t ^ "is polymorhpic"))
|
|
611 |
(Term.fastype_of t); |
|
| 20835 | 612 |
val (_, t') = codegen_term thy t; |
| 20428 | 613 |
in |
| 20699 | 614 |
CodegenSerializer.eval_term thy (ref_spec, t') (Code.get thy) |
| 18217 | 615 |
end; |
616 |
||
617 |
||
| 18516 | 618 |
|
| 20439 | 619 |
(** toplevel interface and setup **) |
| 18756 | 620 |
|
621 |
local |
|
| 19150 | 622 |
|
| 20699 | 623 |
structure P = OuterParse |
624 |
and K = OuterKeyword |
|
625 |
||
| 20439 | 626 |
fun code raw_cs seris thy = |
| 18217 | 627 |
let |
| 20600 | 628 |
val cs = map (CodegenConsts.read_const thy) raw_cs; |
| 20456 | 629 |
val targets = case map fst seris |
630 |
of [] => NONE |
|
631 |
| xs => SOME xs; |
|
| 20896 | 632 |
val seris' = map_filter (fn (target, args as _ :: _) => |
633 |
SOME (CodegenSerializer.get_serializer thy (target, args)) | _ => NONE) seris; |
|
| 20439 | 634 |
fun generate' thy = case cs |
| 20600 | 635 |
of [] => [] |
| 20439 | 636 |
| _ => |
| 20600 | 637 |
generate thy (cs, []) targets NONE (fold_map oooo ensure_def_const') cs; |
| 20896 | 638 |
fun serialize' [] code seri = |
639 |
seri NONE code |
|
640 |
| serialize' cs code seri = |
|
641 |
seri (SOME cs) code; |
|
| 20600 | 642 |
val cs = generate' thy; |
| 20699 | 643 |
val code = Code.get thy; |
| 18217 | 644 |
in |
| 20699 | 645 |
(map (serialize' cs code) seris'; ()) |
| 18217 | 646 |
end; |
647 |
||
| 20931 | 648 |
val (codeK, code_abstypeK, code_constsubstK) = |
649 |
("code_gen", "code_abstype", "code_constsubst");
|
|
| 20699 | 650 |
|
| 18217 | 651 |
in |
652 |
||
| 20439 | 653 |
val codeP = |
| 20835 | 654 |
OuterSyntax.improper_command codeK "generate and serialize executable code for constants" K.diag ( |
| 20439 | 655 |
Scan.repeat P.term |
656 |
-- Scan.repeat (P.$$$ "(" |--
|
|
| 20896 | 657 |
P.name -- P.arguments |
| 20439 | 658 |
--| P.$$$ ")") |
| 20600 | 659 |
>> (fn (raw_cs, seris) => Toplevel.keep (code raw_cs seris o Toplevel.theory_of)) |
| 20439 | 660 |
); |
661 |
||
| 20931 | 662 |
val code_abstypeP = |
663 |
OuterSyntax.command code_abstypeK "axiomatic abstypes for code generation" K.thy_decl ( |
|
664 |
(P.typ -- P.typ -- Scan.optional (P.$$$ "where" |-- Scan.repeat1 |
|
665 |
(P.term --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.term)) []) |
|
666 |
>> (Toplevel.theory o uncurry abstyp_e) |
|
| 20428 | 667 |
); |
668 |
||
| 20931 | 669 |
val code_constsubstP = |
670 |
OuterSyntax.command code_constsubstK "axiomatic constant substitutions for code generation" K.thy_decl ( |
|
671 |
Scan.repeat1 (P.term --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.term) |
|
672 |
>> (Toplevel.theory o constsubst_e) |
|
| 18217 | 673 |
); |
674 |
||
| 20931 | 675 |
val _ = OuterSyntax.add_parsers [codeP, code_abstypeP, code_constsubstP]; |
| 18217 | 676 |
|
677 |
end; (* local *) |
|
678 |
||
679 |
end; (* struct *) |