| author | paulson | 
| Thu, 05 Oct 2006 10:46:26 +0200 | |
| changeset 20864 | bb75b876b260 | 
| parent 20633 | e98f59806244 | 
| child 21199 | 2d83f93c3580 | 
| permissions | -rw-r--r-- | 
| 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$ | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 3 | Author: Gertrud Bauer, Steven Obua, Lawrence C Paulson, and Markus Wenzel, | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 4 | with contributions by Jeremy Avigad | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 5 | *) | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 6 | |
| 14738 | 7 | header {* (Ordered) Rings and Fields *}
 | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 8 | |
| 15229 | 9 | theory Ring_and_Field | 
| 15140 | 10 | imports OrderedGroup | 
| 15131 | 11 | begin | 
| 14504 | 12 | |
| 14738 | 13 | text {*
 | 
| 14 | The theory of partially ordered rings is taken from the books: | |
| 15 |   \begin{itemize}
 | |
| 16 |   \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
 | |
| 17 |   \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
 | |
| 18 |   \end{itemize}
 | |
| 19 | Most of the used notions can also be looked up in | |
| 20 |   \begin{itemize}
 | |
| 14770 | 21 |   \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
 | 
| 14738 | 22 |   \item \emph{Algebra I} by van der Waerden, Springer.
 | 
| 23 |   \end{itemize}
 | |
| 24 | *} | |
| 14504 | 25 | |
| 14738 | 26 | axclass semiring \<subseteq> ab_semigroup_add, semigroup_mult | 
| 27 | left_distrib: "(a + b) * c = a * c + b * c" | |
| 28 | right_distrib: "a * (b + c) = a * b + a * c" | |
| 14504 | 29 | |
| 14738 | 30 | axclass semiring_0 \<subseteq> semiring, comm_monoid_add | 
| 14504 | 31 | |
| 14940 | 32 | axclass semiring_0_cancel \<subseteq> semiring_0, cancel_ab_semigroup_add | 
| 33 | ||
| 14738 | 34 | axclass comm_semiring \<subseteq> ab_semigroup_add, ab_semigroup_mult | 
| 35 | distrib: "(a + b) * c = a * c + b * c" | |
| 14504 | 36 | |
| 14738 | 37 | instance comm_semiring \<subseteq> semiring | 
| 38 | proof | |
| 39 | fix a b c :: 'a | |
| 40 | show "(a + b) * c = a * c + b * c" by (simp add: distrib) | |
| 41 | have "a * (b + c) = (b + c) * a" by (simp add: mult_ac) | |
| 42 | also have "... = b * a + c * a" by (simp only: distrib) | |
| 43 | also have "... = a * b + a * c" by (simp add: mult_ac) | |
| 44 | finally show "a * (b + c) = a * b + a * c" by blast | |
| 14504 | 45 | qed | 
| 46 | ||
| 14738 | 47 | axclass comm_semiring_0 \<subseteq> comm_semiring, comm_monoid_add | 
| 14504 | 48 | |
| 14738 | 49 | instance comm_semiring_0 \<subseteq> semiring_0 .. | 
| 14504 | 50 | |
| 14940 | 51 | axclass comm_semiring_0_cancel \<subseteq> comm_semiring_0, cancel_ab_semigroup_add | 
| 52 | ||
| 53 | instance comm_semiring_0_cancel \<subseteq> semiring_0_cancel .. | |
| 54 | ||
| 20609 | 55 | axclass zero_neq_one \<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 | |
| 20609 | 58 | axclass semiring_1 \<subseteq> zero_neq_one, semiring_0, monoid_mult | 
| 14504 | 59 | |
| 20609 | 60 | axclass comm_semiring_1 \<subseteq> zero_neq_one, comm_semiring_0, comm_monoid_mult (* previously almost_semiring *) | 
| 14738 | 61 | |
| 62 | instance comm_semiring_1 \<subseteq> semiring_1 .. | |
| 14421 
ee97b6463cb4
new Ring_and_Field hierarchy, eliminating redundant axioms
 paulson parents: 
14398diff
changeset | 63 | |
| 20609 | 64 | axclass no_zero_divisors \<subseteq> zero, times | 
| 14738 | 65 | no_zero_divisors: "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a * b \<noteq> 0" | 
| 14504 | 66 | |
| 14940 | 67 | axclass semiring_1_cancel \<subseteq> semiring_1, cancel_ab_semigroup_add | 
| 68 | ||
| 69 | instance semiring_1_cancel \<subseteq> semiring_0_cancel .. | |
| 70 | ||
| 14738 | 71 | axclass comm_semiring_1_cancel \<subseteq> comm_semiring_1, cancel_ab_semigroup_add (* previously semiring *) | 
| 72 | ||
| 14940 | 73 | instance comm_semiring_1_cancel \<subseteq> semiring_1_cancel .. | 
| 74 | ||
| 75 | instance comm_semiring_1_cancel \<subseteq> comm_semiring_0_cancel .. | |
| 76 | ||
| 14738 | 77 | axclass ring \<subseteq> semiring, ab_group_add | 
| 78 | ||
| 14940 | 79 | instance ring \<subseteq> semiring_0_cancel .. | 
| 14504 | 80 | |
| 14738 | 81 | axclass comm_ring \<subseteq> comm_semiring_0, ab_group_add | 
| 82 | ||
| 83 | instance comm_ring \<subseteq> ring .. | |
| 14504 | 84 | |
| 14940 | 85 | instance comm_ring \<subseteq> comm_semiring_0_cancel .. | 
| 14738 | 86 | |
| 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 | 89 | instance ring_1 \<subseteq> semiring_1_cancel .. | 
| 90 | ||
| 14738 | 91 | axclass comm_ring_1 \<subseteq> comm_ring, comm_semiring_1 (* previously ring *) | 
| 92 | ||
| 93 | instance comm_ring_1 \<subseteq> ring_1 .. | |
| 14421 
ee97b6463cb4
new Ring_and_Field hierarchy, eliminating redundant axioms
 paulson parents: 
14398diff
changeset | 94 | |
| 14738 | 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 | |
| 20609 | 97 | axclass idom \<subseteq> comm_ring_1, no_zero_divisors | 
| 14421 
ee97b6463cb4
new Ring_and_Field hierarchy, eliminating redundant axioms
 paulson parents: 
14398diff
changeset | 98 | |
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 99 | axclass division_ring \<subseteq> ring_1, inverse | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 100 | left_inverse [simp]: "a \<noteq> 0 ==> inverse a * a = 1" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 101 | right_inverse [simp]: "a \<noteq> 0 ==> a * inverse a = 1" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 102 | |
| 14738 | 103 | axclass field \<subseteq> comm_ring_1, inverse | 
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 104 | field_left_inverse: "a \<noteq> 0 ==> inverse a * a = 1" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 105 | divide_inverse: "a / b = a * inverse b" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 106 | |
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 107 | lemma field_right_inverse: | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 108 | assumes not0: "a \<noteq> 0" shows "a * inverse (a::'a::field) = 1" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 109 | proof - | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 110 | have "a * inverse a = inverse a * a" by (rule mult_commute) | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 111 | also have "... = 1" using not0 by (rule field_left_inverse) | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 112 | finally show ?thesis . | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 113 | qed | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 114 | |
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 115 | instance field \<subseteq> division_ring | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 116 | by (intro_classes, erule field_left_inverse, erule field_right_inverse) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 117 | |
| 14940 | 118 | lemma mult_zero_left [simp]: "0 * a = (0::'a::semiring_0_cancel)" | 
| 14738 | 119 | proof - | 
| 120 | have "0*a + 0*a = 0*a + 0" | |
| 121 | by (simp add: left_distrib [symmetric]) | |
| 122 | thus ?thesis | |
| 123 | by (simp only: add_left_cancel) | |
| 124 | qed | |
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 125 | |
| 14940 | 126 | lemma mult_zero_right [simp]: "a * 0 = (0::'a::semiring_0_cancel)" | 
| 14738 | 127 | proof - | 
| 128 | have "a*0 + a*0 = a*0 + 0" | |
| 129 | by (simp add: right_distrib [symmetric]) | |
| 130 | thus ?thesis | |
| 131 | by (simp only: add_left_cancel) | |
| 132 | qed | |
| 133 | ||
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 134 | lemma field_mult_eq_0_iff [simp]: | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 135 | "(a*b = (0::'a::division_ring)) = (a = 0 | b = 0)" | 
| 14738 | 136 | proof cases | 
| 137 | assume "a=0" thus ?thesis by simp | |
| 138 | next | |
| 139 | assume anz [simp]: "a\<noteq>0" | |
| 140 |   { assume "a * b = 0"
 | |
| 141 | hence "inverse a * (a * b) = 0" by simp | |
| 142 | hence "b = 0" by (simp (no_asm_use) add: mult_assoc [symmetric])} | |
| 143 | thus ?thesis by force | |
| 144 | qed | |
| 145 | ||
| 146 | instance field \<subseteq> idom | |
| 147 | by (intro_classes, simp) | |
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 148 | |
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 149 | axclass division_by_zero \<subseteq> zero, inverse | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 150 | 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 | 151 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 152 | subsection {* Distribution rules *}
 | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 153 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 154 | 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 | 155 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 156 | text{*For the @{text combine_numerals} simproc*}
 | 
| 14421 
ee97b6463cb4
new Ring_and_Field hierarchy, eliminating redundant axioms
 paulson parents: 
14398diff
changeset | 157 | lemma combine_common_factor: | 
| 14738 | 158 | "a*e + (b*e + c) = (a+b)*e + (c::'a::semiring)" | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 159 | by (simp add: left_distrib add_ac) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 160 | |
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 161 | 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 | 162 | apply (rule equals_zero_I) | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 163 | apply (simp add: left_distrib [symmetric]) | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 164 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 165 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 166 | 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 | 167 | apply (rule equals_zero_I) | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 168 | apply (simp add: right_distrib [symmetric]) | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 169 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 170 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 171 | 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: 
14267diff
changeset | 172 | 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: 
14267diff
changeset | 173 | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 174 | lemma minus_mult_commute: "(- a) * b = a * (- b::'a::ring)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 175 | by (simp add: minus_mult_left [symmetric] minus_mult_right [symmetric]) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 176 | |
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 177 | 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 | 178 | 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 | 179 | 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 | 180 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 181 | lemma left_diff_distrib: "(a - b) * c = a * c - b * (c::'a::ring)" | 
| 14738 | 182 | by (simp add: left_distrib diff_minus | 
| 183 | 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 | 184 | |
| 14738 | 185 | axclass pordered_semiring \<subseteq> semiring_0, pordered_ab_semigroup_add | 
| 186 | mult_left_mono: "a <= b \<Longrightarrow> 0 <= c \<Longrightarrow> c * a <= c * b" | |
| 187 | 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: 
14266diff
changeset | 188 | |
| 14738 | 189 | 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: 
14267diff
changeset | 190 | |
| 14940 | 191 | instance pordered_cancel_semiring \<subseteq> semiring_0_cancel .. | 
| 192 | ||
| 14738 | 193 | axclass ordered_semiring_strict \<subseteq> semiring_0, ordered_cancel_ab_semigroup_add | 
| 194 | mult_strict_left_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b" | |
| 195 | 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: 
14334diff
changeset | 196 | |
| 14940 | 197 | instance ordered_semiring_strict \<subseteq> semiring_0_cancel .. | 
| 198 | ||
| 14738 | 199 | instance ordered_semiring_strict \<subseteq> pordered_cancel_semiring | 
| 200 | apply intro_classes | |
| 201 | apply (case_tac "a < b & 0 < c") | |
| 202 | apply (auto simp add: mult_strict_left_mono order_less_le) | |
| 203 | apply (auto simp add: mult_strict_left_mono order_le_less) | |
| 204 | apply (simp add: mult_strict_right_mono) | |
| 14270 | 205 | done | 
| 206 | ||
| 14738 | 207 | axclass pordered_comm_semiring \<subseteq> comm_semiring_0, pordered_ab_semigroup_add | 
| 208 | mult_mono: "a <= b \<Longrightarrow> 0 <= c \<Longrightarrow> c * a <= c * b" | |
| 14270 | 209 | |
| 14738 | 210 | axclass pordered_cancel_comm_semiring \<subseteq> pordered_comm_semiring, cancel_ab_semigroup_add | 
| 14270 | 211 | |
| 14738 | 212 | instance pordered_cancel_comm_semiring \<subseteq> pordered_comm_semiring .. | 
| 14270 | 213 | |
| 14738 | 214 | axclass ordered_comm_semiring_strict \<subseteq> comm_semiring_0, ordered_cancel_ab_semigroup_add | 
| 215 | 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 | 216 | |
| 14738 | 217 | instance pordered_comm_semiring \<subseteq> pordered_semiring | 
| 218 | 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 | 219 | |
| 14738 | 220 | 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 | 221 | |
| 14738 | 222 | instance ordered_comm_semiring_strict \<subseteq> ordered_semiring_strict | 
| 223 | 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: 
14270diff
changeset | 224 | |
| 14738 | 225 | instance ordered_comm_semiring_strict \<subseteq> pordered_cancel_comm_semiring | 
| 226 | apply (intro_classes) | |
| 227 | apply (case_tac "a < b & 0 < c") | |
| 228 | apply (auto simp add: mult_strict_left_mono order_less_le) | |
| 229 | apply (auto simp add: mult_strict_left_mono order_le_less) | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 230 | done | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 231 | |
| 14738 | 232 | axclass pordered_ring \<subseteq> ring, pordered_semiring | 
| 14270 | 233 | |
| 14738 | 234 | instance pordered_ring \<subseteq> pordered_ab_group_add .. | 
| 14270 | 235 | |
| 14738 | 236 | instance pordered_ring \<subseteq> pordered_cancel_semiring .. | 
| 14270 | 237 | |
| 14738 | 238 | axclass lordered_ring \<subseteq> pordered_ring, lordered_ab_group_abs | 
| 14270 | 239 | |
| 14940 | 240 | instance lordered_ring \<subseteq> lordered_ab_group_meet .. | 
| 241 | ||
| 242 | instance lordered_ring \<subseteq> lordered_ab_group_join .. | |
| 243 | ||
| 20633 | 244 | axclass abs_if \<subseteq> minus, ord, zero | 
| 14738 | 245 | abs_if: "abs a = (if (a < 0) then (-a) else a)" | 
| 14270 | 246 | |
| 20633 | 247 | axclass ordered_ring_strict \<subseteq> ring, ordered_semiring_strict, abs_if | 
| 14270 | 248 | |
| 14738 | 249 | instance ordered_ring_strict \<subseteq> lordered_ab_group .. | 
| 14270 | 250 | |
| 14738 | 251 | instance ordered_ring_strict \<subseteq> lordered_ring | 
| 252 | by (intro_classes, simp add: abs_if join_eq_if) | |
| 14270 | 253 | |
| 14738 | 254 | axclass pordered_comm_ring \<subseteq> comm_ring, pordered_comm_semiring | 
| 14270 | 255 | |
| 14738 | 256 | axclass ordered_semidom \<subseteq> comm_semiring_1_cancel, ordered_comm_semiring_strict (* previously ordered_semiring *) | 
| 257 | zero_less_one [simp]: "0 < 1" | |
| 14270 | 258 | |
| 20633 | 259 | axclass ordered_idom \<subseteq> comm_ring_1, ordered_comm_semiring_strict, abs_if (* previously ordered_ring *) | 
| 14270 | 260 | |
| 14738 | 261 | instance ordered_idom \<subseteq> ordered_ring_strict .. | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 262 | |
| 14738 | 263 | axclass ordered_field \<subseteq> field, ordered_idom | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 264 | |
| 15923 | 265 | lemmas linorder_neqE_ordered_idom = | 
| 266 | linorder_neqE[where 'a = "?'b::ordered_idom"] | |
| 267 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 268 | lemma eq_add_iff1: | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 269 | "(a*e + c = b*e + d) = ((a-b)*e + c = (d::'a::ring))" | 
| 14738 | 270 | apply (simp add: diff_minus left_distrib) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 271 | apply (simp add: diff_minus left_distrib add_ac) | 
| 14738 | 272 | apply (simp add: compare_rls minus_mult_left [symmetric]) | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 273 | done | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 274 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 275 | lemma eq_add_iff2: | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 276 | "(a*e + c = b*e + d) = (c = (b-a)*e + (d::'a::ring))" | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 277 | apply (simp add: diff_minus left_distrib add_ac) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 278 | apply (simp add: compare_rls minus_mult_left [symmetric]) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 279 | done | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 280 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 281 | lemma less_add_iff1: | 
| 14738 | 282 | "(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: 
14270diff
changeset | 283 | apply (simp add: diff_minus left_distrib add_ac) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 284 | apply (simp add: compare_rls minus_mult_left [symmetric]) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 285 | done | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 286 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 287 | lemma less_add_iff2: | 
| 14738 | 288 | "(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: 
14270diff
changeset | 289 | apply (simp add: diff_minus left_distrib add_ac) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 290 | apply (simp add: compare_rls minus_mult_left [symmetric]) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 291 | done | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 292 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 293 | lemma le_add_iff1: | 
| 14738 | 294 | "(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: 
14270diff
changeset | 295 | apply (simp add: diff_minus left_distrib add_ac) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 296 | apply (simp add: compare_rls minus_mult_left [symmetric]) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 297 | done | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 298 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 299 | lemma le_add_iff2: | 
| 14738 | 300 | "(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: 
14270diff
changeset | 301 | apply (simp add: diff_minus left_distrib add_ac) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 302 | apply (simp add: compare_rls minus_mult_left [symmetric]) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 303 | done | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14270diff
changeset | 304 | |
| 14270 | 305 | subsection {* Ordering Rules for Multiplication *}
 | 
| 306 | ||
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 307 | lemma mult_left_le_imp_le: | 
| 14738 | 308 | "[|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: 
14341diff
changeset | 309 | 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: 
14341diff
changeset | 310 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 311 | lemma mult_right_le_imp_le: | 
| 14738 | 312 | "[|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: 
14341diff
changeset | 313 | 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: 
14341diff
changeset | 314 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 315 | lemma mult_left_less_imp_less: | 
| 14738 | 316 | "[|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: 
14341diff
changeset | 317 | by (force simp add: mult_left_mono linorder_not_le [symmetric]) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 318 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 319 | lemma mult_right_less_imp_less: | 
| 14738 | 320 | "[|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: 
14341diff
changeset | 321 | by (force simp add: mult_right_mono linorder_not_le [symmetric]) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 322 | |
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 323 | lemma mult_strict_left_mono_neg: | 
| 14738 | 324 | "[|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 | 325 | 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 | 326 | 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 | 327 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 328 | |
| 14738 | 329 | lemma mult_left_mono_neg: | 
| 330 | "[|b \<le> a; c \<le> 0|] ==> c * a \<le> c * (b::'a::pordered_ring)" | |
| 331 | apply (drule mult_left_mono [of _ _ "-c"]) | |
| 332 | apply (simp_all add: minus_mult_left [symmetric]) | |
| 333 | done | |
| 334 | ||
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 335 | lemma mult_strict_right_mono_neg: | 
| 14738 | 336 | "[|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 | 337 | 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 | 338 | 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 | 339 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 340 | |
| 14738 | 341 | lemma mult_right_mono_neg: | 
| 342 | "[|b \<le> a; c \<le> 0|] ==> a * c \<le> (b::'a::pordered_ring) * c" | |
| 343 | apply (drule mult_right_mono [of _ _ "-c"]) | |
| 344 | apply (simp) | |
| 345 | apply (simp_all add: minus_mult_right [symmetric]) | |
| 346 | done | |
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 347 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 348 | subsection{* Products of Signs *}
 | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 349 | |
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 350 | lemma mult_pos_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 | 351 | 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 | 352 | |
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 353 | lemma mult_nonneg_nonneg: "[| (0::'a::pordered_cancel_semiring) \<le> a; 0 \<le> b |] ==> 0 \<le> a*b" | 
| 14738 | 354 | by (drule mult_left_mono [of 0 b], auto) | 
| 355 | ||
| 356 | 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 | 357 | 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 | 358 | |
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 359 | lemma mult_nonneg_nonpos: "[| (0::'a::pordered_cancel_semiring) \<le> a; b \<le> 0 |] ==> a*b \<le> 0" | 
| 14738 | 360 | by (drule mult_left_mono [of b 0], auto) | 
| 361 | ||
| 362 | lemma mult_pos_neg2: "[| (0::'a::ordered_semiring_strict) < a; b < 0 |] ==> b*a < 0" | |
| 363 | by (drule mult_strict_right_mono[of b 0], auto) | |
| 364 | ||
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 365 | lemma mult_nonneg_nonpos2: "[| (0::'a::pordered_cancel_semiring) \<le> a; b \<le> 0 |] ==> b*a \<le> 0" | 
| 14738 | 366 | by (drule mult_right_mono[of b 0], auto) | 
| 367 | ||
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 368 | lemma mult_neg_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 | 369 | 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 | 370 | |
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 371 | lemma mult_nonpos_nonpos: "[| a \<le> (0::'a::pordered_ring); b \<le> 0 |] ==> 0 \<le> a*b" | 
| 14738 | 372 | by (drule mult_right_mono_neg[of a 0 b ], auto) | 
| 373 | ||
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14334diff
changeset | 374 | lemma zero_less_mult_pos: | 
| 14738 | 375 | "[| 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 | 376 | apply (case_tac "b\<le>0") | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 377 | 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 | 378 | 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 | 379 | 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 | 380 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 381 | |
| 14738 | 382 | lemma zero_less_mult_pos2: | 
| 383 | "[| 0 < b*a; 0 < a|] ==> 0 < (b::'a::ordered_semiring_strict)" | |
| 384 | apply (case_tac "b\<le>0") | |
| 385 | apply (auto simp add: order_le_less linorder_not_less) | |
| 386 | apply (drule_tac mult_pos_neg2 [of a b]) | |
| 387 | apply (auto dest: order_less_not_sym) | |
| 388 | done | |
| 389 | ||
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 390 | lemma zero_less_mult_iff: | 
| 14738 | 391 | "((0::'a::ordered_ring_strict) < a*b) = (0 < a & 0 < b | a < 0 & b < 0)" | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 392 | apply (auto simp add: order_le_less linorder_not_less mult_pos_pos | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 393 | mult_neg_neg) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 394 | apply (blast dest: zero_less_mult_pos) | 
| 14738 | 395 | 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 | 396 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 397 | |
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14334diff
changeset | 398 | 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: 
14272diff
changeset | 399 |       assumption of an ordering.  See @{text field_mult_eq_0_iff} below.*}
 | 
| 14738 | 400 | 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 | 401 | apply (case_tac "a < 0") | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 402 | 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 | 403 | 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 | 404 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 405 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 406 | lemma zero_le_mult_iff: | 
| 14738 | 407 | "((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 | 408 | 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 | 409 | zero_less_mult_iff) | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 410 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 411 | lemma mult_less_0_iff: | 
| 14738 | 412 | "(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 | 413 | 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 | 414 | 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 | 415 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 416 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 417 | lemma mult_le_0_iff: | 
| 14738 | 418 | "(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 | 419 | 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 | 420 | 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 | 421 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 422 | |
| 14738 | 423 | lemma split_mult_pos_le: "(0 \<le> a & 0 \<le> b) | (a \<le> 0 & b \<le> 0) \<Longrightarrow> 0 \<le> a * (b::_::pordered_ring)" | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 424 | by (auto simp add: mult_nonneg_nonneg mult_nonpos_nonpos) | 
| 14738 | 425 | |
| 426 | 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)" | |
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 427 | by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2) | 
| 14738 | 428 | |
| 429 | 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 | 430 | 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 | 431 | |
| 14738 | 432 | text{*Proving axiom @{text zero_less_one} makes all @{text ordered_semidom}
 | 
| 433 |       theorems available to members of @{term ordered_idom} *}
 | |
