| author | wenzelm | 
| Thu, 27 Dec 2001 16:45:19 +0100 | |
| changeset 12600 | 30ec65eaaf5f | 
| parent 11778 | 37efbe093d3c | 
| child 13550 | 5a176b8dda84 | 
| permissions | -rw-r--r-- | 
| 7998 | 1 | (* | 
| 2 | Abstract class ring (commutative, with 1) | |
| 3 | $Id$ | |
| 4 | Author: Clemens Ballarin, started 9 December 1996 | |
| 5 | *) | |
| 6 | ||
| 7 | Blast.overloaded ("Divides.op dvd", domain_type);
 | |
| 8 | ||
| 11778 | 9 | val a_assoc = thm "plus_ac0.assoc"; | 
| 10 | val l_zero = thm "plus_ac0.zero"; | |
| 11 | val a_comm = thm "plus_ac0.commute"; | |
| 12 | ||
| 7998 | 13 | section "Rings"; | 
| 14 | ||
| 15 | fun make_left_commute assoc commute s = | |
| 16 | [rtac (commute RS trans) 1, rtac (assoc RS trans) 1, | |
| 17 | rtac (commute RS arg_cong) 1]; | |
| 18 | ||
| 19 | (* addition *) | |
| 20 | ||
| 21 | qed_goal "a_lcomm" Ring.thy "!!a::'a::ring. a+(b+c) = b+(a+c)" | |
| 22 | (make_left_commute a_assoc a_comm); | |
| 23 | ||
| 24 | val a_ac = [a_assoc, a_comm, a_lcomm]; | |
| 25 | ||
| 11093 | 26 | Goal "!!a::'a::ring. a + 0 = a"; | 
| 9390 
e6b96d953965
renamed // to / (which is what we want anyway) to avoid clash with the new
 paulson parents: 
8707diff
changeset | 27 | by (rtac (a_comm RS trans) 1); | 
| 
e6b96d953965
renamed // to / (which is what we want anyway) to avoid clash with the new
 paulson parents: 
8707diff
changeset | 28 | by (rtac l_zero 1); | 
| 
e6b96d953965
renamed // to / (which is what we want anyway) to avoid clash with the new
 paulson parents: 
8707diff
changeset | 29 | qed "r_zero"; | 
| 7998 | 30 | |
| 11093 | 31 | Goal "!!a::'a::ring. a + (-a) = 0"; | 
| 9390 
e6b96d953965
renamed // to / (which is what we want anyway) to avoid clash with the new
 paulson parents: 
8707diff
changeset | 32 | by (rtac (a_comm RS trans) 1); | 
| 
e6b96d953965
renamed // to / (which is what we want anyway) to avoid clash with the new
 paulson parents: 
8707diff
changeset | 33 | by (rtac l_neg 1); | 
| 
e6b96d953965
renamed // to / (which is what we want anyway) to avoid clash with the new
 paulson parents: 
8707diff
changeset | 34 | qed "r_neg"; | 
| 7998 | 35 | |
| 36 | Goal "!! a::'a::ring. a + b = a + c ==> b = c"; | |
| 37 | by (rtac box_equals 1); | |
| 38 | by (rtac l_zero 2); | |
| 39 | by (rtac l_zero 2); | |
| 40 | by (res_inst_tac [("a1", "a")] (l_neg RS subst) 1);
 | |
| 41 | by (asm_simp_tac (simpset() addsimps [a_assoc]) 1); | |
| 42 | qed "a_lcancel"; | |
| 43 | ||
| 44 | Goal "!! a::'a::ring. b + a = c + a ==> b = c"; | |
| 45 | by (rtac a_lcancel 1); | |
| 46 | by (asm_simp_tac (simpset() addsimps a_ac) 1); | |
| 47 | qed "a_rcancel"; | |
| 48 | ||
| 49 | Goal "!! a::'a::ring. (a + b = a + c) = (b = c)"; | |
| 50 | by (auto_tac (claset() addSDs [a_lcancel], simpset())); | |
| 51 | qed "a_lcancel_eq"; | |
| 52 | ||
| 53 | Goal "!! a::'a::ring. (b + a = c + a) = (b = c)"; | |
| 54 | by (simp_tac (simpset() addsimps [a_lcancel_eq, a_comm]) 1); | |
| 55 | qed "a_rcancel_eq"; | |
| 56 | ||
| 57 | Addsimps [a_lcancel_eq, a_rcancel_eq]; | |
| 58 | ||
| 59 | Goal "!!a::'a::ring. -(-a) = a"; | |
| 60 | by (rtac a_lcancel 1); | |
| 61 | by (rtac (r_neg RS trans) 1); | |
| 62 | by (rtac (l_neg RS sym) 1); | |
| 63 | qed "minus_minus"; | |
| 64 | ||
| 11093 | 65 | Goal "- 0 = (0::'a::ring)"; | 
| 7998 | 66 | by (rtac a_lcancel 1); | 
| 67 | by (rtac (r_neg RS trans) 1); | |
| 68 | by (rtac (l_zero RS sym) 1); | |
| 69 | qed "minus0"; | |
| 70 | ||
| 71 | Goal "!!a::'a::ring. -(a + b) = (-a) + (-b)"; | |
| 72 | by (res_inst_tac [("a", "a+b")] a_lcancel 1);
 | |
