| author | huffman | 
| Sun, 28 Feb 2010 16:11:08 -0800 | |
| changeset 35477 | b972233773dd | 
| parent 35216 | 7641e8d831d2 | 
| child 35410 | 1ea89d2a1bd4 | 
| permissions | -rw-r--r-- | 
| 23252 | 1  | 
(* Title: HOL/Groebner_Basis.thy  | 
2  | 
Author: Amine Chaieb, TU Muenchen  | 
|
3  | 
*)  | 
|
4  | 
||
5  | 
header {* Semiring normalization and Groebner Bases *}
 | 
|
| 28402 | 6  | 
|
| 23252 | 7  | 
theory Groebner_Basis  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33296 
diff
changeset
 | 
8  | 
imports Numeral_Simprocs  | 
| 23252 | 9  | 
uses  | 
10  | 
"Tools/Groebner_Basis/misc.ML"  | 
|
11  | 
"Tools/Groebner_Basis/normalizer_data.ML"  | 
|
12  | 
  ("Tools/Groebner_Basis/normalizer.ML")
 | 
|
| 23312 | 13  | 
  ("Tools/Groebner_Basis/groebner.ML")
 | 
| 23252 | 14  | 
begin  | 
15  | 
||
16  | 
subsection {* Semiring normalization *}
 | 
|
17  | 
||
18  | 
setup NormalizerData.setup  | 
|
19  | 
||
20  | 
||
| 
23258
 
9062e98fdab1
renamed locale ring/semiring to gb_ring/gb_semiring to avoid clash with Ring_and_Field versions;
 
wenzelm 
parents: 
23252 
diff
changeset
 | 
21  | 
locale gb_semiring =  | 
| 23252 | 22  | 
fixes add mul pwr r0 r1  | 
23  | 
assumes add_a:"(add x (add y z) = add (add x y) z)"  | 
|
24  | 
and add_c: "add x y = add y x" and add_0:"add r0 x = x"  | 
|
25  | 
and mul_a:"mul x (mul y z) = mul (mul x y) z" and mul_c:"mul x y = mul y x"  | 
|
26  | 
and mul_1:"mul r1 x = x" and mul_0:"mul r0 x = r0"  | 
|
27  | 
and mul_d:"mul x (add y z) = add (mul x y) (mul x z)"  | 
|
28  | 
and pwr_0:"pwr x 0 = r1" and pwr_Suc:"pwr x (Suc n) = mul x (pwr x n)"  | 
|
29  | 
begin  | 
|
30  | 
||
31  | 
lemma mul_pwr:"mul (pwr x p) (pwr x q) = pwr x (p + q)"  | 
|
32  | 
proof (induct p)  | 
|
33  | 
case 0  | 
|
34  | 
then show ?case by (auto simp add: pwr_0 mul_1)  | 
|
35  | 
next  | 
|
36  | 
case Suc  | 
|
37  | 
from this [symmetric] show ?case  | 
|
38  | 
by (auto simp add: pwr_Suc mul_1 mul_a)  | 
|
39  | 
qed  | 
|
40  | 
||
41  | 
lemma pwr_mul: "pwr (mul x y) q = mul (pwr x q) (pwr y q)"  | 
|
42  | 
proof (induct q arbitrary: x y, auto simp add:pwr_0 pwr_Suc mul_1)  | 
|
43  | 
fix q x y  | 
|
44  | 
assume "\<And>x y. pwr (mul x y) q = mul (pwr x q) (pwr y q)"  | 
|
45  | 
have "mul (mul x y) (mul (pwr x q) (pwr y q)) = mul x (mul y (mul (pwr x q) (pwr y q)))"  | 
|
46  | 
by (simp add: mul_a)  | 
|
47  | 
also have "\<dots> = (mul (mul y (mul (pwr y q) (pwr x q))) x)" by (simp add: mul_c)  | 
|
48  | 
also have "\<dots> = (mul (mul y (pwr y q)) (mul (pwr x q) x))" by (simp add: mul_a)  | 
|
49  | 
finally show "mul (mul x y) (mul (pwr x q) (pwr y q)) =  | 
|
50  | 
mul (mul x (pwr x q)) (mul y (pwr y q))" by (simp add: mul_c)  | 
|
51  | 
qed  | 
|
52  | 
||
53  | 
lemma pwr_pwr: "pwr (pwr x p) q = pwr x (p * q)"  | 
|
54  | 
proof (induct p arbitrary: q)  | 
|
55  | 
case 0  | 
|
56  | 
show ?case using pwr_Suc mul_1 pwr_0 by (induct q) auto  | 
|
57  | 
next  | 
|
58  | 
case Suc  | 
|
59  | 
thus ?case by (auto simp add: mul_pwr [symmetric] pwr_mul pwr_Suc)  | 
|
60  | 
qed  | 
|
61  | 
||
62  | 
||
63  | 
subsubsection {* Declaring the abstract theory *}
 | 
|
64  | 
||
65  | 
lemma semiring_ops:  | 
|
66  | 
shows "TERM (add x y)" and "TERM (mul x y)" and "TERM (pwr x n)"  | 
|
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28823 
diff
changeset
 | 
