src/HOL/Ring_and_Field.thy
author nipkow
Fri, 10 Sep 2004 20:04:14 +0200
changeset 15197 19e735596e51
parent 15178 5f621aa35c25
child 15228 4d332d10fa3d
permissions -rw-r--r--
Added antisymmetry simproc
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$
14770
fe9504ba63d5 removed duplicate thms;
wenzelm
parents: 14754
diff changeset
     3
    Author:  Gertrud Bauer, Steven Obua, Lawrence C Paulson and Markus Wenzel
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     4
*)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     5
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
     6
header {* (Ordered) Rings and Fields *}
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     7
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15077
diff changeset
     8
theory Ring_and_Field 
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
     9
imports OrderedGroup
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15077
diff changeset
    10
begin
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    11
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    12
text {*
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    13
  The theory of partially ordered rings is taken from the books:
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    14
  \begin{itemize}
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    15
  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    16
  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    17
  \end{itemize}
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    18
  Most of the used notions can also be looked up in 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    19
  \begin{itemize}
14770
fe9504ba63d5 removed duplicate thms;
wenzelm
parents: 14754
diff changeset
    20
  \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    21
  \item \emph{Algebra I} by van der Waerden, Springer.
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    22
  \end{itemize}
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    23
*}
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    24
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    25
axclass semiring \<subseteq> ab_semigroup_add, semigroup_mult
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    26
  left_distrib: "(a + b) * c = a * c + b * c"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    27
  right_distrib: "a * (b + c) = a * b + a * c"
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    28
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    29
axclass semiring_0 \<subseteq> semiring, comm_monoid_add
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    30
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    31
axclass semiring_0_cancel \<subseteq> semiring_0, cancel_ab_semigroup_add
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    32
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    33
axclass comm_semiring \<subseteq> ab_semigroup_add, ab_semigroup_mult  
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    34
  mult_commute: "a * b = b * a"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    35
  distrib: "(a + b) * c = a * c + b * c"
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    36
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    37
instance comm_semiring \<subseteq> semiring
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    38
proof
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    39
  fix a b c :: 'a
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    40
  show "(a + b) * c = a * c + b * c" by (simp add: distrib)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    41
  have "a * (b + c) = (b + c) * a" by (simp add: mult_ac)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    42
  also have "... = b * a + c * a" by (simp only: distrib)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    43
  also have "... = a * b + a * c" by (simp add: mult_ac)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    44
  finally show "a * (b + c) = a * b + a * c" by blast
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    45
qed
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    46
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    47
axclass comm_semiring_0 \<subseteq> comm_semiring, comm_monoid_add
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    48
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    49
instance comm_semiring_0 \<subseteq> semiring_0 ..
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    50
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    51
axclass comm_semiring_0_cancel \<subseteq> comm_semiring_0, cancel_ab_semigroup_add
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    52
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    53
instance comm_semiring_0_cancel \<subseteq> semiring_0_cancel ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    54
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    55
axclass axclass_0_neq_1 \<subseteq> zero, one
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    56
  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
    57
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    58
axclass semiring_1 \<subseteq> axclass_0_neq_1, semiring_0, monoid_mult
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    59
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    60
axclass comm_semiring_1 \<subseteq> axclass_0_neq_1, comm_semiring_0, comm_monoid_mult (* previously almost_semiring *)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    61
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    62
instance comm_semiring_1 \<subseteq> semiring_1 ..
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
    63
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    64
axclass axclass_no_zero_divisors \<subseteq> zero, times
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    65
  no_zero_divisors: "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a * b \<noteq> 0"
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    66
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    67
axclass semiring_1_cancel \<subseteq> semiring_1, cancel_ab_semigroup_add
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    68
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    69
instance semiring_1_cancel \<subseteq> semiring_0_cancel ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    70
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    71
axclass comm_semiring_1_cancel \<subseteq> comm_semiring_1, cancel_ab_semigroup_add (* previously semiring *)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    72
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    73
instance comm_semiring_1_cancel \<subseteq> semiring_1_cancel ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    74
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    75
instance comm_semiring_1_cancel \<subseteq> comm_semiring_0_cancel ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    76
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    77
axclass ring \<subseteq> semiring, ab_group_add
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    78
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    79
instance ring \<subseteq> semiring_0_cancel ..
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    80
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    81
axclass comm_ring \<subseteq> comm_semiring_0, ab_group_add
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    82
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    83
instance comm_ring \<subseteq> ring ..
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    84
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    85
instance comm_ring \<subseteq> comm_semiring_0_cancel ..
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    86
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    87
axclass ring_1 \<subseteq> ring, semiring_1
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    88
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    89
instance ring_1 \<subseteq> semiring_1_cancel ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    90
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    91
axclass comm_ring_1 \<subseteq> comm_ring, comm_semiring_1 (* previously ring *)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    92
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    93
instance comm_ring_1 \<subseteq> ring_1 ..
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
    94
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    95
instance comm_ring_1 \<subseteq> comm_semiring_1_cancel ..
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    96
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    97
axclass idom \<subseteq> comm_ring_1, axclass_no_zero_divisors
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
    98
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    99
axclass field \<subseteq> comm_ring_1, inverse
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   100
  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
   101
  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
   102
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   103
lemma mult_zero_left [simp]: "0 * a = (0::'a::semiring_0_cancel)"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   104
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   105
  have "0*a + 0*a = 0*a + 0"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   106
    by (simp add: left_distrib [symmetric])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   107
  thus ?thesis 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   108
    by (simp only: add_left_cancel)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   109