| 73 | by (simp_tac (simpset() addsimps ([r_neg, l_neg, l_zero]@a_ac)) 1); | |
| 74 | qed "minus_add"; | |
| 75 | ||
| 76 | (* multiplication *) | |
| 77 | ||
| 78 | qed_goal "m_lcomm" Ring.thy "!!a::'a::ring. a*(b*c) = b*(a*c)" | |
| 79 | (make_left_commute m_assoc m_comm); | |
| 80 | ||
| 81 | val m_ac = [m_assoc, m_comm, m_lcomm]; | |
| 82 | ||
| 9390 
e6b96d953965
renamed // to / (which is what we want anyway) to avoid clash with the new
 paulson parents: 
8707diff
changeset | 83 | Goal "!!a::'a::ring. a * <1> = a"; | 
| 
e6b96d953965
renamed // to / (which is what we want anyway) to avoid clash with the new
 paulson parents: 
8707diff
changeset | 84 | by (rtac (m_comm RS trans) 1); | 
| 
e6b96d953965
renamed // to / (which is what we want anyway) to avoid clash with the new
 paulson parents: 
8707diff
changeset | 85 | by (rtac l_one 1); | 
| 
e6b96d953965
renamed // to / (which is what we want anyway) to avoid clash with the new
 paulson parents: 
8707diff
changeset | 86 | qed "r_one"; | 
| 7998 | 87 | |
| 88 | (* distributive and derived *) | |
| 89 | ||
| 90 | Goal "!!a::'a::ring. a * (b + c) = a * b + a * c"; | |
| 91 | by (rtac (m_comm RS trans) 1); | |
| 92 | by (rtac (l_distr RS trans) 1); | |
| 93 | by (simp_tac (simpset() addsimps [m_comm]) 1); | |
| 94 | qed "r_distr"; | |
| 95 | ||
| 96 | val m_distr = m_ac @ [l_distr, r_distr]; | |
| 97 | ||
| 98 | (* the following two proofs can be found in | |
| 99 | Jacobson, Basic Algebra I, pp. 88-89 *) | |
| 100 | ||
| 11093 | 101 | Goal "!!a::'a::ring. 0 * a = 0"; | 
| 7998 | 102 | by (rtac a_lcancel 1); | 
| 103 | by (rtac (l_distr RS sym RS trans) 1); | |
| 104 | by (simp_tac (simpset() addsimps [r_zero]) 1); | |
| 105 | qed "l_null"; | |
| 106 | ||
| 11093 | 107 | Goal "!!a::'a::ring. a * 0 = 0"; | 
| 9390 
e6b96d953965
renamed // to / (which is what we want anyway) to avoid clash with the new
 paulson parents: 
8707diff
changeset | 108 | by (rtac (m_comm RS trans) 1); | 
| 
e6b96d953965
renamed // to / (which is what we want anyway) to avoid clash with the new
 paulson parents: 
8707diff
changeset | 109 | by (rtac l_null 1); | 
| 
e6b96d953965
renamed // to / (which is what we want anyway) to avoid clash with the new
 paulson parents: 
