author | haftmann |
Wed, 02 Apr 2008 15:58:40 +0200 | |
changeset 26517 | ef036a63f6e9 |
parent 26331 | 92120667172f |
child 26642 | 454d11701fa4 |
permissions | -rw-r--r-- |
24219 | 1 |
(* Title: Tools/code/code_funcgr.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
5 |
Retrieving, normalizing and structuring defining equations in graph |
|
6 |
with explicit dependencies. |
|
7 |
*) |
|
8 |
||
9 |
signature CODE_FUNCGR = |
|
10 |
sig |
|
11 |
type T |
|
12 |
val timing: bool ref |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
13 |
val funcs: T -> string -> thm list |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
14 |
val typ: T -> string -> typ |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
15 |
val all: T -> string list |
24219 | 16 |
val pretty: theory -> T -> Pretty.T |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
17 |
val make: theory -> string list -> T |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
18 |
val make_consts: theory -> string list -> string list * T |
24283 | 19 |
val eval_conv: theory -> (T -> cterm -> thm) -> cterm -> thm |
24835
8c26128f8997
clarified relationship of code generator conversions and evaluations
haftmann
parents:
24713
diff
changeset
|
20 |
val eval_term: theory -> (T -> term -> 'a) -> term -> 'a |
24219 | 21 |
end |
22 |
||
24283 | 23 |
structure CodeFuncgr : CODE_FUNCGR = |
24219 | 24 |
struct |
25 |
||
26 |
(** the graph type **) |
|
27 |
||
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
28 |
type T = (typ * thm list) Graph.T; |
24219 | 29 |
|
30 |
fun funcs funcgr = |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
31 |
these o Option.map snd o try (Graph.get_node funcgr); |
24219 | 32 |
|
33 |
fun typ funcgr = |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
34 |
fst o Graph.get_node funcgr; |
24219 | 35 |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
36 |
fun all funcgr = Graph.keys funcgr; |
24219 | 37 |
|
38 |
fun pretty thy funcgr = |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
39 |
AList.make (snd o Graph.get_node funcgr) (Graph.keys funcgr) |
24219 | 40 |
|> (map o apfst) (CodeUnit.string_of_const thy) |
41 |
|> sort (string_ord o pairself fst) |
|
42 |
|> map (fn (s, thms) => |
|
43 |
(Pretty.block o Pretty.fbreaks) ( |
|
44 |
Pretty.str s |
|
45 |
:: map Display.pretty_thm thms |
|
46 |
)) |
|
47 |
|> Pretty.chunks; |
|
48 |
||
49 |
||
50 |
(** generic combinators **) |
|
51 |
||
52 |
fun fold_consts f thms = |
|
53 |
thms |
|
54 |
|> maps (op :: o swap o apfst (snd o strip_comb) o Logic.dest_equals o Thm.plain_prop_of) |
|
55 |
|> (fold o fold_aterms) (fn Const c => f c | _ => I); |
|
56 |
||
57 |
fun consts_of (const, []) = [] |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
58 |
| consts_of (const, thms as _ :: _) = |
24219 | 59 |
let |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
60 |
fun the_const (c, _) = if c = const then I else insert (op =) c |
24219 | 61 |
in fold_consts the_const thms [] end; |
62 |
||
63 |
fun insts_of thy algebra c ty_decl ty = |
|
64 |
let |
|
65 |
val tys_decl = Sign.const_typargs thy (c, ty_decl); |
|
66 |
val tys = Sign.const_typargs thy (c, ty); |
|
67 |
fun class_relation (x, _) _ = x; |
|
68 |
fun type_constructor tyco xs class = |
|
69 |
(tyco, class) :: maps (maps fst) xs; |
|
70 |
fun type_variable (TVar (_, sort)) = map (pair []) sort |
|
71 |
| type_variable (TFree (_, sort)) = map (pair []) sort; |
|
72 |
fun mk_inst ty (TVar (_, sort)) = cons (ty, sort) |
|
73 |
| mk_inst ty (TFree (_, sort)) = cons (ty, sort) |
|
74 |
| mk_inst (Type (_, tys1)) (Type (_, tys2)) = fold2 mk_inst tys1 tys2; |
|
75 |
fun of_sort_deriv (ty, sort) = |
|
76 |
Sorts.of_sort_derivation (Sign.pp thy) algebra |
|
77 |
{ class_relation = class_relation, type_constructor = type_constructor, |
|
78 |
type_variable = type_variable } |
|
26517 | 79 |
(ty, sort) handle Sorts.CLASS_ERROR _ => [] (*permissive!*) |
24219 | 80 |
in |
81 |
flat (maps of_sort_deriv (fold2 mk_inst tys tys_decl [])) |
|
82 |
end; |
|
83 |
||
84 |
fun drop_classes thy tfrees thm = |
|
85 |
let |
|
86 |
val (_, thm') = Thm.varifyT' [] thm; |
|
87 |
val tvars = Term.add_tvars (Thm.prop_of thm') []; |
|
88 |
val unconstr = map (Thm.ctyp_of thy o TVar) tvars; |
|
89 |
val instmap = map2 (fn (v_i, _) => fn (v, sort) => pairself (Thm.ctyp_of thy) |
|
90 |
(TVar (v_i, []), TFree (v, sort))) tvars tfrees; |
|
91 |
in |
|
92 |
thm' |
|
93 |
|> fold Thm.unconstrainT unconstr |
|
94 |
|> Thm.instantiate (instmap, []) |
|
95 |
|> Tactic.rule_by_tactic ((REPEAT o CHANGED o ALLGOALS o Tactic.resolve_tac) (AxClass.class_intros thy)) |
|
96 |
end; |
|
97 |
||
98 |
||
99 |
(** graph algorithm **) |
|
100 |
||
101 |
val timing = ref false; |
|
102 |
||
103 |
local |
|
104 |
||
26331 | 105 |
exception CLASS_ERROR of string list * string; |
24219 | 106 |
|
107 |
fun resort_thms algebra tap_typ [] = [] |
|
108 |
| resort_thms algebra tap_typ (thms as thm :: _) = |
|
109 |
let |
|
110 |
val thy = Thm.theory_of_thm thm; |
|
26331 | 111 |
val pp = Sign.pp thy; |
24219 | 112 |
val cs = fold_consts (insert (op =)) thms []; |
113 |
fun match_const c (ty, ty_decl) = |
|
114 |
let |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
115 |
val tys = Sign.const_typargs thy (c, ty); |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
116 |
val sorts = map (snd o dest_TVar) (Sign.const_typargs thy (c, ty_decl)); |
26517 | 117 |
in fn tab => fold2 (curry (Sorts.meet_sort algebra)) tys sorts tab |
26331 | 118 |
handle Sorts.CLASS_ERROR e => raise CLASS_ERROR ([c], Sorts.msg_class_error pp e ^ ",\n" |
119 |
^ "for constant " ^ CodeUnit.string_of_const thy c |
|
120 |
^ "\nin defining equations(s)\n" |
|
121 |
^ (cat_lines o map string_of_thm) thms) |
|
26517 | 122 |
(*handle Sorts.CLASS_ERROR _ => tab (*permissive!*)*) |
26331 | 123 |
end; |
26517 | 124 |
fun match (c, ty) = case tap_typ c |
24219 | 125 |
of SOME ty_decl => match_const c (ty, ty_decl) |
126 |
| NONE => I; |
|
127 |
val tvars = fold match cs Vartab.empty; |
|
128 |
in map (CodeUnit.inst_thm tvars) thms end; |
|
129 |
||
130 |
fun resort_funcss thy algebra funcgr = |
|
131 |
let |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
132 |
val typ_funcgr = try (fst o Graph.get_node funcgr); |
26331 | 133 |
val resort_dep = apsnd (resort_thms algebra typ_funcgr); |
24219 | 134 |
fun resort_rec tap_typ (const, []) = (true, (const, [])) |
135 |
| resort_rec tap_typ (const, thms as thm :: _) = |
|
136 |
let |
|
137 |
val (_, ty) = CodeUnit.head_func thm; |
|
138 |
val thms' as thm' :: _ = resort_thms algebra tap_typ thms |
|
139 |
val (_, ty') = CodeUnit.head_func thm'; |
|
140 |
in (Sign.typ_equiv thy (ty, ty'), (const, thms')) end; |
|
141 |
fun resort_recs funcss = |
|
142 |
let |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
143 |
fun tap_typ c = |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
144 |
AList.lookup (op =) funcss c |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
145 |
|> these |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
146 |
|> try hd |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
147 |
|> Option.map (snd o CodeUnit.head_func); |
24219 | 148 |
val (unchangeds, funcss') = split_list (map (resort_rec tap_typ) funcss); |
149 |
val unchanged = fold (fn x => fn y => x andalso y) unchangeds true; |
|
150 |
in (unchanged, funcss') end; |
|
151 |
fun resort_rec_until funcss = |
|
152 |
let |
|
153 |
val (unchanged, funcss') = resort_recs funcss; |
|
154 |
in if unchanged then funcss' else resort_rec_until funcss' end; |
|
155 |
in map resort_dep #> resort_rec_until end; |
|
156 |
||
157 |
fun instances_of thy algebra insts = |
|
158 |
let |
|
159 |
val thy_classes = (#classes o Sorts.rep_algebra o Sign.classes_of) thy; |
|
24835
8c26128f8997
clarified relationship of code generator conversions and evaluations
haftmann
parents:
24713
diff
changeset
|
160 |
fun all_classparams tyco class = |
24969 | 161 |
these (try (#params o AxClass.get_info thy) class) |
26517 | 162 |
|> map_filter (fn (c, _) => try (AxClass.param_of_inst thy) (c, tyco)) |
24219 | 163 |
in |
164 |
Symtab.empty |
|
165 |
|> fold (fn (tyco, class) => |
|
166 |
Symtab.map_default (tyco, []) (insert (op =) class)) insts |
|
24835
8c26128f8997
clarified relationship of code generator conversions and evaluations
haftmann
parents:
24713
diff
changeset
|
167 |
|> (fn tab => Symtab.fold (fn (tyco, classes) => append (maps (all_classparams tyco) |
24219 | 168 |
(Graph.all_succs thy_classes classes))) tab []) |
169 |
end; |
|
170 |
||
171 |
fun instances_of_consts thy algebra funcgr consts = |
|
172 |
let |
|
173 |
fun inst (cexpr as (c, ty)) = insts_of thy algebra c |
|
26517 | 174 |
((fst o Graph.get_node funcgr) c) ty; |
24219 | 175 |
in |
176 |
[] |
|
177 |
|> fold (fold (insert (op =)) o inst) consts |
|
178 |
|> instances_of thy algebra |
|
179 |
end; |
|
180 |
||
24283 | 181 |
fun ensure_const' thy algebra funcgr const auxgr = |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
182 |
if can (Graph.get_node funcgr) const |
24219 | 183 |
then (NONE, auxgr) |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
184 |
else if can (Graph.get_node auxgr) const |
24219 | 185 |
then (SOME const, auxgr) |
186 |
else if is_some (Code.get_datatype_of_constr thy const) then |
|
187 |
auxgr |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
188 |
|> Graph.new_node (const, []) |
24219 | 189 |
|> pair (SOME const) |
190 |
else let |
|
191 |
val thms = Code.these_funcs thy const |
|
192 |
|> CodeUnit.norm_args |
|
193 |
|> CodeUnit.norm_varnames CodeName.purify_tvar CodeName.purify_var; |
|
194 |
val rhs = consts_of (const, thms); |
|
195 |
in |
|
196 |
auxgr |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
197 |
|> Graph.new_node (const, thms) |
24283 | 198 |
|> fold_map (ensure_const thy algebra funcgr) rhs |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
199 |
|-> (fn rhs' => fold (fn SOME const' => Graph.add_edge (const, const') |
24219 | 200 |
| NONE => I) rhs') |
201 |
|> pair (SOME const) |
|
202 |
end |
|
24283 | 203 |
and ensure_const thy algebra funcgr const = |
24219 | 204 |
let |
205 |
val timeap = if !timing |
|
206 |
then Output.timeap_msg ("time for " ^ CodeUnit.string_of_const thy const) |
|
207 |
else I; |
|
24283 | 208 |
in timeap (ensure_const' thy algebra funcgr const) end; |
24219 | 209 |
|
24283 | 210 |
fun merge_funcss thy algebra raw_funcss funcgr = |
24219 | 211 |
let |
212 |
val funcss = raw_funcss |
|
213 |
|> resort_funcss thy algebra funcgr |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
214 |
|> filter_out (can (Graph.get_node funcgr) o fst); |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
215 |
fun typ_func c [] = Code.default_typ thy c |
25597
34860182b250
moved instance parameter management from class.ML to axclass.ML
haftmann
parents:
25485
diff
changeset
|
216 |
| typ_func c (thms as thm :: _) = case AxClass.inst_of_param thy c |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
217 |
of SOME (c', tyco) => |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
218 |
let |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
219 |
val (_, ty) = CodeUnit.head_func thm; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
220 |
val SOME class = AxClass.class_of_param thy c'; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
221 |
val sorts_decl = Sorts.mg_domain algebra tyco [class]; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
222 |
val tys = Sign.const_typargs thy (c, ty); |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
223 |
val sorts = map (snd o dest_TVar) tys; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
224 |
in if sorts = sorts_decl then ty |
26331 | 225 |
else raise CLASS_ERROR ([c], "Illegal instantation for class operation " |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
226 |
^ CodeUnit.string_of_const thy c |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
227 |
^ "\nin defining equations\n" |
26331 | 228 |
^ (cat_lines o map (string_of_thm o AxClass.overload thy)) thms) |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
229 |
end |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
230 |
| NONE => (snd o CodeUnit.head_func) thm; |
24219 | 231 |
fun add_funcs (const, thms) = |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
232 |
Graph.new_node (const, (typ_func const thms, thms)); |
24219 | 233 |
fun add_deps (funcs as (const, thms)) funcgr = |
234 |
let |
|
235 |
val deps = consts_of funcs; |
|
236 |
val insts = instances_of_consts thy algebra funcgr |
|
237 |
(fold_consts (insert (op =)) thms []); |
|
238 |
in |
|
239 |
funcgr |
|
24283 | 240 |
|> ensure_consts' thy algebra insts |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
241 |
|> fold (curry Graph.add_edge const) deps |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
242 |
|> fold (curry Graph.add_edge const) insts |
24219 | 243 |
end; |
244 |
in |
|
245 |
funcgr |
|
246 |
|> fold add_funcs funcss |
|
247 |
|> fold add_deps funcss |
|
248 |
end |
|
24283 | 249 |
and ensure_consts' thy algebra cs funcgr = |
24219 | 250 |
let |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
251 |
val auxgr = Graph.empty |
24283 | 252 |
|> fold (snd oo ensure_const thy algebra funcgr) cs; |
24219 | 253 |
in |
254 |
funcgr |
|
24283 | 255 |
|> fold (merge_funcss thy algebra) |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
256 |
(map (AList.make (Graph.get_node auxgr)) |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
257 |
(rev (Graph.strong_conn auxgr))) |
26331 | 258 |
end handle CLASS_ERROR (cs', msg) |
259 |
=> raise CLASS_ERROR (fold (insert (op =)) cs' cs, msg); |
|
24219 | 260 |
|
261 |
in |
|
262 |
||
263 |
(** retrieval interfaces **) |
|
264 |
||
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
265 |
fun ensure_consts thy algebra consts funcgr = |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
266 |
ensure_consts' thy algebra consts funcgr |
26331 | 267 |
handle CLASS_ERROR (cs', msg) => error (msg ^ ",\nwhile preprocessing equations for constant(s) " |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
268 |
^ commas (map (CodeUnit.string_of_const thy) cs')); |
24219 | 269 |
|
24283 | 270 |
fun check_consts thy consts funcgr = |
24219 | 271 |
let |
272 |
val algebra = Code.coregular_algebra thy; |
|
273 |
fun try_const const funcgr = |
|
24283 | 274 |
(SOME const, ensure_consts' thy algebra [const] funcgr) |
26331 | 275 |
handle CLASS_ERROR (cs', msg) => (NONE, funcgr); |
24219 | 276 |
val (consts', funcgr') = fold_map try_const consts funcgr; |
277 |
in (map_filter I consts', funcgr') end; |
|
278 |
||
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
279 |
fun raw_eval thy f ct funcgr = |
24219 | 280 |
let |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
281 |
val algebra = Code.coregular_algebra thy; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
282 |
fun consts_of ct = fold_aterms (fn Const c_ty => cons c_ty | _ => I) |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
283 |
(Thm.term_of ct) []; |
24219 | 284 |
val _ = Sign.no_vars (Sign.pp thy) (Thm.term_of ct); |
285 |
val _ = Term.fold_types (Type.no_tvars #> K I) (Thm.term_of ct) (); |
|
24283 | 286 |
val thm1 = Code.preprocess_conv ct; |
24219 | 287 |
val ct' = Thm.rhs_of thm1; |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
288 |
val cs = map fst (consts_of ct'); |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
289 |
val funcgr' = ensure_consts thy algebra cs funcgr; |
24219 | 290 |
val (_, thm2) = Thm.varifyT' [] thm1; |
291 |
val thm3 = Thm.reflexive (Thm.rhs_of thm2); |
|
26331 | 292 |
val [thm4] = resort_thms algebra (try (fst o Graph.get_node funcgr')) [thm3] |
293 |
handle CLASS_ERROR (_, msg) => error msg; |
|
24219 | 294 |
val tfrees = Term.add_tfrees (Thm.prop_of thm1) []; |
295 |
fun inst thm = |
|
296 |
let |
|
297 |
val tvars = Term.add_tvars (Thm.prop_of thm) []; |
|
298 |
val instmap = map2 (fn (v_i, sort) => fn (v, _) => pairself (Thm.ctyp_of thy) |
|
299 |
(TVar (v_i, sort), TFree (v, sort))) tvars tfrees; |
|
300 |
in Thm.instantiate (instmap, []) thm end; |
|
301 |
val thm5 = inst thm2; |
|
302 |
val thm6 = inst thm4; |
|
303 |
val ct'' = Thm.rhs_of thm6; |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
304 |
val c_exprs = consts_of ct''; |
24219 | 305 |
val drop = drop_classes thy tfrees; |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
306 |
val instdefs = instances_of_consts thy algebra funcgr' c_exprs; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
307 |
val funcgr'' = ensure_consts thy algebra instdefs funcgr'; |
25103 | 308 |
in (f drop thm5 funcgr'' ct'', funcgr'') end; |
24219 | 309 |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
310 |
fun raw_eval_conv thy conv = |
24219 | 311 |
let |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
312 |
fun conv' drop_classes thm1 funcgr ct = |
24219 | 313 |
let |
314 |
val thm2 = conv funcgr ct; |
|
315 |
val thm3 = Code.postprocess_conv (Thm.rhs_of thm2); |
|
316 |
val thm23 = drop_classes (Thm.transitive thm2 thm3); |
|
317 |
in |
|
318 |
Thm.transitive thm1 thm23 handle THM _ => |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
319 |
error ("could not construct proof:\n" |
24219 | 320 |
^ (cat_lines o map string_of_thm) [thm1, thm2, thm3]) |
321 |
end; |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
322 |
in raw_eval thy conv' end; |
24283 | 323 |
|
24835
8c26128f8997
clarified relationship of code generator conversions and evaluations
haftmann
parents:
24713
diff
changeset
|
324 |
fun raw_eval_term thy f t = |
24283 | 325 |
let |
24835
8c26128f8997
clarified relationship of code generator conversions and evaluations
haftmann
parents:
24713
diff
changeset
|
326 |
fun f' _ _ funcgr ct = f funcgr (Thm.term_of ct); |
8c26128f8997
clarified relationship of code generator conversions and evaluations
haftmann
parents:
24713
diff
changeset
|
327 |
in raw_eval thy f' (Thm.cterm_of thy t) end; |
24219 | 328 |
|
329 |
end; (*local*) |
|
330 |
||
331 |
structure Funcgr = CodeDataFun |
|
24713 | 332 |
( |
24219 | 333 |
type T = T; |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
334 |
val empty = Graph.empty; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
335 |
fun merge _ _ = Graph.empty; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
336 |
fun purge _ NONE _ = Graph.empty |
24219 | 337 |
| purge _ (SOME cs) funcgr = |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
338 |
Graph.del_nodes ((Graph.all_preds funcgr |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
339 |
o filter (can (Graph.get_node funcgr))) cs) funcgr; |
24713 | 340 |
); |
24219 | 341 |
|
342 |
fun make thy = |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
343 |
Funcgr.change thy o ensure_consts thy (Code.coregular_algebra thy); |
24219 | 344 |
|
345 |
fun make_consts thy = |
|
24283 | 346 |
Funcgr.change_yield thy o check_consts thy; |
24219 | 347 |
|
348 |
fun eval_conv thy f = |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
349 |
fst o Funcgr.change_yield thy o raw_eval_conv thy f; |
24283 | 350 |
|
351 |
fun eval_term thy f = |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24283
diff
changeset
|
352 |
fst o Funcgr.change_yield thy o raw_eval_term thy f; |
24219 | 353 |
|
354 |
end; (*struct*) |