| author | wenzelm | 
| Thu, 20 Aug 2015 17:39:07 +0200 | |
| changeset 60986 | 077f663b6c24 | 
| parent 60758 | d8d85a8172b5 | 
| child 61153 | 3d5e01b427cb | 
| permissions | -rw-r--r-- | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36720diff
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: 
36720diff
changeset | 7 | theory Semiring_Normalization | 
| 36699 
816da1023508
moved nat_arith ot Nat_Numeral: clarified normalizer setup
 haftmann parents: 
36698diff
changeset | 8 | imports Numeral_Simprocs Nat_Transfer | 
| 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: 
36753diff
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: 
36753diff
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: 
36753diff
changeset | 38 | proof | 
| 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 haftmann parents: 
36753diff
changeset | 39 | fix w x y z | 
| 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 haftmann parents: 
36753diff
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: 
36753diff
changeset | 41 | proof | 
| 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 haftmann parents: 
36753diff
changeset | 42 | assume "w * y + x * z = w * z + x * y" | 
| 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 haftmann parents: 
36753diff
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: 
36753diff
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: 
36753diff
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: 
36753diff
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: 
36753diff
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: 
54230diff
changeset | 48 | qed (auto simp add: ac_simps) | 
| 36756 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 haftmann parents: 
36753diff
changeset | 49 | qed | 
| 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 haftmann parents: 
36753diff
changeset | 50 | |
| 36873 | 51 | instance nat :: comm_semiring_1_cancel_crossproduct | 
| 36756 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 haftmann parents: 
36753diff
changeset | 52 | proof | 
| 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 haftmann parents: 
36753diff
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: 
36753diff
changeset | 66 | qed | 
| 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 haftmann parents: 
36753diff
changeset | 67 | |
| 60758 | 68 | text \<open>Semiring normalization proper\<close> | 
| 36871 | 69 | |
| 58826 | 70 | ML_file "Tools/semiring_normalizer.ML" | 
| 23252 | 71 | |
| 36871 | 72 | context comm_semiring_1 | 
| 73 | begin | |
| 74 | ||
| 59554 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 75 | declaration \<open> | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 76 | let | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 77 |   val rules = @{lemma
 | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 78 | "(a * m) + (b * m) = (a + b) * m" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 79 | "(a * m) + m = (a + 1) * m" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 80 | "m + (a * m) = (a + 1) * m" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 81 | "m + m = (1 + 1) * m" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 82 | "0 + a = a" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 83 | "a + 0 = a" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 84 | "a * b = b * a" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 85 | "(a + b) * c = (a * c) + (b * c)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 86 | "0 * a = 0" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 87 | "a * 0 = 0" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 88 | "1 * a = a" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 89 | "a * 1 = a" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 90 | "(lx * ly) * (rx * ry) = (lx * rx) * (ly * ry)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 91 | "(lx * ly) * (rx * ry) = lx * (ly * (rx * ry))" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 92 | "(lx * ly) * (rx * ry) = rx * ((lx * ly) * ry)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 93 | "(lx * ly) * rx = (lx * rx) * ly" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 94 | "(lx * ly) * rx = lx * (ly * rx)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 95 | "lx * (rx * ry) = (lx * rx) * ry" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 96 | "lx * (rx * ry) = rx * (lx * ry)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 97 | "(a + b) + (c + d) = (a + c) + (b + d)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 98 | "(a + b) + c = a + (b + c)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 99 | "a + (c + d) = c + (a + d)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 100 | "(a + b) + c = (a + c) + b" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 101 | "a + c = c + a" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 102 | "a + (c + d) = (a + c) + d" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 103 | "(x ^ p) * (x ^ q) = x ^ (p + q)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 104 | "x * (x ^ q) = x ^ (Suc q)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 105 | "(x ^ q) * x = x ^ (Suc q)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 106 | "x * x = x\<^sup>2" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 107 | "(x * y) ^ q = (x ^ q) * (y ^ q)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 108 | "(x ^ p) ^ q = x ^ (p * q)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 109 | "x ^ 0 = 1" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 110 | "x ^ 1 = x" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 111 | "x * (y + z) = (x * y) + (x * z)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 112 | "x ^ (Suc q) = x * (x ^ q)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 113 | "x ^ (2*n) = (x ^ n) * (x ^ n)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 114 | by (simp_all add: algebra_simps power_add power2_eq_square | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 115 | power_mult_distrib power_mult del: one_add_one)} | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 116 | in | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 117 |   Semiring_Normalizer.declare @{thm comm_semiring_1_axioms}
 | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 118 |     {semiring = ([@{cpat "?x + ?y"}, @{cpat "?x * ?y"}, @{cpat "?x ^ ?n"}, @{cpat 0}, @{cpat 1}],
 | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 119 | rules), ring = ([], []), field = ([], []), idom = [], ideal = []} | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 120 | end\<close> | 
