| author | huffman | 
| Sun, 25 Mar 2012 20:15:39 +0200 | |
| changeset 47108 | 2a1953f0d20d | 
| parent 37946 | be3c0df7bb90 | 
| child 48891 | c0eafbd55de3 | 
| 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))"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
37946 
diff
changeset
 | 
119  | 
by (simp_all add: algebra_simps power_add power2_eq_square  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
37946 
diff
changeset
 | 
120  | 
power_mult_distrib power_mult del: one_add_one)  | 
| 23252 | 121  | 
|
| 36871 | 122  | 
lemmas normalizing_comm_semiring_1_axioms =  | 
| 
36756
 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 
haftmann 
parents: 
36753 
diff
changeset
 | 
123  | 
comm_semiring_1_axioms [normalizer  | 
| 36872 | 124  | 
semiring ops: normalizing_semiring_ops  | 
125  | 
semiring rules: normalizing_semiring_rules]  | 
|
| 
36756
 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 
haftmann 
parents: 
36753 
diff
changeset
 | 
126  | 
|
| 36871 | 127  | 
declaration  | 
| 
36756
 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 
haftmann 
parents: 
36753 
diff
changeset
 | 
128  | 
  {* Semiring_Normalizer.semiring_funs @{thm normalizing_comm_semiring_1_axioms} *}
 | 
| 23573 | 129  | 
|
| 36871 | 130  | 
end  | 
| 23252 | 131  | 
|
| 36871 | 132  | 
context comm_ring_1  | 
133  | 
begin  | 
|
134  | 
||
| 36872 | 135  | 
lemma normalizing_ring_ops: shows "TERM (x- y)" and "TERM (- x)" .  | 
| 36871 | 136  | 
|
| 36872 | 137  | 
lemma normalizing_ring_rules:  | 
| 
36845
 
d778c64fc35d
Add rules directly to the corresponding class locales instead.
 
hoelzl 
parents: 
36756 
diff
changeset
 | 
138  | 
"- x = (- 1) * x"  | 
| 
 
d778c64fc35d
Add rules directly to the corresponding class locales instead.
 
hoelzl 
parents: 
36756 
diff
changeset
 | 
139  | 
"x - y = x + (- y)"  | 
| 
 
d778c64fc35d
Add rules directly to the corresponding class locales instead.
 
hoelzl 
parents: 
36756 
diff
changeset
 | 
140  | 
by (simp_all add: diff_minus)  | 
| 23252 | 141  | 
|
| 36871 | 142  | 
lemmas normalizing_comm_ring_1_axioms =  | 
| 
36756
 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 
haftmann 
parents: 
36753 
diff
changeset
 | 
143  | 
comm_ring_1_axioms [normalizer  | 
| 36872 | 144  | 
semiring ops: normalizing_semiring_ops  | 
145  | 
semiring rules: normalizing_semiring_rules  | 
|
146  | 
ring ops: normalizing_ring_ops  | 
|
147  | 
ring rules: normalizing_ring_rules]  | 
|
| 30866 | 148  | 
|
| 36871 | 149  | 
declaration  | 
| 
36756
 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 
haftmann 
parents: 
36753 
diff
changeset
 | 
150  | 
  {* Semiring_Normalizer.semiring_funs @{thm normalizing_comm_ring_1_axioms} *}
 | 
| 23327 | 151  | 
|
| 36871 | 152  | 
end  | 
153  | 
||
| 36873 | 154  | 
context comm_semiring_1_cancel_crossproduct  | 
| 36871 | 155  | 
begin  | 
156  | 
||
157  | 
declare  | 
|
| 
36756
 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 
haftmann 
parents: 
36753 
diff
changeset
 | 
158  | 
normalizing_comm_semiring_1_axioms [normalizer del]  | 
| 23252 | 159  | 
|
| 36871 | 160  | 
lemmas  | 
| 36873 | 161  | 
normalizing_comm_semiring_1_cancel_crossproduct_axioms =  | 
162  | 
comm_semiring_1_cancel_crossproduct_axioms [normalizer  | 
|
| 36872 | 163  | 
semiring ops: normalizing_semiring_ops  | 
164  | 
semiring rules: normalizing_semiring_rules  | 
|
| 36873 | 165  | 
idom rules: crossproduct_noteq add_scale_eq_noteq]  | 
| 23252 | 166  | 
|
| 36871 | 167  | 
declaration  | 
| 36873 | 168  | 
  {* Semiring_Normalizer.semiring_funs @{thm normalizing_comm_semiring_1_cancel_crossproduct_axioms} *}
 | 