| 434 | ||
| 435 | instance ordered_idom \<subseteq> ordered_semidom | |
| 14421 
ee97b6463cb4
new Ring_and_Field hierarchy, eliminating redundant axioms
 paulson parents: 
14398diff
changeset | 436 | proof | 
| 
ee97b6463cb4
new Ring_and_Field hierarchy, eliminating redundant axioms
 paulson parents: 
14398diff
changeset | 437 | 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: 
14421diff
changeset | 438 | thus "(0::'a) < 1" by (simp add: order_le_less) | 
| 14421 
ee97b6463cb4
new Ring_and_Field hierarchy, eliminating redundant axioms
 paulson parents: 
14398diff
changeset | 439 | qed | 
| 
ee97b6463cb4
new Ring_and_Field hierarchy, eliminating redundant axioms
 paulson parents: 
14398diff
changeset | 440 | |
| 20609 | 441 | instance ordered_ring_strict \<subseteq> no_zero_divisors | 
| 14738 | 442 | by (intro_classes, simp) | 
| 443 | ||
| 444 | instance ordered_idom \<subseteq> idom .. | |
| 445 | ||
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 446 | text{*All three types of comparision involving 0 and 1 are covered.*}
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 447 | |
| 17085 | 448 | lemmas one_neq_zero = zero_neq_one [THEN not_sym] | 
| 449 | declare one_neq_zero [simp] | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 450 | |
| 14738 | 451 | 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: 
14267diff
changeset | 452 | 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: 
14267diff
changeset | 453 | |
| 14738 | 454 | lemma not_one_le_zero [simp]: "~ (1::'a::ordered_semidom) \<le> 0" | 
| 455 | by (simp add: linorder_not_le) | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 456 | |
| 14738 | 457 | lemma not_one_less_zero [simp]: "~ (1::'a::ordered_semidom) < 0" | 
| 458 | by (simp add: linorder_not_less) | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 459 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 460 | subsection{*More Monotonicity*}
 | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 461 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 462 | text{*Strict monotonicity in both arguments*}
 | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 463 | lemma mult_strict_mono: | 
| 14738 | 464 | "[|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: 
14267diff
changeset | 465 | apply (case_tac "c=0") | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 466 | apply (simp add: mult_pos_pos) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 467 | 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: 
14267diff
changeset | 468 | apply (force simp add: order_le_less) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 469 | apply (erule mult_strict_left_mono, assumption) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 470 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 471 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 472 | text{*This weaker variant has more natural premises*}
 | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 473 | lemma mult_strict_mono': | 
| 14738 | 474 | "[| 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: 
14267diff
changeset | 475 | apply (rule mult_strict_mono) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 476 | apply (blast intro: order_le_less_trans)+ | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 477 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 478 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 479 | lemma mult_mono: | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 480 | "[|a \<le> b; c \<le> d; 0 \<le> b; 0 \<le> c|] | 
| 14738 | 481 | ==> a * c \<le> b * (d::'a::pordered_semiring)" | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 482 | apply (erule mult_right_mono [THEN order_trans], assumption) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 483 | apply (erule mult_left_mono, assumption) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 484 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 485 | |
| 14738 | 486 | 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: 
14377diff
changeset | 487 | 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: 
14421diff
changeset | 488 | apply (simp add: order_less_trans [OF zero_less_one]) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 489 | done | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 490 | |
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 491 | lemma mult_less_le_imp_less: "(a::'a::ordered_semiring_strict) < b ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 492 | c <= d ==> 0 <= a ==> 0 < c ==> a * c < b * d" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 493 | apply (subgoal_tac "a * c < b * c") | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 494 | apply (erule order_less_le_trans) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 495 | apply (erule mult_left_mono) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 496 | apply simp | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 497 | apply (erule mult_strict_right_mono) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 498 | apply assumption | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 499 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 500 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 501 | lemma mult_le_less_imp_less: "(a::'a::ordered_semiring_strict) <= b ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 502 | c < d ==> 0 < a ==> 0 <= c ==> a * c < b * d" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 503 | apply (subgoal_tac "a * c <= b * c") | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 504 | apply (erule order_le_less_trans) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 505 | apply (erule mult_strict_left_mono) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 506 | apply simp | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 507 | apply (erule mult_right_mono) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 508 | apply simp | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 509 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 510 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 511 | subsection{*Cancellation Laws for Relationships With a Common Factor*}
 | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 512 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 513 | 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: 
14267diff
changeset | 514 |    also with the relations @{text "\<le>"} and equality.*}
 | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 515 | |
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 516 | text{*These ``disjunction'' versions produce two cases when the comparison is
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 517 | an assumption, but effectively four when the comparison is a goal.*} | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 518 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 519 | lemma mult_less_cancel_right_disj: | 
| 14738 | 520 | "(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: 
14267diff
changeset | 521 | apply (case_tac "c = 0") | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 522 | 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: 
14267diff
changeset | 523 | mult_strict_right_mono_neg) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 524 | apply (auto simp add: linorder_not_less | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 525 | linorder_not_le [symmetric, of "a*c"] | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 526 | linorder_not_le [symmetric, of a]) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 527 | apply (erule_tac [!] notE) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 528 | 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: 
14267diff
changeset | 529 | mult_right_mono_neg) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 530 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 531 | |
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 532 | lemma mult_less_cancel_left_disj: | 
| 14738 | 533 | "(c*a < c*b) = ((0 < c & a < b) | (c < 0 & b < (a::'a::ordered_ring_strict)))" | 
| 534 | apply (case_tac "c = 0") | |
| 535 | apply (auto simp add: linorder_neq_iff mult_strict_left_mono | |
| 536 | mult_strict_left_mono_neg) | |
| 537 | apply (auto simp add: linorder_not_less | |
| 538 | linorder_not_le [symmetric, of "c*a"] | |
| 539 | linorder_not_le [symmetric, of a]) | |
| 540 | apply (erule_tac [!] notE) | |
| 541 | apply (auto simp add: order_less_imp_le mult_left_mono | |
| 542 | mult_left_mono_neg) | |
| 543 | done | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 544 | |
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 545 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 546 | text{*The ``conjunction of implication'' lemmas produce two cases when the
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 547 | comparison is a goal, but give four when the comparison is an assumption.*} | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 548 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 549 | lemma mult_less_cancel_right: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 550 | fixes c :: "'a :: ordered_ring_strict" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 551 | shows "(a*c < b*c) = ((0 \<le> c --> a < b) & (c \<le> 0 --> b < a))" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 552 | by (insert mult_less_cancel_right_disj [of a c b], auto) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 553 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 554 | lemma mult_less_cancel_left: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 555 | fixes c :: "'a :: ordered_ring_strict" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 556 | shows "(c*a < c*b) = ((0 \<le> c --> a < b) & (c \<le> 0 --> b < a))" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 557 | by (insert mult_less_cancel_left_disj [of c a b], auto) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 558 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 559 | lemma mult_le_cancel_right: | 
| 14738 | 560 | "(a*c \<le> b*c) = ((0<c --> a\<le>b) & (c<0 --> b \<le> (a::'a::ordered_ring_strict)))" | 
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 561 | by (simp add: linorder_not_less [symmetric] mult_less_cancel_right_disj) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 562 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 563 | lemma mult_le_cancel_left: | 
| 14738 | 564 | "(c*a \<le> c*b) = ((0<c --> a\<le>b) & (c<0 --> b \<le> (a::'a::ordered_ring_strict)))" | 
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 565 | by (simp add: linorder_not_less [symmetric] mult_less_cancel_left_disj) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 566 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 567 | lemma mult_less_imp_less_left: | 
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14334diff
changeset | 568 | assumes less: "c*a < c*b" and nonneg: "0 \<le> c" | 
| 14738 | 569 | shows "a < (b::'a::ordered_semiring_strict)" | 
| 14377 | 570 | proof (rule ccontr) | 
| 571 | assume "~ a < b" | |
| 572 | hence "b \<le> a" by (simp add: linorder_not_less) | |
| 573 | hence "c*b \<le> c*a" by (rule mult_left_mono) | |
| 574 | with this and less show False | |
| 575 | by (simp add: linorder_not_less [symmetric]) | |
| 576 | qed | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 577 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 578 | lemma mult_less_imp_less_right: | 
| 14738 | 579 | assumes less: "a*c < b*c" and nonneg: "0 <= c" | 
| 580 | shows "a < (b::'a::ordered_semiring_strict)" | |
| 581 | proof (rule ccontr) | |
| 582 | assume "~ a < b" | |
| 583 | hence "b \<le> a" by (simp add: linorder_not_less) | |
| 584 | hence "b*c \<le> a*c" by (rule mult_right_mono) | |
| 585 | with this and less show False | |
| 586 | by (simp add: linorder_not_less [symmetric]) | |
| 587 | qed | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 588 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 589 | text{*Cancellation of equalities with a common factor*}
 | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 590 | lemma mult_cancel_right [simp]: | 
