src/HOL/Ring_and_Field.thy
author paulson
Tue, 02 Dec 2003 11:48:15 +0100
changeset 14270 342451d763f9
parent 14269 502a7c95de73
child 14272 5efbb548107d
permissions -rw-r--r--
More re-organising of numerical theorems
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     1
(*  Title:   HOL/Ring_and_Field.thy
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     2
    ID:      $Id$
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     3
    Author:  Gertrud Bauer and Markus Wenzel, TU Muenchen
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
     4
             Lawrence C Paulson, University of Cambridge
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     5
    License: GPL (GNU GENERAL PUBLIC LICENSE)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     6
*)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     7
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     8
header {*
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     9
  \title{Ring and field structures}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    10
  \author{Gertrud Bauer and Markus Wenzel}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    11
*}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    12
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    13
theory Ring_and_Field = Inductive:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    14
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    15
text{*Lemmas and extension to semirings by L. C. Paulson*}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    16
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    17
subsection {* Abstract algebraic structures *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    18
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    19
axclass semiring \<subseteq> zero, one, plus, times
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    20
  add_assoc: "(a + b) + c = a + (b + c)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    21
  add_commute: "a + b = b + a"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    22
  left_zero [simp]: "0 + a = a"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    23
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    24
  mult_assoc: "(a * b) * c = a * (b * c)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    25
  mult_commute: "a * b = b * a"
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
    26
  mult_1 [simp]: "1 * a = a"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    27
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    28
  left_distrib: "(a + b) * c = a * c + b * c"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    29
  zero_neq_one [simp]: "0 \<noteq> 1"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    30
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    31
axclass ring \<subseteq> semiring, minus
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    32
  left_minus [simp]: "- a + a = 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    33
  diff_minus: "a - b = a + (-b)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    34
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    35
axclass ordered_semiring \<subseteq> semiring, linorder
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    36
  add_left_mono: "a \<le> b ==> c + a \<le> c + b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    37
  mult_strict_left_mono: "a < b ==> 0 < c ==> c * a < c * b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    38
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    39
axclass ordered_ring \<subseteq> ordered_semiring, ring
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    40
  abs_if: "\<bar>a\<bar> = (if a < 0 then -a else a)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    41
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    42
axclass field \<subseteq> ring, inverse
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    43
  left_inverse [simp]: "a \<noteq> 0 ==> inverse a * a = 1"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    44
  divide_inverse:      "b \<noteq> 0 ==> a / b = a * inverse b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    45
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    46
axclass ordered_field \<subseteq> ordered_ring, field
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    47
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    48
axclass division_by_zero \<subseteq> zero, inverse
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
    49
  inverse_zero [simp]: "inverse 0 = 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
    50
  divide_zero [simp]: "a / 0 = 0"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    51
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    52
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    53
subsection {* Derived Rules for Addition *}
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    54
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    55
lemma right_zero [simp]: "a + 0 = (a::'a::semiring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    56
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    57
  have "a + 0 = 0 + a" by (simp only: add_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    58
  also have "... = a" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    59
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    60
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    61
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    62
lemma add_left_commute: "a + (b + c) = b + (a + (c::'a::semiring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    63
  by (rule mk_left_commute [of "op +", OF add_assoc add_commute])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    64
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    65
theorems add_ac = add_assoc add_commute add_left_commute
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    66
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    67
lemma right_minus [simp]: "a + -(a::'a::ring) = 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    68
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    69
  have "a + -a = -a + a" by (simp add: add_ac)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    70
  also have "... = 0" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    71
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    72
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    73
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    74
lemma right_minus_eq: "(a - b = 0) = (a = (b::'a::ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    75
proof
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    76
  have "a = a - b + b" by (simp add: diff_minus add_ac)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    77
  also assume "a - b = 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    78
  finally show "a = b" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    79
next
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    80
  assume "a = b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    81
  thus "a - b = 0" by (simp add: diff_minus)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    82
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    83
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    84
lemma add_left_cancel [simp]:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    85
     "(a + b = a + c) = (b = (c::'a::ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    86
proof
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    87
  assume eq: "a + b = a + c"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
    88
  hence "(-a + a) + b = (-a + a) + c"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    89
    by (simp only: eq add_assoc)
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
    90
  thus "b = c" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    91
next
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    92
  assume eq: "b = c"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
    93
  thus "a + b = a + c" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    94
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    95
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    96
lemma add_right_cancel [simp]:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    97
     "(b + a = c + a) = (b = (c::'a::ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    98
  by (simp add: add_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    99
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   100
lemma minus_minus [simp]: "- (- (a::'a::ring)) = a"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   101
  proof (rule add_left_cancel [of "-a", THEN iffD1])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   102
    show "(-a + -(-a) = -a + a)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   103
    by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   104
  qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   105
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   106
lemma equals_zero_I: "a+b = 0 ==> -a = (b::'a::ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   107
apply (rule right_minus_eq [THEN iffD1, symmetric])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   108
apply (simp add: diff_minus add_commute) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   109
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   110
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   111
lemma minus_zero [simp]: "- 0 = (0::'a::ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   112
by (simp add: equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   113
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   114
lemma diff_self [simp]: "a - (a::'a::ring) = 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   115
  by (simp add: diff_minus)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   116
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   117
lemma diff_0 [simp]: "(0::'a::ring) - a = -a"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   118
by (simp add: diff_minus)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   119
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   120
lemma diff_0_right [simp]: "a - (0::'a::ring) = a" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   121
by (simp add: diff_minus)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   122
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   123
lemma neg_equal_iff_equal [simp]: "(-a = -b) = (a = (b::'a::ring))" 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   124
  proof 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   125
    assume "- a = - b"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   126
    hence "- (- a) = - (- b)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   127
      by simp
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   128
    thus "a=b" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   129
  next
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   130
    assume "a=b"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   131
    thus "-a = -b" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   132
  qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   133
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   134
lemma neg_equal_0_iff_equal [simp]: "(-a = 0) = (a = (0::'a::ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   135
by (subst neg_equal_iff_equal [symmetric], simp)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   136
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   137
lemma neg_0_equal_iff_equal [simp]: "(0 = -a) = (0 = (a::'a::ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   138
by (subst neg_equal_iff_equal [symmetric], simp)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   139
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   140
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   141
subsection {* Derived rules for multiplication *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   142
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   143
lemma mult_1_right [simp]: "a * (1::'a::semiring) = a"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   144
proof -
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   145
  have "a * 1 = 1 * a" by (simp add: mult_commute)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   146
  also have "... = a" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   147
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   148
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   149
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   150
lemma mult_left_commute: "a * (b * c) = b * (a * (c::'a::semiring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   151
  by (rule mk_left_commute [of "op *", OF mult_assoc mult_commute])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   152
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   153
theorems mult_ac = mult_assoc mult_commute mult_left_commute
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   154
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   155
lemma right_inverse [simp]:
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   156
      assumes not0: "a \<noteq> 0" shows "a * inverse (a::'a::field) = 1"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   157
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   158
  have "a * inverse a = inverse a * a" by (simp add: mult_ac)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   159
  also have "... = 1" using not0 by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   160
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   161
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   162
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   163
lemma right_inverse_eq: "b \<noteq> 0 ==> (a / b = 1) = (a = (b::'a::field))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   164
proof
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   165
  assume neq: "b \<noteq> 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   166
  {
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   167
    hence "a = (a / b) * b" by (simp add: divide_inverse mult_ac)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   168
    also assume "a / b = 1"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   169
    finally show "a = b" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   170
  next
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   171
    assume "a = b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   172
    with neq show "a / b = 1" by (simp add: divide_inverse)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   173
  }
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   174
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   175
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   176
lemma divide_self [simp]: "a \<noteq> 0 ==> a / (a::'a::field) = 1"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   177
  by (simp add: divide_inverse)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   178
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   179
lemma mult_left_zero [simp]: "0 * a = (0::'a::ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   180
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   181
  have "0*a + 0*a = 0*a + 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   182
    by (simp add: left_distrib [symmetric])
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   183
  thus ?thesis by (simp only: add_left_cancel)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   184
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   185
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   186
lemma mult_right_zero [simp]: "a * 0 = (0::'a::ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   187
  by (simp add: mult_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   188
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   189
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   190
subsection {* Distribution rules *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   191
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   192
lemma right_distrib: "a * (b + c) = a * b + a * (c::'a::semiring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   193
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   194
  have "a * (b + c) = (b + c) * a" by (simp add: mult_ac)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   195
  also have "... = b * a + c * a" by (simp only: left_distrib)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   196
  also have "... = a * b + a * c" by (simp add: mult_ac)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   197
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   198
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   199
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   200
theorems ring_distrib = right_distrib left_distrib
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   201
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   202
lemma minus_add_distrib [simp]: "- (a + b) = -a + -(b::'a::ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   203
apply (rule equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   204
apply (simp add: add_ac) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   205
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   206
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   207
lemma minus_mult_left: "- (a * b) = (-a) * (b::'a::ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   208
apply (rule equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   209
apply (simp add: left_distrib [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   210
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   211
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   212
lemma minus_mult_right: "- (a * b) = a * -(b::'a::ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   213
apply (rule equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   214
apply (simp add: right_distrib [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   215
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   216
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   217
lemma minus_mult_minus [simp]: "(- a) * (- b) = a * (b::'a::ring)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   218
  by (simp add: minus_mult_left [symmetric] minus_mult_right [symmetric])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   219
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   220
lemma right_diff_distrib: "a * (b - c) = a * b - a * (c::'a::ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   221
by (simp add: right_distrib diff_minus 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   222
              minus_mult_left [symmetric] minus_mult_right [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   223
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   224
lemma minus_diff_eq [simp]: "- (a - b) = b - (a::'a::ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   225
by (simp add: diff_minus add_commute) 
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   226
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   227
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   228
subsection {* Ordering Rules for Addition *}
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   229
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   230
lemma add_right_mono: "a \<le> (b::'a::ordered_semiring) ==> a + c \<le> b + c"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   231
by (simp add: add_commute [of _ c] add_left_mono)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   232
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   233
text {* non-strict, in both arguments *}
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   234
lemma add_mono: "[|a \<le> b;  c \<le> d|] ==> a + c \<le> b + (d::'a::ordered_semiring)"
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   235
  apply (erule add_right_mono [THEN order_trans])
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   236
  apply (simp add: add_commute add_left_mono)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   237
  done
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   238
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   239
lemma add_strict_left_mono:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   240
     "a < b ==> c + a < c + (b::'a::ordered_ring)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   241
 by (simp add: order_less_le add_left_mono) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   242
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   243
lemma add_strict_right_mono:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   244
     "a < b ==> a + c < b + (c::'a::ordered_ring)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   245
 by (simp add: add_commute [of _ c] add_strict_left_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   246
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   247
text{*Strict monotonicity in both arguments*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   248
lemma add_strict_mono: "[|a<b; c<d|] ==> a + c < b + (d::'a::ordered_ring)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   249
apply (erule add_strict_right_mono [THEN order_less_trans])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   250
apply (erule add_strict_left_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   251
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   252
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   253
lemma add_less_imp_less_left:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   254
      assumes less: "c + a < c + b"  shows "a < (b::'a::ordered_ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   255
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   256
  have "-c + (c + a) < -c + (c + b)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   257
    by (rule add_strict_left_mono [OF less])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   258
  thus "a < b" by (simp add: add_assoc [symmetric])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   259
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   260
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   261
lemma add_less_imp_less_right:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   262
      "a + c < b + c ==> a < (b::'a::ordered_ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   263
apply (rule add_less_imp_less_left [of c])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   264
apply (simp add: add_commute)  
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   265
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   266
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   267
lemma add_less_cancel_left [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   268
    "(c+a < c+b) = (a < (b::'a::ordered_ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   269
by (blast intro: add_less_imp_less_left add_strict_left_mono) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   271
lemma add_less_cancel_right [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   272
    "(a+c < b+c) = (a < (b::'a::ordered_ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   273
by (blast intro: add_less_imp_less_right add_strict_right_mono)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   274
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   275
lemma add_le_cancel_left [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   276
    "(c+a \<le> c+b) = (a \<le> (b::'a::ordered_ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   277
by (simp add: linorder_not_less [symmetric]) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   278
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   279
lemma add_le_cancel_right [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   280
    "(a+c \<le> b+c) = (a \<le> (b::'a::ordered_ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   281
by (simp add: linorder_not_less [symmetric]) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   282
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   283
lemma add_le_imp_le_left:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   284
      "c + a \<le> c + b ==> a \<le> (b::'a::ordered_ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   285
by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   286
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   287
lemma add_le_imp_le_right:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   288
      "a + c \<le> b + c ==> a \<le> (b::'a::ordered_ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   289
by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   290
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   291
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   292
subsection {* Ordering Rules for Unary Minus *}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   293
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   294
lemma le_imp_neg_le:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   295
      assumes "a \<le> (b::'a::ordered_ring)" shows "-b \<le> -a"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   296
  proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   297
  have "-a+a \<le> -a+b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   298
    by (rule add_left_mono) 
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   299
  hence "0 \<le> -a+b"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   300
    by simp
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   301
  hence "0 + (-b) \<le> (-a + b) + (-b)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   302
    by (rule add_right_mono) 
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   303
  thus ?thesis
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   304
    by (simp add: add_assoc)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   305
  qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   306
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   307
lemma neg_le_iff_le [simp]: "(-b \<le> -a) = (a \<le> (b::'a::ordered_ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   308
  proof 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   309
    assume "- b \<le> - a"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   310
    hence "- (- a) \<le> - (- b)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   311
      by (rule le_imp_neg_le)
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   312
    thus "a\<le>b" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   313
  next
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   314
    assume "a\<le>b"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   315
    thus "-b \<le> -a" by (rule le_imp_neg_le)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   316
  qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   317
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   318
lemma neg_le_0_iff_le [simp]: "(-a \<le> 0) = (0 \<le> (a::'a::ordered_ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   319
by (subst neg_le_iff_le [symmetric], simp)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   320
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   321
lemma neg_0_le_iff_le [simp]: "(0 \<le> -a) = (a \<le> (0::'a::ordered_ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   322
by (subst neg_le_iff_le [symmetric], simp)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   323
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   324
lemma neg_less_iff_less [simp]: "(-b < -a) = (a < (b::'a::ordered_ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   325
by (force simp add: order_less_le) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   326
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   327
lemma neg_less_0_iff_less [simp]: "(-a < 0) = (0 < (a::'a::ordered_ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   328
by (subst neg_less_iff_less [symmetric], simp)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   329
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   330
lemma neg_0_less_iff_less [simp]: "(0 < -a) = (a < (0::'a::ordered_ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   331
by (subst neg_less_iff_less [symmetric], simp)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   332
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   333
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   334
subsection{*Subtraction Laws*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   335
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   336
lemma add_diff_eq: "a + (b - c) = (a + b) - (c::'a::ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   337
by (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   338
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   339
lemma diff_add_eq: "(a - b) + c = (a + c) - (b::'a::ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   340
by (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   341
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   342
lemma diff_eq_eq: "(a-b = c) = (a = c + (b::'a::ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   343
by (auto simp add: diff_minus add_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   344
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   345
lemma eq_diff_eq: "(a = c-b) = (a + (b::'a::ring) = c)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   346
by (auto simp add: diff_minus add_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   347
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   348
lemma diff_diff_eq: "(a - b) - c = a - (b + (c::'a::ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   349
by (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   350
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   351
lemma diff_diff_eq2: "a - (b - c) = (a + c) - (b::'a::ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   352
by (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   353
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   354
text{*Further subtraction laws for ordered rings*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   355
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   356
lemma less_eq_diff: "(a < b) = (a - b < (0::'a::ordered_ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   357
proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   358
  have  "(a < b) = (a + (- b) < b + (-b))"  
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   359
    by (simp only: add_less_cancel_right)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   360
  also have "... =  (a - b < 0)" by (simp add: diff_minus)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   361
  finally show ?thesis .
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   362
qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   363
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   364
lemma diff_less_eq: "(a-b < c) = (a < c + (b::'a::ordered_ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   365
apply (subst less_eq_diff)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   366
apply (rule less_eq_diff [of _ c, THEN ssubst])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   367
apply (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   368
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   369
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   370
lemma less_diff_eq: "(a < c-b) = (a + (b::'a::ordered_ring) < c)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   371
apply (subst less_eq_diff)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   372
apply (rule less_eq_diff [of _ "c-b", THEN ssubst])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   373
apply (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   374
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   375
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   376
lemma diff_le_eq: "(a-b \<le> c) = (a \<le> c + (b::'a::ordered_ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   377
by (simp add: linorder_not_less [symmetric] less_diff_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   378
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   379
lemma le_diff_eq: "(a \<le> c-b) = (a + (b::'a::ordered_ring) \<le> c)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   380
by (simp add: linorder_not_less [symmetric] diff_less_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   381
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   382
text{*This list of rewrites simplifies (in)equalities by bringing subtractions
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   383
  to the top and then moving negative terms to the other side.
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   384
  Use with @{text add_ac}*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   385
lemmas compare_rls =
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   386
       diff_minus [symmetric]
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   387
       add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   388
       diff_less_eq less_diff_eq diff_le_eq le_diff_eq
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   389
       diff_eq_eq eq_diff_eq
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   390
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   391
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   392
subsection {* Ordering Rules for Multiplication *}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   393
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   394
lemma mult_strict_right_mono:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   395
     "[|a < b; 0 < c|] ==> a * c < b * (c::'a::ordered_semiring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   396
by (simp add: mult_commute [of _ c] mult_strict_left_mono)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   397
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   398
lemma mult_left_mono:
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   399
     "[|a \<le> b; 0 \<le> c|] ==> c * a \<le> c * (b::'a::ordered_ring)"
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   400
  apply (case_tac "c=0", simp)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   401
  apply (force simp add: mult_strict_left_mono order_le_less) 
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   402
  done
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   403
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   404
lemma mult_right_mono:
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   405
     "[|a \<le> b; 0 \<le> c|] ==> a*c \<le> b * (c::'a::ordered_ring)"
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   406
  by (simp add: mult_left_mono mult_commute [of _ c]) 
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   407
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   408
lemma mult_strict_left_mono_neg:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   409
     "[|b < a; c < 0|] ==> c * a < c * (b::'a::ordered_ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   410
apply (drule mult_strict_left_mono [of _ _ "-c"])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   411
apply (simp_all add: minus_mult_left [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   412
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   413
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   414
lemma mult_strict_right_mono_neg:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   415
     "[|b < a; c < 0|] ==> a * c < b * (c::'a::ordered_ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   416
apply (drule mult_strict_right_mono [of _ _ "-c"])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   417
apply (simp_all add: minus_mult_right [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   418
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   419
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   420
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   421
subsection{* Products of Signs *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   422
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   423
lemma mult_pos: "[| (0::'a::ordered_ring) < a; 0 < b |] ==> 0 < a*b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   424
by (drule mult_strict_left_mono [of 0 b], auto)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   425
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   426
lemma mult_pos_neg: "[| (0::'a::ordered_ring) < a; b < 0 |] ==> a*b < 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   427
by (drule mult_strict_left_mono [of b 0], auto)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   428
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   429
lemma mult_neg: "[| a < (0::'a::ordered_ring); b < 0 |] ==> 0 < a*b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   430
by (drule mult_strict_right_mono_neg, auto)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   431
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   432
lemma zero_less_mult_pos: "[| 0 < a*b; 0 < a|] ==> 0 < (b::'a::ordered_ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   433
apply (case_tac "b\<le>0") 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   434
 apply (auto simp add: order_le_less linorder_not_less)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   435
apply (drule_tac mult_pos_neg [of a b]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   436
 apply (auto dest: order_less_not_sym)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   437
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   438
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   439
lemma zero_less_mult_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   440
     "((0::'a::ordered_ring) < a*b) = (0 < a & 0 < b | a < 0 & b < 0)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   441
apply (auto simp add: order_le_less linorder_not_less mult_pos mult_neg)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   442
apply (blast dest: zero_less_mult_pos) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   443
apply (simp add: mult_commute [of a b]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   444
apply (blast dest: zero_less_mult_pos) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   445
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   446
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   447
lemma mult_eq_0_iff [simp]: "(a*b = (0::'a::ordered_ring)) = (a = 0 | b = 0)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   448
apply (case_tac "a < 0")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   449
apply (auto simp add: linorder_not_less order_le_less linorder_neq_iff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   450
apply (force dest: mult_strict_right_mono_neg mult_strict_right_mono)+
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   451
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   452
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   453
lemma zero_le_mult_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   454
     "((0::'a::ordered_ring) \<le> a*b) = (0 \<le> a & 0 \<le> b | a \<le> 0 & b \<le> 0)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   455
by (auto simp add: eq_commute [of 0] order_le_less linorder_not_less
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   456
                   zero_less_mult_iff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   457
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   458
lemma mult_less_0_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   459
     "(a*b < (0::'a::ordered_ring)) = (0 < a & b < 0 | a < 0 & 0 < b)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   460
apply (insert zero_less_mult_iff [of "-a" b]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   461
apply (force simp add: minus_mult_left[symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   462
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   463
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   464
lemma mult_le_0_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   465
     "(a*b \<le> (0::'a::ordered_ring)) = (0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   466
apply (insert zero_le_mult_iff [of "-a" b]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   467
apply (force simp add: minus_mult_left[symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   468
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   469
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   470
lemma zero_le_square: "(0::'a::ordered_ring) \<le> a*a"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   471
by (simp add: zero_le_mult_iff linorder_linear) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   472
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   473
lemma zero_less_one: "(0::'a::ordered_ring) < 1"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   474
apply (insert zero_le_square [of 1]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   475
apply (simp add: order_less_le) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   476
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   477
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   478
lemma zero_le_one: "(0::'a::ordered_ring) \<le> 1"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   479
  by (rule zero_less_one [THEN order_less_imp_le]) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   480
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   481
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   482
subsection{*More Monotonicity*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   483
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   484
lemma mult_left_mono_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   485
     "[|b \<le> a; c \<le> 0|] ==> c * a \<le> c * (b::'a::ordered_ring)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   486
apply (drule mult_left_mono [of _ _ "-c"]) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   487
apply (simp_all add: minus_mult_left [symmetric]) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   488
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   489
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   490
lemma mult_right_mono_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   491
     "[|b \<le> a; c \<le> 0|] ==> a * c \<le> b * (c::'a::ordered_ring)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   492
  by (simp add: mult_left_mono_neg mult_commute [of _ c]) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   493
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   494
text{*Strict monotonicity in both arguments*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   495
lemma mult_strict_mono:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   496
     "[|a<b; c<d; 0<b; 0\<le>c|] ==> a * c < b * (d::'a::ordered_ring)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   497
apply (case_tac "c=0")
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   498
 apply (simp add: mult_pos) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   499
apply (erule mult_strict_right_mono [THEN order_less_trans])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   500
 apply (force simp add: order_le_less) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   501
apply (erule mult_strict_left_mono, assumption)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   502
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   503
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   504
text{*This weaker variant has more natural premises*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   505
lemma mult_strict_mono':
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   506
     "[| a<b; c<d; 0 \<le> a; 0 \<le> c|] ==> a * c < b * (d::'a::ordered_ring)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   507
apply (rule mult_strict_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   508
apply (blast intro: order_le_less_trans)+
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   509
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   510
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   511
lemma mult_mono:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   512
     "[|a \<le> b; c \<le> d; 0 \<le> b; 0 \<le> c|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   513
      ==> a * c  \<le>  b * (d::'a::ordered_ring)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   514
apply (erule mult_right_mono [THEN order_trans], assumption)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   515
apply (erule mult_left_mono, assumption)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   516
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   517
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   518
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   519
subsection{*Cancellation Laws for Relationships With a Common Factor*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   520
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   521
text{*Cancellation laws for @{term "c*a < c*b"} and @{term "a*c < b*c"},
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   522
   also with the relations @{text "\<le>"} and equality.*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   523
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   524
lemma mult_less_cancel_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   525
    "(a*c < b*c) = ((0 < c & a < b) | (c < 0 & b < (a::'a::ordered_ring)))"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   526
apply (case_tac "c = 0")
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   527
apply (auto simp add: linorder_neq_iff mult_strict_right_mono 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   528
                      mult_strict_right_mono_neg)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   529
apply (auto simp add: linorder_not_less 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   530
                      linorder_not_le [symmetric, of "a*c"]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   531
                      linorder_not_le [symmetric, of a])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   532
apply (erule_tac [!] notE)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   533
apply (auto simp add: order_less_imp_le mult_right_mono 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   534
                      mult_right_mono_neg)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   535
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   536
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   537
lemma mult_less_cancel_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   538
    "(c*a < c*b) = ((0 < c & a < b) | (c < 0 & b < (a::'a::ordered_ring)))"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   539
by (simp add: mult_commute [of c] mult_less_cancel_right)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   540
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   541
lemma mult_le_cancel_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   542
     "(a*c \<le> b*c) = ((0<c --> a\<le>b) & (c<0 --> b \<le> (a::'a::ordered_ring)))"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   543
by (simp add: linorder_not_less [symmetric] mult_less_cancel_right)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   544
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   545
lemma mult_le_cancel_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   546
     "(c*a \<le> c*b) = ((0<c --> a\<le>b) & (c<0 --> b \<le> (a::'a::ordered_ring)))"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   547
by (simp add: mult_commute [of c] mult_le_cancel_right)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   548
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   549
lemma mult_less_imp_less_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   550
    "[|c*a < c*b; 0 < c|] ==> a < (b::'a::ordered_ring)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   551
  by (force elim: order_less_asym simp add: mult_less_cancel_left)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   552
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   553
lemma mult_less_imp_less_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   554
    "[|a*c < b*c; 0 < c|] ==> a < (b::'a::ordered_ring)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   555
  by (force elim: order_less_asym simp add: mult_less_cancel_right)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   556
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   557
text{*Cancellation of equalities with a common factor*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   558
lemma mult_cancel_right [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   559
     "(a*c = b*c) = (c = (0::'a::ordered_ring) | a=b)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   560
apply (cut_tac linorder_less_linear [of 0 c])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   561
apply (force dest: mult_strict_right_mono_neg mult_strict_right_mono
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   562
             simp add: linorder_neq_iff)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   563
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   564
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   565
text{*These cancellation theorems require an ordering. Versions are proved
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   566
      below that work for fields without an ordering.*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   567
lemma mult_cancel_left [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   568
     "(c*a = c*b) = (c = (0::'a::ordered_ring) | a=b)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   569
by (simp add: mult_commute [of c] mult_cancel_right)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   570
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   571
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   572
subsection {* Absolute Value *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   573
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   574
text{*But is it really better than just rewriting with @{text abs_if}?*}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   575
lemma abs_split:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   576
     "P(abs(a::'a::ordered_ring)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   577
by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   578
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   579
lemma abs_zero [simp]: "abs 0 = (0::'a::ordered_ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   580
by (simp add: abs_if)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   581
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   582
lemma abs_mult: "abs (x * y) = abs x * abs (y::'a::ordered_ring)" 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   583
apply (case_tac "x=0 | y=0", force) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   584
apply (auto elim: order_less_asym
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   585
            simp add: abs_if mult_less_0_iff linorder_neq_iff
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   586
                  minus_mult_left [symmetric] minus_mult_right [symmetric])  
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   587
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   588
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   589
lemma abs_eq_0 [simp]: "(abs x = 0) = (x = (0::'a::ordered_ring))"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   590
by (simp add: abs_if)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   591
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   592
lemma zero_less_abs_iff [simp]: "(0 < abs x) = (x ~= (0::'a::ordered_ring))"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   593
by (simp add: abs_if linorder_neq_iff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   594
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   595
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   596
subsection {* Fields *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   597
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   598
text{*Compared with @{text mult_eq_0_iff}, this version removes the requirement
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   599
      of an ordering.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   600
lemma field_mult_eq_0_iff: "(a*b = (0::'a::field)) = (a = 0 | b = 0)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   601
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   602
    assume "a=0" thus ?thesis by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   603
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   604
    assume anz [simp]: "a\<noteq>0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   605
    thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   606
    proof auto
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   607
      assume "a * b = 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   608
      hence "inverse a * (a * b) = 0" by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   609
      thus "b = 0"  by (simp (no_asm_use) add: mult_assoc [symmetric])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   610
    qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   611
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   612
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   613
text{*Cancellation of equalities with a common factor*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   614
lemma field_mult_cancel_right_lemma:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   615
      assumes cnz: "c \<noteq> (0::'a::field)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   616
	  and eq:  "a*c = b*c"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   617
	 shows "a=b"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   618
  proof -
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   619
  have "(a * c) * inverse c = (b * c) * inverse c"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   620
    by (simp add: eq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   621
  thus "a=b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   622
    by (simp add: mult_assoc cnz)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   623
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   624
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   625
lemma field_mult_cancel_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   626
     "(a*c = b*c) = (c = (0::'a::field) | a=b)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   627
  proof cases
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   628
    assume "c=0" thus ?thesis by simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   629
  next
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   630
    assume "c\<noteq>0" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   631
    thus ?thesis by (force dest: field_mult_cancel_right_lemma)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   632
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   633
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   634
lemma field_mult_cancel_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   635
     "(c*a = c*b) = (c = (0::'a::field) | a=b)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   636
  by (simp add: mult_commute [of c] field_mult_cancel_right) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   637
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   638
lemma nonzero_imp_inverse_nonzero: "a \<noteq> 0 ==> inverse a \<noteq> (0::'a::field)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   639
  proof
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   640
  assume ianz: "inverse a = 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   641
  assume "a \<noteq> 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   642
  hence "1 = a * inverse a" by simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   643
  also have "... = 0" by (simp add: ianz)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   644
  finally have "1 = (0::'a::field)" .
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   645
  thus False by (simp add: eq_commute)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   646
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   647
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   648
lemma inverse_zero_imp_zero: "inverse a = 0 ==> a = (0::'a::field)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   649
apply (rule ccontr) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   650
apply (blast dest: nonzero_imp_inverse_nonzero) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   651
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   652
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   653
lemma inverse_nonzero_imp_nonzero:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   654
   "inverse a = 0 ==> a = (0::'a::field)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   655
apply (rule ccontr) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   656
apply (blast dest: nonzero_imp_inverse_nonzero) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   657
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   658
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   659
lemma inverse_nonzero_iff_nonzero [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   660
   "(inverse a = 0) = (a = (0::'a::{field,division_by_zero}))"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   661
by (force dest: inverse_nonzero_imp_nonzero) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   662
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   663
lemma nonzero_inverse_minus_eq:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   664
      assumes [simp]: "a\<noteq>0"  shows "inverse(-a) = -inverse(a::'a::field)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   665
  proof -
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   666
    have "-a * inverse (- a) = -a * - inverse a"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   667
      by simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   668
    thus ?thesis 
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   669
      by (simp only: field_mult_cancel_left, simp)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   670
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   671
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   672
lemma inverse_minus_eq [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   673
     "inverse(-a) = -inverse(a::'a::{field,division_by_zero})";
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   674
  proof cases
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   675
    assume "a=0" thus ?thesis by (simp add: inverse_zero)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   676
  next
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   677
    assume "a\<noteq>0" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   678
    thus ?thesis by (simp add: nonzero_inverse_minus_eq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   679
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   680
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   681
lemma nonzero_inverse_eq_imp_eq:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   682
      assumes inveq: "inverse a = inverse b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   683
	  and anz:  "a \<noteq> 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   684
	  and bnz:  "b \<noteq> 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   685
	 shows "a = (b::'a::field)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   686
  proof -
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   687
  have "a * inverse b = a * inverse a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   688
    by (simp add: inveq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   689
  hence "(a * inverse b) * b = (a * inverse a) * b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   690
    by simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   691
  thus "a = b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   692
    by (simp add: mult_assoc anz bnz)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   693
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   694
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   695
lemma inverse_eq_imp_eq:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   696
     "inverse a = inverse b ==> a = (b::'a::{field,division_by_zero})"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   697
apply (case_tac "a=0 | b=0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   698
 apply (force dest!: inverse_zero_imp_zero
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   699
              simp add: eq_commute [of "0::'a"])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   700
apply (force dest!: nonzero_inverse_eq_imp_eq) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   701
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   702
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   703
lemma inverse_eq_iff_eq [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   704
     "(inverse a = inverse b) = (a = (b::'a::{field,division_by_zero}))"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   705
by (force dest!: inverse_eq_imp_eq) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   706
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   707
lemma nonzero_inverse_inverse_eq:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   708
      assumes [simp]: "a \<noteq> 0"  shows "inverse(inverse (a::'a::field)) = a"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   709
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   710
  have "(inverse (inverse a) * inverse a) * a = a" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   711
    by (simp add: nonzero_imp_inverse_nonzero)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   712
  thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   713
    by (simp add: mult_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   714
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   715
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   716
lemma inverse_inverse_eq [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   717
     "inverse(inverse (a::'a::{field,division_by_zero})) = a"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   718
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   719
    assume "a=0" thus ?thesis by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   720
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   721
    assume "a\<noteq>0" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   722
    thus ?thesis by (simp add: nonzero_inverse_inverse_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   723
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   724
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   725
lemma inverse_1 [simp]: "inverse 1 = (1::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   726
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   727
  have "inverse 1 * 1 = (1::'a::field)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   728
    by (rule left_inverse [OF zero_neq_one [symmetric]])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   729
  thus ?thesis  by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   730
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   731
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   732
lemma nonzero_inverse_mult_distrib: 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   733
      assumes anz: "a \<noteq> 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   734
          and bnz: "b \<noteq> 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   735
      shows "inverse(a*b) = inverse(b) * inverse(a::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   736
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   737
  have "inverse(a*b) * (a * b) * inverse(b) = inverse(b)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   738
    by (simp add: field_mult_eq_0_iff anz bnz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   739
  hence "inverse(a*b) * a = inverse(b)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   740
    by (simp add: mult_assoc bnz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   741
  hence "inverse(a*b) * a * inverse(a) = inverse(b) * inverse(a)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   742
    by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   743
  thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   744
    by (simp add: mult_assoc anz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   745
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   746
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   747
text{*This version builds in division by zero while also re-orienting
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   748
      the right-hand side.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   749
lemma inverse_mult_distrib [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   750
     "inverse(a*b) = inverse(a) * inverse(b::'a::{field,division_by_zero})"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   751
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   752
    assume "a \<noteq> 0 & b \<noteq> 0" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   753
    thus ?thesis  by (simp add: nonzero_inverse_mult_distrib mult_commute)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   754
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   755
    assume "~ (a \<noteq> 0 & b \<noteq> 0)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   756
    thus ?thesis  by force
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   757
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   758
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   759
text{*There is no slick version using division by zero.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   760
lemma inverse_add:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   761
     "[|a \<noteq> 0;  b \<noteq> 0|]
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   762
      ==> inverse a + inverse b = (a+b) * inverse a * inverse (b::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   763
apply (simp add: left_distrib mult_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   764
apply (simp add: mult_commute [of "inverse a"]) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   765
apply (simp add: mult_assoc [symmetric] add_commute)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   766
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   767
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   768
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   769
subsection {* Ordered Fields *}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   770
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   771
lemma inverse_gt_0: 
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   772
      assumes a_gt_0: "0 < a"  shows "0 < inverse (a::'a::ordered_field)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   773
  proof -
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   774
  have "0 < a * inverse a" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   775
    by (simp add: a_gt_0 [THEN order_less_imp_not_eq2] zero_less_one)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   776
  thus "0 < inverse a" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   777
    by (simp add: a_gt_0 [THEN order_less_not_sym] zero_less_mult_iff)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   778
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   779
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   780
lemma inverse_less_0:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   781
     "a < 0 ==> inverse a < (0::'a::ordered_field)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   782
  by (insert inverse_gt_0 [of "-a"], 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   783
      simp add: nonzero_inverse_minus_eq order_less_imp_not_eq) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   784
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   785
lemma inverse_le_imp_le:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   786
      assumes invle: "inverse a \<le> inverse b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   787
	  and apos:  "0 < a"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   788
	 shows "b \<le> (a::'a::ordered_field)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   789
  proof (rule classical)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   790
  assume "~ b \<le> a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   791
  hence "a < b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   792
    by (simp add: linorder_not_le)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   793
  hence bpos: "0 < b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   794
    by (blast intro: apos order_less_trans)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   795
  hence "a * inverse a \<le> a * inverse b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   796
    by (simp add: apos invle order_less_imp_le mult_left_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   797
  hence "(a * inverse a) * b \<le> (a * inverse b) * b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   798
    by (simp add: bpos order_less_imp_le mult_right_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   799
  thus "b \<le> a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   800
    by (simp add: mult_assoc apos bpos order_less_imp_not_eq2)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   801
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   802
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   803
lemma less_imp_inverse_less:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   804
      assumes less: "a < b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   805
	  and apos:  "0 < a"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   806
	shows "inverse b < inverse (a::'a::ordered_field)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   807
  proof (rule ccontr)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   808
  assume "~ inverse b < inverse a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   809
  hence "inverse a \<le> inverse b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   810
    by (simp add: linorder_not_less)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   811
  hence "~ (a < b)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   812
    by (simp add: linorder_not_less inverse_le_imp_le [OF _ apos])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   813
  thus False
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   814
    by (rule notE [OF _ less])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   815
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   816
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   817
lemma inverse_less_imp_less:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   818
   "[|inverse a < inverse b; 0 < a|] ==> b < (a::'a::ordered_field)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   819
apply (simp add: order_less_le [of "inverse a"] order_less_le [of "b"])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   820
apply (force dest!: inverse_le_imp_le nonzero_inverse_eq_imp_eq) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   821
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   822
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   823
text{*Both premises are essential. Consider -1 and 1.*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   824
lemma inverse_less_iff_less [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   825
     "[|0 < a; 0 < b|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   826
      ==> (inverse a < inverse b) = (b < (a::'a::ordered_field))"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   827
by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   828
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   829
lemma le_imp_inverse_le:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   830
   "[|a \<le> b; 0 < a|] ==> inverse b \<le> inverse (a::'a::ordered_field)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   831
  by (force simp add: order_le_less less_imp_inverse_less)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   832
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   833
lemma inverse_le_iff_le [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   834
     "[|0 < a; 0 < b|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   835
      ==> (inverse a \<le> inverse b) = (b \<le> (a::'a::ordered_field))"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   836
by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   837
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   838
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   839
text{*These results refer to both operands being negative.  The opposite-sign
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   840
case is trivial, since inverse preserves signs.*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   841
lemma inverse_le_imp_le_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   842
   "[|inverse a \<le> inverse b; b < 0|] ==> b \<le> (a::'a::ordered_field)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   843
  apply (rule classical) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   844
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   845
   prefer 2 apply (force simp add: linorder_not_le intro: order_less_trans) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   846
  apply (insert inverse_le_imp_le [of "-b" "-a"])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   847
  apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   848
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   849
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   850
lemma less_imp_inverse_less_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   851
   "[|a < b; b < 0|] ==> inverse b < inverse (a::'a::ordered_field)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   852
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   853
   prefer 2 apply (blast intro: order_less_trans) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   854
  apply (insert less_imp_inverse_less [of "-b" "-a"])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   855
  apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   856
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   857
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   858
lemma inverse_less_imp_less_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   859
   "[|inverse a < inverse b; b < 0|] ==> b < (a::'a::ordered_field)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   860
  apply (rule classical) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   861
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   862
   prefer 2
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   863
   apply (force simp add: linorder_not_less intro: order_le_less_trans) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   864
  apply (insert inverse_less_imp_less [of "-b" "-a"])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   865
  apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   866
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   867
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   868
lemma inverse_less_iff_less_neg [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   869
     "[|a < 0; b < 0|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   870
      ==> (inverse a < inverse b) = (b < (a::'a::ordered_field))"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   871
  apply (insert inverse_less_iff_less [of "-b" "-a"])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   872
  apply (simp del: inverse_less_iff_less 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   873
	      add: order_less_imp_not_eq nonzero_inverse_minus_eq) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   874
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   875
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   876
lemma le_imp_inverse_le_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   877
   "[|a \<le> b; b < 0|] ==> inverse b \<le> inverse (a::'a::ordered_field)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   878
  by (force simp add: order_le_less less_imp_inverse_less_neg)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   879
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   880
lemma inverse_le_iff_le_neg [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   881
     "[|a < 0; b < 0|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   882
      ==> (inverse a \<le> inverse b) = (b \<le> (a::'a::ordered_field))"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   883
by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) 
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   884
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   885
end