author | haftmann |
Tue, 18 Sep 2007 07:36:10 +0200 | |
changeset 24619 | c2e6a0f8c30b |
parent 24585 | c359896d0f48 |
child 24621 | 97d403d9ab54 |
permissions | -rw-r--r-- |
24219 | 1 |
(* Title: Tools/code/code_package.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
5 |
Code generator translation kernel. Code generator Isar setup. |
|
6 |
*) |
|
7 |
||
8 |
signature CODE_PACKAGE = |
|
9 |
sig |
|
24436 | 10 |
(*interfaces*) |
24219 | 11 |
val eval_conv: theory |
24381 | 12 |
-> (CodeThingol.code -> CodeThingol.typscheme * CodeThingol.iterm |
13 |
-> string list -> cterm -> thm) |
|
24283 | 14 |
-> cterm -> thm; |
15 |
val eval_term: theory |
|
24381 | 16 |
-> (CodeThingol.code -> CodeThingol.typscheme * CodeThingol.iterm |
17 |
-> string list -> cterm -> 'a) |
|
24283 | 18 |
-> cterm -> 'a; |
24585 | 19 |
val satisfies_ref: (unit -> bool) option ref; |
24283 | 20 |
val satisfies: theory -> cterm -> string list -> bool; |
24219 | 21 |
val codegen_command: theory -> string -> unit; |
22 |
||
24436 | 23 |
(*axiomatic interfaces*) |
24219 | 24 |
type appgen; |
25 |
val add_appconst: string * appgen -> theory -> theory; |
|
26 |
val appgen_let: appgen; |
|
27 |
val appgen_if: appgen; |
|
28 |
val appgen_case: (theory -> term |
|
29 |
-> ((string * typ) list * ((term * typ) * (term * term) list)) option) |
|
30 |
-> appgen; |
|
31 |
end; |
|
32 |
||
33 |
structure CodePackage : CODE_PACKAGE = |
|
34 |
struct |
|
35 |
||
36 |
open BasicCodeThingol; |
|
37 |
||
38 |
(** code translation **) |
|
39 |
||
40 |
(* theory data *) |
|
41 |
||
42 |
type appgen = theory -> ((sort -> sort) * Sorts.algebra) * Consts.T |
|
43 |
-> CodeFuncgr.T |
|
44 |
-> (string * typ) * term list -> CodeThingol.transact -> iterm * CodeThingol.transact; |
|
45 |
||
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
46 |
structure Appgens = TheoryDataFun |
24219 | 47 |
( |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
48 |
type T = (int * (appgen * stamp)) Symtab.table; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
49 |
val empty = Symtab.empty; |
24219 | 50 |
val copy = I; |
51 |
val extend = I; |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
52 |
fun merge _ = Symtab.merge (fn ((bounds1, (_, stamp1)), (bounds2, (_, stamp2))) => |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
53 |
bounds1 = bounds2 andalso stamp1 = stamp2); |
24219 | 54 |
); |
55 |
||
24283 | 56 |
fun code_depgr thy [] = CodeFuncgr.make thy [] |
24219 | 57 |
| code_depgr thy consts = |
58 |
let |
|
24283 | 59 |
val gr = CodeFuncgr.make thy consts; |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
60 |
val select = Graph.all_succs gr consts; |
24219 | 61 |
in |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
62 |
gr |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
63 |
|> Graph.subgraph (member (op =) select) |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
64 |
|> Graph.map_nodes ((apsnd o map) (Conv.fconv_rule (Class.overload thy))) |
24219 | 65 |
end; |
66 |
||
67 |
fun code_thms thy = |
|
68 |
Pretty.writeln o CodeFuncgr.pretty thy o code_depgr thy; |
|
69 |
||
70 |
fun code_deps thy consts = |
|
71 |
let |
|
72 |
val gr = code_depgr thy consts; |
|
73 |
fun mk_entry (const, (_, (_, parents))) = |
|
74 |
let |
|
75 |
val name = CodeUnit.string_of_const thy const; |
|
76 |
val nameparents = map (CodeUnit.string_of_const thy) parents; |
|
77 |
in { name = name, ID = name, dir = "", unfold = true, |
|
78 |
path = "", parents = nameparents } |
|
79 |
end; |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
80 |
val prgr = Graph.fold ((fn x => fn xs => xs @ [x]) o mk_entry) gr []; |
24219 | 81 |
in Present.display_graph prgr end; |
82 |
||
83 |
structure Program = CodeDataFun |
|
84 |
( |
|
85 |
type T = CodeThingol.code; |
|
86 |
val empty = CodeThingol.empty_code; |
|
87 |
fun merge _ = CodeThingol.merge_code; |
|
88 |
fun purge _ NONE _ = CodeThingol.empty_code |
|
89 |
| purge NONE _ _ = CodeThingol.empty_code |
|
90 |
| purge (SOME thy) (SOME cs) code = |
|
91 |
let |
|
92 |
val cs_exisiting = |
|
93 |
map_filter (CodeName.const_rev thy) (Graph.keys code); |
|
94 |
val dels = (Graph.all_preds code |
|
95 |
o map (CodeName.const thy) |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
96 |
o filter (member (op =) cs_exisiting) |
24219 | 97 |
) cs; |
98 |
in Graph.del_nodes dels code end; |
|
99 |
); |
|
100 |
||
101 |
||
102 |
(* translation kernel *) |
|
103 |
||
24283 | 104 |
val value_name = "Isabelle_Eval.EVAL.EVAL"; |
105 |
||
106 |
fun ensure_def thy = CodeThingol.ensure_def |
|
107 |
(fn s => if s = value_name then "<term>" else CodeName.labelled_name thy s); |
|
24219 | 108 |
|
109 |
fun ensure_def_class thy (algbr as ((_, algebra), _)) funcgr class = |
|
110 |
let |
|
111 |
val superclasses = (Sorts.certify_sort algebra o Sorts.super_classes algebra) class; |
|
112 |
val (v, cs) = AxClass.params_of_class thy class; |
|
113 |
val class' = CodeName.class thy class; |
|
114 |
val classrels' = map (curry (CodeName.classrel thy) class) superclasses; |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
115 |
val classops' = map (CodeName.const thy o fst) cs; |
24219 | 116 |
val defgen_class = |
117 |
fold_map (ensure_def_class thy algbr funcgr) superclasses |
|
118 |
##>> (fold_map (exprgen_typ thy algbr funcgr) o map snd) cs |
|
24250 | 119 |
#>> (fn (superclasses, classoptyps) => |
120 |
CodeThingol.Class (superclasses ~~ classrels', (unprefix "'" v, classops' ~~ classoptyps))) |
|
24219 | 121 |
in |
122 |
ensure_def thy defgen_class ("generating class " ^ quote class) class' |
|
123 |
#> pair class' |
|
124 |
end |
|
125 |
and ensure_def_classrel thy algbr funcgr (subclass, superclass) = |
|
126 |
ensure_def_class thy algbr funcgr subclass |
|
127 |
#>> (fn _ => CodeName.classrel thy (subclass, superclass)) |
|
24250 | 128 |
and ensure_def_tyco thy algbr funcgr "fun" = |
129 |
pair "fun" |
|
130 |
| ensure_def_tyco thy algbr funcgr tyco = |
|
24219 | 131 |
let |
24250 | 132 |
val defgen_datatype = |
24219 | 133 |
let |
134 |
val (vs, cos) = Code.get_datatype thy tyco; |
|
135 |
in |
|
24250 | 136 |
fold_map (exprgen_tyvar_sort thy algbr funcgr) vs |
137 |
##>> fold_map (fn (c, tys) => |
|
24219 | 138 |
fold_map (exprgen_typ thy algbr funcgr) tys |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
139 |
#>> (fn tys' => (CodeName.const thy c, tys'))) cos |
24250 | 140 |
#>> (fn (vs, cos) => CodeThingol.Datatype (vs, cos)) |
24219 | 141 |
end; |
142 |
val tyco' = CodeName.tyco thy tyco; |
|
143 |
in |
|
24250 | 144 |
ensure_def thy defgen_datatype ("generating type constructor " ^ quote tyco) tyco' |
145 |
#> pair tyco' |
|
24219 | 146 |
end |
24585 | 147 |
and exprgen_tyvar_sort thy (algbr as ((proj_sort, _), _)) funcgr (v, sort) = |
148 |
fold_map (ensure_def_class thy algbr funcgr) (proj_sort sort) |
|
149 |
#>> (fn sort => (unprefix "'" v, sort)) |
|
150 |
and exprgen_typ thy algbr funcgr (TFree vs) = |
|
151 |
exprgen_tyvar_sort thy algbr funcgr vs |
|
152 |
#>> (fn (v, sort) => ITyVar v) |
|
153 |
| exprgen_typ thy algbr funcgr (Type (tyco, tys)) = |
|
154 |
ensure_def_tyco thy algbr funcgr tyco |
|
155 |
##>> fold_map (exprgen_typ thy algbr funcgr) tys |
|
156 |
#>> (fn (tyco, tys) => tyco `%% tys); |
|
24219 | 157 |
|
158 |
exception CONSTRAIN of (string * typ) * typ; |
|
159 |
||
160 |
fun exprgen_dicts thy (algbr as ((proj_sort, algebra), consts)) funcgr (ty_ctxt, sort_decl) = |
|
161 |
let |
|
162 |
val pp = Sign.pp thy; |
|
163 |
datatype typarg = |
|
164 |
Global of (class * string) * typarg list list |
|
165 |
| Local of (class * class) list * (string * (int * sort)); |
|
166 |
fun class_relation (Global ((_, tyco), yss), _) class = |
|
167 |
Global ((class, tyco), yss) |
|
168 |
| class_relation (Local (classrels, v), subclass) superclass = |
|
169 |
Local ((subclass, superclass) :: classrels, v); |
|
170 |
fun type_constructor tyco yss class = |
|
171 |
Global ((class, tyco), (map o map) fst yss); |
|
172 |
fun type_variable (TFree (v, sort)) = |
|
173 |
let |
|
174 |
val sort' = proj_sort sort; |
|
175 |
in map_index (fn (n, class) => (Local ([], (v, (n, sort'))), class)) sort' end; |
|
176 |
val typargs = Sorts.of_sort_derivation pp algebra |
|
177 |
{class_relation = class_relation, type_constructor = type_constructor, |
|
178 |
type_variable = type_variable} |
|
179 |
(ty_ctxt, proj_sort sort_decl); |
|
180 |
fun mk_dict (Global (inst, yss)) = |
|
181 |
ensure_def_inst thy algbr funcgr inst |
|
182 |
##>> (fold_map o fold_map) mk_dict yss |
|
183 |
#>> (fn (inst, dss) => DictConst (inst, dss)) |
|
184 |
| mk_dict (Local (classrels, (v, (k, sort)))) = |
|
185 |
fold_map (ensure_def_classrel thy algbr funcgr) classrels |
|
186 |
#>> (fn classrels => DictVar (classrels, (unprefix "'" v, (k, length sort)))) |
|
187 |
in |
|
188 |
fold_map mk_dict typargs |
|
189 |
end |
|
190 |
and exprgen_dict_parms thy (algbr as (_, consts)) funcgr (c, ty_ctxt) = |
|
191 |
let |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
192 |
val ty_decl = Consts.the_declaration consts c; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
193 |
val (tys, tys_decl) = pairself (curry (Consts.typargs consts) c) (ty_ctxt, ty_decl); |
24219 | 194 |
val sorts = map (snd o dest_TVar) tys_decl; |
195 |
in |
|
196 |
fold_map (exprgen_dicts thy algbr funcgr) (tys ~~ sorts) |
|
197 |
end |
|
24585 | 198 |
and exprgen_eq thy algbr funcgr thm = |
199 |
let |
|
200 |
val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals |
|
201 |
o Logic.unvarify o prop_of) thm; |
|
202 |
in |
|
203 |
fold_map (exprgen_term thy algbr funcgr) args |
|
204 |
##>> exprgen_term thy algbr funcgr rhs |
|
205 |
#>> rpair thm |
|
206 |
end |
|
207 |
and ensure_def_inst thy (algbr as ((_, algebra), _)) funcgr (class, tyco) = |
|
24219 | 208 |
let |
209 |
val superclasses = (Sorts.certify_sort algebra o Sorts.super_classes algebra) class; |
|
210 |
val (var, classops) = try (AxClass.params_of_class thy) class |> the_default ("'a", []) |
|
24585 | 211 |
val vs = Name.names Name.context "'a" (Sorts.mg_domain algebra tyco [class]); |
212 |
val sorts' = Sorts.mg_domain (Sign.classes_of thy) tyco [class]; |
|
213 |
val vs' = map2 (fn (v, sort1) => fn sort2 => (v, |
|
214 |
Sorts.inter_sort (Sign.classes_of thy) (sort1, sort2))) vs sorts'; |
|
24219 | 215 |
val arity_typ = Type (tyco, map TFree vs); |
24585 | 216 |
val arity_typ' = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) vs'); |
217 |
fun gen_superarity superclass = |
|
218 |
ensure_def_class thy algbr funcgr superclass |
|
219 |
##>> ensure_def_classrel thy algbr funcgr (class, superclass) |
|
220 |
##>> exprgen_dicts thy algbr funcgr (arity_typ, [superclass]) |
|
221 |
#>> (fn ((superclass, classrel), [DictConst (inst, dss)]) => |
|
24219 | 222 |
(superclass, (classrel, (inst, dss)))); |
24585 | 223 |
fun gen_classop_def (c, ty) = |
224 |
let |
|
225 |
val c_inst = Const (c, map_type_tfree (K arity_typ') ty); |
|
226 |
val thm = Class.unoverload thy (Thm.cterm_of thy c_inst); |
|
227 |
val c_ty = (apsnd Logic.unvarifyT o dest_Const o snd |
|
228 |
o Logic.dest_equals o Thm.prop_of) thm; |
|
229 |
in |
|
230 |
ensure_def_const thy algbr funcgr c |
|
231 |
##>> exprgen_const thy algbr funcgr c_ty |
|
232 |
#>> (fn (c, IConst c_inst) => ((c, c_inst), thm)) |
|
233 |
end; |
|
234 |
val defgen_inst = |
|
235 |
ensure_def_class thy algbr funcgr class |
|
236 |
##>> ensure_def_tyco thy algbr funcgr tyco |
|
237 |
##>> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs |
|
238 |
##>> fold_map gen_superarity superclasses |
|
239 |
##>> fold_map gen_classop_def classops |
|
240 |
#>> (fn ((((class, tyco), arity), superarities), classops) => |
|
24250 | 241 |
CodeThingol.Classinst ((class, (tyco, arity)), (superarities, classops))); |
24219 | 242 |
val inst = CodeName.instance thy (class, tyco); |
243 |
in |
|
24585 | 244 |
ensure_def thy defgen_inst ("generating instance " ^ quote class ^ " / " ^ quote tyco) inst |
245 |
#> pair inst |
|
24219 | 246 |
end |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
247 |
and ensure_def_const thy (algbr as (_, consts)) funcgr c = |
24219 | 248 |
let |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
249 |
val c' = CodeName.const thy c; |
24585 | 250 |
fun defgen_datatypecons tyco = |
251 |
ensure_def_tyco thy algbr funcgr tyco |
|
252 |
#>> K CodeThingol.Bot; |
|
253 |
fun defgen_classop class = |
|
254 |
ensure_def_class thy algbr funcgr class |
|
255 |
#>> K CodeThingol.Bot; |
|
24219 | 256 |
fun defgen_fun trns = |
257 |
let |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
258 |
val raw_thms = CodeFuncgr.funcs funcgr c; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
259 |
val ty = (Logic.unvarifyT o CodeFuncgr.typ funcgr) c; |
24585 | 260 |
val vs = (map dest_TFree o Consts.typargs consts) (c, ty); |
24219 | 261 |
val thms = if (null o Term.typ_tfrees) ty orelse (null o fst o strip_type) ty |
262 |
then raw_thms |
|
263 |
else map (CodeUnit.expand_eta 1) raw_thms; |
|
264 |
in |
|
265 |
trns |
|
24381 | 266 |
|> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs |
24219 | 267 |
||>> exprgen_typ thy algbr funcgr ty |
24585 | 268 |
||>> fold_map (exprgen_eq thy algbr funcgr) thms |
24381 | 269 |
|>> (fn ((vs, ty), eqs) => CodeThingol.Fun ((vs, ty), eqs)) |
24219 | 270 |
end; |
24585 | 271 |
val defgen = case Code.get_datatype_of_constr thy c |
272 |
of SOME tyco => defgen_datatypecons tyco |
|
273 |
| NONE => (case AxClass.class_of_param thy c |
|
274 |
of SOME class => defgen_classop class |
|
275 |
| NONE => defgen_fun) |
|
24219 | 276 |
in |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
277 |
ensure_def thy defgen ("generating constant " ^ CodeUnit.string_of_const thy c) c' |
24250 | 278 |
#> pair c' |
24219 | 279 |
end |
24585 | 280 |
and exprgen_term thy algbr funcgr (Const (c, ty)) = |
281 |
select_appgen thy algbr funcgr ((c, ty), []) |
|
282 |
| exprgen_term thy algbr funcgr (Free (v, _)) = |
|
283 |
pair (IVar v) |
|
284 |
| exprgen_term thy algbr funcgr (Abs (abs as (_, ty, _))) = |
|
24219 | 285 |
let |
24585 | 286 |
val (v, t) = Syntax.variant_abs abs; |
24219 | 287 |
in |
24585 | 288 |
exprgen_typ thy algbr funcgr ty |
289 |
##>> exprgen_term thy algbr funcgr t |
|
290 |
#>> (fn (ty, t) => (v, ty) `|-> t) |
|
24219 | 291 |
end |
24585 | 292 |
| exprgen_term thy algbr funcgr (t as _ $ _) = |
24219 | 293 |
case strip_comb t |
294 |
of (Const (c, ty), ts) => |
|
24585 | 295 |
select_appgen thy algbr funcgr ((c, ty), ts) |
24219 | 296 |
| (t', ts) => |
24585 | 297 |
exprgen_term thy algbr funcgr t' |
298 |
##>> fold_map (exprgen_term thy algbr funcgr) ts |
|
299 |
#>> (fn (t, ts) => t `$$ ts) |
|
300 |
and exprgen_const thy algbr funcgr (c, ty) = |
|
301 |
ensure_def_const thy algbr funcgr c |
|
302 |
##>> exprgen_dict_parms thy algbr funcgr (c, ty) |
|
303 |
##>> fold_map (exprgen_typ thy algbr funcgr) ((fst o Term.strip_type) ty) |
|
304 |
(*##>> exprgen_typ thy algbr funcgr ((snd o Term.strip_type) ty)*) |
|
305 |
#>> (fn ((c, iss), tys) => IConst (c, (iss, tys))) |
|
306 |
and exprgen_app_default thy algbr funcgr (c_ty, ts) = |
|
307 |
exprgen_const thy algbr funcgr c_ty |
|
308 |
##>> fold_map (exprgen_term thy algbr funcgr) ts |
|
309 |
#>> (fn (t, ts) => t `$$ ts) |
|
310 |
and select_appgen thy algbr funcgr ((c, ty), ts) = |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
311 |
case Symtab.lookup (Appgens.get thy) c |
24219 | 312 |
of SOME (i, (appgen, _)) => |
313 |
if length ts < i then |
|
314 |
let |
|
315 |
val k = length ts; |
|
316 |
val tys = (curry Library.take (i - k) o curry Library.drop k o fst o strip_type) ty; |
|
317 |
val ctxt = (fold o fold_aterms) |
|
318 |
(fn Free (v, _) => Name.declare v | _ => I) ts Name.context; |
|
319 |
val vs = Name.names ctxt "a" tys; |
|
320 |
in |
|
24585 | 321 |
fold_map (exprgen_typ thy algbr funcgr) tys |
322 |
##>> appgen thy algbr funcgr ((c, ty), ts @ map Free vs) |
|
323 |
#>> (fn (tys, t) => map2 (fn (v, _) => pair v) vs tys `|--> t) |
|
24219 | 324 |
end |
325 |
else if length ts > i then |
|
24585 | 326 |
appgen thy algbr funcgr ((c, ty), Library.take (i, ts)) |
327 |
##>> fold_map (exprgen_term thy algbr funcgr) (Library.drop (i, ts)) |
|
328 |
#>> (fn (t, ts) => t `$$ ts) |
|
24219 | 329 |
else |
24585 | 330 |
appgen thy algbr funcgr ((c, ty), ts) |
24219 | 331 |
| NONE => |
24585 | 332 |
exprgen_app_default thy algbr funcgr ((c, ty), ts); |
24219 | 333 |
|
334 |
||
335 |
(* entrance points into translation kernel *) |
|
336 |
||
337 |
fun ensure_def_const' thy algbr funcgr c trns = |
|
338 |
ensure_def_const thy algbr funcgr c trns |
|
339 |
handle CONSTRAIN ((c, ty), ty_decl) => error ( |
|
340 |
"Constant " ^ c ^ " with most general type\n" |
|
341 |
^ CodeUnit.string_of_typ thy ty |
|
342 |
^ "\noccurs with type\n" |
|
343 |
^ CodeUnit.string_of_typ thy ty_decl); |
|
344 |
||
345 |
fun perhaps_def_const thy algbr funcgr c trns = |
|
346 |
case try (ensure_def_const thy algbr funcgr c) trns |
|
347 |
of SOME (c, trns) => (SOME c, trns) |
|
348 |
| NONE => (NONE, trns); |
|
349 |
||
350 |
fun exprgen_term' thy algbr funcgr t trns = |
|
351 |
exprgen_term thy algbr funcgr t trns |
|
352 |
handle CONSTRAIN ((c, ty), ty_decl) => error ("In term " ^ (quote o Sign.string_of_term thy) t |
|
353 |
^ ",\nconstant " ^ c ^ " with most general type\n" |
|
354 |
^ CodeUnit.string_of_typ thy ty |
|
355 |
^ "\noccurs with type\n" |
|
356 |
^ CodeUnit.string_of_typ thy ty_decl); |
|
357 |
||
358 |
||
359 |
(* parametrized application generators, for instantiation in object logic *) |
|
360 |
(* (axiomatic extensions of translation kernel) *) |
|
361 |
||
362 |
fun appgen_case dest_case_expr thy algbr funcgr (app as (c_ty, ts)) = |
|
363 |
let |
|
364 |
val SOME ([], ((st, sty), ds)) = dest_case_expr thy (list_comb (Const c_ty, ts)); |
|
365 |
fun clause_gen (dt, bt) = |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
366 |
exprgen_term thy algbr funcgr |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
367 |
(map_aterms (fn Const (c_ty as (c, ty)) => Const |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
368 |
(Class.unoverload_const thy c_ty, ty) | t => t) dt) |
24219 | 369 |
##>> exprgen_term thy algbr funcgr bt; |
370 |
in |
|
371 |
exprgen_term thy algbr funcgr st |
|
372 |
##>> exprgen_typ thy algbr funcgr sty |
|
373 |
##>> fold_map clause_gen ds |
|
24585 | 374 |
##>> exprgen_app_default thy algbr funcgr app |
24219 | 375 |
#>> (fn (((se, sty), ds), t0) => ICase (((se, sty), ds), t0)) |
376 |
end; |
|
377 |
||
378 |
fun appgen_let thy algbr funcgr (app as (_, [st, ct])) = |
|
379 |
exprgen_term thy algbr funcgr ct |
|
380 |
##>> exprgen_term thy algbr funcgr st |
|
24585 | 381 |
##>> exprgen_app_default thy algbr funcgr app |
24219 | 382 |
#>> (fn (((v, ty) `|-> be, se), t0) => |
383 |
ICase (CodeThingol.collapse_let (((v, ty), se), be), t0) |
|
384 |
| (_, t0) => t0); |
|
385 |
||
386 |
fun appgen_if thy algbr funcgr (app as (_, [tb, tt, tf])) = |
|
387 |
exprgen_term thy algbr funcgr tb |
|
388 |
##>> exprgen_typ thy algbr funcgr (Type ("bool", [])) |
|
389 |
##>> exprgen_term thy algbr funcgr (Const ("True", Type ("bool", []))) |
|
390 |
##>> exprgen_term thy algbr funcgr tt |
|
391 |
##>> exprgen_term thy algbr funcgr (Const ("False", Type ("bool", []))) |
|
392 |
##>> exprgen_term thy algbr funcgr tf |
|
24585 | 393 |
##>> exprgen_app_default thy algbr funcgr app |
24219 | 394 |
#>> (fn ((((((tb, B), T), tt), F), tf), t0) => ICase (((tb, B), [(T, tt), (F, tf)]), t0)); |
395 |
||
396 |
fun add_appconst (c, appgen) thy = |
|
397 |
let |
|
398 |
val i = (length o fst o strip_type o Sign.the_const_type thy) c; |
|
399 |
val _ = Program.change thy (K CodeThingol.empty_code); |
|
400 |
in |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
401 |
Appgens.map (Symtab.update (c, (i, (appgen, stamp ())))) thy |
24219 | 402 |
end; |
403 |
||
404 |
||
405 |
(** code generation interfaces **) |
|
406 |
||
407 |
(* generic generation combinators *) |
|
408 |
||
409 |
fun generate thy funcgr gen it = |
|
410 |
let |
|
411 |
val naming = NameSpace.qualified_names NameSpace.default_naming; |
|
412 |
val consttab = Consts.empty |
|
413 |
|> fold (fn c => Consts.declare naming |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
414 |
((c, CodeFuncgr.typ funcgr c), true)) |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
415 |
(CodeFuncgr.all funcgr); |
24219 | 416 |
val algbr = (Code.operational_algebra thy, consttab); |
417 |
in |
|
418 |
Program.change_yield thy |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
419 |
(CodeThingol.start_transact (gen thy algbr funcgr it)) |
24283 | 420 |
|> fst |
24219 | 421 |
end; |
422 |
||
24436 | 423 |
fun code thy permissive cs seris = |
424 |
let |
|
425 |
val code = Program.get thy; |
|
426 |
val seris' = map (fn (((target, module), file), args) => |
|
427 |
CodeTarget.get_serializer thy target permissive module file args |
|
428 |
CodeName.labelled_name cs) seris; |
|
429 |
in (map (fn f => f code) seris' : unit list; ()) end; |
|
430 |
||
24283 | 431 |
fun raw_eval f thy g = |
24250 | 432 |
let |
433 |
val value_name = "Isabelle_Eval.EVAL.EVAL"; |
|
434 |
fun ensure_eval thy algbr funcgr t = |
|
435 |
let |
|
24381 | 436 |
val ty = fastype_of t; |
24436 | 437 |
val vs = fold_term_types (K (fold_atyps (insert (eq_fst op =) |
438 |
o dest_TFree))) t []; |
|
24250 | 439 |
val defgen_eval = |
24381 | 440 |
fold_map (exprgen_tyvar_sort thy algbr funcgr) vs |
441 |
##>> exprgen_typ thy algbr funcgr ty |
|
442 |
##>> exprgen_term' thy algbr funcgr t |
|
24585 | 443 |
#>> (fn ((vs, ty), t) => CodeThingol.Fun ((vs, ty), [(([], t), Drule.dummy_thm)])); |
24283 | 444 |
fun result (dep, code) = |
445 |
let |
|
24585 | 446 |
val CodeThingol.Fun ((vs, ty), [(([], t), _)]) = Graph.get_node code value_name; |
24283 | 447 |
val deps = Graph.imm_succs code value_name; |
448 |
val code' = Graph.del_nodes [value_name] code; |
|
449 |
val code'' = CodeThingol.project_code false [] (SOME deps) code'; |
|
24381 | 450 |
in ((code'', ((vs, ty), t), deps), (dep, code')) end; |
24250 | 451 |
in |
452 |
ensure_def thy defgen_eval "evaluation" value_name |
|
24283 | 453 |
#> result |
24250 | 454 |
end; |
24283 | 455 |
fun h funcgr ct = |
24219 | 456 |
let |
24381 | 457 |
val (code, vs_ty_t, deps) = generate thy funcgr ensure_eval (Thm.term_of ct); |
458 |
in g code vs_ty_t deps ct end; |
|
24283 | 459 |
in f thy h end; |
24219 | 460 |
|
24283 | 461 |
fun eval_conv thy = raw_eval CodeFuncgr.eval_conv thy; |
462 |
fun eval_term thy = raw_eval CodeFuncgr.eval_term thy; |
|
24219 | 463 |
|
24585 | 464 |
val satisfies_ref : (unit -> bool) option ref = ref NONE; |
24219 | 465 |
|
24283 | 466 |
fun satisfies thy ct witnesses = |
467 |
let |
|
24381 | 468 |
fun evl code ((vs, ty), t) deps ct = |
24283 | 469 |
let |
470 |
val t0 = Thm.term_of ct |
|
471 |
val _ = (Term.map_types o Term.map_atyps) (fn _ => |
|
472 |
error ("Term " ^ Sign.string_of_term thy t0 ^ " contains polymorphic type")) |
|
473 |
t0; |
|
474 |
in |
|
475 |
CodeTarget.eval_term thy ("CodePackage.satisfies_ref", satisfies_ref) |
|
476 |
code (t, ty) witnesses |
|
477 |
end; |
|
478 |
in eval_term thy evl ct end; |
|
24219 | 479 |
|
480 |
fun filter_generatable thy consts = |
|
481 |
let |
|
24283 | 482 |
val (consts', funcgr) = CodeFuncgr.make_consts thy consts; |
483 |
val consts'' = generate thy funcgr (fold_map ooo perhaps_def_const) consts'; |
|
24219 | 484 |
val consts''' = map_filter (fn (const, SOME _) => SOME const | (_, NONE) => NONE) |
485 |
(consts' ~~ consts''); |
|
486 |
in consts''' end; |
|
487 |
||
24436 | 488 |
fun generate_const_exprs thy raw_cs = |
489 |
let |
|
490 |
val (perm1, cs) = CodeUnit.read_const_exprs thy |
|
491 |
(filter_generatable thy) raw_cs; |
|
492 |
val (perm2, cs') = case generate thy (CodeFuncgr.make thy cs) |
|
493 |
(fold_map ooo ensure_def_const') cs |
|
494 |
of [] => (true, NONE) |
|
495 |
| cs => (false, SOME cs); |
|
496 |
in (perm1 orelse perm2, cs') end; |
|
497 |
||
498 |
||
499 |
(** code properties **) |
|
500 |
||
501 |
fun mk_codeprops thy all_cs sel_cs = |
|
502 |
let |
|
503 |
fun select (thmref, thm) = case try (Drule.unvarify o Drule.zero_var_indexes) thm |
|
504 |
of NONE => NONE |
|
505 |
| SOME thm => let |
|
506 |
val t = (ObjectLogic.drop_judgment thy o Thm.prop_of) thm; |
|
507 |
val cs = fold_aterms (fn Const (c, ty) => |
|
508 |
cons (Class.unoverload_const thy (c, ty)) | _ => I) t []; |
|
509 |
in if exists (member (op =) sel_cs) cs |
|
510 |
andalso forall (member (op =) all_cs) cs |
|
511 |
then SOME (thmref, thm) else NONE end; |
|
512 |
fun mk_codeprop (thmref, thm) = |
|
513 |
let |
|
514 |
val t = ObjectLogic.drop_judgment thy (Thm.prop_of thm); |
|
515 |
val ty_judg = fastype_of t; |
|
516 |
val tfrees1 = fold_aterms (fn Const (c, ty) => |
|
517 |
Term.add_tfreesT ty | _ => I) t []; |
|
518 |
val vars = Term.add_frees t []; |
|
519 |
val tfrees2 = fold (Term.add_tfreesT o snd) vars []; |
|
520 |
val tfrees' = subtract (op =) tfrees2 tfrees1 |> map TFree; |
|
521 |
val ty = map Term.itselfT tfrees' @ map snd vars ---> ty_judg; |
|
522 |
val tfree_vars = map Logic.mk_type tfrees'; |
|
523 |
val c = PureThy.string_of_thmref thmref |
|
524 |
|> NameSpace.explode |
|
525 |
|> (fn [x] => [x] | (x::xs) => xs) |
|
526 |
|> space_implode "_" |
|
527 |
val propdef = (((c, ty), tfree_vars @ map Free vars), t); |
|
528 |
in if c = "" then NONE else SOME (thmref, propdef) end; |
|
529 |
in |
|
530 |
PureThy.thms_containing thy ([], []) |
|
531 |
|> maps PureThy.selections |
|
532 |
|> map_filter select |
|
533 |
|> map_filter mk_codeprop |
|
534 |
end; |
|
535 |
||
536 |
fun add_codeprops all_cs sel_cs thy = |
|
537 |
let |
|
538 |
val codeprops = mk_codeprops thy all_cs sel_cs; |
|
539 |
fun lift_name_yield f x = (Name.context, x) |> f ||> snd; |
|
540 |
fun add (thmref, (((raw_c, ty), ts), t)) (names, thy) = |
|
541 |
let |
|
542 |
val _ = warning ("Adding theorem " ^ quote (PureThy.string_of_thmref thmref) |
|
543 |
^ " as code property " ^ quote raw_c); |
|
544 |
val ([raw_c'], names') = Name.variants [raw_c] names; |
|
545 |
in |
|
546 |
thy |
|
547 |
|> PureThy.simple_def ("", []) (((raw_c', ty, Syntax.NoSyn), ts), t) |
|
548 |
||> pair names' |
|
549 |
end; |
|
550 |
in |
|
551 |
thy |
|
552 |
|> Sign.sticky_prefix "codeprop" |
|
553 |
|> lift_name_yield (fold_map add codeprops) |
|
554 |
||> Sign.restore_naming thy |
|
555 |
|-> (fn c_thms => fold (Code.add_func true o snd) c_thms #> pair c_thms) |
|
556 |
end; |
|
557 |
||
558 |
||
24219 | 559 |
|
560 |
(** toplevel interface and setup **) |
|
561 |
||
562 |
local |
|
563 |
||
564 |
structure P = OuterParse |
|
565 |
and K = OuterKeyword |
|
566 |
||
24436 | 567 |
fun code_cmd raw_cs seris thy = |
24219 | 568 |
let |
24436 | 569 |
val (permissive, cs) = generate_const_exprs thy raw_cs; |
570 |
val _ = code thy permissive cs seris; |
|
571 |
in () end; |
|
24219 | 572 |
|
573 |
fun code_thms_cmd thy = |
|
24283 | 574 |
code_thms thy o snd o CodeUnit.read_const_exprs thy (fst o CodeFuncgr.make_consts thy); |
24219 | 575 |
|
576 |
fun code_deps_cmd thy = |
|
24283 | 577 |
code_deps thy o snd o CodeUnit.read_const_exprs thy (fst o CodeFuncgr.make_consts thy); |
24219 | 578 |
|
24436 | 579 |
fun code_props_cmd raw_cs seris thy = |
580 |
let |
|
581 |
val (_, all_cs) = generate_const_exprs thy ["*"]; |
|
582 |
val (permissive, cs) = generate_const_exprs thy raw_cs; |
|
583 |
val (c_thms, thy') = add_codeprops (map (the o CodeName.const_rev thy) (these all_cs)) |
|
584 |
(map (the o CodeName.const_rev thy) (these cs)) thy; |
|
585 |
val prop_cs = (filter_generatable thy' o map fst) c_thms; |
|
586 |
val _ = if null seris then [] else generate thy' (CodeFuncgr.make thy' prop_cs) |
|
587 |
(fold_map ooo ensure_def_const') prop_cs; |
|
588 |
val _ = if null seris then () else code thy' permissive |
|
589 |
(SOME (map (CodeName.const thy') prop_cs)) seris; |
|
590 |
in thy' end; |
|
591 |
||
24250 | 592 |
val (inK, module_nameK, fileK) = ("in", "module_name", "file"); |
24219 | 593 |
|
24436 | 594 |
fun code_exprP cmd = |
24219 | 595 |
(Scan.repeat P.term |
596 |
-- Scan.repeat (P.$$$ inK |-- P.name |
|
24250 | 597 |
-- Scan.option (P.$$$ module_nameK |-- P.name) |
24219 | 598 |
-- Scan.option (P.$$$ fileK |-- P.name) |
599 |
-- Scan.optional (P.$$$ "(" |-- P.arguments --| P.$$$ ")") [] |
|
24436 | 600 |
) >> (fn (raw_cs, seris) => cmd raw_cs seris)); |
24219 | 601 |
|
24250 | 602 |
val _ = OuterSyntax.add_keywords [inK, module_nameK, fileK]; |
24219 | 603 |
|
24436 | 604 |
val (codeK, code_thmsK, code_depsK, code_propsK) = |
605 |
("export_code", "code_thms", "code_deps", "code_props"); |
|
24219 | 606 |
|
607 |
in |
|
608 |
||
609 |
val codeP = |
|
610 |
OuterSyntax.improper_command codeK "generate executable code for constants" |
|
24436 | 611 |
K.diag (P.!!! (code_exprP code_cmd) >> (fn f => Toplevel.keep (f o Toplevel.theory_of))); |
24219 | 612 |
|
613 |
fun codegen_command thy cmd = |
|
24436 | 614 |
case Scan.read OuterLex.stopper (P.!!! (code_exprP code_cmd)) ((filter OuterLex.is_proper o OuterSyntax.scan) cmd) |
24219 | 615 |
of SOME f => (writeln "Now generating code..."; f thy) |
616 |
| NONE => error ("Bad directive " ^ quote cmd); |
|
617 |
||
618 |
val code_thmsP = |
|
619 |
OuterSyntax.improper_command code_thmsK "print system of defining equations for code" OuterKeyword.diag |
|
620 |
(Scan.repeat P.term |
|
621 |
>> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory |
|
622 |
o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of))); |
|
623 |
||
624 |
val code_depsP = |
|
625 |
OuterSyntax.improper_command code_depsK "visualize dependencies of defining equations for code" OuterKeyword.diag |
|
626 |
(Scan.repeat P.term |
|
627 |
>> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory |
|
628 |
o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of))); |
|
629 |
||
24436 | 630 |
val code_propsP = |
631 |
OuterSyntax.command code_propsK "generate characteristic properties for executable constants" |
|
632 |
K.thy_decl (P.!!! (code_exprP code_props_cmd) >> Toplevel.theory); |
|
24219 | 633 |
|
24436 | 634 |
val _ = OuterSyntax.add_parsers [codeP, code_thmsP, code_depsP, code_propsP]; |
24219 | 635 |
|
24436 | 636 |
end; (*local*) |
637 |
||
638 |
end; (*struct*) |