author | haftmann |
Sat, 28 Aug 2010 16:14:32 +0200 | |
changeset 38864 | 4abe644fcea5 |
parent 37946 | be3c0df7bb90 |
child 47108 | 2a1953f0d20d |
permissions | -rw-r--r-- |
36751
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
haftmann
parents:
36720
diff
changeset
|
1 |
(* Title: HOL/Semiring_Normalization.thy |
23252 | 2 |
Author: Amine Chaieb, TU Muenchen |
3 |
*) |
|
4 |
||
36751
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
haftmann
parents:
36720
diff
changeset
|
5 |
header {* Semiring normalization *} |
28402 | 6 |
|
36751
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
haftmann
parents:
36720
diff
changeset
|
7 |
theory Semiring_Normalization |
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 |
36753
5cf4e9128f22
renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents:
36751
diff
changeset
|
10 |
"Tools/semiring_normalizer.ML" |
23252 | 11 |
begin |
12 |
||
36873 | 13 |
text {* Prelude *} |
14 |
||
15 |
class comm_semiring_1_cancel_crossproduct = comm_semiring_1_cancel + |
|
16 |
assumes crossproduct_eq: "w * y + x * z = w * z + x * y \<longleftrightarrow> w = x \<or> y = z" |
|
17 |
begin |
|
18 |
||
19 |
lemma crossproduct_noteq: |
|
20 |
"a \<noteq> b \<and> c \<noteq> d \<longleftrightarrow> a * c + b * d \<noteq> a * d + b * c" |
|
21 |
by (simp add: crossproduct_eq) |
|
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
22 |
|
36873 | 23 |
lemma add_scale_eq_noteq: |
24 |
"r \<noteq> 0 \<Longrightarrow> a = b \<and> c \<noteq> d \<Longrightarrow> a + r * c \<noteq> b + r * d" |
|
25 |
proof (rule notI) |
|
26 |
assume nz: "r\<noteq> 0" and cnd: "a = b \<and> c\<noteq>d" |
|
27 |
and eq: "a + (r * c) = b + (r * d)" |
|
28 |
have "(0 * d) + (r * c) = (0 * c) + (r * d)" |
|
29 |
using add_imp_eq eq mult_zero_left by (simp add: cnd) |
|
30 |
then show False using crossproduct_eq [of 0 d] nz cnd by simp |
|
31 |
qed |
|
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
32 |
|
36873 | 33 |
lemma add_0_iff: |
34 |
"b = b + a \<longleftrightarrow> a = 0" |
|
35 |
using add_imp_eq [of b a 0] by auto |
|
36 |
||
37 |
end |
|
38 |
||
37946 | 39 |
subclass (in idom) comm_semiring_1_cancel_crossproduct |
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
40 |
proof |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
41 |
fix w x y z |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
42 |
show "w * y + x * z = w * z + x * y \<longleftrightarrow> w = x \<or> y = z" |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
43 |
proof |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
44 |
assume "w * y + x * z = w * z + x * y" |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
45 |
then have "w * y + x * z - w * z - x * y = 0" by (simp add: algebra_simps) |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
46 |
then have "w * (y - z) - x * (y - z) = 0" by (simp add: algebra_simps) |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
47 |
then have "(y - z) * (w - x) = 0" by (simp add: algebra_simps) |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
48 |
then have "y - z = 0 \<or> w - x = 0" by (rule divisors_zero) |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
49 |
then show "w = x \<or> y = z" by auto |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
50 |
qed (auto simp add: add_ac) |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
51 |
qed |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
52 |
|
36873 | 53 |
instance nat :: comm_semiring_1_cancel_crossproduct |
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
54 |
proof |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
55 |
fix w x y z :: nat |
36873 | 56 |
have aux: "\<And>y z. y < z \<Longrightarrow> w * y + x * z = w * z + x * y \<Longrightarrow> w = x" |
57 |
proof - |
|
58 |
fix y z :: nat |
|
59 |
assume "y < z" then have "\<exists>k. z = y + k \<and> k \<noteq> 0" by (intro exI [of _ "z - y"]) auto |
|
60 |
then obtain k where "z = y + k" and "k \<noteq> 0" by blast |
|
61 |
assume "w * y + x * z = w * z + x * y" |
|
62 |
then have "(w * y + x * y) + x * k = (w * y + x * y) + w * k" by (simp add: `z = y + k` algebra_simps) |
|
63 |
then have "x * k = w * k" by simp |
|
64 |
then show "w = x" using `k \<noteq> 0` by simp |
|
65 |
qed |
|
66 |
show "w * y + x * z = w * z + x * y \<longleftrightarrow> w = x \<or> y = z" |
|
67 |
by (auto simp add: neq_iff dest!: aux) |
|
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
68 |
qed |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
69 |
|
36873 | 70 |
text {* Semiring normalization proper *} |
36871 | 71 |
|
36753
5cf4e9128f22
renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents:
36751
diff
changeset
|
72 |
setup Semiring_Normalizer.setup |
23252 | 73 |
|
36871 | 74 |
context comm_semiring_1 |
75 |
begin |
|
76 |
||
36872 | 77 |
lemma normalizing_semiring_ops: |
36845
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
78 |
shows "TERM (x + y)" and "TERM (x * y)" and "TERM (x ^ n)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
79 |
and "TERM 0" and "TERM 1" . |
23252 | 80 |
|
36872 | 81 |
lemma normalizing_semiring_rules: |
36845
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
82 |
"(a * m) + (b * m) = (a + b) * m" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
83 |
"(a * m) + m = (a + 1) * m" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
84 |
"m + (a * m) = (a + 1) * m" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
85 |
"m + m = (1 + 1) * m" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
86 |
"0 + a = a" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
87 |
"a + 0 = a" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
88 |
"a * b = b * a" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
89 |
"(a + b) * c = (a * c) + (b * c)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
90 |
"0 * a = 0" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
91 |
"a * 0 = 0" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
92 |
"1 * a = a" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
93 |
"a * 1 = a" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
94 |
"(lx * ly) * (rx * ry) = (lx * rx) * (ly * ry)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
95 |
"(lx * ly) * (rx * ry) = lx * (ly * (rx * ry))" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
96 |
"(lx * ly) * (rx * ry) = rx * ((lx * ly) * ry)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
97 |
"(lx * ly) * rx = (lx * rx) * ly" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
98 |
"(lx * ly) * rx = lx * (ly * rx)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
99 |
"lx * (rx * ry) = (lx * rx) * ry" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
100 |
"lx * (rx * ry) = rx * (lx * ry)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
101 |
"(a + b) + (c + d) = (a + c) + (b + d)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
102 |
"(a + b) + c = a + (b + c)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
103 |
"a + (c + d) = c + (a + d)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
104 |
"(a + b) + c = (a + c) + b" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
105 |
"a + c = c + a" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
106 |
"a + (c + d) = (a + c) + d" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
107 |
"(x ^ p) * (x ^ q) = x ^ (p + q)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
108 |
"x * (x ^ q) = x ^ (Suc q)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
109 |
"(x ^ q) * x = x ^ (Suc q)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
110 |
"x * x = x ^ 2" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
111 |
"(x * y) ^ q = (x ^ q) * (y ^ q)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
112 |
"(x ^ p) ^ q = x ^ (p * q)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
113 |
"x ^ 0 = 1" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
114 |
"x ^ 1 = x" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
115 |
"x * (y + z) = (x * y) + (x * z)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
116 |
"x ^ (Suc q) = x * (x ^ q)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
117 |
"x ^ (2*n) = (x ^ n) * (x ^ n)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
118 |
"x ^ (Suc (2*n)) = x * ((x ^ n) * (x ^ n))" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
119 |
by (simp_all add: algebra_simps power_add power2_eq_square power_mult_distrib power_mult) |
23252 | 120 |
|
36871 | 121 |
lemmas normalizing_comm_semiring_1_axioms = |
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
122 |
comm_semiring_1_axioms [normalizer |
36872 | 123 |
semiring ops: normalizing_semiring_ops |
124 |
semiring rules: normalizing_semiring_rules] |
|
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
125 |
|
36871 | 126 |
declaration |
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
127 |
{* Semiring_Normalizer.semiring_funs @{thm normalizing_comm_semiring_1_axioms} *} |
23573 | 128 |
|
36871 | 129 |
end |
23252 | 130 |
|
36871 | 131 |
context comm_ring_1 |
132 |
begin |
|
133 |
||
36872 | 134 |
lemma normalizing_ring_ops: shows "TERM (x- y)" and "TERM (- x)" . |
36871 | 135 |
|
36872 | 136 |
lemma normalizing_ring_rules: |
36845
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
137 |
"- x = (- 1) * x" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
138 |
"x - y = x + (- y)" |
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
139 |
by (simp_all add: diff_minus) |
23252 | 140 |
|
36871 | 141 |
lemmas normalizing_comm_ring_1_axioms = |
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
142 |
comm_ring_1_axioms [normalizer |
36872 | 143 |
semiring ops: normalizing_semiring_ops |
144 |
semiring rules: normalizing_semiring_rules |
|
145 |
ring ops: normalizing_ring_ops |
|
146 |
ring rules: normalizing_ring_rules] |
|
30866 | 147 |
|
36871 | 148 |
declaration |
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
149 |
{* Semiring_Normalizer.semiring_funs @{thm normalizing_comm_ring_1_axioms} *} |
23327 | 150 |
|
36871 | 151 |
end |
152 |
||
36873 | 153 |
context comm_semiring_1_cancel_crossproduct |
36871 | 154 |
begin |
155 |
||
156 |
declare |
|
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
157 |
normalizing_comm_semiring_1_axioms [normalizer del] |
23252 | 158 |
|
36871 | 159 |
lemmas |
36873 | 160 |
normalizing_comm_semiring_1_cancel_crossproduct_axioms = |
161 |
comm_semiring_1_cancel_crossproduct_axioms [normalizer |
|
36872 | 162 |
semiring ops: normalizing_semiring_ops |
163 |
semiring rules: normalizing_semiring_rules |
|
36873 | 164 |
idom rules: crossproduct_noteq add_scale_eq_noteq] |
23252 | 165 |
|
36871 | 166 |
declaration |
36873 | 167 |
{* Semiring_Normalizer.semiring_funs @{thm normalizing_comm_semiring_1_cancel_crossproduct_axioms} *} |
23252 | 168 |
|
36871 | 169 |
end |
23252 | 170 |
|
36871 | 171 |
context idom |
172 |
begin |
|
173 |
||
174 |
declare normalizing_comm_ring_1_axioms [normalizer del] |
|
175 |
||
176 |
lemmas normalizing_idom_axioms = idom_axioms [normalizer |
|
36872 | 177 |
semiring ops: normalizing_semiring_ops |
178 |
semiring rules: normalizing_semiring_rules |
|
179 |
ring ops: normalizing_ring_ops |
|
180 |
ring rules: normalizing_ring_rules |
|
36873 | 181 |
idom rules: crossproduct_noteq add_scale_eq_noteq |
36845
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
182 |
ideal rules: right_minus_eq add_0_iff] |
23252 | 183 |
|
36871 | 184 |
declaration |
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
185 |
{* Semiring_Normalizer.semiring_funs @{thm normalizing_idom_axioms} *} |
23252 | 186 |
|
36871 | 187 |
end |
188 |
||
189 |
context field |
|
190 |
begin |
|
191 |
||
36872 | 192 |
lemma normalizing_field_ops: |
36845
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
193 |
shows "TERM (x / y)" and "TERM (inverse x)" . |
23327 | 194 |
|
36872 | 195 |
lemmas normalizing_field_rules = divide_inverse inverse_eq_divide |
28402 | 196 |
|
36871 | 197 |
lemmas normalizing_field_axioms = |
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
198 |
field_axioms [normalizer |
36872 | 199 |
semiring ops: normalizing_semiring_ops |
200 |
semiring rules: normalizing_semiring_rules |
|
201 |
ring ops: normalizing_ring_ops |
|
202 |
ring rules: normalizing_ring_rules |
|
203 |
field ops: normalizing_field_ops |
|
204 |
field rules: normalizing_field_rules |
|
36873 | 205 |
idom rules: crossproduct_noteq add_scale_eq_noteq |
36845
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
206 |
ideal rules: right_minus_eq add_0_iff] |
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
207 |
|
36871 | 208 |
declaration |
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
209 |
{* Semiring_Normalizer.field_funs @{thm normalizing_field_axioms} *} |
28402 | 210 |
|
36871 | 211 |
end |
212 |
||
36845
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
213 |
hide_fact (open) normalizing_comm_semiring_1_axioms |
36873 | 214 |
normalizing_comm_semiring_1_cancel_crossproduct_axioms normalizing_semiring_ops normalizing_semiring_rules |
36845
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
215 |
|
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
216 |
hide_fact (open) normalizing_comm_ring_1_axioms |
36872 | 217 |
normalizing_idom_axioms normalizing_ring_ops normalizing_ring_rules |
36845
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
218 |
|
36872 | 219 |
hide_fact (open) normalizing_field_axioms normalizing_field_ops normalizing_field_rules |
36845
d778c64fc35d
Add rules directly to the corresponding class locales instead.
hoelzl
parents:
36756
diff
changeset
|
220 |
|
28402 | 221 |
end |