8707diff
changeset | 110 | qed "r_null"; | 
| 7998 | 111 | |
| 112 | Goal "!!a::'a::ring. (-a) * b = - (a * b)"; | |
| 113 | by (rtac a_lcancel 1); | |
| 114 | by (rtac (r_neg RS sym RSN (2, trans)) 1); | |
| 115 | by (rtac (l_distr RS sym RS trans) 1); | |
| 116 | by (simp_tac (simpset() addsimps [l_null, r_neg]) 1); | |
| 117 | qed "l_minus"; | |
| 118 | ||
| 119 | Goal "!!a::'a::ring. a * (-b) = - (a * b)"; | |
| 120 | by (rtac a_lcancel 1); | |
| 121 | by (rtac (r_neg RS sym RSN (2, trans)) 1); | |
| 122 | by (rtac (r_distr RS sym RS trans) 1); | |
| 123 | by (simp_tac (simpset() addsimps [r_null, r_neg]) 1); | |
| 124 | qed "r_minus"; | |
| 125 | ||
| 126 | val m_minus = [l_minus, r_minus]; | |
| 127 | ||
| 128 | Addsimps [l_zero, r_zero, l_neg, r_neg, minus_minus, minus0, | |
| 11093 | 129 | l_one, r_one, l_null, r_null]; | 
| 7998 | 130 | |
| 131 | (* further rules *) | |
| 132 | ||
| 11093 | 133 | Goal "!!a::'a::ring. -a = 0 ==> a = 0"; | 
| 7998 | 134 | by (res_inst_tac [("t", "a")] (minus_minus RS subst) 1);
 | 
| 135 | by (Asm_simp_tac 1); | |
| 136 | qed "uminus_monom"; | |
| 137 | ||
| 11093 | 138 | Goal "!!a::'a::ring. a ~= 0 ==> -a ~= 0"; | 
| 10198 | 139 | by (blast_tac (claset() addIs [uminus_monom]) 1); | 
| 7998 | 140 | qed "uminus_monom_neq"; | 
| 141 | ||
| 11093 | 142 | Goal "!!a::'a::ring. a * b ~= 0 ==> a ~= 0"; | 
| 10198 | 143 | by Auto_tac; | 
| 7998 | 144 | qed "l_nullD"; | 
| 145 | ||
| 11093 | 146 | Goal "!!a::'a::ring. a * b ~= 0 ==> b ~= 0"; | 
| 10198 | 147 | by Auto_tac; | 
| 7998 | 148 | qed "r_nullD"; | 
| 149 | ||
| 11093 | 150 | (* reflection between a = b and a -- b = 0 *) | 
| 7998 | 151 | |
| 11093 | 152 | Goal "!!a::'a::ring. a = b ==> a + (-b) = 0"; | 
| 7998 | 153 | by (Asm_simp_tac 1); | 
| 154 | qed "eq_imp_diff_zero"; | |
| 155 | ||
| 11093 | 156 | Goal "!!a::'a::ring. a + (-b) = 0 ==> a = b"; | 
| 7998 | 157 | by (res_inst_tac [("a", "-b")] a_rcancel 1);
 | 
| 158 | by (Asm_simp_tac 1); | |
| 159 | qed "diff_zero_imp_eq"; | |
| 160 | ||
| 161 | (* this could be a rewrite rule, but won't terminate | |
| 162 | ==> make it a simproc? | |
| 11093 | 163 | Goal "!!a::'a::ring. (a = b) = (a -- b = 0)"; | 
| 7998 | 164 | *) | 
| 165 | ||
| 166 | (* Power *) | |
| 167 | ||
| 168 | Goal "!!a::'a::ring. a ^ 0 = <1>"; | |
| 169 | by (simp_tac (simpset() addsimps [power_ax]) 1); | |
| 170 | qed "power_0"; | |
| 171 | ||
| 172 | Goal "!!a::'a::ring. a ^ Suc n = a ^ n * a"; | |
| 173 | by (simp_tac (simpset() addsimps [power_ax]) 1); | |
| 174 | qed "power_Suc"; | |
| 175 | ||
| 176 | Addsimps [power_0, power_Suc]; | |
| 177 | ||
| 178 | Goal "<1> ^ n = (<1>::'a::ring)"; | |
| 8707 | 179 | by (induct_tac "n" 1); | 
| 180 | by Auto_tac; | |
| 7998 | 181 | qed "power_one"; | 
| 182 | ||
| 11093 | 183 | Goal "!!n. n ~= 0 ==> 0 ^ n = (0::'a::ring)"; | 
| 7998 | 184 | by (etac rev_mp 1); | 
| 8707 | 185 | by (induct_tac "n" 1); | 
| 186 | by Auto_tac; | |
| 7998 | 187 | qed "power_zero"; | 
| 188 | ||
| 189 | Addsimps [power_zero, power_one]; | |
| 190 | ||
| 191 | Goal "!! a::'a::ring. a ^ m * a ^ n = a ^ (m + n)"; | |
| 8707 | 192 | by (induct_tac "m" 1); | 
| 7998 | 193 | by (Simp_tac 1); | 
| 194 | by (asm_simp_tac (simpset() addsimps m_ac) 1); | |
| 195 | qed "power_mult"; | |
| 196 | ||
| 197 | (* Divisibility *) | |
| 198 | section "Divisibility"; | |
| 199 | ||
| 11093 | 200 | Goalw [dvd_def] "!! a::'a::ring. a dvd 0"; | 
| 201 | by (res_inst_tac [("x", "0")] exI 1);
 | |
