author | haftmann |
Thu, 04 Jan 2007 14:01:39 +0100 | |
changeset 21989 | 0315ecfd3d5d |
parent 21915 | 4e63c55f4cb4 |
child 22021 | 6466a24dee5b |
permissions | -rw-r--r-- |
20600 | 1 |
(* Title: Pure/Tools/codegen_funcgr.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
20855 | 5 |
Retrieving, normalizing and structuring code function theorems |
6 |
in graph with explicit dependencies. |
|
20600 | 7 |
*) |
8 |
||
9 |
signature CODEGEN_FUNCGR = |
|
10 |
sig |
|
11 |
type T; |
|
21120 | 12 |
val make: theory -> CodegenConsts.const list -> T |
21129 | 13 |
val make_term: theory -> ((thm -> thm) -> cterm -> thm -> 'a) -> cterm -> 'a * T |
21120 | 14 |
val funcs: T -> CodegenConsts.const -> thm list |
15 |
val typ: T -> CodegenConsts.const -> typ |
|
16 |
val deps: T -> CodegenConsts.const list -> CodegenConsts.const list list |
|
17 |
val all: T -> CodegenConsts.const list |
|
21989
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
18 |
val norm_varnames: theory -> thm list -> thm list |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
19 |
val expand_eta: theory -> int -> thm -> thm |
20600 | 20 |
val print_codethms: theory -> CodegenConsts.const list -> unit |
20938 | 21 |
structure Constgraph : GRAPH |
20600 | 22 |
end; |
23 |
||
24 |
structure CodegenFuncgr: CODEGEN_FUNCGR = |
|
25 |
struct |
|
26 |
||
27 |
(** code data **) |
|
28 |
||
29 |
structure Consttab = CodegenConsts.Consttab; |
|
30 |
structure Constgraph = GraphFun ( |
|
31 |
type key = CodegenConsts.const; |
|
32 |
val ord = CodegenConsts.const_ord; |
|
33 |
); |
|
34 |
||
35 |
type T = (typ * thm list) Constgraph.T; |
|
36 |
||
37 |
structure Funcgr = CodeDataFun |
|
38 |
(struct |
|
39 |
val name = "Pure/codegen_funcgr"; |
|
40 |
type T = T; |
|
41 |
val empty = Constgraph.empty; |
|
42 |
fun merge _ _ = Constgraph.empty; |
|
20938 | 43 |
fun purge _ NONE _ = Constgraph.empty |
44 |
| purge _ (SOME cs) funcgr = |
|
21463
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21387
diff
changeset
|
45 |
Constgraph.del_nodes ((Constgraph.all_preds funcgr |
20938 | 46 |
o filter (can (Constgraph.get_node funcgr))) cs) funcgr; |
20600 | 47 |
end); |
48 |
||
49 |
val _ = Context.add_setup Funcgr.init; |
|
50 |
||
51 |
||
52 |
(** theorem purification **) |
|
53 |
||
21989
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
54 |
fun expand_eta thy k thm = |
20600 | 55 |
let |
21989
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
56 |
val (lhs, rhs) = (Logic.dest_equals o Drule.plain_prop_of) thm; |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
57 |
val (head, args) = strip_comb lhs; |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
58 |
val l = Int.max (0, k - length args); |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
59 |
val used = Name.make_context (map (fst o fst) (Term.add_vars lhs [])); |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
60 |
fun get_name _ 0 used = ([], used) |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
61 |
| get_name (Abs (v, ty, t)) k used = |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
62 |
used |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
63 |
|> Name.variants [CodegenNames.purify_var v] |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
64 |
||>> get_name t (k - 1) |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
65 |
|>> (fn ([v'], vs') => (v', ty) :: vs') |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
66 |
| get_name t k used = |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
67 |
let |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
68 |
val (tys, _) = (strip_type o fastype_of) t |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
69 |
in case tys |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
70 |
of [] => raise TERM ("expand_eta", [t]) |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
71 |
| ty :: _ => |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
72 |
used |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
73 |
|> Name.variants [""] |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
74 |
|-> (fn [v] => get_name (t $ Var ((v, 0), ty)) (k - 1) |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
75 |
#>> (fn vs' => (v, ty) :: vs')) |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
76 |
end; |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
77 |
val (vs, _) = get_name rhs l used; |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
78 |
val vs_refl = map (fn (v, ty) => Thm.reflexive (Thm.cterm_of thy (Var ((v, 0), ty)))) vs; |
20600 | 79 |
in |
21989
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
80 |
fold (fn refl => fn thm => Thm.combination thm refl) vs_refl thm |
20600 | 81 |
end; |
82 |
||
21989
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
83 |
fun norm_args thy thms = |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
84 |
let |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
85 |
val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals; |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
86 |
val k = fold (curry Int.max o num_args_of o Drule.plain_prop_of) thms 0; |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
87 |
in |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
88 |
thms |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
89 |
|> map (expand_eta thy k) |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
90 |
|> map (Drule.fconv_rule Drule.beta_eta_conversion) |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
91 |
end; |
21120 | 92 |
|
20600 | 93 |
fun canonical_tvars thy thm = |
94 |
let |
|
21159 | 95 |
fun tvars_subst_for thm = (fold_types o fold_atyps) |
96 |
(fn TVar (v_i as (v, _), sort) => let |
|
21915 | 97 |
val v' = CodegenNames.purify_tvar v |
21159 | 98 |
in if v = v' then I |
99 |
else insert (op =) (v_i, (v', sort)) end |
|
20600 | 100 |
| _ => I) (prop_of thm) []; |
21159 | 101 |
fun mk_inst (v_i, (v', sort)) (maxidx, acc) = |
102 |
let |
|
103 |
val ty = TVar (v_i, sort) |
|
104 |
in |
|
105 |
(maxidx + 1, (ctyp_of thy ty, ctyp_of thy (TVar ((v', maxidx), sort))) :: acc) |
|
106 |
end; |
|
20600 | 107 |
val maxidx = Thm.maxidx_of thm + 1; |
21159 | 108 |
val (_, inst) = fold mk_inst (tvars_subst_for thm) (maxidx + 1, []); |
20600 | 109 |
in Thm.instantiate (inst, []) thm end; |
110 |
||
111 |
fun canonical_vars thy thm = |
|
112 |
let |
|
21159 | 113 |
fun vars_subst_for thm = fold_aterms |
114 |
(fn Var (v_i as (v, _), ty) => let |
|
115 |
val v' = CodegenNames.purify_var v |
|
116 |
in if v = v' then I |
|
117 |
else insert (op =) (v_i, (v', ty)) end |
|
20600 | 118 |
| _ => I) (prop_of thm) []; |
21159 | 119 |
fun mk_inst (v_i as (v, i), (v', ty)) (maxidx, acc) = |
120 |
let |
|
121 |
val t = Var (v_i, ty) |
|
122 |
in |
|
123 |
(maxidx + 1, (cterm_of thy t, cterm_of thy (Var ((v', maxidx), ty))) :: acc) |
|
124 |
end; |
|
20600 | 125 |
val maxidx = Thm.maxidx_of thm + 1; |
21159 | 126 |
val (_, inst) = fold mk_inst (vars_subst_for thm) (maxidx + 1, []); |
20600 | 127 |
in Thm.instantiate ([], inst) thm end; |
128 |
||
21989
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
129 |
fun norm_varnames thy thms = |
20600 | 130 |
let |
131 |
fun burrow_thms f [] = [] |
|
132 |
| burrow_thms f thms = |
|
133 |
thms |
|
134 |
|> Conjunction.intr_list |
|
135 |
|> f |
|
136 |
|> Conjunction.elim_list; |
|
137 |
in |
|
138 |
thms |
|
21989
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
139 |
|> norm_args thy |
21159 | 140 |
|> burrow_thms (canonical_tvars thy) |
141 |
|> map (canonical_vars thy) |
|
142 |
|> map Drule.zero_var_indexes |
|
20600 | 143 |
end; |
144 |
||
145 |
||
146 |
||
147 |
(** retrieval **) |
|
148 |
||
21120 | 149 |
fun funcs funcgr = |
150 |
these o Option.map snd o try (Constgraph.get_node funcgr); |
|
20600 | 151 |
|
21120 | 152 |
fun typ funcgr = |
153 |
fst o Constgraph.get_node funcgr; |
|
20600 | 154 |
|
21120 | 155 |
fun deps funcgr cs = |
20835 | 156 |
let |
157 |
val conn = Constgraph.strong_conn funcgr; |
|
158 |
val order = rev conn; |
|
159 |
in |
|
160 |
(map o filter) (member (op =) (Constgraph.all_succs funcgr cs)) order |
|
161 |
|> filter_out null |
|
162 |
end; |
|
163 |
||
21120 | 164 |
fun all funcgr = Constgraph.keys funcgr; |
165 |
||
20600 | 166 |
local |
167 |
||
168 |
fun add_things_of thy f (c, thms) = |
|
169 |
(fold o fold_aterms) |
|
170 |
(fn Const c_ty => let |
|
171 |
val c' = CodegenConsts.norm_of_typ thy c_ty |
|
21120 | 172 |
in if is_some c andalso CodegenConsts.eq_const (the c, c') then I |
20600 | 173 |
else f (c', c_ty) end |
174 |
| _ => I) (maps (op :: o swap o apfst (snd o strip_comb) |
|
175 |
o Logic.dest_equals o Drule.plain_prop_of) thms) |
|
176 |
||
177 |
fun rhs_of thy (c, thms) = |
|
178 |
Consttab.empty |
|
21120 | 179 |
|> add_things_of thy (Consttab.update o rpair () o fst) (SOME c, thms) |
20600 | 180 |
|> Consttab.keys; |
181 |
||
182 |
fun insts_of thy funcgr (c, ty) = |
|
183 |
let |
|
184 |
val tys = Sign.const_typargs thy (c, ty); |
|
185 |
val c' = CodegenConsts.norm thy (c, tys); |
|
20705 | 186 |
val ty_decl = CodegenConsts.disc_typ_of_const thy |
187 |
(fst o Constgraph.get_node funcgr o CodegenConsts.norm thy) (c, tys); |
|
20600 | 188 |
val tys_decl = Sign.const_typargs thy (c, ty_decl); |
189 |
fun constructor tyco xs class = |
|
190 |
(tyco, class) :: maps (maps fst) xs; |
|
191 |
fun variable (TVar (_, sort)) = map (pair []) sort |
|
192 |
| variable (TFree (_, sort)) = map (pair []) sort; |
|
193 |
fun mk_inst ty (TVar (_, sort)) = cons (ty, sort) |
|
194 |
| mk_inst ty (TFree (_, sort)) = cons (ty, sort) |
|
195 |
| mk_inst (Type (tyco1, tys1)) (Type (tyco2, tys2)) = |
|
196 |
if tyco1 <> tyco2 then error "bad instance" |
|
197 |
else fold2 mk_inst tys1 tys2; |
|
21120 | 198 |
val pp = Sign.pp thy; |
199 |
val algebra = Sign.classes_of thy; |
|
200 |
fun classrel (x, _) _ = x; |
|
201 |
fun of_sort_deriv (ty, sort) = |
|
202 |
Sorts.of_sort_derivation pp algebra |
|
203 |
{ classrel = classrel, constructor = constructor, variable = variable } |
|
204 |
(ty, sort) |
|
20600 | 205 |
in |
21120 | 206 |
flat (maps of_sort_deriv (fold2 mk_inst tys tys_decl [])) |
20600 | 207 |
end; |
208 |
||
209 |
fun all_classops thy tyco class = |
|
21463
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21387
diff
changeset
|
210 |
try (AxClass.params_of_class thy) class |
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21387
diff
changeset
|
211 |
|> Option.map snd |
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21387
diff
changeset
|
212 |
|> these |
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21387
diff
changeset
|
213 |
|> map (fn (c, _) => (c, CodegenConsts.disc_typ_of_classop thy (c, [Type (tyco, [])]))) |
20600 | 214 |
|> map (CodegenConsts.norm_of_typ thy); |
215 |
||
216 |
fun instdefs_of thy insts = |
|
217 |
let |
|
218 |
val thy_classes = (#classes o Sorts.rep_algebra o Sign.classes_of) thy; |
|
219 |
in |
|
220 |
Symtab.empty |
|
221 |
|> fold (fn (tyco, class) => |
|
222 |
Symtab.map_default (tyco, []) (insert (op =) class)) insts |
|
223 |
|> (fn tab => Symtab.fold (fn (tyco, classes) => append (maps (all_classops thy tyco) |
|
224 |
(Graph.all_succs thy_classes classes))) tab []) |
|
225 |
end; |
|
226 |
||
21120 | 227 |
fun insts_of_thms thy funcgr (c, thms) = |
20600 | 228 |
let |
229 |
val insts = add_things_of thy (fn (_, c_ty) => fold (insert (op =)) |
|
21120 | 230 |
(insts_of thy funcgr c_ty)) (SOME c, thms) []; |
20600 | 231 |
in instdefs_of thy insts end; |
232 |
||
233 |
fun ensure_const thy funcgr c auxgr = |
|
234 |
if can (Constgraph.get_node funcgr) c |
|
235 |
then (NONE, auxgr) |
|
236 |
else if can (Constgraph.get_node auxgr) c |
|
237 |
then (SOME c, auxgr) |
|
238 |
else if is_some (CodegenData.get_datatype_of_constr thy c) then |
|
239 |
auxgr |
|
240 |
|> Constgraph.new_node (c, []) |
|
241 |
|> pair (SOME c) |
|
242 |
else let |
|
21989
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
243 |
val thms = norm_varnames thy (CodegenData.these_funcs thy c); |
20600 | 244 |
val rhs = rhs_of thy (c, thms); |
245 |
in |
|
246 |
auxgr |
|
247 |
|> Constgraph.new_node (c, thms) |
|
248 |
|> fold_map (ensure_const thy funcgr) rhs |
|
249 |
|-> (fn rhs' => fold (fn SOME c' => Constgraph.add_edge (c, c') |
|
250 |
| NONE => I) rhs') |
|
251 |
|> pair (SOME c) |
|
252 |
end; |
|
253 |
||
21387
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
254 |
exception INVALID of CodegenConsts.const list * string; |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
255 |
|
20600 | 256 |
fun specialize_typs thy funcgr eqss = |
257 |
let |
|
21387
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
258 |
fun max [] = 0 |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
259 |
| max [l] = l |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
260 |
| max (k::l::ls) = max ((if k < l then l else k) :: ls); |
20600 | 261 |
fun typscheme_of (c, ty) = |
262 |
try (Constgraph.get_node funcgr) (CodegenConsts.norm_of_typ thy (c, ty)) |
|
263 |
|> Option.map fst; |
|
21387
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
264 |
fun incr_indices (c, thms) maxidx = |
20600 | 265 |
let |
21387
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
266 |
val thms' = map (Thm.incr_indexes (maxidx + 1)) thms; |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
267 |
val maxidx' = max (maxidx :: map Thm.maxidx_of thms'); |
20600 | 268 |
in ((c, thms'), maxidx') end; |
21387
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
269 |
val (eqss', maxidx) = |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
270 |
fold_map incr_indices eqss 0; |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
271 |
fun unify_const (c, ty) (env, maxidx) = |
20600 | 272 |
case typscheme_of (c, ty) |
273 |
of SOME ty_decl => let |
|
21387
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
274 |
val ty_decl' = Logic.incr_tvar (maxidx + 1) ty_decl; |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
275 |
val maxidx' = max [maxidx, Term.maxidx_of_typ ty_decl']; |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
276 |
in Type.unify (Sign.tsig_of thy) (ty_decl', ty) (env, maxidx') |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
277 |
handle TUNIFY => raise INVALID ([], setmp show_sorts true (setmp show_types true (fn f => f ())) (fn _ => ("Failed to instantiate\n" |
20600 | 278 |
^ (Sign.string_of_typ thy o Envir.norm_type env) ty_decl' ^ "\nto\n" |
21387
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
279 |
^ (Sign.string_of_typ thy o Envir.norm_type env) ty |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
280 |
^ ",\nfor constant " ^ quote c |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
281 |
^ "\nin function theorems\n" |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
282 |
^ (cat_lines o maps (map (Sign.string_of_term thy o map_types (Envir.norm_type env) o Thm.prop_of) o snd)) eqss'))) |
20600 | 283 |
end |
284 |
| NONE => (env, maxidx); |
|
285 |
fun apply_unifier unif (c, []) = (c, []) |
|
286 |
| apply_unifier unif (c, thms as thm :: _) = |
|
287 |
let |
|
21120 | 288 |
val tvars = Term.add_tvars (Thm.prop_of thm) []; |
289 |
fun mk_inst (v_i_sort as (v, _)) = |
|
290 |
let |
|
291 |
val ty = TVar v_i_sort; |
|
292 |
in |
|
293 |
pairself (Thm.ctyp_of thy) (ty, |
|
294 |
TVar (v, (snd o dest_TVar o Envir.norm_type unif) ty)) |
|
295 |
end; |
|
296 |
val instmap = map mk_inst tvars; |
|
21387
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
297 |
val (thms' as thm' :: _) = map (Drule.zero_var_indexes o Thm.instantiate (instmap, [])) thms; |
21989
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
298 |
val _ = if fst c = "" then () |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
299 |
else case pairself (CodegenData.typ_func thy) (thm, thm') |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
300 |
of (ty, ty') => if (is_none o AxClass.class_of_param thy o fst) c |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
301 |
andalso Sign.typ_equiv thy (Type.strip_sorts ty, Type.strip_sorts ty') |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
302 |
orelse Sign.typ_equiv thy (ty, ty') |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
303 |
then () |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
304 |
else raise INVALID ([], "illegal function type instantiation:\n" ^ Sign.string_of_typ thy (CodegenData.typ_func thy thm) |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
305 |
^ "\nto " ^ Sign.string_of_typ thy (CodegenData.typ_func thy thm') |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
306 |
^ ",\nfor constant " ^ CodegenConsts.string_of_const thy c |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
307 |
^ "\nin function theorems\n" |
0315ecfd3d5d
eta-expansion now only to common maximum number of arguments
haftmann
parents:
21915
diff
changeset
|
308 |
^ (cat_lines o map string_of_thm) thms) |
21120 | 309 |
in (c, thms') end; |
310 |
fun rhs_of' thy (("", []), thms as [_]) = |
|
311 |
add_things_of thy (cons o snd) (NONE, thms) [] |
|
312 |
| rhs_of' thy (c, thms) = |
|
313 |
add_things_of thy (cons o snd) (SOME c, thms) []; |
|
20600 | 314 |
val (unif, _) = |
21387
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
315 |
fold (fn (c, thms) => fold unify_const (rhs_of' thy (c, thms))) |
20600 | 316 |
eqss' (Vartab.empty, maxidx); |
317 |
val eqss'' = |
|
318 |
map (apply_unifier unif) eqss'; |
|
319 |
in eqss'' end; |
|
320 |
||
321 |
fun merge_eqsyss thy raw_eqss funcgr = |
|
322 |
let |
|
323 |
val eqss = specialize_typs thy funcgr raw_eqss; |
|
20938 | 324 |
val tys = map (CodegenData.typ_funcs thy) eqss; |
20600 | 325 |
val rhss = map (rhs_of thy) eqss; |
326 |
in |
|
327 |
funcgr |
|
328 |
|> fold2 (fn (c, thms) => fn ty => Constgraph.new_node (c, (ty, thms))) eqss tys |
|
329 |
|> `(fn funcgr => map (insts_of_thms thy funcgr) eqss) |
|
330 |
|-> (fn rhs_insts => fold2 (fn (c, _) => fn rhs_inst => |
|
331 |
ensure_consts thy rhs_inst #> fold (curry Constgraph.add_edge c) rhs_inst) eqss rhs_insts) |
|
332 |
|> fold2 (fn (c, _) => fn rhs => fold (curry Constgraph.add_edge c) rhs) eqss rhss |
|
333 |
end |
|
20705 | 334 |
and merge_new_eqsyss thy raw_eqss funcgr = |
335 |
if exists (member CodegenConsts.eq_const (Constgraph.keys funcgr)) (map fst raw_eqss) |
|
336 |
then funcgr |
|
337 |
else merge_eqsyss thy raw_eqss funcgr |
|
20600 | 338 |
and ensure_consts thy cs funcgr = |
339 |
fold (snd oo ensure_const thy funcgr) cs Constgraph.empty |
|
20705 | 340 |
|> (fn auxgr => fold (merge_new_eqsyss thy) |
20600 | 341 |
(map (AList.make (Constgraph.get_node auxgr)) |
21387
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
342 |
(rev (Constgraph.strong_conn auxgr))) funcgr) |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
343 |
handle INVALID (cs', msg) => raise INVALID (cs @ cs', msg); |
20600 | 344 |
|
21120 | 345 |
fun drop_classes thy tfrees thm = |
346 |
let |
|
347 |
val (_, thm') = Thm.varifyT' [] thm; |
|
348 |
val tvars = Term.add_tvars (Thm.prop_of thm') []; |
|
349 |
val unconstr = map (Thm.ctyp_of thy o TVar) tvars; |
|
350 |
val instmap = map2 (fn (v_i, _) => fn (v, sort) => pairself (Thm.ctyp_of thy) |
|
351 |
(TVar (v_i, []), TFree (v, sort))) tvars tfrees; |
|
352 |
in |
|
353 |
thm' |
|
354 |
|> fold Thm.unconstrainT unconstr |
|
355 |
|> Thm.instantiate (instmap, []) |
|
356 |
|> Tactic.rule_by_tactic ((REPEAT o CHANGED o ALLGOALS o Tactic.resolve_tac) (AxClass.class_intros thy)) |
|
357 |
end; |
|
358 |
||
20600 | 359 |
in |
360 |
||
21387
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
361 |
val ensure_consts = (fn thy => fn cs => fn funcgr => ensure_consts thy |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
362 |
cs funcgr handle INVALID (cs', msg) => error (msg ^ ",\nwhile preprocessing equations for constant(s) " |
5d3d340cb783
clarified code for building function equation system; explicit check of type discipline
haftmann
parents:
21196
diff
changeset
|
363 |
^ commas (map (CodegenConsts.string_of_const thy) cs'))); |
20600 | 364 |
|
21120 | 365 |
fun make thy consts = |
366 |
Funcgr.change thy (ensure_consts thy consts); |
|
367 |
||
21129 | 368 |
fun make_term thy f ct = |
21120 | 369 |
let |
370 |
val _ = Sign.no_vars (Sign.pp thy) (Thm.term_of ct); |
|
371 |
val _ = Term.fold_types (Type.no_tvars #> K I) (Thm.term_of ct) (); |
|
372 |
val thm1 = CodegenData.preprocess_cterm thy ct; |
|
373 |
val ct' = Drule.dest_equals_rhs (Thm.cprop_of thm1); |
|
374 |
val consts = CodegenConsts.consts_of thy (Thm.term_of ct'); |
|
375 |
val funcgr = make thy consts; |
|
376 |
val (_, thm2) = Thm.varifyT' [] thm1; |
|
377 |
val thm3 = Thm.reflexive (Drule.dest_equals_rhs (Thm.cprop_of thm2)); |
|
378 |
val [(_, [thm4])] = specialize_typs thy funcgr [(("", []), [thm3])]; |
|
379 |
val tfrees = Term.add_tfrees (Thm.prop_of thm1) []; |
|
380 |
fun inst thm = |
|
381 |
let |
|
382 |
val tvars = Term.add_tvars (Thm.prop_of thm) []; |
|
383 |
val instmap = map2 (fn (v_i, sort) => fn (v, _) => pairself (Thm.ctyp_of thy) |
|
384 |
(TVar (v_i, sort), TFree (v, sort))) tvars tfrees; |
|
385 |
in Thm.instantiate (instmap, []) thm end; |
|
386 |
val thm5 = inst thm2; |
|
387 |
val thm6 = inst thm4; |
|
388 |
val ct'' = Drule.dest_equals_rhs (Thm.cprop_of thm6); |
|
389 |
val cs = fold_aterms (fn Const c => cons c | _ => I) (Thm.term_of ct'') []; |
|
390 |
val drop = drop_classes thy tfrees; |
|
391 |
val funcgr' = ensure_consts thy |
|
392 |
(instdefs_of thy (fold (fold (insert (op =)) o insts_of thy funcgr) cs [])) funcgr |
|
21129 | 393 |
in (f drop ct'' thm5, Funcgr.change thy (K funcgr')) end; |
20600 | 394 |
|
395 |
end; (*local*) |
|
396 |
||
397 |
fun print_funcgr thy funcgr = |
|
398 |
AList.make (snd o Constgraph.get_node funcgr) (Constgraph.keys funcgr) |
|
399 |
|> (map o apfst) (CodegenConsts.string_of_const thy) |
|
400 |
|> sort (string_ord o pairself fst) |
|
401 |
|> map (fn (s, thms) => |
|
402 |
(Pretty.block o Pretty.fbreaks) ( |
|
403 |
Pretty.str s |
|
404 |
:: map Display.pretty_thm thms |
|
405 |
)) |
|
406 |
|> Pretty.chunks |
|
407 |
|> Pretty.writeln; |
|
408 |
||
409 |
fun print_codethms thy consts = |
|
21120 | 410 |
make thy consts |> print_funcgr thy; |
20600 | 411 |
|
412 |
fun print_codethms_e thy cs = |
|
413 |
print_codethms thy (map (CodegenConsts.read_const thy) cs); |
|
414 |
||
415 |
||
416 |
(** Isar **) |
|
417 |
||
418 |
structure P = OuterParse; |
|
419 |
||
420 |
val print_codethmsK = "print_codethms"; |
|
421 |
||
422 |
val print_codethmsP = |
|
423 |
OuterSyntax.improper_command print_codethmsK "print code theorems of this theory" OuterKeyword.diag |
|
424 |
(Scan.option (P.$$$ "(" |-- Scan.repeat P.term --| P.$$$ ")") |
|
425 |
>> (fn NONE => CodegenData.print_thms |
|
426 |
| SOME cs => fn thy => print_codethms_e thy cs) |
|
427 |
>> (fn f => Toplevel.no_timing o Toplevel.unknown_theory |
|
428 |
o Toplevel.keep (f o Toplevel.theory_of))); |
|
429 |
||
430 |
val _ = OuterSyntax.add_parsers [print_codethmsP]; |
|
431 |
||
432 |
end; (*struct*) |