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