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