qed
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   110
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   111
lemma mult_zero_right [simp]: "a * 0 = (0::'a::semiring_0_cancel)"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   112
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   113
  have "a*0 + a*0 = a*0 + 0"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   114
    by (simp add: right_distrib [symmetric])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   115
  thus ?thesis 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   116
    by (simp only: add_left_cancel)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   117
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   118
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   119
lemma field_mult_eq_0_iff [simp]: "(a*b = (0::'a::field)) = (a = 0 | b = 0)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   120
proof cases
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   121
  assume "a=0" thus ?thesis by simp
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   122
next
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   123
  assume anz [simp]: "a\<noteq>0"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   124
  { assume "a * b = 0"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   125
    hence "inverse a * (a * b) = 0" by simp
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   126
    hence "b = 0"  by (simp (no_asm_use) add: mult_assoc [symmetric])}
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   127
  thus ?thesis by force
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   128
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   129
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   130
instance field \<subseteq> idom
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   131
by (intro_classes, simp)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   132
  
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   133
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
   134
  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
   135
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   136
subsection {* Distribution rules *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   137
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   138
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
   139
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   140
text{*For the @{text combine_numerals} simproc*}
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   141
lemma combine_common_factor:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   142
     "a*e + (b*e + c) = (a+b)*e + (c::'a::semiring)"
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   143
by (simp add: left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   144
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   145
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
   146
apply (rule equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   147
apply (simp add: left_distrib [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   148
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   149
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   150
lemma 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
   151
apply (rule equals_zero_I)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   152
apply (simp add: right_distrib [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   153
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   154
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   155
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
   156
  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
   157
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   158
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
   159
  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
   160
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   161
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
   162
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
   163
              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
   164
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   165
lemma left_diff_distrib: "(a - b) * c = a * c - b * (c::'a::ring)"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   166
by (simp add: left_distrib diff_minus 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   167
              minus_mult_left [symmetric] minus_mult_right [symmetric]) 
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   168
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   169
axclass pordered_semiring \<subseteq> semiring_0, pordered_ab_semigroup_add 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   170
  mult_left_mono: "a <= b \<Longrightarrow> 0 <= c \<Longrightarrow> c * a <= c * b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   171
  mult_right_mono: "a <= b \<Longrightarrow> 0 <= c \<Longrightarrow> a * c <= b * c"
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   172
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   173
axclass pordered_cancel_semiring \<subseteq> pordered_semiring, cancel_ab_semigroup_add
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   174
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   175
instance pordered_cancel_semiring \<subseteq> semiring_0_cancel ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   176
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   177
axclass ordered_semiring_strict \<subseteq> semiring_0, ordered_cancel_ab_semigroup_add
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   178
  mult_strict_left_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   179
  mult_strict_right_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> a * c < b * c"
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   180
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   181
instance ordered_semiring_strict \<subseteq> semiring_0_cancel ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   182
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   183
instance ordered_semiring_strict \<subseteq> pordered_cancel_semiring
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   184
apply intro_classes
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   185
apply (case_tac "a < b & 0 < c")
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   186
apply (auto simp add: mult_strict_left_mono order_less_le)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   187
apply (auto simp add: mult_strict_left_mono order_le_less)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   188
apply (simp add: mult_strict_right_mono)
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   189
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   190
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   191
axclass pordered_comm_semiring \<subseteq> comm_semiring_0, pordered_ab_semigroup_add
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   192
  mult_mono: "a <= b \<Longrightarrow> 0 <= c \<Longrightarrow> c * a <= c * b"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   193
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   194
axclass pordered_cancel_comm_semiring \<subseteq> pordered_comm_semiring, cancel_ab_semigroup_add
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   195
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   196
instance pordered_cancel_comm_semiring \<subseteq> pordered_comm_semiring ..
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   197
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   198
axclass ordered_comm_semiring_strict \<subseteq> comm_semiring_0, ordered_cancel_ab_semigroup_add
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   199
  mult_strict_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   200
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   201
instance pordered_comm_semiring \<subseteq> pordered_semiring
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   202
by (intro_classes, insert mult_mono, simp_all add: mult_commute, blast+)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   203
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   204
instance pordered_cancel_comm_semiring \<subseteq> pordered_cancel_semiring ..
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   205
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   206
instance ordered_comm_semiring_strict \<subseteq> ordered_semiring_strict
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   207
by (intro_classes, insert mult_strict_mono, simp_all add: mult_commute, blast+)
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   208
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   209
instance ordered_comm_semiring_strict \<subseteq> pordered_cancel_comm_semiring
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   210
apply (intro_classes)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   211
apply (case_tac "a < b & 0 < c")
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   212
apply (auto simp add: mult_strict_left_mono order_less_le)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   213
apply (auto simp add: mult_strict_left_mono order_le_less)
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   214
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   215
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   216
axclass pordered_ring \<subseteq> ring, pordered_semiring 
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   217
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   218
instance pordered_ring \<subseteq> pordered_ab_group_add ..
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   219
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   220
instance pordered_ring \<subseteq> pordered_cancel_semiring ..
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   221
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   222
axclass lordered_ring \<subseteq> pordered_ring, lordered_ab_group_abs
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   223
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   224
instance lordered_ring \<subseteq> lordered_ab_group_meet ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   225
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   226
instance lordered_ring \<subseteq> lordered_ab_group_join ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   227
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   228
axclass axclass_abs_if \<subseteq> minus, ord, zero
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   229
  abs_if: "abs a = (if (a < 0) then (-a) else a)"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   230
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   231
axclass ordered_ring_strict \<subseteq> ring, ordered_semiring_strict, axclass_abs_if
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   232
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   233
instance ordered_ring_strict \<subseteq> lordered_ab_group ..
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   234
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   235
instance ordered_ring_strict \<subseteq> lordered_ring
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   236
by (intro_classes, simp add: abs_if join_eq_if)
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   237
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   238
axclass pordered_comm_ring \<subseteq> comm_ring, pordered_comm_semiring
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   239
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   240
axclass ordered_semidom \<subseteq> comm_semiring_1_cancel, ordered_comm_semiring_strict (* previously ordered_semiring *)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   241
  zero_less_one [simp]: "0 < 1"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   242
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   243
axclass ordered_idom \<subseteq> comm_ring_1, ordered_comm_semiring_strict, axclass_abs_if (* previously ordered_ring *)
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   244
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   245
instance ordered_idom \<subseteq> ordered_ring_strict ..
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   246
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   247
axclass ordered_field \<subseteq> field, ordered_idom
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   248
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   249
lemma eq_add_iff1:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   250
     "(a*e + c = b*e + d) = ((a-b)*e + c = (d::'a::ring))"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   251
apply (simp add: diff_minus left_distrib)
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   252
apply (simp add: diff_minus left_distrib add_ac)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   253
apply (simp add: compare_rls minus_mult_left [symmetric])
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   254
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   255
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   256
lemma eq_add_iff2:
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   257
     "(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
   258
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   259
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   260
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   261
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   262
lemma less_add_iff1:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   263
     "(a*e + c < b*e + d) = ((a-b)*e + c < (d::'a::pordered_ring))"
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   264
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   265
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   266
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   267
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   268
lemma less_add_iff2:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   269
     "(a*e + c < b*e + d) = (c < (b-a)*e + (d::'a::pordered_ring))"
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   270
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   271
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   272
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   273
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   274
lemma le_add_iff1:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   275
     "(a*e + c \<le> b*e + d) = ((a-b)*e + c \<le> (d::'a::pordered_ring))"
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   276
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   277
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   278
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   279
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   280
lemma le_add_iff2:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   281
     "(a*e + c \<le> b*e + d) = (c \<le> (b-a)*e + (d::'a::pordered_ring))"
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   282
apply (simp add: diff_minus left_distrib add_ac)
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   283
apply (simp add: compare_rls minus_mult_left [symmetric]) 
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   284
done
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   285
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   286
subsection {* Ordering Rules for Multiplication *}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   287
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   288
lemma mult_left_le_imp_le:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   289
     "[|c*a \<le> c*b; 0 < c|] ==> a \<le> (b::'a::ordered_semiring_strict)"
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   290
  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
   291
 
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   292
lemma mult_right_le_imp_le:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   293
     "[|a*c \<le> b*c; 0 < c|] ==> a \<le> (b::'a::ordered_semiring_strict)"
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   294
  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
   295
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   296
lemma mult_left_less_imp_less:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   297
     "[|c*a < c*b; 0 \<le> c|] ==> a < (b::'a::ordered_semiring_strict)"
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   298
  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
   299
 
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   300
lemma mult_right_less_imp_less:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   301
     "[|a*c < b*c; 0 \<le> c|] ==> a < (b::'a::ordered_semiring_strict)"
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   302
  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
   303
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   304
lemma mult_strict_left_mono_neg:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   305
     "[|b < a; c < 0|] ==> c * a < c * (b::'a::ordered_ring_strict)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   306
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
   307
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
   308
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   309
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   310
lemma mult_left_mono_neg:
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   311
     "[|b \<le> a; c \<le> 0|] ==> c * a \<le>  c * (b::'a::pordered_ring)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   312
apply (drule mult_left_mono [of _ _ "-c"])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   313
apply (simp_all add: minus_mult_left [symmetric]) 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   314
done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   315
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   316
lemma mult_strict_right_mono_neg:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   317
     "[|b < a; c < 0|] ==> a * c < b * (c::'a::ordered_ring_strict)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   318
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
   319
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
   320
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   321
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   322
lemma mult_right_mono_neg:
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   323
     "[|b \<le> a; c \<le> 0|] ==> a * c \<le>  (b::'a::pordered_ring) * c"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   324
apply (drule mult_right_mono [of _ _ "-c"])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   325
apply (simp)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   326
apply (simp_all add: minus_mult_right [symmetric]) 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   327
done
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   328
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   329
subsection{* Products of Signs *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   330
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   331
lemma mult_pos: "[| (0::'a::ordered_semiring_strict) < 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
   332
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
   333
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   334
lemma mult_pos_le: "[| (0::'a::pordered_cancel_semiring) \<le> a; 0 \<le> b |] ==> 0 \<le> a*b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   335
by (drule mult_left_mono [of 0 b], auto)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   336
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   337
lemma mult_pos_neg: "[| (0::'a::ordered_semiring_strict) < 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
   338
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
   339
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   340
lemma mult_pos_neg_le: "[| (0::'a::pordered_cancel_semiring) \<le> a; b \<le> 0 |] ==> a*b \<le> 0"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   341
by (drule mult_left_mono [of b 0], auto)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   342
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   343
lemma mult_pos_neg2: "[| (0::'a::ordered_semiring_strict) < a; b < 0 |] ==> b*a < 0" 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   344
by (drule mult_strict_right_mono[of b 0], auto)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   345
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   346
lemma mult_pos_neg2_le: "[| (0::'a::pordered_cancel_semiring) \<le> a; b \<le> 0 |] ==> b*a \<le> 0" 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   347
by (drule mult_right_mono[of b 0], auto)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   348
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   349
lemma mult_neg: "[| a < (0::'a::ordered_ring_strict); b < 0 |] ==> 0 < a*b"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   350
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
   351
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   352
lemma mult_neg_le: "[| a \<le> (0::'a::pordered_ring); b \<le> 0 |] ==> 0 \<le> a*b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   353
by (drule mult_right_mono_neg[of a 0 b ], auto)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   354
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   355
lemma zero_less_mult_pos:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   356
     "[| 0 < a*b; 0 < a|] ==> 0 < (b::'a::ordered_semiring_strict)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   357
apply (case_tac "b\<le>0") 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   358
 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
   359
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
   360
 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
   361
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   362
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   363
lemma zero_less_mult_pos2:
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   364
     "[| 0 < b*a; 0 < a|] ==> 0 < (b::'a::ordered_semiring_strict)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   365
apply (case_tac "b\<le>0") 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   366
 apply (auto simp add: order_le_less linorder_not_less)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   367
apply (drule_tac mult_pos_neg2 [of a b]) 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   368
 apply (auto dest: order_less_not_sym)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   369
done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   370
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   371
lemma zero_less_mult_iff:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   372
     "((0::'a::ordered_ring_strict) < a*b) = (0 < a & 0 < b | a < 0 & b < 0)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   373
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
   374
apply (blast dest: zero_less_mult_pos) 
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   375
apply (blast dest: zero_less_mult_pos2)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   376
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   377
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   378
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
   379
      assumption of an ordering.  See @{text field_mult_eq_0_iff} below.*}
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   380
lemma mult_eq_0_iff [simp]: "(a*b = (0::'a::ordered_ring_strict)) = (a = 0 | b = 0)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   381
apply (case_tac "a < 0")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   382
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
   383
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
   384
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   385
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   386
lemma zero_le_mult_iff:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   387
     "((0::'a::ordered_ring_strict) \<le> a*b) = (0 \<le> a & 0 \<le> b | a \<le> 0 & b \<le> 0)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   388
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
   389
                   zero_less_mult_iff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   390
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   391
lemma mult_less_0_iff:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   392
     "(a*b < (0::'a::ordered_ring_strict)) = (0 < a & b < 0 | a < 0 & 0 < b)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   393
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
   394
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
   395
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   396
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   397
lemma mult_le_0_iff:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   398
     "(a*b \<le> (0::'a::ordered_ring_strict)) = (0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   399
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
   400
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
   401
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   402
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   403
lemma split_mult_pos_le: "(0 \<le> a & 0 \<le> b) | (a \<le> 0 & b \<le> 0) \<Longrightarrow> 0 \<le> a * (b::_::pordered_ring)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   404
by (auto simp add: mult_pos_le mult_neg_le)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   405
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   406
lemma split_mult_neg_le: "(0 \<le> a & b \<le> 0) | (a \<le> 0 & 0 \<le> b) \<Longrightarrow> a * b \<le> (0::_::pordered_cancel_semiring)" 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   407
by (auto simp add: mult_pos_neg_le mult_pos_neg2_le)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   408
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   409
lemma zero_le_square: "(0::'a::ordered_ring_strict) \<le> a*a"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   410
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
   411
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   412
text{*Proving axiom @{text zero_less_one} makes all @{text ordered_semidom}
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   413
      theorems available to members of @{term ordered_idom} *}
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   414
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   415
instance ordered_idom \<subseteq> ordered_semidom
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   416
proof
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   417
  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
   418
  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
   419
qed
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   420
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   421
instance ordered_ring_strict \<subseteq> axclass_no_zero_divisors 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   422
by (intro_classes, simp)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   423
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   424
instance ordered_idom \<subseteq> idom ..
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   425
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   426
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
   427
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   428
declare zero_neq_one [THEN not_sym, simp]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   429
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   430
lemma zero_le_one [simp]: "(0::'a::ordered_semidom) \<le> 1"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   431
  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
   432
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   433
lemma not_one_le_zero [simp]: "~ (1::'a::ordered_semidom) \<le> 0"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   434
by (simp add: linorder_not_le) 
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   435
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   436
lemma not_one_less_zero [simp]: "~ (1::'a::ordered_semidom) < 0"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   437
by (simp add: linorder_not_less) 
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   438
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   439
subsection{*More Monotonicity*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   440
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   441
text{*Strict monotonicity in both arguments*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   442
lemma mult_strict_mono:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   443
     "[|a<b; c<d; 0<b; 0\<le>c|] ==> a * c < b * (d::'a::ordered_semiring_strict)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   444
apply (case_tac "c=0")
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   445
 apply (simp add: mult_pos) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   446
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
   447
 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
   448
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
   449
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   450
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   451
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
   452
lemma mult_strict_mono':
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   453
     "[| a<b; c<d; 0 \<le> a; 0 \<le> c|] ==> a * c < b * (d::'a::ordered_semiring_strict)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   454
apply (rule mult_strict_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   455
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
   456
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   457
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   458
lemma mult_mono:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   459
     "[|a \<le> b; c \<le> d; 0 \<le> b; 0 \<le> c|] 
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   460
      ==> a * c  \<le>  b * (d::'a::pordered_semiring)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   461
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
   462
apply (erule mult_left_mono, assumption)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   463
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   464
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   465
lemma less_1_mult: "[| 1 < m; 1 < n |] ==> 1 < m*(n::'a::ordered_semidom)"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   466
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
   467
apply (simp add:  order_less_trans [OF zero_less_one]) 
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   468
done
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   469
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   470
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
   471
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   472
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
   473
   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
   474
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   475
lemma mult_less_cancel_right:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   476
    "(a*c < b*c) = ((0 < c & a < b) | (c < 0 & b < (a::'a::ordered_ring_strict)))"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   477
apply (case_tac "c = 0")
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   478
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
   479
                      mult_strict_right_mono_neg)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   480
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
   481
                      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
   482
                      linorder_not_le [symmetric, of a])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   483
apply (erule_tac [!] notE)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   484
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
   485
                      mult_right_mono_neg)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   486
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   487
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   488
lemma mult_less_cancel_left:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   489
    "(c*a < c*b) = ((0 < c & a < b) | (c < 0 & b < (a::'a::ordered_ring_strict)))"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   490
apply (case_tac "c = 0")
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   491
apply (auto simp add: linorder_neq_iff mult_strict_left_mono 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   492
                      mult_strict_left_mono_neg)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   493
apply (auto simp add: linorder_not_less 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   494
                      linorder_not_le [symmetric, of "c*a"]
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   495
                      linorder_not_le [symmetric, of a])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   496
apply (erule_tac [!] notE)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   497
apply (auto simp add: order_less_imp_le mult_left_mono 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   498
                      mult_left_mono_neg)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   499
done
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   500
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   501
lemma mult_le_cancel_right:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   502
     "(a*c \<le> b*c) = ((0<c --> a\<le>b) & (c<0 --> b \<le> (a::'a::ordered_ring_strict)))"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   503
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
   504
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   505
lemma mult_le_cancel_left:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   506
     "(c*a \<le> c*b) = ((0<c --> a\<le>b) & (c<0 --> b \<le> (a::'a::ordered_ring_strict)))"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   507
by (simp add: linorder_not_less [symmetric] mult_less_cancel_left)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   508
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   509
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
   510
      assumes less: "c*a < c*b" and nonneg: "0 \<le> c"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   511
      shows "a < (b::'a::ordered_semiring_strict)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   512
proof (rule ccontr)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   513
  assume "~ a < b"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   514
  hence "b \<le> a" by (simp add: linorder_not_less)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   515
  hence "c*b \<le> c*a" by (rule mult_left_mono)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   516
  with this and less show False 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   517
    by (simp add: linorder_not_less [symmetric])
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   518
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   519
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   520
lemma mult_less_imp_less_right:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   521
  assumes less: "a*c < b*c" and nonneg: "0 <= c"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   522
  shows "a < (b::'a::ordered_semiring_strict)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   523
proof (rule ccontr)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   524
  assume "~ a < b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   525
  hence "b \<le> a" by (simp add: linorder_not_less)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   526
  hence "b*c \<le> a*c" by (rule mult_right_mono)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   527
  with this and less show False 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   528
    by (simp add: linorder_not_less [symmetric])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   529
qed  
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   530
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   531
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
   532
lemma mult_cancel_right [simp]:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   533
     "(a*c = b*c) = (c = (0::'a::ordered_ring_strict) | a=b)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   534
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
   535
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
   536
             simp add: linorder_neq_iff)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   537
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   538
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   539
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
   540
      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
   541
lemma mult_cancel_left [simp]:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   542
     "(c*a = c*b) = (c = (0::'a::ordered_ring_strict) | a=b)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   543
apply (cut_tac linorder_less_linear [of 0 c])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   544
apply (force dest: mult_strict_left_mono_neg mult_strict_left_mono
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   545
             simp add: linorder_neq_iff)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   546
done
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   547
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   548
text{*This list of rewrites decides ring equalities by ordered rewriting.*}
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
   549
lemmas ring_eq_simps =  
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
   550
(*  mult_ac*)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   551
  left_distrib right_distrib left_diff_distrib right_diff_distrib
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
   552
  group_eq_simps
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
   553
(*  add_ac
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   554
  add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
   555
  diff_eq_eq eq_diff_eq *)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   556
    
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   557
subsection {* Fields *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   558
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   559
lemma right_inverse [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   560
      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
   561
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   562
  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
   563
  also have "... = 1" using not0 by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   564
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   565
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   566
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   567
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
   568
proof
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   569
  assume neq: "b \<noteq> 0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   570
  {
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   571
    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
   572
    also assume "a / b = 1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   573
    finally show "a = b" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   574
  next
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   575
    assume "a = b"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   576
    with neq show "a / b = 1" by (simp add: divide_inverse)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   577
  }
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   578
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   579
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   580
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
   581
by (simp add: divide_inverse)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   582
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   583
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
   584
  by (simp add: divide_inverse)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   585
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   586
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
   587
by (simp add: divide_inverse)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   588
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   589
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
   590
by (simp add: divide_inverse)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   591
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   592
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
   593
by (simp add: divide_inverse)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   594
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   595
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
   596
by (simp add: divide_inverse left_distrib) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   597
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   598
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   599
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
   600
      of an ordering.*}
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   601
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
   602
proof cases
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   603
  assume "a=0" thus ?thesis by simp
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   604
next
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   605
  assume anz [simp]: "a\<noteq>0"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   606
  { assume "a * b = 0"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   607
    hence "inverse a * (a * b) = 0" by simp
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   608
    hence "b = 0"  by (simp (no_asm_use) add: mult_assoc [symmetric])}
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   609
  thus ?thesis by force
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   610
qed
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   611
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   612
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
   613
lemma field_mult_cancel_right_lemma:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   614
      assumes cnz: "c \<noteq> (0::'a::field)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   615
	  and eq:  "a*c = b*c"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   616
	 shows "a=b"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   617
proof -
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   618
  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
   619
    by (simp add: eq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   620
  thus "a=b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   621
    by (simp add: mult_assoc cnz)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   622
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   623
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   624
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
   625
     "(a*c = b*c) = (c = (0::'a::field) | a=b)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   626
proof cases
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   627
  assume "c=0" thus ?thesis by simp
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   628
next
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   629
  assume "c\<noteq>0" 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   630
  thus ?thesis by (force dest: field_mult_cancel_right_lemma)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   631
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   632
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   633
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
   634
     "(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
   635
  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
   636
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   637
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
   638
proof
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   639
  assume ianz: "inverse a = 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   640
  assume "a \<noteq> 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   641
  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
   642
  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
   643
  finally have "1 = (0::'a::field)" .
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   644
  thus False by (simp add: eq_commute)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   645
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   646
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   647
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   648
subsection{*Basic Properties of @{term inverse}*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   649
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   650
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
   651
apply (rule ccontr) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   652
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
   653
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   654
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   655
lemma inverse_nonzero_imp_nonzero:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   656
   "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
   657
apply (rule ccontr) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   658
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
   659
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   660
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   661
lemma inverse_nonzero_iff_nonzero [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   662
   "(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
   663
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
   664
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   665
lemma nonzero_inverse_minus_eq:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   666
      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
   667
proof -
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   668
  have "-a * inverse (- a) = -a * - inverse a"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   669
    by simp
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   670
  thus ?thesis 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   671
    by (simp only: field_mult_cancel_left, simp)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   672
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   673
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   674
lemma inverse_minus_eq [simp]:
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   675
   "inverse(-a) = -inverse(a::'a::{field,division_by_zero})";
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   676
proof cases
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   677
  assume "a=0" thus ?thesis by (simp add: inverse_zero)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   678
next
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   679
  assume "a\<noteq>0" 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   680
  thus ?thesis by (simp add: nonzero_inverse_minus_eq)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   681
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   682
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   683
lemma nonzero_inverse_eq_imp_eq:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   684
      assumes inveq: "inverse a = inverse b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   685
	  and anz:  "a \<noteq> 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   686
	  and bnz:  "b \<noteq> 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   687
	 shows "a = (b::'a::field)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   688
proof -
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   689
  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
   690
    by (simp add: inveq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   691
  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
   692
    by simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   693
  thus "a = b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   694
    by (simp add: mult_assoc anz bnz)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   695
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   696
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   697
lemma inverse_eq_imp_eq:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   698
     "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
   699
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
   700
 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
   701
              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
   702
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
   703
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   704
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   705
lemma inverse_eq_iff_eq [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   706
     "(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
   707
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
   708
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   709
lemma nonzero_inverse_inverse_eq:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   710
      assumes [simp]: "a \<noteq> 0"  shows "inverse(inverse (a::'a::field)) = a"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   711
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   712
  have "(inverse (inverse a) * inverse a) * a = a" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   713
    by (simp add: nonzero_imp_inverse_nonzero)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   714
  thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   715
    by (simp add: mult_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   716
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   717
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   718
lemma inverse_inverse_eq [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   719
     "inverse(inverse (a::'a::{field,division_by_zero})) = a"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   720
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   721
    assume "a=0" thus ?thesis by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   722
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   723
    assume "a\<noteq>0" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   724
    thus ?thesis by (simp add: nonzero_inverse_inverse_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   725
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   726
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   727
lemma inverse_1 [simp]: "inverse 1 = (1::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   728
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   729
  have "inverse 1 * 1 = (1::'a::field)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   730
    by (rule left_inverse [OF zero_neq_one [symmetric]])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   731
  thus ?thesis  by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   732
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   733
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
   734
lemma inverse_unique: 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
   735
  assumes ab: "a*b = 1"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
   736
  shows "inverse a = (b::'a::field)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
   737
proof -
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
   738
  have "a \<noteq> 0" using ab by auto
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
   739
  moreover have "inverse a * (a * b) = inverse a" by (simp add: ab) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
   740
  ultimately show ?thesis by (simp add: mult_assoc [symmetric]) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
   741
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
   742
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   743
lemma nonzero_inverse_mult_distrib: 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   744
      assumes anz: "a \<noteq> 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   745
          and bnz: "b \<noteq> 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   746
      shows "inverse(a*b) = inverse(b) * inverse(a::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   747
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   748
  have "inverse(a*b) * (a * b) * inverse(b) = inverse(b)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   749
    by (simp add: field_mult_eq_0_iff anz bnz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   750
  hence "inverse(a*b) * a = inverse(b)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   751
    by (simp add: mult_assoc bnz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   752
  hence "inverse(a*b) * a * inverse(a) = inverse(b) * inverse(a)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   753
    by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   754
  thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   755
    by (simp add: mult_assoc anz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   756
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   757
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   758
text{*This version builds in division by zero while also re-orienting
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   759
      the right-hand side.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   760
lemma inverse_mult_distrib [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   761
     "inverse(a*b) = inverse(a) * inverse(b::'a::{field,division_by_zero})"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   762
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   763
    assume "a \<noteq> 0 & b \<noteq> 0" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   764
    thus ?thesis  by (simp add: nonzero_inverse_mult_distrib mult_commute)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   765
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   766
    assume "~ (a \<noteq> 0 & b \<noteq> 0)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   767
    thus ?thesis  by force
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   768
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   769
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   770
text{*There is no slick version using division by zero.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   771
lemma inverse_add:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   772
     "[|a \<noteq> 0;  b \<noteq> 0|]
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   773
      ==> inverse a + inverse b = (a+b) * inverse a * inverse (b::'a::field)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   774
apply (simp add: left_distrib mult_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   775
apply (simp add: mult_commute [of "inverse a"]) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   776
apply (simp add: mult_assoc [symmetric] add_commute)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   777
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   778
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   779
lemma inverse_divide [simp]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   780
      "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
   781
  by (simp add: divide_inverse mult_commute)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   782
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   783
lemma nonzero_mult_divide_cancel_left:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   784
  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
   785
    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
   786
proof -
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   787
  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
   788
    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
   789
                  nonzero_inverse_mult_distrib)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   790
  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
   791
    by (simp only: mult_ac)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   792
  also have "... =  a * inverse b"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   793
    by simp
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   794
    finally show ?thesis 
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   795
    by (simp add: divide_inverse)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   796
qed
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   797
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   798
lemma mult_divide_cancel_left:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   799
     "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
   800
apply (case_tac "b = 0")
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   801
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
   802
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   803
14321
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
   804
lemma nonzero_mult_divide_cancel_right:
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
   805
     "[|b\<noteq>0; c\<noteq>0|] ==> (a*c) / (b*c) = a/(b::'a::field)"
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
   806
by (simp add: mult_commute [of _ c] nonzero_mult_divide_cancel_left) 
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
   807
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
   808
lemma mult_divide_cancel_right:
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
   809
     "c\<noteq>0 ==> (a*c) / (b*c) = a / (b::'a::{field,division_by_zero})"
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
   810
apply (case_tac "b = 0")
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
   811
apply (simp_all add: nonzero_mult_divide_cancel_right)
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
   812
done
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
   813
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   814
(*For ExtractCommonTerm*)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   815
lemma mult_divide_cancel_eq_if:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   816
     "(c*a) / (c*b) = 
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   817
      (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
   818
  by (simp add: mult_divide_cancel_left)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   819
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
   820
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
   821
  by (simp add: divide_inverse)
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
   822
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   823
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
   824
by (simp add: divide_inverse mult_assoc)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   825
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   826
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
   827
by (simp add: divide_inverse mult_ac)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   828
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   829
lemma divide_divide_eq_right [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   830
     "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
   831
by (simp add: divide_inverse mult_ac)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   832
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   833
lemma divide_divide_eq_left [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   834
     "(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
   835
by (simp add: divide_inverse mult_assoc)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   836
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   837
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   838
subsection {* Division and Unary Minus *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   839
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   840
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
   841
by (simp add: divide_inverse minus_mult_left)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   842
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   843
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
   844
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
   845
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   846
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
   847
by (simp add: divide_inverse nonzero_inverse_minus_eq)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   848
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   849
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
   850
by (simp add: divide_inverse minus_mult_left [symmetric])
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   851
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   852
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
   853
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
   854
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   855
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   856
text{*The effect is to extract signs from divisions*}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   857
declare minus_divide_left  [symmetric, simp]
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   858
declare minus_divide_right [symmetric, simp]
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   859
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   860
text{*Also, extract signs from products*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   861
declare minus_mult_left [symmetric, simp]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   862
declare minus_mult_right [symmetric, simp]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   863
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   864
lemma minus_divide_divide [simp]:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   865
     "(-a)/(-b) = a / (b::'a::{field,division_by_zero})"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   866
apply (case_tac "b=0", simp) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   867
apply (simp add: nonzero_minus_divide_divide) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   868
done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   869
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   870
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
   871
by (simp add: diff_minus add_divide_distrib) 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   872
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   873
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   874
subsection {* Ordered Fields *}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   875
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   876
lemma positive_imp_inverse_positive: 
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   877
      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
   878
  proof -
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   879
  have "0 < a * inverse a" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   880
    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
   881
  thus "0 < inverse a" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   882
    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
   883
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   884
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   885
lemma negative_imp_inverse_negative:
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   886
     "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
   887
  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
   888
      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
   889
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   890
lemma inverse_le_imp_le:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   891
      assumes invle: "inverse a \<le> inverse b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   892
	  and apos:  "0 < a"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   893
	 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
   894
  proof (rule classical)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   895
  assume "~ b \<le> a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   896
  hence "a < b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   897
    by (simp add: linorder_not_le)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   898
  hence bpos: "0 < b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   899
    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
   900
  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
   901
    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
   902
  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
   903
    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
   904
  thus "b \<le> a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   905
    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
   906
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   907
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   908
lemma inverse_positive_imp_positive:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   909
      assumes inv_gt_0: "0 < inverse a"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   910
          and [simp]:   "a \<noteq> 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   911
        shows "0 < (a::'a::ordered_field)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   912
  proof -
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   913
  have "0 < inverse (inverse a)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   914
    by (rule positive_imp_inverse_positive)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   915
  thus "0 < a"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   916
    by (simp add: nonzero_inverse_inverse_eq)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   917
  qed
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   918
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   919
lemma inverse_positive_iff_positive [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   920
      "(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
   921
apply (case_tac "a = 0", simp)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   922
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
   923
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   924
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   925
lemma inverse_negative_imp_negative:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   926
      assumes inv_less_0: "inverse a < 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   927
          and [simp]:   "a \<noteq> 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   928
        shows "a < (0::'a::ordered_field)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   929
  proof -
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   930
  have "inverse (inverse a) < 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   931
    by (rule negative_imp_inverse_negative)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   932
  thus "a < 0"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   933
    by (simp add: nonzero_inverse_inverse_eq)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   934
  qed
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   935
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   936
lemma inverse_negative_iff_negative [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   937
      "(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
   938
apply (case_tac "a = 0", simp)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   939
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
   940
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   941
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   942
lemma inverse_nonnegative_iff_nonnegative [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   943
      "(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
   944
by (simp add: linorder_not_less [symmetric])
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   945
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   946
lemma inverse_nonpositive_iff_nonpositive [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   947
      "(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
   948
by (simp add: linorder_not_less [symmetric])
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   949
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   950
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   951
subsection{*Anti-Monotonicity of @{term inverse}*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   952
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   953
lemma less_imp_inverse_less:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   954
      assumes less: "a < b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   955
	  and apos:  "0 < a"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   956
	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
   957
  proof (rule ccontr)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   958
  assume "~ inverse b < inverse a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   959
  hence "inverse a \<le> inverse b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   960
    by (simp add: linorder_not_less)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   961
  hence "~ (a < b)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   962
    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
   963
  thus False
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   964
    by (rule notE [OF _ less])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   965
  qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   966
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   967
lemma inverse_less_imp_less:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   968
   "[|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
   969
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
   970
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
   971
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   972
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   973
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
   974
lemma inverse_less_iff_less [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   975
     "[|0 < a; 0 < b|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   976
      ==> (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
   977
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
   978
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   979
lemma le_imp_inverse_le:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   980
   "[|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
   981
  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
   982
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   983
lemma inverse_le_iff_le [simp]:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   984
     "[|0 < a; 0 < b|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   985
      ==> (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
   986
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
   987
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   988
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   989
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
   990
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
   991
lemma inverse_le_imp_le_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   992
   "[|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
   993
  apply (rule classical) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   994
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   995
   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
   996
  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
   997
  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
   998
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   999
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1000
lemma less_imp_inverse_less_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1001
   "[|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
  1002
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1003
   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
  1004
  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
  1005
  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
  1006
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1007
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1008
lemma inverse_less_imp_less_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1009
   "[|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
  1010
  apply (rule classical) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1011
  apply (subgoal_tac "a < 0") 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1012
   prefer 2
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1013
   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
  1014
  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
  1015
  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
  1016
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1017
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1018
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
  1019
     "[|a < 0; b < 0|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1020
      ==> (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
  1021
  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
  1022
  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
  1023
	      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
  1024
  done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1025
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1026
lemma le_imp_inverse_le_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1027
   "[|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
  1028
  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
  1029
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1030
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
  1031
     "[|a < 0; b < 0|] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1032
      ==> (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
  1033
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
  1034
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1035
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1036
subsection{*Inverses and the Number One*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1037
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1038
lemma one_less_inverse_iff:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1039
    "(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
  1040
  assume "0 < x"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1041
    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
  1042
    show ?thesis by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1043
next
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1044
  assume notless: "~ (0 < x)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1045
  have "~ (1 < inverse x)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1046
  proof
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1047
    assume "1 < inverse x"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1048
    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
  1049
    also have "... < 1" by (rule zero_less_one) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1050
    finally show False by auto
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1051
  qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1052
  with notless show ?thesis by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1053
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1054
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1055
lemma inverse_eq_1_iff [simp]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1056
    "(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
  1057
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
  1058
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1059
lemma one_le_inverse_iff:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1060
   "(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
  1061
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
  1062
                    eq_commute [of 1]) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1063
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1064
lemma inverse_less_1_iff:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1065
   "(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
  1066
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
  1067
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1068
lemma inverse_le_1_iff:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1069
   "(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
  1070
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
  1071
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1072
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1073
subsection{*Division and Signs*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1074
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1075
lemma zero_less_divide_iff:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1076
     "((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
  1077
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
  1078
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1079
lemma divide_less_0_iff:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1080
     "(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
  1081
      (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
  1082
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
  1083
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1084
lemma zero_le_divide_iff:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1085
     "((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
  1086
      (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
  1087
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
  1088
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1089
lemma divide_le_0_iff:
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1090
     "(a/b \<le> (0::'a::{ordered_field,division_by_zero})) =
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1091
      (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
  1092
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
  1093
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1094
lemma divide_eq_0_iff [simp]:
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1095
     "(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
  1096
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
  1097
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1098
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1099
subsection{*Simplification of Inequalities Involving Literal Divisors*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1100
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1101
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
  1102
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1103
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1104
  hence "(a \<le> b/c) = (a*c \<le> (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1105
    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
  1106
  also have "... = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1107
    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
  1108
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1109
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1110
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1111
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1112
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
  1113
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1114
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1115
  hence "(a \<le> b/c) = ((b/c)*c \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1116
    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
  1117
  also have "... = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1118
    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
  1119
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1120
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1121
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1122
lemma le_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1123
  "(a \<le> b/c) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1124
   (if 0 < c then a*c \<le> b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1125
             else if c < 0 then b \<le> a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1126
             else  a \<le> (0::'a::{ordered_field,division_by_zero}))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1127
apply (case_tac "c=0", simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1128
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
  1129
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1130
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1131
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
  1132
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1133
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1134
  hence "(b/c \<le> a) = ((b/c)*c \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1135
    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
  1136
  also have "... = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1137
    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
  1138
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1139
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1140
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1141
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
  1142
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1143
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1144
  hence "(b/c \<le> a) = (a*c \<le> (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1145
    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
  1146
  also have "... = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1147
    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
  1148
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1149
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1150
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1151
lemma divide_le_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1152
  "(b/c \<le> a) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1153
   (if 0 < c then b \<le> a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1154
             else if c < 0 then a*c \<le> b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1155
             else 0 \<le> (a::'a::{ordered_field,division_by_zero}))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1156
apply (case_tac "c=0", simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1157
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
  1158
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1159
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1160
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1161
lemma pos_less_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1162
     "0 < (c::'a::ordered_field) ==> (a < b/c) = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1163
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1164
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1165
  hence "(a < b/c) = (a*c < (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1166
    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
  1167
  also have "... = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1168
    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1169
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1170
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1171
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1172
lemma neg_less_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1173
 "c < (0::'a::ordered_field) ==> (a < b/c) = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1174
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1175
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1176
  hence "(a < b/c) = ((b/c)*c < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1177
    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
  1178
  also have "... = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1179
    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
  1180
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1181
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1182
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1183
lemma less_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1184
  "(a < b/c) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1185
   (if 0 < c then a*c < b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1186
             else if c < 0 then b < a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1187
             else  a < (0::'a::{ordered_field,division_by_zero}))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1188
apply (case_tac "c=0", simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1189
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
  1190
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1191
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1192
lemma pos_divide_less_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1193
     "0 < (c::'a::ordered_field) ==> (b/c < a) = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1194
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1195
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1196
  hence "(b/c < a) = ((b/c)*c < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1197
    by (simp add: mult_less_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1198
  also have "... = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1199
    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1200
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1201
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1202
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1203
lemma neg_divide_less_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1204
 "c < (0::'a::ordered_field) ==> (b/c < a) = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1205
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1206
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1207
  hence "(b/c < a) = (a*c < (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1208
    by (simp add: mult_less_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1209
  also have "... = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1210
    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1211
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1212
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1213
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1214
lemma divide_less_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1215
  "(b/c < a) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1216
   (if 0 < c then b < a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1217
             else if c < 0 then a*c < b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1218
             else 0 < (a::'a::{ordered_field,division_by_zero}))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1219
apply (case_tac "c=0", simp) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1220
apply (force simp add: pos_divide_less_eq neg_divide_less_eq linorder_neq_iff) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1221
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1222
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1223
lemma 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
  1224
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1225
  assume [simp]: "c\<noteq>0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1226
  have "(a = b/c) = (a*c = (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1227
    by (simp add: field_mult_cancel_right)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1228
  also have "... = (a*c = b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1229
    by (simp add: divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1230
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1231
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1232
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1233
lemma eq_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1234
  "((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
  1235
by (simp add: nonzero_eq_divide_eq) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1236
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1237
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
  1238
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1239
  assume [simp]: "c\<noteq>0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1240
  have "(b/c = a) = ((b/c)*c = a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1241
    by (simp add: field_mult_cancel_right)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1242
  also have "... = (b = a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1243
    by (simp add: divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1244
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1245
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1246
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1247
lemma divide_eq_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1248
  "(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
  1249
by (force simp add: nonzero_divide_eq_eq) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1250
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1251
subsection{*Cancellation Laws for Division*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1252
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1253
lemma divide_cancel_right [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1254
     "(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
  1255
apply (case_tac "c=0", simp) 
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1256
apply (simp add: divide_inverse field_mult_cancel_right) 
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1257
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1258
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1259
lemma divide_cancel_left [simp]:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1260
     "(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
  1261
apply (case_tac "c=0", simp) 
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1262
apply (simp add: divide_inverse field_mult_cancel_left) 
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1263
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1264
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1265
subsection {* Division and the Number One *}
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1266
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1267
text{*Simplify expressions equated with 1*}
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1268
lemma divide_eq_1_iff [simp]:
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1269
     "(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
  1270
apply (case_tac "b=0", simp) 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1271
apply (simp add: right_inverse_eq) 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1272
done
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1273
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1274
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1275
lemma one_eq_divide_iff [simp]:
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1276
     "(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
  1277
by (simp add: eq_commute [of 1])  
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1278
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1279
lemma zero_eq_1_divide_iff [simp]:
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1280
     "((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
  1281
apply (case_tac "a=0", simp) 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1282
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
  1283
done
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1284
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1285
lemma one_divide_eq_0_iff [simp]:
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1286
     "(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
  1287
apply (case_tac "a=0", simp) 
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1288
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
  1289
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
  1290
done
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1291
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1292
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
  1293
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
  1294
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
  1295
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
  1296
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
  1297
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1298
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1299
subsection {* Ordering Rules for Division *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1300
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1301
lemma divide_strict_right_mono:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1302
     "[|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
  1303
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
  1304
              positive_imp_inverse_positive) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1305
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1306
lemma divide_right_mono:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1307
     "[|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
  1308
  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
  1309
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1310
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1311
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
  1312
      have the same sign*}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1313
lemma divide_strict_left_mono:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1314
       "[|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
  1315
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
  1316
      order_less_imp_not_eq order_less_imp_not_eq2  
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1317
      less_imp_inverse_less less_imp_inverse_less_neg) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1318
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1319
lemma divide_left_mono:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1320
     "[|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
  1321
  apply (subgoal_tac "a \<noteq> 0 & b \<noteq> 0") 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1322
   prefer 2 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1323
   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
  1324
  apply (case_tac "c=0", simp add: divide_inverse)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1325
  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
  1326
  done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1327
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1328
lemma divide_strict_left_mono_neg:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1329
     "[|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
  1330
  apply (subgoal_tac "a \<noteq> 0 & b \<noteq> 0") 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1331
   prefer 2 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1332
   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
  1333
  apply (drule divide_strict_left_mono [of _ _ "-c"]) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1334
   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
  1335
  done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1336
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1337
lemma divide_strict_right_mono_neg:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1338
     "[|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
  1339
apply (drule divide_strict_right_mono [of _ _ "-c"], simp) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1340
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
  1341
done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1342
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1343
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1344
subsection {* Ordered Fields are Dense *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1345
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1346
lemma less_add_one: "a < (a+1::'a::ordered_semidom)"
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1347
proof -
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1348
  have "a+0 < (a+1::'a::ordered_semidom)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1349
    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
  1350
  thus ?thesis by simp
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1351
qed
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1352
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1353
lemma zero_less_two: "0 < (1+1::'a::ordered_semidom)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1354
  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
  1355
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1356
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
  1357
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
  1358
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1359
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
  1360
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
  1361
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1362
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
  1363
by (blast intro!: less_half_sum gt_half_sum)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1364
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1365
subsection {* Absolute Value *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1366
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1367
lemma abs_one [simp]: "abs 1 = (1::'a::ordered_idom)"
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1368
  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
  1369
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1370
lemma abs_le_mult: "abs (a * b) \<le> (abs a) * (abs (b::'a::lordered_ring))" 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1371
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1372
  let ?x = "pprt a * pprt b - pprt a * nprt b - nprt a * pprt b + nprt a * nprt b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1373
  let ?y = "pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1374
  have a: "(abs a) * (abs b) = ?x"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1375
    by (simp only: abs_prts[of a] abs_prts[of b] ring_eq_simps)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1376
  {
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1377
    fix u v :: 'a
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1378
    have bh: "\<lbrakk>u = a; v = b\<rbrakk> \<Longrightarrow> u * v = ?y"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1379
      apply (subst prts[of u], subst prts[of v])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1380
      apply (simp add: left_distrib right_distrib add_ac) 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1381
      done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1382
  }
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1383
  note b = this[OF refl[of a] refl[of b]]
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1384
  note addm = add_mono[of "0::'a" _ "0::'a", simplified]
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1385
  note addm2 = add_mono[of _ "0::'a" _ "0::'a", simplified]
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1386
  have xy: "- ?x <= ?y"
14754
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1387
    apply (simp)
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1388
    apply (rule_tac y="0::'a" in order_trans)
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1389
    apply (rule addm2)+
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1390
    apply (simp_all add: mult_pos_le mult_neg_le)
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1391
    apply (rule addm)+
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1392
    apply (simp_all add: mult_pos_le mult_neg_le)
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1393
    done
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1394
  have yx: "?y <= ?x"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1395
    apply (simp add: add_ac)
14754
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1396
    apply (rule_tac y=0 in order_trans)
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1397
    apply (rule addm2, (simp add: mult_pos_neg_le mult_pos_neg2_le)+)
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1398
    apply (rule addm, (simp add: mult_pos_neg_le mult_pos_neg2_le)+)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1399
    done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1400
  have i1: "a*b <= abs a * abs b" by (simp only: a b yx)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1401
  have i2: "- (abs a * abs b) <= a*b" by (simp only: a b xy)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1402
  show ?thesis
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1403
    apply (rule abs_leI)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1404
    apply (simp add: i1)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1405
    apply (simp add: i2[simplified minus_le_iff])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1406
    done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1407
qed
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1408
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1409
lemma abs_eq_mult: 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1410
  assumes "(0 \<le> a \<or> a \<le> 0) \<and> (0 \<le> b \<or> b \<le> 0)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1411
  shows "abs (a*b) = abs a * abs (b::'a::lordered_ring)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1412
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1413
  have s: "(0 <= a*b) | (a*b <= 0)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1414
    apply (auto)    
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1415
    apply (rule_tac split_mult_pos_le)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1416
    apply (rule_tac contrapos_np[of "a*b <= 0"])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1417
    apply (simp)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1418
    apply (rule_tac split_mult_neg_le)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1419
    apply (insert prems)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1420
    apply (blast)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1421
    done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1422
  have mulprts: "a * b = (pprt a + nprt a) * (pprt b + nprt b)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1423
    by (simp add: prts[symmetric])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1424
  show ?thesis
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1425
  proof cases
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1426
    assume "0 <= a * b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1427
    then show ?thesis
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1428
      apply (simp_all add: mulprts abs_prts)
14754
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1429
      apply (simp add: 
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1430
	iff2imp[OF zero_le_iff_zero_nprt]
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1431
	iff2imp[OF le_zero_iff_pprt_id]
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1432
      )
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1433
      apply (insert prems)
14754
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1434
      apply (auto simp add: 
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1435
	ring_eq_simps 
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  1436
	iff2imp[OF zero_le_iff_zero_nprt] iff2imp[OF le_zero_iff_zero_pprt]
15197
19e735596e51 Added antisymmetry simproc
nipkow
parents: 15178
diff changeset
  1437
	iff2imp[OF le_zero_iff_pprt_id] iff2imp[OF zero_le_iff_nprt_id])
19e735596e51 Added antisymmetry simproc
nipkow
parents: 15178
diff changeset
  1438
	apply(drule (1) mult_pos_neg_le[of a b], simp)
19e735596e51 Added antisymmetry simproc
nipkow
parents: 15178
diff changeset
  1439
	apply(drule (1) mult_pos_neg2_le[of b a], simp)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1440
      done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1441
  next
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1442
    assume "~(0 <= a*b)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1443
    with s have "a*b <= 0" by simp
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1444
    then show ?thesis
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1445
      apply (simp_all add: mulprts abs_prts)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1446
      apply (insert prems)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1447
      apply (auto simp add: ring_eq_simps iff2imp[OF zero_le_iff_zero_nprt] iff2imp[OF le_zero_iff_zero_pprt]
15197
19e735596e51 Added antisymmetry simproc
nipkow
parents: 15178
diff changeset
  1448
	iff2imp[OF le_zero_iff_pprt_id] iff2imp[OF zero_le_iff_nprt_id])
19e735596e51 Added antisymmetry simproc
nipkow
parents: 15178
diff changeset
  1449
      apply(drule (1) mult_pos_le[of a b],simp)
19e735596e51 Added antisymmetry simproc
nipkow
parents: 15178
diff changeset
  1450
      apply(drule (1) mult_neg_le[of a b],simp)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1451
      done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1452
  qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1453
qed
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1454
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1455
lemma abs_mult: "abs (a * b) = abs a * abs (b::'a::ordered_idom)" 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1456
by (simp add: abs_eq_mult linorder_linear)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1457
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1458
lemma abs_mult_self: "abs a * abs a = a * (a::'a::ordered_idom)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1459
by (simp add: abs_if) 
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1460
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1461
lemma nonzero_abs_inverse:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1462
     "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
  1463
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
  1464
                      negative_imp_inverse_negative)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1465
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
  1466
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1467
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1468
lemma abs_inverse [simp]:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1469
     "abs (inverse (a::'a::{ordered_field,division_by_zero})) = 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1470
      inverse (abs a)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1471
apply (case_tac "a=0", simp) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1472
apply (simp add: nonzero_abs_inverse) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1473
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1474
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1475
lemma nonzero_abs_divide:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1476
     "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
  1477
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
  1478
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1479
lemma abs_divide:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1480
     "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
  1481
apply (case_tac "b=0", simp) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1482
apply (simp add: nonzero_abs_divide) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1483
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1484
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1485
lemma abs_mult_less:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1486
     "[| abs a < c; abs b < d |] ==> abs a * abs b < c*(d::'a::ordered_idom)"
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1487
proof -
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1488
  assume ac: "abs a < c"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1489
  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
  1490
  assume "abs b < d"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1491
  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
  1492
qed
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1493
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1494
lemma eq_minus_self_iff: "(a = -a) = (a = (0::'a::ordered_idom))"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1495
by (force simp add: order_eq_iff le_minus_self_iff minus_le_self_iff)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1496
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1497
lemma less_minus_self_iff: "(a < -a) = (a < (0::'a::ordered_idom))"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1498
by (simp add: order_less_le le_minus_self_iff eq_minus_self_iff)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1499
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1500
lemma abs_less_iff: "(abs a < b) = (a < b & -a < (b::'a::ordered_idom))" 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1501
apply (simp add: order_less_le abs_le_iff)  
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1502
apply (auto simp add: abs_if minus_le_self_iff eq_minus_self_iff)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1503
apply (simp add: le_minus_self_iff linorder_neq_iff) 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1504
done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1505
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1506
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
  1507
declare times_divide_eq_left [simp]
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1508
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1509
lemma linprog_dual_estimate:
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1510
  assumes
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1511
  "A * x \<le> (b::'a::lordered_ring)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1512
  "0 \<le> y"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1513
  "abs (A - A') \<le> \<delta>A"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1514
  "b \<le> b'"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1515
  "abs (c - c') \<le> \<delta>c"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1516
  "abs x \<le> r"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1517
  shows
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1518
  "c * x \<le> y * b' + (y * \<delta>A + abs (y * A' - c') + \<delta>c) * r"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1519
proof -
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1520
  from prems have 1: "y * b <= y * b'" by (simp add: mult_left_mono)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1521
  from prems have 2: "y * (A * x) <= y * b" by (simp add: mult_left_mono) 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1522
  have 3: "y * (A * x) = c * x + (y * (A - A') + (y * A' - c') + (c'-c)) * x" by (simp add: ring_eq_simps)  
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1523
  from 1 2 3 have 4: "c * x + (y * (A - A') + (y * A' - c') + (c'-c)) * x <= y * b'" by simp
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1524
  have 5: "c * x <= y * b' + abs((y * (A - A') + (y * A' - c') + (c'-c)) * x)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1525
    by (simp only: 4 estimate_by_abs)  
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1526
  have 6: "abs((y * (A - A') + (y * A' - c') + (c'-c)) * x) <= abs (y * (A - A') + (y * A' - c') + (c'-c)) * abs x"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1527
    by (simp add: abs_le_mult)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1528
  have 7: "(abs (y * (A - A') + (y * A' - c') + (c'-c))) * abs x <= (abs (y * (A-A') + (y*A'-c')) + abs(c'-c)) * abs x"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1529
    by (simp add: abs_triangle_ineq mult_right_mono)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1530
  have 8: " (abs (y * (A-A') + (y*A'-c')) + abs(c'-c)) * abs x <=  (abs (y * (A-A')) + abs (y*A'-c') + abs(c'-c)) * abs x"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1531
    by (simp add: abs_triangle_ineq mult_right_mono)    
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1532
  have 9: "(abs (y * (A-A')) + abs (y*A'-c') + abs(c'-c)) * abs x <= (abs y * abs (A-A') + abs (y*A'-c') + abs (c'-c)) * abs x"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1533
    by (simp add: abs_le_mult mult_right_mono)  
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1534
  have 10: "c'-c = -(c-c')" by (simp add: ring_eq_simps)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1535
  have 11: "abs (c'-c) = abs (c-c')" 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1536
    by (subst 10, subst abs_minus_cancel, simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1537
  have 12: "(abs y * abs (A-A') + abs (y*A'-c') + abs (c'-c)) * abs x <= (abs y * abs (A-A') + abs (y*A'-c') + \<delta>c) * abs x"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1538
    by (simp add: 11 prems mult_right_mono)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1539
  have 13: "(abs y * abs (A-A') + abs (y*A'-c') + \<delta>c) * abs x <= (abs y * \<delta>A + abs (y*A'-c') + \<delta>c) * abs x"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1540
    by (simp add: prems mult_right_mono mult_left_mono)  
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1541
  have r: "(abs y * \<delta>A + abs (y*A'-c') + \<delta>c) * abs x <=  (abs y * \<delta>A + abs (y*A'-c') + \<delta>c) * r"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1542
    apply (rule mult_left_mono)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1543
    apply (simp add: prems)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1544
    apply (rule_tac add_mono[of "0::'a" _ "0", simplified])+
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1545
    apply (rule mult_left_mono[of "0" "\<delta>A", simplified])
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1546
    apply (simp_all)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1547
    apply (rule order_trans[where y="abs (A-A')"], simp_all add: prems)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1548
    apply (rule order_trans[where y="abs (c-c')"], simp_all add: prems)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1549
    done    
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1550
  from 6 7 8 9 12 13 r have 14:" abs((y * (A - A') + (y * A' - c') + (c'-c)) * x) <=(abs y * \<delta>A + abs (y*A'-c') + \<delta>c) * r"     
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1551
    by (simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1552
  show ?thesis 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1553
    apply (rule_tac le_add_right_mono[of _ _ "abs((y * (A - A') + (y * A' - c') + (c'-c)) * x)"])
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1554
    apply (simp_all add: 5 14[simplified abs_of_ge_0[of y, simplified prems]])
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1555
    done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1556
qed
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1557
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1558
lemma le_ge_imp_abs_diff_1:
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1559
  assumes
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1560
  "A1 <= (A::'a::lordered_ring)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1561
  "A <= A2" 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1562
  shows "abs (A-A1) <= A2-A1"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1563
proof -
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1564
  have "0 <= A - A1"    
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1565
  proof -
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1566
    have 1: "A - A1 = A + (- A1)" by simp
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1567
    show ?thesis by (simp only: 1 add_right_mono[of A1 A "-A1", simplified, simplified prems])
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1568
  qed
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1569
  then have "abs (A-A1) = A-A1" by (rule abs_of_ge_0)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1570
  with prems show "abs (A-A1) <= (A2-A1)" by simp
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1571
qed
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1572
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1573
lemma linprog_dual_estimate_1:
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1574
  assumes
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1575
  "A * x \<le> (b::'a::lordered_ring)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1576
  "0 \<le> y"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1577
  "A1 <= A"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1578
  "A <= A2"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1579
  "c1 <= c"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1580
  "c <= c2"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1581
  "abs x \<le> r"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1582
  shows
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1583
  "c * x \<le> y * b + (y * (A2 - A1) + abs (y * A1 - c1) + (c2 - c1)) * r"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1584
proof -
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1585
  from prems have delta_A: "abs (A-A1) <= (A2-A1)" by (simp add: le_ge_imp_abs_diff_1)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1586
  from prems have delta_c: "abs (c-c1) <= (c2-c1)" by (simp add: le_ge_imp_abs_diff_1)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1587
  show ?thesis
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1588
    apply (rule_tac linprog_dual_estimate)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1589
    apply (auto intro: delta_A delta_c simp add: prems)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1590
    done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1591
qed
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  1592
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1593
ML {*
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1594
val left_distrib = thm "left_distrib";
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1595
val right_distrib = thm "right_distrib";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1596
val mult_commute = thm "mult_commute";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1597
val distrib = thm "distrib";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1598
val zero_neq_one = thm "zero_neq_one";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1599
val no_zero_divisors = thm "no_zero_divisors";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1600
val left_inverse = thm "left_inverse";
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1601
val divide_inverse = thm "divide_inverse";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1602
val mult_zero_left = thm "mult_zero_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1603
val mult_zero_right = thm "mult_zero_right";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1604
val field_mult_eq_0_iff = thm "field_mult_eq_0_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1605
val inverse_zero = thm "inverse_zero";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1606
val ring_distrib = thms "ring_distrib";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1607
val combine_common_factor = thm "combine_common_factor";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1608
val minus_mult_left = thm "minus_mult_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1609
val minus_mult_right = thm "minus_mult_right";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1610
val minus_mult_minus = thm "minus_mult_minus";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1611
val minus_mult_commute = thm "minus_mult_commute";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1612
val right_diff_distrib = thm "right_diff_distrib";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1613
val left_diff_distrib = thm "left_diff_distrib";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1614
val mult_left_mono = thm "mult_left_mono";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1615
val mult_right_mono = thm "mult_right_mono";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1616
val mult_strict_left_mono = thm "mult_strict_left_mono";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1617
val mult_strict_right_mono = thm "mult_strict_right_mono";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1618
val mult_mono = thm "mult_mono";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1619
val mult_strict_mono = thm "mult_strict_mono";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1620
val abs_if = thm "abs_if";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1621
val zero_less_one = thm "zero_less_one";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1622
val eq_add_iff1 = thm "eq_add_iff1";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1623
val eq_add_iff2 = thm "eq_add_iff2";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1624
val less_add_iff1 = thm "less_add_iff1";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1625
val less_add_iff2 = thm "less_add_iff2";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1626
val le_add_iff1 = thm "le_add_iff1";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1627
val le_add_iff2 = thm "le_add_iff2";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1628
val mult_left_le_imp_le = thm "mult_left_le_imp_le";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1629
val mult_right_le_imp_le = thm "mult_right_le_imp_le";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1630
val mult_left_less_imp_less = thm "mult_left_less_imp_less";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1631
val mult_right_less_imp_less = thm "mult_right_less_imp_less";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1632
val mult_strict_left_mono_neg = thm "mult_strict_left_mono_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1633
val mult_left_mono_neg = thm "mult_left_mono_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1634
val mult_strict_right_mono_neg = thm "mult_strict_right_mono_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1635
val mult_right_mono_neg = thm "mult_right_mono_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1636
val mult_pos = thm "mult_pos";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1637
val mult_pos_le = thm "mult_pos_le";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1638
val mult_pos_neg = thm "mult_pos_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1639
val mult_pos_neg_le = thm "mult_pos_neg_le";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1640
val mult_pos_neg2 = thm "mult_pos_neg2";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1641
val mult_pos_neg2_le = thm "mult_pos_neg2_le";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1642
val mult_neg = thm "mult_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1643
val mult_neg_le = thm "mult_neg_le";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1644
val zero_less_mult_pos = thm "zero_less_mult_pos";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1645
val zero_less_mult_pos2 = thm "zero_less_mult_pos2";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1646
val zero_less_mult_iff = thm "zero_less_mult_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1647
val mult_eq_0_iff = thm "mult_eq_0_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1648
val zero_le_mult_iff = thm "zero_le_mult_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1649
val mult_less_0_iff = thm "mult_less_0_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1650
val mult_le_0_iff = thm "mult_le_0_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1651
val split_mult_pos_le = thm "split_mult_pos_le";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1652
val split_mult_neg_le = thm "split_mult_neg_le";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1653
val zero_le_square = thm "zero_le_square";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1654
val zero_le_one = thm "zero_le_one";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1655
val not_one_le_zero = thm "not_one_le_zero";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1656
val not_one_less_zero = thm "not_one_less_zero";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1657
val mult_left_mono_neg = thm "mult_left_mono_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1658
val mult_right_mono_neg = thm "mult_right_mono_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1659
val mult_strict_mono = thm "mult_strict_mono";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1660
val mult_strict_mono' = thm "mult_strict_mono'";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1661
val mult_mono = thm "mult_mono";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1662
val less_1_mult = thm "less_1_mult";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1663
val mult_less_cancel_right = thm "mult_less_cancel_right";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1664
val mult_less_cancel_left = thm "mult_less_cancel_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1665
val mult_le_cancel_right = thm "mult_le_cancel_right";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1666
val mult_le_cancel_left = thm "mult_le_cancel_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1667
val mult_less_imp_less_left = thm "mult_less_imp_less_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1668
val mult_less_imp_less_right = thm "mult_less_imp_less_right";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1669
val mult_cancel_right = thm "mult_cancel_right";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1670
val mult_cancel_left = thm "mult_cancel_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1671
val ring_eq_simps = thms "ring_eq_simps";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1672
val right_inverse = thm "right_inverse";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1673
val right_inverse_eq = thm "right_inverse_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1674
val nonzero_inverse_eq_divide = thm "nonzero_inverse_eq_divide";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1675
val divide_self = thm "divide_self";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1676
val divide_zero = thm "divide_zero";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1677
val divide_zero_left = thm "divide_zero_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1678
val inverse_eq_divide = thm "inverse_eq_divide";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1679
val add_divide_distrib = thm "add_divide_distrib";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1680
val field_mult_eq_0_iff = thm "field_mult_eq_0_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1681
val field_mult_cancel_right_lemma = thm "field_mult_cancel_right_lemma";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1682
val field_mult_cancel_right = thm "field_mult_cancel_right";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1683
val field_mult_cancel_left = thm "field_mult_cancel_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1684
val nonzero_imp_inverse_nonzero = thm "nonzero_imp_inverse_nonzero";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1685
val inverse_zero_imp_zero = thm "inverse_zero_imp_zero";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1686
val inverse_nonzero_imp_nonzero = thm "inverse_nonzero_imp_nonzero";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1687
val inverse_nonzero_iff_nonzero = thm "inverse_nonzero_iff_nonzero";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1688
val nonzero_inverse_minus_eq = thm "nonzero_inverse_minus_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1689
val inverse_minus_eq = thm "inverse_minus_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1690
val nonzero_inverse_eq_imp_eq = thm "nonzero_inverse_eq_imp_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1691
val inverse_eq_imp_eq = thm "inverse_eq_imp_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1692
val inverse_eq_iff_eq = thm "inverse_eq_iff_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1693
val nonzero_inverse_inverse_eq = thm "nonzero_inverse_inverse_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1694
val inverse_inverse_eq = thm "inverse_inverse_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1695
val inverse_1 = thm "inverse_1";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1696
val nonzero_inverse_mult_distrib = thm "nonzero_inverse_mult_distrib";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1697
val inverse_mult_distrib = thm "inverse_mult_distrib";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1698
val inverse_add = thm "inverse_add";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1699
val inverse_divide = thm "inverse_divide";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1700
val nonzero_mult_divide_cancel_left = thm "nonzero_mult_divide_cancel_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1701
val mult_divide_cancel_left = thm "mult_divide_cancel_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1702
val nonzero_mult_divide_cancel_right = thm "nonzero_mult_divide_cancel_right";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1703
val mult_divide_cancel_right = thm "mult_divide_cancel_right";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1704
val mult_divide_cancel_eq_if = thm "mult_divide_cancel_eq_if";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1705
val divide_1 = thm "divide_1";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1706
val times_divide_eq_right = thm "times_divide_eq_right";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1707
val times_divide_eq_left = thm "times_divide_eq_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1708
val divide_divide_eq_right = thm "divide_divide_eq_right";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1709
val divide_divide_eq_left = thm "divide_divide_eq_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1710
val nonzero_minus_divide_left = thm "nonzero_minus_divide_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1711
val nonzero_minus_divide_right = thm "nonzero_minus_divide_right";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1712
val nonzero_minus_divide_divide = thm "nonzero_minus_divide_divide";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1713
val minus_divide_left = thm "minus_divide_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1714
val minus_divide_right = thm "minus_divide_right";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1715
val minus_divide_divide = thm "minus_divide_divide";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1716
val diff_divide_distrib = thm "diff_divide_distrib";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1717
val positive_imp_inverse_positive = thm "positive_imp_inverse_positive";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1718
val negative_imp_inverse_negative = thm "negative_imp_inverse_negative";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1719
val inverse_le_imp_le = thm "inverse_le_imp_le";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1720
val inverse_positive_imp_positive = thm "inverse_positive_imp_positive";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1721
val inverse_positive_iff_positive = thm "inverse_positive_iff_positive";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1722
val inverse_negative_imp_negative = thm "inverse_negative_imp_negative";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1723
val inverse_negative_iff_negative = thm "inverse_negative_iff_negative";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1724
val inverse_nonnegative_iff_nonnegative = thm "inverse_nonnegative_iff_nonnegative";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1725
val inverse_nonpositive_iff_nonpositive = thm "inverse_nonpositive_iff_nonpositive";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1726
val less_imp_inverse_less = thm "less_imp_inverse_less";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1727
val inverse_less_imp_less = thm "inverse_less_imp_less";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1728
val inverse_less_iff_less = thm "inverse_less_iff_less";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1729
val le_imp_inverse_le = thm "le_imp_inverse_le";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1730
val inverse_le_iff_le = thm "inverse_le_iff_le";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1731
val inverse_le_imp_le_neg = thm "inverse_le_imp_le_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1732
val less_imp_inverse_less_neg = thm "less_imp_inverse_less_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1733
val inverse_less_imp_less_neg = thm "inverse_less_imp_less_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1734
val inverse_less_iff_less_neg = thm "inverse_less_iff_less_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1735
val le_imp_inverse_le_neg = thm "le_imp_inverse_le_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1736
val inverse_le_iff_le_neg = thm "inverse_le_iff_le_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1737
val one_less_inverse_iff = thm "one_less_inverse_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1738
val inverse_eq_1_iff = thm "inverse_eq_1_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1739
val one_le_inverse_iff = thm "one_le_inverse_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1740
val inverse_less_1_iff = thm "inverse_less_1_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1741
val inverse_le_1_iff = thm "inverse_le_1_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1742
val zero_less_divide_iff = thm "zero_less_divide_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1743
val divide_less_0_iff = thm "divide_less_0_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1744
val zero_le_divide_iff = thm "zero_le_divide_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1745
val divide_le_0_iff = thm "divide_le_0_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1746
val divide_eq_0_iff = thm "divide_eq_0_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1747
val pos_le_divide_eq = thm "pos_le_divide_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1748
val neg_le_divide_eq = thm "neg_le_divide_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1749
val le_divide_eq = thm "le_divide_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1750
val pos_divide_le_eq = thm "pos_divide_le_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1751
val neg_divide_le_eq = thm "neg_divide_le_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1752
val divide_le_eq = thm "divide_le_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1753
val pos_less_divide_eq = thm "pos_less_divide_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1754
val neg_less_divide_eq = thm "neg_less_divide_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1755
val less_divide_eq = thm "less_divide_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1756
val pos_divide_less_eq = thm "pos_divide_less_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1757
val neg_divide_less_eq = thm "neg_divide_less_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1758
val divide_less_eq = thm "divide_less_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1759
val nonzero_eq_divide_eq = thm "nonzero_eq_divide_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1760
val eq_divide_eq = thm "eq_divide_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1761
val nonzero_divide_eq_eq = thm "nonzero_divide_eq_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1762
val divide_eq_eq = thm "divide_eq_eq";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1763
val divide_cancel_right = thm "divide_cancel_right";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1764
val divide_cancel_left = thm "divide_cancel_left";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1765
val divide_eq_1_iff = thm "divide_eq_1_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1766
val one_eq_divide_iff = thm "one_eq_divide_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1767
val zero_eq_1_divide_iff = thm "zero_eq_1_divide_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1768
val one_divide_eq_0_iff = thm "one_divide_eq_0_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1769
val divide_strict_right_mono = thm "divide_strict_right_mono";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1770
val divide_right_mono = thm "divide_right_mono";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1771
val divide_strict_left_mono = thm "divide_strict_left_mono";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1772
val divide_left_mono = thm "divide_left_mono";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1773
val divide_strict_left_mono_neg = thm "divide_strict_left_mono_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1774
val divide_strict_right_mono_neg = thm "divide_strict_right_mono_neg";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1775
val less_add_one = thm "less_add_one";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1776
val zero_less_two = thm "zero_less_two";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1777
val less_half_sum = thm "less_half_sum";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1778
val gt_half_sum = thm "gt_half_sum";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1779
val dense = thm "dense";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1780
val abs_one = thm "abs_one";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1781
val abs_le_mult = thm "abs_le_mult";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1782
val abs_eq_mult = thm "abs_eq_mult";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1783
val abs_mult = thm "abs_mult";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1784
val abs_mult_self = thm "abs_mult_self";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1785
val nonzero_abs_inverse = thm "nonzero_abs_inverse";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1786
val abs_inverse = thm "abs_inverse";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1787
val nonzero_abs_divide = thm "nonzero_abs_divide";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1788
val abs_divide = thm "abs_divide";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1789
val abs_mult_less = thm "abs_mult_less";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1790
val eq_minus_self_iff = thm "eq_minus_self_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1791
val less_minus_self_iff = thm "less_minus_self_iff";
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1792
val abs_less_iff = thm "abs_less_iff";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1793
*}
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14321
diff changeset
  1794
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1795
end