| 36756 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 haftmann parents: 
36753diff
changeset | 121 | |
| 36871 | 122 | end | 
| 23252 | 123 | |
| 36871 | 124 | context comm_ring_1 | 
| 125 | begin | |
| 126 | ||
| 59554 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 127 | declaration \<open> | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 128 | let | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 129 |   val rules = @{lemma
 | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 130 | "- x = (- 1) * x" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 131 | "x - y = x + (- y)" | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 132 | by simp_all} | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 133 | in | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 134 |   Semiring_Normalizer.declare @{thm comm_ring_1_axioms}
 | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 135 |     {semiring = Semiring_Normalizer.the_semiring @{context} @{thm comm_semiring_1_axioms},
 | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 136 |       ring = ([@{cpat "?x - ?y"}, @{cpat "- ?x"}], rules), field = ([], []), idom = [], ideal = []}
 | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 137 | end\<close> | 
| 30866 | 138 | |
| 36871 | 139 | end | 
| 140 | ||
| 36873 | 141 | context comm_semiring_1_cancel_crossproduct | 
| 36871 | 142 | begin | 
| 143 | ||
| 59553 
e87974cd9b86
explicit declaration allows cumulative declaration
 haftmann parents: 
59551diff
changeset | 144 | declaration \<open>Semiring_Normalizer.declare @{thm comm_semiring_1_cancel_crossproduct_axioms}
 | 
| 
e87974cd9b86
explicit declaration allows cumulative declaration
 haftmann parents: 
59551diff
changeset | 145 |   {semiring = Semiring_Normalizer.the_semiring @{context} @{thm comm_semiring_1_axioms},
 | 
| 
e87974cd9b86
explicit declaration allows cumulative declaration
 haftmann parents: 
59551diff
changeset | 146 |     ring = ([], []), field = ([], []), idom = @{thms crossproduct_noteq add_scale_eq_noteq}, ideal = []}\<close>
 | 
| 23252 | 147 | |
| 36871 | 148 | end | 
| 23252 | 149 | |
| 36871 | 150 | context idom | 
| 151 | begin | |
| 152 | ||
| 59553 
e87974cd9b86
explicit declaration allows cumulative declaration
 haftmann parents: 
59551diff
changeset | 153 | declaration \<open>Semiring_Normalizer.declare @{thm idom_axioms}
 | 
| 
e87974cd9b86
explicit declaration allows cumulative declaration
 haftmann parents: 
59551diff
changeset | 154 |   {semiring = Semiring_Normalizer.the_semiring @{context} @{thm comm_ring_1_axioms},
 | 
| 
e87974cd9b86
explicit declaration allows cumulative declaration
 haftmann parents: 
59551diff
changeset | 155 |     ring = Semiring_Normalizer.the_ring @{context} @{thm comm_ring_1_axioms},
 | 
| 
e87974cd9b86
explicit declaration allows cumulative declaration
 haftmann parents: 
59551diff
changeset | 156 |     field = ([], []), idom = @{thms crossproduct_noteq add_scale_eq_noteq},
 | 
| 
e87974cd9b86
explicit declaration allows cumulative declaration
 haftmann parents: 
59551diff
changeset | 157 |     ideal = @{thms right_minus_eq add_0_iff}}\<close>
 | 
| 23252 | 158 | |
| 36871 | 159 | end | 
| 160 | ||
| 161 | context field | |
| 162 | begin | |
| 163 | ||
| 59553 
e87974cd9b86
explicit declaration allows cumulative declaration
 haftmann parents: 
59551diff
changeset | 164 | declaration \<open>Semiring_Normalizer.declare @{thm field_axioms}
 | 
| 
e87974cd9b86
explicit declaration allows cumulative declaration
 haftmann parents: 
59551diff
changeset | 165 |   {semiring = Semiring_Normalizer.the_semiring @{context} @{thm idom_axioms},
 | 
| 
e87974cd9b86
explicit declaration allows cumulative declaration
 haftmann parents: 
59551diff
changeset | 166 |     ring = Semiring_Normalizer.the_ring @{context} @{thm idom_axioms},
 | 
| 59554 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59553diff
changeset | 167 |     field = ([@{cpat "?x / ?y"}, @{cpat "inverse ?x"}], @{thms divide_inverse inverse_eq_divide}),
 | 
| 59553 
e87974cd9b86
explicit declaration allows cumulative declaration
 haftmann parents: 
59551diff
changeset | 168 |     idom = Semiring_Normalizer.the_idom @{context} @{thm idom_axioms},
 | 
| 
e87974cd9b86
explicit declaration allows cumulative declaration
 haftmann parents: 
59551diff
changeset | 169 |     ideal = Semiring_Normalizer.the_ideal @{context} @{thm idom_axioms}}\<close>
 | 
| 36756 
c1ae8a0b4265
moved normalization proof tool infrastructure to canonical algebraic classes
 haftmann parents: 
36753diff
changeset | 170 | |
| 36871 | 171 | end | 
| 172 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
48891diff
changeset | 173 | code_identifier | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
48891diff
changeset | 174 | 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: 
37946diff
changeset | 175 | |
| 28402 | 176 | end |