author | wenzelm |
Mon, 12 Apr 2021 14:14:47 +0200 | |
changeset 73563 | 55b66a45bc94 |
parent 70332 | 315489d836d8 |
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 |
||
60758 | 5 |
section \<open>Semiring normalization\<close> |
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 |
66836 | 8 |
imports Numeral_Simprocs |
23252 | 9 |
begin |
10 |
||
60758 | 11 |
text \<open>Prelude\<close> |
36873 | 12 |
|
13 |
class comm_semiring_1_cancel_crossproduct = comm_semiring_1_cancel + |
|
14 |
assumes crossproduct_eq: "w * y + x * z = w * z + x * y \<longleftrightarrow> w = x \<or> y = z" |
|
15 |
begin |
|
16 |
||
17 |
lemma crossproduct_noteq: |
|
18 |
"a \<noteq> b \<and> c \<noteq> d \<longleftrightarrow> a * c + b * d \<noteq> a * d + b * c" |
|
19 |
by (simp add: crossproduct_eq) |
|
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
20 |
|
36873 | 21 |
lemma add_scale_eq_noteq: |
22 |
"r \<noteq> 0 \<Longrightarrow> a = b \<and> c \<noteq> d \<Longrightarrow> a + r * c \<noteq> b + r * d" |
|
23 |
proof (rule notI) |
|
24 |
assume nz: "r\<noteq> 0" and cnd: "a = b \<and> c\<noteq>d" |
|
25 |
and eq: "a + (r * c) = b + (r * d)" |
|
26 |
have "(0 * d) + (r * c) = (0 * c) + (r * d)" |
|
59557 | 27 |
using add_left_imp_eq eq mult_zero_left by (simp add: cnd) |
36873 | 28 |
then show False using crossproduct_eq [of 0 d] nz cnd by simp |
29 |
qed |
|
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
30 |
|
36873 | 31 |
lemma add_0_iff: |
32 |
"b = b + a \<longleftrightarrow> a = 0" |
|
59557 | 33 |
using add_left_imp_eq [of b a 0] by auto |
36873 | 34 |
|
35 |
end |
|
36 |
||
37946 | 37 |
subclass (in idom) comm_semiring_1_cancel_crossproduct |
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
38 |
proof |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
39 |
fix w x y z |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
40 |
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
|
41 |
proof |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
42 |
assume "w * y + x * z = w * z + x * y" |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
43 |
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
|
44 |
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
|
45 |
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
|
46 |
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
|
47 |
then show "w = x \<or> y = z" by auto |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
54230
diff
changeset
|
48 |
qed (auto simp add: ac_simps) |
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
49 |
qed |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
50 |
|
36873 | 51 |
instance nat :: comm_semiring_1_cancel_crossproduct |
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
52 |
proof |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
53 |
fix w x y z :: nat |
36873 | 54 |
have aux: "\<And>y z. y < z \<Longrightarrow> w * y + x * z = w * z + x * y \<Longrightarrow> w = x" |
55 |
proof - |
|
56 |
fix y z :: nat |
|
57 |
assume "y < z" then have "\<exists>k. z = y + k \<and> k \<noteq> 0" by (intro exI [of _ "z - y"]) auto |
|
58 |
then obtain k where "z = y + k" and "k \<noteq> 0" by blast |
|
59 |
assume "w * y + x * z = w * z + x * y" |
|
60758 | 60 |
then have "(w * y + x * y) + x * k = (w * y + x * y) + w * k" by (simp add: \<open>z = y + k\<close> algebra_simps) |
36873 | 61 |
then have "x * k = w * k" by simp |
60758 | 62 |
then show "w = x" using \<open>k \<noteq> 0\<close> by simp |
36873 | 63 |
qed |
64 |
show "w * y + x * z = w * z + x * y \<longleftrightarrow> w = x \<or> y = z" |
|
65 |
by (auto simp add: neq_iff dest!: aux) |
|
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
66 |
qed |
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
67 |
|
60758 | 68 |
text \<open>Semiring normalization proper\<close> |
36871 | 69 |
|
69605 | 70 |
ML_file \<open>Tools/semiring_normalizer.ML\<close> |
23252 | 71 |
|
36871 | 72 |
context comm_semiring_1 |
73 |
begin |
|
74 |
||
70332 | 75 |
lemma semiring_normalization_rules [no_atp]: |
61153 | 76 |
"(a * m) + (b * m) = (a + b) * m" |
77 |
"(a * m) + m = (a + 1) * m" |
|
78 |
"m + (a * m) = (a + 1) * m" |
|
79 |
"m + m = (1 + 1) * m" |
|
80 |
"0 + a = a" |
|
81 |
"a + 0 = a" |
|
82 |
"a * b = b * a" |
|
83 |
"(a + b) * c = (a * c) + (b * c)" |
|
84 |
"0 * a = 0" |
|
85 |
"a * 0 = 0" |
|
86 |
"1 * a = a" |
|
87 |
"a * 1 = a" |
|
88 |
"(lx * ly) * (rx * ry) = (lx * rx) * (ly * ry)" |
|
89 |
"(lx * ly) * (rx * ry) = lx * (ly * (rx * ry))" |
|
90 |
"(lx * ly) * (rx * ry) = rx * ((lx * ly) * ry)" |
|
91 |
"(lx * ly) * rx = (lx * rx) * ly" |
|
92 |
"(lx * ly) * rx = lx * (ly * rx)" |
|
93 |
"lx * (rx * ry) = (lx * rx) * ry" |
|
94 |
"lx * (rx * ry) = rx * (lx * ry)" |
|
95 |
"(a + b) + (c + d) = (a + c) + (b + d)" |
|
96 |
"(a + b) + c = a + (b + c)" |
|
97 |
"a + (c + d) = c + (a + d)" |
|
98 |
"(a + b) + c = (a + c) + b" |
|
99 |
"a + c = c + a" |
|
100 |
"a + (c + d) = (a + c) + d" |
|
101 |
"(x ^ p) * (x ^ q) = x ^ (p + q)" |
|
102 |
"x * (x ^ q) = x ^ (Suc q)" |
|
103 |
"(x ^ q) * x = x ^ (Suc q)" |
|
104 |
"x * x = x\<^sup>2" |
|
105 |
"(x * y) ^ q = (x ^ q) * (y ^ q)" |
|
106 |
"(x ^ p) ^ q = x ^ (p * q)" |
|
107 |
"x ^ 0 = 1" |
|
108 |
"x ^ 1 = x" |
|
109 |
"x * (y + z) = (x * y) + (x * z)" |
|
110 |
"x ^ (Suc q) = x * (x ^ q)" |
|
111 |
"x ^ (2*n) = (x ^ n) * (x ^ n)" |
|
112 |
by (simp_all add: algebra_simps power_add power2_eq_square |
|
113 |
power_mult_distrib power_mult del: one_add_one) |
|
114 |
||
115 |
local_setup \<open> |
|
59554
4044f53326c9
inlined rules to free user-space from technical names
haftmann
parents:
59553
diff
changeset
|
116 |
Semiring_Normalizer.declare @{thm comm_semiring_1_axioms} |
69593 | 117 |
{semiring = ([\<^term>\<open>x + y\<close>, \<^term>\<open>x * y\<close>, \<^term>\<open>x ^ n\<close>, \<^term>\<open>0\<close>, \<^term>\<open>1\<close>], |
61153 | 118 |
@{thms semiring_normalization_rules}), |
119 |
ring = ([], []), |
|
120 |
field = ([], []), |
|
121 |
idom = [], |
|
122 |
ideal = []} |
|
123 |
\<close> |
|
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
124 |
|
36871 | 125 |
end |
23252 | 126 |
|
36871 | 127 |
context comm_ring_1 |
128 |
begin |
|
129 |
||
70332 | 130 |
lemma ring_normalization_rules [no_atp]: |
61153 | 131 |
"- x = (- 1) * x" |
132 |
"x - y = x + (- y)" |
|
133 |
by simp_all |
|
134 |
||
135 |
local_setup \<open> |
|
59554
4044f53326c9
inlined rules to free user-space from technical names
haftmann
parents:
59553
diff
changeset
|
136 |
Semiring_Normalizer.declare @{thm comm_ring_1_axioms} |
69593 | 137 |
{semiring = ([\<^term>\<open>x + y\<close>, \<^term>\<open>x * y\<close>, \<^term>\<open>x ^ n\<close>, \<^term>\<open>0\<close>, \<^term>\<open>1\<close>], |
61153 | 138 |
@{thms semiring_normalization_rules}), |
69593 | 139 |
ring = ([\<^term>\<open>x - y\<close>, \<^term>\<open>- x\<close>], @{thms ring_normalization_rules}), |
61153 | 140 |
field = ([], []), |
141 |
idom = [], |
|
142 |
ideal = []} |
|
143 |
\<close> |
|
30866 | 144 |
|
36871 | 145 |
end |
146 |
||
36873 | 147 |
context comm_semiring_1_cancel_crossproduct |
36871 | 148 |
begin |
149 |
||
61153 | 150 |
local_setup \<open> |
151 |
Semiring_Normalizer.declare @{thm comm_semiring_1_cancel_crossproduct_axioms} |
|
69593 | 152 |
{semiring = ([\<^term>\<open>x + y\<close>, \<^term>\<open>x * y\<close>, \<^term>\<open>x ^ n\<close>, \<^term>\<open>0\<close>, \<^term>\<open>1\<close>], |
61153 | 153 |
@{thms semiring_normalization_rules}), |
154 |
ring = ([], []), |
|
155 |
field = ([], []), |
|
156 |
idom = @{thms crossproduct_noteq add_scale_eq_noteq}, |
|
157 |
ideal = []} |
|
158 |
\<close> |
|
23252 | 159 |
|
36871 | 160 |
end |
23252 | 161 |
|
36871 | 162 |
context idom |
163 |
begin |
|
164 |
||
61153 | 165 |
local_setup \<open> |
166 |
Semiring_Normalizer.declare @{thm idom_axioms} |
|
69593 | 167 |
{semiring = ([\<^term>\<open>x + y\<close>, \<^term>\<open>x * y\<close>, \<^term>\<open>x ^ n\<close>, \<^term>\<open>0\<close>, \<^term>\<open>1\<close>], |
61153 | 168 |
@{thms semiring_normalization_rules}), |
69593 | 169 |
ring = ([\<^term>\<open>x - y\<close>, \<^term>\<open>- x\<close>], @{thms ring_normalization_rules}), |
61153 | 170 |
field = ([], []), |
171 |
idom = @{thms crossproduct_noteq add_scale_eq_noteq}, |
|
172 |
ideal = @{thms right_minus_eq add_0_iff}} |
|
173 |
\<close> |
|
23252 | 174 |
|
36871 | 175 |
end |
176 |
||
177 |
context field |
|
178 |
begin |
|
179 |
||
61153 | 180 |
local_setup \<open> |
181 |
Semiring_Normalizer.declare @{thm field_axioms} |
|
69593 | 182 |
{semiring = ([\<^term>\<open>x + y\<close>, \<^term>\<open>x * y\<close>, \<^term>\<open>x ^ n\<close>, \<^term>\<open>0\<close>, \<^term>\<open>1\<close>], |
61153 | 183 |
@{thms semiring_normalization_rules}), |
69593 | 184 |
ring = ([\<^term>\<open>x - y\<close>, \<^term>\<open>- x\<close>], @{thms ring_normalization_rules}), |
185 |
field = ([\<^term>\<open>x / y\<close>, \<^term>\<open>inverse x\<close>], @{thms divide_inverse inverse_eq_divide}), |
|
61153 | 186 |
idom = @{thms crossproduct_noteq add_scale_eq_noteq}, |
187 |
ideal = @{thms right_minus_eq add_0_iff}} |
|
188 |
\<close> |
|
36756
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
haftmann
parents:
36753
diff
changeset
|
189 |
|
36871 | 190 |
end |
191 |
||
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
48891
diff
changeset
|
192 |
code_identifier |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
48891
diff
changeset
|
193 |
code_module Semiring_Normalization \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
37946
diff
changeset
|
194 |
|
28402 | 195 |
end |