author | wenzelm |
Thu, 31 Aug 2006 23:01:16 +0200 | |
changeset 20452 | 6d8b29c7a960 |
parent 20404 | 1a29e6c3ab04 |
child 20456 | 42be3a46dcd8 |
permissions | -rw-r--r-- |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
1 |
(* Title: Pure/Tools/codegen_theorems.ML |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
2 |
ID: $Id$ |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
3 |
Author: Florian Haftmann, TU Muenchen |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
4 |
|
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
5 |
Theorems used for code generation. |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
6 |
*) |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
7 |
|
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
8 |
signature CODEGEN_THEOREMS = |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
9 |
sig |
19884 | 10 |
val add_notify: ((string * typ) list option -> theory -> theory) -> theory -> theory; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
11 |
val add_preproc: (theory -> thm list -> thm list) -> theory -> theory; |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
12 |
val add_fun_extr: (theory -> string * typ -> thm list) -> theory -> theory; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
13 |
val add_datatype_extr: (theory -> string |
19597 | 14 |
-> (((string * sort) list * (string * typ list) list) * tactic) option) |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
15 |
-> theory -> theory; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
16 |
val add_fun: thm -> theory -> theory; |
19597 | 17 |
val del_fun: thm -> theory -> theory; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
18 |
val add_unfold: thm -> theory -> theory; |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
19 |
val del_unfold: thm -> theory -> theory; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
20 |
val purge_defs: string * typ -> theory -> theory; |
19597 | 21 |
val notify_dirty: theory -> theory; |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
22 |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19897
diff
changeset
|
23 |
val extr_typ: theory -> thm -> typ; |
20404 | 24 |
val rewrite_fun: thm list -> thm -> thm; |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
25 |
val common_typ: theory -> (thm -> typ) -> thm list -> thm list; |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19897
diff
changeset
|
26 |
val preprocess: theory -> thm list -> thm list; |
19597 | 27 |
|
20394 | 28 |
val prove_freeness: theory -> tactic -> string |
29 |
-> (string * sort) list * (string * typ list) list -> thm list; |
|
30 |
||
19597 | 31 |
type thmtab; |
20353 | 32 |
val mk_thmtab: theory -> (string * typ) list -> thmtab; |
20191 | 33 |
val get_sortalgebra: thmtab -> Sorts.algebra; |
34 |
val get_dtyp_of_cons: thmtab -> string * typ -> string option; |
|
35 |
val get_dtyp_spec: thmtab -> string |
|
36 |
-> ((string * sort) list * (string * typ list) list) option; |
|
37 |
val get_fun_thms: thmtab -> string * typ -> thm list; |
|
38 |
||
20394 | 39 |
val pretty_funtab: theory -> thm list CodegenConsts.Consttab.table -> Pretty.T; |
19597 | 40 |
val print_thms: theory -> unit; |
41 |
||
42 |
val init_obj: (thm * thm) * (thm * thm) -> theory -> theory; |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
43 |
val debug: bool ref; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
44 |
val debug_msg: ('a -> string) -> 'a -> 'a; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
45 |
end; |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
46 |
|
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
47 |
structure CodegenTheorems: CODEGEN_THEOREMS = |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
48 |
struct |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
49 |
|
19597 | 50 |
(** preliminaries **) |
51 |
||
52 |
(* diagnostics *) |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
53 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
54 |
val debug = ref false; |
20404 | 55 |
fun debug_msg f x = (if !debug then Output.tracing (f x) else (); x); |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
56 |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
57 |
|
19597 | 58 |
(* auxiliary *) |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
59 |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
60 |
fun getf_first [] _ = NONE |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
61 |
| getf_first (f::fs) x = case f x |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
62 |
of NONE => getf_first fs x |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
63 |
| y as SOME x => y; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
64 |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
65 |
fun getf_first_list [] x = [] |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
66 |
| getf_first_list (f::fs) x = case f x |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
67 |
of [] => getf_first_list fs x |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
68 |
| xs => xs; |
19597 | 69 |
|
70 |
||
71 |
(* object logic setup *) |
|
72 |
||
73 |
structure CodegenTheoremsSetup = TheoryDataFun |
|
74 |
(struct |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19897
diff
changeset
|
75 |
val name = "Pure/codegen_theorems_setup"; |
19597 | 76 |
type T = ((string * thm) * ((string * string) * (string * string))) option; |
77 |
val empty = NONE; |
|
78 |
val copy = I; |
|
79 |
val extend = I; |
|
80 |
fun merge pp = merge_opt (eq_pair (eq_pair (op =) eq_thm) |
|
81 |
(eq_pair (eq_pair (op =) (op =)) (eq_pair (op =) (op =)))) : T * T -> T; |
|
82 |
fun print thy data = (); |
|
83 |
end); |
|
84 |
||
85 |
val _ = Context.add_setup CodegenTheoremsSetup.init; |
|
86 |
||
87 |
fun init_obj ((TrueI, FalseE), (conjI, atomize_eq)) thy = |
|
88 |
case CodegenTheoremsSetup.get thy |
|
20394 | 89 |
of SOME _ => error "Code generator already set up for object logic" |
20191 | 90 |
| NONE => |
19597 | 91 |
let |
92 |
fun strip_implies t = (Logic.strip_imp_prems t, Logic.strip_imp_concl t); |
|
93 |
fun dest_TrueI thm = |
|
94 |
Drule.plain_prop_of thm |
|
95 |
|> ObjectLogic.drop_judgment thy |
|
96 |
|> Term.dest_Const |
|
97 |
|> apsnd ( |
|
98 |
Term.dest_Type |
|
99 |
#> fst |
|
100 |
); |
|
101 |
fun dest_FalseE thm = |
|
102 |
Drule.plain_prop_of thm |
|
103 |
|> Logic.dest_implies |
|
104 |
|> apsnd ( |
|
105 |
ObjectLogic.drop_judgment thy |
|
106 |
#> Term.dest_Var |
|
107 |
) |
|
108 |
|> fst |
|
109 |
|> ObjectLogic.drop_judgment thy |
|
110 |
|> Term.dest_Const |
|
111 |
|> fst; |
|
112 |
fun dest_conjI thm = |
|
113 |
Drule.plain_prop_of thm |
|
114 |
|> strip_implies |
|
115 |
|> apfst (map (ObjectLogic.drop_judgment thy #> Term.dest_Var)) |
|
116 |
|> apsnd ( |
|
117 |
ObjectLogic.drop_judgment thy |
|
118 |
#> Term.strip_comb |
|
119 |
#> apsnd (map Term.dest_Var) |
|
120 |
#> apfst Term.dest_Const |
|
121 |
) |
|
20394 | 122 |
|> (fn (v1, ((conj, _), v2)) => if v1 = v2 then conj else error "Wrong premise") |
20191 | 123 |
fun dest_atomize_eq thm= |
19597 | 124 |
Drule.plain_prop_of thm |
125 |
|> Logic.dest_equals |
|
126 |
|> apfst ( |
|
127 |
ObjectLogic.drop_judgment thy |
|
128 |
#> Term.strip_comb |
|
129 |
#> apsnd (map Term.dest_Var) |
|
130 |
#> apfst Term.dest_Const |
|
131 |
) |
|
132 |
|> apsnd ( |
|
133 |
Logic.dest_equals |
|
134 |
#> apfst Term.dest_Var |
|
135 |
#> apsnd Term.dest_Var |
|
136 |
) |
|
137 |
|> (fn (((eq, _), v2), (v1a as (_, TVar (_, sort)), v1b)) => |
|
20394 | 138 |
if [v1a, v1b] = v2 andalso sort = Sign.defaultS thy then eq else error "Wrong premise") |
19597 | 139 |
in |
140 |
((dest_TrueI TrueI, [dest_FalseE FalseE, dest_conjI conjI, dest_atomize_eq atomize_eq]) |
|
20394 | 141 |
handle _ => error "Bad code generator setup") |
19597 | 142 |
|> (fn ((tr, b), [fl, con, eq]) => CodegenTheoremsSetup.put |
143 |
(SOME ((b, atomize_eq), ((tr, fl), (con, eq)))) thy) |
|
144 |
end; |
|
145 |
||
146 |
fun get_obj thy = |
|
147 |
case CodegenTheoremsSetup.get thy |
|
148 |
of SOME ((b, atomize), x) => ((Type (b, []), atomize) ,x) |
|
20394 | 149 |
| NONE => error "No object logic setup for code theorems"; |
19597 | 150 |
|
151 |
fun mk_true thy = |
|
152 |
let |
|
153 |
val ((b, _), ((tr, fl), (con, eq))) = get_obj thy; |
|
154 |
in Const (tr, b) end; |
|
155 |
||
156 |
fun mk_false thy = |
|
157 |
let |
|
158 |
val ((b, _), ((tr, fl), (con, eq))) = get_obj thy; |
|
159 |
in Const (fl, b) end; |
|
160 |
||
161 |
fun mk_obj_conj thy (x, y) = |
|
162 |
let |
|
163 |
val ((b, _), ((tr, fl), (con, eq))) = get_obj thy; |
|
164 |
in Const (con, b --> b --> b) $ x $ y end; |
|
165 |
||
166 |
fun mk_obj_eq thy (x, y) = |
|
167 |
let |
|
168 |
val ((b, _), ((tr, fl), (con, eq))) = get_obj thy; |
|
20105 | 169 |
in Const (eq, fastype_of x --> fastype_of y --> b) $ x $ y end; |
19597 | 170 |
|
171 |
fun is_obj_eq thy c = |
|
172 |
let |
|
173 |
val ((b, _), ((tr, fl), (con, eq))) = get_obj thy; |
|
174 |
in c = eq end; |
|
175 |
||
176 |
fun mk_func thy ((x, y), rhs) = |
|
177 |
Logic.mk_equals ( |
|
178 |
(mk_obj_eq thy (x, y)), |
|
179 |
rhs |
|
180 |
); |
|
181 |
||
182 |
||
183 |
(* theorem purification *) |
|
184 |
||
185 |
fun err_thm msg thm = |
|
186 |
error (msg ^ ": " ^ string_of_thm thm); |
|
187 |
||
20404 | 188 |
val mk_rule = |
189 |
#mk o #mk_rews o snd o MetaSimplifier.rep_ss o Simplifier.simpset_of; |
|
190 |
||
19597 | 191 |
fun abs_norm thy thm = |
192 |
let |
|
193 |
fun expvars t = |
|
194 |
let |
|
195 |
val lhs = (fst o Logic.dest_equals) t; |
|
20105 | 196 |
val tys = (fst o strip_type o fastype_of) lhs; |
19597 | 197 |
val used = fold_aterms (fn Var ((v, _), _) => insert (op =) v | _ => I) lhs []; |
20076 | 198 |
val vs = Name.invent_list used "x" (length tys); |
19597 | 199 |
in |
200 |
map2 (fn v => fn ty => Var ((v, 0), ty)) vs tys |
|
201 |
end; |
|
202 |
fun expand ct thm = |
|
203 |
Thm.combination thm (Thm.reflexive ct); |
|
204 |
fun beta_norm thm = |
|
205 |
thm |
|
206 |
|> prop_of |
|
207 |
|> Logic.dest_equals |
|
208 |
|> fst |
|
209 |
|> cterm_of thy |
|
210 |
|> Thm.beta_conversion true |
|
211 |
|> Thm.symmetric |
|
212 |
|> (fn thm' => Thm.transitive thm' thm); |
|
213 |
in |
|
214 |
thm |
|
215 |
|> fold (expand o cterm_of thy) ((expvars o prop_of) thm) |
|
216 |
|> beta_norm |
|
217 |
end; |
|
218 |
||
219 |
fun canonical_tvars thy thm = |
|
220 |
let |
|
221 |
fun mk_inst (v_i as (v, i), (v', sort)) (s as (maxidx, set, acc)) = |
|
222 |
if v = v' orelse member (op =) set v then s |
|
223 |
else let |
|
224 |
val ty = TVar (v_i, sort) |
|
225 |
in |
|
226 |
(maxidx + 1, v :: set, |
|
227 |
(ctyp_of thy ty, ctyp_of thy (TVar ((v', maxidx), sort))) :: acc) |
|
228 |
end; |
|
229 |
fun tvars_of thm = (fold_types o fold_atyps) |
|
20386 | 230 |
(fn TVar (v_i as (v, i), sort) => cons (v_i, (CodegenNames.purify_var v, sort)) |
19597 | 231 |
| _ => I) (prop_of thm) []; |
232 |
val maxidx = Thm.maxidx_of thm + 1; |
|
233 |
val (_, _, inst) = fold mk_inst (tvars_of thm) (maxidx + 1, [], []); |
|
234 |
in Thm.instantiate (inst, []) thm end; |
|
235 |
||
236 |
fun canonical_vars thy thm = |
|
237 |
let |
|
238 |
fun mk_inst (v_i as (v, i), (v', ty)) (s as (maxidx, set, acc)) = |
|
239 |
if v = v' orelse member (op =) set v then s |
|
240 |
else let |
|
241 |
val t = if i = ~1 then Free (v, ty) else Var (v_i, ty) |
|
20191 | 242 |
in |
19597 | 243 |
(maxidx + 1, v :: set, |
244 |
(cterm_of thy t, cterm_of thy (Var ((v', maxidx), ty))) :: acc) |
|
245 |
end; |
|
246 |
fun vars_of thm = fold_aterms |
|
20386 | 247 |
(fn Var (v_i as (v, i), ty) => cons (v_i, (CodegenNames.purify_var v, ty)) |
19597 | 248 |
| _ => I) (prop_of thm) []; |
249 |
val maxidx = Thm.maxidx_of thm + 1; |
|
250 |
val (_, _, inst) = fold mk_inst (vars_of thm) (maxidx + 1, [], []); |
|
251 |
in Thm.instantiate ([], inst) thm end; |
|
252 |
||
19884 | 253 |
fun drop_redundant thy eqs = |
254 |
let |
|
255 |
val matches = curry (Pattern.matches thy o |
|
256 |
pairself (fst o Logic.dest_equals o prop_of)) |
|
257 |
fun drop eqs [] = eqs |
|
258 |
| drop eqs (eq::eqs') = |
|
259 |
drop (eq::eqs) (filter_out (matches eq) eqs') |
|
260 |
in drop [] eqs end; |
|
261 |
||
20404 | 262 |
fun drop_refl thy = filter_out (is_equal o Term.fast_term_ord o Logic.dest_equals |
263 |
o ObjectLogic.drop_judgment thy o Drule.plain_prop_of); |
|
264 |
||
20191 | 265 |
fun make_eq thy = |
19597 | 266 |
let |
267 |
val ((_, atomize), _) = get_obj thy; |
|
268 |
in rewrite_rule [atomize] end; |
|
269 |
||
270 |
fun dest_eq thy thm = |
|
271 |
case try (make_eq thy #> Drule.plain_prop_of |
|
272 |
#> ObjectLogic.drop_judgment thy #> Logic.dest_equals) thm |
|
273 |
of SOME eq => (eq, thm) |
|
20394 | 274 |
| NONE => err_thm "Not an equation" thm; |
19597 | 275 |
|
276 |
fun dest_fun thy thm = |
|
277 |
let |
|
278 |
fun dest_fun' ((lhs, _), thm) = |
|
279 |
case try (dest_Const o fst o strip_comb) lhs |
|
280 |
of SOME (c, ty) => (c, (ty, thm)) |
|
20394 | 281 |
| NONE => err_thm "Not a function equation" thm; |
19597 | 282 |
in |
283 |
thm |
|
284 |
|> dest_eq thy |
|
285 |
|> dest_fun' |
|
286 |
end; |
|
287 |
||
288 |
||
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
289 |
|
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
290 |
(** theory data **) |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
291 |
|
19597 | 292 |
(* data structures *) |
293 |
||
20394 | 294 |
structure Consttab = CodegenConsts.Consttab; |
295 |
||
19597 | 296 |
fun merge' eq (xys as (xs, ys)) = |
297 |
if eq_list eq (xs, ys) then (false, xs) else (true, merge eq xys); |
|
298 |
||
299 |
fun alist_merge' eq_key eq (xys as (xs, ys)) = |
|
300 |
if eq_list (eq_pair eq_key eq) (xs, ys) then (false, xs) else (true, AList.merge eq_key eq xys); |
|
301 |
||
20394 | 302 |
fun list_consttab_join' eq (xyt as (xt, yt)) = |
19597 | 303 |
let |
20394 | 304 |
val xc = Consttab.keys xt; |
305 |
val yc = Consttab.keys yt; |
|
306 |
val zc = filter (member CodegenConsts.eq_const yc) xc; |
|
19597 | 307 |
val wc = subtract (op =) zc xc @ subtract (op =) zc yc; |
20394 | 308 |
fun same_thms c = if eq_list eq_thm ((the o Consttab.lookup xt) c, (the o Consttab.lookup yt) c) |
19597 | 309 |
then NONE else SOME c; |
20394 | 310 |
in (wc @ map_filter same_thms zc, Consttab.join (K (merge eq)) xyt) end; |
19597 | 311 |
|
19884 | 312 |
datatype notify = Notify of (serial * ((string * typ) list option -> theory -> theory)) list; |
19597 | 313 |
|
314 |
val mk_notify = Notify; |
|
315 |
fun map_notify f (Notify notify) = mk_notify (f notify); |
|
316 |
fun merge_notify pp (Notify notify1, Notify notify2) = |
|
317 |
mk_notify (AList.merge (op =) (K true) (notify1, notify2)); |
|
318 |
||
319 |
datatype preproc = Preproc of { |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
320 |
preprocs: (serial * (theory -> thm list -> thm list)) list, |
19597 | 321 |
unfolds: thm list |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
322 |
}; |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
323 |
|
19597 | 324 |
fun mk_preproc (preprocs, unfolds) = |
325 |
Preproc { preprocs = preprocs, unfolds = unfolds }; |
|
326 |
fun map_preproc f (Preproc { preprocs, unfolds }) = |
|
327 |
mk_preproc (f (preprocs, unfolds)); |
|
328 |
fun merge_preproc _ (Preproc { preprocs = preprocs1, unfolds = unfolds1 }, |
|
329 |
Preproc { preprocs = preprocs2, unfolds = unfolds2 }) = |
|
330 |
let |
|
331 |
val (dirty1, preprocs) = alist_merge' (op =) (K true) (preprocs1, preprocs2); |
|
332 |
val (dirty2, unfolds) = merge' eq_thm (unfolds1, unfolds2); |
|
333 |
in (dirty1 orelse dirty2, mk_preproc (preprocs, unfolds)) end; |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
334 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
335 |
datatype extrs = Extrs of { |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
336 |
funs: (serial * (theory -> string * typ -> thm list)) list, |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
337 |
datatypes: (serial * (theory -> string -> (((string * sort) list * (string * typ list) list) * tactic) option)) list |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
338 |
}; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
339 |
|
19597 | 340 |
fun mk_extrs (funs, datatypes) = |
341 |
Extrs { funs = funs, datatypes = datatypes }; |
|
342 |
fun map_extrs f (Extrs { funs, datatypes }) = |
|
343 |
mk_extrs (f (funs, datatypes)); |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
344 |
fun merge_extrs _ (Extrs { funs = funs1, datatypes = datatypes1 }, |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
345 |
Extrs { funs = funs2, datatypes = datatypes2 }) = |
19597 | 346 |
let |
347 |
val (dirty1, funs) = alist_merge' (op =) (K true) (funs1, funs2); |
|
348 |
val (dirty2, datatypes) = alist_merge' (op =) (K true) (datatypes1, datatypes2); |
|
349 |
in (dirty1 orelse dirty2, mk_extrs (funs, datatypes)) end; |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
350 |
|
19597 | 351 |
datatype funthms = Funthms of { |
352 |
dirty: string list, |
|
20394 | 353 |
funs: thm list Consttab.table |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
354 |
}; |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
355 |
|
19597 | 356 |
fun mk_funthms (dirty, funs) = |
357 |
Funthms { dirty = dirty, funs = funs }; |
|
358 |
fun map_funthms f (Funthms { dirty, funs }) = |
|
359 |
mk_funthms (f (dirty, funs)); |
|
360 |
fun merge_funthms _ (Funthms { dirty = dirty1, funs = funs1 }, |
|
361 |
Funthms { dirty = dirty2, funs = funs2 }) = |
|
362 |
let |
|
20394 | 363 |
val (dirty3, funs) = list_consttab_join' eq_thm (funs1, funs2); |
364 |
in mk_funthms (merge (op =) (merge (op =) (dirty1, dirty2), map fst dirty3), funs) end; |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
365 |
|
19597 | 366 |
datatype T = T of { |
367 |
dirty: bool, |
|
368 |
notify: notify, |
|
369 |
preproc: preproc, |
|
370 |
extrs: extrs, |
|
371 |
funthms: funthms |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
372 |
}; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
373 |
|
19597 | 374 |
fun mk_T ((dirty, notify), (preproc, (extrs, funthms))) = |
20191 | 375 |
T { dirty = dirty, notify = notify, preproc= preproc, extrs = extrs, funthms = funthms }; |
19597 | 376 |
fun map_T f (T { dirty, notify, preproc, extrs, funthms }) = |
377 |
mk_T (f ((dirty, notify), (preproc, (extrs, funthms)))); |
|
378 |
fun merge_T pp (T { dirty = dirty1, notify = notify1, preproc = preproc1, extrs = extrs1, funthms = funthms1 }, |
|
379 |
T { dirty = dirty2, notify = notify2, preproc = preproc2, extrs = extrs2, funthms = funthms2 }) = |
|
380 |
let |
|
381 |
val (dirty3, preproc) = merge_preproc pp (preproc1, preproc2); |
|
382 |
val (dirty4, extrs) = merge_extrs pp (extrs1, extrs2); |
|
383 |
in |
|
384 |
mk_T ((dirty1 orelse dirty2 orelse dirty3 orelse dirty4, merge_notify pp (notify1, notify2)), |
|
385 |
(preproc, (extrs, merge_funthms pp (funthms1, funthms2)))) |
|
386 |
end; |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
387 |
|
19597 | 388 |
|
389 |
(* setup *) |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
390 |
|
19597 | 391 |
structure CodegenTheoremsData = TheoryDataFun |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
392 |
(struct |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19897
diff
changeset
|
393 |
val name = "Pure/codegen_theorems_data"; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
394 |
type T = T; |
19597 | 395 |
val empty = mk_T ((false, mk_notify []), (mk_preproc ([], []), |
20394 | 396 |
(mk_extrs ([], []), mk_funthms ([], Consttab.empty)))); |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
397 |
val copy = I; |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
398 |
val extend = I; |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
399 |
val merge = merge_T; |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
400 |
fun print (thy : theory) (data : T) = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
401 |
let |
19436 | 402 |
val pretty_thm = ProofContext.pretty_thm (ProofContext.init thy); |
19597 | 403 |
val funthms = (fn T { funthms, ... } => funthms) data; |
20394 | 404 |
val funs = (Consttab.dest o (fn Funthms { funs, ... } => funs)) funthms; |
19597 | 405 |
val preproc = (fn T { preproc, ... } => preproc) data; |
406 |
val unfolds = (fn Preproc { unfolds, ... } => unfolds) preproc; |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
407 |
in |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
408 |
(Pretty.writeln o Pretty.block o Pretty.fbreaks) ([ |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
409 |
Pretty.str "code generation theorems:", |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
410 |
Pretty.str "function theorems:" ] @ |
20191 | 411 |
map (fn (c, thms) => |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
412 |
(Pretty.block o Pretty.fbreaks) ( |
20394 | 413 |
(Pretty.str o CodegenConsts.string_of_const thy) c :: map pretty_thm (rev thms) |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
414 |
) |
20404 | 415 |
) funs @ [ |
19597 | 416 |
Pretty.block ( |
20353 | 417 |
Pretty.str "inlined theorems:" |
19597 | 418 |
:: Pretty.fbrk |
419 |
:: (Pretty.fbreaks o map pretty_thm) unfolds |
|
420 |
)]) |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
421 |
end; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
422 |
end); |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
423 |
|
19597 | 424 |
val _ = Context.add_setup CodegenTheoremsData.init; |
425 |
val print_thms = CodegenTheoremsData.print; |
|
426 |
||
427 |
||
428 |
(* accessors *) |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
429 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
430 |
local |
19597 | 431 |
val the_preproc = (fn T { preproc = Preproc preproc, ... } => preproc) o CodegenTheoremsData.get; |
432 |
val the_extrs = (fn T { extrs = Extrs extrs, ... } => extrs) o CodegenTheoremsData.get; |
|
433 |
val the_funthms = (fn T { funthms = Funthms funthms, ... } => funthms) o CodegenTheoremsData.get; |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
434 |
in |
19597 | 435 |
val is_dirty = (fn T { dirty = dirty, ... } => dirty) o CodegenTheoremsData.get; |
436 |
val the_dirty_consts = (fn { dirty = dirty, ... } => dirty) o the_funthms; |
|
437 |
val the_notify = (fn T { notify = Notify notify, ... } => map snd notify) o CodegenTheoremsData.get; |
|
438 |
val the_preprocs = (fn { preprocs, ... } => map snd preprocs) o the_preproc; |
|
439 |
val the_unfolds = (fn { unfolds, ... } => unfolds) o the_preproc; |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
440 |
val the_funs_extrs = (fn { funs, ... } => map snd funs) o the_extrs; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
441 |
val the_datatypes_extrs = (fn { datatypes, ... } => map snd datatypes) o the_extrs; |
19597 | 442 |
val the_funs = (fn { funs, ... } => funs) o the_funthms; |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
443 |
end (*local*); |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
444 |
|
19597 | 445 |
val map_data = CodegenTheoremsData.map o map_T; |
446 |
||
447 |
(* notifiers *) |
|
448 |
||
19884 | 449 |
fun all_typs thy c = |
20394 | 450 |
let |
451 |
val c_tys = (map (pair c o #lhs o snd) o Defs.specifications_of (Theory.defs_of thy)) c; |
|
452 |
in (c, Sign.the_const_type thy c) :: map (CodegenConsts.typ_of_typinst thy) c_tys end; |
|
19884 | 453 |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
454 |
fun add_notify f = |
19597 | 455 |
map_data (fn ((dirty, notify), x) => |
456 |
((dirty, notify |> map_notify (cons (serial (), f))), x)); |
|
457 |
||
458 |
fun get_reset_dirty thy = |
|
459 |
let |
|
460 |
val dirty = is_dirty thy; |
|
461 |
val dirty_const = if dirty then [] else the_dirty_consts thy; |
|
462 |
in |
|
463 |
thy |
|
464 |
|> map_data (fn ((_, notify), (procs, (extrs, funthms))) => |
|
465 |
((false, notify), (procs, (extrs, funthms |> map_funthms (fn (_, funs) => ([], funs)))))) |
|
466 |
|> pair (dirty, dirty_const) |
|
467 |
end; |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
468 |
|
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
469 |
fun notify_all c thy = |
19597 | 470 |
thy |
471 |
|> get_reset_dirty |
|
472 |
|-> (fn (true, _) => fold (fn f => f NONE) (the_notify thy) |
|
19884 | 473 |
| (false, cs) => let val cs' = case c of NONE => cs | SOME c => insert (op =) c cs |
474 |
in fold (fn f => f (SOME (maps (all_typs thy) cs'))) (the_notify thy) end); |
|
19597 | 475 |
|
476 |
fun notify_dirty thy = |
|
477 |
thy |
|
478 |
|> get_reset_dirty |
|
479 |
|-> (fn (true, _) => fold (fn f => f NONE) (the_notify thy) |
|
19884 | 480 |
| (false, cs) => fold (fn f => f (SOME (maps (all_typs thy) cs))) (the_notify thy)); |
19597 | 481 |
|
482 |
||
483 |
(* modifiers *) |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
484 |
|
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
485 |
fun add_preproc f = |
19597 | 486 |
map_data (fn (x, (preproc, y)) => |
487 |
(x, (preproc |> map_preproc (fn (preprocs, unfolds) => ((serial (), f) :: preprocs, unfolds)), y))) |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
488 |
#> notify_all NONE; |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
489 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
490 |
fun add_fun_extr f = |
19597 | 491 |
map_data (fn (x, (preproc, (extrs, funthms))) => |
492 |
(x, (preproc, (extrs |> map_extrs (fn (funs, datatypes) => |
|
493 |
((serial (), f) :: funs, datatypes)), funthms)))) |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
494 |
#> notify_all NONE; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
495 |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
496 |
fun add_datatype_extr f = |
19597 | 497 |
map_data (fn (x, (preproc, (extrs, funthms))) => |
498 |
(x, (preproc, (extrs |> map_extrs (fn (funs, datatypes) => |
|
499 |
(funs, (serial (), f) :: datatypes)), funthms)))) |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
500 |
#> notify_all NONE; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
501 |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
502 |
fun add_fun thm thy = |
19597 | 503 |
case dest_fun thy thm |
20394 | 504 |
of (c, (ty, _)) => |
19597 | 505 |
thy |
506 |
|> map_data (fn (x, (preproc, (extrs, funthms))) => |
|
507 |
(x, (preproc, (extrs, funthms |> map_funthms (fn (dirty, funs) => |
|
20394 | 508 |
(dirty, funs |> Consttab.map_default (CodegenConsts.norminst_of_typ thy (c, ty), []) (cons thm))))))) |
19597 | 509 |
|> notify_all (SOME c); |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
510 |
|
19597 | 511 |
fun del_fun thm thy = |
512 |
case dest_fun thy thm |
|
20394 | 513 |
of (c, (ty, _)) => |
19597 | 514 |
thy |
515 |
|> map_data (fn (x, (preproc, (extrs, funthms))) => |
|
516 |
(x, (preproc, (extrs, funthms |> map_funthms (fn (dirty, funs) => |
|
20394 | 517 |
(dirty, funs |> Consttab.map_entry (CodegenConsts.norminst_of_typ thy (c, ty)) (remove eq_thm thm))))))) |
19597 | 518 |
|> notify_all (SOME c); |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
519 |
|
19597 | 520 |
fun add_unfold thm thy = |
521 |
thy |
|
522 |
|> tap (fn thy => dest_eq thy thm) |
|
523 |
|> map_data (fn (x, (preproc, y)) => |
|
524 |
(x, (preproc |> map_preproc (fn (preprocs, unfolds) => |
|
525 |
(preprocs, thm :: unfolds)), y))) |
|
526 |
|> notify_all NONE; |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
527 |
|
20191 | 528 |
fun del_unfold thm = |
19597 | 529 |
map_data (fn (x, (preproc, y)) => |
530 |
(x, (preproc |> map_preproc (fn (preprocs, unfolds) => |
|
531 |
(preprocs, remove eq_thm thm unfolds)), y))) |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
532 |
#> notify_all NONE; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
533 |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
534 |
fun purge_defs (c, ty) thy = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
535 |
thy |
19597 | 536 |
|> map_data (fn (x, (preproc, (extrs, funthms))) => |
537 |
(x, (preproc, (extrs, funthms |> map_funthms (fn (dirty, funs) => |
|
20394 | 538 |
(dirty, funs |> Consttab.update (CodegenConsts.norminst_of_typ thy (c, ty), []))))))) |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
539 |
|> notify_all (SOME c); |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
540 |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
541 |
|
19597 | 542 |
|
543 |
(** theorem handling **) |
|
544 |
||
545 |
(* preprocessing *) |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
546 |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19897
diff
changeset
|
547 |
fun extr_typ thy thm = case dest_fun thy thm |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19897
diff
changeset
|
548 |
of (_, (ty, _)) => ty; |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19897
diff
changeset
|
549 |
|
20404 | 550 |
fun rewrite_fun rewrites thm = |
551 |
let |
|
552 |
val rewrite = Tactic.rewrite true rewrites; |
|
553 |
val (ct_eq, [ct_lhs, ct_rhs]) = (Drule.strip_comb o cprop_of) thm; |
|
554 |
val Const ("==", _) = term_of ct_eq; |
|
555 |
val (ct_f, ct_args) = Drule.strip_comb ct_lhs; |
|
556 |
val rhs' = rewrite ct_rhs; |
|
557 |
val args' = map rewrite ct_args; |
|
558 |
val lhs' = Thm.symmetric (fold (fn th1 => fn th2 => Thm.combination th2 th1) |
|
559 |
args' (Thm.reflexive ct_f)); |
|
560 |
in |
|
561 |
Thm.transitive (Thm.transitive lhs' thm) rhs' |
|
562 |
end handle Bind => raise ERROR "rewrite_fun" |
|
20191 | 563 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
564 |
fun common_typ thy _ [] = [] |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
565 |
| common_typ thy _ [thm] = [thm] |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
566 |
| common_typ thy extract_typ thms = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
567 |
let |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
568 |
fun incr_thm thm max = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
569 |
let |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
570 |
val thm' = incr_indexes max thm; |
20105 | 571 |
val max' = (maxidx_of_typ o fastype_of o Drule.plain_prop_of) thm' + 1; |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
572 |
in (thm', max') end; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
573 |
val (thms', maxidx) = fold_map incr_thm thms 0; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
574 |
val (ty1::tys) = map extract_typ thms; |
20394 | 575 |
fun unify ty env = Sign.typ_unify thy (ty1, ty) env |
576 |
handle Type.TUNIFY => |
|
577 |
error ("Type unificaton failed, while unifying function equations\n" |
|
578 |
^ (cat_lines o map Display.string_of_thm) thms |
|
579 |
^ "\nwith types\n" |
|
580 |
^ (cat_lines o map (Sign.string_of_typ thy)) (ty1 :: tys)); |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
581 |
val (env, _) = fold unify tys (Vartab.empty, maxidx) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
582 |
val instT = Vartab.fold (fn (x_i, (sort, ty)) => |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
583 |
cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env []; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
584 |
in map (Thm.instantiate (instT, [])) thms end; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
585 |
|
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
586 |
fun preprocess thy thms = |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
587 |
let |
19597 | 588 |
fun burrow_thms f [] = [] |
20191 | 589 |
| burrow_thms f thms = |
19597 | 590 |
thms |
591 |
|> Conjunction.intr_list |
|
592 |
|> f |
|
593 |
|> Conjunction.elim_list; |
|
19884 | 594 |
fun cmp_thms (thm1, thm2) = |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19897
diff
changeset
|
595 |
not (Sign.typ_instance thy (extr_typ thy thm1, extr_typ thy thm2)); |
19884 | 596 |
fun unvarify thms = |
20218
be3bfb0699ba
Variable.import(T): result includes fixed types/terms;
wenzelm
parents:
20203
diff
changeset
|
597 |
#2 (#1 (Variable.import true thms (ProofContext.init thy))); |
20404 | 598 |
val unfold_thms = map (make_eq thy) (the_unfolds thy); |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
599 |
in |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
600 |
thms |
19597 | 601 |
|> map (make_eq thy) |
602 |
|> map (Thm.transfer thy) |
|
603 |
|> fold (fn f => f thy) (the_preprocs thy) |
|
20404 | 604 |
|> map (rewrite_fun unfold_thms) |
19884 | 605 |
|> debug_msg (fn _ => "[cg_thm] sorting") |
19597 | 606 |
|> debug_msg (commas o map string_of_thm) |
19884 | 607 |
|> sort (make_ord cmp_thms) |
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
608 |
|> debug_msg (fn _ => "[cg_thm] common_typ") |
19597 | 609 |
|> debug_msg (commas o map string_of_thm) |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19897
diff
changeset
|
610 |
|> common_typ thy (extr_typ thy) |
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
611 |
|> debug_msg (fn _ => "[cg_thm] abs_norm") |
19597 | 612 |
|> debug_msg (commas o map string_of_thm) |
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
613 |
|> map (abs_norm thy) |
20404 | 614 |
|> drop_refl thy |
19597 | 615 |
|> burrow_thms ( |
616 |
debug_msg (fn _ => "[cg_thm] canonical tvars") |
|
617 |
#> debug_msg (string_of_thm) |
|
618 |
#> canonical_tvars thy |
|
619 |
#> debug_msg (fn _ => "[cg_thm] canonical vars") |
|
620 |
#> debug_msg (string_of_thm) |
|
621 |
#> canonical_vars thy |
|
622 |
#> debug_msg (fn _ => "[cg_thm] zero indices") |
|
623 |
#> debug_msg (string_of_thm) |
|
624 |
#> Drule.zero_var_indexes |
|
625 |
) |
|
19884 | 626 |
|> drop_redundant thy |
20353 | 627 |
|> debug_msg (fn _ => "[cg_thm] preprocessing done") |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
628 |
end; |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
629 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
630 |
|
19597 | 631 |
(* retrieval *) |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
632 |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
633 |
fun get_funs thy (c, ty) = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
634 |
let |
19597 | 635 |
val _ = debug_msg (fn _ => "[cg_thm] const (1) " ^ c ^ " :: " ^ Sign.string_of_typ thy ty) () |
20394 | 636 |
val postprocess_typ = case AxClass.class_of_param thy c |
637 |
of NONE => map_filter (fn (_, (ty', thm)) => |
|
638 |
if Sign.typ_instance thy (ty, ty') |
|
639 |
then SOME thm else debug_msg (fn _ => "[cg_thm] dropping " ^ string_of_thm thm) NONE) |
|
640 |
| SOME _ => let |
|
641 |
(*FIXME make this more elegant*) |
|
642 |
val ty' = CodegenConsts.typ_of_classop thy (CodegenConsts.norminst_of_typ thy (c, ty)); |
|
643 |
val ct = Thm.cterm_of thy (Const (c, ty')); |
|
644 |
val thm' = Thm.reflexive ct; |
|
645 |
in map (snd o snd) #> cons thm' #> common_typ thy (extr_typ thy) #> tl end; |
|
19597 | 646 |
fun get_funs (c, ty) = |
20394 | 647 |
(these o Consttab.lookup (the_funs thy) o CodegenConsts.norminst_of_typ thy) (c, ty) |
19597 | 648 |
|> debug_msg (fn _ => "[cg_thm] trying funs") |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
649 |
|> map (dest_fun thy) |
20394 | 650 |
|> postprocess_typ; |
19597 | 651 |
fun get_extr (c, ty) = |
652 |
getf_first_list (map (fn f => f thy) (the_funs_extrs thy)) (c, ty) |
|
653 |
|> debug_msg (fn _ => "[cg_thm] trying extr") |
|
654 |
|> map (dest_fun thy) |
|
20394 | 655 |
|> postprocess_typ; |
19597 | 656 |
fun get_spec (c, ty) = |
20394 | 657 |
(CodegenConsts.find_def thy o CodegenConsts.norminst_of_typ thy) (c, ty) |
19597 | 658 |
|> debug_msg (fn _ => "[cg_thm] trying spec") |
20394 | 659 |
|> Option.mapPartial (fn ((_, name), _) => try (Thm.get_axiom_i thy) name) |
660 |
|> the_list |
|
19597 | 661 |
|> map_filter (try (dest_fun thy)) |
20394 | 662 |
|> postprocess_typ; |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
663 |
in |
19597 | 664 |
getf_first_list [get_funs, get_extr, get_spec] (c, ty) |
665 |
|> debug_msg (fn _ => "[cg_thm] const (2) " ^ c ^ " :: " ^ Sign.string_of_typ thy ty) |
|
666 |
|> preprocess thy |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
667 |
end; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
668 |
|
20394 | 669 |
fun prove_freeness thy tac dtco vs_cos = |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
670 |
let |
19597 | 671 |
val truh = mk_true thy; |
672 |
val fals = mk_false thy; |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
673 |
fun mk_lhs vs ((co1, tys1), (co2, tys2)) = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
674 |
let |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
675 |
val dty = Type (dtco, map TFree vs); |
20394 | 676 |
val (xs1, xs2) = chop (length tys1) (Name.invent_list [] "a" (length tys1 + length tys2)); |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
677 |
val frees1 = map2 (fn x => fn ty => Free (x, ty)) xs1 tys1; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
678 |
val frees2 = map2 (fn x => fn ty => Free (x, ty)) xs2 tys2; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
679 |
fun zip_co co xs tys = list_comb (Const (co, |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
680 |
tys ---> dty), map2 (fn x => fn ty => Free (x, ty)) xs tys); |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
681 |
in |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
682 |
((frees1, frees2), (zip_co co1 xs1 tys1, zip_co co2 xs2 tys2)) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
683 |
end; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
684 |
fun mk_rhs [] [] = truh |
19597 | 685 |
| mk_rhs xs ys = foldr1 (mk_obj_conj thy) (map2 (curry (mk_obj_eq thy)) xs ys); |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
686 |
fun mk_eq vs (args as ((co1, _), (co2, _))) (inj, dist) = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
687 |
if co1 = co2 |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
688 |
then let |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
689 |
val ((fs1, fs2), lhs) = mk_lhs vs args; |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
690 |
val rhs = mk_rhs fs1 fs2; |
19597 | 691 |
in (mk_func thy (lhs, rhs) :: inj, dist) end |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
692 |
else let |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
693 |
val (_, lhs) = mk_lhs vs args; |
19597 | 694 |
in (inj, mk_func thy (lhs, fals) :: dist) end; |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
695 |
fun mk_eqs (vs, cos) = |
20191 | 696 |
let val cos' = rev cos |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
697 |
in (op @) (fold (mk_eq vs) (product cos' cos') ([], [])) end; |
20404 | 698 |
val ts = (map (ObjectLogic.ensure_propT thy) o mk_eqs) vs_cos; |
699 |
fun prove t = if !quick_and_dirty then SkipProof.make_thm thy (Logic.varify t) |
|
700 |
else Goal.prove_global thy [] [] t (K tac); |
|
701 |
in map prove ts end; |
|
20394 | 702 |
|
703 |
fun get_datatypes thy dtco = |
|
704 |
let |
|
705 |
val _ = debug_msg (fn _ => "[cg_thm] datatype " ^ dtco) () |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
706 |
in |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
707 |
case getf_first (map (fn f => f thy) (the_datatypes_extrs thy)) dtco |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
708 |
of NONE => NONE |
20394 | 709 |
| SOME (vs_cos, tac) => SOME (vs_cos, prove_freeness thy tac dtco vs_cos) |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
710 |
end; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
711 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
712 |
fun get_eq thy (c, ty) = |
19597 | 713 |
if is_obj_eq thy c |
19816
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
714 |
then case strip_type ty |
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
715 |
of (Type (tyco, _) :: _, _) => |
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
716 |
(case get_datatypes thy tyco |
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
717 |
of SOME (_, thms) => thms |
a8c8ed1c85e0
removed 'primitive definitions' added (non)strict generation, minor fixes
haftmann
parents:
19806
diff
changeset
|
718 |
| _ => []) |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
719 |
| _ => [] |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
720 |
else []; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
721 |
|
20353 | 722 |
fun check_thms c thms = |
723 |
let |
|
724 |
fun check_head_lhs thm (lhs, rhs) = |
|
725 |
case strip_comb lhs |
|
726 |
of (Const (c', _), _) => if c' = c then () |
|
20394 | 727 |
else error ("Illegal function equation for " ^ quote c |
20353 | 728 |
^ ", actually defining " ^ quote c' ^ ": " ^ Display.string_of_thm thm) |
20394 | 729 |
| _ => error ("Illegal function equation: " ^ Display.string_of_thm thm); |
20353 | 730 |
fun check_vars_lhs thm (lhs, rhs) = |
731 |
if has_duplicates (op =) |
|
732 |
(fold_aterms (fn Free (v, _) => cons v | _ => I) lhs []) |
|
20394 | 733 |
then error ("Repeated variables on left hand side of function equation:" |
20353 | 734 |
^ Display.string_of_thm thm) |
735 |
else (); |
|
736 |
fun check_vars_rhs thm (lhs, rhs) = |
|
737 |
if null (subtract (op =) |
|
738 |
(fold_aterms (fn Free (v, _) => cons v | _ => I) lhs []) |
|
739 |
(fold_aterms (fn Free (v, _) => cons v | _ => I) rhs [])) |
|
740 |
then () |
|
20394 | 741 |
else error ("Free variables on right hand side of function equation:" |
20353 | 742 |
^ Display.string_of_thm thm) |
743 |
val tts = map (Logic.dest_equals o Logic.unvarify o Thm.prop_of) thms; |
|
744 |
in |
|
745 |
(map2 check_head_lhs thms tts; map2 check_vars_lhs thms tts; |
|
746 |
map2 check_vars_rhs thms tts; thms) |
|
747 |
end; |
|
748 |
||
20386 | 749 |
structure Consttab = CodegenConsts.Consttab; |
750 |
type thmtab = (theory * (thm list Consttab.table |
|
751 |
* string Consttab.table) |
|
20191 | 752 |
* (Sorts.algebra * ((string * sort) list * (string * typ list) list) Symtab.table)); |
753 |
||
20386 | 754 |
fun thmtab_empty thy = (thy, (Consttab.empty, Consttab.empty), |
20191 | 755 |
(ClassPackage.operational_algebra thy, Symtab.empty)); |
756 |
||
757 |
fun get_sortalgebra (_, _, (algebra, _)) = |
|
758 |
algebra; |
|
19597 | 759 |
|
20386 | 760 |
fun get_dtyp_of_cons (thy, (_, dtcotab), _) = |
761 |
Consttab.lookup dtcotab o CodegenConsts.norminst_of_typ thy; |
|
762 |
||
763 |
fun get_dtyp_spec (_, _, (_, dttab)) = |
|
764 |
Symtab.lookup dttab; |
|
765 |
||
766 |
fun has_fun_thms (thy, (funtab, _), _) = |
|
767 |
is_some o Consttab.lookup funtab o CodegenConsts.norminst_of_typ thy; |
|
20191 | 768 |
|
20386 | 769 |
fun get_fun_thms (thy, (funtab, _), _) (c_ty as (c, _)) = |
770 |
(check_thms c o these o Consttab.lookup funtab |
|
771 |
o CodegenConsts.norminst_of_typ thy) c_ty; |
|
772 |
||
773 |
fun pretty_funtab thy funtab = |
|
774 |
funtab |
|
775 |
|> CodegenConsts.Consttab.dest |
|
776 |
|> map (fn (c, thms) => |
|
777 |
(Pretty.block o Pretty.fbreaks) ( |
|
778 |
(Pretty.str o CodegenConsts.string_of_const thy) c |
|
779 |
:: map Display.pretty_thm thms |
|
780 |
)) |
|
781 |
|> Pretty.chunks; |
|
20191 | 782 |
|
20386 | 783 |
fun constrain_funtab thy funtab = |
20191 | 784 |
let |
20386 | 785 |
fun max k [] = k |
786 |
| max k (l::ls) = max (if k < l then l else k) ls; |
|
787 |
fun mk_consttyps funtab = |
|
788 |
CodegenConsts.Consttab.empty |
|
789 |
|> CodegenConsts.Consttab.fold (fn (c, thm :: _) => |
|
790 |
CodegenConsts.Consttab.update_new (c, extr_typ thy thm) | (_, []) => I) funtab |
|
791 |
fun mk_typescheme_of typtab (c, ty) = |
|
792 |
CodegenConsts.Consttab.lookup typtab (CodegenConsts.norminst_of_typ thy (c, ty)); |
|
793 |
fun incr_indices (c, thms) maxidx = |
|
794 |
let |
|
795 |
val thms' = map (Thm.incr_indexes maxidx) thms; |
|
796 |
val maxidx' = Int.max |
|
797 |
(maxidx, max ~1 (map Thm.maxidx_of thms') + 1); |
|
798 |
in (thms', maxidx') end; |
|
799 |
fun consts_of_eqs thms = |
|
800 |
let |
|
801 |
fun terms_of_eq thm = |
|
802 |
let |
|
803 |
val (lhs, rhs) = (Logic.dest_equals o Drule.plain_prop_of) thm |
|
804 |
in rhs :: (snd o strip_comb) lhs end; |
|
805 |
in (fold o fold_aterms) (fn Const c => insert (eq_pair (op =) (Type.eq_type Vartab.empty)) c | _ => I) |
|
806 |
(maps terms_of_eq thms) [] |
|
807 |
end; |
|
808 |
val typscheme_of = |
|
809 |
mk_typescheme_of (mk_consttyps funtab); |
|
810 |
val tsig = Sign.tsig_of thy; |
|
811 |
fun unify_const (c, ty) (env, maxidx) = |
|
812 |
case typscheme_of (c, ty) |
|
813 |
of SOME ty_decl => let |
|
814 |
(*val _ = writeln "UNIFY"; |
|
815 |
val _ = writeln (CodegenConsts.string_of_const_typ thy (c, ty))*) |
|
816 |
val ty_decl' = Logic.incr_tvar maxidx ty_decl; |
|
817 |
(*val _ = writeln "WITH"; |
|
818 |
val _ = writeln (CodegenConsts.string_of_const_typ thy (c, ty_decl'))*) |
|
819 |
val maxidx' = Int.max (Term.maxidx_of_typ ty_decl' + 1, maxidx); |
|
820 |
(*val _ = writeln (" " ^ string_of_int maxidx ^ " +> " ^ string_of_int maxidx');*) |
|
821 |
in Type.unify tsig (ty_decl', ty) (env, maxidx') end |
|
822 |
| NONE => (env, maxidx); |
|
823 |
fun apply_unifier unif [] = [] |
|
824 |
| apply_unifier unif (thms as thm :: _) = |
|
825 |
let |
|
826 |
val ty = extr_typ thy thm; |
|
827 |
val ty' = Envir.norm_type unif ty; |
|
828 |
val env = Type.typ_match (Sign.tsig_of thy) (ty, ty') Vartab.empty; |
|
829 |
val inst = Thm.instantiate (Vartab.fold (fn (x_i, (sort, ty)) => |
|
830 |
cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env [], []); |
|
831 |
in map (Drule.zero_var_indexes o inst) thms end; |
|
832 |
(* val _ = writeln "(1)"; *) |
|
833 |
(* val _ = (Pretty.writeln o pretty_funtab thy) funtab; *) |
|
834 |
val (funtab', maxidx) = |
|
835 |
CodegenConsts.Consttab.fold_map incr_indices funtab 0; |
|
836 |
(* val _ = writeln "(2)"; |
|
837 |
* val _ = (Pretty.writeln o pretty_funtab thy) funtab'; |
|
838 |
*) |
|
839 |
val (unif, _) = |
|
840 |
CodegenConsts.Consttab.fold (fold unify_const o consts_of_eqs o snd) |
|
841 |
funtab' (Vartab.empty, maxidx); |
|
842 |
(* val _ = writeln "(3)"; *) |
|
843 |
val funtab'' = |
|
844 |
CodegenConsts.Consttab.map (apply_unifier unif) funtab'; |
|
845 |
(* val _ = writeln "(4)"; |
|
846 |
* val _ = (Pretty.writeln o pretty_funtab thy) funtab''; |
|
847 |
*) |
|
848 |
in funtab'' end; |
|
20191 | 849 |
|
20353 | 850 |
fun mk_thmtab thy cs = |
19597 | 851 |
let |
20191 | 852 |
fun add_tycos (Type (tyco, tys)) = insert (op =) tyco #> fold add_tycos tys |
853 |
| add_tycos _ = I; |
|
20386 | 854 |
fun consts_of ts = |
855 |
Consttab.empty |
|
856 |
|> (fold o fold_aterms) |
|
857 |
(fn Const c_ty => Consttab.update (CodegenConsts.norminst_of_typ thy c_ty, ()) |
|
858 |
| _ => I) ts |
|
859 |
|> Consttab.keys; |
|
20191 | 860 |
fun add_dtyps_of_type ty thmtab = |
861 |
let |
|
862 |
val tycos = add_tycos ty []; |
|
20353 | 863 |
val tycos_new = filter (is_none o get_dtyp_spec thmtab) tycos; |
20386 | 864 |
fun add_dtyp_spec dtco (dtyp_spec as (vs, cs)) ((thy, (funtab, dtcotab), (algebra, dttab))) = |
20191 | 865 |
let |
866 |
fun mk_co (c, tys) = |
|
20386 | 867 |
CodegenConsts.norminst_of_typ thy (c, Logic.varifyT (tys ---> Type (dtco, map TFree vs))); |
20191 | 868 |
in |
20386 | 869 |
(thy, (funtab, dtcotab |> fold (fn c_tys => |
870 |
Consttab.update_new (mk_co c_tys, dtco)) cs), |
|
20191 | 871 |
(algebra, dttab |> Symtab.update_new (dtco, dtyp_spec))) |
872 |
end; |
|
873 |
in |
|
874 |
thmtab |
|
875 |
|> fold (fn tyco => case get_datatypes thy tyco |
|
876 |
of SOME (dtyp_spec, _) => add_dtyp_spec tyco dtyp_spec |
|
877 |
| NONE => I) tycos_new |
|
878 |
end; |
|
879 |
fun known thmtab (c, ty) = |
|
880 |
is_some (get_dtyp_of_cons thmtab (c, ty)) orelse has_fun_thms thmtab (c, ty); |
|
20386 | 881 |
fun add_funthms (c, ty) (thmtab as (thy, (funtab, dtcotab), algebra_dttab))= |
882 |
if known thmtab (c, ty) then thmtab |
|
20191 | 883 |
else let |
884 |
val thms = get_funs thy (c, ty) |
|
20386 | 885 |
val cs_dep = (consts_of o map Thm.prop_of) thms; |
20191 | 886 |
in |
20386 | 887 |
(thy, (funtab |> Consttab.update_new (CodegenConsts.norminst_of_typ thy (c, ty), thms) |
20191 | 888 |
, dtcotab), algebra_dttab) |
889 |
|> fold add_c cs_dep |
|
890 |
end |
|
20386 | 891 |
and add_c (c_tys as (c, tys)) thmtab = |
892 |
thmtab |
|
893 |
|> add_dtyps_of_type (snd (CodegenConsts.typ_of_typinst thy c_tys)) |
|
894 |
|> fold (add_funthms o CodegenConsts.typ_of_typinst thy) |
|
895 |
(CodegenConsts.insts_of_classop thy c_tys); |
|
20191 | 896 |
in |
897 |
thmtab_empty thy |
|
20386 | 898 |
|> fold (add_c o CodegenConsts.norminst_of_typ thy) cs |
899 |
|> (fn (a, (funtab, b), c) => (a, (funtab |> constrain_funtab thy, b), c)) |
|
20191 | 900 |
end; |
19597 | 901 |
|
902 |
||
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
903 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
904 |
(** code attributes and setup **) |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
diff
changeset
|
905 |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
906 |
local |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
907 |
fun add_simple_attribute (name, f) = |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
908 |
(Codegen.add_attribute name o (Scan.succeed o Thm.declaration_attribute)) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
909 |
(Context.map_theory o f); |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
910 |
in |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
911 |
val _ = map (Context.add_setup o add_simple_attribute) [ |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
912 |
("fun", add_fun), |
20353 | 913 |
("nofun", del_fun), |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
914 |
("unfold", (fn thm => Codegen.add_unfold thm #> add_unfold thm)), |
20105 | 915 |
("inline", add_unfold), |
20353 | 916 |
("noinline", del_unfold) |
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
917 |
] |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
918 |
end; (*local*) |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
919 |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
920 |
val _ = Context.add_setup (add_fun_extr get_eq); |
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
921 |
|
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
922 |
end; (*struct*) |