src/HOL/Ring_and_Field.thy
changeset 14265 95b42e69436c
child 14266 08b34c902618
equal deleted inserted replaced
14264:3d0c6238162a 14265:95b42e69436c
       
     1 (*  Title:   HOL/Ring_and_Field.thy
       
     2     ID:      $Id$
       
     3     Author:  Gertrud Bauer and Markus Wenzel, TU Muenchen
       
     4     License: GPL (GNU GENERAL PUBLIC LICENSE)
       
     5 *)
       
     6 
       
     7 header {*
       
     8   \title{Ring and field structures}
       
     9   \author{Gertrud Bauer and Markus Wenzel}
       
    10 *}
       
    11 
       
    12 theory Ring_and_Field = Inductive:
       
    13 
       
    14 text{*Lemmas and extension to semirings by L. C. Paulson*}
       
    15 
       
    16 subsection {* Abstract algebraic structures *}
       
    17 
       
    18 axclass semiring \<subseteq> zero, one, plus, times
       
    19   add_assoc: "(a + b) + c = a + (b + c)"
       
    20   add_commute: "a + b = b + a"
       
    21   left_zero [simp]: "0 + a = a"
       
    22 
       
    23   mult_assoc: "(a * b) * c = a * (b * c)"
       
    24   mult_commute: "a * b = b * a"
       
    25   left_one [simp]: "1 * a = a"
       
    26 
       
    27   left_distrib: "(a + b) * c = a * c + b * c"
       
    28   zero_neq_one [simp]: "0 \<noteq> 1"
       
    29 
       
    30 axclass ring \<subseteq> semiring, minus
       
    31   left_minus [simp]: "- a + a = 0"
       
    32   diff_minus: "a - b = a + (-b)"
       
    33 
       
    34 axclass ordered_semiring \<subseteq> semiring, linorder
       
    35   add_left_mono: "a \<le> b ==> c + a \<le> c + b"
       
    36   mult_strict_left_mono: "a < b ==> 0 < c ==> c * a < c * b"
       
    37 
       
    38 axclass ordered_ring \<subseteq> ordered_semiring, ring
       
    39   abs_if: "\<bar>a\<bar> = (if a < 0 then -a else a)"
       
    40 
       
    41 axclass field \<subseteq> ring, inverse
       
    42   left_inverse [simp]: "a \<noteq> 0 ==> inverse a * a = 1"
       
    43   divide_inverse:      "b \<noteq> 0 ==> a / b = a * inverse b"
       
    44 
       
    45 axclass ordered_field \<subseteq> ordered_ring, field
       
    46 
       
    47 axclass division_by_zero \<subseteq> zero, inverse
       
    48   inverse_zero: "inverse 0 = 0"
       
    49   divide_zero: "a / 0 = 0"
       
    50 
       
    51 
       
    52 subsection {* Derived rules for addition *}
       
    53 
       
    54 lemma right_zero [simp]: "a + 0 = (a::'a::semiring)"
       
    55 proof -
       
    56   have "a + 0 = 0 + a" by (simp only: add_commute)
       
    57   also have "... = a" by simp
       
    58   finally show ?thesis .
       
    59 qed
       
    60 
       
    61 lemma add_left_commute: "a + (b + c) = b + (a + (c::'a::semiring))"
       
    62   by (rule mk_left_commute [of "op +", OF add_assoc add_commute])
       
    63 
       
    64 theorems add_ac = add_assoc add_commute add_left_commute
       
    65 
       
    66 lemma right_minus [simp]: "a + -(a::'a::ring) = 0"
       
    67 proof -
       
    68   have "a + -a = -a + a" by (simp add: add_ac)
       
    69   also have "... = 0" by simp
       
    70   finally show ?thesis .
       
    71 qed
       
    72 
       
    73 lemma right_minus_eq: "(a - b = 0) = (a = (b::'a::ring))"
       
    74 proof
       
    75   have "a = a - b + b" by (simp add: diff_minus add_ac)
       
    76   also assume "a - b = 0"
       
    77   finally show "a = b" by simp
       
    78 next
       
    79   assume "a = b"
       
    80   thus "a - b = 0" by (simp add: diff_minus)
       
    81 qed
       
    82 
       
    83 lemma diff_self [simp]: "a - (a::'a::ring) = 0"
       
    84   by (simp add: diff_minus)
       
    85 
       
    86 lemma add_left_cancel [simp]:
       
    87      "(a + b = a + c) = (b = (c::'a::ring))"
       
    88 proof
       
    89   assume eq: "a + b = a + c"
       
    90   then have "(-a + a) + b = (-a + a) + c"
       
    91     by (simp only: eq add_assoc)
       
    92   then show "b = c" by simp
       
    93 next
       
    94   assume eq: "b = c"
       
    95   then show "a + b = a + c" by simp
       
    96 qed
       
    97 
       
    98 lemma add_right_cancel [simp]:
       
    99      "(b + a = c + a) = (b = (c::'a::ring))"
       
   100   by (simp add: add_commute)
       
   101 
       
   102 lemma minus_minus [simp]: "- (- (a::'a::ring)) = a"
       
   103   proof (rule add_left_cancel [of "-a", THEN iffD1])
       
   104     show "(-a + -(-a) = -a + a)"
       
   105     by simp
       
   106   qed
       
   107 
       
   108 lemma equals_zero_I: "a+b = 0 ==> -a = (b::'a::ring)"
       
   109 apply (rule right_minus_eq [THEN iffD1, symmetric])
       
   110 apply (simp add: diff_minus add_commute) 
       
   111 done
       
   112 
       
   113 lemma minus_zero [simp]: "- 0 = (0::'a::ring)"
       
   114 by (simp add: equals_zero_I)
       
   115 
       
   116 lemma neg_equal_iff_equal [simp]: "(-a = -b) = (a = (b::'a::ring))" 
       
   117   proof 
       
   118     assume "- a = - b"
       
   119     then have "- (- a) = - (- b)"
       
   120       by simp
       
   121     then show "a=b"
       
   122       by simp
       
   123   next
       
   124     assume "a=b"
       
   125     then show "-a = -b"
       
   126       by simp
       
   127   qed
       
   128 
       
   129 lemma neg_equal_0_iff_equal [simp]: "(-a = 0) = (a = (0::'a::ring))"
       
   130 by (subst neg_equal_iff_equal [symmetric], simp)
       
   131 
       
   132 lemma neg_0_equal_iff_equal [simp]: "(0 = -a) = (0 = (a::'a::ring))"
       
   133 by (subst neg_equal_iff_equal [symmetric], simp)
       
   134 
       
   135 
       
   136 subsection {* Derived rules for multiplication *}
       
   137 
       
   138 lemma right_one [simp]: "a = a * (1::'a::semiring)"
       
   139 proof -
       
   140   have "a = 1 * a" by simp
       
   141   also have "... = a * 1" by (simp add: mult_commute)
       
   142   finally show ?thesis .
       
   143 qed
       
   144 
       
   145 lemma mult_left_commute: "a * (b * c) = b * (a * (c::'a::semiring))"
       
   146   by (rule mk_left_commute [of "op *", OF mult_assoc mult_commute])
       
   147 
       
   148 theorems mult_ac = mult_assoc mult_commute mult_left_commute
       
   149 
       
   150 lemma right_inverse [simp]: "a \<noteq> 0 ==>  a * inverse (a::'a::field) = 1"
       
   151 proof -
       
   152   have "a * inverse a = inverse a * a" by (simp add: mult_ac)
       
   153   also assume "a \<noteq> 0"
       
   154   hence "inverse a * a = 1" by simp
       
   155   finally show ?thesis .
       
   156 qed
       
   157 
       
   158 lemma right_inverse_eq: "b \<noteq> 0 ==> (a / b = 1) = (a = (b::'a::field))"
       
   159 proof
       
   160   assume neq: "b \<noteq> 0"
       
   161   {
       
   162     hence "a = (a / b) * b" by (simp add: divide_inverse mult_ac)
       
   163     also assume "a / b = 1"
       
   164     finally show "a = b" by simp
       
   165   next
       
   166     assume "a = b"
       
   167     with neq show "a / b = 1" by (simp add: divide_inverse)
       
   168   }
       
   169 qed
       
   170 
       
   171 lemma divide_self [simp]: "a \<noteq> 0 ==> a / (a::'a::field) = 1"
       
   172   by (simp add: divide_inverse)
       
   173 
       
   174 lemma mult_left_zero [simp]: "0 * a = (0::'a::ring)"
       
   175 proof -
       
   176   have "0*a + 0*a = 0*a + 0"
       
   177     by (simp add: left_distrib [symmetric])
       
   178   then show ?thesis by (simp only: add_left_cancel)
       
   179 qed
       
   180 
       
   181 lemma mult_right_zero [simp]: "a * 0 = (0::'a::ring)"
       
   182   by (simp add: mult_commute)
       
   183 
       
   184 
       
   185 subsection {* Distribution rules *}
       
   186 
       
   187 lemma right_distrib: "a * (b + c) = a * b + a * (c::'a::semiring)"
       
   188 proof -
       
   189   have "a * (b + c) = (b + c) * a" by (simp add: mult_ac)
       
   190   also have "... = b * a + c * a" by (simp only: left_distrib)
       
   191   also have "... = a * b + a * c" by (simp add: mult_ac)
       
   192   finally show ?thesis .
       
   193 qed
       
   194 
       
   195 theorems ring_distrib = right_distrib left_distrib
       
   196 
       
   197 lemma minus_add_distrib [simp]: "- (a + b) = -a + -(b::'a::ring)"
       
   198 apply (rule equals_zero_I)
       
   199 apply (simp add: add_ac) 
       
   200 done
       
   201 
       
   202 lemma minus_mult_left: "- (a * b) = (-a) * (b::'a::ring)"
       
   203 apply (rule equals_zero_I)
       
   204 apply (simp add: left_distrib [symmetric]) 
       
   205 done
       
   206 
       
   207 lemma minus_mult_right: "- (a * b) = a * -(b::'a::ring)"
       
   208 apply (rule equals_zero_I)
       
   209 apply (simp add: right_distrib [symmetric]) 
       
   210 done
       
   211 
       
   212 lemma right_diff_distrib: "a * (b - c) = a * b - a * (c::'a::ring)"
       
   213 by (simp add: right_distrib diff_minus 
       
   214               minus_mult_left [symmetric] minus_mult_right [symmetric]) 
       
   215 
       
   216 
       
   217 subsection {* Ordering rules *}
       
   218 
       
   219 lemma add_right_mono: "a \<le> (b::'a::ordered_semiring) ==> a + c \<le> b + c"
       
   220 by (simp add: add_commute [of _ c] add_left_mono)
       
   221 
       
   222 lemma le_imp_neg_le:
       
   223    assumes "a \<le> (b::'a::ordered_ring)" shows "-b \<le> -a"
       
   224   proof -
       
   225   have "-a+a \<le> -a+b"
       
   226     by (rule add_left_mono) 
       
   227   then have "0 \<le> -a+b"
       
   228     by simp
       
   229   then have "0 + (-b) \<le> (-a + b) + (-b)"
       
   230     by (rule add_right_mono) 
       
   231   then show ?thesis
       
   232     by (simp add: add_assoc)
       
   233   qed
       
   234 
       
   235 lemma neg_le_iff_le [simp]: "(-b \<le> -a) = (a \<le> (b::'a::ordered_ring))"
       
   236   proof 
       
   237     assume "- b \<le> - a"
       
   238     then have "- (- a) \<le> - (- b)"
       
   239       by (rule le_imp_neg_le)
       
   240     then show "a\<le>b"
       
   241       by simp
       
   242   next
       
   243     assume "a\<le>b"
       
   244     then show "-b \<le> -a"
       
   245       by (rule le_imp_neg_le)
       
   246   qed
       
   247 
       
   248 lemma neg_le_0_iff_le [simp]: "(-a \<le> 0) = (0 \<le> (a::'a::ordered_ring))"
       
   249 by (subst neg_le_iff_le [symmetric], simp)
       
   250 
       
   251 lemma neg_0_le_iff_le [simp]: "(0 \<le> -a) = (a \<le> (0::'a::ordered_ring))"
       
   252 by (subst neg_le_iff_le [symmetric], simp)
       
   253 
       
   254 lemma neg_less_iff_less [simp]: "(-b < -a) = (a < (b::'a::ordered_ring))"
       
   255 by (force simp add: order_less_le) 
       
   256 
       
   257 lemma neg_less_0_iff_less [simp]: "(-a < 0) = (0 < (a::'a::ordered_ring))"
       
   258 by (subst neg_less_iff_less [symmetric], simp)
       
   259 
       
   260 lemma neg_0_less_iff_less [simp]: "(0 < -a) = (a < (0::'a::ordered_ring))"
       
   261 by (subst neg_less_iff_less [symmetric], simp)
       
   262 
       
   263 lemma mult_strict_right_mono:
       
   264      "[|a < b; 0 < c|] ==> a * c < b * (c::'a::ordered_semiring)"
       
   265 by (simp add: mult_commute [of _ c] mult_strict_left_mono)
       
   266 
       
   267 lemma mult_left_mono:
       
   268      "[|a \<le> b; 0 < c|] ==> c * a \<le> c * (b::'a::ordered_semiring)"
       
   269 by (force simp add: mult_strict_left_mono order_le_less) 
       
   270 
       
   271 lemma mult_right_mono:
       
   272      "[|a \<le> b; 0 < c|] ==> a*c \<le> b * (c::'a::ordered_semiring)"
       
   273 by (force simp add: mult_strict_right_mono order_le_less) 
       
   274 
       
   275 lemma mult_strict_left_mono_neg:
       
   276      "[|b < a; c < 0|] ==> c * a < c * (b::'a::ordered_ring)"
       
   277 apply (drule mult_strict_left_mono [of _ _ "-c"])
       
   278 apply (simp_all add: minus_mult_left [symmetric]) 
       
   279 done
       
   280 
       
   281 lemma mult_strict_right_mono_neg:
       
   282      "[|b < a; c < 0|] ==> a * c < b * (c::'a::ordered_ring)"
       
   283 apply (drule mult_strict_right_mono [of _ _ "-c"])
       
   284 apply (simp_all add: minus_mult_right [symmetric]) 
       
   285 done
       
   286 
       
   287 lemma mult_left_mono_neg:
       
   288      "[|b \<le> a; c < 0|] ==> c * a \<le> c * (b::'a::ordered_ring)"
       
   289 by (force simp add: mult_strict_left_mono_neg order_le_less) 
       
   290 
       
   291 lemma mult_right_mono_neg:
       
   292      "[|b \<le> a; c < 0|] ==> a * c \<le> b * (c::'a::ordered_ring)"
       
   293 by (force simp add: mult_strict_right_mono_neg order_le_less) 
       
   294 
       
   295 
       
   296 subsection{* Products of Signs *}
       
   297 
       
   298 lemma mult_pos: "[| (0::'a::ordered_ring) < a; 0 < b |] ==> 0 < a*b"
       
   299 by (drule mult_strict_left_mono [of 0 b], auto)
       
   300 
       
   301 lemma mult_pos_neg: "[| (0::'a::ordered_ring) < a; b < 0 |] ==> a*b < 0"
       
   302 by (drule mult_strict_left_mono [of b 0], auto)
       
   303 
       
   304 lemma mult_neg: "[| a < (0::'a::ordered_ring); b < 0 |] ==> 0 < a*b"
       
   305 by (drule mult_strict_right_mono_neg, auto)
       
   306 
       
   307 lemma zero_less_mult_pos: "[| 0 < a*b; 0 < a|] ==> 0 < (b::'a::ordered_ring)"
       
   308 apply (case_tac "b\<le>0") 
       
   309  apply (auto simp add: order_le_less linorder_not_less)
       
   310 apply (drule_tac mult_pos_neg [of a b]) 
       
   311  apply (auto dest: order_less_not_sym)
       
   312 done
       
   313 
       
   314 lemma zero_less_mult_iff:
       
   315      "((0::'a::ordered_ring) < a*b) = (0 < a & 0 < b | a < 0 & b < 0)"
       
   316 apply (auto simp add: order_le_less linorder_not_less mult_pos mult_neg)
       
   317 apply (blast dest: zero_less_mult_pos) 
       
   318 apply (simp add: mult_commute [of a b]) 
       
   319 apply (blast dest: zero_less_mult_pos) 
       
   320 done
       
   321 
       
   322 
       
   323 lemma mult_eq_0_iff [iff]: "(a*b = (0::'a::ordered_ring)) = (a = 0 | b = 0)"
       
   324 apply (case_tac "a < 0")
       
   325 apply (auto simp add: linorder_not_less order_le_less linorder_neq_iff)
       
   326 apply (force dest: mult_strict_right_mono_neg mult_strict_right_mono)+
       
   327 done
       
   328 
       
   329 lemma zero_le_mult_iff:
       
   330      "((0::'a::ordered_ring) \<le> a*b) = (0 \<le> a & 0 \<le> b | a \<le> 0 & b \<le> 0)"
       
   331 by (auto simp add: eq_commute [of 0] order_le_less linorder_not_less
       
   332                    zero_less_mult_iff)
       
   333 
       
   334 lemma mult_less_0_iff:
       
   335      "(a*b < (0::'a::ordered_ring)) = (0 < a & b < 0 | a < 0 & 0 < b)"
       
   336 apply (insert zero_less_mult_iff [of "-a" b]) 
       
   337 apply (force simp add: minus_mult_left[symmetric]) 
       
   338 done
       
   339 
       
   340 lemma mult_le_0_iff:
       
   341      "(a*b \<le> (0::'a::ordered_ring)) = (0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)"
       
   342 apply (insert zero_le_mult_iff [of "-a" b]) 
       
   343 apply (force simp add: minus_mult_left[symmetric]) 
       
   344 done
       
   345 
       
   346 lemma zero_le_square: "(0::'a::ordered_ring) \<le> a*a"
       
   347 by (simp add: zero_le_mult_iff linorder_linear) 
       
   348 
       
   349 lemma zero_less_one: "(0::'a::ordered_ring) < 1"
       
   350 apply (insert zero_le_square [of 1]) 
       
   351 apply (simp add: order_less_le) 
       
   352 done
       
   353 
       
   354 
       
   355 subsection {* Absolute Value *}
       
   356 
       
   357 text{*But is it really better than just rewriting with @{text abs_if}?*}
       
   358 lemma abs_split:
       
   359      "P(abs(a::'a::ordered_ring)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
       
   360 by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)
       
   361 
       
   362 lemma abs_zero [simp]: "abs 0 = (0::'a::ordered_ring)"
       
   363 by (simp add: abs_if)
       
   364 
       
   365 lemma abs_mult: "abs (x * y) = abs x * abs (y::'a::ordered_ring)" 
       
   366 apply (case_tac "x=0 | y=0", force) 
       
   367 apply (auto elim: order_less_asym
       
   368             simp add: abs_if mult_less_0_iff linorder_neq_iff
       
   369                   minus_mult_left [symmetric] minus_mult_right [symmetric])  
       
   370 done
       
   371 
       
   372 lemma abs_eq_0 [iff]: "(abs x = 0) = (x = (0::'a::ordered_ring))"
       
   373 by (simp add: abs_if)
       
   374 
       
   375 lemma zero_less_abs_iff [iff]: "(0 < abs x) = (x ~= (0::'a::ordered_ring))"
       
   376 by (simp add: abs_if linorder_neq_iff)
       
   377 
       
   378 
       
   379 subsection {* Fields *}
       
   380 
       
   381 
       
   382 end