src/HOL/Ring_and_Field.thy
author nipkow
Thu, 15 Apr 2004 14:17:45 +0200
changeset 14569 78b75a9eec01
parent 14504 7a3d80e276d4
child 14603 985eb6708207
permissions -rw-r--r--
Added ex/Exceptions.thy
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
14569
78b75a9eec01 Added ex/Exceptions.thy
nipkow
parents: 14504
diff changeset
     8
header {* Ring and field structures *}
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     9
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    10
theory Ring_and_Field = Inductive:
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
subsection {* Abstract algebraic structures *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    13
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    14
subsection {* Types Classes @{text plus_ac0} and @{text times_ac1} *}
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    15
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    16
axclass plus_ac0 \<subseteq> plus, zero
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    17
  commute:     "x + y = y + x"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    18
  assoc:       "(x + y) + z = x + (y + z)"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    19
  zero [simp]: "0 + x = x"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    20
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    21
lemma plus_ac0_left_commute: "x + (y+z) = y + ((x+z)::'a::plus_ac0)"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    22
apply (rule plus_ac0.commute [THEN trans])
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    23
apply (rule plus_ac0.assoc [THEN trans])
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    24
apply (rule plus_ac0.commute [THEN arg_cong])
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    25
done
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    26
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    27
lemma plus_ac0_zero_right [simp]: "x + 0 = (x ::'a::plus_ac0)"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    28
apply (rule plus_ac0.commute [THEN trans])
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    29
apply (rule plus_ac0.zero)
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    30
done
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    31
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    32
lemmas plus_ac0 = plus_ac0.assoc plus_ac0.commute plus_ac0_left_commute
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    33
                  plus_ac0.zero plus_ac0_zero_right
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    34
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    35
axclass times_ac1 \<subseteq> times, one
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    36
  commute:     "x * y = y * x"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    37
  assoc:       "(x * y) * z = x * (y * z)"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    38
  one [simp]:  "1 * x = x"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    39
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    40
theorem times_ac1_left_commute: "(x::'a::times_ac1) * ((y::'a) * z) =
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    41
  y * (x * z)"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    42
proof -
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    43
  have "(x::'a::times_ac1) * (y * z) = (x * y) * z"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    44
    by (rule times_ac1.assoc [THEN sym])
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    45
  also have "x * y = y * x"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    46
    by (rule times_ac1.commute)
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    47
  also have "(y * x) * z = y * (x * z)"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    48
    by (rule times_ac1.assoc)
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    49
  finally show ?thesis .
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    50
qed
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    51
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    52
theorem times_ac1_one_right [simp]: "(x::'a::times_ac1) * 1 = x"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    53
proof -
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    54
  have "x * 1 = 1 * x"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    55
    by (rule times_ac1.commute)
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    56
  also have "... = x"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    57
    by (rule times_ac1.one)
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    58
  finally show ?thesis .
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    59
qed
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    60
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    61
theorems times_ac1 = times_ac1.assoc times_ac1.commute times_ac1_left_commute
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    62
  times_ac1.one times_ac1_one_right
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    63
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    64
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    65
text{*This class is the same as @{text plus_ac0}, while using the same axiom
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    66
names as the other axclasses.*}
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    67
axclass abelian_semigroup \<subseteq> zero, plus
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    68
  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
    69
  add_commute: "a + b = b + a"
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
    70
  add_0 [simp]: "0 + a = a"
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    71
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    72
text{*This class underlies both @{text semiring} and @{text ring}*}
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    73
axclass almost_semiring \<subseteq> abelian_semigroup, one, times
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    74
  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
    75
  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
    76
  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
    77
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    78
  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
    79
  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
    80
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    81
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
    82
axclass semiring \<subseteq> almost_semiring
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
    83
  add_left_imp_eq: "a + b = a + c ==> b=c"
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
    84
    --{*This axiom is needed for semirings only: for rings, etc., it is
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
    85
        redundant. Including it allows many more of the following results
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
    86
        to be proved for semirings too.*}
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
    87
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    88
instance abelian_semigroup \<subseteq> plus_ac0
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    89
proof
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    90
  fix x y z :: 'a
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    91
  show "x + y = y + x" by (rule add_commute)
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    92
  show "x + y + z = x + (y + z)" by (rule add_assoc)
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    93
  show "0+x = x" by (rule add_0)
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    94
qed
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    95
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    96
instance almost_semiring \<subseteq> times_ac1
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    97
proof
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    98
  fix x y z :: 'a
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    99
  show "x * y = y * x" by (rule mult_commute)
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   100
  show "x * y * z = x * (y * z)" by (rule mult_assoc)
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   101
  show "1*x = x" by simp
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   102
qed
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   103
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   104
axclass abelian_group \<subseteq> abelian_semigroup, minus
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   105
   left_minus [simp]: "-a + a = 0"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   106
   diff_minus: "a - b = a + -b"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   107
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   108
axclass ring \<subseteq> almost_semiring, abelian_group
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   109
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   110
text{*Proving axiom @{text add_left_imp_eq} makes all @{text semiring}
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   111
      theorems available to members of @{term ring} *}
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   112
instance ring \<subseteq> semiring
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   113
proof
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   114
  fix a b c :: 'a
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   115
  assume "a + b = a + c"
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   116
  hence  "-a + a + b = -a + a + c" by (simp only: add_assoc)
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   117
  thus "b = c" by simp
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   118
qed
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   119
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   120
text{*This class underlies @{text ordered_semiring} and @{text ordered_ring}*}
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   121
axclass almost_ordered_semiring \<subseteq> semiring, linorder
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   122
  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
   123
  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
   124
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   125
axclass ordered_semiring \<subseteq> almost_ordered_semiring
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   126
  zero_less_one [simp]: "0 < 1" --{*This too is needed for semirings only.*}
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   127
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   128
axclass ordered_ring \<subseteq> almost_ordered_semiring, ring
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   129
  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
   130
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   131
axclass field \<subseteq> ring, inverse
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   132
  left_inverse [simp]: "a \<noteq> 0 ==> inverse a * a = 1"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   133
  divide_inverse:      "a / b = a * inverse b"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   134
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   135
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
   136
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   137
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
   138
  inverse_zero [simp]: "inverse 0 = 0"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   139
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   140
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   141
subsection {* Derived Rules for Addition *}
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   142
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   143
lemma add_0_right [simp]: "a + 0 = (a::'a::plus_ac0)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   144
proof -
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   145
  have "a + 0 = 0 + a" by (rule plus_ac0.commute)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   146
  also have "... = a" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   147
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   148
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   149
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   150
lemma add_left_commute: "a + (b + c) = b + (a + (c::'a::plus_ac0))"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   151
  by (rule mk_left_commute [of "op +", OF plus_ac0.assoc plus_ac0.commute])
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   152
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   153
theorems 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
   154
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   155
lemma right_minus [simp]: "a + -(a::'a::abelian_group) = 0"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   156
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   157
  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
   158
  also have "... = 0" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   159
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   160
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   161
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   162
lemma right_minus_eq: "(a - b = 0) = (a = (b::'a::abelian_group))"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   163
proof
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   164
  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
   165
  also assume "a - b = 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   166
  finally show "a = b" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   167
next
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   168
  assume "a = b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   169
  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
   170
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   171
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   172
lemma add_left_cancel [simp]:
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   173
     "(a + b = a + c) = (b = (c::'a::semiring))"
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   174
by (blast dest: add_left_imp_eq) 
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   175
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   176
lemma add_right_cancel [simp]:
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   177
     "(b + a = c + a) = (b = (c::'a::semiring))"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   178
  by (simp add: add_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   179
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   180
lemma minus_minus [simp]: "- (- (a::'a::abelian_group)) = a" 
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   181
apply (rule right_minus_eq [THEN iffD1]) 
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   182
apply (simp add: diff_minus) 
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   183
done
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   184
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   185
lemma equals_zero_I: "a+b = 0 ==> -a = (b::'a::abelian_group)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   186
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
   187
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
   188
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   189
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   190
lemma minus_zero [simp]: "- 0 = (0::'a::abelian_group)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   191
by (simp add: equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   192
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   193
lemma diff_self [simp]: "a - (a::'a::abelian_group) = 0"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   194
  by (simp add: diff_minus)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   195
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   196
lemma diff_0 [simp]: "(0::'a::abelian_group) - a = -a"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   197
by (simp add: diff_minus)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   198
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   199
lemma diff_0_right [simp]: "a - (0::'a::abelian_group) = a" 
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   200
by (simp add: diff_minus)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   201
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   202
lemma diff_minus_eq_add [simp]: "a - - b = a + (b::'a::abelian_group)"
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   203
by (simp add: diff_minus)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   204
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   205
lemma neg_equal_iff_equal [simp]: "(-a = -b) = (a = (b::'a::abelian_group))" 
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   206
proof 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   207
  assume "- a = - b"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   208
  hence "- (- a) = - (- b)"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   209
    by simp
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   210
  thus "a=b" by simp
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   211
next
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   212
  assume "a=b"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   213
  thus "-a = -b" by simp
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   214
qed
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   215
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   216
lemma neg_equal_0_iff_equal [simp]: "(-a = 0) = (a = (0::'a::abelian_group))"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   217
by (subst neg_equal_iff_equal [symmetric], simp)
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   218
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   219
lemma neg_0_equal_iff_equal [simp]: "(0 = -a) = (0 = (a::'a::abelian_group))"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   220
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
   221
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   222
lemma add_minus_self [simp]: "a + b - b = (a::'a::abelian_group)"; 
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   223
  by (simp add: diff_minus add_assoc)
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   224
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   225
lemma add_minus_self_left [simp]:  "a + (b - a)  = (b::'a::abelian_group)";
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   226
by (simp add: diff_minus add_left_commute [of a]) 
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   227
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   228
lemma add_minus_self_right  [simp]:  "a + b - a  = (b::'a::abelian_group)";
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   229
by (simp add: diff_minus add_left_commute [of a] add_assoc) 
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   230
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   231
lemma minus_add_self [simp]: "a - b + b = (a::'a::abelian_group)"; 
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   232
by (simp add: diff_minus add_assoc) 
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   233
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   234
text{*The next two equations can make the simplifier loop!*}
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   235
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   236
lemma equation_minus_iff: "(a = - b) = (b = - (a::'a::abelian_group))"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   237
proof -
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   238
  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
   239
  thus ?thesis by (simp add: eq_commute)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   240
qed
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   241
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   242
lemma minus_equation_iff: "(- a = b) = (- (b::'a::abelian_group) = a)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   243
proof -
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   244
  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
   245
  thus ?thesis by (simp add: eq_commute)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   246
qed
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   247
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   248
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   249
subsection {* Derived rules for multiplication *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   250
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   251
lemma mult_1_right [simp]: "a * (1::'a::almost_semiring) = a"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   252
proof -
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   253
  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
   254
  also have "... = a" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   255
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   256
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   257
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   258
lemma mult_left_commute: "a * (b * c) = b * (a * (c::'a::almost_semiring))"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   259
  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
   260
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   261
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
   262
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
   263
lemma mult_zero_left [simp]: "0 * a = (0::'a::semiring)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   264
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   265
  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
   266
    by (simp add: left_distrib [symmetric])
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   267
  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
   268
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   269
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
   270
lemma mult_zero_right [simp]: "a * 0 = (0::'a::semiring)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   271
  by (simp add: mult_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   272
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   273
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   274
subsection {* Distribution rules *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   275
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   276
lemma right_distrib: "a * (b + c) = a * b + a * (c::'a::almost_semiring)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   277
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   278
  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
   279
  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
   280
  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
   281
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   282
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   283
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   284
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
   285
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   286
text{*For the @{text combine_numerals} simproc*}
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   287
lemma combine_common_factor:
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   288
     "a*e + (b*e + c) = (a+b)*e + (c::'a::almost_semiring)"
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   289
by (simp add: left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   290
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   291
lemma minus_add_distrib [simp]: "- (a + b) = -a + -(b::'a::abelian_group)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   292
apply (rule equals_zero_I)
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   293
apply (simp add: plus_ac0) 
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   294
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   295
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   296
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
   297
apply (rule equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   298
apply (simp add: left_distrib [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   299
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   300
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   301
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
   302
apply (rule equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   303
apply (simp add: right_distrib [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   304
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   305
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   306
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
   307
  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
   308
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   309
lemma minus_mult_commute: "(- a) * b = a * (- b::'a::ring)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   310
  by (simp add: minus_mult_left [symmetric] minus_mult_right [symmetric])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   311
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   312
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
   313
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
   314
              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
   315
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   316
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
   317
by (simp add: mult_commute [of _ c] right_diff_distrib) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   318
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   319
lemma minus_diff_eq [simp]: "- (a - b) = b - (a::'a::ring)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   320
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
   321
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   322
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   323
subsection {* Ordering Rules for Addition *}
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   324
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   325
lemma add_right_mono: "a \<le> (b::'a::almost_ordered_semiring) ==> a + c \<le> b + c"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   326
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
   327
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   328
text {* non-strict, in both arguments *}
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   329
lemma add_mono:
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   330
     "[|a \<le> b;  c \<le> d|] ==> a + c \<le> b + (d::'a::almost_ordered_semiring)"
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   331
  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
   332
  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
   333
  done
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   334
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   335
lemma add_strict_left_mono:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   336
     "a < b ==> c + a < c + (b::'a::almost_ordered_semiring)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   337
 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
   338
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   339
lemma add_strict_right_mono:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   340
     "a < b ==> a + c < b + (c::'a::almost_ordered_semiring)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   341
 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
   342
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   343
text{*Strict monotonicity in both arguments*}
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   344
lemma add_strict_mono: "[|a<b; c<d|] ==> a + c < b + (d::'a::almost_ordered_semiring)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   345
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
   346
apply (erule add_strict_left_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   347
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   348
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 14368
diff changeset
   349
lemma add_less_le_mono:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   350
     "[| a<b; c\<le>d |] ==> a + c < b + (d::'a::almost_ordered_semiring)"
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   351
apply (erule add_strict_right_mono [THEN order_less_le_trans])
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   352
apply (erule add_left_mono) 
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   353
done
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   354
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   355
lemma add_le_less_mono:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   356
     "[| a\<le>b; c<d |] ==> a + c < b + (d::'a::almost_ordered_semiring)"
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   357
apply (erule add_right_mono [THEN order_le_less_trans])
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   358
apply (erule add_strict_left_mono) 
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   359
done
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   360
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   361
lemma add_less_imp_less_left:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   362
      assumes less: "c + a < c + b"  shows "a < (b::'a::almost_ordered_semiring)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   363
proof (rule ccontr)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   364
  assume "~ a < b"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   365
  hence "b \<le> a" by (simp add: linorder_not_less)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   366
  hence "c+b \<le> c+a" by (rule add_left_mono)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   367
  with this and less show False 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   368
    by (simp add: linorder_not_less [symmetric])
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   369
qed
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   370
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   371
lemma add_less_imp_less_right:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   372
      "a + c < b + c ==> a < (b::'a::almost_ordered_semiring)"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   373
apply (rule add_less_imp_less_left [of c])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   374
apply (simp add: add_commute)  
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   375
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   376
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   377
lemma add_less_cancel_left [simp]:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   378
    "(c+a < c+b) = (a < (b::'a::almost_ordered_semiring))"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   379
by (blast intro: add_less_imp_less_left add_strict_left_mono) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   380
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   381
lemma add_less_cancel_right [simp]:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   382
    "(a+c < b+c) = (a < (b::'a::almost_ordered_semiring))"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   383
by (blast intro: add_less_imp_less_right add_strict_right_mono)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   384
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   385
lemma add_le_cancel_left [simp]:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   386
    "(c+a \<le> c+b) = (a \<le> (b::'a::almost_ordered_semiring))"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   387
by (simp add: linorder_not_less [symmetric]) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   388
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   389
lemma add_le_cancel_right [simp]:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   390
    "(a+c \<le> b+c) = (a \<le> (b::'a::almost_ordered_semiring))"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   391
by (simp add: linorder_not_less [symmetric]) 
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 add_le_imp_le_left:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   394
      "c + a \<le> c + b ==> a \<le> (b::'a::almost_ordered_semiring)"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   395
by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   396
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   397
lemma add_le_imp_le_right:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   398
      "a + c \<le> b + c ==> a \<le> (b::'a::almost_ordered_semiring)"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   399
by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   400
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   401
lemma add_increasing: "[|0\<le>a; b\<le>c|] ==> b \<le> a + (c::'a::almost_ordered_semiring)"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   402
by (insert add_mono [of 0 a b c], simp)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   403
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   404
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   405
subsection {* Ordering Rules for Unary Minus *}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   406
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   407
lemma le_imp_neg_le:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   408
      assumes "a \<le> (b::'a::ordered_ring)" shows "-b \<le> -a"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   409
proof -
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   410
  have "-a+a \<le> -a+b"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   411
    by (rule add_left_mono) 
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   412
  hence "0 \<le> -a+b"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   413
    by simp
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   414
  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
   415
    by (rule add_right_mono) 
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   416
  thus ?thesis
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   417
    by (simp add: add_assoc)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   418
qed
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   419
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   420
lemma neg_le_iff_le [simp]: "(-b \<le> -a) = (a \<le> (b::'a::ordered_ring))"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   421
proof 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   422
  assume "- b \<le> - a"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   423
  hence "- (- a) \<le> - (- b)"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   424
    by (rule le_imp_neg_le)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   425
  thus "a\<le>b" by simp
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   426
next
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   427
  assume "a\<le>b"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   428
  thus "-b \<le> -a" by (rule le_imp_neg_le)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   429
qed
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   430
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   431
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
   432
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
   433
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   434
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
   435
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
   436
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   437
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
   438
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
   439
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   440
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
   441
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
   442
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   443
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
   444
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
   445
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   446
text{*The next several equations can make the simplifier loop!*}
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 less_minus_iff: "(a < - b) = (b < - (a::'a::ordered_ring))"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   449
proof -
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   450
  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
   451
  thus ?thesis by simp
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   452
qed
14272
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 minus_less_iff: "(- a < b) = (- b < (a::'a::ordered_ring))"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   455
proof -
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   456
  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
   457
  thus ?thesis by simp
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   458
qed
14272
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 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
   461
apply (simp add: linorder_not_less [symmetric])
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   462
apply (rule minus_less_iff) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   463
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   464
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   465
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
   466
apply (simp add: linorder_not_less [symmetric])
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   467
apply (rule less_minus_iff) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   468
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   469
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   470
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   471
subsection{*Subtraction Laws*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   472
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   473
lemma add_diff_eq: "a + (b - c) = (a + b) - (c::'a::abelian_group)"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   474
by (simp add: diff_minus plus_ac0)
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   475
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   476
lemma diff_add_eq: "(a - b) + c = (a + c) - (b::'a::abelian_group)"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   477
by (simp add: diff_minus plus_ac0)
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   478
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   479
lemma diff_eq_eq: "(a-b = c) = (a = c + (b::'a::abelian_group))"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   480
by (auto simp add: diff_minus add_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   481
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   482
lemma eq_diff_eq: "(a = c-b) = (a + (b::'a::abelian_group) = c)"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   483
by (auto simp add: diff_minus add_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   484
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   485
lemma diff_diff_eq: "(a - b) - c = a - (b + (c::'a::abelian_group))"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   486
by (simp add: diff_minus plus_ac0)
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   487
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   488
lemma diff_diff_eq2: "a - (b - c) = (a + c) - (b::'a::abelian_group)"
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   489
by (simp add: diff_minus plus_ac0)
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   490
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   491
text{*Further subtraction laws for ordered rings*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   492
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   493
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
   494
proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   495
  have  "(a < b) = (a + (- b) < b + (-b))"  
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   496
    by (simp only: add_less_cancel_right)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   497
  also have "... =  (a - b < 0)" by (simp add: diff_minus)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   498
  finally show ?thesis .
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   499
qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   500
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   501
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
   502
apply (subst less_iff_diff_less_0)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   503
apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   504
apply (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   505
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   506
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   507
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
   508
apply (subst less_iff_diff_less_0)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   509
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
   510
apply (simp add: diff_minus add_ac)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   511
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   512
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   513
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
   514
by (simp add: linorder_not_less [symmetric] less_diff_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   515
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   516
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
   517
by (simp add: linorder_not_less [symmetric] diff_less_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   518
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   519
text{*This list of rewrites simplifies (in)equalities by bringing subtractions
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   520
  to the top and then moving negative terms to the other side.
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   521
  Use with @{text add_ac}*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   522
lemmas compare_rls =
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   523
       diff_minus [symmetric]
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   524
       add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   525
       diff_less_eq less_diff_eq diff_le_eq le_diff_eq
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   526
       diff_eq_eq eq_diff_eq
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   527
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   528
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   529
subsection{*Lemmas for the @{text cancel_numerals} simproc*}
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   530
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   531
lemma eq_iff_diff_eq_0: "(a = b) = (a-b = (0::'a::abelian_group))"
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   532
by (simp add: compare_rls)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   533
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   534
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
   535
by (simp add: compare_rls)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   536
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   537
lemma eq_add_iff1:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   538
     "(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
   539
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   540
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   541
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   542
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   543
lemma eq_add_iff2:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   544
     "(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
   545
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   546
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   547
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   548
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   549
lemma less_add_iff1:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   550
     "(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
   551
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   552
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   553
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   554
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   555
lemma less_add_iff2:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   556
     "(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
   557
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   558
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   559
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   560
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   561
lemma le_add_iff1:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   562
     "(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
   563
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   564
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   565
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   566
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   567
lemma le_add_iff2:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   568
     "(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
   569
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   570
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   571
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   572
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   573
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   574
subsection {* Ordering Rules for Multiplication *}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   575
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   576
lemma mult_strict_right_mono:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   577
     "[|a < b; 0 < c|] ==> a * c < b * (c::'a::almost_ordered_semiring)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   578
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
   579
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   580
lemma mult_left_mono:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   581
     "[|a \<le> b; 0 \<le> c|] ==> c * a \<le> c * (b::'a::almost_ordered_semiring)"
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   582
  apply (case_tac "c=0", simp)
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   583
  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
   584
  done
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   585
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   586
lemma mult_right_mono:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   587
     "[|a \<le> b; 0 \<le> c|] ==> a*c \<le> b * (c::'a::almost_ordered_semiring)"
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   588
  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
   589
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   590
lemma mult_left_le_imp_le:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   591
     "[|c*a \<le> c*b; 0 < c|] ==> a \<le> (b::'a::almost_ordered_semiring)"
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   592
  by (force simp add: mult_strict_left_mono linorder_not_less [symmetric])
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   593
 
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   594
lemma mult_right_le_imp_le:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   595
     "[|a*c \<le> b*c; 0 < c|] ==> a \<le> (b::'a::almost_ordered_semiring)"
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   596
  by (force simp add: mult_strict_right_mono linorder_not_less [symmetric])
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   597
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   598
lemma mult_left_less_imp_less:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   599
     "[|c*a < c*b; 0 \<le> c|] ==> a < (b::'a::almost_ordered_semiring)"
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   600
  by (force simp add: mult_left_mono linorder_not_le [symmetric])
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   601
 
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   602
lemma mult_right_less_imp_less:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   603
     "[|a*c < b*c; 0 \<le> c|] ==> a < (b::'a::almost_ordered_semiring)"
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   604
  by (force simp add: mult_right_mono linorder_not_le [symmetric])
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   605
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   606
lemma mult_strict_left_mono_neg:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   607
     "[|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
   608
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
   609
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
   610
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   611
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   612
lemma mult_strict_right_mono_neg:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   613
     "[|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
   614
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
   615
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
   616
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   617
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   618
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   619
subsection{* Products of Signs *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   620
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   621
lemma mult_pos: "[| (0::'a::almost_ordered_semiring) < a; 0 < b |] ==> 0 < a*b"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   622
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
   623
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   624
lemma mult_pos_neg: "[| (0::'a::almost_ordered_semiring) < a; b < 0 |] ==> a*b < 0"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   625
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
   626
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   627
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
   628
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
   629
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   630
lemma zero_less_mult_pos:
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   631
     "[| 0 < a*b; 0 < a|] ==> 0 < (b::'a::almost_ordered_semiring)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   632
apply (case_tac "b\<le>0") 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   633
 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
   634
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
   635
 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
   636
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   637
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   638
lemma zero_less_mult_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   639
     "((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
   640
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
   641
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
   642
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
   643
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
   644
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   645
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   646
text{*A field has no "zero divisors", and this theorem holds without the
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   647
      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
   648
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
   649
apply (case_tac "a < 0")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   650
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
   651
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
   652
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   653
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   654
lemma zero_le_mult_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   655
     "((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
   656
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
   657
                   zero_less_mult_iff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   658
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   659
lemma mult_less_0_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   660
     "(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
   661
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
   662
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
   663
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   664
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   665
lemma mult_le_0_iff:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   666
     "(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
   667
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
   668
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
   669
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   670
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   671
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
   672
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
   673
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   674
text{*Proving axiom @{text zero_less_one} makes all @{text ordered_semiring}
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   675
      theorems available to members of @{term ordered_ring} *}
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   676
instance ordered_ring \<subseteq> ordered_semiring
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   677
proof
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   678
  have "(0::'a) \<le> 1*1" by (rule zero_le_square)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   679
  thus "(0::'a) < 1" by (simp add: order_le_less) 
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   680
qed
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   681
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   682
text{*All three types of comparision involving 0 and 1 are covered.*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   683
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   684
declare zero_neq_one [THEN not_sym, simp]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   685
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   686
lemma zero_le_one [simp]: "(0::'a::ordered_semiring) \<le> 1"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   687
  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
   688
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   689
lemma not_one_le_zero [simp]: "~ (1::'a::ordered_semiring) \<le> 0"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   690
by (simp add: linorder_not_le zero_less_one) 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   691
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   692
lemma not_one_less_zero [simp]: "~ (1::'a::ordered_semiring) < 0"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   693
by (simp add: linorder_not_less zero_le_one) 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   694
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   695
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   696
subsection{*More Monotonicity*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   697
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   698
lemma mult_left_mono_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   699
     "[|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
   700
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
   701
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
   702
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   703
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   704
lemma mult_right_mono_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   705
     "[|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
   706
  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
   707
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   708
text{*Strict monotonicity in both arguments*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   709
lemma mult_strict_mono:
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   710
     "[|a<b; c<d; 0<b; 0\<le>c|] ==> a * c < b * (d::'a::ordered_semiring)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   711
apply (case_tac "c=0")
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   712
 apply (simp add: mult_pos) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   713
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
   714
 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
   715
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
   716
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   717
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   718
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
   719
lemma mult_strict_mono':
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   720
     "[| a<b; c<d; 0 \<le> a; 0 \<le> c|] ==> a * c < b * (d::'a::ordered_semiring)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   721
apply (rule mult_strict_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   722
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
   723
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   724
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   725
lemma mult_mono:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   726
     "[|a \<le> b; c \<le> d; 0 \<le> b; 0 \<le> c|] 
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   727
      ==> a * c  \<le>  b * (d::'a::ordered_semiring)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   728
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
   729
apply (erule mult_left_mono, assumption)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   730
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   731
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   732
lemma less_1_mult: "[| 1 < m; 1 < n |] ==> 1 < m*(n::'a::ordered_semiring)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   733
apply (insert mult_strict_mono [of 1 m 1 n]) 
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   734
apply (simp add:  order_less_trans [OF zero_less_one]) 
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   735
done
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   736
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   737
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   738
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
   739
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   740
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
   741
   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
   742
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   743
lemma mult_less_cancel_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   744
    "(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
   745
apply (case_tac "c = 0")
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   746
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
   747
                      mult_strict_right_mono_neg)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   748
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
   749
                      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
   750
                      linorder_not_le [symmetric, of a])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   751
apply (erule_tac [!] notE)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   752
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
   753
                      mult_right_mono_neg)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   754
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   755
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   756
lemma mult_less_cancel_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   757
    "(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
   758
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
   759
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   760
lemma mult_le_cancel_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   761
     "(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
   762
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
   763
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   764
lemma mult_le_cancel_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   765
     "(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
   766
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
   767
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   768
lemma mult_less_imp_less_left:
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   769
      assumes less: "c*a < c*b" and nonneg: "0 \<le> c"
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   770
      shows "a < (b::'a::ordered_semiring)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   771
proof (rule ccontr)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   772
  assume "~ a < b"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   773
  hence "b \<le> a" by (simp add: linorder_not_less)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   774
  hence "c*b \<le> c*a" by (rule mult_left_mono)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   775
  with this and less show False 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   776
    by (simp add: linorder_not_less [symmetric])
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   777
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   778
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   779
lemma mult_less_imp_less_right:
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   780
    "[|a*c < b*c; 0 \<le> c|] ==> a < (b::'a::ordered_semiring)"
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   781
  by (rule mult_less_imp_less_left, simp add: mult_commute)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   782
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   783
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
   784
lemma mult_cancel_right [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   785
     "(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
   786
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
   787
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
   788
             simp add: linorder_neq_iff)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   789
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   790
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   791
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
   792
      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
   793
lemma mult_cancel_left [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   794
     "(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
   795
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
   796
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   797
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   798
subsection {* Fields *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   799
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   800
lemma right_inverse [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   801
      assumes not0: "a \<noteq> 0" shows "a * inverse (a::'a::field) = 1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   802
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   803
  have "a * inverse a = inverse a * a" by (simp add: mult_ac)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   804
  also have "... = 1" using not0 by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   805
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   806
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   807
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   808
lemma right_inverse_eq: "b \<noteq> 0 ==> (a / b = 1) = (a = (b::'a::field))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   809
proof
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   810
  assume neq: "b \<noteq> 0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   811
  {
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   812
    hence "a = (a / b) * b" by (simp add: divide_inverse mult_ac)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   813
    also assume "a / b = 1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   814
    finally show "a = b" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   815
  next
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   816
    assume "a = b"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   817
    with neq show "a / b = 1" by (simp add: divide_inverse)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   818
  }
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   819
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   820
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   821
lemma nonzero_inverse_eq_divide: "a \<noteq> 0 ==> inverse (a::'a::field) = 1/a"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   822
by (simp add: divide_inverse)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   823
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   824
lemma divide_self [simp]: "a \<noteq> 0 ==> a / (a::'a::field) = 1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   825
  by (simp add: divide_inverse)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   826
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   827
lemma divide_zero [simp]: "a / 0 = (0::'a::{field,division_by_zero})"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   828
by (simp add: divide_inverse)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   829
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   830
lemma divide_zero_left [simp]: "0/a = (0::'a::field)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   831
by (simp add: divide_inverse)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   832
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   833
lemma inverse_eq_divide: "inverse (a::'a::field) = 1/a"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   834
by (simp add: divide_inverse)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   835
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   836
lemma add_divide_distrib: "(a+b)/(c::'a::field) = a/c + b/c"
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   837
by (simp add: divide_inverse left_distrib) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   838
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   839
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   840
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
   841
      of an ordering.*}
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   842
lemma field_mult_eq_0_iff [simp]: "(a*b = (0::'a::field)) = (a = 0 | b = 0)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   843
proof cases
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   844
  assume "a=0" thus ?thesis by simp
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   845
next
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   846
  assume anz [simp]: "a\<noteq>0"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   847
  { assume "a * b = 0"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   848
    hence "inverse a * (a * b) = 0" by simp
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   849
    hence "b = 0"  by (simp (no_asm_use) add: mult_assoc [symmetric])}
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   850
  thus ?thesis by force
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   851
qed
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   852
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   853
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
   854
lemma field_mult_cancel_right_lemma:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   855
      assumes cnz: "c \<noteq> (0::'a::field)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   856
	  and eq:  "a*c = b*c"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   857
	 shows "a=b"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   858
proof -
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   859
  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
   860
    by (simp add: eq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   861
  thus "a=b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   862
    by (simp add: mult_assoc cnz)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   863
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   864
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   865
lemma field_mult_cancel_right [simp]:
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   866
     "(a*c = b*c) = (c = (0::'a::field) | a=b)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   867
proof cases
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   868
  assume "c=0" thus ?thesis by simp
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   869
next
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   870
  assume "c\<noteq>0" 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   871
  thus ?thesis by (force dest: field_mult_cancel_right_lemma)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   872
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   873
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   874
lemma field_mult_cancel_left [simp]:
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   875
     "(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
   876
  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
   877
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   878
lemma nonzero_imp_inverse_nonzero: "a \<noteq> 0 ==> inverse a \<noteq> (0::'a::field)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   879
proof
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   880
  assume ianz: "inverse a = 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   881
  assume "a \<noteq> 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   882
  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
   883
  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
   884
  finally have "1 = (0::'a::field)" .
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   885
  thus False by (simp add: eq_commute)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   886
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   887
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   888
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   889
subsection{*Basic Properties of @{term inverse}*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   890
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   891
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
   892
apply (rule ccontr) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   893
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
   894
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   895
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   896
lemma inverse_nonzero_imp_nonzero:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   897
   "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
   898
apply (rule ccontr) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   899
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
   900
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   901
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   902
lemma inverse_nonzero_iff_nonzero [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   903
   "(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
   904
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
   905
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   906
lemma nonzero_inverse_minus_eq:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   907
      assumes [simp]: "a\<noteq>0"  shows "inverse(-a) = -inverse(a::'a::field)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   908
proof -
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   909
  have "-a * inverse (- a) = -a * - inverse a"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   910
    by simp
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   911
  thus ?thesis 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   912
    by (simp only: field_mult_cancel_left, simp)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   913
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   914
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   915
lemma inverse_minus_eq [simp]:
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   916
   "inverse(-a) = -inverse(a::'a::{field,division_by_zero})";
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   917
proof cases
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   918
  assume "a=0" thus ?thesis by (simp add: inverse_zero)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   919
next
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   920
  assume "a\<noteq>0" 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   921
  thus ?thesis by (simp add: nonzero_inverse_minus_eq)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   922
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   923
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   924
lemma nonzero_inverse_eq_imp_eq:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   925
      assumes inveq: "inverse a = inverse b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   926
	  and anz:  "a \<noteq> 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   927
	  and bnz:  "b \<noteq> 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   928
	 shows "a = (b::'a::field)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   929
proof -
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   930
  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
   931
    by (simp add: inveq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   932
  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
   933
    by simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   934
  thus "a = b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   935
    by (simp add: mult_assoc anz bnz)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   936
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   937
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   938
lemma inverse_eq_imp_eq:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   939
     "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
   940
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
   941
 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
   942
              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
   943
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
   944
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   945
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   946
lemma inverse_eq_iff_eq [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   947
     "(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
   948
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
   949
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   950
lemma nonzero_inverse_inverse_eq:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   951
      assumes [simp]: "a \<noteq> 0"  shows "inverse(inverse (a::'a::field)) = a"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   952
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   953
  have "(inverse (inverse a) * inverse a) * a = a" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   954
    by (simp add: nonzero_imp_inverse_nonzero)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   955
  thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   956
    by (simp add: mult_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   957
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   958
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   959
lemma inverse_inverse_eq [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   960
     "inverse(inverse (a::'a::{field,division_by_zero})) = a"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   961
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   962
    assume "a=0" thus ?thesis by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   963
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   964
    assume "a\<noteq>0" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   965
    thus ?thesis by (simp add: nonzero_inverse_inverse_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   966
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   967
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   968
lemma inverse_1 [simp]: "inverse 1 = (1::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   969
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   970
  have "inverse 1 * 1 = (1::'a::field)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   971
    by (rule left_inverse [OF zero_neq_one [symmetric]])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   972
  thus ?thesis  by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   973
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   974
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   975
lemma nonzero_inverse_mult_distrib: 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   976
      assumes anz: "a \<noteq> 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   977
          and bnz: "b \<noteq> 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   978
      shows "inverse(a*b) = inverse(b) * inverse(a::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   979
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   980
  have "inverse(a*b) * (a * b) * inverse(b) = inverse(b)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   981
    by (simp add: field_mult_eq_0_iff anz bnz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   982
  hence "inverse(a*b) * a = inverse(b)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   983
    by (simp add: mult_assoc bnz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   984
  hence "inverse(a*b) * a * inverse(a) = inverse(b) * inverse(a)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   985
    by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   986
  thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   987
    by (simp add: mult_assoc anz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   988
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   989
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   990
text{*This version builds in division by zero while also re-orienting
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   991
      the right-hand side.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   992
lemma inverse_mult_distrib [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   993
     "inverse(a*b) = inverse(a) * inverse(b::'a::{field,division_by_zero})"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   994
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   995
    assume "a \<noteq> 0 & b \<noteq> 0" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   996
    thus ?thesis  by (simp add: nonzero_inverse_mult_distrib mult_commute)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   997
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   998
    assume "~ (a \<noteq> 0 & b \<noteq> 0)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   999
    thus ?thesis  by force
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1000
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1001
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1002
text{*There is no slick version using division by zero.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1003
lemma inverse_add:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1004
     "[|a \<noteq> 0;  b \<noteq> 0|]
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1005
      ==> inverse a + inverse b = (a+b) * inverse a * inverse (b::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1006
apply (simp add: left_distrib mult_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1007
apply (simp add: mult_commute [of "inverse a"]) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1008
apply (simp add: mult_assoc [symmetric] add_commute)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1009
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1010
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1011
lemma inverse_divide [simp]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1012
      "inverse (a/b) = b / (a::'a::{field,division_by_zero})"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1013
  by (simp add: divide_inverse mult_commute)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1014
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1015
lemma nonzero_mult_divide_cancel_left:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1016
  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
  1017
    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
  1018
proof -
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1019
  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
  1020
    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
  1021
                  nonzero_inverse_mult_distrib)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1022
  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
  1023
    by (simp only: mult_ac)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1024
  also have "... =  a * inverse b"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1025
    by simp
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1026
    finally show ?thesis 
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1027
    by (simp add: divide_inverse)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1028
qed
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1029
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1030
lemma mult_divide_cancel_left:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1031
     "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
  1032
apply (case_tac "b = 0")
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1033
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
  1034
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1035
14321
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
  1036
lemma nonzero_mult_divide_cancel_right:
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
  1037
     "[|b\<noteq>0; c\<noteq>0|] ==> (a*c) / (b*c) = a/(b::'a::field)"
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
  1038
by (simp add: mult_commute [of _ c] nonzero_mult_divide_cancel_left) 
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
  1039
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
  1040
lemma mult_divide_cancel_right:
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
  1041
     "c\<noteq>0 ==> (a*c) / (b*c) = a / (b::'a::{field,division_by_zero})"
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
  1042
apply (case_tac "b = 0")
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
  1043
apply (simp_all add: nonzero_mult_divide_cancel_right)
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
  1044
done
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
  1045
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1046
(*For ExtractCommonTerm*)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1047
lemma mult_divide_cancel_eq_if:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1048
     "(c*a) / (c*b) = 
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1049
      (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
  1050
  by (simp add: mult_divide_cancel_left)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1051
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1052
lemma divide_1 [simp]: "a/1 = (a::'a::field)"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1053
  by (simp add: divide_inverse)
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1054
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1055
lemma times_divide_eq_right [simp]: "a * (b/c) = (a*b) / (c::'a::field)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1056
by (simp add: divide_inverse mult_assoc)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1057
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1058
lemma times_divide_eq_left: "(b/c) * a = (b*a) / (c::'a::field)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1059
by (simp add: divide_inverse mult_ac)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1060
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1061
lemma divide_divide_eq_right [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1062
     "a / (b/c) = (a*c) / (b::'a::{field,division_by_zero})"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1063
by (simp add: divide_inverse mult_ac)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1064
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1065
lemma divide_divide_eq_left [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1066
     "(a / b) / (c::'a::{field,division_by_zero}) = a / (b*c)"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1067
by (simp add: divide_inverse mult_assoc)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1068
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1069
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1070
subsection {* Division and Unary Minus *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1071
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1072
lemma nonzero_minus_divide_left: "b \<noteq> 0 ==> - (a/b) = (-a) / (b::'a::field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1073
by (simp add: divide_inverse minus_mult_left)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1074
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1075
lemma nonzero_minus_divide_right: "b \<noteq> 0 ==> - (a/b) = a / -(b::'a::field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1076
by (simp add: divide_inverse nonzero_inverse_minus_eq minus_mult_right)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1077
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1078
lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a)/(-b) = a / (b::'a::field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1079
by (simp add: divide_inverse nonzero_inverse_minus_eq)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1080
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1081
lemma minus_divide_left: "- (a/b) = (-a) / (b::'a::field)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1082
by (simp add: divide_inverse minus_mult_left [symmetric])
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1083
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1084
lemma minus_divide_right: "- (a/b) = a / -(b::'a::{field,division_by_zero})"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1085
by (simp add: divide_inverse minus_mult_right [symmetric])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1086
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1087
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1088
text{*The effect is to extract signs from divisions*}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1089
declare minus_divide_left  [symmetric, simp]
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1090
declare minus_divide_right [symmetric, simp]
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1091
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
  1092
text{*Also, extract signs from products*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
  1093
declare minus_mult_left [symmetric, simp]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
  1094
declare minus_mult_right [symmetric, simp]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
  1095
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1096
lemma minus_divide_divide [simp]:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1097
     "(-a)/(-b) = a / (b::'a::{field,division_by_zero})"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1098
apply (case_tac "b=0", simp) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1099
apply (simp add: nonzero_minus_divide_divide) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1100
done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1101
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1102
lemma diff_divide_distrib: "(a-b)/(c::'a::field) = a/c - b/c"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
  1103
by (simp add: diff_minus add_divide_distrib) 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
  1104
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1105
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1106
subsection {* Ordered Fields *}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1107
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1108
lemma positive_imp_inverse_positive: 
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  1109
      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
  1110
  proof -
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1111
  have "0 < a * inverse a" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1112
    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
  1113
  thus "0 < inverse a" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1114
    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
  1115
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1116
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1117
lemma negative_imp_inverse_negative:
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1118
     "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
  1119
  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
  1120
      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
  1121
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1122
lemma inverse_le_imp_le:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  1123
      assumes invle: "inverse a \<le> inverse b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  1124
	  and apos:  "0 < a"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  1125
	 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
  1126
  proof (rule classical)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1127
  assume "~ b \<le> a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1128
  hence "a < b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1129
    by (simp add: linorder_not_le)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1130
  hence bpos: "0 < b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1131
    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
  1132
  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
  1133
    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
  1134
  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
  1135
    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
  1136
  thus "b \<le> a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1137
    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
  1138
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1139
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1140
lemma inverse_positive_imp_positive:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1141
      assumes inv_gt_0: "0 < inverse a"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1142
          and [simp]:   "a \<noteq> 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1143
        shows "0 < (a::'a::ordered_field)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1144
  proof -
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1145
  have "0 < inverse (inverse a)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1146
    by (rule positive_imp_inverse_positive)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1147
  thus "0 < a"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1148
    by (simp add: nonzero_inverse_inverse_eq)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1149
  qed
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1150
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1151
lemma inverse_positive_iff_positive [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1152
      "(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
  1153
apply (case_tac "a = 0", simp)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1154
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
  1155
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1156
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1157
lemma inverse_negative_imp_negative:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1158
      assumes inv_less_0: "inverse a < 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1159
          and [simp]:   "a \<noteq> 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1160
        shows "a < (0::'a::ordered_field)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1161
  proof -
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1162
  have "inverse (inverse a) < 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1163
    by (rule negative_imp_inverse_negative)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1164
  thus "a < 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1165
    by (simp add: nonzero_inverse_inverse_eq)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1166
  qed
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1167
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1168
lemma inverse_negative_iff_negative [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1169
      "(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
  1170
apply (case_tac "a = 0", simp)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1171
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
  1172
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1173
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1174
lemma inverse_nonnegative_iff_nonnegative [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1175
      "(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
  1176
by (simp add: linorder_not_less [symmetric])
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1177
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1178
lemma inverse_nonpositive_iff_nonpositive [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1179
      "(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
  1180
by (simp add: linorder_not_less [symmetric])
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1181
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1182
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1183
subsection{*Anti-Monotonicity of @{term inverse}*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1184
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1185
lemma less_imp_inverse_less:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  1186
      assumes less: "a < b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  1187
	  and apos:  "0 < a"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  1188
	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
  1189
  proof (rule ccontr)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1190
  assume "~ inverse b < inverse a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1191
  hence "inverse a \<le> inverse b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1192
    by (simp add: linorder_not_less)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1193
  hence "~ (a < b)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1194
    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
  1195
  thus False
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1196
    by (rule notE [OF _ less])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1197
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1198
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1199
lemma inverse_less_imp_less:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1200
   "[|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
  1201
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
  1202
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
  1203
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1204
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1205
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
  1206
lemma inverse_less_iff_less [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1207
     "[|0 < a; 0 < b|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1208
      ==> (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
  1209
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
  1210
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1211
lemma le_imp_inverse_le:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1212
   "[|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
  1213
  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
  1214
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1215
lemma inverse_le_iff_le [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1216
     "[|0 < a; 0 < b|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1217
      ==> (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
  1218
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
  1219
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1220
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1221
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
  1222
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
  1223
lemma inverse_le_imp_le_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1224
   "[|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
  1225
  apply (rule classical) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1226
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1227
   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
  1228
  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
  1229
  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
  1230
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1231
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1232
lemma less_imp_inverse_less_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1233
   "[|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
  1234
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1235
   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
  1236
  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
  1237
  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
  1238
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1239
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1240
lemma inverse_less_imp_less_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1241
   "[|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
  1242
  apply (rule classical) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1243
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1244
   prefer 2
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1245
   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
  1246
  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
  1247
  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
  1248
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1249
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1250
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
  1251
     "[|a < 0; b < 0|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1252
      ==> (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
  1253
  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
  1254
  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
  1255
	      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
  1256
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1257
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1258
lemma le_imp_inverse_le_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1259
   "[|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
  1260
  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
  1261
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1262
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
  1263
     "[|a < 0; b < 0|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1264
      ==> (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
  1265
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
  1266
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1267
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1268
subsection{*Inverses and the Number One*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1269
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1270
lemma one_less_inverse_iff:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1271
    "(1 < inverse x) = (0 < x & x < (1::'a::{ordered_field,division_by_zero}))"proof cases
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1272
  assume "0 < x"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1273
    with inverse_less_iff_less [OF zero_less_one, of x]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1274
    show ?thesis by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1275
next
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1276
  assume notless: "~ (0 < x)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1277
  have "~ (1 < inverse x)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1278
  proof
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1279
    assume "1 < inverse x"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1280
    also with notless have "... \<le> 0" by (simp add: linorder_not_less)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1281
    also have "... < 1" by (rule zero_less_one) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1282
    finally show False by auto
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1283
  qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1284
  with notless show ?thesis by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1285
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1286
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1287
lemma inverse_eq_1_iff [simp]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1288
    "(inverse x = 1) = (x = (1::'a::{field,division_by_zero}))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1289
by (insert inverse_eq_iff_eq [of x 1], simp) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1290
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1291
lemma one_le_inverse_iff:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1292
   "(1 \<le> inverse x) = (0 < x & x \<le> (1::'a::{ordered_field,division_by_zero}))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1293
by (force simp add: order_le_less one_less_inverse_iff zero_less_one 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1294
                    eq_commute [of 1]) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1295
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1296
lemma inverse_less_1_iff:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1297
   "(inverse x < 1) = (x \<le> 0 | 1 < (x::'a::{ordered_field,division_by_zero}))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1298
by (simp add: linorder_not_le [symmetric] one_le_inverse_iff) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1299
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1300
lemma inverse_le_1_iff:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1301
   "(inverse x \<le> 1) = (x \<le> 0 | 1 \<le> (x::'a::{ordered_field,division_by_zero}))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1302
by (simp add: linorder_not_less [symmetric] one_less_inverse_iff) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1303
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1304
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1305
subsection{*Division and Signs*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1306
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1307
lemma zero_less_divide_iff:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1308
     "((0::'a::{ordered_field,division_by_zero}) < a/b) = (0 < a & 0 < b | a < 0 & b < 0)"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1309
by (simp add: divide_inverse zero_less_mult_iff)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1310
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1311
lemma divide_less_0_iff:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1312
     "(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
  1313
      (0 < a & b < 0 | a < 0 & 0 < b)"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1314
by (simp add: divide_inverse mult_less_0_iff)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1315
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1316
lemma zero_le_divide_iff:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1317
     "((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
  1318
      (0 \<le> a & 0 \<le> b | a \<le> 0 & b \<le> 0)"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1319
by (simp add: divide_inverse zero_le_mult_iff)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1320
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1321
lemma divide_le_0_iff:
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1322
     "(a/b \<le> (0::'a::{ordered_field,division_by_zero})) =
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1323
      (0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1324
by (simp add: divide_inverse mult_le_0_iff)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1325
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1326
lemma divide_eq_0_iff [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1327
     "(a/b = 0) = (a=0 | b=(0::'a::{field,division_by_zero}))"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1328
by (simp add: divide_inverse field_mult_eq_0_iff)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1329
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1330
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1331
subsection{*Simplification of Inequalities Involving Literal Divisors*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1332
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1333
lemma pos_le_divide_eq: "0 < (c::'a::ordered_field) ==> (a \<le> b/c) = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1334
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1335
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1336
  hence "(a \<le> b/c) = (a*c \<le> (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1337
    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1338
  also have "... = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1339
    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1340
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1341
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1342
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1343
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1344
lemma neg_le_divide_eq: "c < (0::'a::ordered_field) ==> (a \<le> b/c) = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1345
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1346
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1347
  hence "(a \<le> b/c) = ((b/c)*c \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1348
    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1349
  also have "... = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1350
    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1351
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1352
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1353
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1354
lemma le_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1355
  "(a \<le> b/c) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1356
   (if 0 < c then a*c \<le> b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1357
             else if c < 0 then b \<le> a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1358
             else  a \<le> (0::'a::{ordered_field,division_by_zero}))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1359
apply (case_tac "c=0", simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1360
apply (force simp add: pos_le_divide_eq neg_le_divide_eq linorder_neq_iff) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1361
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1362
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1363
lemma pos_divide_le_eq: "0 < (c::'a::ordered_field) ==> (b/c \<le> a) = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1364
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1365
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1366
  hence "(b/c \<le> a) = ((b/c)*c \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1367
    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1368
  also have "... = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1369
    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1370
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1371
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1372
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1373
lemma neg_divide_le_eq: "c < (0::'a::ordered_field) ==> (b/c \<le> a) = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1374
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1375
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1376
  hence "(b/c \<le> a) = (a*c \<le> (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1377
    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1378
  also have "... = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1379
    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1380
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1381
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1382
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1383
lemma divide_le_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1384
  "(b/c \<le> a) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1385
   (if 0 < c then b \<le> a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1386
             else if c < 0 then a*c \<le> b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1387
             else 0 \<le> (a::'a::{ordered_field,division_by_zero}))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1388
apply (case_tac "c=0", simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1389
apply (force simp add: pos_divide_le_eq neg_divide_le_eq linorder_neq_iff) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1390
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1391
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1392
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1393
lemma pos_less_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1394
     "0 < (c::'a::ordered_field) ==> (a < b/c) = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1395
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1396
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1397
  hence "(a < b/c) = (a*c < (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1398
    by (simp add: mult_less_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1399
  also have "... = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1400
    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1401
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1402
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1403
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1404
lemma neg_less_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1405
 "c < (0::'a::ordered_field) ==> (a < b/c) = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1406
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1407
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1408
  hence "(a < b/c) = ((b/c)*c < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1409
    by (simp add: mult_less_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1410
  also have "... = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1411
    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1412
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1413
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1414
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1415
lemma less_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1416
  "(a < b/c) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1417
   (if 0 < c then a*c < b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1418
             else if c < 0 then b < a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1419
             else  a < (0::'a::{ordered_field,division_by_zero}))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1420
apply (case_tac "c=0", simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1421
apply (force simp add: pos_less_divide_eq neg_less_divide_eq linorder_neq_iff) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1422
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1423
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1424
lemma pos_divide_less_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1425
     "0 < (c::'a::ordered_field) ==> (b/c < a) = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1426
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1427
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1428
  hence "(b/c < a) = ((b/c)*c < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1429
    by (simp add: mult_less_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1430
  also have "... = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1431
    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1432
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1433
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1434
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1435
lemma neg_divide_less_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1436
 "c < (0::'a::ordered_field) ==> (b/c < a) = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1437
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1438
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1439
  hence "(b/c < a) = (a*c < (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1440
    by (simp add: mult_less_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1441
  also have "... = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1442
    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1443
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1444
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1445
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1446
lemma divide_less_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1447
  "(b/c < a) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1448
   (if 0 < c then b < a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1449
             else if c < 0 then a*c < b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1450
             else 0 < (a::'a::{ordered_field,division_by_zero}))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1451
apply (case_tac "c=0", simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1452
apply (force simp add: pos_divide_less_eq neg_divide_less_eq linorder_neq_iff) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1453
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1454
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1455
lemma nonzero_eq_divide_eq: "c\<noteq>0 ==> ((a::'a::field) = b/c) = (a*c = b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1456
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1457
  assume [simp]: "c\<noteq>0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1458
  have "(a = b/c) = (a*c = (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1459
    by (simp add: field_mult_cancel_right)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1460
  also have "... = (a*c = b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1461
    by (simp add: divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1462
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1463
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1464
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1465
lemma eq_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1466
  "((a::'a::{field,division_by_zero}) = b/c) = (if c\<noteq>0 then a*c = b else a=0)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1467
by (simp add: nonzero_eq_divide_eq) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1468
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1469
lemma nonzero_divide_eq_eq: "c\<noteq>0 ==> (b/c = (a::'a::field)) = (b = a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1470
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1471
  assume [simp]: "c\<noteq>0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1472
  have "(b/c = a) = ((b/c)*c = a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1473
    by (simp add: field_mult_cancel_right)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1474
  also have "... = (b = a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1475
    by (simp add: divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1476
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1477
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1478
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1479
lemma divide_eq_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1480
  "(b/c = (a::'a::{field,division_by_zero})) = (if c\<noteq>0 then b = a*c else a=0)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1481
by (force simp add: nonzero_divide_eq_eq) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1482
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1483
subsection{*Cancellation Laws for Division*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1484
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1485
lemma divide_cancel_right [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1486
     "(a/c = b/c) = (c = 0 | a = (b::'a::{field,division_by_zero}))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1487
apply (case_tac "c=0", simp) 
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1488
apply (simp add: divide_inverse field_mult_cancel_right) 
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1489
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1490
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1491
lemma divide_cancel_left [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1492
     "(c/a = c/b) = (c = 0 | a = (b::'a::{field,division_by_zero}))" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1493
apply (case_tac "c=0", simp) 
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1494
apply (simp add: divide_inverse field_mult_cancel_left) 
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1495
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1496
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1497
subsection {* Division and the Number One *}
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1498
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1499
text{*Simplify expressions equated with 1*}
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1500
lemma divide_eq_1_iff [simp]:
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1501
     "(a/b = 1) = (b \<noteq> 0 & a = (b::'a::{field,division_by_zero}))"
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1502
apply (case_tac "b=0", simp) 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1503
apply (simp add: right_inverse_eq) 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1504
done
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1505
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1506
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1507
lemma one_eq_divide_iff [simp]:
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1508
     "(1 = a/b) = (b \<noteq> 0 & a = (b::'a::{field,division_by_zero}))"
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1509
by (simp add: eq_commute [of 1])  
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1510
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1511
lemma zero_eq_1_divide_iff [simp]:
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1512
     "((0::'a::{ordered_field,division_by_zero}) = 1/a) = (a = 0)"
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1513
apply (case_tac "a=0", simp) 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1514
apply (auto simp add: nonzero_eq_divide_eq) 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1515
done
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1516
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1517
lemma one_divide_eq_0_iff [simp]:
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1518
     "(1/a = (0::'a::{ordered_field,division_by_zero})) = (a = 0)"
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1519
apply (case_tac "a=0", simp) 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1520
apply (insert zero_neq_one [THEN not_sym]) 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1521
apply (auto simp add: nonzero_divide_eq_eq) 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1522
done
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1523
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1524
text{*Simplify expressions such as @{text "0 < 1/x"} to @{text "0 < x"}*}
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1525
declare zero_less_divide_iff [of "1", simp]
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1526
declare divide_less_0_iff [of "1", simp]
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1527
declare zero_le_divide_iff [of "1", simp]
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1528
declare divide_le_0_iff [of "1", simp]
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1529
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1530
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1531
subsection {* Ordering Rules for Division *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1532
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1533
lemma divide_strict_right_mono:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1534
     "[|a < b; 0 < c|] ==> a / c < b / (c::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1535
by (simp add: order_less_imp_not_eq2 divide_inverse mult_strict_right_mono 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1536
              positive_imp_inverse_positive) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1537
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1538
lemma divide_right_mono:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1539
     "[|a \<le> b; 0 \<le> c|] ==> a/c \<le> b/(c::'a::{ordered_field,division_by_zero})"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1540
  by (force simp add: divide_strict_right_mono order_le_less) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1541
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1542
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1543
text{*The last premise ensures that @{term a} and @{term b} 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1544
      have the same sign*}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1545
lemma divide_strict_left_mono:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1546
       "[|b < a; 0 < c; 0 < a*b|] ==> c / a < c / (b::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1547
by (force simp add: zero_less_mult_iff divide_inverse mult_strict_left_mono 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1548
      order_less_imp_not_eq order_less_imp_not_eq2  
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1549
      less_imp_inverse_less less_imp_inverse_less_neg) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1550
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1551
lemma divide_left_mono:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1552
     "[|b \<le> a; 0 \<le> c; 0 < a*b|] ==> c / a \<le> c / (b::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1553
  apply (subgoal_tac "a \<noteq> 0 & b \<noteq> 0") 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1554
   prefer 2 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1555
   apply (force simp add: zero_less_mult_iff order_less_imp_not_eq) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1556
  apply (case_tac "c=0", simp add: divide_inverse)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1557
  apply (force simp add: divide_strict_left_mono order_le_less) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1558
  done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1559
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1560
lemma divide_strict_left_mono_neg:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1561
     "[|a < b; c < 0; 0 < a*b|] ==> c / a < c / (b::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1562
  apply (subgoal_tac "a \<noteq> 0 & b \<noteq> 0") 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1563
   prefer 2 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1564
   apply (force simp add: zero_less_mult_iff order_less_imp_not_eq) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1565
  apply (drule divide_strict_left_mono [of _ _ "-c"]) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1566
   apply (simp_all add: mult_commute nonzero_minus_divide_left [symmetric]) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1567
  done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1568
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1569
lemma divide_strict_right_mono_neg:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1570
     "[|b < a; c < 0|] ==> a / c < b / (c::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1571
apply (drule divide_strict_right_mono [of _ _ "-c"], simp) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1572
apply (simp add: order_less_imp_not_eq nonzero_minus_divide_right [symmetric]) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1573
done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1574
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1575
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1576
subsection {* Ordered Fields are Dense *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1577
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1578
lemma less_add_one: "a < (a+1::'a::ordered_semiring)"
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1579
proof -
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1580
  have "a+0 < (a+1::'a::ordered_semiring)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1581
    by (blast intro: zero_less_one add_strict_left_mono) 
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1582
  thus ?thesis by simp
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1583
qed
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1584
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1585
lemma zero_less_two: "0 < (1+1::'a::ordered_semiring)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1586
  by (blast intro: order_less_trans zero_less_one less_add_one) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1587
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1588
lemma less_half_sum: "a < b ==> a < (a+b) / (1+1::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1589
by (simp add: zero_less_two pos_less_divide_eq right_distrib) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1590
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1591
lemma gt_half_sum: "a < b ==> (a+b)/(1+1::'a::ordered_field) < b"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1592
by (simp add: zero_less_two pos_divide_less_eq right_distrib) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1593
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1594
lemma dense: "a < b ==> \<exists>r::'a::ordered_field. a < r & r < b"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1595
by (blast intro!: less_half_sum gt_half_sum)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1596
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1597
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1598
subsection {* Absolute Value *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1599
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1600
lemma abs_zero [simp]: "abs 0 = (0::'a::ordered_ring)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1601
by (simp add: abs_if)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1602
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1603
lemma abs_one [simp]: "abs 1 = (1::'a::ordered_ring)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1604
  by (simp add: abs_if zero_less_one [THEN order_less_not_sym]) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1605
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1606
lemma abs_mult: "abs (a * b) = abs a * abs (b::'a::ordered_ring)" 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1607
apply (case_tac "a=0 | b=0", force) 
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1608
apply (auto elim: order_less_asym
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1609
            simp add: abs_if mult_less_0_iff linorder_neq_iff
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1610
                  minus_mult_left [symmetric] minus_mult_right [symmetric])  
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1611
done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1612
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
  1613
lemma abs_mult_self: "abs a * abs a = a * (a::'a::ordered_ring)"
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
  1614
by (simp add: abs_if) 
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
  1615
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1616
lemma abs_eq_0 [simp]: "(abs a = 0) = (a = (0::'a::ordered_ring))"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1617
by (simp add: abs_if)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1618
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1619
lemma zero_less_abs_iff [simp]: "(0 < abs a) = (a \<noteq> (0::'a::ordered_ring))"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1620
by (simp add: abs_if linorder_neq_iff)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1621
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1622
lemma abs_not_less_zero [simp]: "~ abs a < (0::'a::ordered_ring)"
14398
c5c47703f763 Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents: 14387
diff changeset
  1623
apply (simp add: abs_if)
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1624
by (simp add: abs_if  order_less_not_sym [of a 0])
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1625
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1626
lemma abs_le_zero_iff [simp]: "(abs a \<le> (0::'a::ordered_ring)) = (a = 0)" 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1627
by (simp add: order_le_less) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1628
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1629
lemma abs_minus_cancel [simp]: "abs (-a) = abs(a::'a::ordered_ring)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1630
apply (auto simp add: abs_if linorder_not_less order_less_not_sym [of 0 a])  
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1631
apply (drule order_antisym, assumption, simp) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1632
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1633
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1634
lemma abs_ge_zero [simp]: "(0::'a::ordered_ring) \<le> abs a"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1635
apply (simp add: abs_if order_less_imp_le)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1636
apply (simp add: linorder_not_less) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1637
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1638
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1639
lemma abs_idempotent [simp]: "abs (abs a) = abs (a::'a::ordered_ring)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1640
  by (force elim: order_less_asym simp add: abs_if)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1641
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14295
diff changeset
  1642
lemma abs_zero_iff [simp]: "(abs a = 0) = (a = (0::'a::ordered_ring))"
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1643
by (simp add: abs_if)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1644
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1645
lemma abs_ge_self: "a \<le> abs (a::'a::ordered_ring)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1646
apply (simp add: abs_if)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1647
apply (simp add: order_less_imp_le order_trans [of _ 0])
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1648
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1649
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1650
lemma abs_ge_minus_self: "-a \<le> abs (a::'a::ordered_ring)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1651
by (insert abs_ge_self [of "-a"], simp)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1652
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1653
lemma nonzero_abs_inverse:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1654
     "a \<noteq> 0 ==> abs (inverse (a::'a::ordered_field)) = inverse (abs a)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1655
apply (auto simp add: linorder_neq_iff abs_if nonzero_inverse_minus_eq 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1656
                      negative_imp_inverse_negative)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1657
apply (blast intro: positive_imp_inverse_positive elim: order_less_asym) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1658
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1659
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1660
lemma abs_inverse [simp]:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1661
     "abs (inverse (a::'a::{ordered_field,division_by_zero})) = 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1662
      inverse (abs a)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1663
apply (case_tac "a=0", simp) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1664
apply (simp add: nonzero_abs_inverse) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1665
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1666
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1667
lemma nonzero_abs_divide:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1668
     "b \<noteq> 0 ==> abs (a / (b::'a::ordered_field)) = abs a / abs b"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1669
by (simp add: divide_inverse abs_mult nonzero_abs_inverse) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1670
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1671
lemma abs_divide:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1672
     "abs (a / (b::'a::{ordered_field,division_by_zero})) = abs a / abs b"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1673
apply (case_tac "b=0", simp) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1674
apply (simp add: nonzero_abs_divide) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1675
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1676
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1677
lemma abs_leI: "[|a \<le> b; -a \<le> b|] ==> abs a \<le> (b::'a::ordered_ring)"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1678
by (simp add: abs_if)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1679
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1680
lemma le_minus_self_iff: "(a \<le> -a) = (a \<le> (0::'a::ordered_ring))"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1681
proof 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1682
  assume ale: "a \<le> -a"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1683
  show "a\<le>0"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1684
    apply (rule classical) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1685
    apply (simp add: linorder_not_le) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1686
    apply (blast intro: ale order_trans order_less_imp_le
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1687
                        neg_0_le_iff_le [THEN iffD1]) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1688
    done
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1689
next
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1690
  assume "a\<le>0"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1691
  hence "0 \<le> -a" by (simp only: neg_0_le_iff_le)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1692
  thus "a \<le> -a"  by (blast intro: prems order_trans) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1693
qed
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1694
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1695
lemma minus_le_self_iff: "(-a \<le> a) = (0 \<le> (a::'a::ordered_ring))"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1696
by (insert le_minus_self_iff [of "-a"], simp)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1697
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1698
lemma eq_minus_self_iff: "(a = -a) = (a = (0::'a::ordered_ring))"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1699
by (force simp add: order_eq_iff le_minus_self_iff minus_le_self_iff)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1700
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1701
lemma less_minus_self_iff: "(a < -a) = (a < (0::'a::ordered_ring))"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1702
by (simp add: order_less_le le_minus_self_iff eq_minus_self_iff)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1703
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1704
lemma abs_le_D1: "abs a \<le> b ==> a \<le> (b::'a::ordered_ring)"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1705
apply (simp add: abs_if split: split_if_asm)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1706
apply (rule order_trans [of _ "-a"]) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1707
 apply (simp add: less_minus_self_iff order_less_imp_le, assumption)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1708
done
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1709
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1710
lemma abs_le_D2: "abs a \<le> b ==> -a \<le> (b::'a::ordered_ring)"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1711
by (insert abs_le_D1 [of "-a"], simp)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1712
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1713
lemma abs_le_iff: "(abs a \<le> b) = (a \<le> b & -a \<le> (b::'a::ordered_ring))"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1714
by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2)
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1715
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1716
lemma abs_less_iff: "(abs a < b) = (a < b & -a < (b::'a::ordered_ring))" 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1717
apply (simp add: order_less_le abs_le_iff)  
14398
c5c47703f763 Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents: 14387
diff changeset
  1718
apply (auto simp add: abs_if minus_le_self_iff eq_minus_self_iff)
c5c47703f763 Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents: 14387
diff changeset
  1719
apply (simp add: le_minus_self_iff linorder_neq_iff) 
c5c47703f763 Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents: 14387
diff changeset
  1720
done
c5c47703f763 Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents: 14387
diff changeset
  1721
(*
c5c47703f763 Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents: 14387
diff changeset
  1722
apply (simp add: order_less_le abs_le_iff)  
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1723
apply (auto simp add: abs_if minus_le_self_iff eq_minus_self_iff) 
14398
c5c47703f763 Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents: 14387
diff changeset
  1724
 apply (simp add:  linorder_not_less [symmetric])
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1725
apply (simp add: le_minus_self_iff linorder_neq_iff) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1726
apply (simp add:  linorder_not_less [symmetric]) 
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1727
done
14398
c5c47703f763 Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents: 14387
diff changeset
  1728
*)
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1729
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1730
lemma abs_triangle_ineq: "abs (a+b) \<le> abs a + abs (b::'a::ordered_ring)"
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
  1731
by (force simp add: abs_le_iff abs_ge_self abs_ge_minus_self add_mono)
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1732
14475
aa973ba84f69 new thms
paulson
parents: 14430
diff changeset
  1733
lemma abs_diff_triangle_ineq:
aa973ba84f69 new thms
paulson
parents: 14430
diff changeset
  1734
     "\<bar>(a::'a::ordered_ring) + b - (c+d)\<bar> \<le> \<bar>a-c\<bar> + \<bar>b-d\<bar>"
aa973ba84f69 new thms
paulson
parents: 14430
diff changeset
  1735
proof -
aa973ba84f69 new thms
paulson
parents: 14430
diff changeset
  1736
  have "\<bar>a + b - (c+d)\<bar> = \<bar>(a-c) + (b-d)\<bar>" by (simp add: diff_minus add_ac)
aa973ba84f69 new thms
paulson
parents: 14430
diff changeset
  1737
  also have "... \<le> \<bar>a-c\<bar> + \<bar>b-d\<bar>" by (rule abs_triangle_ineq)
aa973ba84f69 new thms
paulson
parents: 14430
diff changeset
  1738
  finally show ?thesis .
aa973ba84f69 new thms
paulson
parents: 14430
diff changeset
  1739
qed
aa973ba84f69 new thms
paulson
parents: 14430
diff changeset
  1740
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1741
lemma abs_mult_less:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1742
     "[| abs a < c; abs b < d |] ==> abs a * abs b < c*(d::'a::ordered_ring)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1743
proof -
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1744
  assume ac: "abs a < c"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1745
  hence cpos: "0<c" by (blast intro: order_le_less_trans abs_ge_zero)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1746
  assume "abs b < d"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1747
  thus ?thesis by (simp add: ac cpos mult_strict_mono) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1748
qed
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1749
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1750
text{*Moving this up spoils many proofs using @{text mult_le_cancel_right}*}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1751
declare times_divide_eq_left [simp]
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1752
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1753
ML
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1754
{*
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1755
val add_assoc = thm"add_assoc";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1756
val add_commute = thm"add_commute";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1757
val mult_assoc = thm"mult_assoc";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1758
val mult_commute = thm"mult_commute";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1759
val zero_neq_one = thm"zero_neq_one";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1760
val diff_minus = thm"diff_minus";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1761
val abs_if = thm"abs_if";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1762
val divide_inverse = thm"divide_inverse";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1763
val inverse_zero = thm"inverse_zero";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1764
val divide_zero = thm"divide_zero";
14368
2763da611ad9 converted Real/Lubs to Isar script. Converting arithmetic setup
paulson
parents: 14365
diff changeset
  1765
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1766
val add_0 = thm"add_0";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1767
val add_0_right = thm"add_0_right";
14368
2763da611ad9 converted Real/Lubs to Isar script. Converting arithmetic setup
paulson
parents: 14365
diff changeset
  1768
val add_zero_left = thm"add_0";
2763da611ad9 converted Real/Lubs to Isar script. Converting arithmetic setup
paulson
parents: 14365
diff changeset
  1769
val add_zero_right = thm"add_0_right";
2763da611ad9 converted Real/Lubs to Isar script. Converting arithmetic setup
paulson
parents: 14365
diff changeset
  1770
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1771
val add_left_commute = thm"add_left_commute";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1772
val left_minus = thm"left_minus";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1773
val right_minus = thm"right_minus";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1774
val right_minus_eq = thm"right_minus_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1775
val add_left_cancel = thm"add_left_cancel";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1776
val add_right_cancel = thm"add_right_cancel";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1777
val minus_minus = thm"minus_minus";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1778
val equals_zero_I = thm"equals_zero_I";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1779
val minus_zero = thm"minus_zero";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1780
val diff_self = thm"diff_self";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1781
val diff_0 = thm"diff_0";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1782
val diff_0_right = thm"diff_0_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1783
val diff_minus_eq_add = thm"diff_minus_eq_add";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1784
val neg_equal_iff_equal = thm"neg_equal_iff_equal";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1785
val neg_equal_0_iff_equal = thm"neg_equal_0_iff_equal";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1786
val neg_0_equal_iff_equal = thm"neg_0_equal_iff_equal";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1787
val equation_minus_iff = thm"equation_minus_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1788
val minus_equation_iff = thm"minus_equation_iff";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1789
val mult_1 = thm"mult_1";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1790
val mult_1_right = thm"mult_1_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1791
val mult_left_commute = thm"mult_left_commute";
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1792
val mult_zero_left = thm"mult_zero_left";
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1793
val mult_zero_right = thm"mult_zero_right";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1794
val left_distrib = thm "left_distrib";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1795
val right_distrib = thm"right_distrib";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1796
val combine_common_factor = thm"combine_common_factor";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1797
val minus_add_distrib = thm"minus_add_distrib";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1798
val minus_mult_left = thm"minus_mult_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1799
val minus_mult_right = thm"minus_mult_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1800
val minus_mult_minus = thm"minus_mult_minus";
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1801
val minus_mult_commute = thm"minus_mult_commute";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1802
val right_diff_distrib = thm"right_diff_distrib";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1803
val left_diff_distrib = thm"left_diff_distrib";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1804
val minus_diff_eq = thm"minus_diff_eq";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1805
val add_left_mono = thm"add_left_mono";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1806
val add_right_mono = thm"add_right_mono";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1807
val add_mono = thm"add_mono";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1808
val add_strict_left_mono = thm"add_strict_left_mono";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1809
val add_strict_right_mono = thm"add_strict_right_mono";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1810
val add_strict_mono = thm"add_strict_mono";
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
  1811
val add_less_le_mono = thm"add_less_le_mono";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
  1812
val add_le_less_mono = thm"add_le_less_mono";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1813
val add_less_imp_less_left = thm"add_less_imp_less_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1814
val add_less_imp_less_right = thm"add_less_imp_less_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1815
val add_less_cancel_left = thm"add_less_cancel_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1816
val add_less_cancel_right = thm"add_less_cancel_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1817
val add_le_cancel_left = thm"add_le_cancel_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1818
val add_le_cancel_right = thm"add_le_cancel_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1819
val add_le_imp_le_left = thm"add_le_imp_le_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1820
val add_le_imp_le_right = thm"add_le_imp_le_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1821
val le_imp_neg_le = thm"le_imp_neg_le";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1822
val neg_le_iff_le = thm"neg_le_iff_le";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1823
val neg_le_0_iff_le = thm"neg_le_0_iff_le";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1824
val neg_0_le_iff_le = thm"neg_0_le_iff_le";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1825
val neg_less_iff_less = thm"neg_less_iff_less";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1826
val neg_less_0_iff_less = thm"neg_less_0_iff_less";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1827
val neg_0_less_iff_less = thm"neg_0_less_iff_less";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1828
val less_minus_iff = thm"less_minus_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1829
val minus_less_iff = thm"minus_less_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1830
val le_minus_iff = thm"le_minus_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1831
val minus_le_iff = thm"minus_le_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1832
val add_diff_eq = thm"add_diff_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1833
val diff_add_eq = thm"diff_add_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1834
val diff_eq_eq = thm"diff_eq_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1835
val eq_diff_eq = thm"eq_diff_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1836
val diff_diff_eq = thm"diff_diff_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1837
val diff_diff_eq2 = thm"diff_diff_eq2";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1838
val less_iff_diff_less_0 = thm"less_iff_diff_less_0";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1839
val diff_less_eq = thm"diff_less_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1840
val less_diff_eq = thm"less_diff_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1841
val diff_le_eq = thm"diff_le_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1842
val le_diff_eq = thm"le_diff_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1843
val eq_iff_diff_eq_0 = thm"eq_iff_diff_eq_0";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1844
val le_iff_diff_le_0 = thm"le_iff_diff_le_0";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1845
val eq_add_iff1 = thm"eq_add_iff1";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1846
val eq_add_iff2 = thm"eq_add_iff2";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1847
val less_add_iff1 = thm"less_add_iff1";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1848
val less_add_iff2 = thm"less_add_iff2";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1849
val le_add_iff1 = thm"le_add_iff1";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1850
val le_add_iff2 = thm"le_add_iff2";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1851
val mult_strict_left_mono = thm"mult_strict_left_mono";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1852
val mult_strict_right_mono = thm"mult_strict_right_mono";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1853
val mult_left_mono = thm"mult_left_mono";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1854
val mult_right_mono = thm"mult_right_mono";
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
  1855
val mult_left_le_imp_le = thm"mult_left_le_imp_le";
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
  1856
val mult_right_le_imp_le = thm"mult_right_le_imp_le";
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
  1857
val mult_left_less_imp_less = thm"mult_left_less_imp_less";
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
  1858
val mult_right_less_imp_less = thm"mult_right_less_imp_less";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1859
val mult_strict_left_mono_neg = thm"mult_strict_left_mono_neg";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1860
val mult_strict_right_mono_neg = thm"mult_strict_right_mono_neg";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1861
val mult_pos = thm"mult_pos";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1862
val mult_pos_neg = thm"mult_pos_neg";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1863
val mult_neg = thm"mult_neg";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1864
val zero_less_mult_pos = thm"zero_less_mult_pos";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1865
val zero_less_mult_iff = thm"zero_less_mult_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1866
val mult_eq_0_iff = thm"mult_eq_0_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1867
val zero_le_mult_iff = thm"zero_le_mult_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1868
val mult_less_0_iff = thm"mult_less_0_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1869
val mult_le_0_iff = thm"mult_le_0_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1870
val zero_le_square = thm"zero_le_square";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1871
val zero_less_one = thm"zero_less_one";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1872
val zero_le_one = thm"zero_le_one";
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
  1873
val not_one_less_zero = thm"not_one_less_zero";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
  1874
val not_one_le_zero = thm"not_one_le_zero";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1875
val mult_left_mono_neg = thm"mult_left_mono_neg";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1876
val mult_right_mono_neg = thm"mult_right_mono_neg";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1877
val mult_strict_mono = thm"mult_strict_mono";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1878
val mult_strict_mono' = thm"mult_strict_mono'";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1879
val mult_mono = thm"mult_mono";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1880
val mult_less_cancel_right = thm"mult_less_cancel_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1881
val mult_less_cancel_left = thm"mult_less_cancel_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1882
val mult_le_cancel_right = thm"mult_le_cancel_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1883
val mult_le_cancel_left = thm"mult_le_cancel_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1884
val mult_less_imp_less_left = thm"mult_less_imp_less_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1885
val mult_less_imp_less_right = thm"mult_less_imp_less_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1886
val mult_cancel_right = thm"mult_cancel_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1887
val mult_cancel_left = thm"mult_cancel_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1888
val left_inverse = thm "left_inverse";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1889
val right_inverse = thm"right_inverse";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1890
val right_inverse_eq = thm"right_inverse_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1891
val nonzero_inverse_eq_divide = thm"nonzero_inverse_eq_divide";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1892
val divide_self = thm"divide_self";
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1893
val inverse_divide = thm"inverse_divide";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1894
val divide_zero_left = thm"divide_zero_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1895
val inverse_eq_divide = thm"inverse_eq_divide";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1896
val add_divide_distrib = thm"add_divide_distrib";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1897
val field_mult_eq_0_iff = thm"field_mult_eq_0_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1898
val field_mult_cancel_right = thm"field_mult_cancel_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1899
val field_mult_cancel_left = thm"field_mult_cancel_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1900
val nonzero_imp_inverse_nonzero = thm"nonzero_imp_inverse_nonzero";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1901
val inverse_zero_imp_zero = thm"inverse_zero_imp_zero";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1902
val inverse_nonzero_imp_nonzero = thm"inverse_nonzero_imp_nonzero";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1903
val inverse_nonzero_iff_nonzero = thm"inverse_nonzero_iff_nonzero";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1904
val nonzero_inverse_minus_eq = thm"nonzero_inverse_minus_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1905
val inverse_minus_eq = thm"inverse_minus_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1906
val nonzero_inverse_eq_imp_eq = thm"nonzero_inverse_eq_imp_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1907
val inverse_eq_imp_eq = thm"inverse_eq_imp_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1908
val inverse_eq_iff_eq = thm"inverse_eq_iff_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1909
val nonzero_inverse_inverse_eq = thm"nonzero_inverse_inverse_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1910
val inverse_inverse_eq = thm"inverse_inverse_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1911
val inverse_1 = thm"inverse_1";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1912
val nonzero_inverse_mult_distrib = thm"nonzero_inverse_mult_distrib";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1913
val inverse_mult_distrib = thm"inverse_mult_distrib";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1914
val inverse_add = thm"inverse_add";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1915
val nonzero_mult_divide_cancel_left = thm"nonzero_mult_divide_cancel_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1916
val mult_divide_cancel_left = thm"mult_divide_cancel_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1917
val nonzero_mult_divide_cancel_right = thm"nonzero_mult_divide_cancel_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1918
val mult_divide_cancel_right = thm"mult_divide_cancel_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1919
val mult_divide_cancel_eq_if = thm"mult_divide_cancel_eq_if";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1920
val divide_1 = thm"divide_1";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1921
val times_divide_eq_right = thm"times_divide_eq_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1922
val times_divide_eq_left = thm"times_divide_eq_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1923
val divide_divide_eq_right = thm"divide_divide_eq_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1924
val divide_divide_eq_left = thm"divide_divide_eq_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1925
val nonzero_minus_divide_left = thm"nonzero_minus_divide_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1926
val nonzero_minus_divide_right = thm"nonzero_minus_divide_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1927
val nonzero_minus_divide_divide = thm"nonzero_minus_divide_divide";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1928
val minus_divide_left = thm"minus_divide_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1929
val minus_divide_right = thm"minus_divide_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1930
val minus_divide_divide = thm"minus_divide_divide";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1931
val positive_imp_inverse_positive = thm"positive_imp_inverse_positive";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1932
val negative_imp_inverse_negative = thm"negative_imp_inverse_negative";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1933
val inverse_le_imp_le = thm"inverse_le_imp_le";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1934
val inverse_positive_imp_positive = thm"inverse_positive_imp_positive";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1935
val inverse_positive_iff_positive = thm"inverse_positive_iff_positive";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1936
val inverse_negative_imp_negative = thm"inverse_negative_imp_negative";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1937
val inverse_negative_iff_negative = thm"inverse_negative_iff_negative";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1938
val inverse_nonnegative_iff_nonnegative = thm"inverse_nonnegative_iff_nonnegative";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1939
val inverse_nonpositive_iff_nonpositive = thm"inverse_nonpositive_iff_nonpositive";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1940
val less_imp_inverse_less = thm"less_imp_inverse_less";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1941
val inverse_less_imp_less = thm"inverse_less_imp_less";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1942
val inverse_less_iff_less = thm"inverse_less_iff_less";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1943
val le_imp_inverse_le = thm"le_imp_inverse_le";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1944
val inverse_le_iff_le = thm"inverse_le_iff_le";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1945
val inverse_le_imp_le_neg = thm"inverse_le_imp_le_neg";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1946
val less_imp_inverse_less_neg = thm"less_imp_inverse_less_neg";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1947
val inverse_less_imp_less_neg = thm"inverse_less_imp_less_neg";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1948
val inverse_less_iff_less_neg = thm"inverse_less_iff_less_neg";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1949
val le_imp_inverse_le_neg = thm"le_imp_inverse_le_neg";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1950
val inverse_le_iff_le_neg = thm"inverse_le_iff_le_neg";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1951
val zero_less_divide_iff = thm"zero_less_divide_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1952
val divide_less_0_iff = thm"divide_less_0_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1953
val zero_le_divide_iff = thm"zero_le_divide_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1954
val divide_le_0_iff = thm"divide_le_0_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1955
val divide_eq_0_iff = thm"divide_eq_0_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1956
val pos_le_divide_eq = thm"pos_le_divide_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1957
val neg_le_divide_eq = thm"neg_le_divide_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1958
val le_divide_eq = thm"le_divide_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1959
val pos_divide_le_eq = thm"pos_divide_le_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1960
val neg_divide_le_eq = thm"neg_divide_le_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1961
val divide_le_eq = thm"divide_le_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1962
val pos_less_divide_eq = thm"pos_less_divide_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1963
val neg_less_divide_eq = thm"neg_less_divide_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1964
val less_divide_eq = thm"less_divide_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1965
val pos_divide_less_eq = thm"pos_divide_less_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1966
val neg_divide_less_eq = thm"neg_divide_less_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1967
val divide_less_eq = thm"divide_less_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1968
val nonzero_eq_divide_eq = thm"nonzero_eq_divide_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1969
val eq_divide_eq = thm"eq_divide_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1970
val nonzero_divide_eq_eq = thm"nonzero_divide_eq_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1971
val divide_eq_eq = thm"divide_eq_eq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1972
val divide_cancel_right = thm"divide_cancel_right";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1973
val divide_cancel_left = thm"divide_cancel_left";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1974
val divide_strict_right_mono = thm"divide_strict_right_mono";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1975
val divide_right_mono = thm"divide_right_mono";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1976
val divide_strict_left_mono = thm"divide_strict_left_mono";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1977
val divide_left_mono = thm"divide_left_mono";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1978
val divide_strict_left_mono_neg = thm"divide_strict_left_mono_neg";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1979
val divide_strict_right_mono_neg = thm"divide_strict_right_mono_neg";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1980
val zero_less_two = thm"zero_less_two";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1981
val less_half_sum = thm"less_half_sum";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1982
val gt_half_sum = thm"gt_half_sum";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1983
val dense = thm"dense";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1984
val abs_zero = thm"abs_zero";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1985
val abs_one = thm"abs_one";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1986
val abs_mult = thm"abs_mult";
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
  1987
val abs_mult_self = thm"abs_mult_self";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1988
val abs_eq_0 = thm"abs_eq_0";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1989
val zero_less_abs_iff = thm"zero_less_abs_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1990
val abs_not_less_zero = thm"abs_not_less_zero";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1991
val abs_le_zero_iff = thm"abs_le_zero_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1992
val abs_minus_cancel = thm"abs_minus_cancel";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1993
val abs_ge_zero = thm"abs_ge_zero";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1994
val abs_idempotent = thm"abs_idempotent";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1995
val abs_zero_iff = thm"abs_zero_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1996
val abs_ge_self = thm"abs_ge_self";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1997
val abs_ge_minus_self = thm"abs_ge_minus_self";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1998
val nonzero_abs_inverse = thm"nonzero_abs_inverse";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1999
val abs_inverse = thm"abs_inverse";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2000
val nonzero_abs_divide = thm"nonzero_abs_divide";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2001
val abs_divide = thm"abs_divide";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2002
val abs_leI = thm"abs_leI";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2003
val le_minus_self_iff = thm"le_minus_self_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2004
val minus_le_self_iff = thm"minus_le_self_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2005
val eq_minus_self_iff = thm"eq_minus_self_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2006
val less_minus_self_iff = thm"less_minus_self_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2007
val abs_le_D1 = thm"abs_le_D1";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2008
val abs_le_D2 = thm"abs_le_D2";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2009
val abs_le_iff = thm"abs_le_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2010
val abs_less_iff = thm"abs_less_iff";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2011
val abs_triangle_ineq = thm"abs_triangle_ineq";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2012
val abs_mult_less = thm"abs_mult_less";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2013
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2014
val compare_rls = thms"compare_rls";
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2015
*}
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  2016
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  2017
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  2018
end