| 7998 | 202 | by (Simp_tac 1); | 
| 203 | qed "dvd_zero_right"; | |
| 204 | ||
| 11093 | 205 | Goalw [dvd_def] "!! a::'a::ring. 0 dvd a ==> a = 0"; | 
| 7998 | 206 | by Auto_tac; | 
| 207 | qed "dvd_zero_left"; | |
| 208 | ||
| 209 | Goalw [dvd_def] "!! a::'a::ring. a dvd a"; | |
| 210 | by (res_inst_tac [("x", "<1>")] exI 1);
 | |
| 211 | by (Simp_tac 1); | |
| 212 | qed "dvd_refl_ring"; | |
| 213 | ||
| 214 | Goalw [dvd_def] "!! a::'a::ring. [| a dvd b; b dvd c |] ==> a dvd c"; | |
| 215 | by (Step_tac 1); | |
| 216 | by (res_inst_tac [("x", "k * ka")] exI 1);
 | |
| 217 | by (simp_tac (simpset() addsimps m_ac) 1); | |
| 218 | qed "dvd_trans_ring"; | |
| 219 | ||
| 220 | Addsimps [dvd_zero_right, dvd_refl_ring]; | |
| 221 | ||
| 222 | Goalw [dvd_def] | |
| 223 | "!!a::'a::ring. [| a dvd <1>; b dvd <1> |] ==> a * b dvd <1>"; | |
| 224 | by (Clarify_tac 1); | |
| 225 | by (res_inst_tac [("x", "k * ka")] exI 1);
 | |
| 226 | by (asm_full_simp_tac (simpset() addsimps m_ac) 1); | |
| 227 | qed "unit_mult"; | |
| 228 | ||
| 229 | Goal "!!a::'a::ring. a dvd <1> ==> a^n dvd <1>"; | |
| 230 | by (induct_tac "n" 1); | |
| 231 | by (Simp_tac 1); | |
| 232 | by (asm_simp_tac (simpset() addsimps [unit_mult]) 1); | |
| 233 | qed "unit_power"; | |
| 234 | ||
| 235 | Goalw [dvd_def] | |
| 10789 
260fa2c67e3e
Changed priority of dvd from 70 to 50 as befits a relation.
 nipkow parents: 
10448diff
changeset | 236 | "!! a::'a::ring. [| a dvd b; a dvd c |] ==> a dvd b + c"; | 
| 7998 | 237 | by (Clarify_tac 1); | 
| 238 | by (res_inst_tac [("x", "k + ka")] exI 1);
 | |
| 239 | by (simp_tac (simpset() addsimps [r_distr]) 1); | |
| 240 | qed "dvd_add_right"; | |
| 241 | ||
| 242 | Goalw [dvd_def] | |
| 10789 
260fa2c67e3e
Changed priority of dvd from 70 to 50 as befits a relation.
 nipkow parents: 
10448diff
changeset | 243 | "!! a::'a::ring. a dvd b ==> a dvd -b"; | 
| 7998 | 244 | by (Clarify_tac 1); | 
| 245 | by (res_inst_tac [("x", "-k")] exI 1);
 | |
| 246 | by (simp_tac (simpset() addsimps [r_minus]) 1); | |
| 247 | qed "dvd_uminus_right"; | |
| 248 | ||
| 249 | Goalw [dvd_def] | |
| 10789 
260fa2c67e3e
Changed priority of dvd from 70 to 50 as befits a relation.
 nipkow parents: 
10448diff
changeset | 250 | "!! a::'a::ring. a dvd b ==> a dvd c*b"; | 
| 7998 | 251 | by (Clarify_tac 1); | 
| 252 | by (res_inst_tac [("x", "c * k")] exI 1);
 | |
| 253 | by (simp_tac (simpset() addsimps m_ac) 1); | |
| 254 | qed "dvd_l_mult_right"; | |
| 255 | ||
| 256 | Goalw [dvd_def] | |
| 10789 
260fa2c67e3e
Changed priority of dvd from 70 to 50 as befits a relation.
 nipkow parents: 