67  | 
and "TERM r0" and "TERM r1" .  | 
| 23252 | 68  | 
|
69  | 
lemma semiring_rules:  | 
|
70  | 
"add (mul a m) (mul b m) = mul (add a b) m"  | 
|
71  | 
"add (mul a m) m = mul (add a r1) m"  | 
|
72  | 
"add m (mul a m) = mul (add a r1) m"  | 
|
73  | 
"add m m = mul (add r1 r1) m"  | 
|
74  | 
"add r0 a = a"  | 
|
75  | 
"add a r0 = a"  | 
|
76  | 
"mul a b = mul b a"  | 
|
77  | 
"mul (add a b) c = add (mul a c) (mul b c)"  | 
|
78  | 
"mul r0 a = r0"  | 
|
79  | 
"mul a r0 = r0"  | 
|
80  | 
"mul r1 a = a"  | 
|
81  | 
"mul a r1 = a"  | 
|
82  | 
"mul (mul lx ly) (mul rx ry) = mul (mul lx rx) (mul ly ry)"  | 
|
83  | 
"mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry))"  | 
|
84  | 
"mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry)"  | 
|
85  | 
"mul (mul lx ly) rx = mul (mul lx rx) ly"  | 
|
86  | 
"mul (mul lx ly) rx = mul lx (mul ly rx)"  | 
|
87  | 
"mul lx (mul rx ry) = mul (mul lx rx) ry"  | 
|
88  | 
"mul lx (mul rx ry) = mul rx (mul lx ry)"  | 
|
89  | 
"add (add a b) (add c d) = add (add a c) (add b d)"  | 
|
90  | 
"add (add a b) c = add a (add b c)"  | 
|
91  | 
"add a (add c d) = add c (add a d)"  | 
|
92  | 
"add (add a b) c = add (add a c) b"  | 
|
93  | 
"add a c = add c a"  | 
|
94  | 
"add a (add c d) = add (add a c) d"  | 
|
95  | 
"mul (pwr x p) (pwr x q) = pwr x (p + q)"  | 
|
96  | 
"mul x (pwr x q) = pwr x (Suc q)"  | 
|
97  | 
"mul (pwr x q) x = pwr x (Suc q)"  | 
|
98  | 
"mul x x = pwr x 2"  | 
|
99  | 
"pwr (mul x y) q = mul (pwr x q) (pwr y q)"  | 
|
100  | 
"pwr (pwr x p) q = pwr x (p * q)"  | 
|
101  | 
"pwr x 0 = r1"  | 
|
102  | 
"pwr x 1 = x"  | 
|
103  | 
"mul x (add y z) = add (mul x y) (mul x z)"  | 
|
104  | 
"pwr x (Suc q) = mul x (pwr x q)"  | 
|
105  | 
"pwr x (2*n) = mul (pwr x n) (pwr x n)"  | 
|
106  | 
"pwr x (Suc (2*n)) = mul x (mul (pwr x n) (pwr x n))"  | 
|
107  | 
proof -  | 
|
108  | 
show "add (mul a m) (mul b m) = mul (add a b) m" using mul_d mul_c by simp  | 
|
109  | 
next show"add (mul a m) m = mul (add a r1) m" using mul_d mul_c mul_1 by simp  | 
|
110  | 
next show "add m (mul a m) = mul (add a r1) m" using mul_c mul_d mul_1 add_c by simp  | 
|
111  | 
next show "add m m = mul (add r1 r1) m" using mul_c mul_d mul_1 by simp  | 
|
112  | 
next show "add r0 a = a" using add_0 by simp  | 
|
113  | 
next show "add a r0 = a" using add_0 add_c by simp  | 
|
114  | 
next show "mul a b = mul b a" using mul_c by simp  | 
|
115  | 
next show "mul (add a b) c = add (mul a c) (mul b c)" using mul_c mul_d by simp  | 
|
116  | 
next show "mul r0 a = r0" using mul_0 by simp  | 
|
117  | 
next show "mul a r0 = r0" using mul_0 mul_c by simp  | 
|
118  | 
next show "mul r1 a = a" using mul_1 by simp  | 
|
119  | 
next show "mul a r1 = a" using mul_1 mul_c by simp  | 
|
120  | 
next show "mul (mul lx ly) (mul rx ry) = mul (mul lx rx) (mul ly ry)"  | 
|
121  | 
using mul_c mul_a by simp  | 
|
122  | 
next show "mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry))"  | 
|
123  | 
using mul_a by simp  | 
|
124  | 
next  | 
|
125  | 
have "mul (mul lx ly) (mul rx ry) = mul (mul rx ry) (mul lx ly)" by (rule mul_c)  | 
|
126  | 
also have "\<dots> = mul rx (mul ry (mul lx ly))" using mul_a by simp  | 
|
127  | 
finally  | 
|
128  | 
show "mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry)"  | 
|
129  | 
using mul_c by simp  | 
|
130  | 
next show "mul (mul lx ly) rx = mul (mul lx rx) ly" using mul_c mul_a by simp  | 
|
131  | 
next  | 
|
132  | 
show "mul (mul lx ly) rx = mul lx (mul ly rx)" by (simp add: mul_a)  | 
|
133  | 
next show "mul lx (mul rx ry) = mul (mul lx rx) ry" by (simp add: mul_a )  | 
|
134  | 
next show "mul lx (mul rx ry) = mul rx (mul lx ry)" by (simp add: mul_a,simp add: mul_c)  | 
|
135  | 
next show "add (add a b) (add c d) = add (add a c) (add b d)"  | 
|
136  | 
using add_c add_a by simp  | 
|
137  | 
next show "add (add a b) c = add a (add b c)" using add_a by simp  | 
|
138  | 
next show "add a (add c d) = add c (add a d)"  | 
|
139  | 
apply (simp add: add_a) by (simp only: add_c)  | 
|
140  | 
next show "add (add a b) c = add (add a c) b" using add_a add_c by simp  | 
|
141  | 
next show "add a c = add c a" by (rule add_c)  | 
|
142  | 
next show "add a (add c d) = add (add a c) d" using add_a by simp  | 
|
143  | 
next show "mul (pwr x p) (pwr x q) = pwr x (p + q)" by (rule mul_pwr)  | 
|
144  | 
next show "mul x (pwr x q) = pwr x (Suc q)" using pwr_Suc by simp  | 
|
145  | 
next show "mul (pwr x q) x = pwr x (Suc q)" using pwr_Suc mul_c by simp  | 
|
| 35216 | 146  | 
next show "mul x x = pwr x 2" by (simp add: nat_number' pwr_Suc pwr_0 mul_1 mul_c)  | 
| 23252 | 147  | 
next show "pwr (mul x y) q = mul (pwr x q) (pwr y q)" by (rule pwr_mul)  | 
148  | 
next show "pwr (pwr x p) q = pwr x (p * q)" by (rule pwr_pwr)  | 
|
149  | 
next show "pwr x 0 = r1" using pwr_0 .  | 
|
| 35216 | 150  | 
next show "pwr x 1 = x" unfolding One_nat_def by (simp add: nat_number' pwr_Suc pwr_0 mul_1 mul_c)  | 
| 23252 | 151  | 
next show "mul x (add y z) = add (mul x y) (mul x z)" using mul_d by simp  | 
152  | 
next show "pwr x (Suc q) = mul x (pwr x q)" using pwr_Suc by simp  | 
|
| 35216 | 153  | 
next show "pwr x (2 * n) = mul (pwr x n) (pwr x n)" by (simp add: nat_number' mul_pwr)  | 
| 23252 | 154  | 
next show "pwr x (Suc (2 * n)) = mul x (mul (pwr x n) (pwr x n))"  | 
| 35216 | 155  | 
by (simp add: nat_number' pwr_Suc mul_pwr)  | 
| 23252 | 156  | 
qed  | 
157  | 
||
158  | 
||
| 26462 | 159  | 
lemmas gb_semiring_axioms' =  | 
| 26314 | 160  | 
gb_semiring_axioms [normalizer  | 
| 23252 | 161  | 
semiring ops: semiring_ops  | 
| 26314 | 162  | 
semiring rules: semiring_rules]  | 
| 23252 | 163  | 
|
164  | 
end  | 
|
165  | 
||
| 
30729
 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 
wenzelm 
parents: 
30654 
diff
changeset
 | 
166  | 
interpretation class_semiring: gb_semiring  | 
| 31017 | 167  | 
    "op +" "op *" "op ^" "0::'a::{comm_semiring_1}" "1"
 | 
| 35216 | 168  | 
proof qed (auto simp add: algebra_simps)  | 
| 23252 | 169  | 
|
170  | 
lemmas nat_arith =  | 
|
| 28987 | 171  | 
add_nat_number_of  | 
172  | 
diff_nat_number_of  | 
|
173  | 
mult_nat_number_of  | 
|
174  | 
eq_nat_number_of  | 
|
175  | 
less_nat_number_of  | 
|
| 23252 | 176  | 
|
177  | 
lemma not_iszero_Numeral1: "\<not> iszero (Numeral1::'a::number_ring)"  | 
|
| 35216 | 178  | 
by simp  | 
| 28986 | 179  | 
|
| 29039 | 180  | 
lemmas comp_arith =  | 
181  | 
Let_def arith_simps nat_arith rel_simps neg_simps if_False  | 
|
| 23252 | 182  | 
if_True add_0 add_Suc add_number_of_left mult_number_of_left  | 
| 31790 | 183  | 
numeral_1_eq_1[symmetric] Suc_eq_plus1  | 
| 28986 | 184  | 
numeral_0_eq_0[symmetric] numerals[symmetric]  | 
185  | 
iszero_simps not_iszero_Numeral1  | 
|
| 23252 | 186  | 
|
187  | 
lemmas semiring_norm = comp_arith  | 
|
188  | 
||
189  | 
ML {*
 | 
|
| 23573 | 190  | 
local  | 
| 23252 | 191  | 
|
| 23573 | 192  | 
open Conv;  | 
193  | 
||
| 30866 | 194  | 
fun numeral_is_const ct = can HOLogic.dest_number (Thm.term_of ct);  | 
| 23252 | 195  | 
|
| 23573 | 196  | 
fun int_of_rat x =  | 
197  | 
(case Rat.quotient_of_rat x of (i, 1) => i  | 
|
198  | 
| _ => error "int_of_rat: bad int");  | 
|
| 23252 | 199  | 
|
| 23573 | 200  | 
val numeral_conv =  | 
201  | 
  Simplifier.rewrite (HOL_basic_ss addsimps @{thms semiring_norm}) then_conv
 | 
|
202  | 
Simplifier.rewrite (HOL_basic_ss addsimps  | 
|
203  | 
    (@{thms numeral_1_eq_1} @ @{thms numeral_0_eq_0} @ @{thms numerals(1-2)}));
 | 
|
204  | 
||
205  | 
in  | 
|
206  | 
||
207  | 
fun normalizer_funs key =  | 
|
208  | 
NormalizerData.funs key  | 
|
| 23252 | 209  | 
   {is_const = fn phi => numeral_is_const,
 | 
210  | 
dest_const = fn phi => fn ct =>  | 
|
211  | 
Rat.rat_of_int (snd  | 
|
212  | 
(HOLogic.dest_number (Thm.term_of ct)  | 
|
213  | 
handle TERM _ => error "ring_dest_const")),  | 
|
| 23573 | 214  | 
mk_const = fn phi => fn cT => fn x => Numeral.mk_cnumber cT (int_of_rat x),  | 
| 
23330
 
01c09922ce59
Conversion for computation on constants now depends on the context
 
chaieb 
parents: 
23327 
diff
changeset
 | 
215  | 
conv = fn phi => K numeral_conv}  | 
| 23573 | 216  | 
|
217  | 
end  | 
|
| 23252 | 218  | 
*}  | 
219  | 
||
| 26462 | 220  | 
declaration {* normalizer_funs @{thm class_semiring.gb_semiring_axioms'} *}
 | 
| 23573 | 221  | 
|
| 23252 | 222  | 
|
| 
23258
 
9062e98fdab1
renamed locale ring/semiring to gb_ring/gb_semiring to avoid clash with Ring_and_Field versions;
 
wenzelm 
parents: 
23252 
diff
changeset
 | 
223  | 
locale gb_ring = gb_semiring +  | 
| 23252 | 224  | 
fixes sub :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  | 
225  | 
and neg :: "'a \<Rightarrow> 'a"  | 
|
226  | 
assumes neg_mul: "neg x = mul (neg r1) x"  | 
|
227  | 
and sub_add: "sub x y = add x (neg y)"  | 
|
228  | 
begin  | 
|
229  | 
||
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28823 
diff
changeset
 | 
230  | 
lemma ring_ops: shows "TERM (sub x y)" and "TERM (neg x)" .  | 
| 23252 | 231  | 
|
232  | 
lemmas ring_rules = neg_mul sub_add  | 
|
233  | 
||
| 26462 | 234  | 
lemmas gb_ring_axioms' =  | 
| 26314 | 235  | 
gb_ring_axioms [normalizer  | 
236  | 
semiring ops: semiring_ops  | 
|
237  | 
semiring rules: semiring_rules  | 
|
238  | 
ring ops: ring_ops  | 
|
239  | 
ring rules: ring_rules]  | 
|
| 23252 | 240  | 
|
241  | 
end  | 
|
242  | 
||
243  | 
||
| 
30729
 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 
wenzelm 
parents: 
30654 
diff
changeset
 | 
244  | 
interpretation class_ring: gb_ring "op +" "op *" "op ^"  | 
| 31017 | 245  | 
    "0::'a::{comm_semiring_1,number_ring}" 1 "op -" "uminus"
 | 
| 28823 | 246  | 
proof qed simp_all  | 
| 23252 | 247  | 
|
248  | 
||
| 26462 | 249  | 
declaration {* normalizer_funs @{thm class_ring.gb_ring_axioms'} *}
 | 
| 23252 | 250  | 
|
251  | 
use "Tools/Groebner_Basis/normalizer.ML"  | 
|
252  | 
||
| 27666 | 253  | 
|
| 23252 | 254  | 
method_setup sring_norm = {*
 | 
| 30549 | 255  | 
Scan.succeed (SIMPLE_METHOD' o Normalizer.semiring_normalize_tac)  | 
| 23458 | 256  | 
*} "semiring normalizer"  | 
| 23252 | 257  | 
|
258  | 
||
| 23327 | 259  | 
locale gb_field = gb_ring +  | 
260  | 
fixes divide :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  | 
|
261  | 
and inverse:: "'a \<Rightarrow> 'a"  | 
|
| 30866 | 262  | 
assumes divide_inverse: "divide x y = mul x (inverse y)"  | 
263  | 
and inverse_divide: "inverse x = divide r1 x"  | 
|
| 23327 | 264  | 
begin  | 
265  | 
||
| 30866 | 266  | 
lemma field_ops: shows "TERM (divide x y)" and "TERM (inverse x)" .  | 
267  | 
||
268  | 
lemmas field_rules = divide_inverse inverse_divide  | 
|
269  | 
||
| 26462 | 270  | 
lemmas gb_field_axioms' =  | 
| 26314 | 271  | 
gb_field_axioms [normalizer  | 
272  | 
semiring ops: semiring_ops  | 
|
273  | 
semiring rules: semiring_rules  | 
|
274  | 
ring ops: ring_ops  | 
|
| 30866 | 275  | 
ring rules: ring_rules  | 
276  | 
field ops: field_ops  | 
|
277  | 
field rules: field_rules]  | 
|
| 23327 | 278  | 
|
279  | 
end  | 
|
280  | 
||
| 23458 | 281  | 
|
| 23266 | 282  | 
subsection {* Groebner Bases *}
 | 
| 23252 | 283  | 
|
| 
23258
 
9062e98fdab1
renamed locale ring/semiring to gb_ring/gb_semiring to avoid clash with Ring_and_Field versions;
 
wenzelm 
parents: 
23252 
diff
changeset
 | 
284  | 
locale semiringb = gb_semiring +  | 
| 23252 | 285  | 
assumes add_cancel: "add (x::'a) y = add x z \<longleftrightarrow> y = z"  | 
286  | 
and add_mul_solve: "add (mul w y) (mul x z) =  | 
|
287  | 
add (mul w z) (mul x y) \<longleftrightarrow> w = x \<or> y = z"  | 
|
288  | 
begin  | 
|
289  | 
||
290  | 
lemma noteq_reduce: "a \<noteq> b \<and> c \<noteq> d \<longleftrightarrow> add (mul a c) (mul b d) \<noteq> add (mul a d) (mul b c)"  | 
|
291  | 
proof-  | 
|
292  | 
have "a \<noteq> b \<and> c \<noteq> d \<longleftrightarrow> \<not> (a = b \<or> c = d)" by simp  | 
|
293  | 
also have "\<dots> \<longleftrightarrow> add (mul a c) (mul b d) \<noteq> add (mul a d) (mul b c)"  | 
|
294  | 
using add_mul_solve by blast  | 
|
295  | 
finally show "a \<noteq> b \<and> c \<noteq> d \<longleftrightarrow> add (mul a c) (mul b d) \<noteq> add (mul a d) (mul b c)"  | 
|
296  | 
by simp  | 
|
297  | 
qed  | 
|
298  | 
||
299  | 
lemma add_scale_eq_noteq: "\<lbrakk>r \<noteq> r0 ; (a = b) \<and> ~(c = d)\<rbrakk>  | 
|
300  | 
\<Longrightarrow> add a (mul r c) \<noteq> add b (mul r d)"  | 
|
301  | 
proof(clarify)  | 
|
302  | 
assume nz: "r\<noteq> r0" and cnd: "c\<noteq>d"  | 
|
303  | 
and eq: "add b (mul r c) = add b (mul r d)"  | 
|
304  | 
hence "mul r c = mul r d" using cnd add_cancel by simp  | 
|
305  | 
hence "add (mul r0 d) (mul r c) = add (mul r0 c) (mul r d)"  | 
|
306  | 
using mul_0 add_cancel by simp  | 
|
307  | 
thus "False" using add_mul_solve nz cnd by simp  | 
|
308  | 
qed  | 
|
309  | 
||
| 
25250
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
310  | 
lemma add_r0_iff: " x = add x a \<longleftrightarrow> a = r0"  | 
| 
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
311  | 
proof-  | 
| 
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
312  | 
have "a = r0 \<longleftrightarrow> add x a = add x r0" by (simp add: add_cancel)  | 
| 
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
313  | 
thus "x = add x a \<longleftrightarrow> a = r0" by (auto simp add: add_c add_0)  | 
| 
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
314  | 
qed  | 
| 
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
315  | 
|
| 26462 | 316  | 
declare gb_semiring_axioms' [normalizer del]  | 
| 23252 | 317  | 
|
| 26462 | 318  | 
lemmas semiringb_axioms' = semiringb_axioms [normalizer  | 
| 23252 | 319  | 
semiring ops: semiring_ops  | 
320  | 
semiring rules: semiring_rules  | 
|
| 26314 | 321  | 
idom rules: noteq_reduce add_scale_eq_noteq]  | 
| 23252 | 322  | 
|
323  | 
end  | 
|
324  | 
||
| 
25250
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
325  | 
locale ringb = semiringb + gb_ring +  | 
| 
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
326  | 
assumes subr0_iff: "sub x y = r0 \<longleftrightarrow> x = y"  | 
| 23252 | 327  | 
begin  | 
328  | 
||
| 26462 | 329  | 
declare gb_ring_axioms' [normalizer del]  | 
| 23252 | 330  | 
|
| 26462 | 331  | 
lemmas ringb_axioms' = ringb_axioms [normalizer  | 
| 23252 | 332  | 
semiring ops: semiring_ops  | 
333  | 
semiring rules: semiring_rules  | 
|
334  | 
ring ops: ring_ops  | 
|
335  | 
ring rules: ring_rules  | 
|
| 
25250
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
336  | 
idom rules: noteq_reduce add_scale_eq_noteq  | 
| 26314 | 337  | 
ideal rules: subr0_iff add_r0_iff]  | 
| 23252 | 338  | 
|
339  | 
end  | 
|
340  | 
||
| 
25250
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
341  | 
|
| 23252 | 342  | 
lemma no_zero_divirors_neq0:  | 
343  | 
assumes az: "(a::'a::no_zero_divisors) \<noteq> 0"  | 
|
344  | 
and ab: "a*b = 0" shows "b = 0"  | 
|
345  | 
proof -  | 
|
346  | 
  { assume bz: "b \<noteq> 0"
 | 
|
347  | 
from no_zero_divisors [OF az bz] ab have False by blast }  | 
|
348  | 
thus "b = 0" by blast  | 
|
349  | 
qed  | 
|
350  | 
||
| 
30729
 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 
wenzelm 
parents: 
30654 
diff
changeset
 | 
351  | 
interpretation class_ringb: ringb  | 
| 31017 | 352  | 
  "op +" "op *" "op ^" "0::'a::{idom,number_ring}" "1" "op -" "uminus"
 | 
| 35216 | 353  | 
proof(unfold_locales, simp add: algebra_simps, auto)  | 
| 31017 | 354  | 
  fix w x y z ::"'a::{idom,number_ring}"
 | 
| 23252 | 355  | 
assume p: "w * y + x * z = w * z + x * y" and ynz: "y \<noteq> z"  | 
356  | 
hence ynz': "y - z \<noteq> 0" by simp  | 
|
357  | 
from p have "w * y + x* z - w*z - x*y = 0" by simp  | 
|
| 29667 | 358  | 
hence "w* (y - z) - x * (y - z) = 0" by (simp add: algebra_simps)  | 
359  | 
hence "(y - z) * (w - x) = 0" by (simp add: algebra_simps)  | 
|
| 23252 | 360  | 
with no_zero_divirors_neq0 [OF ynz']  | 
361  | 
have "w - x = 0" by blast  | 
|
362  | 
thus "w = x" by simp  | 
|
363  | 
qed  | 
|
364  | 
||
| 26462 | 365  | 
declaration {* normalizer_funs @{thm class_ringb.ringb_axioms'} *}
 | 
| 23252 | 366  | 
|
| 
30729
 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 
wenzelm 
parents: 
30654 
diff
changeset
 | 
367  | 
interpretation natgb: semiringb  | 
| 29223 | 368  | 
"op +" "op *" "op ^" "0::nat" "1"  | 
| 35216 | 369  | 
proof (unfold_locales, simp add: algebra_simps)  | 
| 23252 | 370  | 
fix w x y z ::"nat"  | 
371  | 
  { assume p: "w * y + x * z = w * z + x * y" and ynz: "y \<noteq> z"
 | 
|
372  | 
hence "y < z \<or> y > z" by arith  | 
|
373  | 
    moreover {
 | 
|
374  | 
assume lt:"y <z" hence "\<exists>k. z = y + k \<and> k > 0" by (rule_tac x="z - y" in exI, auto)  | 
|
375  | 
then obtain k where kp: "k>0" and yz:"z = y + k" by blast  | 
|
| 29667 | 376  | 
from p have "(w * y + x *y) + x*k = (w * y + x*y) + w*k" by (simp add: yz algebra_simps)  | 
| 23252 | 377  | 
hence "x*k = w*k" by simp  | 
| 35216 | 378  | 
hence "w = x" using kp by simp }  | 
| 23252 | 379  | 
    moreover {
 | 
380  | 
assume lt: "y >z" hence "\<exists>k. y = z + k \<and> k>0" by (rule_tac x="y - z" in exI, auto)  | 
|
381  | 
then obtain k where kp: "k>0" and yz:"y = z + k" by blast  | 
|
| 29667 | 382  | 
from p have "(w * z + x *z) + w*k = (w * z + x*z) + x*k" by (simp add: yz algebra_simps)  | 
| 23252 | 383  | 
hence "w*k = x*k" by simp  | 
| 35216 | 384  | 
hence "w = x" using kp by simp }  | 
| 23252 | 385  | 
ultimately have "w=x" by blast }  | 
386  | 
thus "(w * y + x * z = w * z + x * y) = (w = x \<or> y = z)" by auto  | 
|
387  | 
qed  | 
|
388  | 
||
| 26462 | 389  | 
declaration {* normalizer_funs @{thm natgb.semiringb_axioms'} *}
 | 
| 23252 | 390  | 
|
| 23327 | 391  | 
locale fieldgb = ringb + gb_field  | 
392  | 
begin  | 
|
393  | 
||
| 26462 | 394  | 
declare gb_field_axioms' [normalizer del]  | 
| 23327 | 395  | 
|
| 26462 | 396  | 
lemmas fieldgb_axioms' = fieldgb_axioms [normalizer  | 
| 23327 | 397  | 
semiring ops: semiring_ops  | 
398  | 
semiring rules: semiring_rules  | 
|
399  | 
ring ops: ring_ops  | 
|
400  | 
ring rules: ring_rules  | 
|
| 30866 | 401  | 
field ops: field_ops  | 
402  | 
field rules: field_rules  | 
|
| 
25250
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
403  | 
idom rules: noteq_reduce add_scale_eq_noteq  | 
| 26314 | 404  | 
ideal rules: subr0_iff add_r0_iff]  | 
405  | 
||
| 23327 | 406  | 
end  | 
407  | 
||
408  | 
||
| 
23258
 
9062e98fdab1
renamed locale ring/semiring to gb_ring/gb_semiring to avoid clash with Ring_and_Field versions;
 
wenzelm 
parents: 
23252 
diff
changeset
 | 
409  | 
lemmas bool_simps = simp_thms(1-34)  | 
| 23252 | 410  | 
lemma dnf:  | 
411  | 
"(P & (Q | R)) = ((P&Q) | (P&R))" "((Q | R) & P) = ((Q&P) | (R&P))"  | 
|
412  | 
"(P \<and> Q) = (Q \<and> P)" "(P \<or> Q) = (Q \<or> P)"  | 
|
413  | 
by blast+  | 
|
414  | 
||
415  | 
lemmas weak_dnf_simps = dnf bool_simps  | 
|
416  | 
||
417  | 
lemma nnf_simps:  | 
|
418  | 
"(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)" "(P \<longrightarrow> Q) = (\<not>P \<or> Q)"  | 
|
419  | 
"(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"  | 
|
420  | 
by blast+  | 
|
421  | 
||
422  | 
lemma PFalse:  | 
|
423  | 
"P \<equiv> False \<Longrightarrow> \<not> P"  | 
|
424  | 
"\<not> P \<Longrightarrow> (P \<equiv> False)"  | 
|
425  | 
by auto  | 
|
426  | 
use "Tools/Groebner_Basis/groebner.ML"  | 
|
427  | 
||
| 
23332
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
428  | 
method_setup algebra =  | 
| 23458 | 429  | 
{*
 | 
| 
23332
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
430  | 
let  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
431  | 
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
432  | 
val addN = "add"  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
433  | 
val delN = "del"  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
434  | 
val any_keyword = keyword addN || keyword delN  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
435  | 
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
436  | 
in  | 
| 30549 | 437  | 
((Scan.optional (keyword addN |-- thms) []) --  | 
438  | 
(Scan.optional (keyword delN |-- thms) [])) >>  | 
|
439  | 
(fn (add_ths, del_ths) => fn ctxt =>  | 
|
| 
30510
 
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
 
wenzelm 
parents: 
30242 
diff
changeset
 | 
440  | 
SIMPLE_METHOD' (Groebner.algebra_tac add_ths del_ths ctxt))  | 
| 
23332
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
441  | 
end  | 
| 
25250
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
442  | 
*} "solve polynomial equations over (semi)rings and ideal membership problems using Groebner bases"  | 
| 27666 | 443  | 
declare dvd_def[algebra]  | 
444  | 
declare dvd_eq_mod_eq_0[symmetric, algebra]  | 
|
| 30027 | 445  | 
declare mod_div_trivial[algebra]  | 
446  | 
declare mod_mod_trivial[algebra]  | 
|
| 27666 | 447  | 
declare conjunct1[OF DIVISION_BY_ZERO, algebra]  | 
448  | 
declare conjunct2[OF DIVISION_BY_ZERO, algebra]  | 
|
449  | 
declare zmod_zdiv_equality[symmetric,algebra]  | 
|
450  | 
declare zdiv_zmod_equality[symmetric, algebra]  | 
|
451  | 
declare zdiv_zminus_zminus[algebra]  | 
|
452  | 
declare zmod_zminus_zminus[algebra]  | 
|
453  | 
declare zdiv_zminus2[algebra]  | 
|
454  | 
declare zmod_zminus2[algebra]  | 
|
455  | 
declare zdiv_zero[algebra]  | 
|
456  | 
declare zmod_zero[algebra]  | 
|
| 30031 | 457  | 
declare mod_by_1[algebra]  | 
458  | 
declare div_by_1[algebra]  | 
|
| 27666 | 459  | 
declare zmod_minus1_right[algebra]  | 
460  | 
declare zdiv_minus1_right[algebra]  | 
|
461  | 
declare mod_div_trivial[algebra]  | 
|
462  | 
declare mod_mod_trivial[algebra]  | 
|
| 30034 | 463  | 
declare mod_mult_self2_is_0[algebra]  | 
464  | 
declare mod_mult_self1_is_0[algebra]  | 
|
| 27666 | 465  | 
declare zmod_eq_0_iff[algebra]  | 
| 30042 | 466  | 
declare dvd_0_left_iff[algebra]  | 
| 27666 | 467  | 
declare zdvd1_eq[algebra]  | 
468  | 
declare zmod_eq_dvd_iff[algebra]  | 
|
469  | 
declare nat_mod_eq_iff[algebra]  | 
|
| 23252 | 470  | 
|
| 28402 | 471  | 
subsection{* Groebner Bases for fields *}
 | 
472  | 
||
| 
30729
 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 
wenzelm 
parents: 
30654 
diff
changeset
 | 
473  | 
interpretation class_fieldgb:  | 
| 31017 | 474  | 
  fieldgb "op +" "op *" "op ^" "0::'a::{field,number_ring}" "1" "op -" "uminus" "op /" "inverse" apply (unfold_locales) by (simp_all add: divide_inverse)
 | 
| 28402 | 475  | 
|
476  | 
lemma divide_Numeral1: "(x::'a::{field,number_ring}) / Numeral1 = x" by simp
 | 
|
477  | 
lemma divide_Numeral0: "(x::'a::{field,number_ring, division_by_zero}) / Numeral0 = 0"
 | 
|
478  | 
by simp  | 
|
479  | 
lemma mult_frac_frac: "((x::'a::{field,division_by_zero}) / y) * (z / w) = (x*z) / (y*w)"
 | 
|
480  | 
by simp  | 
|
481  | 
lemma mult_frac_num: "((x::'a::{field, division_by_zero}) / y) * z  = (x*z) / y"
 | 
|
482  | 
by simp  | 
|
483  | 
lemma mult_num_frac: "((x::'a::{field, division_by_zero}) / y) * z  = (x*z) / y"
 | 
|
484  | 
by simp  | 
|
485  | 
||
486  | 
lemma Numeral1_eq1_nat: "(1::nat) = Numeral1" by simp  | 
|
487  | 
||
488  | 
lemma add_frac_num: "y\<noteq> 0 \<Longrightarrow> (x::'a::{field, division_by_zero}) / y + z = (x + z*y) / y"
 | 
|
489  | 
by (simp add: add_divide_distrib)  | 
|
490  | 
lemma add_num_frac: "y\<noteq> 0 \<Longrightarrow> z + (x::'a::{field, division_by_zero}) / y = (x + z*y) / y"
 | 
|
491  | 
by (simp add: add_divide_distrib)  | 
|
| 35084 | 492  | 
|
493  | 
ML {*
 | 
|
494  | 
let open Conv  | 
|
495  | 
in fconv_rule (arg_conv (arg1_conv (rewr_conv (mk_meta_eq @{thm mult_commute})))) (@{thm field_divide_inverse} RS sym)
 | 
|
496  | 
end  | 
|
497  | 
*}  | 
|
498  | 
||
| 28402 | 499  | 
ML{* 
 | 
500  | 
local  | 
|
501  | 
 val zr = @{cpat "0"}
 | 
|
502  | 
val zT = ctyp_of_term zr  | 
|
503  | 
 val geq = @{cpat "op ="}
 | 
|
504  | 
val eqT = Thm.dest_ctyp (ctyp_of_term geq) |> hd  | 
|
505  | 
 val add_frac_eq = mk_meta_eq @{thm "add_frac_eq"}
 | 
|
506  | 
 val add_frac_num = mk_meta_eq @{thm "add_frac_num"}
 | 
|
507  | 
 val add_num_frac = mk_meta_eq @{thm "add_num_frac"}
 | 
|
508  | 
||
509  | 
fun prove_nz ss T t =  | 
|
510  | 
let  | 
|
511  | 
val z = instantiate_cterm ([(zT,T)],[]) zr  | 
|
512  | 
val eq = instantiate_cterm ([(eqT,T)],[]) geq  | 
|
513  | 
val th = Simplifier.rewrite (ss addsimps simp_thms)  | 
|
514  | 
           (Thm.capply @{cterm "Trueprop"} (Thm.capply @{cterm "Not"}
 | 
|
515  | 
(Thm.capply (Thm.capply eq t) z)))  | 
|
516  | 
in equal_elim (symmetric th) TrueI  | 
|
517  | 
end  | 
|
518  | 
||
519  | 
fun proc phi ss ct =  | 
|
520  | 
let  | 
|
521  | 
val ((x,y),(w,z)) =  | 
|
522  | 
(Thm.dest_binop #> (fn (a,b) => (Thm.dest_binop a, Thm.dest_binop b))) ct  | 
|
523  | 
val _ = map (HOLogic.dest_number o term_of) [x,y,z,w]  | 
|
524  | 
val T = ctyp_of_term x  | 
|
525  | 
val [y_nz, z_nz] = map (prove_nz ss T) [y, z]  | 
|
526  | 
val th = instantiate' [SOME T] (map SOME [y,z,x,w]) add_frac_eq  | 
|
527  | 
in SOME (implies_elim (implies_elim th y_nz) z_nz)  | 
|
528  | 
end  | 
|
529  | 
handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE  | 
|
530  | 
||
531  | 
fun proc2 phi ss ct =  | 
|
532  | 
let  | 
|
533  | 
val (l,r) = Thm.dest_binop ct  | 
|
534  | 
val T = ctyp_of_term l  | 
|
535  | 
in (case (term_of l, term_of r) of  | 
|
| 35084 | 536  | 
      (Const(@{const_name Rings.divide},_)$_$_, _) =>
 | 
| 28402 | 537  | 
let val (x,y) = Thm.dest_binop l val z = r  | 
538  | 
val _ = map (HOLogic.dest_number o term_of) [x,y,z]  | 
|
539  | 
val ynz = prove_nz ss T y  | 
|
540  | 
in SOME (implies_elim (instantiate' [SOME T] (map SOME [y,x,z]) add_frac_num) ynz)  | 
|
541  | 
end  | 
|
| 35084 | 542  | 
     | (_, Const (@{const_name Rings.divide},_)$_$_) =>
 | 
| 28402 | 543  | 
let val (x,y) = Thm.dest_binop r val z = l  | 
544  | 
val _ = map (HOLogic.dest_number o term_of) [x,y,z]  | 
|
545  | 
val ynz = prove_nz ss T y  | 
|
546  | 
in SOME (implies_elim (instantiate' [SOME T] (map SOME [y,z,x]) add_num_frac) ynz)  | 
|
547  | 
end  | 
|
548  | 
| _ => NONE)  | 
|
549  | 
end  | 
|
550  | 
handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE  | 
|
551  | 
||
| 35084 | 552  | 
 fun is_number (Const(@{const_name Rings.divide},_)$a$b) = is_number a andalso is_number b
 | 
| 28402 | 553  | 
| is_number t = can HOLogic.dest_number t  | 
554  | 
||
555  | 
val is_number = is_number o term_of  | 
|
556  | 
||
557  | 
fun proc3 phi ss ct =  | 
|
558  | 
(case term_of ct of  | 
|
| 
35092
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35084 
diff
changeset
 | 
559  | 
    Const(@{const_name Orderings.less},_)$(Const(@{const_name Rings.divide},_)$_$_)$_ =>
 | 
| 28402 | 560  | 
let  | 
561  | 
val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop  | 
|
562  | 
val _ = map is_number [a,b,c]  | 
|
563  | 
val T = ctyp_of_term c  | 
|
564  | 
        val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_less_eq"}
 | 
|
565  | 
in SOME (mk_meta_eq th) end  | 
|
| 
35092
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35084 
diff
changeset
 | 
566  | 
  | Const(@{const_name Orderings.less_eq},_)$(Const(@{const_name Rings.divide},_)$_$_)$_ =>
 | 
| 28402 | 567  | 
let  | 
568  | 
val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop  | 
|
569  | 
val _ = map is_number [a,b,c]  | 
|
570  | 
val T = ctyp_of_term c  | 
|
571  | 
        val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_le_eq"}
 | 
|
572  | 
in SOME (mk_meta_eq th) end  | 
|
| 35084 | 573  | 
  | Const("op =",_)$(Const(@{const_name Rings.divide},_)$_$_)$_ =>
 | 
| 28402 | 574  | 
let  | 
575  | 
val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop  | 
|
576  | 
val _ = map is_number [a,b,c]  | 
|
577  | 
val T = ctyp_of_term c  | 
|
578  | 
        val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_eq_eq"}
 | 
|
579  | 
in SOME (mk_meta_eq th) end  | 
|
| 
35092
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35084 
diff
changeset
 | 
580  | 
  | Const(@{const_name Orderings.less},_)$_$(Const(@{const_name Rings.divide},_)$_$_) =>
 | 
| 28402 | 581  | 
let  | 
582  | 
val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop  | 
|
583  | 
val _ = map is_number [a,b,c]  | 
|
584  | 
val T = ctyp_of_term c  | 
|
585  | 
        val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "less_divide_eq"}
 | 
|
586  | 
in SOME (mk_meta_eq th) end  | 
|
| 
35092
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35084 
diff
changeset
 | 
587  | 
  | Const(@{const_name Orderings.less_eq},_)$_$(Const(@{const_name Rings.divide},_)$_$_) =>
 | 
| 28402 | 588  | 
let  | 
589  | 
val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop  | 
|
590  | 
val _ = map is_number [a,b,c]  | 
|
591  | 
val T = ctyp_of_term c  | 
|
592  | 
        val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "le_divide_eq"}
 | 
|
593  | 
in SOME (mk_meta_eq th) end  | 
|
| 35084 | 594  | 
  | Const("op =",_)$_$(Const(@{const_name Rings.divide},_)$_$_) =>
 | 
| 28402 | 595  | 
let  | 
596  | 
val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop  | 
|
597  | 
val _ = map is_number [a,b,c]  | 
|
598  | 
val T = ctyp_of_term c  | 
|
599  | 
        val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "eq_divide_eq"}
 | 
|
600  | 
in SOME (mk_meta_eq th) end  | 
|
601  | 
| _ => NONE)  | 
|
602  | 
handle TERM _ => NONE | CTERM _ => NONE | THM _ => NONE  | 
|
603  | 
||
604  | 
val add_frac_frac_simproc =  | 
|
605  | 
       make_simproc {lhss = [@{cpat "(?x::?'a::field)/?y + (?w::?'a::field)/?z"}],
 | 
|
606  | 
name = "add_frac_frac_simproc",  | 
|
607  | 
proc = proc, identifier = []}  | 
|
608  | 
||
609  | 
val add_frac_num_simproc =  | 
|
610  | 
       make_simproc {lhss = [@{cpat "(?x::?'a::field)/?y + ?z"}, @{cpat "?z + (?x::?'a::field)/?y"}],
 | 
|
611  | 
name = "add_frac_num_simproc",  | 
|
612  | 
proc = proc2, identifier = []}  | 
|
613  | 
||
614  | 
val ord_frac_simproc =  | 
|
615  | 
make_simproc  | 
|
616  | 
    {lhss = [@{cpat "(?a::(?'a::{field, ord}))/?b < ?c"},
 | 
|
617  | 
             @{cpat "(?a::(?'a::{field, ord}))/?b \<le> ?c"},
 | 
|
618  | 
             @{cpat "?c < (?a::(?'a::{field, ord}))/?b"},
 | 
|
619  | 
             @{cpat "?c \<le> (?a::(?'a::{field, ord}))/?b"},
 | 
|
620  | 
             @{cpat "?c = ((?a::(?'a::{field, ord}))/?b)"},
 | 
|
621  | 
             @{cpat "((?a::(?'a::{field, ord}))/ ?b) = ?c"}],
 | 
|
622  | 
name = "ord_frac_simproc", proc = proc3, identifier = []}  | 
|
623  | 
||
| 
30869
 
71fde5b7b43c
More precise treatement of rational constants by the normalizer for fields
 
chaieb 
parents: 
30866 
diff
changeset
 | 
624  | 
local  | 
| 
 
71fde5b7b43c
More precise treatement of rational constants by the normalizer for fields
 
chaieb 
parents: 
30866 
diff
changeset
 | 
625  | 
open Conv  | 
| 
 
71fde5b7b43c
More precise treatement of rational constants by the normalizer for fields
 
chaieb 
parents: 
30866 
diff
changeset
 | 
626  | 
in  | 
| 
 
71fde5b7b43c
More precise treatement of rational constants by the normalizer for fields
 
chaieb 
parents: 
30866 
diff
changeset
 | 
627  | 
|
| 28402 | 628  | 
val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"},
 | 
629  | 
           @{thm "divide_Numeral1"},
 | 
|
| 
35050
 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 
haftmann 
parents: 
34974 
diff
changeset
 | 
630  | 
           @{thm "Fields.divide_zero"}, @{thm "divide_Numeral0"},
 | 
| 28402 | 631  | 
           @{thm "divide_divide_eq_left"}, @{thm "mult_frac_frac"},
 | 
632  | 
           @{thm "mult_num_frac"}, @{thm "mult_frac_num"},
 | 
|
633  | 
           @{thm "mult_frac_frac"}, @{thm "times_divide_eq_right"},
 | 
|
634  | 
           @{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"},
 | 
|
635  | 
           @{thm "diff_def"}, @{thm "minus_divide_left"},
 | 
|
| 
30869
 
71fde5b7b43c
More precise treatement of rational constants by the normalizer for fields
 
chaieb 
parents: 
30866 
diff
changeset
 | 
636  | 
           @{thm "Numeral1_eq1_nat"}, @{thm "add_divide_distrib"} RS sym,
 | 
| 35084 | 637  | 
           @{thm field_divide_inverse} RS sym, @{thm inverse_divide}, 
 | 
| 
30869
 
71fde5b7b43c
More precise treatement of rational constants by the normalizer for fields
 
chaieb 
parents: 
30866 
diff
changeset
 | 
638  | 
           fconv_rule (arg_conv (arg1_conv (rewr_conv (mk_meta_eq @{thm mult_commute}))))   
 | 
| 35084 | 639  | 
           (@{thm field_divide_inverse} RS sym)]
 | 
| 28402 | 640  | 
|
641  | 
val comp_conv = (Simplifier.rewrite  | 
|
642  | 
(HOL_basic_ss addsimps @{thms "Groebner_Basis.comp_arith"}
 | 
|
| 28987 | 643  | 
addsimps ths addsimps simp_thms  | 
| 
31068
 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 
haftmann 
parents: 
31017 
diff
changeset
 | 
644  | 
addsimprocs Numeral_Simprocs.field_cancel_numeral_factors  | 
| 28402 | 645  | 
addsimprocs [add_frac_frac_simproc, add_frac_num_simproc,  | 
646  | 
ord_frac_simproc]  | 
|
647  | 
                addcongs [@{thm "if_weak_cong"}]))
 | 
|
648  | 
then_conv (Simplifier.rewrite (HOL_basic_ss addsimps  | 
|
649  | 
  [@{thm numeral_1_eq_1},@{thm numeral_0_eq_0}] @ @{thms numerals(1-2)}))
 | 
|
| 23252 | 650  | 
end  | 
| 28402 | 651  | 
|
652  | 
fun numeral_is_const ct =  | 
|
653  | 
case term_of ct of  | 
|
| 35084 | 654  | 
   Const (@{const_name Rings.divide},_) $ a $ b =>
 | 
| 30866 | 655  | 
can HOLogic.dest_number a andalso can HOLogic.dest_number b  | 
| 35084 | 656  | 
 | Const (@{const_name Rings.inverse},_)$t => can HOLogic.dest_number t
 | 
| 28402 | 657  | 
| t => can HOLogic.dest_number t  | 
658  | 
||
659  | 
fun dest_const ct = ((case term_of ct of  | 
|
| 35084 | 660  | 
   Const (@{const_name Rings.divide},_) $ a $ b=>
 | 
| 28402 | 661  | 
Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b))  | 
| 35084 | 662  | 
 | Const (@{const_name Rings.inverse},_)$t => 
 | 
| 
30869
 
71fde5b7b43c
More precise treatement of rational constants by the normalizer for fields
 
chaieb 
parents: 
30866 
diff
changeset
 | 
663  | 
Rat.inv (Rat.rat_of_int (snd (HOLogic.dest_number t)))  | 
| 28402 | 664  | 
| t => Rat.rat_of_int (snd (HOLogic.dest_number t)))  | 
665  | 
handle TERM _ => error "ring_dest_const")  | 
|
666  | 
||
667  | 
fun mk_const phi cT x =  | 
|
668  | 
let val (a, b) = Rat.quotient_of_rat x  | 
|
669  | 
in if b = 1 then Numeral.mk_cnumber cT a  | 
|
670  | 
else Thm.capply  | 
|
671  | 
         (Thm.capply (Drule.cterm_rule (instantiate' [SOME cT] []) @{cpat "op /"})
 | 
|
672  | 
(Numeral.mk_cnumber cT a))  | 
|
673  | 
(Numeral.mk_cnumber cT b)  | 
|
674  | 
end  | 
|
675  | 
||
676  | 
in  | 
|
677  | 
val field_comp_conv = comp_conv;  | 
|
678  | 
val fieldgb_declaration =  | 
|
679  | 
  NormalizerData.funs @{thm class_fieldgb.fieldgb_axioms'}
 | 
|
680  | 
   {is_const = K numeral_is_const,
 | 
|
681  | 
dest_const = K dest_const,  | 
|
682  | 
mk_const = mk_const,  | 
|
683  | 
conv = K (K comp_conv)}  | 
|
684  | 
end;  | 
|
685  | 
*}  | 
|
686  | 
||
687  | 
declaration fieldgb_declaration  | 
|
688  | 
||
689  | 
end  |