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