| author | haftmann | 
| Sat, 19 Oct 2019 09:15:41 +0000 | |
| changeset 70903 | c550368a4e29 | 
| parent 70019 | 095dce9892e8 | 
| child 74397 | e80c4cde6064 | 
| permissions | -rw-r--r-- | 
| 64962 | 1  | 
(* Title: HOL/Decision_Procs/Reflective_Field.thy  | 
2  | 
Author: Stefan Berghofer  | 
|
3  | 
||
4  | 
Reducing equalities in fields to equalities in rings.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
theory Reflective_Field  | 
|
| 67123 | 8  | 
imports Commutative_Ring  | 
| 64962 | 9  | 
begin  | 
10  | 
||
11  | 
datatype fexpr =  | 
|
12  | 
FCnst int  | 
|
13  | 
| FVar nat  | 
|
14  | 
| FAdd fexpr fexpr  | 
|
15  | 
| FSub fexpr fexpr  | 
|
16  | 
| FMul fexpr fexpr  | 
|
17  | 
| FNeg fexpr  | 
|
18  | 
| FDiv fexpr fexpr  | 
|
19  | 
| FPow fexpr nat  | 
|
20  | 
||
| 67123 | 21  | 
fun (in field) nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a"  | 
22  | 
where  | 
|
23  | 
"nth_el [] n = \<zero>"  | 
|
24  | 
| "nth_el (x # xs) 0 = x"  | 
|
25  | 
| "nth_el (x # xs) (Suc n) = nth_el xs n"  | 
|
| 64962 | 26  | 
|
| 67123 | 27  | 
lemma (in field) nth_el_Cons: "nth_el (x # xs) n = (if n = 0 then x else nth_el xs (n - 1))"  | 
| 64962 | 28  | 
by (cases n) simp_all  | 
29  | 
||
| 67123 | 30  | 
lemma (in field) nth_el_closed [simp]: "in_carrier xs \<Longrightarrow> nth_el xs n \<in> carrier R"  | 
| 64962 | 31  | 
by (induct xs n rule: nth_el.induct) (simp_all add: in_carrier_def)  | 
32  | 
||
33  | 
primrec (in field) feval :: "'a list \<Rightarrow> fexpr \<Rightarrow> 'a"  | 
|
| 67123 | 34  | 
where  | 
35  | 
"feval xs (FCnst c) = \<guillemotleft>c\<guillemotright>"  | 
|
36  | 
| "feval xs (FVar n) = nth_el xs n"  | 
|
37  | 
| "feval xs (FAdd a b) = feval xs a \<oplus> feval xs b"  | 
|
38  | 
| "feval xs (FSub a b) = feval xs a \<ominus> feval xs b"  | 
|
39  | 
| "feval xs (FMul a b) = feval xs a \<otimes> feval xs b"  | 
|
40  | 
| "feval xs (FNeg a) = \<ominus> feval xs a"  | 
|
41  | 
| "feval xs (FDiv a b) = feval xs a \<oslash> feval xs b"  | 
|
| 
67341
 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 
nipkow 
parents: 
67123 
diff
changeset
 | 
42  | 
| "feval xs (FPow a n) = feval xs a [^] n"  | 
| 64962 | 43  | 
|
44  | 
lemma (in field) feval_Cnst:  | 
|
45  | 
"feval xs (FCnst 0) = \<zero>"  | 
|
46  | 
"feval xs (FCnst 1) = \<one>"  | 
|
47  | 
"feval xs (FCnst (numeral n)) = \<guillemotleft>numeral n\<guillemotright>"  | 
|
48  | 
by simp_all  | 
|
49  | 
||
50  | 
datatype pexpr =  | 
|
51  | 
PExpr1 pexpr1  | 
|
52  | 
| PExpr2 pexpr2  | 
|
53  | 
and pexpr1 =  | 
|
54  | 
PCnst int  | 
|
55  | 
| PVar nat  | 
|
56  | 
| PAdd pexpr pexpr  | 
|
57  | 
| PSub pexpr pexpr  | 
|
58  | 
| PNeg pexpr  | 
|
59  | 
and pexpr2 =  | 
|
60  | 
PMul pexpr pexpr  | 
|
61  | 
| PPow pexpr nat  | 
|
62  | 
||
63  | 
lemma pexpr_cases [case_names PCnst PVar PAdd PSub PNeg PMul PPow]:  | 
|
64  | 
assumes  | 
|
65  | 
"\<And>c. e = PExpr1 (PCnst c) \<Longrightarrow> P"  | 
|
66  | 
"\<And>n. e = PExpr1 (PVar n) \<Longrightarrow> P"  | 
|
67  | 
"\<And>e1 e2. e = PExpr1 (PAdd e1 e2) \<Longrightarrow> P"  | 
|
68  | 
"\<And>e1 e2. e = PExpr1 (PSub e1 e2) \<Longrightarrow> P"  | 
|
69  | 
"\<And>e'. e = PExpr1 (PNeg e') \<Longrightarrow> P"  | 
|
70  | 
"\<And>e1 e2. e = PExpr2 (PMul e1 e2) \<Longrightarrow> P"  | 
|
71  | 
"\<And>e' n. e = PExpr2 (PPow e' n) \<Longrightarrow> P"  | 
|
72  | 
shows P  | 
|
73  | 
proof (cases e)  | 
|
74  | 
case (PExpr1 e')  | 
|
75  | 
then show ?thesis  | 
|
76  | 
apply (cases e')  | 
|
| 67123 | 77  | 
apply simp_all  | 
78  | 
apply (erule assms)+  | 
|
| 64962 | 79  | 
done  | 
80  | 
next  | 
|
81  | 
case (PExpr2 e')  | 
|
82  | 
then show ?thesis  | 
|
83  | 
apply (cases e')  | 
|
| 67123 | 84  | 
apply simp_all  | 
85  | 
apply (erule assms)+  | 
|
| 64962 | 86  | 
done  | 
87  | 
qed  | 
|
88  | 
||
89  | 
lemmas pexpr_cases2 = pexpr_cases [case_product pexpr_cases]  | 
|
90  | 
||
91  | 
fun (in field) peval :: "'a list \<Rightarrow> pexpr \<Rightarrow> 'a"  | 
|
| 67123 | 92  | 
where  | 
93  | 
"peval xs (PExpr1 (PCnst c)) = \<guillemotleft>c\<guillemotright>"  | 
|
94  | 
| "peval xs (PExpr1 (PVar n)) = nth_el xs n"  | 
|
95  | 
| "peval xs (PExpr1 (PAdd a b)) = peval xs a \<oplus> peval xs b"  | 
|
96  | 
| "peval xs (PExpr1 (PSub a b)) = peval xs a \<ominus> peval xs b"  | 
|
97  | 
| "peval xs (PExpr1 (PNeg a)) = \<ominus> peval xs a"  | 
|
98  | 
| "peval xs (PExpr2 (PMul a b)) = peval xs a \<otimes> peval xs b"  | 
|
| 
67341
 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 
nipkow 
parents: 
67123 
diff
changeset
 | 
99  | 
| "peval xs (PExpr2 (PPow a n)) = peval xs a [^] n"  | 
| 64962 | 100  | 
|
101  | 
lemma (in field) peval_Cnst:  | 
|
102  | 
"peval xs (PExpr1 (PCnst 0)) = \<zero>"  | 
|
103  | 
"peval xs (PExpr1 (PCnst 1)) = \<one>"  | 
|
104  | 
"peval xs (PExpr1 (PCnst (numeral n))) = \<guillemotleft>numeral n\<guillemotright>"  | 
|
105  | 
"peval xs (PExpr1 (PCnst (- numeral n))) = \<ominus> \<guillemotleft>numeral n\<guillemotright>"  | 
|
106  | 
by simp_all  | 
|
107  | 
||
108  | 
lemma (in field) peval_closed [simp]:  | 
|
109  | 
"in_carrier xs \<Longrightarrow> peval xs e \<in> carrier R"  | 
|
110  | 
"in_carrier xs \<Longrightarrow> peval xs (PExpr1 e1) \<in> carrier R"  | 
|
111  | 
"in_carrier xs \<Longrightarrow> peval xs (PExpr2 e2) \<in> carrier R"  | 
|
112  | 
by (induct e and e1 and e2) simp_all  | 
|
113  | 
||
114  | 
definition npepow :: "pexpr \<Rightarrow> nat \<Rightarrow> pexpr"  | 
|
| 67123 | 115  | 
where "npepow e n =  | 
116  | 
(if n = 0 then PExpr1 (PCnst 1)  | 
|
117  | 
else if n = 1 then e  | 
|
118  | 
else  | 
|
119  | 
(case e of  | 
|
120  | 
PExpr1 (PCnst c) \<Rightarrow> PExpr1 (PCnst (c ^ n))  | 
|
121  | 
| _ \<Rightarrow> PExpr2 (PPow e n)))"  | 
|
| 64962 | 122  | 
|
123  | 
lemma (in field) npepow_correct:  | 
|
124  | 
"in_carrier xs \<Longrightarrow> peval xs (npepow e n) = peval xs (PExpr2 (PPow e n))"  | 
|
| 67123 | 125  | 
by (cases e rule: pexpr_cases) (simp_all add: npepow_def)  | 
| 64962 | 126  | 
|
127  | 
fun npemul :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"  | 
|
| 67123 | 128  | 
where "npemul x y =  | 
129  | 
(case x of  | 
|
130  | 
PExpr1 (PCnst c) \<Rightarrow>  | 
|
131  | 
if c = 0 then x  | 
|
132  | 
else if c = 1 then y else  | 
|
133  | 
(case y of  | 
|
134  | 
PExpr1 (PCnst d) \<Rightarrow> PExpr1 (PCnst (c * d))  | 
|
| 64962 | 135  | 
| _ \<Rightarrow> PExpr2 (PMul x y))  | 
| 67123 | 136  | 
| PExpr2 (PPow e1 n) \<Rightarrow>  | 
137  | 
(case y of  | 
|
138  | 
PExpr2 (PPow e2 m) \<Rightarrow>  | 
|
139  | 
if n = m then npepow (npemul e1 e2) n  | 
|
140  | 
else PExpr2 (PMul x y)  | 
|
141  | 
| PExpr1 (PCnst d) \<Rightarrow>  | 
|
142  | 
if d = 0 then y  | 
|
143  | 
else if d = 1 then x  | 
|
144  | 
else PExpr2 (PMul x y)  | 
|
145  | 
| _ \<Rightarrow> PExpr2 (PMul x y))  | 
|
146  | 
| _ \<Rightarrow>  | 
|
147  | 
(case y of  | 
|
148  | 
PExpr1 (PCnst d) \<Rightarrow>  | 
|
149  | 
if d = 0 then y  | 
|
150  | 
else if d = 1 then x  | 
|
151  | 
else PExpr2 (PMul x y)  | 
|
152  | 
| _ \<Rightarrow> PExpr2 (PMul x y)))"  | 
|
| 64962 | 153  | 
|
154  | 
lemma (in field) npemul_correct:  | 
|
155  | 
"in_carrier xs \<Longrightarrow> peval xs (npemul e1 e2) = peval xs (PExpr2 (PMul e1 e2))"  | 
|
156  | 
proof (induct e1 e2 rule: npemul.induct)  | 
|
157  | 
case (1 x y)  | 
|
158  | 
then show ?case  | 
|
159  | 
proof (cases x y rule: pexpr_cases2)  | 
|
160  | 
case (PPow_PPow e n e' m)  | 
|
161  | 
then show ?thesis  | 
|
| 
70019
 
095dce9892e8
A few results in Algebra, and bits for Analysis
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
162  | 
by (simp del: npemul.simps add: 1 npepow_correct nat_pow_distrib  | 
| 67123 | 163  | 
npemul.simps [of "PExpr2 (PPow e n)" "PExpr2 (PPow e' m)"])  | 
| 64962 | 164  | 
qed simp_all  | 
165  | 
qed  | 
|
166  | 
||
167  | 
declare npemul.simps [simp del]  | 
|
168  | 
||
169  | 
definition npeadd :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"  | 
|
| 67123 | 170  | 
where "npeadd x y =  | 
171  | 
(case x of  | 
|
172  | 
PExpr1 (PCnst c) \<Rightarrow>  | 
|
173  | 
if c = 0 then y  | 
|
174  | 
else  | 
|
175  | 
(case y of  | 
|
176  | 
PExpr1 (PCnst d) \<Rightarrow> PExpr1 (PCnst (c + d))  | 
|
177  | 
| _ \<Rightarrow> PExpr1 (PAdd x y))  | 
|
178  | 
| _ \<Rightarrow>  | 
|
179  | 
(case y of  | 
|
180  | 
PExpr1 (PCnst d) \<Rightarrow>  | 
|
181  | 
if d = 0 then x  | 
|
182  | 
else PExpr1 (PAdd x y)  | 
|
183  | 
| _ \<Rightarrow> PExpr1 (PAdd x y)))"  | 
|
| 64962 | 184  | 
|
185  | 
lemma (in field) npeadd_correct:  | 
|
186  | 
"in_carrier xs \<Longrightarrow> peval xs (npeadd e1 e2) = peval xs (PExpr1 (PAdd e1 e2))"  | 
|
187  | 
by (cases e1 e2 rule: pexpr_cases2) (simp_all add: npeadd_def)  | 
|
188  | 
||
189  | 
definition npesub :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"  | 
|
| 67123 | 190  | 
where "npesub x y =  | 
191  | 
(case y of  | 
|
192  | 
PExpr1 (PCnst d) \<Rightarrow>  | 
|
193  | 
if d = 0 then x  | 
|
194  | 
else  | 
|
195  | 
(case x of  | 
|
196  | 
PExpr1 (PCnst c) \<Rightarrow> PExpr1 (PCnst (c - d))  | 
|
197  | 
| _ \<Rightarrow> PExpr1 (PSub x y))  | 
|
198  | 
| _ \<Rightarrow>  | 
|
199  | 
(case x of  | 
|
200  | 
PExpr1 (PCnst c) \<Rightarrow>  | 
|
201  | 
if c = 0 then PExpr1 (PNeg y)  | 
|
202  | 
else PExpr1 (PSub x y)  | 
|
203  | 
| _ \<Rightarrow> PExpr1 (PSub x y)))"  | 
|
| 64962 | 204  | 
|
205  | 
lemma (in field) npesub_correct:  | 
|
206  | 
"in_carrier xs \<Longrightarrow> peval xs (npesub e1 e2) = peval xs (PExpr1 (PSub e1 e2))"  | 
|
207  | 
by (cases e1 e2 rule: pexpr_cases2) (simp_all add: npesub_def)  | 
|
208  | 
||
209  | 
definition npeneg :: "pexpr \<Rightarrow> pexpr"  | 
|
| 67123 | 210  | 
where "npeneg e =  | 
211  | 
(case e of  | 
|
212  | 
PExpr1 (PCnst c) \<Rightarrow> PExpr1 (PCnst (- c))  | 
|
213  | 
| _ \<Rightarrow> PExpr1 (PNeg e))"  | 
|
| 64962 | 214  | 
|
| 67123 | 215  | 
lemma (in field) npeneg_correct: "peval xs (npeneg e) = peval xs (PExpr1 (PNeg e))"  | 
| 64962 | 216  | 
by (cases e rule: pexpr_cases) (simp_all add: npeneg_def)  | 
217  | 
||
218  | 
lemma option_pair_cases [case_names None Some]:  | 
|
| 67123 | 219  | 
obtains (None) "x = None" | (Some) p q where "x = Some (p, q)"  | 
| 64962 | 220  | 
proof (cases x)  | 
221  | 
case None  | 
|
| 67123 | 222  | 
then show ?thesis by (rule that)  | 
| 64962 | 223  | 
next  | 
224  | 
case (Some r)  | 
|
225  | 
then show ?thesis  | 
|
| 67123 | 226  | 
by (cases r, simp) (rule that)  | 
| 64962 | 227  | 
qed  | 
228  | 
||
| 67123 | 229  | 
fun isin :: "pexpr \<Rightarrow> nat \<Rightarrow> pexpr \<Rightarrow> nat \<Rightarrow> (nat \<times> pexpr) option"  | 
230  | 
where  | 
|
231  | 
"isin e n (PExpr2 (PMul e1 e2)) m =  | 
|
232  | 
(case isin e n e1 m of  | 
|
| 64962 | 233  | 
Some (k, e3) \<Rightarrow>  | 
234  | 
if k = 0 then Some (0, npemul e3 (npepow e2 m))  | 
|
| 67123 | 235  | 
else  | 
236  | 
(case isin e k e2 m of  | 
|
| 64962 | 237  | 
Some (l, e4) \<Rightarrow> Some (l, npemul e3 e4)  | 
238  | 
| None \<Rightarrow> Some (k, npemul e3 (npepow e2 m)))  | 
|
| 67123 | 239  | 
| None \<Rightarrow>  | 
240  | 
(case isin e n e2 m of  | 
|
241  | 
Some (k, e3) \<Rightarrow> Some (k, npemul (npepow e1 m) e3)  | 
|
242  | 
| None \<Rightarrow> None))"  | 
|
243  | 
| "isin e n (PExpr2 (PPow e' k)) m =  | 
|
244  | 
(if k = 0 then None else isin e n e' (k * m))"  | 
|
245  | 
| "isin (PExpr1 e) n (PExpr1 e') m =  | 
|
246  | 
(if e = e' then  | 
|
247  | 
if n \<ge> m then Some (n - m, PExpr1 (PCnst 1))  | 
|
| 64962 | 248  | 
else Some (0, npepow (PExpr1 e) (m - n))  | 
| 67123 | 249  | 
else None)"  | 
250  | 
| "isin (PExpr2 e) n (PExpr1 e') m = None"  | 
|
| 64962 | 251  | 
|
252  | 
lemma (in field) isin_correct:  | 
|
253  | 
assumes "in_carrier xs"  | 
|
| 67123 | 254  | 
and "isin e n e' m = Some (p, e'')"  | 
255  | 
shows "peval xs (PExpr2 (PPow e' m)) = peval xs (PExpr2 (PMul (PExpr2 (PPow e (n - p))) e''))"  | 
|
256  | 
and "p \<le> n"  | 
|
| 64962 | 257  | 
using assms  | 
258  | 
by (induct e n e' m arbitrary: p e'' rule: isin.induct)  | 
|
259  | 
(force  | 
|
260  | 
simp add:  | 
|
| 
70019
 
095dce9892e8
A few results in Algebra, and bits for Analysis
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
261  | 
nat_pow_distrib nat_pow_pow nat_pow_mult m_ac  | 
| 64962 | 262  | 
npemul_correct npepow_correct  | 
263  | 
split: option.split_asm prod.split_asm if_split_asm)+  | 
|
264  | 
||
265  | 
lemma (in field) isin_correct':  | 
|
266  | 
"in_carrier xs \<Longrightarrow> isin e n e' 1 = Some (p, e'') \<Longrightarrow>  | 
|
| 
67341
 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 
nipkow 
parents: 
67123 
diff
changeset
 | 
267  | 
peval xs e' = peval xs e [^] (n - p) \<otimes> peval xs e''"  | 
| 64962 | 268  | 
"in_carrier xs \<Longrightarrow> isin e n e' 1 = Some (p, e'') \<Longrightarrow> p \<le> n"  | 
| 67123 | 269  | 
using isin_correct [where m = 1] by simp_all  | 
| 64962 | 270  | 
|
271  | 
fun split_aux :: "pexpr \<Rightarrow> nat \<Rightarrow> pexpr \<Rightarrow> pexpr \<times> pexpr \<times> pexpr"  | 
|
| 67123 | 272  | 
where  | 
273  | 
"split_aux (PExpr2 (PMul e1 e2)) n e =  | 
|
274  | 
(let  | 
|
275  | 
(left1, common1, right1) = split_aux e1 n e;  | 
|
276  | 
(left2, common2, right2) = split_aux e2 n right1  | 
|
277  | 
in (npemul left1 left2, npemul common1 common2, right2))"  | 
|
278  | 
| "split_aux (PExpr2 (PPow e' m)) n e =  | 
|
279  | 
(if m = 0 then (PExpr1 (PCnst 1), PExpr1 (PCnst 1), e)  | 
|
280  | 
else split_aux e' (m * n) e)"  | 
|
281  | 
| "split_aux (PExpr1 e') n e =  | 
|
282  | 
(case isin (PExpr1 e') n e 1 of  | 
|
283  | 
Some (m, e'') \<Rightarrow>  | 
|
284  | 
(if m = 0 then (PExpr1 (PCnst 1), npepow (PExpr1 e') n, e'')  | 
|
285  | 
else (npepow (PExpr1 e') m, npepow (PExpr1 e') (n - m), e''))  | 
|
286  | 
| None \<Rightarrow> (npepow (PExpr1 e') n, PExpr1 (PCnst 1), e))"  | 
|
| 64962 | 287  | 
|
| 67123 | 288  | 
hide_const Left Right (* FIXME !? *)  | 
| 64962 | 289  | 
|
| 67123 | 290  | 
abbreviation Left :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"  | 
291  | 
where "Left e1 e2 \<equiv> fst (split_aux e1 (Suc 0) e2)"  | 
|
| 64962 | 292  | 
|
| 67123 | 293  | 
abbreviation Common :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"  | 
294  | 
where "Common e1 e2 \<equiv> fst (snd (split_aux e1 (Suc 0) e2))"  | 
|
| 64962 | 295  | 
|
| 67123 | 296  | 
abbreviation Right :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"  | 
297  | 
where "Right e1 e2 \<equiv> snd (snd (split_aux e1 (Suc 0) e2))"  | 
|
| 64962 | 298  | 
|
299  | 
lemma split_aux_induct [case_names 1 2 3]:  | 
|
300  | 
assumes I1: "\<And>e1 e2 n e. P e1 n e \<Longrightarrow> P e2 n (snd (snd (split_aux e1 n e))) \<Longrightarrow>  | 
|
301  | 
P (PExpr2 (PMul e1 e2)) n e"  | 
|
302  | 
and I2: "\<And>e' m n e. (m \<noteq> 0 \<Longrightarrow> P e' (m * n) e) \<Longrightarrow> P (PExpr2 (PPow e' m)) n e"  | 
|
303  | 
and I3: "\<And>e' n e. P (PExpr1 e') n e"  | 
|
304  | 
shows "P x y z"  | 
|
305  | 
proof (induct x y z rule: split_aux.induct)  | 
|
306  | 
case 1  | 
|
307  | 
from 1(1) 1(2) [OF refl prod.collapse prod.collapse]  | 
|
308  | 
show ?case by (rule I1)  | 
|
309  | 
next  | 
|
310  | 
case 2  | 
|
311  | 
then show ?case by (rule I2)  | 
|
312  | 
next  | 
|
313  | 
case 3  | 
|
314  | 
then show ?case by (rule I3)  | 
|
315  | 
qed  | 
|
316  | 
||
317  | 
lemma (in field) split_aux_correct:  | 
|
318  | 
"in_carrier xs \<Longrightarrow>  | 
|
| 67123 | 319  | 
peval xs (PExpr2 (PPow e\<^sub>1 n)) =  | 
320  | 
peval xs (PExpr2 (PMul (fst (split_aux e\<^sub>1 n e\<^sub>2)) (fst (snd (split_aux e\<^sub>1 n e\<^sub>2)))))"  | 
|
| 64962 | 321  | 
"in_carrier xs \<Longrightarrow>  | 
| 67123 | 322  | 
peval xs e\<^sub>2 =  | 
323  | 
peval xs (PExpr2 (PMul (snd (snd (split_aux e\<^sub>1 n e\<^sub>2))) (fst (snd (split_aux e\<^sub>1 n e\<^sub>2)))))"  | 
|
| 64962 | 324  | 
by (induct e\<^sub>1 n e\<^sub>2 rule: split_aux_induct)  | 
325  | 
(auto simp add: split_beta  | 
|
| 
70019
 
095dce9892e8
A few results in Algebra, and bits for Analysis
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
326  | 
nat_pow_distrib nat_pow_pow nat_pow_mult m_ac  | 
| 64962 | 327  | 
npemul_correct npepow_correct isin_correct'  | 
328  | 
split: option.split)  | 
|
329  | 
||
330  | 
lemma (in field) split_aux_correct':  | 
|
331  | 
"in_carrier xs \<Longrightarrow>  | 
|
| 67123 | 332  | 
peval xs e\<^sub>1 = peval xs (Left e\<^sub>1 e\<^sub>2) \<otimes> peval xs (Common e\<^sub>1 e\<^sub>2)"  | 
| 64962 | 333  | 
"in_carrier xs \<Longrightarrow>  | 
| 67123 | 334  | 
peval xs e\<^sub>2 = peval xs (Right e\<^sub>1 e\<^sub>2) \<otimes> peval xs (Common e\<^sub>1 e\<^sub>2)"  | 
335  | 
using split_aux_correct [where n = 1] by simp_all  | 
|
| 64962 | 336  | 
|
337  | 
fun fnorm :: "fexpr \<Rightarrow> pexpr \<times> pexpr \<times> pexpr list"  | 
|
| 67123 | 338  | 
where  | 
339  | 
"fnorm (FCnst c) = (PExpr1 (PCnst c), PExpr1 (PCnst 1), [])"  | 
|
340  | 
| "fnorm (FVar n) = (PExpr1 (PVar n), PExpr1 (PCnst 1), [])"  | 
|
341  | 
| "fnorm (FAdd e1 e2) =  | 
|
342  | 
(let  | 
|
343  | 
(xn, xd, xc) = fnorm e1;  | 
|
344  | 
(yn, yd, yc) = fnorm e2;  | 
|
345  | 
(left, common, right) = split_aux xd 1 yd  | 
|
346  | 
in  | 
|
347  | 
(npeadd (npemul xn right) (npemul yn left),  | 
|
348  | 
npemul left (npemul right common),  | 
|
349  | 
List.union xc yc))"  | 
|
350  | 
| "fnorm (FSub e1 e2) =  | 
|
351  | 
(let  | 
|
352  | 
(xn, xd, xc) = fnorm e1;  | 
|
353  | 
(yn, yd, yc) = fnorm e2;  | 
|
354  | 
(left, common, right) = split_aux xd 1 yd  | 
|
355  | 
in  | 
|
356  | 
(npesub (npemul xn right) (npemul yn left),  | 
|
357  | 
npemul left (npemul right common),  | 
|
358  | 
List.union xc yc))"  | 
|
359  | 
| "fnorm (FMul e1 e2) =  | 
|
360  | 
(let  | 
|
361  | 
(xn, xd, xc) = fnorm e1;  | 
|
362  | 
(yn, yd, yc) = fnorm e2;  | 
|
363  | 
(left1, common1, right1) = split_aux xn 1 yd;  | 
|
364  | 
(left2, common2, right2) = split_aux yn 1 xd  | 
|
365  | 
in  | 
|
366  | 
(npemul left1 left2,  | 
|
367  | 
npemul right2 right1,  | 
|
368  | 
List.union xc yc))"  | 
|
369  | 
| "fnorm (FNeg e) =  | 
|
370  | 
(let (n, d, c) = fnorm e  | 
|
371  | 
in (npeneg n, d, c))"  | 
|
372  | 
| "fnorm (FDiv e1 e2) =  | 
|
373  | 
(let  | 
|
374  | 
(xn, xd, xc) = fnorm e1;  | 
|
375  | 
(yn, yd, yc) = fnorm e2;  | 
|
376  | 
(left1, common1, right1) = split_aux xn 1 yn;  | 
|
377  | 
(left2, common2, right2) = split_aux xd 1 yd  | 
|
378  | 
in  | 
|
379  | 
(npemul left1 right2,  | 
|
380  | 
npemul left2 right1,  | 
|
381  | 
List.insert yn (List.union xc yc)))"  | 
|
382  | 
| "fnorm (FPow e m) =  | 
|
383  | 
(let (n, d, c) = fnorm e  | 
|
384  | 
in (npepow n m, npepow d m, c))"  | 
|
| 64962 | 385  | 
|
| 67123 | 386  | 
abbreviation Num :: "fexpr \<Rightarrow> pexpr"  | 
387  | 
where "Num e \<equiv> fst (fnorm e)"  | 
|
| 64962 | 388  | 
|
| 67123 | 389  | 
abbreviation Denom :: "fexpr \<Rightarrow> pexpr"  | 
390  | 
where "Denom e \<equiv> fst (snd (fnorm e))"  | 
|
| 64962 | 391  | 
|
| 67123 | 392  | 
abbreviation Cond :: "fexpr \<Rightarrow> pexpr list"  | 
393  | 
where "Cond e \<equiv> snd (snd (fnorm e))"  | 
|
| 64962 | 394  | 
|
395  | 
primrec (in field) nonzero :: "'a list \<Rightarrow> pexpr list \<Rightarrow> bool"  | 
|
| 67123 | 396  | 
where  | 
397  | 
"nonzero xs [] \<longleftrightarrow> True"  | 
|
398  | 
| "nonzero xs (p # ps) \<longleftrightarrow> peval xs p \<noteq> \<zero> \<and> nonzero xs ps"  | 
|
| 64962 | 399  | 
|
| 67123 | 400  | 
lemma (in field) nonzero_singleton: "nonzero xs [p] = (peval xs p \<noteq> \<zero>)"  | 
| 64962 | 401  | 
by simp  | 
402  | 
||
| 67123 | 403  | 
lemma (in field) nonzero_append: "nonzero xs (ps @ qs) = (nonzero xs ps \<and> nonzero xs qs)"  | 
| 64962 | 404  | 
by (induct ps) simp_all  | 
405  | 
||
406  | 
lemma (in field) nonzero_idempotent:  | 
|
407  | 
"p \<in> set ps \<Longrightarrow> (peval xs p \<noteq> \<zero> \<and> nonzero xs ps) = nonzero xs ps"  | 
|
408  | 
by (induct ps) auto  | 
|
409  | 
||
410  | 
lemma (in field) nonzero_insert:  | 
|
411  | 
"nonzero xs (List.insert p ps) = (peval xs p \<noteq> \<zero> \<and> nonzero xs ps)"  | 
|
412  | 
by (simp add: List.insert_def nonzero_idempotent)  | 
|
413  | 
||
414  | 
lemma (in field) nonzero_union:  | 
|
415  | 
"nonzero xs (List.union ps qs) = (nonzero xs ps \<and> nonzero xs qs)"  | 
|
416  | 
by (induct ps rule: rev_induct)  | 
|
417  | 
(auto simp add: List.union_def nonzero_insert nonzero_append)  | 
|
418  | 
||
419  | 
lemma (in field) fnorm_correct:  | 
|
420  | 
assumes "in_carrier xs"  | 
|
| 67123 | 421  | 
and "nonzero xs (Cond e)"  | 
| 64962 | 422  | 
shows "feval xs e = peval xs (Num e) \<oslash> peval xs (Denom e)"  | 
| 67123 | 423  | 
and "peval xs (Denom e) \<noteq> \<zero>"  | 
| 64962 | 424  | 
using assms  | 
425  | 
proof (induct e)  | 
|
| 67123 | 426  | 
case (FCnst c)  | 
427  | 
  {
 | 
|
| 64962 | 428  | 
case 1  | 
429  | 
show ?case by simp  | 
|
430  | 
next  | 
|
431  | 
case 2  | 
|
432  | 
show ?case by simp  | 
|
433  | 
}  | 
|
434  | 
next  | 
|
| 67123 | 435  | 
case (FVar n)  | 
436  | 
  {
 | 
|
| 64962 | 437  | 
case 1  | 
438  | 
then show ?case by simp  | 
|
439  | 
next  | 
|
440  | 
case 2  | 
|
441  | 
show ?case by simp  | 
|
442  | 
}  | 
|
443  | 
next  | 
|
444  | 
case (FAdd e1 e2)  | 
|
| 67123 | 445  | 
note split = split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"]  | 
| 64962 | 446  | 
  {
 | 
447  | 
case 1  | 
|
448  | 
let ?left = "peval xs (Left (Denom e1) (Denom e2))"  | 
|
449  | 
let ?common = "peval xs (Common (Denom e1) (Denom e2))"  | 
|
450  | 
let ?right = "peval xs (Right (Denom e1) (Denom e2))"  | 
|
| 67123 | 451  | 
from 1 FAdd have "feval xs (FAdd e1 e2) =  | 
| 64962 | 452  | 
(?common \<otimes> (peval xs (Num e1) \<otimes> ?right \<oplus> peval xs (Num e2) \<otimes> ?left)) \<oslash>  | 
453  | 
(?common \<otimes> (?left \<otimes> (?right \<otimes> ?common)))"  | 
|
| 67123 | 454  | 
by (simp add: split_beta split nonzero_union add_frac_eq r_distr m_ac)  | 
455  | 
also from 1 FAdd have "\<dots> = peval xs (Num (FAdd e1 e2)) \<oslash> peval xs (Denom (FAdd e1 e2))"  | 
|
| 64962 | 456  | 
by (simp add: split_beta split nonzero_union npeadd_correct npemul_correct integral_iff)  | 
457  | 
finally show ?case .  | 
|
458  | 
next  | 
|
459  | 
case 2  | 
|
460  | 
with FAdd show ?case  | 
|
461  | 
by (simp add: split_beta split nonzero_union npemul_correct integral_iff)  | 
|
462  | 
}  | 
|
463  | 
next  | 
|
464  | 
case (FSub e1 e2)  | 
|
| 67123 | 465  | 
note split = split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"]  | 
| 64962 | 466  | 
  {
 | 
467  | 
case 1  | 
|
468  | 
let ?left = "peval xs (Left (Denom e1) (Denom e2))"  | 
|
469  | 
let ?common = "peval xs (Common (Denom e1) (Denom e2))"  | 
|
470  | 
let ?right = "peval xs (Right (Denom e1) (Denom e2))"  | 
|
471  | 
from 1 FSub  | 
|
472  | 
have "feval xs (FSub e1 e2) =  | 
|
473  | 
(?common \<otimes> (peval xs (Num e1) \<otimes> ?right \<ominus> peval xs (Num e2) \<otimes> ?left)) \<oslash>  | 
|
474  | 
(?common \<otimes> (?left \<otimes> (?right \<otimes> ?common)))"  | 
|
| 67123 | 475  | 
by (simp add: split_beta split nonzero_union diff_frac_eq r_diff_distr m_ac)  | 
476  | 
also from 1 FSub have "\<dots> = peval xs (Num (FSub e1 e2)) \<oslash> peval xs (Denom (FSub e1 e2))"  | 
|
| 64962 | 477  | 
by (simp add: split_beta split nonzero_union npesub_correct npemul_correct integral_iff)  | 
478  | 
finally show ?case .  | 
|
479  | 
next  | 
|
480  | 
case 2  | 
|
481  | 
with FSub show ?case  | 
|
482  | 
by (simp add: split_beta split nonzero_union npemul_correct integral_iff)  | 
|
483  | 
}  | 
|
484  | 
next  | 
|
485  | 
case (FMul e1 e2)  | 
|
486  | 
note split =  | 
|
| 67123 | 487  | 
split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e1" and e\<^sub>2 = "Denom e2"]  | 
488  | 
split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e2" and e\<^sub>2 = "Denom e1"]  | 
|
| 64962 | 489  | 
  {
 | 
490  | 
case 1  | 
|
491  | 
let ?left\<^sub>1 = "peval xs (Left (Num e1) (Denom e2))"  | 
|
492  | 
let ?common\<^sub>1 = "peval xs (Common (Num e1) (Denom e2))"  | 
|
493  | 
let ?right\<^sub>1 = "peval xs (Right (Num e1) (Denom e2))"  | 
|
494  | 
let ?left\<^sub>2 = "peval xs (Left (Num e2) (Denom e1))"  | 
|
495  | 
let ?common\<^sub>2 = "peval xs (Common (Num e2) (Denom e1))"  | 
|
496  | 
let ?right\<^sub>2 = "peval xs (Right (Num e2) (Denom e1))"  | 
|
| 67123 | 497  | 
from 1 FMul have "feval xs (FMul e1 e2) =  | 
| 64962 | 498  | 
((?common\<^sub>1 \<otimes> ?common\<^sub>2) \<otimes> (?left\<^sub>1 \<otimes> ?left\<^sub>2)) \<oslash>  | 
499  | 
((?common\<^sub>1 \<otimes> ?common\<^sub>2) \<otimes> (?right\<^sub>2 \<otimes> ?right\<^sub>1))"  | 
|
500  | 
by (simp add: split_beta split nonzero_union  | 
|
501  | 
nonzero_divide_divide_eq_left m_ac)  | 
|
| 67123 | 502  | 
also from 1 FMul have "\<dots> = peval xs (Num (FMul e1 e2)) \<oslash> peval xs (Denom (FMul e1 e2))"  | 
| 64962 | 503  | 
by (simp add: split_beta split nonzero_union npemul_correct integral_iff)  | 
504  | 
finally show ?case .  | 
|
505  | 
next  | 
|
506  | 
case 2  | 
|
507  | 
with FMul show ?case  | 
|
508  | 
by (simp add: split_beta split nonzero_union npemul_correct integral_iff)  | 
|
509  | 
}  | 
|
510  | 
next  | 
|
511  | 
case (FNeg e)  | 
|
512  | 
  {
 | 
|
513  | 
case 1  | 
|
514  | 
with FNeg show ?case  | 
|
515  | 
by (simp add: split_beta npeneg_correct)  | 
|
516  | 
next  | 
|
517  | 
case 2  | 
|
518  | 
with FNeg show ?case  | 
|
519  | 
by (simp add: split_beta)  | 
|
520  | 
}  | 
|
521  | 
next  | 
|
522  | 
case (FDiv e1 e2)  | 
|
523  | 
note split =  | 
|
| 67123 | 524  | 
split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e1" and e\<^sub>2 = "Num e2"]  | 
525  | 
split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"]  | 
|
| 64962 | 526  | 
  {
 | 
527  | 
case 1  | 
|
528  | 
let ?left\<^sub>1 = "peval xs (Left (Num e1) (Num e2))"  | 
|
529  | 
let ?common\<^sub>1 = "peval xs (Common (Num e1) (Num e2))"  | 
|
530  | 
let ?right\<^sub>1 = "peval xs (Right (Num e1) (Num e2))"  | 
|
531  | 
let ?left\<^sub>2 = "peval xs (Left (Denom e1) (Denom e2))"  | 
|
532  | 
let ?common\<^sub>2 = "peval xs (Common (Denom e1) (Denom e2))"  | 
|
533  | 
let ?right\<^sub>2 = "peval xs (Right (Denom e1) (Denom e2))"  | 
|
534  | 
from 1 FDiv  | 
|
535  | 
have "feval xs (FDiv e1 e2) =  | 
|
536  | 
((?common\<^sub>1 \<otimes> ?common\<^sub>2) \<otimes> (?left\<^sub>1 \<otimes> ?right\<^sub>2)) \<oslash>  | 
|
537  | 
((?common\<^sub>1 \<otimes> ?common\<^sub>2) \<otimes> (?left\<^sub>2 \<otimes> ?right\<^sub>1))"  | 
|
538  | 
by (simp add: split_beta split nonzero_union nonzero_insert  | 
|
539  | 
nonzero_divide_divide_eq m_ac)  | 
|
| 67123 | 540  | 
also from 1 FDiv have "\<dots> = peval xs (Num (FDiv e1 e2)) \<oslash> peval xs (Denom (FDiv e1 e2))"  | 
| 64962 | 541  | 
by (simp add: split_beta split nonzero_union nonzero_insert npemul_correct integral_iff)  | 
542  | 
finally show ?case .  | 
|
543  | 
next  | 
|
544  | 
case 2  | 
|
545  | 
with FDiv show ?case  | 
|
546  | 
by (simp add: split_beta split nonzero_union nonzero_insert npemul_correct integral_iff)  | 
|
547  | 
}  | 
|
548  | 
next  | 
|
549  | 
case (FPow e n)  | 
|
550  | 
  {
 | 
|
551  | 
case 1  | 
|
552  | 
with FPow show ?case  | 
|
553  | 
by (simp add: split_beta nonzero_power_divide npepow_correct)  | 
|
554  | 
next  | 
|
555  | 
case 2  | 
|
556  | 
with FPow show ?case  | 
|
557  | 
by (simp add: split_beta npepow_correct)  | 
|
558  | 
}  | 
|
559  | 
qed  | 
|
560  | 
||
561  | 
lemma (in field) feval_eq0:  | 
|
562  | 
assumes "in_carrier xs"  | 
|
| 67123 | 563  | 
and "fnorm e = (n, d, c)"  | 
564  | 
and "nonzero xs c"  | 
|
565  | 
and "peval xs n = \<zero>"  | 
|
| 64962 | 566  | 
shows "feval xs e = \<zero>"  | 
| 67123 | 567  | 
using assms fnorm_correct [of xs e] by simp  | 
| 64962 | 568  | 
|
569  | 
lemma (in field) fexpr_in_carrier:  | 
|
570  | 
assumes "in_carrier xs"  | 
|
| 67123 | 571  | 
and "nonzero xs (Cond e)"  | 
| 64962 | 572  | 
shows "feval xs e \<in> carrier R"  | 
573  | 
using assms  | 
|
574  | 
proof (induct e)  | 
|
575  | 
case (FDiv e1 e2)  | 
|
576  | 
then have "feval xs e1 \<in> carrier R" "feval xs e2 \<in> carrier R"  | 
|
577  | 
"peval xs (Num e2) \<noteq> \<zero>" "nonzero xs (Cond e2)"  | 
|
578  | 
by (simp_all add: nonzero_union nonzero_insert split: prod.split_asm)  | 
|
| 64998 | 579  | 
from \<open>in_carrier xs\<close> \<open>nonzero xs (Cond e2)\<close>  | 
| 64962 | 580  | 
have "feval xs e2 = peval xs (Num e2) \<oslash> peval xs (Denom e2)"  | 
581  | 
by (rule fnorm_correct)  | 
|
| 64998 | 582  | 
moreover from \<open>in_carrier xs\<close> \<open>nonzero xs (Cond e2)\<close>  | 
| 64962 | 583  | 
have "peval xs (Denom e2) \<noteq> \<zero>" by (rule fnorm_correct)  | 
| 64998 | 584  | 
ultimately have "feval xs e2 \<noteq> \<zero>" using \<open>peval xs (Num e2) \<noteq> \<zero>\<close> \<open>in_carrier xs\<close>  | 
| 64962 | 585  | 
by (simp add: divide_eq_0_iff)  | 
| 64998 | 586  | 
with \<open>feval xs e1 \<in> carrier R\<close> \<open>feval xs e2 \<in> carrier R\<close>  | 
| 64962 | 587  | 
show ?case by simp  | 
588  | 
qed (simp_all add: nonzero_union split: prod.split_asm)  | 
|
589  | 
||
590  | 
lemma (in field) feval_eq:  | 
|
591  | 
assumes "in_carrier xs"  | 
|
| 67123 | 592  | 
and "fnorm (FSub e e') = (n, d, c)"  | 
593  | 
and "nonzero xs c"  | 
|
| 64962 | 594  | 
shows "(feval xs e = feval xs e') = (peval xs n = \<zero>)"  | 
595  | 
proof -  | 
|
596  | 
from assms have "nonzero xs (Cond e)" "nonzero xs (Cond e')"  | 
|
597  | 
by (auto simp add: nonzero_union split: prod.split_asm)  | 
|
598  | 
with assms fnorm_correct [of xs "FSub e e'"]  | 
|
599  | 
have "feval xs e \<ominus> feval xs e' = peval xs n \<oslash> peval xs d"  | 
|
600  | 
"peval xs d \<noteq> \<zero>"  | 
|
601  | 
by simp_all  | 
|
602  | 
show ?thesis  | 
|
603  | 
proof  | 
|
604  | 
assume "feval xs e = feval xs e'"  | 
|
| 64998 | 605  | 
with \<open>feval xs e \<ominus> feval xs e' = peval xs n \<oslash> peval xs d\<close>  | 
606  | 
\<open>in_carrier xs\<close> \<open>nonzero xs (Cond e')\<close>  | 
|
| 64962 | 607  | 
have "peval xs n \<oslash> peval xs d = \<zero>"  | 
608  | 
by (simp add: fexpr_in_carrier minus_eq r_neg)  | 
|
| 64998 | 609  | 
with \<open>peval xs d \<noteq> \<zero>\<close> \<open>in_carrier xs\<close>  | 
| 64962 | 610  | 
show "peval xs n = \<zero>"  | 
611  | 
by (simp add: divide_eq_0_iff)  | 
|
612  | 
next  | 
|
613  | 
assume "peval xs n = \<zero>"  | 
|
| 64998 | 614  | 
with \<open>feval xs e \<ominus> feval xs e' = peval xs n \<oslash> peval xs d\<close> \<open>peval xs d \<noteq> \<zero>\<close>  | 
615  | 
\<open>nonzero xs (Cond e)\<close> \<open>nonzero xs (Cond e')\<close> \<open>in_carrier xs\<close>  | 
|
| 64962 | 616  | 
show "feval xs e = feval xs e'"  | 
617  | 
by (simp add: eq_diff0 fexpr_in_carrier)  | 
|
618  | 
qed  | 
|
619  | 
qed  | 
|
620  | 
||
| 64999 | 621  | 
ML \<open>  | 
| 69597 | 622  | 
val term_of_nat = HOLogic.mk_number \<^typ>\<open>nat\<close> o @{code integer_of_nat};
 | 
| 64999 | 623  | 
|
| 69597 | 624  | 
val term_of_int = HOLogic.mk_number \<^typ>\<open>int\<close> o @{code integer_of_int};
 | 
| 64999 | 625  | 
|
| 69597 | 626  | 
fun term_of_pexpr (@{code PExpr1} x) = \<^term>\<open>PExpr1\<close> $ term_of_pexpr1 x
 | 
627  | 
  | term_of_pexpr (@{code PExpr2} x) = \<^term>\<open>PExpr2\<close> $ term_of_pexpr2 x
 | 
|
628  | 
and term_of_pexpr1 (@{code PCnst} k) = \<^term>\<open>PCnst\<close> $ term_of_int k
 | 
|
629  | 
  | term_of_pexpr1 (@{code PVar} n) = \<^term>\<open>PVar\<close> $ term_of_nat n
 | 
|
630  | 
  | term_of_pexpr1 (@{code PAdd} (x, y)) = \<^term>\<open>PAdd\<close> $ term_of_pexpr x $ term_of_pexpr y
 | 
|
631  | 
  | term_of_pexpr1 (@{code PSub} (x, y)) = \<^term>\<open>PSub\<close> $ term_of_pexpr x $ term_of_pexpr y
 | 
|
632  | 
  | term_of_pexpr1 (@{code PNeg} x) = \<^term>\<open>PNeg\<close> $ term_of_pexpr x
 | 
|
633  | 
and term_of_pexpr2 (@{code PMul} (x, y)) = \<^term>\<open>PMul\<close> $ term_of_pexpr x $ term_of_pexpr y
 | 
|
634  | 
  | term_of_pexpr2 (@{code PPow} (x, n)) = \<^term>\<open>PPow\<close> $ term_of_pexpr x $ term_of_nat n
 | 
|
| 64999 | 635  | 
|
636  | 
fun term_of_result (x, (y, zs)) =  | 
|
637  | 
HOLogic.mk_prod (term_of_pexpr x, HOLogic.mk_prod  | 
|
| 69597 | 638  | 
(term_of_pexpr y, HOLogic.mk_list \<^typ>\<open>pexpr\<close> (map term_of_pexpr zs)));  | 
| 64962 | 639  | 
|
| 64999 | 640  | 
local  | 
641  | 
||
| 69597 | 642  | 
fun fnorm (ctxt, ct, t) = Thm.mk_binop \<^cterm>\<open>Pure.eq :: pexpr \<times> pexpr \<times> pexpr list \<Rightarrow> pexpr \<times> pexpr \<times> pexpr list \<Rightarrow> prop\<close>  | 
| 64999 | 643  | 
ct (Thm.cterm_of ctxt t);  | 
644  | 
||
645  | 
val (_, raw_fnorm_oracle) = Context.>>> (Context.map_theory_result  | 
|
| 69597 | 646  | 
(Thm.add_oracle (\<^binding>\<open>fnorm\<close>, fnorm)));  | 
| 64999 | 647  | 
|
648  | 
fun fnorm_oracle ctxt ct t = raw_fnorm_oracle (ctxt, ct, t);  | 
|
649  | 
||
650  | 
in  | 
|
651  | 
||
652  | 
val cv = @{computation_conv "pexpr \<times> pexpr \<times> pexpr list"
 | 
|
653  | 
terms: fnorm nat_of_integer Code_Target_Nat.natural  | 
|
654  | 
"0::nat" "1::nat" "2::nat" "3::nat"  | 
|
655  | 
"0::int" "1::int" "2::int" "3::int" "-1::int"  | 
|
656  | 
datatypes: fexpr int integer num}  | 
|
| 
65043
 
fd753468786f
explicit dynamic context for gap-bridging function
 
haftmann 
parents: 
64999 
diff
changeset
 | 
657  | 
(fn ctxt => fn result => fn ct => fnorm_oracle ctxt ct (term_of_result result))  | 
| 64999 | 658  | 
|
659  | 
end  | 
|
660  | 
\<close>  | 
|
| 64962 | 661  | 
|
| 64998 | 662  | 
ML \<open>  | 
| 64962 | 663  | 
signature FIELD_TAC =  | 
664  | 
sig  | 
|
665  | 
structure Field_Simps:  | 
|
666  | 
sig  | 
|
667  | 
type T  | 
|
668  | 
val get: Context.generic -> T  | 
|
669  | 
val put: T -> Context.generic -> Context.generic  | 
|
670  | 
val map: (T -> T) -> Context.generic -> Context.generic  | 
|
671  | 
end  | 
|
672  | 
val eq_field_simps:  | 
|
673  | 
(term * (thm list * thm list * thm list * thm * thm)) *  | 
|
674  | 
(term * (thm list * thm list * thm list * thm * thm)) -> bool  | 
|
675  | 
val field_tac: bool -> Proof.context -> int -> tactic  | 
|
676  | 
end  | 
|
677  | 
||
678  | 
structure Field_Tac : FIELD_TAC =  | 
|
679  | 
struct  | 
|
680  | 
||
681  | 
open Ring_Tac;  | 
|
682  | 
||
| 69597 | 683  | 
fun field_struct (Const (\<^const_name>\<open>Ring.ring.add\<close>, _) $ R $ _ $ _) = SOME R  | 
684  | 
| field_struct (Const (\<^const_name>\<open>Ring.a_minus\<close>, _) $ R $ _ $ _) = SOME R  | 
|
685  | 
| field_struct (Const (\<^const_name>\<open>Group.monoid.mult\<close>, _) $ R $ _ $ _) = SOME R  | 
|
686  | 
| field_struct (Const (\<^const_name>\<open>Ring.a_inv\<close>, _) $ R $ _) = SOME R  | 
|
687  | 
| field_struct (Const (\<^const_name>\<open>Group.pow\<close>, _) $ R $ _ $ _) = SOME R  | 
|
688  | 
| field_struct (Const (\<^const_name>\<open>Algebra_Aux.m_div\<close>, _) $ R $ _ $ _) = SOME R  | 
|
689  | 
| field_struct (Const (\<^const_name>\<open>Ring.ring.zero\<close>, _) $ R) = SOME R  | 
|
690  | 
| field_struct (Const (\<^const_name>\<open>Group.monoid.one\<close>, _) $ R) = SOME R  | 
|
691  | 
| field_struct (Const (\<^const_name>\<open>Algebra_Aux.of_integer\<close>, _) $ R $ _) = SOME R  | 
|
| 64962 | 692  | 
| field_struct _ = NONE;  | 
693  | 
||
| 69597 | 694  | 
fun reif_fexpr vs (Const (\<^const_name>\<open>Ring.ring.add\<close>, _) $ _ $ a $ b) =  | 
695  | 
\<^const>\<open>FAdd\<close> $ reif_fexpr vs a $ reif_fexpr vs b  | 
|
696  | 
| reif_fexpr vs (Const (\<^const_name>\<open>Ring.a_minus\<close>, _) $ _ $ a $ b) =  | 
|
697  | 
\<^const>\<open>FSub\<close> $ reif_fexpr vs a $ reif_fexpr vs b  | 
|
698  | 
| reif_fexpr vs (Const (\<^const_name>\<open>Group.monoid.mult\<close>, _) $ _ $ a $ b) =  | 
|
699  | 
\<^const>\<open>FMul\<close> $ reif_fexpr vs a $ reif_fexpr vs b  | 
|
700  | 
| reif_fexpr vs (Const (\<^const_name>\<open>Ring.a_inv\<close>, _) $ _ $ a) =  | 
|
701  | 
\<^const>\<open>FNeg\<close> $ reif_fexpr vs a  | 
|
702  | 
| reif_fexpr vs (Const (\<^const_name>\<open>Group.pow\<close>, _) $ _ $ a $ n) =  | 
|
703  | 
\<^const>\<open>FPow\<close> $ reif_fexpr vs a $ n  | 
|
704  | 
| reif_fexpr vs (Const (\<^const_name>\<open>Algebra_Aux.m_div\<close>, _) $ _ $ a $ b) =  | 
|
705  | 
\<^const>\<open>FDiv\<close> $ reif_fexpr vs a $ reif_fexpr vs b  | 
|
| 64962 | 706  | 
| reif_fexpr vs (Free x) =  | 
| 69597 | 707  | 
\<^const>\<open>FVar\<close> $ HOLogic.mk_number HOLogic.natT (find_index (equal x) vs)  | 
708  | 
| reif_fexpr vs (Const (\<^const_name>\<open>Ring.ring.zero\<close>, _) $ _) =  | 
|
709  | 
\<^term>\<open>FCnst 0\<close>  | 
|
710  | 
| reif_fexpr vs (Const (\<^const_name>\<open>Group.monoid.one\<close>, _) $ _) =  | 
|
711  | 
\<^term>\<open>FCnst 1\<close>  | 
|
712  | 
| reif_fexpr vs (Const (\<^const_name>\<open>Algebra_Aux.of_integer\<close>, _) $ _ $ n) =  | 
|
713  | 
\<^const>\<open>FCnst\<close> $ n  | 
|
| 64962 | 714  | 
| reif_fexpr _ _ = error "reif_fexpr: bad expression";  | 
715  | 
||
| 69597 | 716  | 
fun reif_fexpr' vs (Const (\<^const_name>\<open>Groups.plus\<close>, _) $ a $ b) =  | 
717  | 
\<^const>\<open>FAdd\<close> $ reif_fexpr' vs a $ reif_fexpr' vs b  | 
|
718  | 
| reif_fexpr' vs (Const (\<^const_name>\<open>Groups.minus\<close>, _) $ a $ b) =  | 
|
719  | 
\<^const>\<open>FSub\<close> $ reif_fexpr' vs a $ reif_fexpr' vs b  | 
|
720  | 
| reif_fexpr' vs (Const (\<^const_name>\<open>Groups.times\<close>, _) $ a $ b) =  | 
|
721  | 
\<^const>\<open>FMul\<close> $ reif_fexpr' vs a $ reif_fexpr' vs b  | 
|
722  | 
| reif_fexpr' vs (Const (\<^const_name>\<open>Groups.uminus\<close>, _) $ a) =  | 
|
723  | 
\<^const>\<open>FNeg\<close> $ reif_fexpr' vs a  | 
|
724  | 
| reif_fexpr' vs (Const (\<^const_name>\<open>Power.power\<close>, _) $ a $ n) =  | 
|
725  | 
\<^const>\<open>FPow\<close> $ reif_fexpr' vs a $ n  | 
|
726  | 
| reif_fexpr' vs (Const (\<^const_name>\<open>divide\<close>, _) $ a $ b) =  | 
|
727  | 
\<^const>\<open>FDiv\<close> $ reif_fexpr' vs a $ reif_fexpr' vs b  | 
|
| 64962 | 728  | 
| reif_fexpr' vs (Free x) =  | 
| 69597 | 729  | 
\<^const>\<open>FVar\<close> $ HOLogic.mk_number HOLogic.natT (find_index (equal x) vs)  | 
730  | 
| reif_fexpr' vs (Const (\<^const_name>\<open>zero_class.zero\<close>, _)) =  | 
|
731  | 
\<^term>\<open>FCnst 0\<close>  | 
|
732  | 
| reif_fexpr' vs (Const (\<^const_name>\<open>one_class.one\<close>, _)) =  | 
|
733  | 
\<^term>\<open>FCnst 1\<close>  | 
|
734  | 
| reif_fexpr' vs (Const (\<^const_name>\<open>numeral\<close>, _) $ b) =  | 
|
735  | 
      \<^const>\<open>FCnst\<close> $ (@{const numeral (int)} $ b)
 | 
|
| 64962 | 736  | 
| reif_fexpr' _ _ = error "reif_fexpr: bad expression";  | 
737  | 
||
738  | 
fun eq_field_simps  | 
|
739  | 
((t, (ths1, ths2, ths3, th4, th)),  | 
|
740  | 
(t', (ths1', ths2', ths3', th4', th'))) =  | 
|
741  | 
t aconv t' andalso  | 
|
742  | 
eq_list Thm.eq_thm (ths1, ths1') andalso  | 
|
743  | 
eq_list Thm.eq_thm (ths2, ths2') andalso  | 
|
744  | 
eq_list Thm.eq_thm (ths3, ths3') andalso  | 
|
745  | 
Thm.eq_thm (th4, th4') andalso  | 
|
746  | 
Thm.eq_thm (th, th');  | 
|
747  | 
||
748  | 
structure Field_Simps = Generic_Data  | 
|
749  | 
(struct  | 
|
750  | 
type T = (term * (thm list * thm list * thm list * thm * thm)) Net.net  | 
|
751  | 
val empty = Net.empty  | 
|
752  | 
val extend = I  | 
|
753  | 
val merge = Net.merge eq_field_simps  | 
|
754  | 
end);  | 
|
755  | 
||
756  | 
fun get_field_simps ctxt optcT t =  | 
|
757  | 
(case get_matching_rules ctxt (Field_Simps.get (Context.Proof ctxt)) t of  | 
|
758  | 
SOME (ths1, ths2, ths3, th4, th) =>  | 
|
759  | 
let val tr =  | 
|
| 67649 | 760  | 
Thm.transfer' ctxt #>  | 
| 64962 | 761  | 
(case optcT of NONE => I | SOME cT => inst [cT] [] #> norm)  | 
762  | 
in (map tr ths1, map tr ths2, map tr ths3, tr th4, tr th) end  | 
|
763  | 
| NONE => error "get_field_simps: lookup failed");  | 
|
764  | 
||
765  | 
fun nth_el_conv (_, _, _, nth_el_Cons, _) =  | 
|
766  | 
let  | 
|
767  | 
val a = type_of_eqn nth_el_Cons;  | 
|
768  | 
val If_conv_a = If_conv a;  | 
|
769  | 
||
770  | 
fun conv ys n = (case strip_app ys of  | 
|
| 69597 | 771  | 
(\<^const_name>\<open>Cons\<close>, [x, xs]) =>  | 
| 64962 | 772  | 
transitive'  | 
773  | 
(inst [] [x, xs, n] nth_el_Cons)  | 
|
774  | 
(If_conv_a (args2 nat_eq_conv)  | 
|
775  | 
Thm.reflexive  | 
|
776  | 
(cong2' conv Thm.reflexive (args2 nat_minus_conv))))  | 
|
777  | 
in conv end;  | 
|
778  | 
||
779  | 
fun feval_conv (rls as  | 
|
780  | 
([feval_simps_1, feval_simps_2, feval_simps_3,  | 
|
781  | 
feval_simps_4, feval_simps_5, feval_simps_6,  | 
|
782  | 
feval_simps_7, feval_simps_8, feval_simps_9,  | 
|
783  | 
feval_simps_10, feval_simps_11],  | 
|
784  | 
_, _, _, _)) =  | 
|
785  | 
let  | 
|
786  | 
val nth_el_conv' = nth_el_conv rls;  | 
|
787  | 
||
788  | 
fun conv xs x = (case strip_app x of  | 
|
| 69597 | 789  | 
(\<^const_name>\<open>FCnst\<close>, [c]) => (case strip_app c of  | 
790  | 
(\<^const_name>\<open>zero_class.zero\<close>, _) => inst [] [xs] feval_simps_9  | 
|
791  | 
| (\<^const_name>\<open>one_class.one\<close>, _) => inst [] [xs] feval_simps_10  | 
|
792  | 
| (\<^const_name>\<open>numeral\<close>, [n]) => inst [] [xs, n] feval_simps_11  | 
|
| 64962 | 793  | 
| _ => inst [] [xs, c] feval_simps_1)  | 
| 69597 | 794  | 
| (\<^const_name>\<open>FVar\<close>, [n]) =>  | 
| 64962 | 795  | 
transitive' (inst [] [xs, n] feval_simps_2) (args2 nth_el_conv')  | 
| 69597 | 796  | 
| (\<^const_name>\<open>FAdd\<close>, [a, b]) =>  | 
| 64962 | 797  | 
transitive' (inst [] [xs, a, b] feval_simps_3)  | 
798  | 
(cong2 (args2 conv) (args2 conv))  | 
|
| 69597 | 799  | 
| (\<^const_name>\<open>FSub\<close>, [a, b]) =>  | 
| 64962 | 800  | 
transitive' (inst [] [xs, a, b] feval_simps_4)  | 
801  | 
(cong2 (args2 conv) (args2 conv))  | 
|
| 69597 | 802  | 
| (\<^const_name>\<open>FMul\<close>, [a, b]) =>  | 
| 64962 | 803  | 
transitive' (inst [] [xs, a, b] feval_simps_5)  | 
804  | 
(cong2 (args2 conv) (args2 conv))  | 
|
| 69597 | 805  | 
| (\<^const_name>\<open>FNeg\<close>, [a]) =>  | 
| 64962 | 806  | 
transitive' (inst [] [xs, a] feval_simps_6)  | 
807  | 
(cong1 (args2 conv))  | 
|
| 69597 | 808  | 
| (\<^const_name>\<open>FDiv\<close>, [a, b]) =>  | 
| 64962 | 809  | 
transitive' (inst [] [xs, a, b] feval_simps_7)  | 
810  | 
(cong2 (args2 conv) (args2 conv))  | 
|
| 69597 | 811  | 
| (\<^const_name>\<open>FPow\<close>, [a, n]) =>  | 
| 64962 | 812  | 
transitive' (inst [] [xs, a, n] feval_simps_8)  | 
813  | 
(cong2 (args2 conv) Thm.reflexive))  | 
|
814  | 
in conv end;  | 
|
815  | 
||
816  | 
fun peval_conv (rls as  | 
|
817  | 
(_,  | 
|
818  | 
[peval_simps_1, peval_simps_2, peval_simps_3,  | 
|
819  | 
peval_simps_4, peval_simps_5, peval_simps_6,  | 
|
820  | 
peval_simps_7, peval_simps_8, peval_simps_9,  | 
|
821  | 
peval_simps_10, peval_simps_11],  | 
|
822  | 
_, _, _)) =  | 
|
823  | 
let  | 
|
824  | 
val nth_el_conv' = nth_el_conv rls;  | 
|
825  | 
||
826  | 
fun conv xs x = (case strip_app x of  | 
|
| 69597 | 827  | 
(\<^const_name>\<open>PExpr1\<close>, [e]) => (case strip_app e of  | 
828  | 
(\<^const_name>\<open>PCnst\<close>, [c]) => (case strip_numeral c of  | 
|
829  | 
(\<^const_name>\<open>zero_class.zero\<close>, _) => inst [] [xs] peval_simps_8  | 
|
830  | 
| (\<^const_name>\<open>one_class.one\<close>, _) => inst [] [xs] peval_simps_9  | 
|
831  | 
| (\<^const_name>\<open>numeral\<close>, [n]) => inst [] [xs, n] peval_simps_10  | 
|
832  | 
| (\<^const_name>\<open>uminus\<close>, [n]) => inst [] [xs, n] peval_simps_11  | 
|
| 64962 | 833  | 
| _ => inst [] [xs, c] peval_simps_1)  | 
| 69597 | 834  | 
| (\<^const_name>\<open>PVar\<close>, [n]) =>  | 
| 64962 | 835  | 
transitive' (inst [] [xs, n] peval_simps_2) (args2 nth_el_conv')  | 
| 69597 | 836  | 
| (\<^const_name>\<open>PAdd\<close>, [a, b]) =>  | 
| 64962 | 837  | 
transitive' (inst [] [xs, a, b] peval_simps_3)  | 
838  | 
(cong2 (args2 conv) (args2 conv))  | 
|
| 69597 | 839  | 
| (\<^const_name>\<open>PSub\<close>, [a, b]) =>  | 
| 64962 | 840  | 
transitive' (inst [] [xs, a, b] peval_simps_4)  | 
841  | 
(cong2 (args2 conv) (args2 conv))  | 
|
| 69597 | 842  | 
| (\<^const_name>\<open>PNeg\<close>, [a]) =>  | 
| 64962 | 843  | 
transitive' (inst [] [xs, a] peval_simps_5)  | 
844  | 
(cong1 (args2 conv)))  | 
|
| 69597 | 845  | 
| (\<^const_name>\<open>PExpr2\<close>, [e]) => (case strip_app e of  | 
846  | 
(\<^const_name>\<open>PMul\<close>, [a, b]) =>  | 
|
| 64962 | 847  | 
transitive' (inst [] [xs, a, b] peval_simps_6)  | 
848  | 
(cong2 (args2 conv) (args2 conv))  | 
|
| 69597 | 849  | 
| (\<^const_name>\<open>PPow\<close>, [a, n]) =>  | 
| 64962 | 850  | 
transitive' (inst [] [xs, a, n] peval_simps_7)  | 
851  | 
(cong2 (args2 conv) Thm.reflexive)))  | 
|
852  | 
in conv end;  | 
|
853  | 
||
854  | 
fun nonzero_conv (rls as  | 
|
855  | 
(_, _,  | 
|
856  | 
[nonzero_Nil, nonzero_Cons, nonzero_singleton],  | 
|
857  | 
_, _)) =  | 
|
858  | 
let  | 
|
859  | 
val peval_conv' = peval_conv rls;  | 
|
860  | 
||
861  | 
fun conv xs qs = (case strip_app qs of  | 
|
| 69597 | 862  | 
(\<^const_name>\<open>Nil\<close>, []) => inst [] [xs] nonzero_Nil  | 
863  | 
| (\<^const_name>\<open>Cons\<close>, [p, ps]) => (case Thm.term_of ps of  | 
|
864  | 
Const (\<^const_name>\<open>Nil\<close>, _) =>  | 
|
| 64962 | 865  | 
transitive' (inst [] [xs, p] nonzero_singleton)  | 
866  | 
(cong1 (cong2 (args2 peval_conv') Thm.reflexive))  | 
|
867  | 
| _ => transitive' (inst [] [xs, p, ps] nonzero_Cons)  | 
|
868  | 
(cong2 (cong1 (cong2 (args2 peval_conv') Thm.reflexive)) (args2 conv))))  | 
|
869  | 
in conv end;  | 
|
870  | 
||
871  | 
fun field_tac in_prem ctxt =  | 
|
872  | 
SUBGOAL (fn (g, i) =>  | 
|
873  | 
let  | 
|
874  | 
val (prems, concl) = Logic.strip_horn g;  | 
|
875  | 
fun find_eq s = (case s of  | 
|
| 69597 | 876  | 
(_ $ (Const (\<^const_name>\<open>HOL.eq\<close>, Type (_, [T, _])) $ t $ u)) =>  | 
| 64962 | 877  | 
(case (field_struct t, field_struct u) of  | 
878  | 
(SOME R, _) => SOME ((t, u), R, T, NONE, mk_in_carrier ctxt R [], reif_fexpr)  | 
|
879  | 
| (_, SOME R) => SOME ((t, u), R, T, NONE, mk_in_carrier ctxt R [], reif_fexpr)  | 
|
880  | 
| _ =>  | 
|
| 69597 | 881  | 
if Sign.of_sort (Proof_Context.theory_of ctxt) (T, \<^sort>\<open>field\<close>)  | 
| 64962 | 882  | 
                 then SOME ((t, u), mk_ring T, T, SOME T, K @{thm in_carrier_trivial}, reif_fexpr')
 | 
883  | 
else NONE)  | 
|
884  | 
| _ => NONE);  | 
|
885  | 
val ((t, u), R, T, optT, mkic, reif) =  | 
|
886  | 
(case get_first find_eq  | 
|
887  | 
(if in_prem then prems else [concl]) of  | 
|
888  | 
SOME q => q  | 
|
889  | 
| NONE => error "cannot determine field");  | 
|
890  | 
val rls as (_, _, _, _, feval_eq) =  | 
|
891  | 
get_field_simps ctxt (Option.map (Thm.ctyp_of ctxt) optT) R;  | 
|
892  | 
val xs = [] |> Term.add_frees t |> Term.add_frees u |> filter (equal T o snd);  | 
|
893  | 
val cxs = Thm.cterm_of ctxt (HOLogic.mk_list T (map Free xs));  | 
|
894  | 
val ce = Thm.cterm_of ctxt (reif xs t);  | 
|
895  | 
val ce' = Thm.cterm_of ctxt (reif xs u);  | 
|
896  | 
val fnorm = cv ctxt  | 
|
| 69597 | 897  | 
(Thm.apply \<^cterm>\<open>fnorm\<close> (Thm.apply (Thm.apply \<^cterm>\<open>FSub\<close> ce) ce'));  | 
| 64962 | 898  | 
val (_, [n, dc]) = strip_app (Thm.rhs_of fnorm);  | 
899  | 
val (_, [_, c]) = strip_app dc;  | 
|
900  | 
val th =  | 
|
901  | 
Conv.fconv_rule (Conv.concl_conv 1 (Conv.arg_conv  | 
|
902  | 
(binop_conv  | 
|
903  | 
(binop_conv  | 
|
904  | 
(K (feval_conv rls cxs ce)) (K (feval_conv rls cxs ce')))  | 
|
905  | 
(Conv.arg1_conv (K (peval_conv rls cxs n))))))  | 
|
906  | 
([mkic xs,  | 
|
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
907  | 
HOLogic.mk_obj_eq fnorm,  | 
| 
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
908  | 
          HOLogic.mk_obj_eq (nonzero_conv rls cxs c) RS @{thm iffD2}] MRS
 | 
| 64962 | 909  | 
feval_eq);  | 
910  | 
val th' = Drule.rotate_prems 1  | 
|
911  | 
        (th RS (if in_prem then @{thm iffD1} else @{thm iffD2}));
 | 
|
912  | 
in  | 
|
913  | 
if in_prem then  | 
|
914  | 
dresolve_tac ctxt [th'] 1 THEN defer_tac 1  | 
|
915  | 
else  | 
|
916  | 
resolve_tac ctxt [th'] 1  | 
|
917  | 
end);  | 
|
918  | 
||
919  | 
end  | 
|
| 64998 | 920  | 
\<close>  | 
| 64962 | 921  | 
|
| 67123 | 922  | 
context field  | 
923  | 
begin  | 
|
| 64962 | 924  | 
|
| 64998 | 925  | 
local_setup \<open>  | 
| 64962 | 926  | 
Local_Theory.declaration {syntax = false, pervasive = false}
 | 
927  | 
(fn phi => Field_Tac.Field_Simps.map (Ring_Tac.insert_rules Field_Tac.eq_field_simps  | 
|
| 69597 | 928  | 
(Morphism.term phi \<^term>\<open>R\<close>,  | 
| 64962 | 929  | 
     (Morphism.fact phi @{thms feval.simps [meta] feval_Cnst [meta]},
 | 
930  | 
      Morphism.fact phi @{thms peval.simps [meta] peval_Cnst [meta]},
 | 
|
931  | 
      Morphism.fact phi @{thms nonzero.simps [meta] nonzero_singleton [meta]},
 | 
|
932  | 
      singleton (Morphism.fact phi) @{thm nth_el_Cons [meta]},
 | 
|
933  | 
      singleton (Morphism.fact phi) @{thm feval_eq}))))
 | 
|
| 64998 | 934  | 
\<close>  | 
| 64962 | 935  | 
|
936  | 
end  | 
|
937  | 
||
| 64998 | 938  | 
method_setup field = \<open>  | 
| 64962 | 939  | 
Scan.lift (Args.mode "prems") -- Attrib.thms >> (fn (in_prem, thms) => fn ctxt =>  | 
940  | 
SIMPLE_METHOD' (Field_Tac.field_tac in_prem ctxt THEN' Ring_Tac.ring_tac in_prem thms ctxt))  | 
|
| 64998 | 941  | 
\<close> "reduce equations over fields to equations over rings"  | 
| 64962 | 942  | 
|
943  | 
end  |