| 23252 | 169  | 
|
| 36871 | 170  | 
end  | 
| 23252 | 171  | 
|
| 36871 | 172  | 
context idom  | 
173  | 
begin  | 
|
174  | 
||
175  | 
declare normalizing_comm_ring_1_axioms [normalizer del]  | 
|
176  | 
||
177  | 
lemmas normalizing_idom_axioms = idom_axioms [normalizer  | 
|
| 36872 | 178  | 
semiring ops: normalizing_semiring_ops  | 
179  | 
semiring rules: normalizing_semiring_rules  | 
|
180  | 
ring ops: normalizing_ring_ops  | 
|
181  | 
ring rules: normalizing_ring_rules  | 
|
| 36873 | 182  | 
idom rules: crossproduct_noteq add_scale_eq_noteq  | 
| 
36845
 
d778c64fc35d
Add rules directly to the corresponding class locales instead.
 
hoelzl 
parents: 
36756 
diff
changeset
 | 
183  | 
ideal rules: right_minus_eq add_0_iff]  | 
| 23252 | 184  | 
|
| 36871 | 185  | 
declaration  | 
| 
36756
 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 
haftmann 
parents: 
36753 
diff
changeset
 | 
186  | 
  {* Semiring_Normalizer.semiring_funs @{thm normalizing_idom_axioms} *}
 | 
| 23252 | 187  | 
|
| 36871 | 188  | 
end  | 
189  | 
||
190  | 
context field  | 
|
191  | 
begin  | 
|
192  | 
||
| 36872 | 193  | 
lemma normalizing_field_ops:  | 
| 
36845
 
d778c64fc35d
Add rules directly to the corresponding class locales instead.
 
hoelzl 
parents: 
36756 
diff
changeset
 | 
194  | 
shows "TERM (x / y)" and "TERM (inverse x)" .  | 
| 23327 | 195  | 
|
| 36872 | 196  | 
lemmas normalizing_field_rules = divide_inverse inverse_eq_divide  | 
| 28402 | 197  | 
|
| 36871 | 198  | 
lemmas normalizing_field_axioms =  | 
| 
36756
 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 
haftmann 
parents: 
36753 
diff
changeset
 | 
199  | 
field_axioms [normalizer  | 
| 36872 | 200  | 
semiring ops: normalizing_semiring_ops  | 
201  | 
semiring rules: normalizing_semiring_rules  | 
|
202  | 
ring ops: normalizing_ring_ops  | 
|
203  | 
ring rules: normalizing_ring_rules  | 
|
204  | 
field ops: normalizing_field_ops  | 
|
205  | 
field rules: normalizing_field_rules  | 
|
| 36873 | 206  | 
idom rules: crossproduct_noteq add_scale_eq_noteq  | 
| 
36845
 
d778c64fc35d
Add rules directly to the corresponding class locales instead.
 
hoelzl 
parents: 
36756 
diff
changeset
 | 
207  | 
ideal rules: right_minus_eq add_0_iff]  | 
| 
36756
 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 
haftmann 
parents: 
36753 
diff
changeset
 | 
208  | 
|
| 36871 | 209  | 
declaration  | 
| 
36756
 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 
haftmann 
parents: 
36753 
diff
changeset
 | 
210  | 
  {* Semiring_Normalizer.field_funs @{thm normalizing_field_axioms} *}
 | 
| 28402 | 211  | 
|
| 36871 | 212  | 
end  | 
213  | 
||
| 
36845
 
d778c64fc35d
Add rules directly to the corresponding class locales instead.
 
hoelzl 
parents: 
36756 
diff
changeset
 | 
214  | 
hide_fact (open) normalizing_comm_semiring_1_axioms  | 
| 36873 | 215  | 
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
 | 
216  | 
|
| 
 
d778c64fc35d
Add rules directly to the corresponding class locales instead.
 
hoelzl 
parents: 
36756 
diff
changeset
 | 
217  | 
hide_fact (open) normalizing_comm_ring_1_axioms  | 
| 36872 | 218  | 
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
 | 
219  | 
|
| 36872 | 220  | 
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
 | 
221  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
37946 
diff
changeset
 | 
222  | 
code_modulename SML  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
37946 
diff
changeset
 | 
223  | 
Semiring_Normalization Arith  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
37946 
diff
changeset
 | 
224  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
37946 
diff
changeset
 | 
225  | 
code_modulename OCaml  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
37946 
diff
changeset
 | 
226  | 
Semiring_Normalization Arith  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
37946 
diff
changeset
 | 
227  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
37946 
diff
changeset
 | 
228  | 
code_modulename Haskell  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
37946 
diff
changeset
 | 
229  | 
Semiring_Normalization Arith  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
37946 
diff
changeset
 | 
230  | 
|
| 28402 | 231  | 
end  |