src/HOL/Ring_and_Field.thy
author paulson
Mon, 22 Dec 2003 12:50:01 +0100
changeset 14308 c0489217deb2
parent 14305 f17ca9f6dc8c
child 14321 55c688d2eefa
permissions -rw-r--r--
removing obsolete bindings
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"
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
    22
  add_0 [simp]: "0 + a = a"
14265
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
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
    55
lemma add_0_right [simp]: "a + 0 = (a::'a::semiring)"
14265
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
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   123
lemma diff_minus_eq_add [simp]: "a - - b = a + (b::'a::ring)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   124
by (simp add: diff_minus)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   125
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   126
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
   127
  proof 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   128
    assume "- a = - b"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   129
    hence "- (- a) = - (- b)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   130
      by simp
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
  next
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   133
    assume "a=b"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   134
    thus "-a = -b" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   135
  qed
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_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
   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
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
   141
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
   142
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   143
text{*The next two equations can make the simplifier loop!*}
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   144
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   145
lemma equation_minus_iff: "(a = - b) = (b = - (a::'a::ring))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   146
  proof -
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   147
  have "(- (-a) = - b) = (- a = b)" by (rule neg_equal_iff_equal)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   148
  thus ?thesis by (simp add: eq_commute)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   149
  qed
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   150
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   151
lemma minus_equation_iff: "(- a = b) = (- (b::'a::ring) = a)"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   152
  proof -
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   153
  have "(- a = - (-b)) = (a = -b)" by (rule neg_equal_iff_equal)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   154
  thus ?thesis by (simp add: eq_commute)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   155
  qed
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   156
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   157
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   158
subsection {* Derived rules for multiplication *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   159
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   160
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
   161
proof -
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   162
  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
   163
  also have "... = a" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   164
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   165
qed
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
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
   168
  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
   169
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   170
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
   171
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   172
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
   173
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   174
  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
   175
    by (simp add: left_distrib [symmetric])
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   176
  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
   177
qed
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_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
   180
  by (simp add: mult_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   181
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   182
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   183
subsection {* Distribution rules *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   184
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   185
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
   186
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   187
  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
   188
  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
   189
  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
   190
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   191
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   192
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   193
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
   194
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   195
text{*For the @{text combine_numerals} simproc*}
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   196
lemma combine_common_factor: "a*e + (b*e + c) = (a+b)*e + (c::'a::semiring)"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   197
by (simp add: left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   198
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   199
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
   200
apply (rule equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   201
apply (simp add: add_ac) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   202
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   203
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   204
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
   205
apply (rule equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   206
apply (simp add: left_distrib [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   207
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   208
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   209
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
   210
apply (rule equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   211
apply (simp add: right_distrib [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   212
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   213
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   214
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
   215
  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
   216
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   217
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
   218
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
   219
              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
   220
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   221
lemma left_diff_distrib: "(a - b) * c = a * c - b * (c::'a::ring)"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   222
by (simp add: mult_commute [of _ c] right_diff_distrib) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
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
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   333
text{*The next several equations can make the simplifier loop!*}
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   334
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   335
lemma less_minus_iff: "(a < - b) = (b < - (a::'a::ordered_ring))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   336
  proof -
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   337
  have "(- (-a) < - b) = (b < - a)" by (rule neg_less_iff_less)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   338
  thus ?thesis by simp
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   339
  qed
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   340
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   341
lemma minus_less_iff: "(- a < b) = (- b < (a::'a::ordered_ring))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   342
  proof -
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   343
  have "(- a < - (-b)) = (- b < a)" by (rule neg_less_iff_less)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   344
  thus ?thesis by simp
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   345
  qed
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   346
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   347
lemma le_minus_iff: "(a \<le> - b) = (b \<le> - (a::'a::ordered_ring))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   348
apply (simp add: linorder_not_less [symmetric])
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   349
apply (rule minus_less_iff) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   350
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   351
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   352
lemma minus_le_iff: "(- a \<le> b) = (- b \<le> (a::'a::ordered_ring))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   353
apply (simp add: linorder_not_less [symmetric])
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   354
apply (rule less_minus_iff) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   355
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   356
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   357
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   358
subsection{*Subtraction Laws*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   359
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   360
lemma add_diff_eq: "a + (b - c) = (a + b) - (c::'a::ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   361
by (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   362
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   363
lemma diff_add_eq: "(a - b) + c = (a + c) - (b::'a::ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   364
by (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   365
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   366
lemma diff_eq_eq: "(a-b = c) = (a = c + (b::'a::ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   367
by (auto simp add: diff_minus add_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   368
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   369
lemma eq_diff_eq: "(a = c-b) = (a + (b::'a::ring) = c)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   370
by (auto simp add: diff_minus add_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   371
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   372
lemma diff_diff_eq: "(a - b) - c = a - (b + (c::'a::ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   373
by (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   374
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   375
lemma diff_diff_eq2: "a - (b - c) = (a + c) - (b::'a::ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   376
by (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   377
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   378
text{*Further subtraction laws for ordered rings*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   379
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   380
lemma less_iff_diff_less_0: "(a < b) = (a - b < (0::'a::ordered_ring))"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   381
proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   382
  have  "(a < b) = (a + (- b) < b + (-b))"  
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   383
    by (simp only: add_less_cancel_right)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   384
  also have "... =  (a - b < 0)" by (simp add: diff_minus)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   385
  finally show ?thesis .
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   386
qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   387
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   388
lemma diff_less_eq: "(a-b < c) = (a < c + (b::'a::ordered_ring))"
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   389
apply (subst less_iff_diff_less_0)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   390
apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   391
apply (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   392
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   393
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   394
lemma less_diff_eq: "(a < c-b) = (a + (b::'a::ordered_ring) < c)"
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   395
apply (subst less_iff_diff_less_0)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   396
apply (rule less_iff_diff_less_0 [of _ "c-b", THEN ssubst])
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   397
apply (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   398
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   399
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   400
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
   401
by (simp add: linorder_not_less [symmetric] less_diff_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   402
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   403
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
   404
by (simp add: linorder_not_less [symmetric] diff_less_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   405
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   406
text{*This list of rewrites simplifies (in)equalities by bringing subtractions
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   407
  to the top and then moving negative terms to the other side.
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   408
  Use with @{text add_ac}*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   409
lemmas compare_rls =
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   410
       diff_minus [symmetric]
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   411
       add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   412
       diff_less_eq less_diff_eq diff_le_eq le_diff_eq
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   413
       diff_eq_eq eq_diff_eq
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   414
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   415
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   416
subsection{*Lemmas for the @{text cancel_numerals} simproc*}
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   417
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   418
lemma eq_iff_diff_eq_0: "(a = b) = (a-b = (0::'a::ring))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   419
by (simp add: compare_rls)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   420
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   421
lemma le_iff_diff_le_0: "(a \<le> b) = (a-b \<le> (0::'a::ordered_ring))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   422
by (simp add: compare_rls)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   423
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   424
lemma eq_add_iff1:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   425
     "(a*e + c = b*e + d) = ((a-b)*e + c = (d::'a::ring))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   426
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   427
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   428
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   429
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   430
lemma eq_add_iff2:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   431
     "(a*e + c = b*e + d) = (c = (b-a)*e + (d::'a::ring))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   432
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   433
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   434
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   435
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   436
lemma less_add_iff1:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   437
     "(a*e + c < b*e + d) = ((a-b)*e + c < (d::'a::ordered_ring))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   438
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   439
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   440
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   441
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   442
lemma less_add_iff2:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   443
     "(a*e + c < b*e + d) = (c < (b-a)*e + (d::'a::ordered_ring))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   444
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   445
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   446
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   447
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   448
lemma le_add_iff1:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   449
     "(a*e + c \<le> b*e + d) = ((a-b)*e + c \<le> (d::'a::ordered_ring))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   450
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   451
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   452
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   453
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   454
lemma le_add_iff2:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   455
     "(a*e + c \<le> b*e + d) = (c \<le> (b-a)*e + (d::'a::ordered_ring))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   456
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   457
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   458
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   459
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   460
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   461
subsection {* Ordering Rules for Multiplication *}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   462
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   463
lemma mult_strict_right_mono:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   464
     "[|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
   465
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
   466
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   467
lemma mult_left_mono:
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   468
     "[|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
   469
  apply (case_tac "c=0", simp)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   470
  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
   471
  done
14265
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 mult_right_mono:
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   474
     "[|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
   475
  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
   476
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   477
lemma mult_strict_left_mono_neg:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   478
     "[|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
   479
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
   480
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
   481
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   482
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   483
lemma mult_strict_right_mono_neg:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   484
     "[|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
   485
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
   486
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
   487
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   488
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   489
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   490
subsection{* Products of Signs *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   491
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   492
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
   493
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
   494
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   495
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
   496
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
   497
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   498
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
   499
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
   500
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   501
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
   502
apply (case_tac "b\<le>0") 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   503
 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
   504
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
   505
 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
   506
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   507
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   508
lemma zero_less_mult_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   509
     "((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
   510
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
   511
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
   512
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
   513
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
   514
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   515
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   516
text{*A field has no "zero divisors", so this theorem should hold without the
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   517
      assumption of an ordering.  See @{text field_mult_eq_0_iff} below.*}
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   518
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
   519
apply (case_tac "a < 0")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   520
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
   521
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
   522
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   523
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   524
lemma zero_le_mult_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   525
     "((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
   526
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
   527
                   zero_less_mult_iff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   528
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   529
lemma mult_less_0_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   530
     "(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
   531
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
   532
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
   533
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   534
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   535
lemma mult_le_0_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   536
     "(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
   537
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
   538
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
   539
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   540
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   541
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
   542
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
   543
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   544
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
   545
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
   546
apply (simp add: order_less_le) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   547
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   548
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   549
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
   550
  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
   551
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
subsection{*More Monotonicity*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   554
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   555
lemma mult_left_mono_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   556
     "[|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
   557
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
   558
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
   559
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   560
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   561
lemma mult_right_mono_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   562
     "[|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
   563
  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
   564
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   565
text{*Strict monotonicity in both arguments*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   566
lemma mult_strict_mono:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   567
     "[|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
   568
apply (case_tac "c=0")
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   569
 apply (simp add: mult_pos) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   570
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
   571
 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
   572
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
   573
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   574
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   575
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
   576
lemma mult_strict_mono':
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   577
     "[| 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
   578
apply (rule mult_strict_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   579
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
   580
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   581
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   582
lemma mult_mono:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   583
     "[|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
   584
      ==> 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
   585
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
   586
apply (erule mult_left_mono, assumption)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   587
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   588
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   589
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   590
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
   591
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   592
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
   593
   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
   594
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   595
lemma mult_less_cancel_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   596
    "(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
   597
apply (case_tac "c = 0")
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   598
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
   599
                      mult_strict_right_mono_neg)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   600
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
   601
                      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
   602
                      linorder_not_le [symmetric, of a])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   603
apply (erule_tac [!] notE)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   604
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
   605
                      mult_right_mono_neg)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   606
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   607
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   608
lemma mult_less_cancel_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   609
    "(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
   610
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
   611
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   612
lemma mult_le_cancel_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   613
     "(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
   614
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
   615
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   616
lemma mult_le_cancel_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   617
     "(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
   618
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
   619
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   620
lemma mult_less_imp_less_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   621
    "[|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
   622
  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
   623
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   624
lemma mult_less_imp_less_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   625
    "[|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
   626
  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
   627
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   628
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
   629
lemma mult_cancel_right [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   630
     "(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
   631
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
   632
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
   633
             simp add: linorder_neq_iff)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   634
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   635
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   636
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
   637
      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
   638
lemma mult_cancel_left [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   639
     "(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
   640
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
   641
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   642
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   643
subsection {* Fields *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   644
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   645
lemma right_inverse [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   646
      assumes not0: "a \<noteq> 0" shows "a * inverse (a::'a::field) = 1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   647
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   648
  have "a * inverse a = inverse a * a" by (simp add: mult_ac)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   649
  also have "... = 1" using not0 by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   650
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   651
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   652
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   653
lemma right_inverse_eq: "b \<noteq> 0 ==> (a / b = 1) = (a = (b::'a::field))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   654
proof
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   655
  assume neq: "b \<noteq> 0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   656
  {
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   657
    hence "a = (a / b) * b" by (simp add: divide_inverse mult_ac)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   658
    also assume "a / b = 1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   659
    finally show "a = b" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   660
  next
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   661
    assume "a = b"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   662
    with neq show "a / b = 1" by (simp add: divide_inverse)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   663
  }
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   664
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   665
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   666
lemma nonzero_inverse_eq_divide: "a \<noteq> 0 ==> inverse (a::'a::field) = 1/a"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   667
by (simp add: divide_inverse)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   668
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   669
lemma divide_self [simp]: "a \<noteq> 0 ==> a / (a::'a::field) = 1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   670
  by (simp add: divide_inverse)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   671
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   672
lemma divide_inverse_zero: "a/b = a * inverse(b::'a::{field,division_by_zero})"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   673
apply (case_tac "b = 0")
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   674
apply (simp_all add: divide_inverse)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   675
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   676
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   677
lemma divide_zero_left [simp]: "0/a = (0::'a::{field,division_by_zero})"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   678
by (simp add: divide_inverse_zero)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   679
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   680
lemma inverse_eq_divide: "inverse (a::'a::{field,division_by_zero}) = 1/a"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   681
by (simp add: divide_inverse_zero)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   682
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   683
lemma nonzero_add_divide_distrib: "c \<noteq> 0 ==> (a+b)/(c::'a::field) = a/c + b/c"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   684
by (simp add: divide_inverse left_distrib) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   685
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   686
lemma add_divide_distrib: "(a+b)/(c::'a::{field,division_by_zero}) = a/c + b/c"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   687
apply (case_tac "c=0", simp) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   688
apply (simp add: nonzero_add_divide_distrib) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   689
done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   690
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   691
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
   692
      of an ordering.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   693
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
   694
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   695
    assume "a=0" thus ?thesis by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   696
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   697
    assume anz [simp]: "a\<noteq>0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   698
    thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   699
    proof auto
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   700
      assume "a * b = 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   701
      hence "inverse a * (a * b) = 0" by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   702
      thus "b = 0"  by (simp (no_asm_use) add: mult_assoc [symmetric])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   703
    qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   704
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   705
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   706
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
   707
lemma field_mult_cancel_right_lemma:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   708
      assumes cnz: "c \<noteq> (0::'a::field)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   709
	  and eq:  "a*c = b*c"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   710
	 shows "a=b"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   711
  proof -
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   712
  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
   713
    by (simp add: eq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   714
  thus "a=b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   715
    by (simp add: mult_assoc cnz)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   716
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   717
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   718
lemma field_mult_cancel_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   719
     "(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
   720
  proof cases
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   721
    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
   722
  next
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   723
    assume "c\<noteq>0" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   724
    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
   725
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   726
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   727
lemma field_mult_cancel_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   728
     "(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
   729
  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
   730
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   731
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
   732
  proof
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   733
  assume ianz: "inverse a = 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   734
  assume "a \<noteq> 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   735
  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
   736
  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
   737
  finally have "1 = (0::'a::field)" .
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   738
  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
   739
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   740
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   741
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   742
subsection{*Basic Properties of @{term inverse}*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   743
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   744
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
   745
apply (rule ccontr) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   746
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
   747
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   748
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   749
lemma inverse_nonzero_imp_nonzero:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   750
   "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
   751
apply (rule ccontr) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   752
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
   753
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   754
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   755
lemma inverse_nonzero_iff_nonzero [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   756
   "(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
   757
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
   758
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   759
lemma nonzero_inverse_minus_eq:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   760
      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
   761
  proof -
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   762
    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
   763
      by simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   764
    thus ?thesis 
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   765
      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
   766
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   767
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   768
lemma inverse_minus_eq [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   769
     "inverse(-a) = -inverse(a::'a::{field,division_by_zero})";
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   770
  proof cases
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   771
    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
   772
  next
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   773
    assume "a\<noteq>0" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   774
    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
   775
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   776
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   777
lemma nonzero_inverse_eq_imp_eq:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   778
      assumes inveq: "inverse a = inverse b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   779
	  and anz:  "a \<noteq> 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   780
	  and bnz:  "b \<noteq> 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   781
	 shows "a = (b::'a::field)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   782
  proof -
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   783
  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
   784
    by (simp add: inveq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   785
  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
   786
    by simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   787
  thus "a = b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   788
    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
   789
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   790
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   791
lemma inverse_eq_imp_eq:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   792
     "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
   793
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
   794
 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
   795
              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
   796
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
   797
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   798
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   799
lemma inverse_eq_iff_eq [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   800
     "(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
   801
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
   802
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   803
lemma nonzero_inverse_inverse_eq:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   804
      assumes [simp]: "a \<noteq> 0"  shows "inverse(inverse (a::'a::field)) = a"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   805
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   806
  have "(inverse (inverse a) * inverse a) * a = a" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   807
    by (simp add: nonzero_imp_inverse_nonzero)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   808
  thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   809
    by (simp add: mult_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   810
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   811
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   812
lemma inverse_inverse_eq [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   813
     "inverse(inverse (a::'a::{field,division_by_zero})) = a"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   814
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   815
    assume "a=0" thus ?thesis by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   816
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   817
    assume "a\<noteq>0" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   818
    thus ?thesis by (simp add: nonzero_inverse_inverse_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   819
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   820
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   821
lemma inverse_1 [simp]: "inverse 1 = (1::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   822
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   823
  have "inverse 1 * 1 = (1::'a::field)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   824
    by (rule left_inverse [OF zero_neq_one [symmetric]])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   825
  thus ?thesis  by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   826
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   827
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   828
lemma nonzero_inverse_mult_distrib: 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   829
      assumes anz: "a \<noteq> 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   830
          and bnz: "b \<noteq> 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   831
      shows "inverse(a*b) = inverse(b) * inverse(a::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   832
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   833
  have "inverse(a*b) * (a * b) * inverse(b) = inverse(b)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   834
    by (simp add: field_mult_eq_0_iff anz bnz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   835
  hence "inverse(a*b) * a = inverse(b)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   836
    by (simp add: mult_assoc bnz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   837
  hence "inverse(a*b) * a * inverse(a) = inverse(b) * inverse(a)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   838
    by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   839
  thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   840
    by (simp add: mult_assoc anz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   841
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   842
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   843
text{*This version builds in division by zero while also re-orienting
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   844
      the right-hand side.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   845
lemma inverse_mult_distrib [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   846
     "inverse(a*b) = inverse(a) * inverse(b::'a::{field,division_by_zero})"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   847
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   848
    assume "a \<noteq> 0 & b \<noteq> 0" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   849
    thus ?thesis  by (simp add: nonzero_inverse_mult_distrib mult_commute)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   850
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   851
    assume "~ (a \<noteq> 0 & b \<noteq> 0)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   852
    thus ?thesis  by force
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   853
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   854
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   855
text{*There is no slick version using division by zero.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   856
lemma inverse_add:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   857
     "[|a \<noteq> 0;  b \<noteq> 0|]
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   858
      ==> inverse a + inverse b = (a+b) * inverse a * inverse (b::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   859
apply (simp add: left_distrib mult_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   860
apply (simp add: mult_commute [of "inverse a"]) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   861
apply (simp add: mult_assoc [symmetric] add_commute)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   862
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   863
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   864
lemma nonzero_mult_divide_cancel_left:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   865
  assumes [simp]: "b\<noteq>0" and [simp]: "c\<noteq>0" 
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   866
    shows "(c*a)/(c*b) = a/(b::'a::field)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   867
proof -
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   868
  have "(c*a)/(c*b) = c * a * (inverse b * inverse c)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   869
    by (simp add: field_mult_eq_0_iff divide_inverse 
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   870
                  nonzero_inverse_mult_distrib)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   871
  also have "... =  a * inverse b * (inverse c * c)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   872
    by (simp only: mult_ac)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   873
  also have "... =  a * inverse b"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   874
    by simp
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   875
    finally show ?thesis 
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   876
    by (simp add: divide_inverse)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   877
qed
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   878
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   879
lemma mult_divide_cancel_left:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   880
     "c\<noteq>0 ==> (c*a) / (c*b) = a / (b::'a::{field,division_by_zero})"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   881
apply (case_tac "b = 0")
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   882
apply (simp_all add: nonzero_mult_divide_cancel_left)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   883
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   884
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   885
(*For ExtractCommonTerm*)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   886
lemma mult_divide_cancel_eq_if:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   887
     "(c*a) / (c*b) = 
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   888
      (if c=0 then 0 else a / (b::'a::{field,division_by_zero}))"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   889
  by (simp add: mult_divide_cancel_left)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   890
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
   891
lemma divide_1 [simp]: "a/1 = (a::'a::field)"
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
   892
  by (simp add: divide_inverse [OF not_sym])
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
   893
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   894
lemma times_divide_eq_right [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   895
     "a * (b/c) = (a*b) / (c::'a::{field,division_by_zero})"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   896
by (simp add: divide_inverse_zero mult_assoc)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   897
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   898
lemma times_divide_eq_left [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   899
     "(b/c) * a = (b*a) / (c::'a::{field,division_by_zero})"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   900
by (simp add: divide_inverse_zero mult_ac)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   901
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   902
lemma divide_divide_eq_right [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   903
     "a / (b/c) = (a*c) / (b::'a::{field,division_by_zero})"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   904
by (simp add: divide_inverse_zero mult_ac)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   905
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   906
lemma divide_divide_eq_left [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   907
     "(a / b) / (c::'a::{field,division_by_zero}) = a / (b*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   908
by (simp add: divide_inverse_zero mult_assoc)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   909
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   910
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   911
subsection {* Division and Unary Minus *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   912
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   913
lemma nonzero_minus_divide_left: "b \<noteq> 0 ==> - (a/b) = (-a) / (b::'a::field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   914
by (simp add: divide_inverse minus_mult_left)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   915
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   916
lemma nonzero_minus_divide_right: "b \<noteq> 0 ==> - (a/b) = a / -(b::'a::field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   917
by (simp add: divide_inverse nonzero_inverse_minus_eq minus_mult_right)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   918
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   919
lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a)/(-b) = a / (b::'a::field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   920
by (simp add: divide_inverse nonzero_inverse_minus_eq)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   921
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   922
lemma minus_divide_left: "- (a/b) = (-a) / (b::'a::{field,division_by_zero})"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   923
apply (case_tac "b=0", simp) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   924
apply (simp add: nonzero_minus_divide_left) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   925
done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   926
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   927
lemma minus_divide_right: "- (a/b) = a / -(b::'a::{field,division_by_zero})"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   928
apply (case_tac "b=0", simp) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   929
by (rule nonzero_minus_divide_right) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   930
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   931
text{*The effect is to extract signs from divisions*}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   932
declare minus_divide_left  [symmetric, simp]
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   933
declare minus_divide_right [symmetric, simp]
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   934
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   935
lemma minus_divide_divide [simp]:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   936
     "(-a)/(-b) = a / (b::'a::{field,division_by_zero})"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   937
apply (case_tac "b=0", simp) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   938
apply (simp add: nonzero_minus_divide_divide) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   939
done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   940
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   941
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   942
subsection {* Ordered Fields *}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   943
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   944
lemma positive_imp_inverse_positive: 
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   945
      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
   946
  proof -
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   947
  have "0 < a * inverse a" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   948
    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
   949
  thus "0 < inverse a" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   950
    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
   951
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   952
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   953
lemma negative_imp_inverse_negative:
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   954
     "a < 0 ==> inverse a < (0::'a::ordered_field)"
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   955
  by (insert positive_imp_inverse_positive [of "-a"], 
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   956
      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
   957
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   958
lemma inverse_le_imp_le:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   959
      assumes invle: "inverse a \<le> inverse b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   960
	  and apos:  "0 < a"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   961
	 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
   962
  proof (rule classical)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   963
  assume "~ b \<le> a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   964
  hence "a < b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   965
    by (simp add: linorder_not_le)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   966
  hence bpos: "0 < b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   967
    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
   968
  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
   969
    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
   970
  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
   971
    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
   972
  thus "b \<le> a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   973
    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
   974
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   975
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   976
lemma inverse_positive_imp_positive:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   977
      assumes inv_gt_0: "0 < inverse a"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   978
          and [simp]:   "a \<noteq> 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   979
        shows "0 < (a::'a::ordered_field)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   980
  proof -
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   981
  have "0 < inverse (inverse a)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   982
    by (rule positive_imp_inverse_positive)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   983
  thus "0 < a"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   984
    by (simp add: nonzero_inverse_inverse_eq)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   985
  qed
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   986
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   987
lemma inverse_positive_iff_positive [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   988
      "(0 < inverse a) = (0 < (a::'a::{ordered_field,division_by_zero}))"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   989
apply (case_tac "a = 0", simp)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   990
apply (blast intro: inverse_positive_imp_positive positive_imp_inverse_positive)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   991
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   992
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   993
lemma inverse_negative_imp_negative:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   994
      assumes inv_less_0: "inverse a < 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   995
          and [simp]:   "a \<noteq> 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   996
        shows "a < (0::'a::ordered_field)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   997
  proof -
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   998
  have "inverse (inverse a) < 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   999
    by (rule negative_imp_inverse_negative)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1000
  thus "a < 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1001
    by (simp add: nonzero_inverse_inverse_eq)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1002
  qed
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1003
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1004
lemma inverse_negative_iff_negative [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1005
      "(inverse a < 0) = (a < (0::'a::{ordered_field,division_by_zero}))"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1006
apply (case_tac "a = 0", simp)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1007
apply (blast intro: inverse_negative_imp_negative negative_imp_inverse_negative)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1008
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1009
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1010
lemma inverse_nonnegative_iff_nonnegative [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1011
      "(0 \<le> inverse a) = (0 \<le> (a::'a::{ordered_field,division_by_zero}))"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1012
by (simp add: linorder_not_less [symmetric])
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1013
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1014
lemma inverse_nonpositive_iff_nonpositive [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1015
      "(inverse a \<le> 0) = (a \<le> (0::'a::{ordered_field,division_by_zero}))"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1016
by (simp add: linorder_not_less [symmetric])
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1017
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1018
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1019
subsection{*Anti-Monotonicity of @{term inverse}*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1020
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1021
lemma less_imp_inverse_less:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  1022
      assumes less: "a < b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  1023
	  and apos:  "0 < a"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  1024
	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
  1025
  proof (rule ccontr)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1026
  assume "~ inverse b < inverse a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1027
  hence "inverse a \<le> inverse b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1028
    by (simp add: linorder_not_less)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1029
  hence "~ (a < b)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1030
    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
  1031
  thus False
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1032
    by (rule notE [OF _ less])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1033
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1034
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1035
lemma inverse_less_imp_less:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1036
   "[|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
  1037
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
  1038
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
  1039
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1040
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1041
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
  1042
lemma inverse_less_iff_less [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1043
     "[|0 < a; 0 < b|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1044
      ==> (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
  1045
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
  1046
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1047
lemma le_imp_inverse_le:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1048
   "[|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
  1049
  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
  1050
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1051
lemma inverse_le_iff_le [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1052
     "[|0 < a; 0 < b|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1053
      ==> (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
  1054
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
  1055
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1056
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1057
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
  1058
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
  1059
lemma inverse_le_imp_le_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1060
   "[|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
  1061
  apply (rule classical) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1062
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1063
   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
  1064
  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
  1065
  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
  1066
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1067
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1068
lemma less_imp_inverse_less_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1069
   "[|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
  1070
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1071
   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
  1072
  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
  1073
  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
  1074
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1075
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1076
lemma inverse_less_imp_less_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1077
   "[|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
  1078
  apply (rule classical) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1079
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1080
   prefer 2
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1081
   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
  1082
  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
  1083
  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
  1084
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1085
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1086
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
  1087
     "[|a < 0; b < 0|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1088
      ==> (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
  1089
  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
  1090
  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
  1091
	      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
  1092
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1093
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1094
lemma le_imp_inverse_le_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1095
   "[|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
  1096
  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
  1097
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1098
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
  1099
     "[|a < 0; b < 0|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1100
      ==> (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
  1101
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
  1102
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1103
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1104
subsection{*Division and Signs*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1105
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1106
lemma zero_less_divide_iff:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1107
     "((0::'a::{ordered_field,division_by_zero}) < a/b) = (0 < a & 0 < b | a < 0 & b < 0)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1108
by (simp add: divide_inverse_zero zero_less_mult_iff)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1109
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1110
lemma divide_less_0_iff:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1111
     "(a/b < (0::'a::{ordered_field,division_by_zero})) = 
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1112
      (0 < a & b < 0 | a < 0 & 0 < b)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1113
by (simp add: divide_inverse_zero mult_less_0_iff)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1114
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1115
lemma zero_le_divide_iff:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1116
     "((0::'a::{ordered_field,division_by_zero}) \<le> a/b) =
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1117
      (0 \<le> a & 0 \<le> b | a \<le> 0 & b \<le> 0)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1118
by (simp add: divide_inverse_zero zero_le_mult_iff)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1119
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1120
lemma divide_le_0_iff:
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1121
     "(a/b \<le> (0::'a::{ordered_field,division_by_zero})) =
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1122
      (0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)"
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1123
by (simp add: divide_inverse_zero mult_le_0_iff)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1124
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1125
lemma divide_eq_0_iff [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1126
     "(a/b = 0) = (a=0 | b=(0::'a::{field,division_by_zero}))"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1127
by (simp add: divide_inverse_zero field_mult_eq_0_iff)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1128
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1129
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1130
subsection{*Simplification of Inequalities Involving Literal Divisors*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1131
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1132
lemma pos_le_divide_eq: "0 < (c::'a::ordered_field) ==> (a \<le> b/c) = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1133
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1134
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1135
  hence "(a \<le> b/c) = (a*c \<le> (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1136
    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1137
  also have "... = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1138
    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1139
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1140
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1141
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1142
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1143
lemma neg_le_divide_eq: "c < (0::'a::ordered_field) ==> (a \<le> b/c) = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1144
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1145
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1146
  hence "(a \<le> b/c) = ((b/c)*c \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1147
    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1148
  also have "... = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1149
    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1150
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1151
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1152
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1153
lemma le_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1154
  "(a \<le> b/c) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1155
   (if 0 < c then a*c \<le> b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1156
             else if c < 0 then b \<le> a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1157
             else  a \<le> (0::'a::{ordered_field,division_by_zero}))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1158
apply (case_tac "c=0", simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1159
apply (force simp add: pos_le_divide_eq neg_le_divide_eq linorder_neq_iff) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1160
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1161
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1162
lemma pos_divide_le_eq: "0 < (c::'a::ordered_field) ==> (b/c \<le> a) = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1163
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1164
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1165
  hence "(b/c \<le> a) = ((b/c)*c \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1166
    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1167
  also have "... = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1168
    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1169
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1170
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1171
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1172
lemma neg_divide_le_eq: "c < (0::'a::ordered_field) ==> (b/c \<le> a) = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1173
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1174
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1175
  hence "(b/c \<le> a) = (a*c \<le> (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1176
    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1177
  also have "... = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1178
    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1179
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1180
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1181
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1182
lemma divide_le_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1183
  "(b/c \<le> a) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1184
   (if 0 < c then b \<le> a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1185
             else if c < 0 then a*c \<le> b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1186
             else 0 \<le> (a::'a::{ordered_field,division_by_zero}))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1187
apply (case_tac "c=0", simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1188
apply (force simp add: pos_divide_le_eq neg_divide_le_eq linorder_neq_iff) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1189
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1190
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1191
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1192
lemma pos_less_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1193
     "0 < (c::'a::ordered_field) ==> (a < b/c) = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1194
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1195
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1196
  hence "(a < b/c) = (a*c < (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1197
    by (simp add: mult_less_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1198
  also have "... = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1199
    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1200
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1201
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1202
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1203
lemma neg_less_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1204
 "c < (0::'a::ordered_field) ==> (a < b/c) = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1205
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1206
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1207
  hence "(a < b/c) = ((b/c)*c < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1208
    by (simp add: mult_less_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1209
  also have "... = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1210
    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1211
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1212
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1213
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1214
lemma less_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1215
  "(a < b/c) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1216
   (if 0 < c then a*c < b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1217
             else if c < 0 then b < a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1218
             else  a < (0::'a::{ordered_field,division_by_zero}))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1219
apply (case_tac "c=0", simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1220
apply (force simp add: pos_less_divide_eq neg_less_divide_eq linorder_neq_iff) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1221
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1222
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1223
lemma pos_divide_less_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1224
     "0 < (c::'a::ordered_field) ==> (b/c < a) = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1225
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1226
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1227
  hence "(b/c < a) = ((b/c)*c < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1228
    by (simp add: mult_less_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1229
  also have "... = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1230
    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1231
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1232
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1233
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1234
lemma neg_divide_less_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1235
 "c < (0::'a::ordered_field) ==> (b/c < a) = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1236
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1237
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1238
  hence "(b/c < a) = (a*c < (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1239
    by (simp add: mult_less_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1240
  also have "... = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1241
    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1242
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1243
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1244
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1245
lemma divide_less_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1246
  "(b/c < a) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1247
   (if 0 < c then b < a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1248
             else if c < 0 then a*c < b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1249
             else 0 < (a::'a::{ordered_field,division_by_zero}))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1250
apply (case_tac "c=0", simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1251
apply (force simp add: pos_divide_less_eq neg_divide_less_eq linorder_neq_iff) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1252
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1253
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1254
lemma nonzero_eq_divide_eq: "c\<noteq>0 ==> ((a::'a::field) = b/c) = (a*c = b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1255
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1256
  assume [simp]: "c\<noteq>0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1257
  have "(a = b/c) = (a*c = (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1258
    by (simp add: field_mult_cancel_right)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1259
  also have "... = (a*c = b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1260
    by (simp add: divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1261
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1262
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1263
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1264
lemma eq_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1265
  "((a::'a::{field,division_by_zero}) = b/c) = (if c\<noteq>0 then a*c = b else a=0)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1266
by (simp add: nonzero_eq_divide_eq) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1267
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1268
lemma nonzero_divide_eq_eq: "c\<noteq>0 ==> (b/c = (a::'a::field)) = (b = a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1269
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1270
  assume [simp]: "c\<noteq>0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1271
  have "(b/c = a) = ((b/c)*c = a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1272
    by (simp add: field_mult_cancel_right)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1273
  also have "... = (b = a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1274
    by (simp add: divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1275
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1276
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1277
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1278
lemma divide_eq_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1279
  "(b/c = (a::'a::{field,division_by_zero})) = (if c\<noteq>0 then b = a*c else a=0)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1280
by (force simp add: nonzero_divide_eq_eq) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1281
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1282
subsection{*Cancellation Laws for Division*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1283
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1284
lemma divide_cancel_right [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1285
     "(a/c = b/c) = (c = 0 | a = (b::'a::{field,division_by_zero}))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1286
apply (case_tac "c=0", simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1287
apply (simp add: divide_inverse_zero field_mult_cancel_right) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1288
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1289
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1290
lemma divide_cancel_left [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1291
     "(c/a = c/b) = (c = 0 | a = (b::'a::{field,division_by_zero}))" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1292
apply (case_tac "c=0", simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1293
apply (simp add: divide_inverse_zero field_mult_cancel_left) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1294
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1295
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1296
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1297
subsection {* Ordering Rules for Division *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1298
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1299
lemma divide_strict_right_mono:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1300
     "[|a < b; 0 < c|] ==> a / c < b / (c::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1301
by (simp add: order_less_imp_not_eq2 divide_inverse mult_strict_right_mono 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1302
              positive_imp_inverse_positive) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1303
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1304
lemma divide_right_mono:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1305
     "[|a \<le> b; 0 \<le> c|] ==> a/c \<le> b/(c::'a::{ordered_field,division_by_zero})"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1306
  by (force simp add: divide_strict_right_mono order_le_less) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1307
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1308
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1309
text{*The last premise ensures that @{term a} and @{term b} 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1310
      have the same sign*}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1311
lemma divide_strict_left_mono:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1312
       "[|b < a; 0 < c; 0 < a*b|] ==> c / a < c / (b::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1313
by (force simp add: zero_less_mult_iff divide_inverse mult_strict_left_mono 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1314
      order_less_imp_not_eq order_less_imp_not_eq2  
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1315
      less_imp_inverse_less less_imp_inverse_less_neg) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1316
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1317
lemma divide_left_mono:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1318
     "[|b \<le> a; 0 \<le> c; 0 < a*b|] ==> c / a \<le> c / (b::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1319
  apply (subgoal_tac "a \<noteq> 0 & b \<noteq> 0") 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1320
   prefer 2 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1321
   apply (force simp add: zero_less_mult_iff order_less_imp_not_eq) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1322
  apply (case_tac "c=0", simp add: divide_inverse)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1323
  apply (force simp add: divide_strict_left_mono order_le_less) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1324
  done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1325
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1326
lemma divide_strict_left_mono_neg:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1327
     "[|a < b; c < 0; 0 < a*b|] ==> c / a < c / (b::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1328
  apply (subgoal_tac "a \<noteq> 0 & b \<noteq> 0") 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1329
   prefer 2 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1330
   apply (force simp add: zero_less_mult_iff order_less_imp_not_eq) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1331
  apply (drule divide_strict_left_mono [of _ _ "-c"]) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1332
   apply (simp_all add: mult_commute nonzero_minus_divide_left [symmetric]) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1333
  done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1334
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1335
lemma divide_strict_right_mono_neg:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1336
     "[|b < a; c < 0|] ==> a / c < b / (c::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1337
apply (drule divide_strict_right_mono [of _ _ "-c"], simp) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1338
apply (simp add: order_less_imp_not_eq nonzero_minus_divide_right [symmetric]) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1339
done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1340
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1341
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1342
subsection {* Ordered Fields are Dense *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1343
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1344
lemma zero_less_two: "0 < (1+1::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1345
proof -
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1346
  have "0 + 0 <  (1+1::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1347
    by (blast intro: zero_less_one add_strict_mono) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1348
  thus ?thesis by simp
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1349
qed
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1350
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1351
lemma less_half_sum: "a < b ==> a < (a+b) / (1+1::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1352
by (simp add: zero_less_two pos_less_divide_eq right_distrib) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1353
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1354
lemma gt_half_sum: "a < b ==> (a+b)/(1+1::'a::ordered_field) < b"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1355
by (simp add: zero_less_two pos_divide_less_eq right_distrib) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1356
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1357
lemma dense: "a < b ==> \<exists>r::'a::ordered_field. a < r & r < b"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1358
by (blast intro!: less_half_sum gt_half_sum)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1359
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1360
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1361
subsection {* Absolute Value *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1362
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1363
lemma abs_zero [simp]: "abs 0 = (0::'a::ordered_ring)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1364
by (simp add: abs_if)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1365
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1366
lemma abs_one [simp]: "abs 1 = (1::'a::ordered_ring)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1367
  by (simp add: abs_if zero_less_one [THEN order_less_not_sym]) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1368
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1369
lemma abs_mult: "abs (a * b) = abs a * abs (b::'a::ordered_ring)" 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1370
apply (case_tac "a=0 | b=0", force) 
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1371
apply (auto elim: order_less_asym
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1372
            simp add: abs_if mult_less_0_iff linorder_neq_iff
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1373
                  minus_mult_left [symmetric] minus_mult_right [symmetric])  
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1374
done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1375
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1376
lemma abs_eq_0 [simp]: "(abs a = 0) = (a = (0::'a::ordered_ring))"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1377
by (simp add: abs_if)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1378
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1379
lemma zero_less_abs_iff [simp]: "(0 < abs a) = (a \<noteq> (0::'a::ordered_ring))"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1380
by (simp add: abs_if linorder_neq_iff)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1381
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1382
lemma abs_not_less_zero [simp]: "~ abs a < (0::'a::ordered_ring)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1383
by (simp add: abs_if  order_less_not_sym [of a 0])
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1384
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1385
lemma abs_le_zero_iff [simp]: "(abs a \<le> (0::'a::ordered_ring)) = (a = 0)" 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1386
by (simp add: order_le_less) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1387
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1388
lemma abs_minus_cancel [simp]: "abs (-a) = abs(a::'a::ordered_ring)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1389
apply (auto simp add: abs_if linorder_not_less order_less_not_sym [of 0 a])  
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1390
apply (drule order_antisym, assumption, simp) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1391
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1392
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1393
lemma abs_ge_zero [simp]: "(0::'a::ordered_ring) \<le> abs a"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1394
apply (simp add: abs_if order_less_imp_le)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1395
apply (simp add: linorder_not_less) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1396
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1397
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1398
lemma abs_idempotent [simp]: "abs (abs a) = abs (a::'a::ordered_ring)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1399
  by (force elim: order_less_asym simp add: abs_if)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1400
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14295
diff changeset
  1401
lemma abs_zero_iff [simp]: "(abs a = 0) = (a = (0::'a::ordered_ring))"
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1402
by (simp add: abs_if)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1403
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1404
lemma abs_ge_self: "a \<le> abs (a::'a::ordered_ring)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1405
apply (simp add: abs_if)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1406
apply (simp add: order_less_imp_le order_trans [of _ 0])
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1407
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1408
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1409
lemma abs_ge_minus_self: "-a \<le> abs (a::'a::ordered_ring)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1410
by (insert abs_ge_self [of "-a"], simp)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1411
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1412
lemma nonzero_abs_inverse:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1413
     "a \<noteq> 0 ==> abs (inverse (a::'a::ordered_field)) = inverse (abs a)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1414
apply (auto simp add: linorder_neq_iff abs_if nonzero_inverse_minus_eq 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1415
                      negative_imp_inverse_negative)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1416
apply (blast intro: positive_imp_inverse_positive elim: order_less_asym) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1417
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1418
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1419
lemma abs_inverse [simp]:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1420
     "abs (inverse (a::'a::{ordered_field,division_by_zero})) = 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1421
      inverse (abs a)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1422
apply (case_tac "a=0", simp) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1423
apply (simp add: nonzero_abs_inverse) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1424
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1425
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1426
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1427
lemma nonzero_abs_divide:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1428
     "b \<noteq> 0 ==> abs (a / (b::'a::ordered_field)) = abs a / abs b"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1429
by (simp add: divide_inverse abs_mult nonzero_abs_inverse) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1430
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1431
lemma abs_divide:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1432
     "abs (a / (b::'a::{ordered_field,division_by_zero})) = abs a / abs b"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1433
apply (case_tac "b=0", simp) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1434
apply (simp add: nonzero_abs_divide) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1435
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1436
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1437
lemma abs_leI: "[|a \<le> b; -a \<le> b|] ==> abs a \<le> (b::'a::ordered_ring)"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1438
by (simp add: abs_if)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1439
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1440
lemma le_minus_self_iff: "(a \<le> -a) = (a \<le> (0::'a::ordered_ring))"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1441
proof 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1442
  assume ale: "a \<le> -a"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1443
  show "a\<le>0"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1444
    apply (rule classical) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1445
    apply (simp add: linorder_not_le) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1446
    apply (blast intro: ale order_trans order_less_imp_le
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1447
                        neg_0_le_iff_le [THEN iffD1]) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1448
    done
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1449
next
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1450
  assume "a\<le>0"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1451
  hence "0 \<le> -a" by (simp only: neg_0_le_iff_le)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1452
  thus "a \<le> -a"  by (blast intro: prems order_trans) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1453
qed
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1454
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1455
lemma minus_le_self_iff: "(-a \<le> a) = (0 \<le> (a::'a::ordered_ring))"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1456
by (insert le_minus_self_iff [of "-a"], simp)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1457
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1458
lemma eq_minus_self_iff: "(a = -a) = (a = (0::'a::ordered_ring))"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1459
by (force simp add: order_eq_iff le_minus_self_iff minus_le_self_iff)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1460
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1461
lemma less_minus_self_iff: "(a < -a) = (a < (0::'a::ordered_ring))"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1462
by (simp add: order_less_le le_minus_self_iff eq_minus_self_iff)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1463
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1464
lemma abs_le_D1: "abs a \<le> b ==> a \<le> (b::'a::ordered_ring)"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1465
apply (simp add: abs_if split: split_if_asm)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1466
apply (rule order_trans [of _ "-a"]) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1467
 apply (simp add: less_minus_self_iff order_less_imp_le, assumption)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1468
done
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1469
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1470
lemma abs_le_D2: "abs a \<le> b ==> -a \<le> (b::'a::ordered_ring)"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1471
by (insert abs_le_D1 [of "-a"], simp)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1472
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1473
lemma abs_le_iff: "(abs a \<le> b) = (a \<le> b & -a \<le> (b::'a::ordered_ring))"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1474
by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1475
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1476
lemma abs_less_iff: "(abs a < b) = (a < b & -a < (b::'a::ordered_ring))" 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1477
apply (simp add: order_less_le abs_le_iff)  
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1478
apply (auto simp add: abs_if minus_le_self_iff eq_minus_self_iff) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1479
 apply (simp add:  linorder_not_less [symmetric]) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1480
apply (simp add: le_minus_self_iff linorder_neq_iff) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1481
apply (simp add:  linorder_not_less [symmetric]) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1482
done
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1483
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1484
lemma abs_triangle_ineq: "abs (a+b) \<le> abs a + abs (b::'a::ordered_ring)"
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1485
by (force simp add: abs_le_iff abs_ge_self abs_ge_minus_self add_mono)
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1486
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1487
lemma abs_mult_less:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1488
     "[| abs a < c; abs b < d |] ==> abs a * abs b < c*(d::'a::ordered_ring)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1489
proof -
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1490
  assume ac: "abs a < c"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1491
  hence cpos: "0<c" by (blast intro: order_le_less_trans abs_ge_zero)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1492
  assume "abs b < d"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1493
  thus ?thesis by (simp add: ac cpos mult_strict_mono) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1494
qed
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1495
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1496
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1497
end