| author | wenzelm | 
| Thu, 28 Aug 2008 19:29:57 +0200 | |
| changeset 28044 | e4b569b53e10 | 
| parent 27651 | 16a26996c30e | 
| child 29269 | 5c25a2012975 | 
| permissions | -rw-r--r-- | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 1 | (* | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 2 | Title: The algebraic hierarchy of rings as axiomatic classes | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 3 | Id: $Id$ | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 4 | Author: Clemens Ballarin, started 9 December 1996 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 5 | Copyright: Clemens Ballarin | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 6 | *) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 7 | |
| 27542 | 8 | header {* The algebraic hierarchy of rings as type classes *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 9 | |
| 27542 | 10 | theory Ring2 | 
| 11 | imports Main | |
| 12 | begin | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 13 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 14 | subsection {* Ring axioms *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 15 | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27542diff
changeset | 16 | class ring = zero + one + plus + minus + uminus + times + inverse + power + Ring_and_Field.dvd + | 
| 27542 | 17 | assumes a_assoc: "(a + b) + c = a + (b + c)" | 
| 18 | and l_zero: "0 + a = a" | |
| 19 | and l_neg: "(-a) + a = 0" | |
| 20 | and a_comm: "a + b = b + a" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 21 | |
| 27542 | 22 | assumes m_assoc: "(a * b) * c = a * (b * c)" | 
| 23 | and l_one: "1 * a = a" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 24 | |
| 27542 | 25 | assumes l_distr: "(a + b) * c = a * c + b * c" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 26 | |
| 27542 | 27 | assumes m_comm: "a * b = b * a" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 28 | |
| 27542 | 29 | assumes minus_def: "a - b = a + (-b)" | 
| 30 | and inverse_def: "inverse a = (if a dvd 1 then THE x. a*x = 1 else 0)" | |
| 31 | and divide_def: "a / b = a * inverse b" | |
| 32 | and power_0 [simp]: "a ^ 0 = 1" | |
| 33 | and power_Suc [simp]: "a ^ Suc n = a ^ n * a" | |
| 34 | begin | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 35 | |
| 27542 | 36 | definition assoc :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "assoc" 50) where | 
| 37 | assoc_def: "a assoc b \<longleftrightarrow> a dvd b & b dvd a" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 38 | |
| 27542 | 39 | definition irred :: "'a \<Rightarrow> bool" where | 
| 40 | irred_def: "irred a \<longleftrightarrow> a ~= 0 & ~ a dvd 1 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 41 | & (ALL d. d dvd a --> d dvd 1 | a dvd d)" | 
| 27542 | 42 | |
| 43 | definition prime :: "'a \<Rightarrow> bool" where | |
| 44 | prime_def: "prime p \<longleftrightarrow> p ~= 0 & ~ p dvd 1 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 45 | & (ALL a b. p dvd (a*b) --> p dvd a | p dvd b)" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 46 | |
| 27542 | 47 | end | 
| 48 | ||
| 49 | ||
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 50 | subsection {* Integral domains *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 51 | |
| 27542 | 52 | class "domain" = ring + | 
| 53 | assumes one_not_zero: "1 ~= 0" | |
| 54 | and integral: "a * b = 0 ==> a = 0 | b = 0" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 55 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 56 | subsection {* Factorial domains *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 57 | |
| 27542 | 58 | class factorial = "domain" + | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 59 | (* | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 60 | Proper definition using divisor chain condition currently not supported. | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 61 |   factorial_divisor:    "wf {(a, b). a dvd b & ~ (b dvd a)}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 62 | *) | 
| 27542 | 63 | assumes factorial_divisor: "True" | 
| 64 | and factorial_prime: "irred a ==> prime a" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 65 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 66 | subsection {* Euclidean domains *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 67 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 68 | (* | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 69 | axclass | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 70 | euclidean < "domain" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 71 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 72 |   euclidean_ax:  "b ~= 0 ==> Ex (% (q, r, e_size::('a::ringS)=>nat).
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 73 | a = b * q + r & e_size r < e_size b)" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 74 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 75 | Nothing has been proved about Euclidean domains, yet. | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 76 | Design question: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 77 | Fix quo, rem and e_size as constants that are axiomatised with | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 78 | euclidean_ax? | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 79 | - advantage: more pragmatic and easier to use | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 80 | - disadvantage: for every type, one definition of quo and rem will | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 81 | be fixed, users may want to use differing ones; | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 82 | also, it seems not possible to prove that fields are euclidean | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 83 | domains, because that would require generic (type-independent) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 84 | definitions of quo and rem. | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 85 | *) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 86 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 87 | subsection {* Fields *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 88 | |
| 27542 | 89 | class field = ring + | 
| 90 | assumes field_one_not_zero: "1 ~= 0" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 91 | (* Avoid a common superclass as the first thing we will | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 92 | prove about fields is that they are domains. *) | 
| 27542 | 93 | and field_ax: "a ~= 0 ==> a dvd 1" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 94 | |
| 21423 | 95 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 96 | section {* Basic facts *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 97 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 98 | subsection {* Normaliser for rings *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 99 | |
| 21423 | 100 | (* derived rewrite rules *) | 
| 101 | ||
| 102 | lemma a_lcomm: "(a::'a::ring)+(b+c) = b+(a+c)" | |
| 103 | apply (rule a_comm [THEN trans]) | |
| 104 | apply (rule a_assoc [THEN trans]) | |
| 105 | apply (rule a_comm [THEN arg_cong]) | |
| 106 | done | |
| 107 | ||
| 108 | lemma r_zero: "(a::'a::ring) + 0 = a" | |
| 109 | apply (rule a_comm [THEN trans]) | |
| 110 | apply (rule l_zero) | |
| 111 | done | |
| 112 | ||
| 113 | lemma r_neg: "(a::'a::ring) + (-a) = 0" | |
| 114 | apply (rule a_comm [THEN trans]) | |
| 115 | apply (rule l_neg) | |
| 116 | done | |
| 117 | ||
| 118 | lemma r_neg2: "(a::'a::ring) + (-a + b) = b" | |
| 119 | apply (rule a_assoc [symmetric, THEN trans]) | |
| 120 | apply (simp add: r_neg l_zero) | |
| 121 | done | |
| 122 | ||
| 123 | lemma r_neg1: "-(a::'a::ring) + (a + b) = b" | |
| 124 | apply (rule a_assoc [symmetric, THEN trans]) | |
| 125 | apply (simp add: l_neg l_zero) | |
| 126 | done | |
| 127 | ||
| 128 | ||
| 129 | (* auxiliary *) | |
| 130 | ||
| 131 | lemma a_lcancel: "!! a::'a::ring. a + b = a + c ==> b = c" | |
| 132 | apply (rule box_equals) | |
| 133 | prefer 2 | |
| 134 | apply (rule l_zero) | |
| 135 | prefer 2 | |
| 136 | apply (rule l_zero) | |
| 137 | apply (rule_tac a1 = a in l_neg [THEN subst]) | |
| 138 | apply (simp add: a_assoc) | |
| 139 | done | |
| 140 | ||
| 141 | lemma minus_add: "-((a::'a::ring) + b) = (-a) + (-b)" | |
| 142 | apply (rule_tac a = "a + b" in a_lcancel) | |
| 143 | apply (simp add: r_neg l_neg l_zero a_assoc a_comm a_lcomm) | |
| 144 | done | |
| 145 | ||
| 146 | lemma minus_minus: "-(-(a::'a::ring)) = a" | |
| 147 | apply (rule a_lcancel) | |
| 148 | apply (rule r_neg [THEN trans]) | |
| 149 | apply (rule l_neg [symmetric]) | |
| 150 | done | |
| 151 | ||
| 152 | lemma minus0: "- 0 = (0::'a::ring)" | |
| 153 | apply (rule a_lcancel) | |
| 154 | apply (rule r_neg [THEN trans]) | |
| 155 | apply (rule l_zero [symmetric]) | |
| 156 | done | |
| 157 | ||
| 158 | ||
| 159 | (* derived rules for multiplication *) | |
| 160 | ||
| 161 | lemma m_lcomm: "(a::'a::ring)*(b*c) = b*(a*c)" | |
| 162 | apply (rule m_comm [THEN trans]) | |
| 163 | apply (rule m_assoc [THEN trans]) | |
| 164 | apply (rule m_comm [THEN arg_cong]) | |
| 165 | done | |
| 166 | ||
| 167 | lemma r_one: "(a::'a::ring) * 1 = a" | |
| 168 | apply (rule m_comm [THEN trans]) | |
| 169 | apply (rule l_one) | |
| 170 | done | |
| 171 | ||
| 172 | lemma r_distr: "(a::'a::ring) * (b + c) = a * b + a * c" | |
| 173 | apply (rule m_comm [THEN trans]) | |
| 174 | apply (rule l_distr [THEN trans]) | |
| 175 | apply (simp add: m_comm) | |
| 176 | done | |
| 177 | ||
| 178 | ||
| 179 | (* the following proof is from Jacobson, Basic Algebra I, pp. 88-89 *) | |
| 180 | lemma l_null: "0 * (a::'a::ring) = 0" | |
| 181 | apply (rule a_lcancel) | |
| 182 | apply (rule l_distr [symmetric, THEN trans]) | |
| 183 | apply (simp add: r_zero) | |
| 184 | done | |
| 185 | ||
| 186 | lemma r_null: "(a::'a::ring) * 0 = 0" | |
| 187 | apply (rule m_comm [THEN trans]) | |
| 188 | apply (rule l_null) | |
| 189 | done | |
| 190 | ||
| 191 | lemma l_minus: "(-(a::'a::ring)) * b = - (a * b)" | |
| 192 | apply (rule a_lcancel) | |
| 193 | apply (rule r_neg [symmetric, THEN [2] trans]) | |
| 194 | apply (rule l_distr [symmetric, THEN trans]) | |
| 195 | apply (simp add: l_null r_neg) | |
| 196 | done | |
| 197 | ||
| 198 | lemma r_minus: "(a::'a::ring) * (-b) = - (a * b)" | |
| 199 | apply (rule a_lcancel) | |
| 200 | apply (rule r_neg [symmetric, THEN [2] trans]) | |
| 201 | apply (rule r_distr [symmetric, THEN trans]) | |
| 202 | apply (simp add: r_null r_neg) | |
| 203 | done | |
| 204 | ||
| 205 | (*** Term order for commutative rings ***) | |
| 206 | ||
| 207 | ML {*
 | |
| 208 | fun ring_ord (Const (a, _)) = | |
| 209 | find_index (fn a' => a = a') | |
| 22997 | 210 |       [@{const_name HOL.zero}, @{const_name HOL.plus}, @{const_name HOL.uminus},
 | 
| 211 |         @{const_name HOL.minus}, @{const_name HOL.one}, @{const_name HOL.times}]
 | |
| 21423 | 212 | | ring_ord _ = ~1; | 
| 213 | ||
| 214 | fun termless_ring (a, b) = (Term.term_lpo ring_ord (a, b) = LESS); | |
| 215 | ||
| 216 | val ring_ss = HOL_basic_ss settermless termless_ring addsimps | |
| 217 | [thm "a_assoc", thm "l_zero", thm "l_neg", thm "a_comm", thm "m_assoc", | |
| 218 | thm "l_one", thm "l_distr", thm "m_comm", thm "minus_def", | |
| 219 | thm "r_zero", thm "r_neg", thm "r_neg2", thm "r_neg1", thm "minus_add", | |
| 220 | thm "minus_minus", thm "minus0", thm "a_lcomm", thm "m_lcomm", (*thm "r_one",*) | |
| 221 | thm "r_distr", thm "l_null", thm "r_null", thm "l_minus", thm "r_minus"]; | |
| 222 | *} (* Note: r_one is not necessary in ring_ss *) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 223 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 224 | method_setup ring = | 
| 21588 | 225 |   {* Method.no_args (Method.SIMPLE_METHOD' (full_simp_tac ring_ss)) *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 226 |   {* computes distributive normal form in rings *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 227 | |
| 21423 | 228 | lemmas ring_simps = | 
| 229 | l_zero r_zero l_neg r_neg minus_minus minus0 | |
| 230 | l_one r_one l_null r_null l_minus r_minus | |
| 231 | ||
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 232 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 233 | subsection {* Rings and the summation operator *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 234 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 235 | (* Basic facts --- move to HOL!!! *) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 236 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 237 | (* needed because natsum_cong (below) disables atMost_0 *) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 238 | lemma natsum_0 [simp]: "setsum f {..(0::nat)} = (f 0::'a::comm_monoid_add)"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 239 | by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 240 | (* | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 241 | lemma natsum_Suc [simp]: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 242 |   "setsum f {..Suc n} = (f (Suc n) + setsum f {..n}::'a::comm_monoid_add)"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 243 | by (simp add: atMost_Suc) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 244 | *) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 245 | lemma natsum_Suc2: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 246 |   "setsum f {..Suc n} = (f 0::'a::comm_monoid_add) + (setsum (%i. f (Suc i)) {..n})"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 247 | proof (induct n) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 248 | case 0 show ?case by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 249 | next | 
| 22384 
33a46e6c7f04
prefix of class interpretation not mandatory any longer
 haftmann parents: 
21588diff
changeset | 250 | case Suc thus ?case by (simp add: add_assoc) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 251 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 252 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 253 | lemma natsum_cong [cong]: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 254 | "!!k. [| j = k; !!i::nat. i <= k ==> f i = (g i::'a::comm_monoid_add) |] ==> | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 255 |         setsum f {..j} = setsum g {..k}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 256 | by (induct j) auto | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 257 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 258 | lemma natsum_zero [simp]: "setsum (%i. 0) {..n::nat} = (0::'a::comm_monoid_add)"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 259 | by (induct n) simp_all | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 260 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 261 | lemma natsum_add [simp]: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 262 | "!!f::nat=>'a::comm_monoid_add. | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 263 |    setsum (%i. f i + g i) {..n::nat} = setsum f {..n} + setsum g {..n}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 264 | by (induct n) (simp_all add: add_ac) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 265 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 266 | (* Facts specific to rings *) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 267 | |
| 27542 | 268 | subclass (in ring) comm_monoid_add | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 269 | proof | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 270 | fix x y z | 
| 27542 | 271 | show "x + y = y + x" by (rule a_comm) | 
| 272 | show "(x + y) + z = x + (y + z)" by (rule a_assoc) | |
| 273 | show "0 + x = x" by (rule l_zero) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 274 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 275 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 276 | ML {*
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 277 | local | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 278 | val lhss = | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 279 | ["t + u::'a::ring", | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 280 | "t - u::'a::ring", | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 281 | "t * u::'a::ring", | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 282 | "- t::'a::ring"]; | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 283 | fun proc ss t = | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 284 | let val rew = Goal.prove (Simplifier.the_context ss) [] [] | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 285 | (HOLogic.mk_Trueprop | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 286 |               (HOLogic.mk_eq (t, Var (("x", Term.maxidx_of_term t + 1), fastype_of t))))
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 287 | (fn _ => simp_tac (Simplifier.inherit_context ss ring_ss) 1) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 288 | |> mk_meta_eq; | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 289 | val (t', u) = Logic.dest_equals (Thm.prop_of rew); | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 290 | in if t' aconv u | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 291 | then NONE | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 292 | else SOME rew | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 293 | end; | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 294 | in | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 295 | val ring_simproc = Simplifier.simproc (the_context ()) "ring" lhss (K proc); | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 296 | end; | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 297 | *} | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 298 | |
| 26480 | 299 | ML {* Addsimprocs [ring_simproc] *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 300 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 301 | lemma natsum_ldistr: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 302 |   "!!a::'a::ring. setsum f {..n::nat} * a = setsum (%i. f i * a) {..n}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 303 | by (induct n) simp_all | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 304 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 305 | lemma natsum_rdistr: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 306 |   "!!a::'a::ring. a * setsum f {..n::nat} = setsum (%i. a * f i) {..n}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 307 | by (induct n) simp_all | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 308 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 309 | subsection {* Integral Domains *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 310 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 311 | declare one_not_zero [simp] | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 312 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 313 | lemma zero_not_one [simp]: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 314 | "0 ~= (1::'a::domain)" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 315 | by (rule not_sym) simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 316 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 317 | lemma integral_iff: (* not by default a simp rule! *) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 318 | "(a * b = (0::'a::domain)) = (a = 0 | b = 0)" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 319 | proof | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 320 | assume "a * b = 0" then show "a = 0 | b = 0" by (simp add: integral) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 321 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 322 | assume "a = 0 | b = 0" then show "a * b = 0" by auto | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 323 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 324 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 325 | (* | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 326 | lemma "(a::'a::ring) - (a - b) = b" apply simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 327 | simproc seems to fail on this example (fixed with new term order) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 328 | *) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 329 | (* | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 330 | lemma bug: "(b::'a::ring) - (b - a) = a" by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 331 | simproc for rings cannot prove "(a::'a::ring) - (a - b) = b" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 332 | *) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 333 | lemma m_lcancel: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 334 | assumes prem: "(a::'a::domain) ~= 0" shows conc: "(a * b = a * c) = (b = c)" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 335 | proof | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 336 | assume eq: "a * b = a * c" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 337 | then have "a * (b - c) = 0" by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 338 | then have "a = 0 | (b - c) = 0" by (simp only: integral_iff) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 339 | with prem have "b - c = 0" by auto | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 340 | then have "b = b - (b - c)" by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 341 | also have "b - (b - c) = c" by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 342 | finally show "b = c" . | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 343 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 344 | assume "b = c" then show "a * b = a * c" by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 345 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 346 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 347 | lemma m_rcancel: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 348 | "(a::'a::domain) ~= 0 ==> (b * a = c * a) = (b = c)" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 349 | by (simp add: m_lcancel) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 350 | |
| 27542 | 351 | declare power_Suc [simp] | 
| 21416 | 352 | |
| 353 | lemma power_one [simp]: | |
| 354 | "1 ^ n = (1::'a::ring)" by (induct n) simp_all | |
| 355 | ||
| 356 | lemma power_zero [simp]: | |
| 357 | "n \<noteq> 0 \<Longrightarrow> 0 ^ n = (0::'a::ring)" by (induct n) simp_all | |
| 358 | ||
| 359 | lemma power_mult [simp]: | |
| 360 | "(a::'a::ring) ^ m * a ^ n = a ^ (m + n)" | |
| 361 | by (induct m) simp_all | |
| 362 | ||
| 363 | ||
| 364 | section "Divisibility" | |
| 365 | ||
| 366 | lemma dvd_zero_right [simp]: | |
| 367 | "(a::'a::ring) dvd 0" | |
| 368 | proof | |
| 369 | show "0 = a * 0" by simp | |
| 370 | qed | |
| 371 | ||
| 372 | lemma dvd_zero_left: | |
| 373 | "0 dvd (a::'a::ring) \<Longrightarrow> a = 0" unfolding dvd_def by simp | |
| 374 | ||
| 375 | lemma dvd_refl_ring [simp]: | |
| 376 | "(a::'a::ring) dvd a" | |
| 377 | proof | |
| 378 | show "a = a * 1" by simp | |
| 379 | qed | |
| 380 | ||
| 381 | lemma dvd_trans_ring: | |
| 382 | fixes a b c :: "'a::ring" | |
| 383 | assumes a_dvd_b: "a dvd b" | |
| 384 | and b_dvd_c: "b dvd c" | |
| 385 | shows "a dvd c" | |
| 386 | proof - | |
| 387 | from a_dvd_b obtain l where "b = a * l" using dvd_def by blast | |
| 388 | moreover from b_dvd_c obtain j where "c = b * j" using dvd_def by blast | |
| 389 | ultimately have "c = a * (l * j)" by simp | |
| 390 | then have "\<exists>k. c = a * k" .. | |
| 391 | then show ?thesis using dvd_def by blast | |
| 392 | qed | |
| 393 | ||
| 21423 | 394 | |
| 395 | lemma unit_mult: | |
| 396 | "!!a::'a::ring. [| a dvd 1; b dvd 1 |] ==> a * b dvd 1" | |
| 397 | apply (unfold dvd_def) | |
| 398 | apply clarify | |
| 399 | apply (rule_tac x = "k * ka" in exI) | |
| 400 | apply simp | |
| 401 | done | |
| 402 | ||
| 403 | lemma unit_power: "!!a::'a::ring. a dvd 1 ==> a^n dvd 1" | |
| 404 | apply (induct_tac n) | |
| 405 | apply simp | |
| 406 | apply (simp add: unit_mult) | |
| 407 | done | |
| 408 | ||
| 409 | lemma dvd_add_right [simp]: | |
| 410 | "!! a::'a::ring. [| a dvd b; a dvd c |] ==> a dvd b + c" | |
| 411 | apply (unfold dvd_def) | |
| 412 | apply clarify | |
| 413 | apply (rule_tac x = "k + ka" in exI) | |
| 414 | apply (simp add: r_distr) | |
| 415 | done | |
| 416 | ||
| 417 | lemma dvd_uminus_right [simp]: | |
| 418 | "!! a::'a::ring. a dvd b ==> a dvd -b" | |
| 419 | apply (unfold dvd_def) | |
| 420 | apply clarify | |
| 421 | apply (rule_tac x = "-k" in exI) | |
| 422 | apply (simp add: r_minus) | |
| 423 | done | |
| 424 | ||
| 425 | lemma dvd_l_mult_right [simp]: | |
| 426 | "!! a::'a::ring. a dvd b ==> a dvd c*b" | |
| 427 | apply (unfold dvd_def) | |
| 428 | apply clarify | |
| 429 | apply (rule_tac x = "c * k" in exI) | |
| 430 | apply simp | |
| 431 | done | |
| 432 | ||
| 433 | lemma dvd_r_mult_right [simp]: | |
| 434 | "!! a::'a::ring. a dvd b ==> a dvd b*c" | |
| 435 | apply (unfold dvd_def) | |
| 436 | apply clarify | |
| 437 | apply (rule_tac x = "k * c" in exI) | |
| 438 | apply simp | |
| 439 | done | |
| 440 | ||
| 441 | ||
| 442 | (* Inverse of multiplication *) | |
| 443 | ||
| 444 | section "inverse" | |
| 445 | ||
| 446 | lemma inverse_unique: "!! a::'a::ring. [| a * x = 1; a * y = 1 |] ==> x = y" | |
| 447 | apply (rule_tac a = "(a*y) * x" and b = "y * (a*x)" in box_equals) | |
| 448 | apply (simp (no_asm)) | |
| 449 | apply auto | |
| 450 | done | |
| 451 | ||
| 452 | lemma r_inverse_ring: "!! a::'a::ring. a dvd 1 ==> a * inverse a = 1" | |
| 453 | apply (unfold inverse_def dvd_def) | |
| 26342 | 454 |   apply (tactic {* asm_full_simp_tac (@{simpset} delsimprocs [ring_simproc]) 1 *})
 | 
| 21423 | 455 | apply clarify | 
| 456 | apply (rule theI) | |
| 457 | apply assumption | |
| 458 | apply (rule inverse_unique) | |
| 459 | apply assumption | |
| 460 | apply assumption | |
| 461 | done | |
| 462 | ||
| 463 | lemma l_inverse_ring: "!! a::'a::ring. a dvd 1 ==> inverse a * a = 1" | |
| 464 | by (simp add: r_inverse_ring) | |
| 465 | ||
| 466 | ||
| 467 | (* Fields *) | |
| 468 | ||
| 469 | section "Fields" | |
| 470 | ||
| 471 | lemma field_unit [simp]: "!! a::'a::field. (a dvd 1) = (a ~= 0)" | |
| 472 | by (auto dest: field_ax dvd_zero_left simp add: field_one_not_zero) | |
| 473 | ||
| 474 | lemma r_inverse [simp]: "!! a::'a::field. a ~= 0 ==> a * inverse a = 1" | |
| 475 | by (simp add: r_inverse_ring) | |
| 476 | ||
| 477 | lemma l_inverse [simp]: "!! a::'a::field. a ~= 0 ==> inverse a * a= 1" | |
| 478 | by (simp add: l_inverse_ring) | |
| 479 | ||
| 480 | ||
| 481 | (* fields are integral domains *) | |
| 482 | ||
| 483 | lemma field_integral: "!! a::'a::field. a * b = 0 ==> a = 0 | b = 0" | |
| 23894 
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
 wenzelm parents: 
22997diff
changeset | 484 |   apply (tactic "step_tac @{claset} 1")
 | 
| 21423 | 485 | apply (rule_tac a = " (a*b) * inverse b" in box_equals) | 
| 486 | apply (rule_tac [3] refl) | |
| 487 | prefer 2 | |
| 488 | apply (simp (no_asm)) | |
| 489 | apply auto | |
| 490 | done | |
| 491 | ||
| 492 | ||
| 493 | (* fields are factorial domains *) | |
| 494 | ||
| 495 | lemma field_fact_prime: "!! a::'a::field. irred a ==> prime a" | |
| 496 | unfolding prime_def irred_def by (blast intro: field_ax) | |
| 21416 | 497 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 498 | end |