author | haftmann |
Tue, 24 Jul 2007 15:20:53 +0200 | |
changeset 23955 | f1ba12c117ec |
parent 23812 | f935b85fbb4c |
child 24166 | 7b28dc69bdbb |
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 |
|
23516 | 5 |
Code generator translation 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 |
23955 | 10 |
val compile_term: theory -> term -> CodegenThingol.code * CodegenThingol.iterm; |
20600 | 11 |
val eval_term: theory -> (string (*reference name!*) * 'a option ref) * term -> 'a; |
22464 | 12 |
val satisfies_ref: bool option ref; |
13 |
val satisfies: theory -> term -> string list -> bool; |
|
21881 | 14 |
val codegen_command: theory -> string -> unit; |
18217 | 15 |
|
23516 | 16 |
(*axiomatic interfaces*) |
19884 | 17 |
type appgen; |
20439 | 18 |
val add_appconst: string * appgen -> theory -> theory; |
23516 | 19 |
val appgen_let: appgen; |
20105 | 20 |
val appgen_case: (theory -> term |
21 |
-> ((string * typ) list * ((term * typ) * (term * term) list)) option) |
|
22 |
-> appgen; |
|
21012 | 23 |
|
24 |
val timing: bool ref; |
|
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
25 |
end; |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
26 |
|
18217 | 27 |
structure CodegenPackage : CODEGEN_PACKAGE = |
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
28 |
struct |
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
29 |
|
20896 | 30 |
open BasicCodegenThingol; |
31 |
val tracing = CodegenThingol.tracing; |
|
32 |
val succeed = CodegenThingol.succeed; |
|
33 |
val fail = CodegenThingol.fail; |
|
20699 | 34 |
|
23516 | 35 |
(** code translation **) |
20699 | 36 |
|
37 |
(* theory data *) |
|
18217 | 38 |
|
20600 | 39 |
type appgen = theory -> ((sort -> sort) * Sorts.algebra) * Consts.T |
40 |
-> CodegenFuncgr.T |
|
20896 | 41 |
-> (string * typ) * term list -> CodegenThingol.transact -> iterm * CodegenThingol.transact; |
20439 | 42 |
|
43 |
type appgens = (int * (appgen * stamp)) Symtab.table; |
|
20931 | 44 |
val merge_appgens : appgens * appgens -> appgens = |
20105 | 45 |
Symtab.merge (fn ((bounds1, (_, stamp1)), (bounds2, (_, stamp2))) => |
20931 | 46 |
bounds1 = bounds2 andalso stamp1 = stamp2); |
18217 | 47 |
|
20931 | 48 |
structure Consttab = CodegenConsts.Consttab; |
49 |
type abstypes = typ Symtab.table * CodegenConsts.const Consttab.table; |
|
50 |
fun merge_abstypes ((typs1, consts1) : abstypes, (typs2, consts2) : abstypes) = |
|
51 |
(Symtab.merge (Type.eq_type Vartab.empty) (typs1, typs2), |
|
52 |
Consttab.merge CodegenConsts.eq_const (consts1, consts2)); |
|
20456 | 53 |
|
23812 | 54 |
structure Translation = TheoryDataFun |
22846 | 55 |
( |
20931 | 56 |
type T = appgens * abstypes; |
57 |
val empty = (Symtab.empty, (Symtab.empty, Consttab.empty)); |
|
18217 | 58 |
val copy = I; |
59 |
val extend = I; |
|
20931 | 60 |
fun merge _ ((appgens1, abstypes1), (appgens2, abstypes2)) = |
61 |
(merge_appgens (appgens1, appgens2), merge_abstypes (abstypes1, abstypes2)); |
|
22846 | 62 |
); |
18217 | 63 |
|
23136 | 64 |
structure Funcgr = CodegenFuncgrRetrieval (fun rewrites thy = []); |
22213 | 65 |
|
22737 | 66 |
fun code_depgr thy [] = Funcgr.make thy [] |
67 |
| code_depgr thy consts = |
|
68 |
let |
|
69 |
val gr = Funcgr.make thy consts; |
|
70 |
val select = CodegenFuncgr.Constgraph.all_succs gr consts; |
|
71 |
in |
|
72 |
CodegenFuncgr.Constgraph.subgraph |
|
73 |
(member CodegenConsts.eq_const select) gr |
|
74 |
end; |
|
75 |
||
76 |
fun code_thms thy = |
|
77 |
Pretty.writeln o CodegenFuncgr.pretty thy o code_depgr thy; |
|
22213 | 78 |
|
22507 | 79 |
fun code_deps thy consts = |
80 |
let |
|
22737 | 81 |
val gr = code_depgr thy consts; |
22507 | 82 |
fun mk_entry (const, (_, (_, parents))) = |
83 |
let |
|
84 |
val name = CodegenConsts.string_of_const thy const; |
|
85 |
val nameparents = map (CodegenConsts.string_of_const thy) parents; |
|
86 |
in { name = name, ID = name, dir = "", unfold = true, |
|
87 |
path = "", parents = nameparents } |
|
88 |
end; |
|
89 |
val prgr = CodegenFuncgr.Constgraph.fold ((fn x => fn xs => xs @ [x]) o mk_entry) gr []; |
|
90 |
in Present.display_graph prgr end; |
|
91 |
||
22213 | 92 |
structure Code = CodeDataFun |
22846 | 93 |
( |
22213 | 94 |
type T = CodegenThingol.code; |
95 |
val empty = CodegenThingol.empty_code; |
|
96 |
fun merge _ = CodegenThingol.merge_code; |
|
97 |
fun purge _ NONE _ = CodegenThingol.empty_code |
|
98 |
| purge NONE _ _ = CodegenThingol.empty_code |
|
99 |
| purge (SOME thy) (SOME cs) code = |
|
100 |
let |
|
101 |
val cs_exisiting = |
|
102 |
map_filter (CodegenNames.const_rev thy) (Graph.keys code); |
|
103 |
val dels = (Graph.all_preds code |
|
104 |
o map (CodegenNames.const thy) |
|
105 |
o filter (member CodegenConsts.eq_const cs_exisiting) |
|
106 |
) cs; |
|
107 |
in Graph.del_nodes dels code end; |
|
22846 | 108 |
); |
22213 | 109 |
|
18708 | 110 |
|
23516 | 111 |
(* translation kernel *) |
18865 | 112 |
|
23812 | 113 |
fun get_abstype thy (tyco, tys) = case Symtab.lookup ((fst o snd o Translation.get) thy) tyco |
20931 | 114 |
of SOME ty => SOME ((map_atyps (fn TFree (n, _) => nth tys (the (Int.fromString n)))) ty) |
115 |
| NONE => NONE; |
|
116 |
||
22035 | 117 |
fun ensure_def thy = CodegenThingol.ensure_def (CodegenNames.labelled_name thy); |
118 |
||
23691 | 119 |
fun ensure_def_class thy (algbr as ((_, algebra), _)) funcgr class = |
20386 | 120 |
let |
22185 | 121 |
val superclasses = (Sorts.certify_sort algebra o Sorts.super_classes algebra) class; |
21895 | 122 |
val (v, cs) = AxClass.params_of_class thy class; |
22076
42ae57200d96
changed dictionary representation to explicit classrel witnesses
haftmann
parents:
22035
diff
changeset
|
123 |
val class' = CodegenNames.class thy class; |
42ae57200d96
changed dictionary representation to explicit classrel witnesses
haftmann
parents:
22035
diff
changeset
|
124 |
val classrels' = map (curry (CodegenNames.classrel thy) class) superclasses; |
22554
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
125 |
val classops' = map (CodegenNames.const thy o CodegenConsts.const_of_cexpr thy) cs; |
22197 | 126 |
val defgen_class = |
23691 | 127 |
fold_map (ensure_def_class thy algbr funcgr) superclasses |
128 |
##>> (fold_map (exprgen_type thy algbr funcgr) o map snd) cs |
|
22197 | 129 |
#-> (fn (superclasses, classoptyps) => succeed |
22076
42ae57200d96
changed dictionary representation to explicit classrel witnesses
haftmann
parents:
22035
diff
changeset
|
130 |
(CodegenThingol.Class (superclasses ~~ classrels', (unprefix "'" v, classops' ~~ classoptyps)))) |
18454 | 131 |
in |
22197 | 132 |
tracing (fn _ => "generating class " ^ quote class) |
23691 | 133 |
#> ensure_def thy defgen_class ("generating class " ^ quote class) class' |
22197 | 134 |
#> pair class' |
18865 | 135 |
end |
23691 | 136 |
and ensure_def_classrel thy algbr funcgr (subclass, superclass) = |
137 |
ensure_def_class thy algbr funcgr subclass |
|
22197 | 138 |
#>> (fn _ => CodegenNames.classrel thy (subclass, superclass)) |
23691 | 139 |
and ensure_def_tyco thy algbr funcgr "fun" trns = |
22354 | 140 |
trns |
141 |
|> pair "fun" |
|
23691 | 142 |
| ensure_def_tyco thy algbr funcgr tyco trns = |
22354 | 143 |
let |
144 |
fun defgen_datatype trns = |
|
22396 | 145 |
let |
22423 | 146 |
val (vs, cos) = CodegenData.get_datatype thy tyco; |
22396 | 147 |
in |
148 |
trns |
|
23691 | 149 |
|> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs |
22396 | 150 |
||>> fold_map (fn (c, tys) => |
23691 | 151 |
fold_map (exprgen_type thy algbr funcgr) tys |
22396 | 152 |
#-> (fn tys' => |
22554
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
153 |
pair ((CodegenNames.const thy o CodegenConsts.const_of_cexpr thy) |
22396 | 154 |
(c, tys ---> Type (tyco, map TFree vs)), tys'))) cos |
155 |
|-> (fn (vs, cos) => succeed (CodegenThingol.Datatype (vs, cos))) |
|
156 |
end; |
|
22354 | 157 |
val tyco' = CodegenNames.tyco thy tyco; |
158 |
in |
|
159 |
trns |
|
160 |
|> tracing (fn _ => "generating type constructor " ^ quote tyco) |
|
23691 | 161 |
|> ensure_def thy defgen_datatype ("generating type constructor " ^ quote tyco) tyco' |
22354 | 162 |
|> pair tyco' |
163 |
end |
|
23691 | 164 |
and exprgen_tyvar_sort thy (algbr as ((proj_sort, _), _)) funcgr (v, sort) trns = |
18516 | 165 |
trns |
23691 | 166 |
|> fold_map (ensure_def_class thy algbr funcgr) (proj_sort sort) |
22185 | 167 |
|>> (fn sort => (unprefix "'" v, sort)) |
23691 | 168 |
and exprgen_type thy algbr funcgr (TFree vs) trns = |
18516 | 169 |
trns |
23691 | 170 |
|> exprgen_tyvar_sort thy algbr funcgr vs |
22185 | 171 |
|>> (fn (v, sort) => ITyVar v) |
23691 | 172 |
| exprgen_type thy algbr funcgr (Type (tyco, tys)) trns = |
20931 | 173 |
case get_abstype thy (tyco, tys) |
174 |
of SOME ty => |
|
175 |
trns |
|
23691 | 176 |
|> exprgen_type thy algbr funcgr ty |
20931 | 177 |
| NONE => |
178 |
trns |
|
23691 | 179 |
|> ensure_def_tyco thy algbr funcgr tyco |
180 |
||>> fold_map (exprgen_type thy algbr funcgr) tys |
|
22185 | 181 |
|>> (fn (tyco, tys) => tyco `%% tys); |
18516 | 182 |
|
20835 | 183 |
exception CONSTRAIN of (string * typ) * typ; |
21012 | 184 |
val timing = ref false; |
20600 | 185 |
|
23691 | 186 |
fun exprgen_dicts thy (algbr as ((proj_sort, algebra), consts)) funcgr (ty_ctxt, sort_decl) = |
20456 | 187 |
let |
188 |
val pp = Sign.pp thy; |
|
22076
42ae57200d96
changed dictionary representation to explicit classrel witnesses
haftmann
parents:
22035
diff
changeset
|
189 |
datatype typarg = |
42ae57200d96
changed dictionary representation to explicit classrel witnesses
haftmann
parents:
22035
diff
changeset
|
190 |
Global of (class * string) * typarg list list |
42ae57200d96
changed dictionary representation to explicit classrel witnesses
haftmann
parents:
22035
diff
changeset
|
191 |
| Local of (class * class) list * (string * (int * sort)); |
22570
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
wenzelm
parents:
22554
diff
changeset
|
192 |
fun class_relation (Global ((_, tyco), yss), _) class = |
22076
42ae57200d96
changed dictionary representation to explicit classrel witnesses
haftmann
parents:
22035
diff
changeset
|
193 |
Global ((class, tyco), yss) |
22570
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
wenzelm
parents:
22554
diff
changeset
|
194 |
| class_relation (Local (classrels, v), subclass) superclass = |
22076
42ae57200d96
changed dictionary representation to explicit classrel witnesses
haftmann
parents:
22035
diff
changeset
|
195 |
Local ((subclass, superclass) :: classrels, v); |
22570
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
wenzelm
parents:
22554
diff
changeset
|
196 |
fun type_constructor tyco yss class = |
22076
42ae57200d96
changed dictionary representation to explicit classrel witnesses
haftmann
parents:
22035
diff
changeset
|
197 |
Global ((class, tyco), (map o map) fst yss); |
22570
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
wenzelm
parents:
22554
diff
changeset
|
198 |
fun type_variable (TFree (v, sort)) = |
20456 | 199 |
let |
200 |
val sort' = proj_sort sort; |
|
22076
42ae57200d96
changed dictionary representation to explicit classrel witnesses
haftmann
parents:
22035
diff
changeset
|
201 |
in map_index (fn (n, class) => (Local ([], (v, (n, sort'))), class)) sort' end; |
42ae57200d96
changed dictionary representation to explicit classrel witnesses
haftmann
parents:
22035
diff
changeset
|
202 |
val typargs = Sorts.of_sort_derivation pp algebra |
22570
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
wenzelm
parents:
22554
diff
changeset
|
203 |
{class_relation = class_relation, type_constructor = type_constructor, |
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
wenzelm
parents:
22554
diff
changeset
|
204 |
type_variable = type_variable} |
20456 | 205 |
(ty_ctxt, proj_sort sort_decl); |
22197 | 206 |
fun mk_dict (Global (inst, yss)) = |
23691 | 207 |
ensure_def_inst thy algbr funcgr inst |
22197 | 208 |
##>> (fold_map o fold_map) mk_dict yss |
209 |
#>> (fn (inst, dss) => DictConst (inst, dss)) |
|
210 |
| mk_dict (Local (classrels, (v, (k, sort)))) = |
|
23691 | 211 |
fold_map (ensure_def_classrel thy algbr funcgr) classrels |
22197 | 212 |
#>> (fn classrels => DictVar (classrels, (unprefix "'" v, (k, length sort)))) |
20456 | 213 |
in |
22197 | 214 |
fold_map mk_dict typargs |
20456 | 215 |
end |
23691 | 216 |
and exprgen_dict_parms thy (algbr as (_, consts)) funcgr (c, ty_ctxt) = |
20600 | 217 |
let |
22554
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
218 |
val c' = CodegenConsts.const_of_cexpr thy (c, ty_ctxt) |
20699 | 219 |
val idf = CodegenNames.const thy c'; |
21722 | 220 |
val ty_decl = Consts.the_declaration consts idf; |
22197 | 221 |
val (tys, tys_decl) = pairself (curry (Consts.typargs consts) idf) (ty_ctxt, ty_decl); |
222 |
val sorts = map (snd o dest_TVar) tys_decl; |
|
20456 | 223 |
in |
23691 | 224 |
fold_map (exprgen_dicts thy algbr funcgr) (tys ~~ sorts) |
20456 | 225 |
end |
23691 | 226 |
and ensure_def_inst thy (algbr as ((_, algebra), _)) funcgr (class, tyco) trns = |
20456 | 227 |
let |
22185 | 228 |
val superclasses = (Sorts.certify_sort algebra o Sorts.super_classes algebra) class; |
229 |
val (var, classops) = try (AxClass.params_of_class thy) class |> the_default ("'a", []) |
|
230 |
val vs = Name.names (Name.declare var Name.context) "'a" (Sorts.mg_domain algebra tyco [class]); |
|
231 |
val arity_typ = Type (tyco, map TFree vs); |
|
20896 | 232 |
fun gen_superarity superclass trns = |
233 |
trns |
|
23691 | 234 |
|> ensure_def_class thy algbr funcgr superclass |
235 |
||>> ensure_def_classrel thy algbr funcgr (class, superclass) |
|
236 |
||>> exprgen_dicts thy algbr funcgr (arity_typ, [superclass]) |
|
22185 | 237 |
|>> (fn ((superclass, classrel), [DictConst (inst, dss)]) => |
238 |
(superclass, (classrel, (inst, dss)))); |
|
239 |
fun gen_classop_def (classop as (c, ty)) trns = |
|
20896 | 240 |
trns |
23691 | 241 |
|> ensure_def_const thy algbr funcgr (CodegenConsts.const_of_cexpr thy classop) |
242 |
||>> exprgen_term thy algbr funcgr (Const (c, map_type_tfree (K arity_typ) ty)); |
|
20896 | 243 |
fun defgen_inst trns = |
244 |
trns |
|
23691 | 245 |
|> ensure_def_class thy algbr funcgr class |
246 |
||>> ensure_def_tyco thy algbr funcgr tyco |
|
247 |
||>> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs |
|
20896 | 248 |
||>> fold_map gen_superarity superclasses |
22185 | 249 |
||>> fold_map gen_classop_def classops |
20896 | 250 |
|-> (fn ((((class, tyco), arity), superarities), classops) => |
251 |
succeed (CodegenThingol.Classinst ((class, (tyco, arity)), (superarities, classops)))); |
|
252 |
val inst = CodegenNames.instance thy (class, tyco); |
|
18865 | 253 |
in |
254 |
trns |
|
20896 | 255 |
|> tracing (fn _ => "generating instance " ^ quote class ^ " / " ^ quote tyco) |
23691 | 256 |
|> ensure_def thy defgen_inst ("generating instance " ^ quote class ^ " / " ^ quote tyco) inst |
18865 | 257 |
|> pair inst |
258 |
end |
|
23691 | 259 |
and ensure_def_const thy (algbr as (_, consts)) funcgr (const as (c, opt_tyco)) trns = |
18865 | 260 |
let |
22554
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
261 |
val c' = CodegenNames.const thy const; |
20896 | 262 |
fun defgen_datatypecons trns = |
263 |
trns |
|
23691 | 264 |
|> ensure_def_tyco thy algbr funcgr |
22554
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
265 |
((the o CodegenData.get_datatype_of_constr thy) const) |
20896 | 266 |
|-> (fn _ => succeed CodegenThingol.Bot); |
267 |
fun defgen_classop trns = |
|
268 |
trns |
|
23691 | 269 |
|> ensure_def_class thy algbr funcgr ((the o AxClass.class_of_param thy) c) |
20896 | 270 |
|-> (fn _ => succeed CodegenThingol.Bot); |
271 |
fun defgen_fun trns = |
|
23691 | 272 |
let |
23812 | 273 |
val const' = perhaps (Consttab.lookup ((snd o snd o Translation.get) thy)) const; |
23691 | 274 |
val raw_thms = CodegenFuncgr.funcs funcgr const'; |
275 |
val ty = (Logic.unvarifyT o CodegenFuncgr.typ funcgr) const'; |
|
276 |
val thms = if (null o Term.typ_tfrees) ty orelse (null o fst o strip_type) ty |
|
277 |
then raw_thms |
|
278 |
else map (CodegenFunc.expand_eta 1) raw_thms; |
|
279 |
val timeap = if !timing then Output.timeap_msg ("time for " ^ c') |
|
280 |
else I; |
|
281 |
val msg = cat_lines ("generating code for theorems " :: map string_of_thm thms); |
|
282 |
val vs = (map dest_TFree o Consts.typargs consts) (c', ty); |
|
283 |
val dest_eqthm = |
|
284 |
apfst (snd o strip_comb) o Logic.dest_equals o Logic.unvarify o prop_of; |
|
285 |
fun exprgen_eq (args, rhs) trns = |
|
286 |
trns |
|
287 |
|> fold_map (exprgen_term thy algbr funcgr) args |
|
288 |
||>> exprgen_term thy algbr funcgr rhs; |
|
289 |
in |
|
290 |
trns |
|
291 |
|> CodegenThingol.message msg (fn trns => trns |
|
292 |
|> timeap (fold_map (exprgen_eq o dest_eqthm) thms) |
|
293 |
||>> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs |
|
294 |
||>> exprgen_type thy algbr funcgr ty |
|
295 |
|-> (fn ((eqs, vs), ty) => succeed (CodegenThingol.Fun (eqs, (vs, ty))))) |
|
296 |
end; |
|
22554
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
297 |
val defgen = if (is_some o CodegenData.get_datatype_of_constr thy) const |
21081 | 298 |
then defgen_datatypecons |
22554
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
299 |
else if is_some opt_tyco |
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
300 |
orelse (not o is_some o AxClass.class_of_param thy) c |
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
301 |
then defgen_fun |
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
302 |
else defgen_classop |
18865 | 303 |
in |
304 |
trns |
|
20835 | 305 |
|> tracing (fn _ => "generating constant " |
22554
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
306 |
^ (quote o CodegenConsts.string_of_const thy) const) |
23691 | 307 |
|> ensure_def thy defgen ("generating constant " ^ CodegenConsts.string_of_const thy const) c' |
20896 | 308 |
|> pair c' |
18865 | 309 |
end |
23691 | 310 |
and exprgen_term thy algbr funcgr (Const (c, ty)) trns = |
18517 | 311 |
trns |
23691 | 312 |
|> select_appgen thy algbr funcgr ((c, ty), []) |
313 |
| exprgen_term thy algbr funcgr (Free (v, ty)) trns = |
|
18516 | 314 |
trns |
23691 | 315 |
|> exprgen_type thy algbr funcgr ty |
22185 | 316 |
|>> (fn _ => IVar v) |
23691 | 317 |
| exprgen_term thy algbr funcgr (Abs (raw_v, ty, raw_t)) trns = |
19136 | 318 |
let |
22035 | 319 |
val (v, t) = Syntax.variant_abs (raw_v, ty, raw_t); |
19136 | 320 |
in |
321 |
trns |
|
23691 | 322 |
|> exprgen_type thy algbr funcgr ty |
323 |
||>> exprgen_term thy algbr funcgr t |
|
22185 | 324 |
|>> (fn (ty, t) => (v, ty) `|-> t) |
19136 | 325 |
end |
23691 | 326 |
| exprgen_term thy algbr funcgr (t as _ $ _) trns = |
20896 | 327 |
case strip_comb t |
328 |
of (Const (c, ty), ts) => |
|
18516 | 329 |
trns |
23691 | 330 |
|> select_appgen thy algbr funcgr ((c, ty), ts) |
20896 | 331 |
| (t', ts) => |
18516 | 332 |
trns |
23691 | 333 |
|> exprgen_term thy algbr funcgr t' |
334 |
||>> fold_map (exprgen_term thy algbr funcgr) ts |
|
22185 | 335 |
|>> (fn (t, ts) => t `$$ ts) |
23691 | 336 |
and appgen_default thy algbr funcgr ((c, ty), ts) trns = |
18865 | 337 |
trns |
23691 | 338 |
|> ensure_def_const thy algbr funcgr (CodegenConsts.const_of_cexpr thy (c, ty)) |
339 |
||>> fold_map (exprgen_type thy algbr funcgr) ((fst o Term.strip_type) ty) |
|
340 |
||>> exprgen_type thy algbr funcgr ((snd o Term.strip_type) ty) |
|
341 |
||>> exprgen_dict_parms thy algbr funcgr (c, ty) |
|
342 |
||>> fold_map (exprgen_term thy algbr funcgr) ts |
|
22305 | 343 |
|>> (fn ((((c, tys), ty), iss), ts) => IConst (c, (iss, tys)) `$$ ts) |
23691 | 344 |
and select_appgen thy algbr funcgr ((c, ty), ts) trns = |
23812 | 345 |
case Symtab.lookup (fst (Translation.get thy)) c |
20896 | 346 |
of SOME (i, (appgen, _)) => |
20105 | 347 |
if length ts < i then |
18865 | 348 |
let |
21161 | 349 |
val k = length ts; |
350 |
val tys = (curry Library.take (i - k) o curry Library.drop k o fst o strip_type) ty; |
|
351 |
val ctxt = (fold o fold_aterms) |
|
352 |
(fn Free (v, _) => Name.declare v | _ => I) ts Name.context; |
|
353 |
val vs = Name.names ctxt "a" tys; |
|
18865 | 354 |
in |
355 |
trns |
|
23691 | 356 |
|> fold_map (exprgen_type thy algbr funcgr) tys |
357 |
||>> appgen thy algbr funcgr ((c, ty), ts @ map Free vs) |
|
22185 | 358 |
|>> (fn (tys, t) => map2 (fn (v, _) => pair v) vs tys `|--> t) |
18865 | 359 |
end |
20105 | 360 |
else if length ts > i then |
18865 | 361 |
trns |
23691 | 362 |
|> appgen thy algbr funcgr ((c, ty), Library.take (i, ts)) |
363 |
||>> fold_map (exprgen_term thy algbr funcgr) (Library.drop (i, ts)) |
|
22185 | 364 |
|>> (fn (t, ts) => t `$$ ts) |
18865 | 365 |
else |
366 |
trns |
|
23691 | 367 |
|> appgen thy algbr funcgr ((c, ty), ts) |
18865 | 368 |
| NONE => |
369 |
trns |
|
23691 | 370 |
|> appgen_default thy algbr funcgr ((c, ty), ts); |
20600 | 371 |
|
372 |
||
23516 | 373 |
(* entrance points into translation kernel *) |
20600 | 374 |
|
23691 | 375 |
fun ensure_def_const' thy algbr funcgr c trns = |
376 |
ensure_def_const thy algbr funcgr c trns |
|
20835 | 377 |
handle CONSTRAIN ((c, ty), ty_decl) => error ( |
20600 | 378 |
"Constant " ^ c ^ " with most general type\n" |
22197 | 379 |
^ CodegenConsts.string_of_typ thy ty |
20600 | 380 |
^ "\noccurs with type\n" |
22197 | 381 |
^ CodegenConsts.string_of_typ thy ty_decl); |
22185 | 382 |
|
23691 | 383 |
fun perhaps_def_const thy algbr funcgr c trns = |
384 |
case try (ensure_def_const thy algbr funcgr c) trns |
|
22185 | 385 |
of SOME (c, trns) => (SOME c, trns) |
386 |
| NONE => (NONE, trns); |
|
387 |
||
23691 | 388 |
fun exprgen_term' thy algbr funcgr t trns = |
389 |
exprgen_term thy algbr funcgr t trns |
|
20835 | 390 |
handle CONSTRAIN ((c, ty), ty_decl) => error ("In term " ^ (quote o Sign.string_of_term thy) t |
20600 | 391 |
^ ",\nconstant " ^ c ^ " with most general type\n" |
22197 | 392 |
^ CodegenConsts.string_of_typ thy ty |
20600 | 393 |
^ "\noccurs with type\n" |
22197 | 394 |
^ CodegenConsts.string_of_typ thy ty_decl); |
18516 | 395 |
|
18702 | 396 |
|
20439 | 397 |
(* parametrized application generators, for instantiation in object logic *) |
23516 | 398 |
(* (axiomatic extensions of translation kernel) *) |
18217 | 399 |
|
23691 | 400 |
fun appgen_case dest_case_expr thy algbr funcgr (app as (c_ty, ts)) trns = |
20105 | 401 |
let |
402 |
val SOME ([], ((st, sty), ds)) = dest_case_expr thy (list_comb (Const c_ty, ts)); |
|
22035 | 403 |
fun clausegen (dt, bt) trns = |
404 |
trns |
|
23691 | 405 |
|> exprgen_term thy algbr funcgr dt |
406 |
||>> exprgen_term thy algbr funcgr bt; |
|
20105 | 407 |
in |
408 |
trns |
|
23691 | 409 |
|> exprgen_term thy algbr funcgr st |
410 |
||>> exprgen_type thy algbr funcgr sty |
|
20105 | 411 |
||>> fold_map clausegen ds |
22185 | 412 |
|>> (fn ((se, sty), ds) => ICase ((se, sty), ds)) |
20105 | 413 |
end; |
414 |
||
23691 | 415 |
fun appgen_let thy algbr funcgr (app as (_, [st, ct])) trns = |
20105 | 416 |
trns |
23691 | 417 |
|> exprgen_term thy algbr funcgr ct |
418 |
||>> exprgen_term thy algbr funcgr st |
|
23516 | 419 |
|-> (fn ((v, ty) `|-> be, se) => pair (CodegenThingol.collapse_let (((v, ty), se), be)) |
23691 | 420 |
| _ => appgen_default thy algbr funcgr app); |
20105 | 421 |
|
20439 | 422 |
fun add_appconst (c, appgen) thy = |
423 |
let |
|
20931 | 424 |
val i = (length o fst o strip_type o Sign.the_const_type thy) c; |
425 |
val _ = Code.change thy (K CodegenThingol.empty_code); |
|
426 |
in |
|
23812 | 427 |
(Translation.map o apfst) |
20931 | 428 |
(Symtab.update (c, (i, (appgen, stamp ())))) thy |
429 |
end; |
|
430 |
||
431 |
||
432 |
||
433 |
(** abstype and constsubst interface **) |
|
434 |
||
21916 | 435 |
local |
436 |
||
22554
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
437 |
fun add_consts thy f (c1, c2 as (c, opt_tyco)) = |
21916 | 438 |
let |
22554
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
439 |
val _ = if |
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
440 |
is_some (AxClass.class_of_param thy c) andalso is_none opt_tyco |
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
441 |
orelse is_some (CodegenData.get_datatype_of_constr thy c2) |
22197 | 442 |
then error ("Not a function: " ^ CodegenConsts.string_of_const thy c2) |
21916 | 443 |
else (); |
22213 | 444 |
val funcgr = Funcgr.make thy [c1, c2]; |
21916 | 445 |
val ty1 = (f o CodegenFuncgr.typ funcgr) c1; |
446 |
val ty2 = CodegenFuncgr.typ funcgr c2; |
|
447 |
val _ = if Sign.typ_equiv thy (ty1, ty2) then () else |
|
448 |
error ("Incompatiable type signatures of " ^ CodegenConsts.string_of_const thy c1 |
|
449 |
^ " and " ^ CodegenConsts.string_of_const thy c2 ^ ":\n" |
|
22197 | 450 |
^ CodegenConsts.string_of_typ thy ty1 ^ "\n" ^ CodegenConsts.string_of_typ thy ty2); |
21916 | 451 |
in Consttab.update (c1, c2) end; |
452 |
||
20931 | 453 |
fun gen_abstyp prep_const prep_typ (raw_abstyp, raw_substtyp) raw_absconsts thy = |
454 |
let |
|
455 |
val abstyp = Type.no_tvars (prep_typ thy raw_abstyp); |
|
456 |
val substtyp = Type.no_tvars (prep_typ thy raw_substtyp); |
|
457 |
val absconsts = (map o pairself) (prep_const thy) raw_absconsts; |
|
22197 | 458 |
val Type (abstyco, tys) = abstyp handle BIND => error ("Bad type: " ^ Sign.string_of_typ thy abstyp); |
459 |
val typarms = map (fst o dest_TFree) tys handle MATCH => error ("Bad type: " ^ Sign.string_of_typ thy abstyp); |
|
20931 | 460 |
fun mk_index v = |
461 |
let |
|
462 |
val k = find_index (fn w => v = w) typarms; |
|
463 |
in if k = ~1 |
|
22197 | 464 |
then error ("Free type variable: " ^ quote v) |
20931 | 465 |
else TFree (string_of_int k, []) |
466 |
end; |
|
467 |
val typpat = map_atyps (fn TFree (v, _) => mk_index v) substtyp; |
|
468 |
fun apply_typpat (Type (tyco, tys)) = |
|
469 |
let |
|
470 |
val tys' = map apply_typpat tys; |
|
471 |
in if tyco = abstyco then |
|
472 |
(map_atyps (fn TFree (n, _) => nth tys' (the (Int.fromString n)))) typpat |
|
473 |
else |
|
474 |
Type (tyco, tys') |
|
475 |
end |
|
476 |
| apply_typpat ty = ty; |
|
477 |
val _ = Code.change thy (K CodegenThingol.empty_code); |
|
478 |
in |
|
479 |
thy |
|
23812 | 480 |
|> (Translation.map o apsnd) (fn (abstypes, abscs) => |
20931 | 481 |
(abstypes |
482 |
|> Symtab.update (abstyco, typpat), |
|
483 |
abscs |
|
21916 | 484 |
|> fold (add_consts thy apply_typpat) absconsts) |
20931 | 485 |
) |
486 |
end; |
|
487 |
||
488 |
fun gen_constsubst prep_const raw_constsubsts thy = |
|
489 |
let |
|
490 |
val constsubsts = (map o pairself) (prep_const thy) raw_constsubsts; |
|
491 |
val _ = Code.change thy (K CodegenThingol.empty_code); |
|
492 |
in |
|
493 |
thy |
|
23812 | 494 |
|> (Translation.map o apsnd o apsnd) (fold (add_consts thy I) constsubsts) |
20931 | 495 |
end; |
496 |
||
21916 | 497 |
in |
498 |
||
22554
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
499 |
val abstyp = gen_abstyp (K I) Sign.certify_typ; |
22687
53943f4dab21
cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
wenzelm
parents:
22570
diff
changeset
|
500 |
val abstyp_e = gen_abstyp CodegenConsts.read_const Sign.read_typ; |
20931 | 501 |
|
22554
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22507
diff
changeset
|
502 |
val constsubst = gen_constsubst (K I); |
20931 | 503 |
val constsubst_e = gen_constsubst CodegenConsts.read_const; |
20439 | 504 |
|
21916 | 505 |
end; (*local*) |
18217 | 506 |
|
18516 | 507 |
|
20439 | 508 |
(** code generation interfaces **) |
18516 | 509 |
|
21881 | 510 |
(* generic generation combinators *) |
511 |
||
23691 | 512 |
fun generate thy funcgr gen it = |
20466 | 513 |
let |
23812 | 514 |
val cs = map_filter (Consttab.lookup ((snd o snd o Translation.get) thy)) |
21121 | 515 |
(CodegenFuncgr.all funcgr); |
22213 | 516 |
val funcgr' = Funcgr.make thy cs; |
23089 | 517 |
val naming = NameSpace.qualified_names NameSpace.default_naming; |
20466 | 518 |
val consttab = Consts.empty |
23089 | 519 |
|> fold (fn c => Consts.declare naming |
21121 | 520 |
((CodegenNames.const thy c, CodegenFuncgr.typ funcgr' c), true)) |
521 |
(CodegenFuncgr.all funcgr'); |
|
22185 | 522 |
val algbr = (CodegenData.operational_algebra thy, consttab); |
20466 | 523 |
in |
23691 | 524 |
Code.change_yield thy |
525 |
(CodegenThingol.start_transact (gen thy algbr funcgr' it)) |
|
21121 | 526 |
|> fst |
20466 | 527 |
end; |
18516 | 528 |
|
20600 | 529 |
fun codegen_term thy t = |
20353 | 530 |
let |
20600 | 531 |
val ct = Thm.cterm_of thy t; |
22213 | 532 |
val (ct', funcgr) = Funcgr.make_term thy (K (K K)) ct; |
20835 | 533 |
val t' = Thm.term_of ct'; |
23691 | 534 |
in generate thy funcgr exprgen_term' t' end; |
19136 | 535 |
|
23955 | 536 |
fun compile_term thy t = |
537 |
let |
|
538 |
val ct = Thm.cterm_of thy t; |
|
539 |
val (ct', funcgr) = Funcgr.make_term thy (K (K K)) ct; |
|
540 |
val t' = Thm.term_of ct'; |
|
541 |
val t'' = generate thy funcgr exprgen_term' t'; |
|
542 |
val consts = CodegenThingol.fold_constnames (insert (op =)) t'' []; |
|
543 |
val code = Code.get thy |
|
544 |
|> CodegenThingol.project_code true [] (SOME consts) |
|
545 |
in (code, t'') end; |
|
546 |
||
22464 | 547 |
fun raw_eval_term thy (ref_spec, t) args = |
20213 | 548 |
let |
22507 | 549 |
val _ = (Term.map_types o Term.map_atyps) (fn _ => |
550 |
error ("Term " ^ Sign.string_of_term thy t ^ " contains polymorphic type")) |
|
21388 | 551 |
t; |
21121 | 552 |
val t' = codegen_term thy t; |
22464 | 553 |
in |
554 |
CodegenSerializer.eval_term thy CodegenNames.labelled_name |
|
555 |
(Code.get thy) (ref_spec, t') args |
|
556 |
end; |
|
557 |
||
558 |
val satisfies_ref : bool option ref = ref NONE; |
|
559 |
||
560 |
fun eval_term thy t = raw_eval_term thy t []; |
|
561 |
fun satisfies thy t witnesses = raw_eval_term thy |
|
562 |
(("CodegenPackage.satisfies_ref", satisfies_ref), t) witnesses; |
|
18217 | 563 |
|
23691 | 564 |
fun filter_generatable thy consts = |
21881 | 565 |
let |
22213 | 566 |
val (consts', funcgr) = Funcgr.make_consts thy consts; |
23691 | 567 |
val consts'' = generate thy funcgr (fold_map ooo perhaps_def_const) consts'; |
22185 | 568 |
val consts''' = map_filter (fn (const, SOME _) => SOME const | (_, NONE) => NONE) |
569 |
(consts' ~~ consts''); |
|
570 |
in consts''' end; |
|
21881 | 571 |
|
18516 | 572 |
|
20439 | 573 |
(** toplevel interface and setup **) |
18756 | 574 |
|
575 |
local |
|
19150 | 576 |
|
20699 | 577 |
structure P = OuterParse |
578 |
and K = OuterKeyword |
|
579 |
||
20439 | 580 |
fun code raw_cs seris thy = |
18217 | 581 |
let |
23691 | 582 |
val (perm1, cs) = CodegenConsts.read_const_exprs thy |
583 |
(filter_generatable thy) raw_cs; |
|
584 |
val (perm2, cs') = case generate thy (Funcgr.make thy cs) (fold_map ooo ensure_def_const') cs |
|
585 |
of [] => (true, NONE) |
|
586 |
| cs => (false, SOME cs); |
|
20699 | 587 |
val code = Code.get thy; |
23691 | 588 |
val seris' = map (fn (((target, module), file), args) => |
589 |
CodegenSerializer.get_serializer thy target (perm1 orelse perm2) module file args |
|
590 |
CodegenNames.labelled_name cs') seris; |
|
18217 | 591 |
in |
23691 | 592 |
(map (fn f => f code) seris' : unit list; ()) |
18217 | 593 |
end; |
594 |
||
22737 | 595 |
fun code_thms_cmd thy = |
23691 | 596 |
code_thms thy o snd o CodegenConsts.read_const_exprs thy (fst o Funcgr.make_consts thy); |
20699 | 597 |
|
22507 | 598 |
fun code_deps_cmd thy = |
23691 | 599 |
code_deps thy o snd o CodegenConsts.read_const_exprs thy (fst o Funcgr.make_consts thy); |
600 |
||
601 |
val (inK, toK, fileK) = ("in", "to", "file"); |
|
18217 | 602 |
|
22845 | 603 |
val code_exprP = |
604 |
(Scan.repeat P.term |
|
23691 | 605 |
-- Scan.repeat (P.$$$ inK |-- P.name |
606 |
-- Scan.option (P.$$$ toK |-- P.name) |
|
607 |
-- Scan.option (P.$$$ fileK |-- P.name) |
|
22845 | 608 |
-- Scan.optional (P.$$$ "(" |-- P.arguments --| P.$$$ ")") [] |
609 |
) >> (fn (raw_cs, seris) => code raw_cs seris)); |
|
20439 | 610 |
|
23691 | 611 |
val _ = OuterSyntax.add_keywords [inK, toK, fileK]; |
612 |
||
22507 | 613 |
val (codeK, code_abstypeK, code_axiomsK, code_thmsK, code_depsK) = |
614 |
("code_gen", "code_abstype", "code_axioms", "code_thms", "code_deps"); |
|
22213 | 615 |
|
616 |
in |
|
617 |
||
21881 | 618 |
val codeP = |
22845 | 619 |
OuterSyntax.improper_command codeK "generate executable code for constants" |
22213 | 620 |
K.diag (P.!!! code_exprP >> (fn f => Toplevel.keep (f o Toplevel.theory_of))); |
21881 | 621 |
|
622 |
fun codegen_command thy cmd = |
|
22213 | 623 |
case Scan.read OuterLex.stopper (P.!!! code_exprP) ((filter OuterLex.is_proper o OuterSyntax.scan) cmd) |
21916 | 624 |
of SOME f => (writeln "Now generating code..."; f thy) |
22197 | 625 |
| NONE => error ("Bad directive " ^ quote cmd); |
21881 | 626 |
|
20931 | 627 |
val code_abstypeP = |
628 |
OuterSyntax.command code_abstypeK "axiomatic abstypes for code generation" K.thy_decl ( |
|
629 |
(P.typ -- P.typ -- Scan.optional (P.$$$ "where" |-- Scan.repeat1 |
|
630 |
(P.term --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.term)) []) |
|
631 |
>> (Toplevel.theory o uncurry abstyp_e) |
|
20428 | 632 |
); |
633 |
||
21062 | 634 |
val code_axiomsP = |
635 |
OuterSyntax.command code_axiomsK "axiomatic constant equalities for code generation" K.thy_decl ( |
|
20931 | 636 |
Scan.repeat1 (P.term --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.term) |
637 |
>> (Toplevel.theory o constsubst_e) |
|
18217 | 638 |
); |
639 |
||
22305 | 640 |
val code_thmsP = |
22737 | 641 |
OuterSyntax.improper_command code_thmsK "print system of defining equations for code" OuterKeyword.diag |
22305 | 642 |
(Scan.repeat P.term |
643 |
>> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory |
|
22737 | 644 |
o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of))); |
22213 | 645 |
|
22507 | 646 |
val code_depsP = |
22737 | 647 |
OuterSyntax.improper_command code_depsK "visualize dependencies of defining equations for code" OuterKeyword.diag |
22507 | 648 |
(Scan.repeat P.term |
649 |
>> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory |
|
650 |
o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of))); |
|
651 |
||
652 |
val _ = OuterSyntax.add_parsers [codeP, code_abstypeP, code_axiomsP, code_thmsP, code_depsP]; |
|
18217 | 653 |
|
654 |
end; (* local *) |
|
655 |
||
656 |
end; (* struct *) |