| 14738 | 591 | "(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: 
14267diff
changeset | 592 | apply (cut_tac linorder_less_linear [of 0 c]) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 593 | 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: 
14267diff
changeset | 594 | simp add: linorder_neq_iff) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 595 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 596 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 597 | 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: 
14267diff
changeset | 598 | below that work for fields without an ordering.*} | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 599 | lemma mult_cancel_left [simp]: | 
| 14738 | 600 | "(c*a = c*b) = (c = (0::'a::ordered_ring_strict) | a=b)" | 
| 601 | apply (cut_tac linorder_less_linear [of 0 c]) | |
| 602 | apply (force dest: mult_strict_left_mono_neg mult_strict_left_mono | |
| 603 | simp add: linorder_neq_iff) | |
| 604 | done | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 605 | |
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 606 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 607 | subsubsection{*Special Cancellation Simprules for Multiplication*}
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 608 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 609 | text{*These also produce two cases when the comparison is a goal.*}
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 610 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 611 | lemma mult_le_cancel_right1: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 612 | fixes c :: "'a :: ordered_idom" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 613 | shows "(c \<le> b*c) = ((0<c --> 1\<le>b) & (c<0 --> b \<le> 1))" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 614 | by (insert mult_le_cancel_right [of 1 c b], simp) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 615 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 616 | lemma mult_le_cancel_right2: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 617 | fixes c :: "'a :: ordered_idom" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 618 | shows "(a*c \<le> c) = ((0<c --> a\<le>1) & (c<0 --> 1 \<le> a))" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 619 | by (insert mult_le_cancel_right [of a c 1], simp) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 620 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 621 | lemma mult_le_cancel_left1: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 622 | fixes c :: "'a :: ordered_idom" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 623 | shows "(c \<le> c*b) = ((0<c --> 1\<le>b) & (c<0 --> b \<le> 1))" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 624 | by (insert mult_le_cancel_left [of c 1 b], simp) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 625 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 626 | lemma mult_le_cancel_left2: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 627 | fixes c :: "'a :: ordered_idom" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 628 | shows "(c*a \<le> c) = ((0<c --> a\<le>1) & (c<0 --> 1 \<le> a))" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 629 | by (insert mult_le_cancel_left [of c a 1], simp) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 630 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 631 | lemma mult_less_cancel_right1: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 632 | fixes c :: "'a :: ordered_idom" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 633 | shows "(c < b*c) = ((0 \<le> c --> 1<b) & (c \<le> 0 --> b < 1))" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 634 | by (insert mult_less_cancel_right [of 1 c b], simp) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 635 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 636 | lemma mult_less_cancel_right2: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 637 | fixes c :: "'a :: ordered_idom" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 638 | shows "(a*c < c) = ((0 \<le> c --> a<1) & (c \<le> 0 --> 1 < a))" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 639 | by (insert mult_less_cancel_right [of a c 1], simp) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 640 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 641 | lemma mult_less_cancel_left1: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 642 | fixes c :: "'a :: ordered_idom" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 643 | shows "(c < c*b) = ((0 \<le> c --> 1<b) & (c \<le> 0 --> b < 1))" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 644 | by (insert mult_less_cancel_left [of c 1 b], simp) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 645 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 646 | lemma mult_less_cancel_left2: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 647 | fixes c :: "'a :: ordered_idom" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 648 | shows "(c*a < c) = ((0 \<le> c --> a<1) & (c \<le> 0 --> 1 < a))" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 649 | by (insert mult_less_cancel_left [of c a 1], simp) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 650 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 651 | lemma mult_cancel_right1 [simp]: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 652 | fixes c :: "'a :: ordered_idom" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 653 | shows "(c = b*c) = (c = 0 | b=1)" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 654 | by (insert mult_cancel_right [of 1 c b], force) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 655 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 656 | lemma mult_cancel_right2 [simp]: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 657 | fixes c :: "'a :: ordered_idom" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 658 | shows "(a*c = c) = (c = 0 | a=1)" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 659 | by (insert mult_cancel_right [of a c 1], simp) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 660 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 661 | lemma mult_cancel_left1 [simp]: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 662 | fixes c :: "'a :: ordered_idom" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 663 | shows "(c = c*b) = (c = 0 | b=1)" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 664 | by (insert mult_cancel_left [of c 1 b], force) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 665 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 666 | lemma mult_cancel_left2 [simp]: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 667 | fixes c :: "'a :: ordered_idom" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 668 | shows "(c*a = c) = (c = 0 | a=1)" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 669 | by (insert mult_cancel_left [of c a 1], simp) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 670 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 671 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 672 | text{*Simprules for comparisons where common factors can be cancelled.*}
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 673 | lemmas mult_compare_simps = | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 674 | mult_le_cancel_right mult_le_cancel_left | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 675 | mult_le_cancel_right1 mult_le_cancel_right2 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 676 | mult_le_cancel_left1 mult_le_cancel_left2 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 677 | mult_less_cancel_right mult_less_cancel_left | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 678 | mult_less_cancel_right1 mult_less_cancel_right2 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 679 | mult_less_cancel_left1 mult_less_cancel_left2 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 680 | mult_cancel_right mult_cancel_left | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 681 | mult_cancel_right1 mult_cancel_right2 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 682 | mult_cancel_left1 mult_cancel_left2 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 683 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 684 | |
| 14738 | 685 | text{*This list of rewrites decides ring equalities by ordered rewriting.*}
 | 
| 15178 | 686 | lemmas ring_eq_simps = | 
| 687 | (* mult_ac*) | |
| 14738 | 688 | left_distrib right_distrib left_diff_distrib right_diff_distrib | 
| 15178 | 689 | group_eq_simps | 
| 690 | (* add_ac | |
| 14738 | 691 | add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2 | 
| 15178 | 692 | diff_eq_eq eq_diff_eq *) | 
| 14738 | 693 | |
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 694 | subsection {* Fields *}
 | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 695 | |
| 14288 | 696 | lemma right_inverse_eq: "b \<noteq> 0 ==> (a / b = 1) = (a = (b::'a::field))" | 
| 697 | proof | |
| 698 | assume neq: "b \<noteq> 0" | |
| 699 |   {
 | |
| 700 | hence "a = (a / b) * b" by (simp add: divide_inverse mult_ac) | |
| 701 | also assume "a / b = 1" | |
| 702 | finally show "a = b" by simp | |
| 703 | next | |
| 704 | assume "a = b" | |
| 705 | with neq show "a / b = 1" by (simp add: divide_inverse) | |
| 706 | } | |
| 707 | qed | |
| 708 | ||
| 709 | lemma nonzero_inverse_eq_divide: "a \<noteq> 0 ==> inverse (a::'a::field) = 1/a" | |
| 710 | by (simp add: divide_inverse) | |
| 711 | ||
| 15228 | 712 | lemma divide_self: "a \<noteq> 0 ==> a / (a::'a::field) = 1" | 
| 14288 | 713 | by (simp add: divide_inverse) | 
| 714 | ||
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 715 | 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: 
14421diff
changeset | 716 | by (simp add: divide_inverse) | 
| 14277 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 717 | |
| 15228 | 718 | lemma divide_self_if [simp]: | 
| 719 |      "a / (a::'a::{field,division_by_zero}) = (if a=0 then 0 else 1)"
 | |
| 720 | by (simp add: divide_self) | |
| 721 | ||
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 722 | lemma divide_zero_left [simp]: "0/a = (0::'a::field)" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 723 | by (simp add: divide_inverse) | 
| 14277 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 724 | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 725 | lemma inverse_eq_divide: "inverse (a::'a::field) = 1/a" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 726 | by (simp add: divide_inverse) | 
| 14277 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 727 | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 728 | lemma add_divide_distrib: "(a+b)/(c::'a::field) = a/c + b/c" | 
| 14293 | 729 | by (simp add: divide_inverse left_distrib) | 
| 730 | ||
| 731 | ||
| 14270 | 732 | text{*Compared with @{text mult_eq_0_iff}, this version removes the requirement
 | 
| 733 | of an ordering.*} | |
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 734 | lemma field_mult_eq_0_iff [simp]: | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 735 | "(a*b = (0::'a::division_ring)) = (a = 0 | b = 0)" | 
| 14377 | 736 | proof cases | 
| 737 | assume "a=0" thus ?thesis by simp | |
| 738 | next | |
| 739 | assume anz [simp]: "a\<noteq>0" | |
| 740 |   { assume "a * b = 0"
 | |
| 741 | hence "inverse a * (a * b) = 0" by simp | |
| 742 | hence "b = 0" by (simp (no_asm_use) add: mult_assoc [symmetric])} | |
| 743 | thus ?thesis by force | |
| 744 | qed | |
| 14270 | 745 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 746 | text{*Cancellation of equalities with a common factor*}
 | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 747 | lemma field_mult_cancel_right_lemma: | 
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 748 | assumes cnz: "c \<noteq> (0::'a::division_ring)" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 749 | and eq: "a*c = b*c" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 750 | shows "a=b" | 
| 14377 | 751 | proof - | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 752 | have "(a * c) * inverse c = (b * c) * inverse c" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 753 | by (simp add: eq) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 754 | thus "a=b" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 755 | by (simp add: mult_assoc cnz) | 
| 14377 | 756 | qed | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 757 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 758 | lemma field_mult_cancel_right [simp]: | 
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 759 | "(a*c = b*c) = (c = (0::'a::division_ring) | a=b)" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 760 | proof - | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 761 | have "(a*c = b*c) = (a*c - b*c = 0)" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 762 | by simp | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 763 | also have "\<dots> = ((a - b)*c = 0)" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 764 | by (simp only: left_diff_distrib) | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 765 | also have "\<dots> = (c = 0 \<or> a = b)" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 766 | by (simp add: disj_commute) | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 767 | finally show ?thesis . | 
| 14377 | 768 | qed | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 769 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 770 | lemma field_mult_cancel_left [simp]: | 
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 771 | "(c*a = c*b) = (c = (0::'a::division_ring) | a=b)" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 772 | proof - | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 773 | have "(c*a = c*b) = (c*a - c*b = 0)" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 774 | by simp | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 775 | also have "\<dots> = (c*(a - b) = 0)" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 776 | by (simp only: right_diff_distrib) | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 777 | also have "\<dots> = (c = 0 \<or> a = b)" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 778 | by simp | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 779 | finally show ?thesis . | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 780 | qed | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 781 | |
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 782 | lemma nonzero_imp_inverse_nonzero: | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 783 | "a \<noteq> 0 ==> inverse a \<noteq> (0::'a::division_ring)" | 
| 14377 | 784 | proof | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 785 | assume ianz: "inverse a = 0" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 786 | assume "a \<noteq> 0" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 787 | hence "1 = a * inverse a" by simp | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 788 | also have "... = 0" by (simp add: ianz) | 
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 789 | finally have "1 = (0::'a::division_ring)" . | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 790 | thus False by (simp add: eq_commute) | 
| 14377 | 791 | qed | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 792 | |
| 14277 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 793 | |
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 794 | subsection{*Basic Properties of @{term inverse}*}
 | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 795 | |
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 796 | lemma inverse_zero_imp_zero: "inverse a = 0 ==> a = (0::'a::division_ring)" | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 797 | apply (rule ccontr) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 798 | apply (blast dest: nonzero_imp_inverse_nonzero) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 799 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 800 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 801 | lemma inverse_nonzero_imp_nonzero: | 
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 802 | "inverse a = 0 ==> a = (0::'a::division_ring)" | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 803 | apply (rule ccontr) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 804 | apply (blast dest: nonzero_imp_inverse_nonzero) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 805 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 806 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 807 | lemma inverse_nonzero_iff_nonzero [simp]: | 
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 808 |    "(inverse a = 0) = (a = (0::'a::{division_ring,division_by_zero}))"
 | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 809 | by (force dest: inverse_nonzero_imp_nonzero) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 810 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 811 | lemma nonzero_inverse_minus_eq: | 
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 812 | assumes [simp]: "a\<noteq>0" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 813 | shows "inverse(-a) = -inverse(a::'a::division_ring)" | 
| 14377 | 814 | proof - | 
| 815 | have "-a * inverse (- a) = -a * - inverse a" | |
| 816 | by simp | |
| 817 | thus ?thesis | |
| 818 | by (simp only: field_mult_cancel_left, simp) | |
| 819 | qed | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 820 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 821 | lemma inverse_minus_eq [simp]: | 
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 822 |    "inverse(-a) = -inverse(a::'a::{division_ring,division_by_zero})"
 | 
| 14377 | 823 | proof cases | 
| 824 | assume "a=0" thus ?thesis by (simp add: inverse_zero) | |
| 825 | next | |
| 826 | assume "a\<noteq>0" | |
| 827 | thus ?thesis by (simp add: nonzero_inverse_minus_eq) | |
| 828 | qed | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 829 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 830 | lemma nonzero_inverse_eq_imp_eq: | 
| 14269 | 831 | assumes inveq: "inverse a = inverse b" | 
| 832 | and anz: "a \<noteq> 0" | |
| 833 | and bnz: "b \<noteq> 0" | |
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 834 | shows "a = (b::'a::division_ring)" | 
| 14377 | 835 | proof - | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 836 | have "a * inverse b = a * inverse a" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 837 | by (simp add: inveq) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 838 | hence "(a * inverse b) * b = (a * inverse a) * b" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 839 | by simp | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 840 | thus "a = b" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 841 | by (simp add: mult_assoc anz bnz) | 
| 14377 | 842 | qed | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 843 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 844 | lemma inverse_eq_imp_eq: | 
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 845 |   "inverse a = inverse b ==> a = (b::'a::{division_ring,division_by_zero})"
 | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 846 | apply (case_tac "a=0 | b=0") | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 847 | apply (force dest!: inverse_zero_imp_zero | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 848 | simp add: eq_commute [of "0::'a"]) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 849 | apply (force dest!: nonzero_inverse_eq_imp_eq) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 850 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 851 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 852 | lemma inverse_eq_iff_eq [simp]: | 
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 853 |   "(inverse a = inverse b) = (a = (b::'a::{division_ring,division_by_zero}))"
 | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 854 | by (force dest!: inverse_eq_imp_eq) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 855 | |
| 14270 | 856 | lemma nonzero_inverse_inverse_eq: | 
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 857 | assumes [simp]: "a \<noteq> 0" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 858 | shows "inverse(inverse (a::'a::division_ring)) = a" | 
| 14270 | 859 | proof - | 
| 860 | have "(inverse (inverse a) * inverse a) * a = a" | |
| 861 | by (simp add: nonzero_imp_inverse_nonzero) | |
| 862 | thus ?thesis | |
| 863 | by (simp add: mult_assoc) | |
| 864 | qed | |
| 865 | ||
| 866 | lemma inverse_inverse_eq [simp]: | |
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 867 |      "inverse(inverse (a::'a::{division_ring,division_by_zero})) = a"
 | 
| 14270 | 868 | proof cases | 
| 869 | assume "a=0" thus ?thesis by simp | |
| 870 | next | |
| 871 | assume "a\<noteq>0" | |
| 872 | thus ?thesis by (simp add: nonzero_inverse_inverse_eq) | |
| 873 | qed | |
| 874 | ||
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 875 | lemma inverse_1 [simp]: "inverse 1 = (1::'a::division_ring)" | 
| 14270 | 876 | proof - | 
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 877 | have "inverse 1 * 1 = (1::'a::division_ring)" | 
| 14270 | 878 | by (rule left_inverse [OF zero_neq_one [symmetric]]) | 
| 879 | thus ?thesis by simp | |
| 880 | qed | |
| 881 | ||
| 15077 
89840837108e
converting Hyperreal/Transcendental to Isar script
 paulson parents: 
15010diff
changeset | 882 | lemma inverse_unique: | 
| 
89840837108e
converting Hyperreal/Transcendental to Isar script
 paulson parents: 
15010diff
changeset | 883 | assumes ab: "a*b = 1" | 
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 884 | shows "inverse a = (b::'a::division_ring)" | 
| 15077 
89840837108e
converting Hyperreal/Transcendental to Isar script
 paulson parents: 
15010diff
changeset | 885 | proof - | 
| 
89840837108e
converting Hyperreal/Transcendental to Isar script
 paulson parents: 
15010diff
changeset | 886 | have "a \<noteq> 0" using ab by auto | 
| 
89840837108e
converting Hyperreal/Transcendental to Isar script
 paulson parents: 
15010diff
changeset | 887 | moreover have "inverse a * (a * b) = inverse a" by (simp add: ab) | 
| 
89840837108e
converting Hyperreal/Transcendental to Isar script
 paulson parents: 
15010diff
changeset | 888 | ultimately show ?thesis by (simp add: mult_assoc [symmetric]) | 
| 
89840837108e
converting Hyperreal/Transcendental to Isar script
 paulson parents: 
15010diff
changeset | 889 | qed | 
| 
89840837108e
converting Hyperreal/Transcendental to Isar script
 paulson parents: 
15010diff
changeset | 890 | |
| 14270 | 891 | lemma nonzero_inverse_mult_distrib: | 
| 892 | assumes anz: "a \<noteq> 0" | |
| 893 | and bnz: "b \<noteq> 0" | |
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 894 | shows "inverse(a*b) = inverse(b) * inverse(a::'a::division_ring)" | 
| 14270 | 895 | proof - | 
| 896 | have "inverse(a*b) * (a * b) * inverse(b) = inverse(b)" | |
| 897 | by (simp add: field_mult_eq_0_iff anz bnz) | |
| 898 | hence "inverse(a*b) * a = inverse(b)" | |
| 899 | by (simp add: mult_assoc bnz) | |
| 900 | hence "inverse(a*b) * a * inverse(a) = inverse(b) * inverse(a)" | |
| 901 | by simp | |
| 902 | thus ?thesis | |
| 903 | by (simp add: mult_assoc anz) | |
| 904 | qed | |
| 905 | ||
| 906 | text{*This version builds in division by zero while also re-orienting
 | |
| 907 | the right-hand side.*} | |
| 908 | lemma inverse_mult_distrib [simp]: | |
| 909 |      "inverse(a*b) = inverse(a) * inverse(b::'a::{field,division_by_zero})"
 | |
