src/HOL/Ring_and_Field.thy
author paulson
Fri, 05 Dec 2003 18:10:59 +0100
changeset 14277 ad66687ece6e
parent 14272 5efbb548107d
child 14284 f1abe67c448a
permissions -rw-r--r--
more field division lemmas transferred from Real to Ring_and_Field
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     1
(*  Title:   HOL/Ring_and_Field.thy
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     2
    ID:      $Id$
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     3
    Author:  Gertrud Bauer and Markus Wenzel, TU Muenchen
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
     4
             Lawrence C Paulson, University of Cambridge
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     5
    License: GPL (GNU GENERAL PUBLIC LICENSE)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     6
*)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     7
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     8
header {*
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     9
  \title{Ring and field structures}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    10
  \author{Gertrud Bauer and Markus Wenzel}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    11
*}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    12
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    13
theory Ring_and_Field = Inductive:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    14
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    15
text{*Lemmas and extension to semirings by L. C. Paulson*}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    16
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    17
subsection {* Abstract algebraic structures *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    18
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    19
axclass semiring \<subseteq> zero, one, plus, times
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    20
  add_assoc: "(a + b) + c = a + (b + c)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    21
  add_commute: "a + b = b + a"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    22
  left_zero [simp]: "0 + a = a"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    23
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    24
  mult_assoc: "(a * b) * c = a * (b * c)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    25
  mult_commute: "a * b = b * a"
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
    26
  mult_1 [simp]: "1 * a = a"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    27
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    28
  left_distrib: "(a + b) * c = a * c + b * c"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    29
  zero_neq_one [simp]: "0 \<noteq> 1"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    30
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    31
axclass ring \<subseteq> semiring, minus
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    32
  left_minus [simp]: "- a + a = 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    33
  diff_minus: "a - b = a + (-b)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    34
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    35
axclass ordered_semiring \<subseteq> semiring, linorder
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    36
  add_left_mono: "a \<le> b ==> c + a \<le> c + b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    37
  mult_strict_left_mono: "a < b ==> 0 < c ==> c * a < c * b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    38
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    39
axclass ordered_ring \<subseteq> ordered_semiring, ring
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    40
  abs_if: "\<bar>a\<bar> = (if a < 0 then -a else a)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    41
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    42
axclass field \<subseteq> ring, inverse
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    43
  left_inverse [simp]: "a \<noteq> 0 ==> inverse a * a = 1"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    44
  divide_inverse:      "b \<noteq> 0 ==> a / b = a * inverse b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    45
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    46
axclass ordered_field \<subseteq> ordered_ring, field
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    47
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    48
axclass division_by_zero \<subseteq> zero, inverse
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
    49
  inverse_zero [simp]: "inverse 0 = 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
    50
  divide_zero [simp]: "a / 0 = 0"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    51
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    52
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    53
subsection {* Derived Rules for Addition *}
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    54
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    55
lemma right_zero [simp]: "a + 0 = (a::'a::semiring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    56
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    57
  have "a + 0 = 0 + a" by (simp only: add_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    58
  also have "... = a" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    59
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    60
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    61
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    62
lemma add_left_commute: "a + (b + c) = b + (a + (c::'a::semiring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    63
  by (rule mk_left_commute [of "op +", OF add_assoc add_commute])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    64
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    65
theorems add_ac = add_assoc add_commute add_left_commute
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    66
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    67
lemma right_minus [simp]: "a + -(a::'a::ring) = 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    68
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    69
  have "a + -a = -a + a" by (simp add: add_ac)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    70
  also have "... = 0" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    71
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    72
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    73
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    74
lemma right_minus_eq: "(a - b = 0) = (a = (b::'a::ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    75
proof
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    76
  have "a = a - b + b" by (simp add: diff_minus add_ac)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    77
  also assume "a - b = 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    78
  finally show "a = b" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    79
next
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    80
  assume "a = b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    81
  thus "a - b = 0" by (simp add: diff_minus)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    82
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    83
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    84
lemma add_left_cancel [simp]:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    85
     "(a + b = a + c) = (b = (c::'a::ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    86
proof
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    87
  assume eq: "a + b = a + c"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
    88
  hence "(-a + a) + b = (-a + a) + c"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    89
    by (simp only: eq add_assoc)
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
    90
  thus "b = c" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    91
next
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    92
  assume eq: "b = c"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
    93
  thus "a + b = a + c" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    94
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    95
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    96
lemma add_right_cancel [simp]:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    97
     "(b + a = c + a) = (b = (c::'a::ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    98
  by (simp add: add_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    99
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   100
lemma minus_minus [simp]: "- (- (a::'a::ring)) = a"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   101
  proof (rule add_left_cancel [of "-a", THEN iffD1])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   102
    show "(-a + -(-a) = -a + a)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   103
    by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   104
  qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   105
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   106
lemma equals_zero_I: "a+b = 0 ==> -a = (b::'a::ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   107
apply (rule right_minus_eq [THEN iffD1, symmetric])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   108
apply (simp add: diff_minus add_commute) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   109
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   110
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   111
lemma minus_zero [simp]: "- 0 = (0::'a::ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   112
by (simp add: equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   113
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   114
lemma diff_self [simp]: "a - (a::'a::ring) = 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   115
  by (simp add: diff_minus)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   116
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   117
lemma diff_0 [simp]: "(0::'a::ring) - a = -a"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   118
by (simp add: diff_minus)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   119
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   120
lemma diff_0_right [simp]: "a - (0::'a::ring) = a" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   121
by (simp add: diff_minus)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   122
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   123
lemma neg_equal_iff_equal [simp]: "(-a = -b) = (a = (b::'a::ring))" 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   124
  proof 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   125
    assume "- a = - b"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   126
    hence "- (- a) = - (- b)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   127
      by simp
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   128
    thus "a=b" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   129
  next
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   130
    assume "a=b"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   131
    thus "-a = -b" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   132
  qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   133
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   134
lemma neg_equal_0_iff_equal [simp]: "(-a = 0) = (a = (0::'a::ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   135
by (subst neg_equal_iff_equal [symmetric], simp)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   136
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   137
lemma neg_0_equal_iff_equal [simp]: "(0 = -a) = (0 = (a::'a::ring))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   138
by (subst neg_equal_iff_equal [symmetric], simp)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   139
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   140
text{*The next two equations can make the simplifier loop!*}
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   141
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   142
lemma equation_minus_iff: "(a = - b) = (b = - (a::'a::ring))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   143
  proof -
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   144
  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
   145
  thus ?thesis by (simp add: eq_commute)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   146
  qed
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   147
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   148
lemma minus_equation_iff: "(- a = b) = (- (b::'a::ring) = a)"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   149
  proof -
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   150
  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
   151
  thus ?thesis by (simp add: eq_commute)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   152
  qed
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   153
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   154
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   155
subsection {* Derived rules for multiplication *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   156
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   157
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
   158
proof -
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   159
  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
   160
  also have "... = a" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   161
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   162
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   163
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   164
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
   165
  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
   166
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   167
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
   168
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   169
lemma right_inverse [simp]:
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   170
      assumes not0: "a \<noteq> 0" shows "a * inverse (a::'a::field) = 1"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   171
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   172
  have "a * inverse a = inverse a * a" by (simp add: mult_ac)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   173
  also have "... = 1" using not0 by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   174
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   175
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   176
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   177
lemma right_inverse_eq: "b \<noteq> 0 ==> (a / b = 1) = (a = (b::'a::field))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   178
proof
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   179
  assume neq: "b \<noteq> 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   180
  {
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   181
    hence "a = (a / b) * b" by (simp add: divide_inverse mult_ac)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   182
    also assume "a / b = 1"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   183
    finally show "a = b" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   184
  next
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   185
    assume "a = b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   186
    with neq show "a / b = 1" by (simp add: divide_inverse)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   187
  }
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   188
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   189
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   190
lemma nonzero_inverse_eq_divide: "a \<noteq> 0 ==> inverse (a::'a::field) = 1/a"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   191
by (simp add: divide_inverse)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   192
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   193
lemma divide_self [simp]: "a \<noteq> 0 ==> a / (a::'a::field) = 1"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   194
  by (simp add: divide_inverse)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   195
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   196
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
   197
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   198
  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
   199
    by (simp add: left_distrib [symmetric])
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   200
  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
   201
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   202
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   203
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
   204
  by (simp add: mult_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   205
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   206
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   207
subsection {* Distribution rules *}
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 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
   210
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   211
  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
   212
  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
   213
  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
   214
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   215
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   216
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   217
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
   218
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   219
text{*For the @{text combine_numerals} simproc*}
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   220
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
   221
by (simp add: left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   222
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   223
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
   224
apply (rule equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   225
apply (simp add: add_ac) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   226
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   227
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   228
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
   229
apply (rule equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   230
apply (simp add: left_distrib [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   231
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   232
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   233
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
   234
apply (rule equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   235
apply (simp add: right_distrib [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   236
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   237
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   238
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
   239
  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
   240
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   241
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
   242
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
   243
              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
   244
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   245
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
   246
by (simp add: mult_commute [of _ c] right_diff_distrib) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   247
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   248
lemma minus_diff_eq [simp]: "- (a - b) = b - (a::'a::ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   249
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
   250
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   251
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   252
subsection {* Ordering Rules for Addition *}
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   253
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   254
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
   255
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
   256
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   257
text {* non-strict, in both arguments *}
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   258
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
   259
  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
   260
  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
   261
  done
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   262
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   263
lemma add_strict_left_mono:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   264
     "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
   265
 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
   266
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   267
lemma add_strict_right_mono:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   268
     "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
   269
 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
   270
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   271
text{*Strict monotonicity in both arguments*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   272
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
   273
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
   274
apply (erule add_strict_left_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   275
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   276
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   277
lemma add_less_imp_less_left:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   278
      assumes less: "c + a < c + b"  shows "a < (b::'a::ordered_ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   279
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   280
  have "-c + (c + a) < -c + (c + b)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   281
    by (rule add_strict_left_mono [OF less])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   282
  thus "a < b" by (simp add: add_assoc [symmetric])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   283
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   284
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   285
lemma add_less_imp_less_right:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   286
      "a + c < b + c ==> a < (b::'a::ordered_ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   287
apply (rule add_less_imp_less_left [of c])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   288
apply (simp add: add_commute)  
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   289
done
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
lemma add_less_cancel_left [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   292
    "(c+a < c+b) = (a < (b::'a::ordered_ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   293
by (blast intro: add_less_imp_less_left add_strict_left_mono) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   294
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   295
lemma add_less_cancel_right [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   296
    "(a+c < b+c) = (a < (b::'a::ordered_ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   297
by (blast intro: add_less_imp_less_right add_strict_right_mono)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   298
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   299
lemma add_le_cancel_left [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   300
    "(c+a \<le> c+b) = (a \<le> (b::'a::ordered_ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   301
by (simp add: linorder_not_less [symmetric]) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   302
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   303
lemma add_le_cancel_right [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   304
    "(a+c \<le> b+c) = (a \<le> (b::'a::ordered_ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   305
by (simp add: linorder_not_less [symmetric]) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   306
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   307
lemma add_le_imp_le_left:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   308
      "c + a \<le> c + b ==> a \<le> (b::'a::ordered_ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   309
by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   310
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   311
lemma add_le_imp_le_right:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   312
      "a + c \<le> b + c ==> a \<le> (b::'a::ordered_ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   313
by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   314
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   315
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   316
subsection {* Ordering Rules for Unary Minus *}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   317
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   318
lemma le_imp_neg_le:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   319
      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
   320
  proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   321
  have "-a+a \<le> -a+b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   322
    by (rule add_left_mono) 
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   323
  hence "0 \<le> -a+b"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   324
    by simp
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   325
  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
   326
    by (rule add_right_mono) 
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   327
  thus ?thesis
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   328
    by (simp add: add_assoc)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   329
  qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   330
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   331
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
   332
  proof 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   333
    assume "- b \<le> - a"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   334
    hence "- (- a) \<le> - (- b)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   335
      by (rule le_imp_neg_le)
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   336
    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
   337
  next
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   338
    assume "a\<le>b"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   339
    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
   340
  qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   341
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   342
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
   343
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
   344
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   345
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
   346
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
   347
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   348
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
   349
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
   350
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   351
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
   352
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
   353
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   354
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
   355
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
   356
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   357
text{*The next several equations can make the simplifier loop!*}
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   358
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   359
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
   360
  proof -
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   361
  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
   362
  thus ?thesis by simp
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   363
  qed
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   364
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   365
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
   366
  proof -
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   367
  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
   368
  thus ?thesis by simp
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   369
  qed
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   370
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   371
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
   372
apply (simp add: linorder_not_less [symmetric])
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   373
apply (rule minus_less_iff) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   374
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   375
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   376
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
   377
apply (simp add: linorder_not_less [symmetric])
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   378
apply (rule less_minus_iff) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   379
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   380
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   381
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   382
subsection{*Subtraction Laws*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   383
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   384
lemma add_diff_eq: "a + (b - c) = (a + b) - (c::'a::ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   385
by (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   386
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   387
lemma diff_add_eq: "(a - b) + c = (a + c) - (b::'a::ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   388
by (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   389
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   390
lemma diff_eq_eq: "(a-b = c) = (a = c + (b::'a::ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   391
by (auto simp add: diff_minus add_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   392
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   393
lemma eq_diff_eq: "(a = c-b) = (a + (b::'a::ring) = c)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   394
by (auto simp add: diff_minus add_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   395
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   396
lemma diff_diff_eq: "(a - b) - c = a - (b + (c::'a::ring))"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   397
by (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   398
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   399
lemma diff_diff_eq2: "a - (b - c) = (a + c) - (b::'a::ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   400
by (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   401
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   402
text{*Further subtraction laws for ordered rings*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   403
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   404
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
   405
proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   406
  have  "(a < b) = (a + (- b) < b + (-b))"  
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   407
    by (simp only: add_less_cancel_right)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   408
  also have "... =  (a - b < 0)" by (simp add: diff_minus)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   409
  finally show ?thesis .
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   410
qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   411
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   412
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
   413
apply (subst less_iff_diff_less_0)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   414
apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   415
apply (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   416
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   417
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   418
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
   419
apply (subst less_iff_diff_less_0)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   420
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
   421
apply (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   422
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   423
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   424
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
   425
by (simp add: linorder_not_less [symmetric] less_diff_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   426
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   427
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
   428
by (simp add: linorder_not_less [symmetric] diff_less_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   429
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   430
text{*This list of rewrites simplifies (in)equalities by bringing subtractions
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   431
  to the top and then moving negative terms to the other side.
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   432
  Use with @{text add_ac}*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   433
lemmas compare_rls =
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   434
       diff_minus [symmetric]
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   435
       add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   436
       diff_less_eq less_diff_eq diff_le_eq le_diff_eq
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   437
       diff_eq_eq eq_diff_eq
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   438
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   439
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   440
subsection{*Lemmas for the @{text cancel_numerals} simproc*}
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 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
   443
by (simp add: compare_rls)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   444
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   445
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
   446
by (simp add: compare_rls)
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 eq_add_iff1:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   449
     "(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
   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 eq_add_iff2:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   455
     "(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
   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
lemma less_add_iff1:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   461
     "(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
   462
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   463
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   464
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   465
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   466
lemma less_add_iff2:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   467
     "(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
   468
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   469
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   470
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   471
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   472
lemma le_add_iff1:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   473
     "(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
   474
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   475
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   476
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   477
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   478
lemma le_add_iff2:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   479
     "(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
   480
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   481
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   482
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   483
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   484
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   485
subsection {* Ordering Rules for Multiplication *}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   486
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   487
lemma mult_strict_right_mono:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   488
     "[|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
   489
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
   490
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   491
lemma mult_left_mono:
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   492
     "[|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
   493
  apply (case_tac "c=0", simp)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   494
  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
   495
  done
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   496
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   497
lemma mult_right_mono:
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   498
     "[|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
   499
  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
   500
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   501
lemma mult_strict_left_mono_neg:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   502
     "[|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
   503
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
   504
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
   505
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   506
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   507
lemma mult_strict_right_mono_neg:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   508
     "[|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
   509
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
   510
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
   511
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   512
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   513
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   514
subsection{* Products of Signs *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   515
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   516
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
   517
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
   518
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   519
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
   520
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
   521
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   522
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
   523
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
   524
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   525
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
   526
apply (case_tac "b\<le>0") 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   527
 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
   528
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
   529
 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
   530
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   531
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   532
lemma zero_less_mult_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   533
     "((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
   534
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
   535
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
   536
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
   537
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
   538
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   539
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   540
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
   541
      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
   542
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
   543
apply (case_tac "a < 0")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   544
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
   545
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
   546
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   547
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   548
lemma zero_le_mult_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   549
     "((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
   550
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
   551
                   zero_less_mult_iff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   552
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   553
lemma mult_less_0_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   554
     "(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
   555
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
   556
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
   557
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   558
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   559
lemma mult_le_0_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   560
     "(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
   561
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
   562
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
   563
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   564
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   565
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
   566
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
   567
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   568
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
   569
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
   570
apply (simp add: order_less_le) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   571
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   572
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   573
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
   574
  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
   575
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   576
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   577
subsection{*More Monotonicity*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   578
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   579
lemma mult_left_mono_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   580
     "[|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
   581
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
   582
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
   583
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   584
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   585
lemma mult_right_mono_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   586
     "[|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
   587
  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
   588
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   589
text{*Strict monotonicity in both arguments*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   590
lemma mult_strict_mono:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   591
     "[|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
   592
apply (case_tac "c=0")
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   593
 apply (simp add: mult_pos) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   594
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
   595
 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
   596
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
   597
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   598
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   599
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
   600
lemma mult_strict_mono':
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   601
     "[| 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
   602
apply (rule mult_strict_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   603
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
   604
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   605
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   606
lemma mult_mono:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   607
     "[|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
   608
      ==> 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
   609
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
   610
apply (erule mult_left_mono, assumption)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   611
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   612
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   613
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   614
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
   615
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   616
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
   617
   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
   618
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   619
lemma mult_less_cancel_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   620
    "(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
   621
apply (case_tac "c = 0")
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   622
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
   623
                      mult_strict_right_mono_neg)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   624
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
   625
                      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
   626
                      linorder_not_le [symmetric, of a])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   627
apply (erule_tac [!] notE)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   628
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
   629
                      mult_right_mono_neg)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   630
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   631
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   632
lemma mult_less_cancel_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   633
    "(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
   634
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
   635
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   636
lemma mult_le_cancel_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   637
     "(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
   638
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
   639
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   640
lemma mult_le_cancel_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   641
     "(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
   642
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
   643
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   644
lemma mult_less_imp_less_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   645
    "[|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
   646
  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
   647
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   648
lemma mult_less_imp_less_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   649
    "[|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
   650
  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
   651
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   652
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
   653
lemma mult_cancel_right [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   654
     "(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
   655
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
   656
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
   657
             simp add: linorder_neq_iff)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   658
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   659
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   660
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
   661
      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
   662
lemma mult_cancel_left [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   663
     "(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
   664
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
   665
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   666
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   667
subsection {* Absolute Value *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   668
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   669
text{*But is it really better than just rewriting with @{text abs_if}?*}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   670
lemma abs_split:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   671
     "P(abs(a::'a::ordered_ring)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   672
by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   673
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   674
lemma abs_zero [simp]: "abs 0 = (0::'a::ordered_ring)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   675
by (simp add: abs_if)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   676
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   677
lemma abs_mult: "abs (x * y) = abs x * abs (y::'a::ordered_ring)" 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   678
apply (case_tac "x=0 | y=0", force) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   679
apply (auto elim: order_less_asym
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   680
            simp add: abs_if mult_less_0_iff linorder_neq_iff
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   681
                  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
   682
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   683
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   684
lemma abs_eq_0 [simp]: "(abs x = 0) = (x = (0::'a::ordered_ring))"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   685
by (simp add: abs_if)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   686
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   687
lemma zero_less_abs_iff [simp]: "(0 < abs x) = (x ~= (0::'a::ordered_ring))"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   688
by (simp add: abs_if linorder_neq_iff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   689
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   690
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   691
subsection {* Fields *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   692
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   693
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
   694
apply (case_tac "b = 0")
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   695
apply (simp_all add: divide_inverse)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   696
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   697
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   698
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
   699
by (simp add: divide_inverse_zero)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   700
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   701
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
   702
by (simp add: divide_inverse_zero)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   703
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   704
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
   705
      of an ordering.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   706
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
   707
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   708
    assume "a=0" thus ?thesis by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   709
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   710
    assume anz [simp]: "a\<noteq>0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   711
    thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   712
    proof auto
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   713
      assume "a * b = 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   714
      hence "inverse a * (a * b) = 0" by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   715
      thus "b = 0"  by (simp (no_asm_use) add: mult_assoc [symmetric])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   716
    qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   717
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   718
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   719
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
   720
lemma field_mult_cancel_right_lemma:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   721
      assumes cnz: "c \<noteq> (0::'a::field)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   722
	  and eq:  "a*c = b*c"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   723
	 shows "a=b"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   724
  proof -
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   725
  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
   726
    by (simp add: eq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   727
  thus "a=b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   728
    by (simp add: mult_assoc cnz)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   729
  qed
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 field_mult_cancel_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   732
     "(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
   733
  proof cases
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   734
    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
   735
  next
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   736
    assume "c\<noteq>0" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   737
    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
   738
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   739
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   740
lemma field_mult_cancel_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   741
     "(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
   742
  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
   743
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   744
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
   745
  proof
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   746
  assume ianz: "inverse a = 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   747
  assume "a \<noteq> 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   748
  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
   749
  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
   750
  finally have "1 = (0::'a::field)" .
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   751
  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
   752
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   753
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   754
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   755
subsection{*Basic Properties of @{term inverse}*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   756
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   757
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
   758
apply (rule ccontr) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   759
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
   760
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   761
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   762
lemma inverse_nonzero_imp_nonzero:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   763
   "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
   764
apply (rule ccontr) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   765
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
   766
done
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_nonzero_iff_nonzero [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   769
   "(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
   770
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
   771
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   772
lemma nonzero_inverse_minus_eq:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   773
      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
   774
  proof -
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   775
    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
   776
      by simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   777
    thus ?thesis 
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   778
      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
   779
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   780
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   781
lemma inverse_minus_eq [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   782
     "inverse(-a) = -inverse(a::'a::{field,division_by_zero})";
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   783
  proof cases
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   784
    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
   785
  next
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   786
    assume "a\<noteq>0" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   787
    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
   788
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   789
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   790
lemma nonzero_inverse_eq_imp_eq:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   791
      assumes inveq: "inverse a = inverse b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   792
	  and anz:  "a \<noteq> 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   793
	  and bnz:  "b \<noteq> 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   794
	 shows "a = (b::'a::field)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   795
  proof -
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   796
  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
   797
    by (simp add: inveq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   798
  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
   799
    by simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   800
  thus "a = b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   801
    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
   802
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   803
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   804
lemma inverse_eq_imp_eq:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   805
     "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
   806
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
   807
 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
   808
              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
   809
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
   810
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   811
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   812
lemma inverse_eq_iff_eq [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   813
     "(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
   814
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
   815
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   816
lemma nonzero_inverse_inverse_eq:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   817
      assumes [simp]: "a \<noteq> 0"  shows "inverse(inverse (a::'a::field)) = a"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   818
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   819
  have "(inverse (inverse a) * inverse a) * a = a" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   820
    by (simp add: nonzero_imp_inverse_nonzero)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   821
  thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   822
    by (simp add: mult_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   823
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   824
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   825
lemma inverse_inverse_eq [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   826
     "inverse(inverse (a::'a::{field,division_by_zero})) = a"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   827
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   828
    assume "a=0" thus ?thesis by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   829
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   830
    assume "a\<noteq>0" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   831
    thus ?thesis by (simp add: nonzero_inverse_inverse_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   832
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   833
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   834
lemma inverse_1 [simp]: "inverse 1 = (1::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   835
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   836
  have "inverse 1 * 1 = (1::'a::field)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   837
    by (rule left_inverse [OF zero_neq_one [symmetric]])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   838
  thus ?thesis  by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   839
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   840
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   841
lemma nonzero_inverse_mult_distrib: 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   842
      assumes anz: "a \<noteq> 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   843
          and bnz: "b \<noteq> 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   844
      shows "inverse(a*b) = inverse(b) * inverse(a::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   845
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   846
  have "inverse(a*b) * (a * b) * inverse(b) = inverse(b)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   847
    by (simp add: field_mult_eq_0_iff anz bnz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   848
  hence "inverse(a*b) * a = inverse(b)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   849
    by (simp add: mult_assoc bnz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   850
  hence "inverse(a*b) * a * inverse(a) = inverse(b) * inverse(a)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   851
    by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   852
  thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   853
    by (simp add: mult_assoc anz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   854
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   855
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   856
text{*This version builds in division by zero while also re-orienting
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   857
      the right-hand side.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   858
lemma inverse_mult_distrib [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   859
     "inverse(a*b) = inverse(a) * inverse(b::'a::{field,division_by_zero})"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   860
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   861
    assume "a \<noteq> 0 & b \<noteq> 0" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   862
    thus ?thesis  by (simp add: nonzero_inverse_mult_distrib mult_commute)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   863
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   864
    assume "~ (a \<noteq> 0 & b \<noteq> 0)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   865
    thus ?thesis  by force
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   866
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   867
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   868
text{*There is no slick version using division by zero.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   869
lemma inverse_add:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   870
     "[|a \<noteq> 0;  b \<noteq> 0|]
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   871
      ==> inverse a + inverse b = (a+b) * inverse a * inverse (b::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   872
apply (simp add: left_distrib mult_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   873
apply (simp add: mult_commute [of "inverse a"]) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   874
apply (simp add: mult_assoc [symmetric] add_commute)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   875
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   876
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   877
lemma nonzero_mult_divide_cancel_left:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   878
  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
   879
    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
   880
proof -
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   881
  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
   882
    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
   883
                  nonzero_inverse_mult_distrib)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   884
  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
   885
    by (simp only: mult_ac)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   886
  also have "... =  a * inverse b"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   887
    by simp
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   888
    finally show ?thesis 
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   889
    by (simp add: divide_inverse)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   890
qed
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   891
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   892
lemma mult_divide_cancel_left:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   893
     "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
   894
apply (case_tac "b = 0")
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   895
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
   896
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   897
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   898
(*For ExtractCommonTerm*)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   899
lemma mult_divide_cancel_eq_if:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   900
     "(c*a) / (c*b) = 
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   901
      (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
   902
  by (simp add: mult_divide_cancel_left)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   903
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   904
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   905
subsection {* Ordered Fields *}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   906
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   907
lemma positive_imp_inverse_positive: 
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   908
      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
   909
  proof -
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   910
  have "0 < a * inverse a" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   911
    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
   912
  thus "0 < inverse a" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   913
    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
   914
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   915
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   916
lemma negative_imp_inverse_negative:
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   917
     "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
   918
  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
   919
      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
   920
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   921
lemma inverse_le_imp_le:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   922
      assumes invle: "inverse a \<le> inverse b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   923
	  and apos:  "0 < a"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   924
	 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
   925
  proof (rule classical)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   926
  assume "~ b \<le> a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   927
  hence "a < b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   928
    by (simp add: linorder_not_le)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   929
  hence bpos: "0 < b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   930
    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
   931
  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
   932
    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
   933
  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
   934
    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
   935
  thus "b \<le> a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   936
    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
   937
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   938
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   939
lemma inverse_positive_imp_positive:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   940
      assumes inv_gt_0: "0 < inverse a"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   941
          and [simp]:   "a \<noteq> 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   942
        shows "0 < (a::'a::ordered_field)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   943
  proof -
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   944
  have "0 < inverse (inverse a)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   945
    by (rule positive_imp_inverse_positive)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   946
  thus "0 < a"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   947
    by (simp add: nonzero_inverse_inverse_eq)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   948
  qed
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   949
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   950
lemma inverse_positive_iff_positive [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   951
      "(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
   952
apply (case_tac "a = 0", simp)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   953
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
   954
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   955
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   956
lemma inverse_negative_imp_negative:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   957
      assumes inv_less_0: "inverse a < 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   958
          and [simp]:   "a \<noteq> 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   959
        shows "a < (0::'a::ordered_field)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   960
  proof -
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   961
  have "inverse (inverse a) < 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   962
    by (rule negative_imp_inverse_negative)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   963
  thus "a < 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   964
    by (simp add: nonzero_inverse_inverse_eq)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   965
  qed
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   966
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   967
lemma inverse_negative_iff_negative [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   968
      "(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
   969
apply (case_tac "a = 0", simp)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   970
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
   971
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   972
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   973
lemma inverse_nonnegative_iff_nonnegative [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   974
      "(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
   975
by (simp add: linorder_not_less [symmetric])
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   976
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   977
lemma inverse_nonpositive_iff_nonpositive [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   978
      "(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
   979
by (simp add: linorder_not_less [symmetric])
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   980
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   981
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   982
subsection{*Anti-Monotonicity of @{term inverse}*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   983
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   984
lemma less_imp_inverse_less:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   985
      assumes less: "a < b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   986
	  and apos:  "0 < a"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   987
	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
   988
  proof (rule ccontr)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   989
  assume "~ inverse b < inverse a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   990
  hence "inverse a \<le> inverse b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   991
    by (simp add: linorder_not_less)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   992
  hence "~ (a < b)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   993
    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
   994
  thus False
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   995
    by (rule notE [OF _ less])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   996
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   997
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   998
lemma inverse_less_imp_less:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   999
   "[|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
  1000
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
  1001
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
  1002
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1003
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1004
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
  1005
lemma inverse_less_iff_less [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1006
     "[|0 < a; 0 < b|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1007
      ==> (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
  1008
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
  1009
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1010
lemma le_imp_inverse_le:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1011
   "[|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
  1012
  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
  1013
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1014
lemma inverse_le_iff_le [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1015
     "[|0 < a; 0 < b|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1016
      ==> (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
  1017
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
  1018
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1019
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1020
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
  1021
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
  1022
lemma inverse_le_imp_le_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1023
   "[|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
  1024
  apply (rule classical) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1025
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1026
   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
  1027
  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
  1028
  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
  1029
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1030
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1031
lemma less_imp_inverse_less_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1032
   "[|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
  1033
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1034
   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
  1035
  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
  1036
  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
  1037
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1038
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1039
lemma inverse_less_imp_less_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1040
   "[|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
  1041
  apply (rule classical) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1042
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1043
   prefer 2
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1044
   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
  1045
  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
  1046
  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
  1047
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1048
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1049
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
  1050
     "[|a < 0; b < 0|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1051
      ==> (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
  1052
  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
  1053
  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
  1054
	      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
  1055
  done
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
lemma le_imp_inverse_le_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1058
   "[|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
  1059
  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
  1060
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1061
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
  1062
     "[|a < 0; b < 0|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1063
      ==> (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
  1064
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
  1065
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1066
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1067
subsection{*Division and Signs*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1068
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1069
lemma zero_less_divide_iff:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1070
     "((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
  1071
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
  1072
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1073
lemma divide_less_0_iff:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1074
     "(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
  1075
      (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
  1076
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
  1077
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1078
lemma zero_le_divide_iff:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1079
     "((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
  1080
      (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
  1081
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
  1082
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1083
lemma divide_le_0_iff:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1084
     "(a/b \<le> (0::'a::{ordered_field,division_by_zero})) = (0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1085
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
  1086
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1087
lemma divide_eq_0_iff [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1088
     "(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
  1089
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
  1090
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1091
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1092
end