10448diff
changeset | 257 | "!! a::'a::ring. a dvd b ==> a dvd b*c"; | 
| 7998 | 258 | by (Clarify_tac 1); | 
| 259 | by (res_inst_tac [("x", "k * c")] exI 1);
 | |
| 260 | by (simp_tac (simpset() addsimps m_ac) 1); | |
| 261 | qed "dvd_r_mult_right"; | |
| 262 | ||
| 263 | Addsimps [dvd_add_right, dvd_uminus_right, dvd_l_mult_right, dvd_r_mult_right]; | |
| 264 | ||
| 265 | (* Inverse of multiplication *) | |
| 266 | ||
| 267 | section "inverse"; | |
| 268 | ||
| 269 | Goal "!! a::'a::ring. [| a * x = <1>; a * y = <1> |] ==> x = y"; | |
| 270 | by (res_inst_tac [("a", "(a*y)*x"), ("b", "y*(a*x)")] box_equals 1);
 | |
| 271 | by (simp_tac (simpset() addsimps m_ac) 1); | |
| 272 | by Auto_tac; | |
| 273 | qed "inverse_unique"; | |
| 274 | ||
| 10448 | 275 | Goal "!! a::'a::ring. a dvd <1> ==> a * inverse a = <1>"; | 
| 276 | by (asm_full_simp_tac (simpset () addsimps [inverse_ax, dvd_def]) 1); | |
| 7998 | 277 | by (Clarify_tac 1); | 
| 9970 | 278 | by (rtac someI 1); | 
| 7998 | 279 | by (rtac sym 1); | 
| 280 | by (assume_tac 1); | |
| 281 | qed "r_inverse_ring"; | |
| 282 | ||
| 283 | Goal "!! a::'a::ring. a dvd <1> ==> inverse a * a= <1>"; | |
| 284 | by (asm_simp_tac (simpset() addsimps r_inverse_ring::m_ac) 1); | |
| 285 | qed "l_inverse_ring"; | |
| 286 | ||
| 287 | (* Integral domain *) | |
| 288 | ||
| 289 | section "Integral domains"; | |
| 290 | ||
| 11093 | 291 | Goal "0 ~= (<1>::'a::domain)"; | 
| 292 | by (rtac not_sym 1); | |
| 293 | by (rtac one_not_zero 1); | |
| 294 | qed "zero_not_one"; | |
| 295 | ||
| 296 | Goal "!! a::'a::domain. a dvd <1> ==> a ~= 0"; | |
| 297 | by (auto_tac (claset() addDs [dvd_zero_left], | |
| 298 | simpset() addsimps [one_not_zero] )); | |
| 299 | qed "unit_imp_nonzero"; | |
| 300 | ||
| 301 | Goal "[| a * b = 0; a ~= 0 |] ==> (b::'a::domain) = 0"; | |
| 7998 | 302 | by (dtac integral 1); | 
| 303 | by (Fast_tac 1); | |
| 304 | qed "r_integral"; | |
| 305 | ||
| 11093 | 306 | Goal "[| a * b = 0; b ~= 0 |] ==> (a::'a::domain) = 0"; | 
| 7998 | 307 | by (dtac integral 1); | 
| 308 | by (Fast_tac 1); | |
| 309 | qed "l_integral"; | |
| 310 | ||
| 11093 | 311 | Goal "!! a::'a::domain. [| a ~= 0; b ~= 0 |] ==> a * b ~= 0"; | 
| 10230 | 312 | by (blast_tac (claset() addIs [l_integral]) 1); | 
| 7998 | 313 | qed "not_integral"; | 
| 314 | ||
| 11093 | 315 | Addsimps [not_integral, one_not_zero, zero_not_one]; | 
| 7998 | 316 | |
| 11093 | 317 | Goal "!! a::'a::domain. [| a * x = x; x ~= 0 |] ==> a = <1>"; | 
| 7998 | 318 | by (res_inst_tac [("a", "- <1>")] a_lcancel 1);
 | 
| 319 | by (Simp_tac 1); | |
| 320 | by (rtac l_integral 1); | |
| 321 | by (assume_tac 2); | |
| 322 | by (asm_simp_tac (simpset() addsimps [l_distr, l_minus]) 1); | |
| 323 | qed "l_one_integral"; | |
| 324 | ||
| 11093 | 325 | Goal "!! a::'a::domain. [| x * a = x; x ~= 0 |] ==> a = <1>"; | 
| 7998 | 326 | by (res_inst_tac [("a", "- <1>")] a_rcancel 1);
 | 
