author | haftmann |
Thu, 06 May 2010 23:11:57 +0200 | |
changeset 36720 | 41da7025e59c |
parent 36716 | b09f3ad3208f |
child 36751 | 7f1da69cacb3 |
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 |
36699
816da1023508
moved nat_arith ot Nat_Numeral: clarified normalizer setup
haftmann
parents:
36698
diff
changeset
|
8 |
imports Numeral_Simprocs Nat_Transfer |
23252 | 9 |
uses |
36699
816da1023508
moved nat_arith ot Nat_Numeral: clarified normalizer setup
haftmann
parents:
36698
diff
changeset
|
10 |
"Tools/Groebner_Basis/normalizer.ML" |
23312 | 11 |
("Tools/Groebner_Basis/groebner.ML") |
23252 | 12 |
begin |
13 |
||
14 |
subsection {* Semiring normalization *} |
|
15 |
||
36700
9b85b9d74b83
dropped auxiliary method sring_norm; integrated normalizer.ML and normalizer_data.ML
haftmann
parents:
36699
diff
changeset
|
16 |
setup Normalizer.setup |
23252 | 17 |
|
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
18 |
locale normalizing_semiring = |
23252 | 19 |
fixes add mul pwr r0 r1 |
20 |
assumes add_a:"(add x (add y z) = add (add x y) z)" |
|
21 |
and add_c: "add x y = add y x" and add_0:"add r0 x = x" |
|
22 |
and mul_a:"mul x (mul y z) = mul (mul x y) z" and mul_c:"mul x y = mul y x" |
|
23 |
and mul_1:"mul r1 x = x" and mul_0:"mul r0 x = r0" |
|
24 |
and mul_d:"mul x (add y z) = add (mul x y) (mul x z)" |
|
25 |
and pwr_0:"pwr x 0 = r1" and pwr_Suc:"pwr x (Suc n) = mul x (pwr x n)" |
|
26 |
begin |
|
27 |
||
28 |
lemma mul_pwr:"mul (pwr x p) (pwr x q) = pwr x (p + q)" |
|
29 |
proof (induct p) |
|
30 |
case 0 |
|
31 |
then show ?case by (auto simp add: pwr_0 mul_1) |
|
32 |
next |
|
33 |
case Suc |
|
34 |
from this [symmetric] show ?case |
|
35 |
by (auto simp add: pwr_Suc mul_1 mul_a) |
|
36 |
qed |
|
37 |
||
38 |
lemma pwr_mul: "pwr (mul x y) q = mul (pwr x q) (pwr y q)" |
|
39 |
proof (induct q arbitrary: x y, auto simp add:pwr_0 pwr_Suc mul_1) |
|
40 |
fix q x y |
|
41 |
assume "\<And>x y. pwr (mul x y) q = mul (pwr x q) (pwr y q)" |
|
42 |
have "mul (mul x y) (mul (pwr x q) (pwr y q)) = mul x (mul y (mul (pwr x q) (pwr y q)))" |
|
43 |
by (simp add: mul_a) |
|
44 |
also have "\<dots> = (mul (mul y (mul (pwr y q) (pwr x q))) x)" by (simp add: mul_c) |
|
45 |
also have "\<dots> = (mul (mul y (pwr y q)) (mul (pwr x q) x))" by (simp add: mul_a) |
|
46 |
finally show "mul (mul x y) (mul (pwr x q) (pwr y q)) = |
|
47 |
mul (mul x (pwr x q)) (mul y (pwr y q))" by (simp add: mul_c) |
|
48 |
qed |
|
49 |
||
50 |
lemma pwr_pwr: "pwr (pwr x p) q = pwr x (p * q)" |
|
51 |
proof (induct p arbitrary: q) |
|
52 |
case 0 |
|
53 |
show ?case using pwr_Suc mul_1 pwr_0 by (induct q) auto |
|
54 |
next |
|
55 |
case Suc |
|
56 |
thus ?case by (auto simp add: mul_pwr [symmetric] pwr_mul pwr_Suc) |
|
57 |
qed |
|
58 |
||
59 |
lemma semiring_ops: |
|
60 |
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
|
61 |
and "TERM r0" and "TERM r1" . |
23252 | 62 |
|
63 |
lemma semiring_rules: |
|
64 |
"add (mul a m) (mul b m) = mul (add a b) m" |
|
65 |
"add (mul a m) m = mul (add a r1) m" |
|
66 |
"add m (mul a m) = mul (add a r1) m" |
|
67 |
"add m m = mul (add r1 r1) m" |
|
68 |
"add r0 a = a" |
|
69 |
"add a r0 = a" |
|
70 |
"mul a b = mul b a" |
|
71 |
"mul (add a b) c = add (mul a c) (mul b c)" |
|
72 |
"mul r0 a = r0" |
|
73 |
"mul a r0 = r0" |
|
74 |
"mul r1 a = a" |
|
75 |
"mul a r1 = a" |
|
76 |
"mul (mul lx ly) (mul rx ry) = mul (mul lx rx) (mul ly ry)" |
|
77 |
"mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry))" |
|
78 |
"mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry)" |
|
79 |
"mul (mul lx ly) rx = mul (mul lx rx) ly" |
|
80 |
"mul (mul lx ly) rx = mul lx (mul ly rx)" |
|
81 |
"mul lx (mul rx ry) = mul (mul lx rx) ry" |
|
82 |
"mul lx (mul rx ry) = mul rx (mul lx ry)" |
|
83 |
"add (add a b) (add c d) = add (add a c) (add b d)" |
|
84 |
"add (add a b) c = add a (add b c)" |
|
85 |
"add a (add c d) = add c (add a d)" |
|
86 |
"add (add a b) c = add (add a c) b" |
|
87 |
"add a c = add c a" |
|
88 |
"add a (add c d) = add (add a c) d" |
|
89 |
"mul (pwr x p) (pwr x q) = pwr x (p + q)" |
|
90 |
"mul x (pwr x q) = pwr x (Suc q)" |
|
91 |
"mul (pwr x q) x = pwr x (Suc q)" |
|
92 |
"mul x x = pwr x 2" |
|
93 |
"pwr (mul x y) q = mul (pwr x q) (pwr y q)" |
|
94 |
"pwr (pwr x p) q = pwr x (p * q)" |
|
95 |
"pwr x 0 = r1" |
|
96 |
"pwr x 1 = x" |
|
97 |
"mul x (add y z) = add (mul x y) (mul x z)" |
|
98 |
"pwr x (Suc q) = mul x (pwr x q)" |
|
99 |
"pwr x (2*n) = mul (pwr x n) (pwr x n)" |
|
100 |
"pwr x (Suc (2*n)) = mul x (mul (pwr x n) (pwr x n))" |
|
101 |
proof - |
|
102 |
show "add (mul a m) (mul b m) = mul (add a b) m" using mul_d mul_c by simp |
|
103 |
next show"add (mul a m) m = mul (add a r1) m" using mul_d mul_c mul_1 by simp |
|
104 |
next show "add m (mul a m) = mul (add a r1) m" using mul_c mul_d mul_1 add_c by simp |
|
105 |
next show "add m m = mul (add r1 r1) m" using mul_c mul_d mul_1 by simp |
|
106 |
next show "add r0 a = a" using add_0 by simp |
|
107 |
next show "add a r0 = a" using add_0 add_c by simp |
|
108 |
next show "mul a b = mul b a" using mul_c by simp |
|
109 |
next show "mul (add a b) c = add (mul a c) (mul b c)" using mul_c mul_d by simp |
|
110 |
next show "mul r0 a = r0" using mul_0 by simp |
|
111 |
next show "mul a r0 = r0" using mul_0 mul_c by simp |
|
112 |
next show "mul r1 a = a" using mul_1 by simp |
|
113 |
next show "mul a r1 = a" using mul_1 mul_c by simp |
|
114 |
next show "mul (mul lx ly) (mul rx ry) = mul (mul lx rx) (mul ly ry)" |
|
115 |
using mul_c mul_a by simp |
|
116 |
next show "mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry))" |
|
117 |
using mul_a by simp |
|
118 |
next |
|
119 |
have "mul (mul lx ly) (mul rx ry) = mul (mul rx ry) (mul lx ly)" by (rule mul_c) |
|
120 |
also have "\<dots> = mul rx (mul ry (mul lx ly))" using mul_a by simp |
|
121 |
finally |
|
122 |
show "mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry)" |
|
123 |
using mul_c by simp |
|
124 |
next show "mul (mul lx ly) rx = mul (mul lx rx) ly" using mul_c mul_a by simp |
|
125 |
next |
|
126 |
show "mul (mul lx ly) rx = mul lx (mul ly rx)" by (simp add: mul_a) |
|
127 |
next show "mul lx (mul rx ry) = mul (mul lx rx) ry" by (simp add: mul_a ) |
|
128 |
next show "mul lx (mul rx ry) = mul rx (mul lx ry)" by (simp add: mul_a,simp add: mul_c) |
|
129 |
next show "add (add a b) (add c d) = add (add a c) (add b d)" |
|
130 |
using add_c add_a by simp |
|
131 |
next show "add (add a b) c = add a (add b c)" using add_a by simp |
|
132 |
next show "add a (add c d) = add c (add a d)" |
|
133 |
apply (simp add: add_a) by (simp only: add_c) |
|
134 |
next show "add (add a b) c = add (add a c) b" using add_a add_c by simp |
|
135 |
next show "add a c = add c a" by (rule add_c) |
|
136 |
next show "add a (add c d) = add (add a c) d" using add_a by simp |
|
137 |
next show "mul (pwr x p) (pwr x q) = pwr x (p + q)" by (rule mul_pwr) |
|
138 |
next show "mul x (pwr x q) = pwr x (Suc q)" using pwr_Suc by simp |
|
139 |
next show "mul (pwr x q) x = pwr x (Suc q)" using pwr_Suc mul_c by simp |
|
35216 | 140 |
next show "mul x x = pwr x 2" by (simp add: nat_number' pwr_Suc pwr_0 mul_1 mul_c) |
23252 | 141 |
next show "pwr (mul x y) q = mul (pwr x q) (pwr y q)" by (rule pwr_mul) |
142 |
next show "pwr (pwr x p) q = pwr x (p * q)" by (rule pwr_pwr) |
|
143 |
next show "pwr x 0 = r1" using pwr_0 . |
|
35216 | 144 |
next show "pwr x 1 = x" unfolding One_nat_def by (simp add: nat_number' pwr_Suc pwr_0 mul_1 mul_c) |
23252 | 145 |
next show "mul x (add y z) = add (mul x y) (mul x z)" using mul_d by simp |
146 |
next show "pwr x (Suc q) = mul x (pwr x q)" using pwr_Suc by simp |
|
35216 | 147 |
next show "pwr x (2 * n) = mul (pwr x n) (pwr x n)" by (simp add: nat_number' mul_pwr) |
23252 | 148 |
next show "pwr x (Suc (2 * n)) = mul x (mul (pwr x n) (pwr x n))" |
35216 | 149 |
by (simp add: nat_number' pwr_Suc mul_pwr) |
23252 | 150 |
qed |
151 |
||
152 |
||
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
153 |
lemmas normalizing_semiring_axioms' = |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
154 |
normalizing_semiring_axioms [normalizer |
23252 | 155 |
semiring ops: semiring_ops |
26314 | 156 |
semiring rules: semiring_rules] |
23252 | 157 |
|
158 |
end |
|
159 |
||
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
160 |
sublocale comm_semiring_1 |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
161 |
< normalizing!: normalizing_semiring plus times power zero one |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
162 |
proof |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
163 |
qed (simp_all add: algebra_simps) |
23252 | 164 |
|
36720 | 165 |
declaration {* Normalizer.semiring_funs @{thm normalizing.normalizing_semiring_axioms'} *} |
23573 | 166 |
|
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
167 |
locale normalizing_ring = normalizing_semiring + |
23252 | 168 |
fixes sub :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" |
169 |
and neg :: "'a \<Rightarrow> 'a" |
|
170 |
assumes neg_mul: "neg x = mul (neg r1) x" |
|
171 |
and sub_add: "sub x y = add x (neg y)" |
|
172 |
begin |
|
173 |
||
28856
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28823
diff
changeset
|
174 |
lemma ring_ops: shows "TERM (sub x y)" and "TERM (neg x)" . |
23252 | 175 |
|
176 |
lemmas ring_rules = neg_mul sub_add |
|
177 |
||
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
178 |
lemmas normalizing_ring_axioms' = |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
179 |
normalizing_ring_axioms [normalizer |
26314 | 180 |
semiring ops: semiring_ops |
181 |
semiring rules: semiring_rules |
|
182 |
ring ops: ring_ops |
|
183 |
ring rules: ring_rules] |
|
23252 | 184 |
|
185 |
end |
|
186 |
||
36720 | 187 |
sublocale comm_ring_1 |
188 |
< normalizing!: normalizing_ring plus times power zero one minus uminus |
|
189 |
proof |
|
190 |
qed (simp_all add: diff_minus) |
|
23252 | 191 |
|
36720 | 192 |
declaration {* Normalizer.semiring_funs @{thm normalizing.normalizing_ring_axioms'} *} |
23252 | 193 |
|
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
194 |
locale normalizing_field = normalizing_ring + |
23327 | 195 |
fixes divide :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" |
196 |
and inverse:: "'a \<Rightarrow> 'a" |
|
30866 | 197 |
assumes divide_inverse: "divide x y = mul x (inverse y)" |
198 |
and inverse_divide: "inverse x = divide r1 x" |
|
23327 | 199 |
begin |
200 |
||
30866 | 201 |
lemma field_ops: shows "TERM (divide x y)" and "TERM (inverse x)" . |
202 |
||
203 |
lemmas field_rules = divide_inverse inverse_divide |
|
204 |
||
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
205 |
lemmas normalizing_field_axioms' = |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
206 |
normalizing_field_axioms [normalizer |
26314 | 207 |
semiring ops: semiring_ops |
208 |
semiring rules: semiring_rules |
|
209 |
ring ops: ring_ops |
|
30866 | 210 |
ring rules: ring_rules |
211 |
field ops: field_ops |
|
212 |
field rules: field_rules] |
|
23327 | 213 |
|
214 |
end |
|
215 |
||
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
216 |
locale normalizing_semiring_cancel = normalizing_semiring + |
23252 | 217 |
assumes add_cancel: "add (x::'a) y = add x z \<longleftrightarrow> y = z" |
218 |
and add_mul_solve: "add (mul w y) (mul x z) = |
|
219 |
add (mul w z) (mul x y) \<longleftrightarrow> w = x \<or> y = z" |
|
220 |
begin |
|
221 |
||
222 |
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)" |
|
223 |
proof- |
|
224 |
have "a \<noteq> b \<and> c \<noteq> d \<longleftrightarrow> \<not> (a = b \<or> c = d)" by simp |
|
225 |
also have "\<dots> \<longleftrightarrow> add (mul a c) (mul b d) \<noteq> add (mul a d) (mul b c)" |
|
226 |
using add_mul_solve by blast |
|
227 |
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)" |
|
228 |
by simp |
|
229 |
qed |
|
230 |
||
231 |
lemma add_scale_eq_noteq: "\<lbrakk>r \<noteq> r0 ; (a = b) \<and> ~(c = d)\<rbrakk> |
|
232 |
\<Longrightarrow> add a (mul r c) \<noteq> add b (mul r d)" |
|
233 |
proof(clarify) |
|
234 |
assume nz: "r\<noteq> r0" and cnd: "c\<noteq>d" |
|
235 |
and eq: "add b (mul r c) = add b (mul r d)" |
|
236 |
hence "mul r c = mul r d" using cnd add_cancel by simp |
|
237 |
hence "add (mul r0 d) (mul r c) = add (mul r0 c) (mul r d)" |
|
238 |
using mul_0 add_cancel by simp |
|
239 |
thus "False" using add_mul_solve nz cnd by simp |
|
240 |
qed |
|
241 |
||
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
|
242 |
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
|
243 |
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
|
244 |
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
|
245 |
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
|
246 |
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
|
247 |
|
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
248 |
declare normalizing_semiring_axioms' [normalizer del] |
23252 | 249 |
|
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
250 |
lemmas normalizing_semiring_cancel_axioms' = |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
251 |
normalizing_semiring_cancel_axioms [normalizer |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
252 |
semiring ops: semiring_ops |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
253 |
semiring rules: semiring_rules |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
254 |
idom rules: noteq_reduce add_scale_eq_noteq] |
23252 | 255 |
|
256 |
end |
|
257 |
||
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
258 |
locale normalizing_ring_cancel = normalizing_semiring_cancel + normalizing_ring + |
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
|
259 |
assumes subr0_iff: "sub x y = r0 \<longleftrightarrow> x = y" |
23252 | 260 |
begin |
261 |
||
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
262 |
declare normalizing_ring_axioms' [normalizer del] |
23252 | 263 |
|
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
264 |
lemmas normalizing_ring_cancel_axioms' = normalizing_ring_cancel_axioms [normalizer |
23252 | 265 |
semiring ops: semiring_ops |
266 |
semiring rules: semiring_rules |
|
267 |
ring ops: ring_ops |
|
268 |
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
|
269 |
idom rules: noteq_reduce add_scale_eq_noteq |
26314 | 270 |
ideal rules: subr0_iff add_r0_iff] |
23252 | 271 |
|
272 |
end |
|
273 |
||
36720 | 274 |
sublocale idom |
275 |
< normalizing!: normalizing_ring_cancel plus times power zero one minus uminus |
|
276 |
proof |
|
277 |
fix w x y z |
|
278 |
show "w * y + x * z = w * z + x * y \<longleftrightarrow> w = x \<or> y = z" |
|
279 |
proof |
|
280 |
assume "w * y + x * z = w * z + x * y" |
|
281 |
then have "w * y + x * z - w * z - x * y = 0" by (simp add: algebra_simps) |
|
282 |
then have "w * (y - z) - x * (y - z) = 0" by (simp add: algebra_simps) |
|
283 |
then have "(y - z) * (w - x) = 0" by (simp add: algebra_simps) |
|
284 |
then have "y - z = 0 \<or> w - x = 0" by (rule divisors_zero) |
|
285 |
then show "w = x \<or> y = z" by auto |
|
286 |
qed (auto simp add: add_ac) |
|
287 |
qed (simp_all add: algebra_simps) |
|
23252 | 288 |
|
36720 | 289 |
declaration {* Normalizer.semiring_funs @{thm normalizing.normalizing_ring_cancel_axioms'} *} |
23252 | 290 |
|
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
291 |
interpretation normalizing_nat!: normalizing_semiring_cancel |
29223 | 292 |
"op +" "op *" "op ^" "0::nat" "1" |
35216 | 293 |
proof (unfold_locales, simp add: algebra_simps) |
23252 | 294 |
fix w x y z ::"nat" |
295 |
{ assume p: "w * y + x * z = w * z + x * y" and ynz: "y \<noteq> z" |
|
296 |
hence "y < z \<or> y > z" by arith |
|
297 |
moreover { |
|
298 |
assume lt:"y <z" hence "\<exists>k. z = y + k \<and> k > 0" by (rule_tac x="z - y" in exI, auto) |
|
299 |
then obtain k where kp: "k>0" and yz:"z = y + k" by blast |
|
29667 | 300 |
from p have "(w * y + x *y) + x*k = (w * y + x*y) + w*k" by (simp add: yz algebra_simps) |
23252 | 301 |
hence "x*k = w*k" by simp |
35216 | 302 |
hence "w = x" using kp by simp } |
23252 | 303 |
moreover { |
304 |
assume lt: "y >z" hence "\<exists>k. y = z + k \<and> k>0" by (rule_tac x="y - z" in exI, auto) |
|
305 |
then obtain k where kp: "k>0" and yz:"y = z + k" by blast |
|
29667 | 306 |
from p have "(w * z + x *z) + w*k = (w * z + x*z) + x*k" by (simp add: yz algebra_simps) |
23252 | 307 |
hence "w*k = x*k" by simp |
35216 | 308 |
hence "w = x" using kp by simp } |
23252 | 309 |
ultimately have "w=x" by blast } |
310 |
thus "(w * y + x * z = w * z + x * y) = (w = x \<or> y = z)" by auto |
|
311 |
qed |
|
312 |
||
36720 | 313 |
declaration {* Normalizer.semiring_funs @{thm normalizing_nat.normalizing_semiring_cancel_axioms'} *} |
23252 | 314 |
|
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
315 |
locale normalizing_field_cancel = normalizing_ring_cancel + normalizing_field |
23327 | 316 |
begin |
317 |
||
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
318 |
declare normalizing_field_axioms' [normalizer del] |
23327 | 319 |
|
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
320 |
lemmas normalizing_field_cancel_axioms' = normalizing_field_cancel_axioms [normalizer |
23327 | 321 |
semiring ops: semiring_ops |
322 |
semiring rules: semiring_rules |
|
323 |
ring ops: ring_ops |
|
324 |
ring rules: ring_rules |
|
30866 | 325 |
field ops: field_ops |
326 |
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
|
327 |
idom rules: noteq_reduce add_scale_eq_noteq |
26314 | 328 |
ideal rules: subr0_iff add_r0_iff] |
329 |
||
23327 | 330 |
end |
331 |
||
36720 | 332 |
sublocale field |
333 |
< normalizing!: normalizing_field_cancel plus times power zero one minus uminus divide inverse |
|
334 |
proof |
|
335 |
qed (simp_all add: divide_inverse) |
|
28402 | 336 |
|
36720 | 337 |
declaration {* Normalizer.field_funs @{thm normalizing.normalizing_field_cancel_axioms'} *} |
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
338 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
339 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
340 |
subsection {* Groebner Bases *} |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
341 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
342 |
lemmas bool_simps = simp_thms(1-34) |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
343 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
344 |
lemma dnf: |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
345 |
"(P & (Q | R)) = ((P&Q) | (P&R))" "((Q | R) & P) = ((Q&P) | (R&P))" |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
346 |
"(P \<and> Q) = (Q \<and> P)" "(P \<or> Q) = (Q \<or> P)" |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
347 |
by blast+ |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
348 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
349 |
lemmas weak_dnf_simps = dnf bool_simps |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
350 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
351 |
lemma nnf_simps: |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
352 |
"(\<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)" |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
353 |
"(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P" |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
354 |
by blast+ |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
355 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
356 |
lemma PFalse: |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
357 |
"P \<equiv> False \<Longrightarrow> \<not> P" |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
358 |
"\<not> P \<Longrightarrow> (P \<equiv> False)" |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
359 |
by auto |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
360 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
361 |
ML {* |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
362 |
structure Algebra_Simplification = Named_Thms( |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
363 |
val name = "algebra" |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
364 |
val description = "pre-simplification rules for algebraic methods" |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
365 |
) |
28402 | 366 |
*} |
367 |
||
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
368 |
setup Algebra_Simplification.setup |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
369 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
370 |
declare dvd_def[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
371 |
declare dvd_eq_mod_eq_0[symmetric, algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
372 |
declare mod_div_trivial[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
373 |
declare mod_mod_trivial[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
374 |
declare conjunct1[OF DIVISION_BY_ZERO, algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
375 |
declare conjunct2[OF DIVISION_BY_ZERO, algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
376 |
declare zmod_zdiv_equality[symmetric,algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
377 |
declare zdiv_zmod_equality[symmetric, algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
378 |
declare zdiv_zminus_zminus[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
379 |
declare zmod_zminus_zminus[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
380 |
declare zdiv_zminus2[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
381 |
declare zmod_zminus2[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
382 |
declare zdiv_zero[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
383 |
declare zmod_zero[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
384 |
declare mod_by_1[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
385 |
declare div_by_1[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
386 |
declare zmod_minus1_right[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
387 |
declare zdiv_minus1_right[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
388 |
declare mod_div_trivial[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
389 |
declare mod_mod_trivial[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
390 |
declare mod_mult_self2_is_0[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
391 |
declare mod_mult_self1_is_0[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
392 |
declare zmod_eq_0_iff[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
393 |
declare dvd_0_left_iff[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
394 |
declare zdvd1_eq[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
395 |
declare zmod_eq_dvd_iff[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
396 |
declare nat_mod_eq_iff[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
397 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
398 |
use "Tools/Groebner_Basis/groebner.ML" |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
399 |
|
36720 | 400 |
method_setup algebra = Groebner.algebra_method |
401 |
"solve polynomial equations over (semi)rings and ideal membership problems using Groebner bases" |
|
28402 | 402 |
|
403 |
end |