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