| 327 | by (Simp_tac 1); | |
| 328 | by (rtac r_integral 1); | |
| 329 | by (assume_tac 2); | |
| 330 | by (asm_simp_tac (simpset() addsimps [r_distr, r_minus]) 1); | |
| 331 | qed "r_one_integral"; | |
| 332 | ||
| 333 | (* cancellation laws for multiplication *) | |
| 334 | ||
| 11093 | 335 | Goal "!! a::'a::domain. [| a ~= 0; a * b = a * c |] ==> b = c"; | 
| 7998 | 336 | by (rtac diff_zero_imp_eq 1); | 
| 337 | by (dtac eq_imp_diff_zero 1); | |
| 338 | by (full_simp_tac (simpset() addsimps [r_minus RS sym, r_distr RS sym]) 1); | |
| 339 | by (fast_tac (claset() addIs [l_integral]) 1); | |
| 340 | qed "m_lcancel"; | |
| 341 | ||
| 11093 | 342 | Goal "!! a::'a::domain. [| a ~= 0; b * a = c * a |] ==> b = c"; | 
| 7998 | 343 | by (rtac m_lcancel 1); | 
| 344 | by (assume_tac 1); | |
| 345 | by (asm_full_simp_tac (simpset() addsimps m_ac) 1); | |
| 346 | qed "m_rcancel"; | |
| 347 | ||
| 11093 | 348 | Goal "!! a::'a::domain. a ~= 0 ==> (a * b = a * c) = (b = c)"; | 
| 7998 | 349 | by (auto_tac (claset() addDs [m_lcancel], simpset())); | 
| 350 | qed "m_lcancel_eq"; | |
| 351 | ||
| 11093 | 352 | Goal "!! a::'a::domain. a ~= 0 ==> (b * a = c * a) = (b = c)"; | 
| 7998 | 353 | by (asm_simp_tac (simpset() addsimps [m_lcancel_eq, m_comm]) 1); | 
| 354 | qed "m_rcancel_eq"; | |
| 355 | ||
| 356 | Addsimps [m_lcancel_eq, m_rcancel_eq]; | |
| 357 | ||
| 358 | (* Fields *) | |
| 359 | ||
| 360 | section "Fields"; | |
| 361 | ||
| 11093 | 362 | Goal "!! a::'a::field. (a dvd <1>) = (a ~= 0)"; | 
| 363 | by (auto_tac (claset() addDs [field_ax, dvd_zero_left], | |
| 364 | simpset() addsimps [field_one_not_zero])); | |
| 7998 | 365 | qed "field_unit"; | 
| 366 | ||
| 367 | Addsimps [field_unit]; | |
| 368 | ||
| 11093 | 369 | Goal "!! a::'a::field. a ~= 0 ==> a * inverse a = <1>"; | 
| 7998 | 370 | by (asm_full_simp_tac (simpset() addsimps [r_inverse_ring]) 1); | 
| 371 | qed "r_inverse"; | |
| 372 | ||
| 11093 | 373 | Goal "!! a::'a::field. a ~= 0 ==> inverse a * a= <1>"; | 
| 7998 | 374 | by (asm_full_simp_tac (simpset() addsimps [l_inverse_ring]) 1); | 
| 375 | qed "l_inverse"; | |
| 376 | ||
| 377 | Addsimps [l_inverse, r_inverse]; | |
| 378 | ||
| 11093 | 379 | (* fields are integral domains *) | 
| 7998 | 380 | |
| 11093 | 381 | Goal "!! a::'a::field. a * b = 0 ==> a = 0 | b = 0"; | 
| 7998 | 382 | by (Step_tac 1); | 
| 383 | by (res_inst_tac [("a", "(a*b)*inverse b")] box_equals 1);
 | |
| 384 | by (rtac refl 3); | |
| 385 | by (simp_tac (simpset() addsimps m_ac) 2); | |
| 386 | by Auto_tac; | |
| 387 | qed "field_integral"; | |
| 388 | ||
| 11093 | 389 | (* fields are factorial domains *) | 
| 390 | ||
| 7998 | 391 | Goalw [prime_def, irred_def] "!! a::'a::field. irred a ==> prime a"; | 
| 392 | by (blast_tac (claset() addIs [field_ax]) 1); | |
| 393 | qed "field_fact_prime"; |