| 910 | proof cases | |
| 911 | assume "a \<noteq> 0 & b \<noteq> 0" | |
| 912 | thus ?thesis by (simp add: nonzero_inverse_mult_distrib mult_commute) | |
| 913 | next | |
| 914 | assume "~ (a \<noteq> 0 & b \<noteq> 0)" | |
| 915 | thus ?thesis by force | |
| 916 | qed | |
| 917 | ||
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 918 | lemma division_ring_inverse_add: | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 919 | "[|(a::'a::division_ring) \<noteq> 0; b \<noteq> 0|] | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 920 | ==> inverse a + inverse b = inverse a * (a+b) * inverse b" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 921 | by (simp add: right_distrib left_distrib mult_assoc) | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 922 | |
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 923 | lemma division_ring_inverse_diff: | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 924 | "[|(a::'a::division_ring) \<noteq> 0; b \<noteq> 0|] | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 925 | ==> inverse a - inverse b = inverse a * (b-a) * inverse b" | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 926 | by (simp add: right_diff_distrib left_diff_distrib mult_assoc) | 
| 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 927 | |
| 14270 | 928 | text{*There is no slick version using division by zero.*}
 | 
| 929 | lemma inverse_add: | |
| 930 | "[|a \<noteq> 0; b \<noteq> 0|] | |
| 931 | ==> inverse a + inverse b = (a+b) * inverse a * inverse (b::'a::field)" | |
| 20496 
23eb6034c06d
added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
 huffman parents: 
19404diff
changeset | 932 | by (simp add: division_ring_inverse_add mult_ac) | 
| 14270 | 933 | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 934 | lemma inverse_divide [simp]: | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 935 |       "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: 
14421diff
changeset | 936 | by (simp add: divide_inverse mult_commute) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 937 | |
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 938 | subsection {* Calculations with fractions *}
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 939 | |
| 14277 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 940 | lemma nonzero_mult_divide_cancel_left: | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 941 | assumes [simp]: "b\<noteq>0" and [simp]: "c\<noteq>0" | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 942 | shows "(c*a)/(c*b) = a/(b::'a::field)" | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 943 | proof - | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 944 | 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: 
14272diff
changeset | 945 | by (simp add: field_mult_eq_0_iff divide_inverse | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 946 | nonzero_inverse_mult_distrib) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 947 | also have "... = a * inverse b * (inverse c * c)" | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 948 | by (simp only: mult_ac) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 949 | also have "... = a * inverse b" | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 950 | by simp | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 951 | finally show ?thesis | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 952 | by (simp add: divide_inverse) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 953 | qed | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 954 | |
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 955 | lemma mult_divide_cancel_left: | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 956 |      "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: 
14272diff
changeset | 957 | apply (case_tac "b = 0") | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 958 | apply (simp_all add: nonzero_mult_divide_cancel_left) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 959 | done | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 960 | |
| 14321 | 961 | lemma nonzero_mult_divide_cancel_right: | 
| 962 | "[|b\<noteq>0; c\<noteq>0|] ==> (a*c) / (b*c) = a/(b::'a::field)" | |
| 963 | by (simp add: mult_commute [of _ c] nonzero_mult_divide_cancel_left) | |
| 964 | ||
| 965 | lemma mult_divide_cancel_right: | |
| 966 |      "c\<noteq>0 ==> (a*c) / (b*c) = a / (b::'a::{field,division_by_zero})"
 | |
| 967 | apply (case_tac "b = 0") | |
| 968 | apply (simp_all add: nonzero_mult_divide_cancel_right) | |
| 969 | done | |
| 970 | ||
| 14277 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 971 | (*For ExtractCommonTerm*) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 972 | lemma mult_divide_cancel_eq_if: | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 973 | "(c*a) / (c*b) = | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 974 |       (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: 
14272diff
changeset | 975 | by (simp add: mult_divide_cancel_left) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 976 | |
| 14284 
f1abe67c448a
re-organisation of Real/RealArith0.ML; more `Isar scripts
 paulson parents: 
14277diff
changeset | 977 | lemma divide_1 [simp]: "a/1 = (a::'a::field)" | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 978 | by (simp add: divide_inverse) | 
| 14284 
f1abe67c448a
re-organisation of Real/RealArith0.ML; more `Isar scripts
 paulson parents: 
14277diff
changeset | 979 | |
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 980 | lemma times_divide_eq_right: "a * (b/c) = (a*b) / (c::'a::field)" | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 981 | by (simp add: divide_inverse mult_assoc) | 
| 14288 | 982 | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 983 | 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: 
14421diff
changeset | 984 | by (simp add: divide_inverse mult_ac) | 
| 14288 | 985 | |
| 986 | lemma divide_divide_eq_right [simp]: | |
| 987 |      "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: 
14421diff
changeset | 988 | by (simp add: divide_inverse mult_ac) | 
| 14288 | 989 | |
| 990 | lemma divide_divide_eq_left [simp]: | |
| 991 |      "(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: 
14421diff
changeset | 992 | by (simp add: divide_inverse mult_assoc) | 
| 14288 | 993 | |
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 994 | lemma add_frac_eq: "(y::'a::field) ~= 0 ==> z ~= 0 ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 995 | x / y + w / z = (x * z + w * y) / (y * z)" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 996 | apply (subgoal_tac "x / y = (x * z) / (y * z)") | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 997 | apply (erule ssubst) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 998 | apply (subgoal_tac "w / z = (w * y) / (y * z)") | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 999 | apply (erule ssubst) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1000 | apply (rule add_divide_distrib [THEN sym]) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1001 | apply (subst mult_commute) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1002 | apply (erule nonzero_mult_divide_cancel_left [THEN sym]) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1003 | apply assumption | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1004 | apply (erule nonzero_mult_divide_cancel_right [THEN sym]) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1005 | apply assumption | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1006 | done | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1007 | |
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1008 | subsubsection{*Special Cancellation Simprules for Division*}
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1009 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1010 | lemma mult_divide_cancel_left_if [simp]: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1011 |   fixes c :: "'a :: {field,division_by_zero}"
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1012 | shows "(c*a) / (c*b) = (if c=0 then 0 else a/b)" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1013 | by (simp add: mult_divide_cancel_left) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1014 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1015 | lemma mult_divide_cancel_right_if [simp]: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1016 |   fixes c :: "'a :: {field,division_by_zero}"
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1017 | shows "(a*c) / (b*c) = (if c=0 then 0 else a/b)" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1018 | by (simp add: mult_divide_cancel_right) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1019 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1020 | lemma mult_divide_cancel_left_if1 [simp]: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1021 |   fixes c :: "'a :: {field,division_by_zero}"
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1022 | shows "c / (c*b) = (if c=0 then 0 else 1/b)" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1023 | apply (insert mult_divide_cancel_left_if [of c 1 b]) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1024 | apply (simp del: mult_divide_cancel_left_if) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1025 | done | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1026 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1027 | lemma mult_divide_cancel_left_if2 [simp]: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1028 |   fixes c :: "'a :: {field,division_by_zero}"
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1029 | shows "(c*a) / c = (if c=0 then 0 else a)" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1030 | apply (insert mult_divide_cancel_left_if [of c a 1]) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1031 | apply (simp del: mult_divide_cancel_left_if) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1032 | done | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1033 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1034 | lemma mult_divide_cancel_right_if1 [simp]: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1035 |   fixes c :: "'a :: {field,division_by_zero}"
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1036 | shows "c / (b*c) = (if c=0 then 0 else 1/b)" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1037 | apply (insert mult_divide_cancel_right_if [of 1 c b]) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1038 | apply (simp del: mult_divide_cancel_right_if) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1039 | done | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1040 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1041 | lemma mult_divide_cancel_right_if2 [simp]: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1042 |   fixes c :: "'a :: {field,division_by_zero}"
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1043 | shows "(a*c) / c = (if c=0 then 0 else a)" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1044 | apply (insert mult_divide_cancel_right_if [of a c 1]) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1045 | apply (simp del: mult_divide_cancel_right_if) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1046 | done | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1047 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1048 | text{*Two lemmas for cancelling the denominator*}
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1049 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1050 | lemma times_divide_self_right [simp]: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1051 |   fixes a :: "'a :: {field,division_by_zero}"
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1052 | shows "a * (b/a) = (if a=0 then 0 else b)" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1053 | by (simp add: times_divide_eq_right) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1054 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1055 | lemma times_divide_self_left [simp]: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1056 |   fixes a :: "'a :: {field,division_by_zero}"
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1057 | shows "(b/a) * a = (if a=0 then 0 else b)" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1058 | by (simp add: times_divide_eq_left) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1059 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1060 | |
| 14293 | 1061 | subsection {* Division and Unary Minus *}
 | 
| 1062 | ||
| 1063 | lemma nonzero_minus_divide_left: "b \<noteq> 0 ==> - (a/b) = (-a) / (b::'a::field)" | |
| 1064 | by (simp add: divide_inverse minus_mult_left) | |
| 1065 | ||
| 1066 | lemma nonzero_minus_divide_right: "b \<noteq> 0 ==> - (a/b) = a / -(b::'a::field)" | |
| 1067 | by (simp add: divide_inverse nonzero_inverse_minus_eq minus_mult_right) | |
| 1068 | ||
| 1069 | lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a)/(-b) = a / (b::'a::field)" | |
| 1070 | by (simp add: divide_inverse nonzero_inverse_minus_eq) | |
| 1071 | ||
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 1072 | lemma minus_divide_left: "- (a/b) = (-a) / (b::'a::field)" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 1073 | by (simp add: divide_inverse minus_mult_left [symmetric]) | 
| 14293 | 1074 | |
| 1075 | 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: 
14421diff
changeset | 1076 | by (simp add: divide_inverse minus_mult_right [symmetric]) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 1077 | |
| 14293 | 1078 | |
| 1079 | text{*The effect is to extract signs from divisions*}
 | |
| 17085 | 1080 | lemmas divide_minus_left = minus_divide_left [symmetric] | 
| 1081 | lemmas divide_minus_right = minus_divide_right [symmetric] | |
| 1082 | declare divide_minus_left [simp] divide_minus_right [simp] | |
| 14293 | 1083 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 1084 | text{*Also, extract signs from products*}
 | 
| 17085 | 1085 | lemmas mult_minus_left = minus_mult_left [symmetric] | 
| 1086 | lemmas mult_minus_right = minus_mult_right [symmetric] | |
| 1087 | declare mult_minus_left [simp] mult_minus_right [simp] | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 1088 | |
| 14293 | 1089 | lemma minus_divide_divide [simp]: | 
| 1090 |      "(-a)/(-b) = a / (b::'a::{field,division_by_zero})"
 | |
| 1091 | apply (case_tac "b=0", simp) | |
| 1092 | apply (simp add: nonzero_minus_divide_divide) | |
| 1093 | done | |
| 1094 | ||
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 1095 | lemma diff_divide_distrib: "(a-b)/(c::'a::field) = a/c - b/c" | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 1096 | by (simp add: diff_minus add_divide_distrib) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 1097 | |
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1098 | lemma diff_frac_eq: "(y::'a::field) ~= 0 ==> z ~= 0 ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1099 | x / y - w / z = (x * z - w * y) / (y * z)" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1100 | apply (subst diff_def)+ | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1101 | apply (subst minus_divide_left) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1102 | apply (subst add_frac_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1103 | apply simp_all | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1104 | done | 
| 14293 | 1105 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1106 | subsection {* Ordered Fields *}
 | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1107 | |
| 14277 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1108 | lemma positive_imp_inverse_positive: | 
| 14269 | 1109 | assumes a_gt_0: "0 < a" shows "0 < inverse (a::'a::ordered_field)" | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1110 | proof - | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1111 | have "0 < a * inverse a" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1112 | by (simp add: a_gt_0 [THEN order_less_imp_not_eq2] zero_less_one) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1113 | thus "0 < inverse a" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1114 | by (simp add: a_gt_0 [THEN order_less_not_sym] zero_less_mult_iff) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1115 | qed | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1116 | |
| 14277 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1117 | lemma negative_imp_inverse_negative: | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1118 | "a < 0 ==> inverse a < (0::'a::ordered_field)" | 
| 14277 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1119 | by (insert positive_imp_inverse_positive [of "-a"], | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1120 | simp add: nonzero_inverse_minus_eq order_less_imp_not_eq) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1121 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1122 | lemma inverse_le_imp_le: | 
| 14269 | 1123 | assumes invle: "inverse a \<le> inverse b" | 
| 1124 | and apos: "0 < a" | |
| 1125 | shows "b \<le> (a::'a::ordered_field)" | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1126 | proof (rule classical) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1127 | assume "~ b \<le> a" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1128 | hence "a < b" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1129 | by (simp add: linorder_not_le) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1130 | hence bpos: "0 < b" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1131 | by (blast intro: apos order_less_trans) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1132 | hence "a * inverse a \<le> a * inverse b" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1133 | by (simp add: apos invle order_less_imp_le mult_left_mono) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1134 | hence "(a * inverse a) * b \<le> (a * inverse b) * b" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1135 | by (simp add: bpos order_less_imp_le mult_right_mono) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1136 | thus "b \<le> a" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1137 | by (simp add: mult_assoc apos bpos order_less_imp_not_eq2) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1138 | qed | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1139 | |
| 14277 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1140 | lemma inverse_positive_imp_positive: | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1141 | assumes inv_gt_0: "0 < inverse a" | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1142 | and [simp]: "a \<noteq> 0" | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1143 | shows "0 < (a::'a::ordered_field)" | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1144 | proof - | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1145 | have "0 < inverse (inverse a)" | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1146 | by (rule positive_imp_inverse_positive) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1147 | thus "0 < a" | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1148 | by (simp add: nonzero_inverse_inverse_eq) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1149 | qed | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1150 | |
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1151 | lemma inverse_positive_iff_positive [simp]: | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1152 |       "(0 < inverse a) = (0 < (a::'a::{ordered_field,division_by_zero}))"
 | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1153 | apply (case_tac "a = 0", simp) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1154 | apply (blast intro: inverse_positive_imp_positive positive_imp_inverse_positive) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1155 | done | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1156 | |
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1157 | lemma inverse_negative_imp_negative: | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1158 | assumes inv_less_0: "inverse a < 0" | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1159 | and [simp]: "a \<noteq> 0" | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1160 | shows "a < (0::'a::ordered_field)" | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1161 | proof - | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1162 | have "inverse (inverse a) < 0" | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1163 | by (rule negative_imp_inverse_negative) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1164 | thus "a < 0" | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1165 | by (simp add: nonzero_inverse_inverse_eq) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1166 | qed | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1167 | |
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1168 | lemma inverse_negative_iff_negative [simp]: | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1169 |       "(inverse a < 0) = (a < (0::'a::{ordered_field,division_by_zero}))"
 | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1170 | apply (case_tac "a = 0", simp) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1171 | apply (blast intro: inverse_negative_imp_negative negative_imp_inverse_negative) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1172 | done | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1173 | |
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1174 | lemma inverse_nonnegative_iff_nonnegative [simp]: | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1175 |       "(0 \<le> inverse a) = (0 \<le> (a::'a::{ordered_field,division_by_zero}))"
 | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1176 | by (simp add: linorder_not_less [symmetric]) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1177 | |
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1178 | lemma inverse_nonpositive_iff_nonpositive [simp]: | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1179 |       "(inverse a \<le> 0) = (a \<le> (0::'a::{ordered_field,division_by_zero}))"
 | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1180 | by (simp add: linorder_not_less [symmetric]) | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1181 | |
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1182 | |
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1183 | subsection{*Anti-Monotonicity of @{term inverse}*}
 | 
| 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1184 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1185 | lemma less_imp_inverse_less: | 
| 14269 | 1186 | assumes less: "a < b" | 
| 1187 | and apos: "0 < a" | |
| 1188 | shows "inverse b < inverse (a::'a::ordered_field)" | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1189 | proof (rule ccontr) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1190 | assume "~ inverse b < inverse a" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1191 | hence "inverse a \<le> inverse b" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1192 | by (simp add: linorder_not_less) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1193 | hence "~ (a < b)" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1194 | by (simp add: linorder_not_less inverse_le_imp_le [OF _ apos]) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1195 | thus False | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1196 | by (rule notE [OF _ less]) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1197 | qed | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1198 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1199 | lemma inverse_less_imp_less: | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1200 | "[|inverse a < inverse b; 0 < a|] ==> b < (a::'a::ordered_field)" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1201 | apply (simp add: order_less_le [of "inverse a"] order_less_le [of "b"]) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1202 | apply (force dest!: inverse_le_imp_le nonzero_inverse_eq_imp_eq) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1203 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1204 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1205 | text{*Both premises are essential. Consider -1 and 1.*}
 | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1206 | lemma inverse_less_iff_less [simp]: | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1207 | "[|0 < a; 0 < b|] | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1208 | ==> (inverse a < inverse b) = (b < (a::'a::ordered_field))" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1209 | by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1210 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1211 | lemma le_imp_inverse_le: | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1212 | "[|a \<le> b; 0 < a|] ==> inverse b \<le> inverse (a::'a::ordered_field)" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1213 | by (force simp add: order_le_less less_imp_inverse_less) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1214 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1215 | lemma inverse_le_iff_le [simp]: | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1216 | "[|0 < a; 0 < b|] | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1217 | ==> (inverse a \<le> inverse b) = (b \<le> (a::'a::ordered_field))" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1218 | by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1219 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1220 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1221 | text{*These results refer to both operands being negative.  The opposite-sign
 | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1222 | case is trivial, since inverse preserves signs.*} | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1223 | lemma inverse_le_imp_le_neg: | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1224 | "[|inverse a \<le> inverse b; b < 0|] ==> b \<le> (a::'a::ordered_field)" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1225 | apply (rule classical) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1226 | apply (subgoal_tac "a < 0") | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1227 | prefer 2 apply (force simp add: linorder_not_le intro: order_less_trans) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1228 | apply (insert inverse_le_imp_le [of "-b" "-a"]) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1229 | apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1230 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1231 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1232 | lemma less_imp_inverse_less_neg: | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1233 | "[|a < b; b < 0|] ==> inverse b < inverse (a::'a::ordered_field)" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1234 | apply (subgoal_tac "a < 0") | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1235 | prefer 2 apply (blast intro: order_less_trans) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1236 | apply (insert less_imp_inverse_less [of "-b" "-a"]) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1237 | apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1238 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1239 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1240 | lemma inverse_less_imp_less_neg: | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1241 | "[|inverse a < inverse b; b < 0|] ==> b < (a::'a::ordered_field)" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1242 | apply (rule classical) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1243 | apply (subgoal_tac "a < 0") | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1244 | prefer 2 | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1245 | apply (force simp add: linorder_not_less intro: order_le_less_trans) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1246 | apply (insert inverse_less_imp_less [of "-b" "-a"]) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1247 | apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1248 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1249 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1250 | lemma inverse_less_iff_less_neg [simp]: | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1251 | "[|a < 0; b < 0|] | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1252 | ==> (inverse a < inverse b) = (b < (a::'a::ordered_field))" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1253 | apply (insert inverse_less_iff_less [of "-b" "-a"]) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1254 | apply (simp del: inverse_less_iff_less | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1255 | add: order_less_imp_not_eq nonzero_inverse_minus_eq) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1256 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1257 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1258 | lemma le_imp_inverse_le_neg: | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1259 | "[|a \<le> b; b < 0|] ==> inverse b \<le> inverse (a::'a::ordered_field)" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1260 | by (force simp add: order_le_less less_imp_inverse_less_neg) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1261 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1262 | lemma inverse_le_iff_le_neg [simp]: | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1263 | "[|a < 0; b < 0|] | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1264 | ==> (inverse a \<le> inverse b) = (b \<le> (a::'a::ordered_field))" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14267diff
changeset | 1265 | by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 1266 | |
| 14277 
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
 paulson parents: 
14272diff
changeset | 1267 | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1268 | subsection{*Inverses and the Number One*}
 | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1269 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1270 | lemma one_less_inverse_iff: | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1271 |     "(1 < inverse x) = (0 < x & x < (1::'a::{ordered_field,division_by_zero}))"proof cases
 | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1272 | assume "0 < x" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1273 | with inverse_less_iff_less [OF zero_less_one, of x] | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1274 | show ?thesis by simp | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1275 | next | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1276 | assume notless: "~ (0 < x)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1277 | have "~ (1 < inverse x)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1278 | proof | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1279 | assume "1 < inverse x" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1280 | also with notless have "... \<le> 0" by (simp add: linorder_not_less) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1281 | also have "... < 1" by (rule zero_less_one) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1282 | finally show False by auto | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1283 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1284 | with notless show ?thesis by simp | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1285 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1286 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1287 | lemma inverse_eq_1_iff [simp]: | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1288 |     "(inverse x = 1) = (x = (1::'a::{field,division_by_zero}))"
 | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1289 | by (insert inverse_eq_iff_eq [of x 1], simp) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1290 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1291 | lemma one_le_inverse_iff: | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1292 |    "(1 \<le> inverse x) = (0 < x & x \<le> (1::'a::{ordered_field,division_by_zero}))"
 | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1293 | by (force simp add: order_le_less one_less_inverse_iff zero_less_one | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1294 | eq_commute [of 1]) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1295 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1296 | lemma inverse_less_1_iff: | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1297 |    "(inverse x < 1) = (x \<le> 0 | 1 < (x::'a::{ordered_field,division_by_zero}))"
 | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1298 | by (simp add: linorder_not_le [symmetric] one_le_inverse_iff) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1299 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1300 | lemma inverse_le_1_iff: | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1301 |    "(inverse x \<le> 1) = (x \<le> 0 | 1 \<le> (x::'a::{ordered_field,division_by_zero}))"
 | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1302 | by (simp add: linorder_not_less [symmetric] one_less_inverse_iff) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1303 | |
| 14288 | 1304 | subsection{*Simplification of Inequalities Involving Literal Divisors*}
 | 
| 1305 | ||
| 1306 | lemma pos_le_divide_eq: "0 < (c::'a::ordered_field) ==> (a \<le> b/c) = (a*c \<le> b)" | |
| 1307 | proof - | |
| 1308 | assume less: "0<c" | |
| 1309 | hence "(a \<le> b/c) = (a*c \<le> (b/c)*c)" | |
| 1310 | by (simp add: mult_le_cancel_right order_less_not_sym [OF less]) | |
| 1311 | also have "... = (a*c \<le> b)" | |
| 1312 | by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) | |
| 1313 | finally show ?thesis . | |
| 1314 | qed | |
| 1315 | ||
| 1316 | lemma neg_le_divide_eq: "c < (0::'a::ordered_field) ==> (a \<le> b/c) = (b \<le> a*c)" | |
| 1317 | proof - | |
| 1318 | assume less: "c<0" | |
| 1319 | hence "(a \<le> b/c) = ((b/c)*c \<le> a*c)" | |
| 1320 | by (simp add: mult_le_cancel_right order_less_not_sym [OF less]) | |
| 1321 | also have "... = (b \<le> a*c)" | |
| 1322 | by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) | |
| 1323 | finally show ?thesis . | |
| 1324 | qed | |
| 1325 | ||
| 1326 | lemma le_divide_eq: | |
| 1327 | "(a \<le> b/c) = | |
| 1328 | (if 0 < c then a*c \<le> b | |
| 1329 | else if c < 0 then b \<le> a*c | |
| 1330 |              else  a \<le> (0::'a::{ordered_field,division_by_zero}))"
 | |
| 1331 | apply (case_tac "c=0", simp) | |
| 1332 | apply (force simp add: pos_le_divide_eq neg_le_divide_eq linorder_neq_iff) | |
| 1333 | done | |
| 1334 | ||
| 1335 | lemma pos_divide_le_eq: "0 < (c::'a::ordered_field) ==> (b/c \<le> a) = (b \<le> a*c)" | |
| 1336 | proof - | |
| 1337 | assume less: "0<c" | |
| 1338 | hence "(b/c \<le> a) = ((b/c)*c \<le> a*c)" | |
| 1339 | by (simp add: mult_le_cancel_right order_less_not_sym [OF less]) | |
| 1340 | also have "... = (b \<le> a*c)" | |
| 1341 | by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) | |
| 1342 | finally show ?thesis . | |
| 1343 | qed | |
| 1344 | ||
| 1345 | lemma neg_divide_le_eq: "c < (0::'a::ordered_field) ==> (b/c \<le> a) = (a*c \<le> b)" | |
| 1346 | proof - | |
| 1347 | assume less: "c<0" | |
| 1348 | hence "(b/c \<le> a) = (a*c \<le> (b/c)*c)" | |
| 1349 | by (simp add: mult_le_cancel_right order_less_not_sym [OF less]) | |
| 1350 | also have "... = (a*c \<le> b)" | |
| 1351 | by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) | |
| 1352 | finally show ?thesis . | |
| 1353 | qed | |
| 1354 | ||
| 1355 | lemma divide_le_eq: | |
| 1356 | "(b/c \<le> a) = | |
| 1357 | (if 0 < c then b \<le> a*c | |
| 1358 | else if c < 0 then a*c \<le> b | |
| 1359 |              else 0 \<le> (a::'a::{ordered_field,division_by_zero}))"
 | |
| 1360 | apply (case_tac "c=0", simp) | |
| 1361 | apply (force simp add: pos_divide_le_eq neg_divide_le_eq linorder_neq_iff) | |
| 1362 | done | |
| 1363 | ||
| 1364 | lemma pos_less_divide_eq: | |
| 1365 | "0 < (c::'a::ordered_field) ==> (a < b/c) = (a*c < b)" | |
| 1366 | proof - | |
| 1367 | assume less: "0<c" | |
| 1368 | hence "(a < b/c) = (a*c < (b/c)*c)" | |
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1369 | by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less]) | 
| 14288 | 1370 | also have "... = (a*c < b)" | 
| 1371 | by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) | |
| 1372 | finally show ?thesis . | |
| 1373 | qed | |
| 1374 | ||
| 1375 | lemma neg_less_divide_eq: | |
| 1376 | "c < (0::'a::ordered_field) ==> (a < b/c) = (b < a*c)" | |
| 1377 | proof - | |
| 1378 | assume less: "c<0" | |
| 1379 | hence "(a < b/c) = ((b/c)*c < a*c)" | |
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1380 | by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less]) | 
| 14288 | 1381 | also have "... = (b < a*c)" | 
| 1382 | by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) | |
| 1383 | finally show ?thesis . | |
| 1384 | qed | |
| 1385 | ||
| 1386 | lemma less_divide_eq: | |
| 1387 | "(a < b/c) = | |
| 1388 | (if 0 < c then a*c < b | |
| 1389 | else if c < 0 then b < a*c | |
| 1390 |              else  a < (0::'a::{ordered_field,division_by_zero}))"
 | |
| 1391 | apply (case_tac "c=0", simp) | |
| 1392 | apply (force simp add: pos_less_divide_eq neg_less_divide_eq linorder_neq_iff) | |
| 1393 | done | |
| 1394 | ||
| 1395 | lemma pos_divide_less_eq: | |
| 1396 | "0 < (c::'a::ordered_field) ==> (b/c < a) = (b < a*c)" | |
| 1397 | proof - | |
| 1398 | assume less: "0<c" | |
| 1399 | hence "(b/c < a) = ((b/c)*c < a*c)" | |
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1400 | by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less]) | 
| 14288 | 1401 | also have "... = (b < a*c)" | 
| 1402 | by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) | |
| 1403 | finally show ?thesis . | |
| 1404 | qed | |
| 1405 | ||
| 1406 | lemma neg_divide_less_eq: | |
| 1407 | "c < (0::'a::ordered_field) ==> (b/c < a) = (a*c < b)" | |
| 1408 | proof - | |
| 1409 | assume less: "c<0" | |
| 1410 | hence "(b/c < a) = (a*c < (b/c)*c)" | |
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1411 | by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less]) | 
| 14288 | 1412 | also have "... = (a*c < b)" | 
| 1413 | by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) | |
| 1414 | finally show ?thesis . | |
| 1415 | qed | |
| 1416 | ||
| 1417 | lemma divide_less_eq: | |
| 1418 | "(b/c < a) = | |
| 1419 | (if 0 < c then b < a*c | |
| 1420 | else if c < 0 then a*c < b | |
| 1421 |              else 0 < (a::'a::{ordered_field,division_by_zero}))"
 | |
| 1422 | apply (case_tac "c=0", simp) | |
| 1423 | apply (force simp add: pos_divide_less_eq neg_divide_less_eq linorder_neq_iff) | |
| 1424 | done | |
| 1425 | ||
| 1426 | lemma nonzero_eq_divide_eq: "c\<noteq>0 ==> ((a::'a::field) = b/c) = (a*c = b)" | |
| 1427 | proof - | |
| 1428 | assume [simp]: "c\<noteq>0" | |
| 1429 | have "(a = b/c) = (a*c = (b/c)*c)" | |
| 1430 | by (simp add: field_mult_cancel_right) | |
| 1431 | also have "... = (a*c = b)" | |
| 1432 | by (simp add: divide_inverse mult_assoc) | |
| 1433 | finally show ?thesis . | |
| 1434 | qed | |
| 1435 | ||
| 1436 | lemma eq_divide_eq: | |
| 1437 |   "((a::'a::{field,division_by_zero}) = b/c) = (if c\<noteq>0 then a*c = b else a=0)"
 | |
| 1438 | by (simp add: nonzero_eq_divide_eq) | |
| 1439 | ||
| 1440 | lemma nonzero_divide_eq_eq: "c\<noteq>0 ==> (b/c = (a::'a::field)) = (b = a*c)" | |
| 1441 | proof - | |
| 1442 | assume [simp]: "c\<noteq>0" | |
| 1443 | have "(b/c = a) = ((b/c)*c = a*c)" | |
| 1444 | by (simp add: field_mult_cancel_right) | |
| 1445 | also have "... = (b = a*c)" | |
| 1446 | by (simp add: divide_inverse mult_assoc) | |
| 1447 | finally show ?thesis . | |
| 1448 | qed | |
| 1449 | ||
| 1450 | lemma divide_eq_eq: | |
| 1451 |   "(b/c = (a::'a::{field,division_by_zero})) = (if c\<noteq>0 then b = a*c else a=0)"
 | |
| 1452 | by (force simp add: nonzero_divide_eq_eq) | |
| 1453 | ||
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1454 | lemma divide_eq_imp: "(c::'a::{division_by_zero,field}) ~= 0 ==>
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1455 | b = a * c ==> b / c = a" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1456 | by (subst divide_eq_eq, simp) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1457 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1458 | lemma eq_divide_imp: "(c::'a::{division_by_zero,field}) ~= 0 ==>
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1459 | a * c = b ==> a = b / c" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1460 | by (subst eq_divide_eq, simp) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1461 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1462 | lemma frac_eq_eq: "(y::'a::field) ~= 0 ==> z ~= 0 ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1463 | (x / y = w / z) = (x * z = w * y)" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1464 | apply (subst nonzero_eq_divide_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1465 | apply assumption | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1466 | apply (subst times_divide_eq_left) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1467 | apply (erule nonzero_divide_eq_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1468 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1469 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1470 | subsection{*Division and Signs*}
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1471 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1472 | lemma zero_less_divide_iff: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1473 |      "((0::'a::{ordered_field,division_by_zero}) < a/b) = (0 < a & 0 < b | a < 0 & b < 0)"
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1474 | by (simp add: divide_inverse zero_less_mult_iff) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1475 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1476 | lemma divide_less_0_iff: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1477 |      "(a/b < (0::'a::{ordered_field,division_by_zero})) = 
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1478 | (0 < a & b < 0 | a < 0 & 0 < b)" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1479 | by (simp add: divide_inverse mult_less_0_iff) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1480 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1481 | lemma zero_le_divide_iff: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1482 |      "((0::'a::{ordered_field,division_by_zero}) \<le> a/b) =
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1483 | (0 \<le> a & 0 \<le> b | a \<le> 0 & b \<le> 0)" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1484 | by (simp add: divide_inverse zero_le_mult_iff) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1485 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1486 | lemma divide_le_0_iff: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1487 |      "(a/b \<le> (0::'a::{ordered_field,division_by_zero})) =
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1488 | (0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1489 | by (simp add: divide_inverse mult_le_0_iff) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1490 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1491 | lemma divide_eq_0_iff [simp]: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1492 |      "(a/b = 0) = (a=0 | b=(0::'a::{field,division_by_zero}))"
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1493 | by (simp add: divide_inverse field_mult_eq_0_iff) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1494 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1495 | lemma divide_pos_pos: "0 < (x::'a::ordered_field) ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1496 | 0 < y ==> 0 < x / y" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1497 | apply (subst pos_less_divide_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1498 | apply assumption | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1499 | apply simp | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1500 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1501 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1502 | lemma divide_nonneg_pos: "0 <= (x::'a::ordered_field) ==> 0 < y ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1503 | 0 <= x / y" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1504 | apply (subst pos_le_divide_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1505 | apply assumption | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1506 | apply simp | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1507 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1508 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1509 | lemma divide_neg_pos: "(x::'a::ordered_field) < 0 ==> 0 < y ==> x / y < 0" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1510 | apply (subst pos_divide_less_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1511 | apply assumption | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1512 | apply simp | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1513 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1514 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1515 | lemma divide_nonpos_pos: "(x::'a::ordered_field) <= 0 ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1516 | 0 < y ==> x / y <= 0" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1517 | apply (subst pos_divide_le_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1518 | apply assumption | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1519 | apply simp | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1520 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1521 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1522 | lemma divide_pos_neg: "0 < (x::'a::ordered_field) ==> y < 0 ==> x / y < 0" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1523 | apply (subst neg_divide_less_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1524 | apply assumption | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1525 | apply simp | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1526 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1527 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1528 | lemma divide_nonneg_neg: "0 <= (x::'a::ordered_field) ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1529 | y < 0 ==> x / y <= 0" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1530 | apply (subst neg_divide_le_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1531 | apply assumption | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1532 | apply simp | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1533 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1534 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1535 | lemma divide_neg_neg: "(x::'a::ordered_field) < 0 ==> y < 0 ==> 0 < x / y" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1536 | apply (subst neg_less_divide_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1537 | apply assumption | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1538 | apply simp | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1539 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1540 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1541 | lemma divide_nonpos_neg: "(x::'a::ordered_field) <= 0 ==> y < 0 ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1542 | 0 <= x / y" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1543 | apply (subst neg_le_divide_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1544 | apply assumption | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1545 | apply simp | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1546 | done | 
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1547 | |
| 14288 | 1548 | subsection{*Cancellation Laws for Division*}
 | 
| 1549 | ||
| 1550 | lemma divide_cancel_right [simp]: | |
| 1551 |      "(a/c = b/c) = (c = 0 | a = (b::'a::{field,division_by_zero}))"
 | |
| 1552 | apply (case_tac "c=0", simp) | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 1553 | apply (simp add: divide_inverse field_mult_cancel_right) | 
| 14288 | 1554 | done | 
| 1555 | ||
| 1556 | lemma divide_cancel_left [simp]: | |
| 1557 |      "(c/a = c/b) = (c = 0 | a = (b::'a::{field,division_by_zero}))" 
 | |
| 1558 | apply (case_tac "c=0", simp) | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 1559 | apply (simp add: divide_inverse field_mult_cancel_left) | 
| 14288 | 1560 | done | 
| 1561 | ||
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1562 | subsection {* Division and the Number One *}
 | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1563 | |
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1564 | text{*Simplify expressions equated with 1*}
 | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1565 | lemma divide_eq_1_iff [simp]: | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1566 |      "(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: 
14348diff
changeset | 1567 | apply (case_tac "b=0", simp) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1568 | apply (simp add: right_inverse_eq) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1569 | done | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1570 | |
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1571 | lemma one_eq_divide_iff [simp]: | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1572 |      "(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: 
14348diff
changeset | 1573 | by (simp add: eq_commute [of 1]) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1574 | |
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1575 | lemma zero_eq_1_divide_iff [simp]: | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1576 |      "((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: 
14348diff
changeset | 1577 | apply (case_tac "a=0", simp) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1578 | apply (auto simp add: nonzero_eq_divide_eq) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1579 | done | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1580 | |
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1581 | lemma one_divide_eq_0_iff [simp]: | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1582 |      "(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: 
14348diff
changeset | 1583 | apply (case_tac "a=0", simp) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1584 | apply (insert zero_neq_one [THEN not_sym]) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1585 | apply (auto simp add: nonzero_divide_eq_eq) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1586 | done | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1587 | |
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1588 | text{*Simplify expressions such as @{text "0 < 1/x"} to @{text "0 < x"}*}
 | 
| 18623 | 1589 | lemmas zero_less_divide_1_iff = zero_less_divide_iff [of 1, simplified] | 
| 1590 | lemmas divide_less_0_1_iff = divide_less_0_iff [of 1, simplified] | |
| 1591 | lemmas zero_le_divide_1_iff = zero_le_divide_iff [of 1, simplified] | |
| 1592 | lemmas divide_le_0_1_iff = divide_le_0_iff [of 1, simplified] | |
| 17085 | 1593 | |
| 1594 | declare zero_less_divide_1_iff [simp] | |
| 1595 | declare divide_less_0_1_iff [simp] | |
| 1596 | declare zero_le_divide_1_iff [simp] | |
| 1597 | declare divide_le_0_1_iff [simp] | |
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 1598 | |
| 14293 | 1599 | subsection {* Ordering Rules for Division *}
 | 
| 1600 | ||
| 1601 | lemma divide_strict_right_mono: | |
| 1602 | "[|a < b; 0 < c|] ==> a / c < b / (c::'a::ordered_field)" | |
| 1603 | by (simp add: order_less_imp_not_eq2 divide_inverse mult_strict_right_mono | |
| 1604 | positive_imp_inverse_positive) | |
| 1605 | ||
| 1606 | lemma divide_right_mono: | |
| 1607 |      "[|a \<le> b; 0 \<le> c|] ==> a/c \<le> b/(c::'a::{ordered_field,division_by_zero})"
 | |
| 1608 | by (force simp add: divide_strict_right_mono order_le_less) | |
| 1609 | ||
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1610 | lemma divide_right_mono_neg: "(a::'a::{division_by_zero,ordered_field}) <= b 
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1611 | ==> c <= 0 ==> b / c <= a / c" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1612 | apply (drule divide_right_mono [of _ _ "- c"]) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1613 | apply auto | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1614 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1615 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1616 | lemma divide_strict_right_mono_neg: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1617 | "[|b < a; c < 0|] ==> a / c < b / (c::'a::ordered_field)" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1618 | apply (drule divide_strict_right_mono [of _ _ "-c"], simp) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1619 | apply (simp add: order_less_imp_not_eq nonzero_minus_divide_right [symmetric]) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1620 | done | 
| 14293 | 1621 | |
| 1622 | text{*The last premise ensures that @{term a} and @{term b} 
 | |
| 1623 | have the same sign*} | |
| 1624 | lemma divide_strict_left_mono: | |
| 1625 | "[|b < a; 0 < c; 0 < a*b|] ==> c / a < c / (b::'a::ordered_field)" | |
| 1626 | by (force simp add: zero_less_mult_iff divide_inverse mult_strict_left_mono | |
| 1627 | order_less_imp_not_eq order_less_imp_not_eq2 | |
| 1628 | less_imp_inverse_less less_imp_inverse_less_neg) | |
| 1629 | ||
| 1630 | lemma divide_left_mono: | |
| 1631 | "[|b \<le> a; 0 \<le> c; 0 < a*b|] ==> c / a \<le> c / (b::'a::ordered_field)" | |
| 1632 | apply (subgoal_tac "a \<noteq> 0 & b \<noteq> 0") | |
| 1633 | prefer 2 | |
| 1634 | apply (force simp add: zero_less_mult_iff order_less_imp_not_eq) | |
| 1635 | apply (case_tac "c=0", simp add: divide_inverse) | |
| 1636 | apply (force simp add: divide_strict_left_mono order_le_less) | |
| 1637 | done | |
| 1638 | ||
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1639 | lemma divide_left_mono_neg: "(a::'a::{division_by_zero,ordered_field}) <= b 
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1640 | ==> c <= 0 ==> 0 < a * b ==> c / a <= c / b" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1641 | apply (drule divide_left_mono [of _ _ "- c"]) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1642 | apply (auto simp add: mult_commute) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1643 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1644 | |
| 14293 | 1645 | lemma divide_strict_left_mono_neg: | 
| 1646 | "[|a < b; c < 0; 0 < a*b|] ==> c / a < c / (b::'a::ordered_field)" | |
| 1647 | apply (subgoal_tac "a \<noteq> 0 & b \<noteq> 0") | |
| 1648 | prefer 2 | |
| 1649 | apply (force simp add: zero_less_mult_iff order_less_imp_not_eq) | |
| 1650 | apply (drule divide_strict_left_mono [of _ _ "-c"]) | |
| 1651 | apply (simp_all add: mult_commute nonzero_minus_divide_left [symmetric]) | |
| 1652 | done | |
| 1653 | ||
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1654 | text{*Simplify quotients that are compared with the value 1.*}
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1655 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1656 | lemma le_divide_eq_1: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1657 |   fixes a :: "'a :: {ordered_field,division_by_zero}"
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1658 | shows "(1 \<le> b / a) = ((0 < a & a \<le> b) | (a < 0 & b \<le> a))" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1659 | by (auto simp add: le_divide_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1660 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1661 | lemma divide_le_eq_1: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1662 |   fixes a :: "'a :: {ordered_field,division_by_zero}"
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1663 | shows "(b / a \<le> 1) = ((0 < a & b \<le> a) | (a < 0 & a \<le> b) | a=0)" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1664 | by (auto simp add: divide_le_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1665 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1666 | lemma less_divide_eq_1: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1667 |   fixes a :: "'a :: {ordered_field,division_by_zero}"
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1668 | shows "(1 < b / a) = ((0 < a & a < b) | (a < 0 & b < a))" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1669 | by (auto simp add: less_divide_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1670 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1671 | lemma divide_less_eq_1: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1672 |   fixes a :: "'a :: {ordered_field,division_by_zero}"
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1673 | shows "(b / a < 1) = ((0 < a & b < a) | (a < 0 & a < b) | a=0)" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1674 | by (auto simp add: divide_less_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1675 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1676 | subsection{*Conditional Simplification Rules: No Case Splits*}
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1677 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1678 | lemma le_divide_eq_1_pos [simp]: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1679 |   fixes a :: "'a :: {ordered_field,division_by_zero}"
 | 
| 18649 
bb99c2e705ca
tidied, and added missing thm divide_less_eq_1_neg
 paulson parents: 
18623diff
changeset | 1680 | shows "0 < a \<Longrightarrow> (1 \<le> b/a) = (a \<le> b)" | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1681 | by (auto simp add: le_divide_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1682 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1683 | lemma le_divide_eq_1_neg [simp]: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1684 |   fixes a :: "'a :: {ordered_field,division_by_zero}"
 | 
| 18649 
bb99c2e705ca
tidied, and added missing thm divide_less_eq_1_neg
 paulson parents: 
18623diff
changeset | 1685 | shows "a < 0 \<Longrightarrow> (1 \<le> b/a) = (b \<le> a)" | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1686 | by (auto simp add: le_divide_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1687 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1688 | lemma divide_le_eq_1_pos [simp]: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1689 |   fixes a :: "'a :: {ordered_field,division_by_zero}"
 | 
| 18649 
bb99c2e705ca
tidied, and added missing thm divide_less_eq_1_neg
 paulson parents: 
18623diff
changeset | 1690 | shows "0 < a \<Longrightarrow> (b/a \<le> 1) = (b \<le> a)" | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1691 | by (auto simp add: divide_le_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1692 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1693 | lemma divide_le_eq_1_neg [simp]: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1694 |   fixes a :: "'a :: {ordered_field,division_by_zero}"
 | 
| 18649 
bb99c2e705ca
tidied, and added missing thm divide_less_eq_1_neg
 paulson parents: 
18623diff
changeset | 1695 | shows "a < 0 \<Longrightarrow> (b/a \<le> 1) = (a \<le> b)" | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1696 | by (auto simp add: divide_le_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1697 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1698 | lemma less_divide_eq_1_pos [simp]: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1699 |   fixes a :: "'a :: {ordered_field,division_by_zero}"
 | 
| 18649 
bb99c2e705ca
tidied, and added missing thm divide_less_eq_1_neg
 paulson parents: 
18623diff
changeset | 1700 | shows "0 < a \<Longrightarrow> (1 < b/a) = (a < b)" | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1701 | by (auto simp add: less_divide_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1702 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1703 | lemma less_divide_eq_1_neg [simp]: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1704 |   fixes a :: "'a :: {ordered_field,division_by_zero}"
 | 
| 18649 
bb99c2e705ca
tidied, and added missing thm divide_less_eq_1_neg
 paulson parents: 
18623diff
changeset | 1705 | shows "a < 0 \<Longrightarrow> (1 < b/a) = (b < a)" | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1706 | by (auto simp add: less_divide_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1707 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1708 | lemma divide_less_eq_1_pos [simp]: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1709 |   fixes a :: "'a :: {ordered_field,division_by_zero}"
 | 
| 18649 
bb99c2e705ca
tidied, and added missing thm divide_less_eq_1_neg
 paulson parents: 
18623diff
changeset | 1710 | shows "0 < a \<Longrightarrow> (b/a < 1) = (b < a)" | 
| 
bb99c2e705ca
tidied, and added missing thm divide_less_eq_1_neg
 paulson parents: 
18623diff
changeset | 1711 | by (auto simp add: divide_less_eq) | 
| 
bb99c2e705ca
tidied, and added missing thm divide_less_eq_1_neg
 paulson parents: 
18623diff
changeset | 1712 | |
| 
bb99c2e705ca
tidied, and added missing thm divide_less_eq_1_neg
 paulson parents: 
18623diff
changeset | 1713 | lemma divide_less_eq_1_neg [simp]: | 
| 
bb99c2e705ca
tidied, and added missing thm divide_less_eq_1_neg
 paulson parents: 
18623diff
changeset | 1714 |   fixes a :: "'a :: {ordered_field,division_by_zero}"
 | 
| 
bb99c2e705ca
tidied, and added missing thm divide_less_eq_1_neg
 paulson parents: 
18623diff
changeset | 1715 | shows "a < 0 \<Longrightarrow> b/a < 1 <-> a < b" | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1716 | by (auto simp add: divide_less_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1717 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1718 | lemma eq_divide_eq_1 [simp]: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1719 |   fixes a :: "'a :: {ordered_field,division_by_zero}"
 | 
| 18649 
bb99c2e705ca
tidied, and added missing thm divide_less_eq_1_neg
 paulson parents: 
18623diff
changeset | 1720 | shows "(1 = b/a) = ((a \<noteq> 0 & a = b))" | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1721 | by (auto simp add: eq_divide_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1722 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1723 | lemma divide_eq_eq_1 [simp]: | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1724 |   fixes a :: "'a :: {ordered_field,division_by_zero}"
 | 
| 18649 
bb99c2e705ca
tidied, and added missing thm divide_less_eq_1_neg
 paulson parents: 
18623diff
changeset | 1725 | shows "(b/a = 1) = ((a \<noteq> 0 & a = b))" | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1726 | by (auto simp add: divide_eq_eq) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1727 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1728 | subsection {* Reasoning about inequalities with division *}
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1729 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1730 | lemma mult_right_le_one_le: "0 <= (x::'a::ordered_idom) ==> 0 <= y ==> y <= 1 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1731 | ==> x * y <= x" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1732 | by (auto simp add: mult_compare_simps); | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1733 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1734 | lemma mult_left_le_one_le: "0 <= (x::'a::ordered_idom) ==> 0 <= y ==> y <= 1 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1735 | ==> y * x <= x" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1736 | by (auto simp add: mult_compare_simps); | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1737 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1738 | lemma mult_imp_div_pos_le: "0 < (y::'a::ordered_field) ==> x <= z * y ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1739 | x / y <= z"; | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1740 | by (subst pos_divide_le_eq, assumption+); | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1741 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1742 | lemma mult_imp_le_div_pos: "0 < (y::'a::ordered_field) ==> z * y <= x ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1743 | z <= x / y"; | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1744 | by (subst pos_le_divide_eq, assumption+) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1745 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1746 | lemma mult_imp_div_pos_less: "0 < (y::'a::ordered_field) ==> x < z * y ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1747 | x / y < z" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1748 | by (subst pos_divide_less_eq, assumption+) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1749 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1750 | lemma mult_imp_less_div_pos: "0 < (y::'a::ordered_field) ==> z * y < x ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1751 | z < x / y" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1752 | by (subst pos_less_divide_eq, assumption+) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1753 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1754 | lemma frac_le: "(0::'a::ordered_field) <= x ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1755 | x <= y ==> 0 < w ==> w <= z ==> x / z <= y / w" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1756 | apply (rule mult_imp_div_pos_le) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1757 | apply simp; | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1758 | apply (subst times_divide_eq_left); | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1759 | apply (rule mult_imp_le_div_pos, assumption) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1760 | apply (rule mult_mono) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1761 | apply simp_all | 
| 14293 | 1762 | done | 
| 1763 | ||
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1764 | lemma frac_less: "(0::'a::ordered_field) <= x ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1765 | x < y ==> 0 < w ==> w <= z ==> x / z < y / w" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1766 | apply (rule mult_imp_div_pos_less) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1767 | apply simp; | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1768 | apply (subst times_divide_eq_left); | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1769 | apply (rule mult_imp_less_div_pos, assumption) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1770 | apply (erule mult_less_le_imp_less) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1771 | apply simp_all | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1772 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1773 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1774 | lemma frac_less2: "(0::'a::ordered_field) < x ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1775 | x <= y ==> 0 < w ==> w < z ==> x / z < y / w" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1776 | apply (rule mult_imp_div_pos_less) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1777 | apply simp_all | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1778 | apply (subst times_divide_eq_left); | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1779 | apply (rule mult_imp_less_div_pos, assumption) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1780 | apply (erule mult_le_less_imp_less) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1781 | apply simp_all | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1782 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1783 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1784 | lemmas times_divide_eq = times_divide_eq_right times_divide_eq_left | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1785 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1786 | text{*It's not obvious whether these should be simprules or not. 
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1787 | Their effect is to gather terms into one big fraction, like | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1788 | a*b*c / x*y*z. The rationale for that is unclear, but many proofs | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1789 | seem to need them.*} | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1790 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1791 | declare times_divide_eq [simp] | 
| 14293 | 1792 | |
| 1793 | subsection {* Ordered Fields are Dense *}
 | |
| 1794 | ||
| 14738 | 1795 | lemma less_add_one: "a < (a+1::'a::ordered_semidom)" | 
| 14293 | 1796 | proof - | 
| 14738 | 1797 | have "a+0 < (a+1::'a::ordered_semidom)" | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1798 | by (blast intro: zero_less_one add_strict_left_mono) | 
| 14293 | 1799 | thus ?thesis by simp | 
| 1800 | qed | |
| 1801 | ||
| 14738 | 1802 | lemma zero_less_two: "0 < (1+1::'a::ordered_semidom)" | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14353diff
changeset | 1803 | 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: 
14353diff
changeset | 1804 | |
| 14293 | 1805 | lemma less_half_sum: "a < b ==> a < (a+b) / (1+1::'a::ordered_field)" | 
| 1806 | by (simp add: zero_less_two pos_less_divide_eq right_distrib) | |
| 1807 | ||
| 1808 | lemma gt_half_sum: "a < b ==> (a+b)/(1+1::'a::ordered_field) < b" | |
| 1809 | by (simp add: zero_less_two pos_divide_less_eq right_distrib) | |
| 1810 | ||
| 1811 | lemma dense: "a < b ==> \<exists>r::'a::ordered_field. a < r & r < b" | |
| 1812 | by (blast intro!: less_half_sum gt_half_sum) | |
| 1813 | ||
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1814 | |
| 14293 | 1815 | subsection {* Absolute Value *}
 | 
| 1816 | ||
| 14738 | 1817 | lemma abs_one [simp]: "abs 1 = (1::'a::ordered_idom)" | 
| 14294 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1818 | 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: 
14293diff
changeset | 1819 | |
| 14738 | 1820 | lemma abs_le_mult: "abs (a * b) \<le> (abs a) * (abs (b::'a::lordered_ring))" | 
| 1821 | proof - | |
| 1822 | let ?x = "pprt a * pprt b - pprt a * nprt b - nprt a * pprt b + nprt a * nprt b" | |
| 1823 | let ?y = "pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b" | |
| 1824 | have a: "(abs a) * (abs b) = ?x" | |
| 1825 | by (simp only: abs_prts[of a] abs_prts[of b] ring_eq_simps) | |
| 1826 |   {
 | |
| 1827 | fix u v :: 'a | |
| 15481 | 1828 | have bh: "\<lbrakk>u = a; v = b\<rbrakk> \<Longrightarrow> | 
| 1829 | u * v = pprt a * pprt b + pprt a * nprt b + | |
| 1830 | nprt a * pprt b + nprt a * nprt b" | |
| 14738 | 1831 | apply (subst prts[of u], subst prts[of v]) | 
| 1832 | apply (simp add: left_distrib right_distrib add_ac) | |
| 1833 | done | |
| 1834 | } | |
| 1835 | note b = this[OF refl[of a] refl[of b]] | |
| 1836 | note addm = add_mono[of "0::'a" _ "0::'a", simplified] | |
| 1837 | note addm2 = add_mono[of _ "0::'a" _ "0::'a", simplified] | |
| 1838 | have xy: "- ?x <= ?y" | |
| 14754 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 obua parents: 
14738diff
changeset | 1839 | apply (simp) | 
| 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 obua parents: 
14738diff
changeset | 1840 | apply (rule_tac y="0::'a" in order_trans) | 
| 16568 | 1841 | apply (rule addm2) | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1842 | apply (simp_all add: mult_nonneg_nonneg mult_nonpos_nonpos) | 
| 16568 | 1843 | apply (rule addm) | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1844 | apply (simp_all add: mult_nonneg_nonneg mult_nonpos_nonpos) | 
| 14754 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 obua parents: 
14738diff
changeset | 1845 | done | 
| 14738 | 1846 | have yx: "?y <= ?x" | 
| 16568 | 1847 | apply (simp add:diff_def) | 
| 14754 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 obua parents: 
14738diff
changeset | 1848 | apply (rule_tac y=0 in order_trans) | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1849 | apply (rule addm2, (simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)+) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1850 | apply (rule addm, (simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)+) | 
| 14738 | 1851 | done | 
| 1852 | have i1: "a*b <= abs a * abs b" by (simp only: a b yx) | |
| 1853 | have i2: "- (abs a * abs b) <= a*b" by (simp only: a b xy) | |
| 1854 | show ?thesis | |
| 1855 | apply (rule abs_leI) | |
| 1856 | apply (simp add: i1) | |
| 1857 | apply (simp add: i2[simplified minus_le_iff]) | |
| 1858 | done | |
| 1859 | qed | |
| 14294 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1860 | |
| 14738 | 1861 | lemma abs_eq_mult: | 
| 1862 | assumes "(0 \<le> a \<or> a \<le> 0) \<and> (0 \<le> b \<or> b \<le> 0)" | |
| 1863 | shows "abs (a*b) = abs a * abs (b::'a::lordered_ring)" | |
| 1864 | proof - | |
| 1865 | have s: "(0 <= a*b) | (a*b <= 0)" | |
| 1866 | apply (auto) | |
| 1867 | apply (rule_tac split_mult_pos_le) | |
| 1868 | apply (rule_tac contrapos_np[of "a*b <= 0"]) | |
| 1869 | apply (simp) | |
| 1870 | apply (rule_tac split_mult_neg_le) | |
| 1871 | apply (insert prems) | |
| 1872 | apply (blast) | |
| 1873 | done | |
| 1874 | have mulprts: "a * b = (pprt a + nprt a) * (pprt b + nprt b)" | |
| 1875 | by (simp add: prts[symmetric]) | |
| 1876 | show ?thesis | |
| 1877 | proof cases | |
| 1878 | assume "0 <= a * b" | |
| 1879 | then show ?thesis | |
| 1880 | apply (simp_all add: mulprts abs_prts) | |
| 1881 | apply (insert prems) | |
| 14754 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 obua parents: 
14738diff
changeset | 1882 | apply (auto simp add: | 
| 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 obua parents: 
14738diff
changeset | 1883 | ring_eq_simps | 
| 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 obua parents: 
14738diff
changeset | 1884 | iff2imp[OF zero_le_iff_zero_nprt] iff2imp[OF le_zero_iff_zero_pprt] | 
| 15197 | 1885 | iff2imp[OF le_zero_iff_pprt_id] iff2imp[OF zero_le_iff_nprt_id]) | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1886 | apply(drule (1) mult_nonneg_nonpos[of a b], simp) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1887 | apply(drule (1) mult_nonneg_nonpos2[of b a], simp) | 
| 14738 | 1888 | done | 
| 1889 | next | |
| 1890 | assume "~(0 <= a*b)" | |
| 1891 | with s have "a*b <= 0" by simp | |
| 1892 | then show ?thesis | |
| 1893 | apply (simp_all add: mulprts abs_prts) | |
| 1894 | apply (insert prems) | |
| 15580 | 1895 | apply (auto simp add: ring_eq_simps) | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1896 | apply(drule (1) mult_nonneg_nonneg[of a b],simp) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1897 | apply(drule (1) mult_nonpos_nonpos[of a b],simp) | 
| 14738 | 1898 | done | 
| 1899 | qed | |
| 1900 | qed | |
| 14294 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1901 | |
| 14738 | 1902 | lemma abs_mult: "abs (a * b) = abs a * abs (b::'a::ordered_idom)" | 
| 1903 | by (simp add: abs_eq_mult linorder_linear) | |
| 14293 | 1904 | |
| 14738 | 1905 | lemma abs_mult_self: "abs a * abs a = a * (a::'a::ordered_idom)" | 
| 1906 | by (simp add: abs_if) | |
| 14294 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1907 | |
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1908 | lemma nonzero_abs_inverse: | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1909 | "a \<noteq> 0 ==> abs (inverse (a::'a::ordered_field)) = inverse (abs a)" | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1910 | 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: 
14293diff
changeset | 1911 | negative_imp_inverse_negative) | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1912 | apply (blast intro: positive_imp_inverse_positive elim: order_less_asym) | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1913 | done | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1914 | |
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1915 | lemma abs_inverse [simp]: | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1916 |      "abs (inverse (a::'a::{ordered_field,division_by_zero})) = 
 | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1917 | inverse (abs a)" | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1918 | apply (case_tac "a=0", simp) | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1919 | apply (simp add: nonzero_abs_inverse) | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1920 | done | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1921 | |
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1922 | lemma nonzero_abs_divide: | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1923 | "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: 
14293diff
changeset | 1924 | by (simp add: divide_inverse abs_mult nonzero_abs_inverse) | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1925 | |
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 1926 | lemma abs_divide [simp]: | 
| 14294 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1927 |      "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: 
14293diff
changeset | 1928 | apply (case_tac "b=0", simp) | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1929 | apply (simp add: nonzero_abs_divide) | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1930 | done | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1931 | |
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1932 | lemma abs_mult_less: | 
| 14738 | 1933 | "[| 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: 
14293diff
changeset | 1934 | proof - | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1935 | assume ac: "abs a < c" | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1936 | 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: 
14293diff
changeset | 1937 | assume "abs b < d" | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1938 | thus ?thesis by (simp add: ac cpos mult_strict_mono) | 
| 
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
 paulson parents: 
14293diff
changeset | 1939 | qed | 
| 14293 | 1940 | |
| 14738 | 1941 | lemma eq_minus_self_iff: "(a = -a) = (a = (0::'a::ordered_idom))" | 
| 1942 | by (force simp add: order_eq_iff le_minus_self_iff minus_le_self_iff) | |
| 1943 | ||
| 1944 | lemma less_minus_self_iff: "(a < -a) = (a < (0::'a::ordered_idom))" | |
| 1945 | by (simp add: order_less_le le_minus_self_iff eq_minus_self_iff) | |
| 1946 | ||
| 1947 | lemma abs_less_iff: "(abs a < b) = (a < b & -a < (b::'a::ordered_idom))" | |
| 1948 | apply (simp add: order_less_le abs_le_iff) | |
| 1949 | apply (auto simp add: abs_if minus_le_self_iff eq_minus_self_iff) | |
| 1950 | apply (simp add: le_minus_self_iff linorder_neq_iff) | |
| 1951 | done | |
| 1952 | ||
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1953 | lemma abs_mult_pos: "(0::'a::ordered_idom) <= x ==> | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1954 | (abs y) * x = abs (y * x)"; | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1955 | apply (subst abs_mult); | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1956 | apply simp; | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1957 | done; | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1958 | |
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1959 | lemma abs_div_pos: "(0::'a::{division_by_zero,ordered_field}) < y ==> 
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1960 | abs x / y = abs (x / y)"; | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1961 | apply (subst abs_divide); | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1962 | apply (simp add: order_less_imp_le); | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1963 | done; | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 1964 | |
| 19404 | 1965 | subsection {* Bounds of products via negative and positive Part *}
 | 
| 15178 | 1966 | |
| 15580 | 1967 | lemma mult_le_prts: | 
| 1968 | assumes | |
| 1969 | "a1 <= (a::'a::lordered_ring)" | |
| 1970 | "a <= a2" | |
| 1971 | "b1 <= b" | |
| 1972 | "b <= b2" | |
| 1973 | shows | |
| 1974 | "a * b <= pprt a2 * pprt b2 + pprt a1 * nprt b2 + nprt a2 * pprt b1 + nprt a1 * nprt b1" | |
| 1975 | proof - | |
| 1976 | have "a * b = (pprt a + nprt a) * (pprt b + nprt b)" | |
| 1977 | apply (subst prts[symmetric])+ | |
| 1978 | apply simp | |
| 1979 | done | |
| 1980 | then have "a * b = pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b" | |
| 1981 | by (simp add: ring_eq_simps) | |
| 1982 | moreover have "pprt a * pprt b <= pprt a2 * pprt b2" | |
| 1983 | by (simp_all add: prems mult_mono) | |
| 1984 | moreover have "pprt a * nprt b <= pprt a1 * nprt b2" | |
| 1985 | proof - | |
| 1986 | have "pprt a * nprt b <= pprt a * nprt b2" | |
| 1987 | by (simp add: mult_left_mono prems) | |
| 1988 | moreover have "pprt a * nprt b2 <= pprt a1 * nprt b2" | |
| 1989 | by (simp add: mult_right_mono_neg prems) | |
| 1990 | ultimately show ?thesis | |
| 1991 | by simp | |
| 1992 | qed | |
| 1993 | moreover have "nprt a * pprt b <= nprt a2 * pprt b1" | |
| 1994 | proof - | |
| 1995 | have "nprt a * pprt b <= nprt a2 * pprt b" | |
| 1996 | by (simp add: mult_right_mono prems) | |
| 1997 | moreover have "nprt a2 * pprt b <= nprt a2 * pprt b1" | |
| 1998 | by (simp add: mult_left_mono_neg prems) | |
| 1999 | ultimately show ?thesis | |
| 2000 | by simp | |
| 2001 | qed | |
| 2002 | moreover have "nprt a * nprt b <= nprt a1 * nprt b1" | |
| 2003 | proof - | |
| 2004 | have "nprt a * nprt b <= nprt a * nprt b1" | |
| 2005 | by (simp add: mult_left_mono_neg prems) | |
| 2006 | moreover have "nprt a * nprt b1 <= nprt a1 * nprt b1" | |
| 2007 | by (simp add: mult_right_mono_neg prems) | |
| 2008 | ultimately show ?thesis | |
| 2009 | by simp | |
| 2010 | qed | |
| 2011 | ultimately show ?thesis | |
| 2012 | by - (rule add_mono | simp)+ | |
| 2013 | qed | |
| 19404 | 2014 | |
| 2015 | lemma mult_ge_prts: | |
| 15178 | 2016 | assumes | 
| 19404 | 2017 | "a1 <= (a::'a::lordered_ring)" | 
| 2018 | "a <= a2" | |
| 2019 | "b1 <= b" | |
| 2020 | "b <= b2" | |
| 15178 | 2021 | shows | 
| 19404 | 2022 | "a * b >= nprt a1 * pprt b2 + nprt a2 * nprt b2 + pprt a1 * pprt b1 + pprt a2 * nprt b1" | 
| 2023 | proof - | |
| 2024 | from prems have a1:"- a2 <= -a" by auto | |
| 2025 | from prems have a2: "-a <= -a1" by auto | |
| 2026 | from mult_le_prts[of "-a2" "-a" "-a1" "b1" b "b2", OF a1 a2 prems(3) prems(4), simplified nprt_neg pprt_neg] | |
| 2027 | have le: "- (a * b) <= - nprt a1 * pprt b2 + - nprt a2 * nprt b2 + - pprt a1 * pprt b1 + - pprt a2 * nprt b1" by simp | |
| 2028 | then have "-(- nprt a1 * pprt b2 + - nprt a2 * nprt b2 + - pprt a1 * pprt b1 + - pprt a2 * nprt b1) <= a * b" | |
| 2029 | by (simp only: minus_le_iff) | |
| 2030 | then show ?thesis by simp | |
| 15178 | 2031 | qed | 
| 2032 | ||
| 14738 | 2033 | ML {*
 | 
| 14334 | 2034 | val left_distrib = thm "left_distrib"; | 
| 14738 | 2035 | val right_distrib = thm "right_distrib"; | 
| 2036 | val mult_commute = thm "mult_commute"; | |
| 2037 | val distrib = thm "distrib"; | |
| 2038 | val zero_neq_one = thm "zero_neq_one"; | |
| 2039 | val no_zero_divisors = thm "no_zero_divisors"; | |
| 14331 | 2040 | val left_inverse = thm "left_inverse"; | 
| 14738 | 2041 | val divide_inverse = thm "divide_inverse"; | 
| 2042 | val mult_zero_left = thm "mult_zero_left"; | |
| 2043 | val mult_zero_right = thm "mult_zero_right"; | |
| 2044 | val field_mult_eq_0_iff = thm "field_mult_eq_0_iff"; | |
| 2045 | val inverse_zero = thm "inverse_zero"; | |
| 2046 | val ring_distrib = thms "ring_distrib"; | |
| 2047 | val combine_common_factor = thm "combine_common_factor"; | |
| 2048 | val minus_mult_left = thm "minus_mult_left"; | |
| 2049 | val minus_mult_right = thm "minus_mult_right"; | |
| 2050 | val minus_mult_minus = thm "minus_mult_minus"; | |
| 2051 | val minus_mult_commute = thm "minus_mult_commute"; | |
| 2052 | val right_diff_distrib = thm "right_diff_distrib"; | |
| 2053 | val left_diff_distrib = thm "left_diff_distrib"; | |
| 2054 | val mult_left_mono = thm "mult_left_mono"; | |
| 2055 | val mult_right_mono = thm "mult_right_mono"; | |
| 2056 | val mult_strict_left_mono = thm "mult_strict_left_mono"; | |
| 2057 | val mult_strict_right_mono = thm "mult_strict_right_mono"; | |
| 2058 | val mult_mono = thm "mult_mono"; | |
| 2059 | val mult_strict_mono = thm "mult_strict_mono"; | |
| 2060 | val abs_if = thm "abs_if"; | |
| 2061 | val zero_less_one = thm "zero_less_one"; | |
| 2062 | val eq_add_iff1 = thm "eq_add_iff1"; | |
| 2063 | val eq_add_iff2 = thm "eq_add_iff2"; | |
| 2064 | val less_add_iff1 = thm "less_add_iff1"; | |
| 2065 | val less_add_iff2 = thm "less_add_iff2"; | |
| 2066 | val le_add_iff1 = thm "le_add_iff1"; | |
| 2067 | val le_add_iff2 = thm "le_add_iff2"; | |
| 2068 | val mult_left_le_imp_le = thm "mult_left_le_imp_le"; | |
| 2069 | val mult_right_le_imp_le = thm "mult_right_le_imp_le"; | |
| 2070 | val mult_left_less_imp_less = thm "mult_left_less_imp_less"; | |
| 2071 | val mult_right_less_imp_less = thm "mult_right_less_imp_less"; | |
| 2072 | val mult_strict_left_mono_neg = thm "mult_strict_left_mono_neg"; | |
| 2073 | val mult_left_mono_neg = thm "mult_left_mono_neg"; | |
| 2074 | val mult_strict_right_mono_neg = thm "mult_strict_right_mono_neg"; | |
| 2075 | val mult_right_mono_neg = thm "mult_right_mono_neg"; | |
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 2076 | (* | 
| 14738 | 2077 | val mult_pos = thm "mult_pos"; | 
| 2078 | val mult_pos_le = thm "mult_pos_le"; | |
| 2079 | val mult_pos_neg = thm "mult_pos_neg"; | |
| 2080 | val mult_pos_neg_le = thm "mult_pos_neg_le"; | |
| 2081 | val mult_pos_neg2 = thm "mult_pos_neg2"; | |
| 2082 | val mult_pos_neg2_le = thm "mult_pos_neg2_le"; | |
| 2083 | val mult_neg = thm "mult_neg"; | |
| 2084 | val mult_neg_le = thm "mult_neg_le"; | |
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16568diff
changeset | 2085 | *) | 
| 14738 | 2086 | val zero_less_mult_pos = thm "zero_less_mult_pos"; | 
| 2087 | val zero_less_mult_pos2 = thm "zero_less_mult_pos2"; | |
| 2088 | val zero_less_mult_iff = thm "zero_less_mult_iff"; | |
| 2089 | val mult_eq_0_iff = thm "mult_eq_0_iff"; | |
| 2090 | val zero_le_mult_iff = thm "zero_le_mult_iff"; | |
| 2091 | val mult_less_0_iff = thm "mult_less_0_iff"; | |
| 2092 | val mult_le_0_iff = thm "mult_le_0_iff"; | |
| 2093 | val split_mult_pos_le = thm "split_mult_pos_le"; | |
| 2094 | val split_mult_neg_le = thm "split_mult_neg_le"; | |
| 2095 | val zero_le_square = thm "zero_le_square"; | |
| 2096 | val zero_le_one = thm "zero_le_one"; | |
| 2097 | val not_one_le_zero = thm "not_one_le_zero"; | |
| 2098 | val not_one_less_zero = thm "not_one_less_zero"; | |
| 2099 | val mult_left_mono_neg = thm "mult_left_mono_neg"; | |
| 2100 | val mult_right_mono_neg = thm "mult_right_mono_neg"; | |
| 2101 | val mult_strict_mono = thm "mult_strict_mono"; | |
| 2102 | val mult_strict_mono' = thm "mult_strict_mono'"; | |
| 2103 | val mult_mono = thm "mult_mono"; | |
| 2104 | val less_1_mult = thm "less_1_mult"; | |
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 2105 | val mult_less_cancel_right_disj = thm "mult_less_cancel_right_disj"; | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 2106 | val mult_less_cancel_left_disj = thm "mult_less_cancel_left_disj"; | 
| 14738 | 2107 | val mult_less_cancel_right = thm "mult_less_cancel_right"; | 
| 2108 | val mult_less_cancel_left = thm "mult_less_cancel_left"; | |
| 2109 | val mult_le_cancel_right = thm "mult_le_cancel_right"; | |
| 2110 | val mult_le_cancel_left = thm "mult_le_cancel_left"; | |
| 2111 | val mult_less_imp_less_left = thm "mult_less_imp_less_left"; | |
| 2112 | val mult_less_imp_less_right = thm "mult_less_imp_less_right"; | |
| 2113 | val mult_cancel_right = thm "mult_cancel_right"; | |
| 2114 | val mult_cancel_left = thm "mult_cancel_left"; | |
| 2115 | val ring_eq_simps = thms "ring_eq_simps"; | |
| 2116 | val right_inverse = thm "right_inverse"; | |
| 2117 | val right_inverse_eq = thm "right_inverse_eq"; | |
| 2118 | val nonzero_inverse_eq_divide = thm "nonzero_inverse_eq_divide"; | |
| 2119 | val divide_self = thm "divide_self"; | |
| 2120 | val divide_zero = thm "divide_zero"; | |
| 2121 | val divide_zero_left = thm "divide_zero_left"; | |
| 2122 | val inverse_eq_divide = thm "inverse_eq_divide"; | |
| 2123 | val add_divide_distrib = thm "add_divide_distrib"; | |
| 2124 | val field_mult_eq_0_iff = thm "field_mult_eq_0_iff"; | |
| 2125 | val field_mult_cancel_right_lemma = thm "field_mult_cancel_right_lemma"; | |
| 2126 | val field_mult_cancel_right = thm "field_mult_cancel_right"; | |
| 2127 | val field_mult_cancel_left = thm "field_mult_cancel_left"; | |
| 2128 | val nonzero_imp_inverse_nonzero = thm "nonzero_imp_inverse_nonzero"; | |
| 2129 | val inverse_zero_imp_zero = thm "inverse_zero_imp_zero"; | |
| 2130 | val inverse_nonzero_imp_nonzero = thm "inverse_nonzero_imp_nonzero"; | |
| 2131 | val inverse_nonzero_iff_nonzero = thm "inverse_nonzero_iff_nonzero"; | |
| 2132 | val nonzero_inverse_minus_eq = thm "nonzero_inverse_minus_eq"; | |
| 2133 | val inverse_minus_eq = thm "inverse_minus_eq"; | |
| 2134 | val nonzero_inverse_eq_imp_eq = thm "nonzero_inverse_eq_imp_eq"; | |
| 2135 | val inverse_eq_imp_eq = thm "inverse_eq_imp_eq"; | |
| 2136 | val inverse_eq_iff_eq = thm "inverse_eq_iff_eq"; | |
| 2137 | val nonzero_inverse_inverse_eq = thm "nonzero_inverse_inverse_eq"; | |
| 2138 | val inverse_inverse_eq = thm "inverse_inverse_eq"; | |
| 2139 | val inverse_1 = thm "inverse_1"; | |
| 2140 | val nonzero_inverse_mult_distrib = thm "nonzero_inverse_mult_distrib"; | |
| 2141 | val inverse_mult_distrib = thm "inverse_mult_distrib"; | |
| 2142 | val inverse_add = thm "inverse_add"; | |
| 2143 | val inverse_divide = thm "inverse_divide"; | |
| 2144 | val nonzero_mult_divide_cancel_left = thm "nonzero_mult_divide_cancel_left"; | |
| 2145 | val mult_divide_cancel_left = thm "mult_divide_cancel_left"; | |
| 2146 | val nonzero_mult_divide_cancel_right = thm "nonzero_mult_divide_cancel_right"; | |
| 2147 | val mult_divide_cancel_right = thm "mult_divide_cancel_right"; | |
| 2148 | val mult_divide_cancel_eq_if = thm "mult_divide_cancel_eq_if"; | |
| 2149 | val divide_1 = thm "divide_1"; | |
| 2150 | val times_divide_eq_right = thm "times_divide_eq_right"; | |
| 2151 | val times_divide_eq_left = thm "times_divide_eq_left"; | |
| 2152 | val divide_divide_eq_right = thm "divide_divide_eq_right"; | |
| 2153 | val divide_divide_eq_left = thm "divide_divide_eq_left"; | |
| 2154 | val nonzero_minus_divide_left = thm "nonzero_minus_divide_left"; | |
| 2155 | val nonzero_minus_divide_right = thm "nonzero_minus_divide_right"; | |
| 2156 | val nonzero_minus_divide_divide = thm "nonzero_minus_divide_divide"; | |
| 2157 | val minus_divide_left = thm "minus_divide_left"; | |
| 2158 | val minus_divide_right = thm "minus_divide_right"; | |
| 2159 | val minus_divide_divide = thm "minus_divide_divide"; | |
| 2160 | val diff_divide_distrib = thm "diff_divide_distrib"; | |
| 2161 | val positive_imp_inverse_positive = thm "positive_imp_inverse_positive"; | |
| 2162 | val negative_imp_inverse_negative = thm "negative_imp_inverse_negative"; | |
| 2163 | val inverse_le_imp_le = thm "inverse_le_imp_le"; | |
| 2164 | val inverse_positive_imp_positive = thm "inverse_positive_imp_positive"; | |
| 2165 | val inverse_positive_iff_positive = thm "inverse_positive_iff_positive"; | |
| 2166 | val inverse_negative_imp_negative = thm "inverse_negative_imp_negative"; | |
| 2167 | val inverse_negative_iff_negative = thm "inverse_negative_iff_negative"; | |
| 2168 | val inverse_nonnegative_iff_nonnegative = thm "inverse_nonnegative_iff_nonnegative"; | |
| 2169 | val inverse_nonpositive_iff_nonpositive = thm "inverse_nonpositive_iff_nonpositive"; | |
| 2170 | val less_imp_inverse_less = thm "less_imp_inverse_less"; | |
| 2171 | val inverse_less_imp_less = thm "inverse_less_imp_less"; | |
| 2172 | val inverse_less_iff_less = thm "inverse_less_iff_less"; | |
| 2173 | val le_imp_inverse_le = thm "le_imp_inverse_le"; | |
| 2174 | val inverse_le_iff_le = thm "inverse_le_iff_le"; | |
| 2175 | val inverse_le_imp_le_neg = thm "inverse_le_imp_le_neg"; | |
| 2176 | val less_imp_inverse_less_neg = thm "less_imp_inverse_less_neg"; | |
| 2177 | val inverse_less_imp_less_neg = thm "inverse_less_imp_less_neg"; | |
| 2178 | val inverse_less_iff_less_neg = thm "inverse_less_iff_less_neg"; | |
| 2179 | val le_imp_inverse_le_neg = thm "le_imp_inverse_le_neg"; | |
| 2180 | val inverse_le_iff_le_neg = thm "inverse_le_iff_le_neg"; | |
| 2181 | val one_less_inverse_iff = thm "one_less_inverse_iff"; | |
| 2182 | val inverse_eq_1_iff = thm "inverse_eq_1_iff"; | |
| 2183 | val one_le_inverse_iff = thm "one_le_inverse_iff"; | |
| 2184 | val inverse_less_1_iff = thm "inverse_less_1_iff"; | |
| 2185 | val inverse_le_1_iff = thm "inverse_le_1_iff"; | |
| 2186 | val zero_less_divide_iff = thm "zero_less_divide_iff"; | |
| 2187 | val divide_less_0_iff = thm "divide_less_0_iff"; | |
| 2188 | val zero_le_divide_iff = thm "zero_le_divide_iff"; | |
| 2189 | val divide_le_0_iff = thm "divide_le_0_iff"; | |
| 2190 | val divide_eq_0_iff = thm "divide_eq_0_iff"; | |
| 2191 | val pos_le_divide_eq = thm "pos_le_divide_eq"; | |
| 2192 | val neg_le_divide_eq = thm "neg_le_divide_eq"; | |
| 2193 | val le_divide_eq = thm "le_divide_eq"; | |
| 2194 | val pos_divide_le_eq = thm "pos_divide_le_eq"; | |
| 2195 | val neg_divide_le_eq = thm "neg_divide_le_eq"; | |
| 2196 | val divide_le_eq = thm "divide_le_eq"; | |
| 2197 | val pos_less_divide_eq = thm "pos_less_divide_eq"; | |
| 2198 | val neg_less_divide_eq = thm "neg_less_divide_eq"; | |
| 2199 | val less_divide_eq = thm "less_divide_eq"; | |
| 2200 | val pos_divide_less_eq = thm "pos_divide_less_eq"; | |
| 2201 | val neg_divide_less_eq = thm "neg_divide_less_eq"; | |
| 2202 | val divide_less_eq = thm "divide_less_eq"; | |
| 2203 | val nonzero_eq_divide_eq = thm "nonzero_eq_divide_eq"; | |
| 2204 | val eq_divide_eq = thm "eq_divide_eq"; | |
| 2205 | val nonzero_divide_eq_eq = thm "nonzero_divide_eq_eq"; | |
| 2206 | val divide_eq_eq = thm "divide_eq_eq"; | |
| 2207 | val divide_cancel_right = thm "divide_cancel_right"; | |
| 2208 | val divide_cancel_left = thm "divide_cancel_left"; | |
| 2209 | val divide_eq_1_iff = thm "divide_eq_1_iff"; | |
| 2210 | val one_eq_divide_iff = thm "one_eq_divide_iff"; | |
| 2211 | val zero_eq_1_divide_iff = thm "zero_eq_1_divide_iff"; | |
| 2212 | val one_divide_eq_0_iff = thm "one_divide_eq_0_iff"; | |
| 2213 | val divide_strict_right_mono = thm "divide_strict_right_mono"; | |
| 2214 | val divide_right_mono = thm "divide_right_mono"; | |
| 2215 | val divide_strict_left_mono = thm "divide_strict_left_mono"; | |
| 2216 | val divide_left_mono = thm "divide_left_mono"; | |
| 2217 | val divide_strict_left_mono_neg = thm "divide_strict_left_mono_neg"; | |
| 2218 | val divide_strict_right_mono_neg = thm "divide_strict_right_mono_neg"; | |
| 2219 | val less_add_one = thm "less_add_one"; | |
| 2220 | val zero_less_two = thm "zero_less_two"; | |
| 2221 | val less_half_sum = thm "less_half_sum"; | |
| 2222 | val gt_half_sum = thm "gt_half_sum"; | |
| 2223 | val dense = thm "dense"; | |
| 2224 | val abs_one = thm "abs_one"; | |
| 2225 | val abs_le_mult = thm "abs_le_mult"; | |
| 2226 | val abs_eq_mult = thm "abs_eq_mult"; | |
| 2227 | val abs_mult = thm "abs_mult"; | |
| 2228 | val abs_mult_self = thm "abs_mult_self"; | |
| 2229 | val nonzero_abs_inverse = thm "nonzero_abs_inverse"; | |
| 2230 | val abs_inverse = thm "abs_inverse"; | |
| 2231 | val nonzero_abs_divide = thm "nonzero_abs_divide"; | |
| 2232 | val abs_divide = thm "abs_divide"; | |
| 2233 | val abs_mult_less = thm "abs_mult_less"; | |
| 2234 | val eq_minus_self_iff = thm "eq_minus_self_iff"; | |
| 2235 | val less_minus_self_iff = thm "less_minus_self_iff"; | |
| 2236 | val abs_less_iff = thm "abs_less_iff"; | |
| 14331 | 2237 | *} | 
| 2238 | ||
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: diff
changeset | 2239 | end |