author | haftmann |
Thu, 26 Apr 2007 13:33:12 +0200 | |
changeset 22804 | d3c23b90c6c6 |
parent 22737 | d87ccbcc2702 |
child 22900 | f8a7c10e1bd0 |
permissions | -rw-r--r-- |
22023 | 1 |
(* Title: Pure/Tools/codegen_func.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
22185 | 5 |
Basic handling of defining equations ("func"s) for code generator framework. |
22023 | 6 |
*) |
7 |
||
8 |
signature CODEGEN_FUNC = |
|
9 |
sig |
|
22185 | 10 |
val assert_rew: thm -> thm |
22737 | 11 |
val mk_rew: thm -> thm |
12 |
val mk_func: thm -> thm |
|
13 |
val head_func: thm -> CodegenConsts.const * typ |
|
14 |
val bad_thm: string -> 'a |
|
15 |
val error_thm: (thm -> thm) -> thm -> thm |
|
16 |
val warning_thm: (thm -> thm) -> thm -> thm option |
|
22185 | 17 |
|
22197 | 18 |
val inst_thm: sort Vartab.table -> thm -> thm |
22033 | 19 |
val expand_eta: int -> thm -> thm |
20 |
val rewrite_func: thm list -> thm -> thm |
|
22211 | 21 |
val norm_args: thm list -> thm list |
22 |
val norm_varnames: (string -> string) -> (string -> string) -> thm list -> thm list |
|
22023 | 23 |
end; |
24 |
||
25 |
structure CodegenFunc : CODEGEN_FUNC = |
|
26 |
struct |
|
27 |
||
22737 | 28 |
|
29 |
(* auxiliary *) |
|
22033 | 30 |
|
22737 | 31 |
exception BAD_THM of string; |
32 |
fun bad_thm msg = raise BAD_THM msg; |
|
33 |
fun error_thm f thm = f thm handle BAD_THM msg => error msg; |
|
34 |
fun warning_thm f thm = SOME (f thm) handle BAD_THM msg |
|
35 |
=> (warning msg; NONE); |
|
22033 | 36 |
|
37 |
||
38 |
(* making rewrite theorems *) |
|
39 |
||
22185 | 40 |
fun assert_rew thm = |
22033 | 41 |
let |
42 |
val thy = Thm.theory_of_thm thm; |
|
22737 | 43 |
val (lhs, rhs) = (Logic.dest_equals o Thm.prop_of) thm |
44 |
handle TERM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm); |
|
22033 | 45 |
fun vars_of t = fold_aterms |
46 |
(fn Var (v, _) => insert (op =) v |
|
22737 | 47 |
| Free _ => bad_thm ("Illegal free variable in rewrite theorem\n" |
48 |
^ Display.string_of_thm thm) |
|
22033 | 49 |
| _ => I) t []; |
50 |
fun tvars_of t = fold_term_types |
|
51 |
(fn _ => fold_atyps (fn TVar (v, _) => insert (op =) v |
|
22737 | 52 |
| TFree _ => bad_thm |
53 |
("Illegal free type variable in rewrite theorem\n" ^ Display.string_of_thm thm))) t []; |
|
22033 | 54 |
val lhs_vs = vars_of lhs; |
55 |
val rhs_vs = vars_of rhs; |
|
56 |
val lhs_tvs = tvars_of lhs; |
|
57 |
val rhs_tvs = tvars_of lhs; |
|
58 |
val _ = if null (subtract (op =) lhs_vs rhs_vs) |
|
59 |
then () |
|
22737 | 60 |
else bad_thm ("Free variables on right hand side of rewrite theorem\n" |
61 |
^ Display.string_of_thm thm); |
|
22033 | 62 |
val _ = if null (subtract (op =) lhs_tvs rhs_tvs) |
63 |
then () |
|
22737 | 64 |
else bad_thm ("Free type variables on right hand side of rewrite theorem\n" |
65 |
^ Display.string_of_thm thm) |
|
22033 | 66 |
in thm end; |
67 |
||
68 |
fun mk_rew thm = |
|
22023 | 69 |
let |
22033 | 70 |
val thy = Thm.theory_of_thm thm; |
22737 | 71 |
val ctxt = ProofContext.init thy; |
22033 | 72 |
in |
22737 | 73 |
thm |
74 |
|> LocalDefs.meta_rewrite_rule ctxt |
|
75 |
|> assert_rew |
|
22033 | 76 |
end; |
77 |
||
78 |
||
22197 | 79 |
(* making defining equations *) |
22033 | 80 |
|
22737 | 81 |
fun assert_func thm = |
82 |
let |
|
83 |
val thy = Thm.theory_of_thm thm; |
|
22804 | 84 |
val (head, args) = (strip_comb o fst o Logic.dest_equals |
22737 | 85 |
o ObjectLogic.drop_judgment thy o Thm.plain_prop_of) thm; |
22804 | 86 |
val _ = case head of Const _ => () | _ => |
87 |
bad_thm ("Equation not headed by constant\n" ^ Display.string_of_thm thm); |
|
22737 | 88 |
val _ = |
89 |
if has_duplicates (op =) |
|
90 |
((fold o fold_aterms) (fn Var (v, _) => cons v |
|
91 |
| _ => I |
|
92 |
) args []) |
|
93 |
then bad_thm ("Duplicated variables on left hand side of equation\n" |
|
94 |
^ Display.string_of_thm thm) |
|
95 |
else () |
|
96 |
fun check _ (Abs _) = bad_thm |
|
97 |
("Abstraction on left hand side of equation\n" |
|
98 |
^ Display.string_of_thm thm) |
|
99 |
| check 0 (Var _) = () |
|
100 |
| check _ (Var _) = bad_thm |
|
101 |
("Variable with application on left hand side of defining equation\n" |
|
102 |
^ Display.string_of_thm thm) |
|
103 |
| check n (t1 $ t2) = (check (n+1) t1; check 0 t2) |
|
104 |
| check n (Const (_, ty)) = if n <> (length o fst o strip_type) ty |
|
105 |
then bad_thm |
|
106 |
("Partially applied constant on left hand side of equation" |
|
107 |
^ Display.string_of_thm thm) |
|
108 |
else (); |
|
109 |
val _ = map (check 0) args; |
|
110 |
in thm end; |
|
22484
25dfebd7b4c8
improved treatment of defining equations stemming from specification tools
haftmann
parents:
22475
diff
changeset
|
111 |
|
22737 | 112 |
val mk_func = assert_func o Drule.fconv_rule Drule.beta_eta_conversion o mk_rew; |
22023 | 113 |
|
22737 | 114 |
fun head_func thm = |
115 |
let |
|
116 |
val thy = Thm.theory_of_thm thm; |
|
117 |
val (Const (c_ty as (_, ty))) = (fst o strip_comb o fst o Logic.dest_equals |
|
118 |
o ObjectLogic.drop_judgment thy o Thm.plain_prop_of) thm; |
|
119 |
val const = CodegenConsts.const_of_cexpr thy c_ty; |
|
120 |
in (const, ty) end; |
|
22033 | 121 |
|
122 |
||
123 |
(* utilities *) |
|
124 |
||
22197 | 125 |
fun inst_thm tvars' thm = |
126 |
let |
|
127 |
val thy = Thm.theory_of_thm thm; |
|
128 |
val tvars = (Term.add_tvars o Thm.prop_of) thm []; |
|
129 |
fun mk_inst (tvar as (v, _)) = case Vartab.lookup tvars' v |
|
130 |
of SOME sort => SOME (pairself (Thm.ctyp_of thy o TVar) (tvar, (v, sort))) |
|
131 |
| NONE => NONE; |
|
132 |
val insts = map_filter mk_inst tvars; |
|
133 |
in Thm.instantiate (insts, []) thm end; |
|
134 |
||
22033 | 135 |
fun expand_eta k thm = |
136 |
let |
|
137 |
val thy = Thm.theory_of_thm thm; |
|
22705 | 138 |
val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm; |
22023 | 139 |
val (head, args) = strip_comb lhs; |
140 |
val l = if k = ~1 |
|
141 |
then (length o fst o strip_abs) rhs |
|
142 |
else Int.max (0, k - length args); |
|
143 |
val used = Name.make_context (map (fst o fst) (Term.add_vars lhs [])); |
|
144 |
fun get_name _ 0 used = ([], used) |
|
145 |
| get_name (Abs (v, ty, t)) k used = |
|
146 |
used |
|
22033 | 147 |
|> Name.variants [v] |
22023 | 148 |
||>> get_name t (k - 1) |
149 |
|>> (fn ([v'], vs') => (v', ty) :: vs') |
|
150 |
| get_name t k used = |
|
151 |
let |
|
152 |
val (tys, _) = (strip_type o fastype_of) t |
|
153 |
in case tys |
|
154 |
of [] => raise TERM ("expand_eta", [t]) |
|
155 |
| ty :: _ => |
|
156 |
used |
|
157 |
|> Name.variants [""] |
|
158 |
|-> (fn [v] => get_name (t $ Var ((v, 0), ty)) (k - 1) |
|
159 |
#>> (fn vs' => (v, ty) :: vs')) |
|
160 |
end; |
|
161 |
val (vs, _) = get_name rhs l used; |
|
162 |
val vs_refl = map (fn (v, ty) => Thm.reflexive (Thm.cterm_of thy (Var ((v, 0), ty)))) vs; |
|
163 |
in |
|
164 |
fold (fn refl => fn thm => Thm.combination thm refl) vs_refl thm |
|
165 |
end; |
|
166 |
||
22033 | 167 |
fun rewrite_func rewrites thm = |
168 |
let |
|
169 |
val rewrite = MetaSimplifier.rewrite false rewrites; |
|
170 |
val (ct_eq, [ct_lhs, ct_rhs]) = (Drule.strip_comb o Thm.cprop_of) thm; |
|
171 |
val Const ("==", _) = Thm.term_of ct_eq; |
|
172 |
val (ct_f, ct_args) = Drule.strip_comb ct_lhs; |
|
173 |
val rhs' = rewrite ct_rhs; |
|
174 |
val args' = map rewrite ct_args; |
|
175 |
val lhs' = Thm.symmetric (fold (fn th1 => fn th2 => Thm.combination th2 th1) |
|
176 |
args' (Thm.reflexive ct_f)); |
|
22197 | 177 |
in Thm.transitive (Thm.transitive lhs' thm) rhs' end; |
22033 | 178 |
|
22211 | 179 |
fun norm_args thms = |
180 |
let |
|
181 |
val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals; |
|
22705 | 182 |
val k = fold (curry Int.max o num_args_of o Thm.plain_prop_of) thms 0; |
22211 | 183 |
in |
184 |
thms |
|
185 |
|> map (expand_eta k) |
|
186 |
|> map (Drule.fconv_rule Drule.beta_eta_conversion) |
|
187 |
end; |
|
188 |
||
189 |
fun canonical_tvars purify_tvar thm = |
|
190 |
let |
|
191 |
val ctyp = Thm.ctyp_of (Thm.theory_of_thm thm); |
|
192 |
fun tvars_subst_for thm = (fold_types o fold_atyps) |
|
193 |
(fn TVar (v_i as (v, _), sort) => let |
|
194 |
val v' = purify_tvar v |
|
195 |
in if v = v' then I |
|
196 |
else insert (op =) (v_i, (v', sort)) end |
|
197 |
| _ => I) (prop_of thm) []; |
|
198 |
fun mk_inst (v_i, (v', sort)) (maxidx, acc) = |
|
199 |
let |
|
200 |
val ty = TVar (v_i, sort) |
|
201 |
in |
|
202 |
(maxidx + 1, (ctyp ty, ctyp (TVar ((v', maxidx), sort))) :: acc) |
|
203 |
end; |
|
204 |
val maxidx = Thm.maxidx_of thm + 1; |
|
205 |
val (_, inst) = fold mk_inst (tvars_subst_for thm) (maxidx + 1, []); |
|
206 |
in Thm.instantiate (inst, []) thm end; |
|
207 |
||
208 |
fun canonical_vars purify_var thm = |
|
209 |
let |
|
210 |
val cterm = Thm.cterm_of (Thm.theory_of_thm thm); |
|
211 |
fun vars_subst_for thm = fold_aterms |
|
212 |
(fn Var (v_i as (v, _), ty) => let |
|
213 |
val v' = purify_var v |
|
214 |
in if v = v' then I |
|
215 |
else insert (op =) (v_i, (v', ty)) end |
|
216 |
| _ => I) (prop_of thm) []; |
|
217 |
fun mk_inst (v_i as (v, i), (v', ty)) (maxidx, acc) = |
|
218 |
let |
|
219 |
val t = Var (v_i, ty) |
|
220 |
in |
|
221 |
(maxidx + 1, (cterm t, cterm (Var ((v', maxidx), ty))) :: acc) |
|
222 |
end; |
|
223 |
val maxidx = Thm.maxidx_of thm + 1; |
|
224 |
val (_, inst) = fold mk_inst (vars_subst_for thm) (maxidx + 1, []); |
|
225 |
in Thm.instantiate ([], inst) thm end; |
|
226 |
||
227 |
fun canonical_absvars purify_var thm = |
|
228 |
let |
|
229 |
val t = Thm.prop_of thm; |
|
230 |
val t' = Term.map_abs_vars purify_var t; |
|
231 |
in Thm.rename_boundvars t t' thm end; |
|
232 |
||
233 |
fun norm_varnames purify_tvar purify_var thms = |
|
234 |
let |
|
235 |
fun burrow_thms f [] = [] |
|
236 |
| burrow_thms f thms = |
|
237 |
thms |
|
238 |
|> Conjunction.intr_list |
|
239 |
|> f |
|
240 |
|> Conjunction.elim_list; |
|
241 |
in |
|
242 |
thms |
|
243 |
|> burrow_thms (canonical_tvars purify_tvar) |
|
244 |
|> map (canonical_vars purify_var) |
|
245 |
|> map (canonical_absvars purify_var) |
|
246 |
|> map Drule.zero_var_indexes |
|
247 |
end; |
|
248 |
||
22023 | 249 |
end; |