| author | chaieb | 
| Wed, 22 Aug 2007 17:13:41 +0200 | |
| changeset 24403 | b7c3ee2ca184 | 
| parent 24380 | c215e256beca | 
| child 24748 | ee0a0eb6b738 | 
| permissions | -rw-r--r-- | 
| 14770 | 1  | 
(* Title: HOL/OrderedGroup.thy  | 
| 14738 | 2  | 
ID: $Id$  | 
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
3  | 
Author: Gertrud Bauer, Steven Obua, Lawrence C Paulson, and Markus Wenzel,  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
4  | 
with contributions by Jeremy Avigad  | 
| 14738 | 5  | 
*)  | 
6  | 
||
7  | 
header {* Ordered Groups *}
 | 
|
8  | 
||
| 15131 | 9  | 
theory OrderedGroup  | 
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
10  | 
imports Lattices  | 
| 19798 | 11  | 
uses "~~/src/Provers/Arith/abel_cancel.ML"  | 
| 15131 | 12  | 
begin  | 
| 14738 | 13  | 
|
14  | 
text {*
 | 
|
15  | 
The theory of partially ordered groups is taken from the books:  | 
|
16  | 
  \begin{itemize}
 | 
|
17  | 
  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
 | 
|
18  | 
  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
 | 
|
19  | 
  \end{itemize}
 | 
|
20  | 
Most of the used notions can also be looked up in  | 
|
21  | 
  \begin{itemize}
 | 
|
| 14770 | 22  | 
  \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
 | 
| 14738 | 23  | 
  \item \emph{Algebra I} by van der Waerden, Springer.
 | 
24  | 
  \end{itemize}
 | 
|
25  | 
*}  | 
|
26  | 
||
| 23085 | 27  | 
subsection {* Semigroups and Monoids *}
 | 
| 14738 | 28  | 
|
| 22390 | 29  | 
class semigroup_add = plus +  | 
30  | 
assumes add_assoc: "(a \<^loc>+ b) \<^loc>+ c = a \<^loc>+ (b \<^loc>+ c)"  | 
|
31  | 
||
32  | 
class ab_semigroup_add = semigroup_add +  | 
|
33  | 
assumes add_commute: "a \<^loc>+ b = b \<^loc>+ a"  | 
|
| 14738 | 34  | 
|
35  | 
lemma add_left_commute: "a + (b + c) = b + (a + (c::'a::ab_semigroup_add))"  | 
|
36  | 
by (rule mk_left_commute [of "op +", OF add_assoc add_commute])  | 
|
37  | 
||
38  | 
theorems add_ac = add_assoc add_commute add_left_commute  | 
|
39  | 
||
| 22390 | 40  | 
class semigroup_mult = times +  | 
41  | 
assumes mult_assoc: "(a \<^loc>* b) \<^loc>* c = a \<^loc>* (b \<^loc>* c)"  | 
|
| 14738 | 42  | 
|
| 22390 | 43  | 
class ab_semigroup_mult = semigroup_mult +  | 
44  | 
assumes mult_commute: "a \<^loc>* b = b \<^loc>* a"  | 
|
| 23181 | 45  | 
begin  | 
| 14738 | 46  | 
|
| 23181 | 47  | 
lemma mult_left_commute: "a \<^loc>* (b \<^loc>* c) = b \<^loc>* (a \<^loc>* c)"  | 
48  | 
by (rule mk_left_commute [of "op \<^loc>*", OF mult_assoc mult_commute])  | 
|
49  | 
||
50  | 
end  | 
|
| 14738 | 51  | 
|
52  | 
theorems mult_ac = mult_assoc mult_commute mult_left_commute  | 
|
53  | 
||
| 23085 | 54  | 
class monoid_add = zero + semigroup_add +  | 
55  | 
assumes add_0_left [simp]: "\<^loc>0 \<^loc>+ a = a" and add_0_right [simp]: "a \<^loc>+ \<^loc>0 = a"  | 
|
56  | 
||
| 22390 | 57  | 
class comm_monoid_add = zero + ab_semigroup_add +  | 
| 23085 | 58  | 
assumes add_0: "\<^loc>0 \<^loc>+ a = a"  | 
59  | 
||
60  | 
instance comm_monoid_add < monoid_add  | 
|
61  | 
by intro_classes (insert comm_monoid_add_class.zero_plus.add_0, simp_all add: add_commute, auto)  | 
|
| 14738 | 62  | 
|
| 22390 | 63  | 
class monoid_mult = one + semigroup_mult +  | 
64  | 
assumes mult_1_left [simp]: "\<^loc>1 \<^loc>* a = a"  | 
|
65  | 
assumes mult_1_right [simp]: "a \<^loc>* \<^loc>1 = a"  | 
|
| 14738 | 66  | 
|
| 22390 | 67  | 
class comm_monoid_mult = one + ab_semigroup_mult +  | 
68  | 
assumes mult_1: "\<^loc>1 \<^loc>* a = a"  | 
|
| 14738 | 69  | 
|
70  | 
instance comm_monoid_mult \<subseteq> monoid_mult  | 
|
| 22390 | 71  | 
by intro_classes (insert mult_1, simp_all add: mult_commute, auto)  | 
| 14738 | 72  | 
|
| 22390 | 73  | 
class cancel_semigroup_add = semigroup_add +  | 
74  | 
assumes add_left_imp_eq: "a \<^loc>+ b = a \<^loc>+ c \<Longrightarrow> b = c"  | 
|
75  | 
assumes add_right_imp_eq: "b \<^loc>+ a = c \<^loc>+ a \<Longrightarrow> b = c"  | 
|
| 14738 | 76  | 
|
| 22390 | 77  | 
class cancel_ab_semigroup_add = ab_semigroup_add +  | 
78  | 
assumes add_imp_eq: "a \<^loc>+ b = a \<^loc>+ c \<Longrightarrow> b = c"  | 
|
| 14738 | 79  | 
|
80  | 
instance cancel_ab_semigroup_add \<subseteq> cancel_semigroup_add  | 
|
| 22390 | 81  | 
proof intro_classes  | 
82  | 
fix a b c :: 'a  | 
|
83  | 
assume "a + b = a + c"  | 
|
84  | 
then show "b = c" by (rule add_imp_eq)  | 
|
85  | 
next  | 
|
| 14738 | 86  | 
fix a b c :: 'a  | 
87  | 
assume "b + a = c + a"  | 
|
| 22390 | 88  | 
then have "a + b = a + c" by (simp only: add_commute)  | 
89  | 
then show "b = c" by (rule add_imp_eq)  | 
|
| 14738 | 90  | 
qed  | 
91  | 
||
| 23085 | 92  | 
lemma add_left_cancel [simp]:  | 
93  | 
"a + b = a + c \<longleftrightarrow> b = (c \<Colon> 'a\<Colon>cancel_semigroup_add)"  | 
|
94  | 
by (blast dest: add_left_imp_eq)  | 
|
95  | 
||
96  | 
lemma add_right_cancel [simp]:  | 
|
97  | 
"b + a = c + a \<longleftrightarrow> b = (c \<Colon> 'a\<Colon>cancel_semigroup_add)"  | 
|
98  | 
by (blast dest: add_right_imp_eq)  | 
|
99  | 
||
100  | 
subsection {* Groups *}
 | 
|
101  | 
||
| 22390 | 102  | 
class ab_group_add = minus + comm_monoid_add +  | 
| 23085 | 103  | 
assumes ab_left_minus: "uminus a \<^loc>+ a = \<^loc>0"  | 
104  | 
assumes ab_diff_minus: "a \<^loc>- b = a \<^loc>+ (uminus b)"  | 
|
105  | 
||
106  | 
class group_add = minus + monoid_add +  | 
|
| 22390 | 107  | 
assumes left_minus [simp]: "uminus a \<^loc>+ a = \<^loc>0"  | 
108  | 
assumes diff_minus: "a \<^loc>- b = a \<^loc>+ (uminus b)"  | 
|
| 14738 | 109  | 
|
| 23085 | 110  | 
instance ab_group_add < group_add  | 
111  | 
by intro_classes (simp_all add: ab_left_minus ab_diff_minus)  | 
|
112  | 
||
| 14738 | 113  | 
instance ab_group_add \<subseteq> cancel_ab_semigroup_add  | 
| 22390 | 114  | 
proof intro_classes  | 
| 14738 | 115  | 
fix a b c :: 'a  | 
116  | 
assume "a + b = a + c"  | 
|
| 22390 | 117  | 
then have "uminus a + a + b = uminus a + a + c" unfolding add_assoc by simp  | 
| 23085 | 118  | 
then show "b = c" by simp  | 
| 14738 | 119  | 
qed  | 
120  | 
||
| 23085 | 121  | 
lemma minus_add_cancel: "-(a::'a::group_add) + (a+b) = b"  | 
122  | 
by(simp add:add_assoc[symmetric])  | 
|
123  | 
||
124  | 
lemma minus_zero[simp]: "-(0::'a::group_add) = 0"  | 
|
| 14738 | 125  | 
proof -  | 
| 23085 | 126  | 
have "-(0::'a::group_add) = - 0 + (0+0)" by(simp only: add_0_right)  | 
127  | 
also have "\<dots> = 0" by(rule minus_add_cancel)  | 
|
| 14738 | 128  | 
finally show ?thesis .  | 
129  | 
qed  | 
|
130  | 
||
| 23085 | 131  | 
lemma minus_minus[simp]: "- (-(a::'a::group_add)) = a"  | 
132  | 
proof -  | 
|
133  | 
have "-(-a) = -(-a) + (-a + a)" by simp  | 
|
134  | 
also have "\<dots> = a" by(rule minus_add_cancel)  | 
|
135  | 
finally show ?thesis .  | 
|
136  | 
qed  | 
|
| 14738 | 137  | 
|
| 23085 | 138  | 
lemma right_minus[simp]: "a + - a = (0::'a::group_add)"  | 
| 14738 | 139  | 
proof -  | 
| 23085 | 140  | 
have "a + -a = -(-a) + -a" by simp  | 
141  | 
also have "\<dots> = 0" thm group_add.left_minus by(rule left_minus)  | 
|
| 14738 | 142  | 
finally show ?thesis .  | 
143  | 
qed  | 
|
144  | 
||
| 23085 | 145  | 
lemma right_minus_eq: "(a - b = 0) = (a = (b::'a::group_add))"  | 
| 14738 | 146  | 
proof  | 
| 23085 | 147  | 
assume "a - b = 0"  | 
148  | 
have "a = (a - b) + b" by (simp add:diff_minus add_assoc)  | 
|
149  | 
also have "\<dots> = b" using `a - b = 0` by simp  | 
|
150  | 
finally show "a = b" .  | 
|
| 14738 | 151  | 
next  | 
| 23085 | 152  | 
assume "a = b" thus "a - b = 0" by (simp add: diff_minus)  | 
| 14738 | 153  | 
qed  | 
154  | 
||
| 23085 | 155  | 
lemma equals_zero_I: assumes "a+b = 0" shows "-a = (b::'a::group_add)"  | 
156  | 
proof -  | 
|
157  | 
have "- a = -a + (a+b)" using assms by simp  | 
|
158  | 
also have "\<dots> = b" by(simp add:add_assoc[symmetric])  | 
|
159  | 
finally show ?thesis .  | 
|
160  | 
qed  | 
|
| 14738 | 161  | 
|
| 23085 | 162  | 
lemma diff_self[simp]: "(a::'a::group_add) - a = 0"  | 
163  | 
by(simp add: diff_minus)  | 
|
| 14738 | 164  | 
|
| 23085 | 165  | 
lemma diff_0 [simp]: "(0::'a::group_add) - a = -a"  | 
| 14738 | 166  | 
by (simp add: diff_minus)  | 
167  | 
||
| 23085 | 168  | 
lemma diff_0_right [simp]: "a - (0::'a::group_add) = a"  | 
| 14738 | 169  | 
by (simp add: diff_minus)  | 
170  | 
||
| 23085 | 171  | 
lemma diff_minus_eq_add [simp]: "a - - b = a + (b::'a::group_add)"  | 
| 14738 | 172  | 
by (simp add: diff_minus)  | 
173  | 
||
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23389 
diff
changeset
 | 
174  | 
lemma uminus_add_conv_diff: "-a + b = b - (a::'a::ab_group_add)"  | 
| 
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23389 
diff
changeset
 | 
175  | 
by(simp add:diff_minus add_commute)  | 
| 
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23389 
diff
changeset
 | 
176  | 
|
| 23085 | 177  | 
lemma neg_equal_iff_equal [simp]: "(-a = -b) = (a = (b::'a::group_add))"  | 
| 14738 | 178  | 
proof  | 
179  | 
assume "- a = - b"  | 
|
180  | 
hence "- (- a) = - (- b)"  | 
|
181  | 
by simp  | 
|
182  | 
thus "a=b" by simp  | 
|
183  | 
next  | 
|
184  | 
assume "a=b"  | 
|
185  | 
thus "-a = -b" by simp  | 
|
186  | 
qed  | 
|
187  | 
||
| 23085 | 188  | 
lemma neg_equal_0_iff_equal [simp]: "(-a = 0) = (a = (0::'a::group_add))"  | 
| 14738 | 189  | 
by (subst neg_equal_iff_equal [symmetric], simp)  | 
190  | 
||
| 23085 | 191  | 
lemma neg_0_equal_iff_equal [simp]: "(0 = -a) = (0 = (a::'a::group_add))"  | 
| 14738 | 192  | 
by (subst neg_equal_iff_equal [symmetric], simp)  | 
193  | 
||
194  | 
text{*The next two equations can make the simplifier loop!*}
 | 
|
195  | 
||
| 23085 | 196  | 
lemma equation_minus_iff: "(a = - b) = (b = - (a::'a::group_add))"  | 
| 14738 | 197  | 
proof -  | 
198  | 
have "(- (-a) = - b) = (- a = b)" by (rule neg_equal_iff_equal)  | 
|
199  | 
thus ?thesis by (simp add: eq_commute)  | 
|
200  | 
qed  | 
|
201  | 
||
| 23085 | 202  | 
lemma minus_equation_iff: "(- a = b) = (- (b::'a::group_add) = a)"  | 
| 14738 | 203  | 
proof -  | 
204  | 
have "(- a = - (-b)) = (a = -b)" by (rule neg_equal_iff_equal)  | 
|
205  | 
thus ?thesis by (simp add: eq_commute)  | 
|
206  | 
qed  | 
|
207  | 
||
208  | 
lemma minus_add_distrib [simp]: "- (a + b) = -a + -(b::'a::ab_group_add)"  | 
|
209  | 
apply (rule equals_zero_I)  | 
|
| 23085 | 210  | 
apply (simp add: add_ac)  | 
| 14738 | 211  | 
done  | 
212  | 
||
213  | 
lemma minus_diff_eq [simp]: "- (a - b) = b - (a::'a::ab_group_add)"  | 
|
214  | 
by (simp add: diff_minus add_commute)  | 
|
215  | 
||
216  | 
subsection {* (Partially) Ordered Groups *} 
 | 
|
217  | 
||
| 22390 | 218  | 
class pordered_ab_semigroup_add = order + ab_semigroup_add +  | 
219  | 
assumes add_left_mono: "a \<sqsubseteq> b \<Longrightarrow> c \<^loc>+ a \<sqsubseteq> c \<^loc>+ b"  | 
|
| 14738 | 220  | 
|
| 22390 | 221  | 
class pordered_cancel_ab_semigroup_add =  | 
222  | 
pordered_ab_semigroup_add + cancel_ab_semigroup_add  | 
|
| 14738 | 223  | 
|
| 22390 | 224  | 
class pordered_ab_semigroup_add_imp_le = pordered_cancel_ab_semigroup_add +  | 
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
225  | 
assumes add_le_imp_le_left: "c \<^loc>+ a \<sqsubseteq> c \<^loc>+ b \<Longrightarrow> a \<sqsubseteq> b"  | 
| 14738 | 226  | 
|
| 22390 | 227  | 
class pordered_ab_group_add = ab_group_add + pordered_ab_semigroup_add  | 
| 14738 | 228  | 
|
229  | 
instance pordered_ab_group_add \<subseteq> pordered_ab_semigroup_add_imp_le  | 
|
230  | 
proof  | 
|
231  | 
fix a b c :: 'a  | 
|
232  | 
assume "c + a \<le> c + b"  | 
|
233  | 
hence "(-c) + (c + a) \<le> (-c) + (c + b)" by (rule add_left_mono)  | 
|
234  | 
hence "((-c) + c) + a \<le> ((-c) + c) + b" by (simp only: add_assoc)  | 
|
235  | 
thus "a \<le> b" by simp  | 
|
236  | 
qed  | 
|
237  | 
||
| 
24380
 
c215e256beca
moved ordered_ab_semigroup_add to OrderedGroup.thy
 
haftmann 
parents: 
24286 
diff
changeset
 | 
238  | 
class ordered_ab_semigroup_add =  | 
| 
 
c215e256beca
moved ordered_ab_semigroup_add to OrderedGroup.thy
 
haftmann 
parents: 
24286 
diff
changeset
 | 
239  | 
linorder + pordered_ab_semigroup_add  | 
| 
 
c215e256beca
moved ordered_ab_semigroup_add to OrderedGroup.thy
 
haftmann 
parents: 
24286 
diff
changeset
 | 
240  | 
|
| 
 
c215e256beca
moved ordered_ab_semigroup_add to OrderedGroup.thy
 
haftmann 
parents: 
24286 
diff
changeset
 | 
241  | 
class ordered_cancel_ab_semigroup_add =  | 
| 
 
c215e256beca
moved ordered_ab_semigroup_add to OrderedGroup.thy
 
haftmann 
parents: 
24286 
diff
changeset
 | 
242  | 
linorder + pordered_cancel_ab_semigroup_add  | 
| 
 
c215e256beca
moved ordered_ab_semigroup_add to OrderedGroup.thy
 
haftmann 
parents: 
24286 
diff
changeset
 | 
243  | 
|
| 
 
c215e256beca
moved ordered_ab_semigroup_add to OrderedGroup.thy
 
haftmann 
parents: 
24286 
diff
changeset
 | 
244  | 
instance ordered_cancel_ab_semigroup_add \<subseteq> ordered_ab_semigroup_add ..  | 
| 14738 | 245  | 
|
246  | 
instance ordered_cancel_ab_semigroup_add \<subseteq> pordered_ab_semigroup_add_imp_le  | 
|
247  | 
proof  | 
|
248  | 
fix a b c :: 'a  | 
|
249  | 
assume le: "c + a <= c + b"  | 
|
250  | 
show "a <= b"  | 
|
251  | 
proof (rule ccontr)  | 
|
252  | 
assume w: "~ a \<le> b"  | 
|
253  | 
hence "b <= a" by (simp add: linorder_not_le)  | 
|
254  | 
hence le2: "c+b <= c+a" by (rule add_left_mono)  | 
|
255  | 
have "a = b"  | 
|
256  | 
apply (insert le)  | 
|
257  | 
apply (insert le2)  | 
|
258  | 
apply (drule order_antisym, simp_all)  | 
|
259  | 
done  | 
|
260  | 
with w show False  | 
|
261  | 
by (simp add: linorder_not_le [symmetric])  | 
|
262  | 
qed  | 
|
263  | 
qed  | 
|
264  | 
||
265  | 
lemma add_right_mono: "a \<le> (b::'a::pordered_ab_semigroup_add) ==> a + c \<le> b + c"  | 
|
| 22390 | 266  | 
by (simp add: add_commute [of _ c] add_left_mono)  | 
| 14738 | 267  | 
|
268  | 
text {* non-strict, in both arguments *}
 | 
|
269  | 
lemma add_mono:  | 
|
270  | 
"[|a \<le> b; c \<le> d|] ==> a + c \<le> b + (d::'a::pordered_ab_semigroup_add)"  | 
|
271  | 
apply (erule add_right_mono [THEN order_trans])  | 
|
272  | 
apply (simp add: add_commute add_left_mono)  | 
|
273  | 
done  | 
|
274  | 
||
275  | 
lemma add_strict_left_mono:  | 
|
276  | 
"a < b ==> c + a < c + (b::'a::pordered_cancel_ab_semigroup_add)"  | 
|
277  | 
by (simp add: order_less_le add_left_mono)  | 
|
278  | 
||
279  | 
lemma add_strict_right_mono:  | 
|
280  | 
"a < b ==> a + c < b + (c::'a::pordered_cancel_ab_semigroup_add)"  | 
|
281  | 
by (simp add: add_commute [of _ c] add_strict_left_mono)  | 
|
282  | 
||
283  | 
text{*Strict monotonicity in both arguments*}
 | 
|
284  | 
lemma add_strict_mono: "[|a<b; c<d|] ==> a + c < b + (d::'a::pordered_cancel_ab_semigroup_add)"  | 
|
285  | 
apply (erule add_strict_right_mono [THEN order_less_trans])  | 
|
286  | 
apply (erule add_strict_left_mono)  | 
|
287  | 
done  | 
|
288  | 
||
289  | 
lemma add_less_le_mono:  | 
|
290  | 
"[| a<b; c\<le>d |] ==> a + c < b + (d::'a::pordered_cancel_ab_semigroup_add)"  | 
|
291  | 
apply (erule add_strict_right_mono [THEN order_less_le_trans])  | 
|
292  | 
apply (erule add_left_mono)  | 
|
293  | 
done  | 
|
294  | 
||
295  | 
lemma add_le_less_mono:  | 
|
296  | 
"[| a\<le>b; c<d |] ==> a + c < b + (d::'a::pordered_cancel_ab_semigroup_add)"  | 
|
297  | 
apply (erule add_right_mono [THEN order_le_less_trans])  | 
|
298  | 
apply (erule add_strict_left_mono)  | 
|
299  | 
done  | 
|
300  | 
||
301  | 
lemma add_less_imp_less_left:  | 
|
302  | 
assumes less: "c + a < c + b" shows "a < (b::'a::pordered_ab_semigroup_add_imp_le)"  | 
|
303  | 
proof -  | 
|
304  | 
from less have le: "c + a <= c + b" by (simp add: order_le_less)  | 
|
305  | 
have "a <= b"  | 
|
306  | 
apply (insert le)  | 
|
307  | 
apply (drule add_le_imp_le_left)  | 
|
308  | 
by (insert le, drule add_le_imp_le_left, assumption)  | 
|
309  | 
moreover have "a \<noteq> b"  | 
|
310  | 
proof (rule ccontr)  | 
|
311  | 
assume "~(a \<noteq> b)"  | 
|
312  | 
then have "a = b" by simp  | 
|
313  | 
then have "c + a = c + b" by simp  | 
|
314  | 
with less show "False"by simp  | 
|
315  | 
qed  | 
|
316  | 
ultimately show "a < b" by (simp add: order_le_less)  | 
|
317  | 
qed  | 
|
318  | 
||
319  | 
lemma add_less_imp_less_right:  | 
|
320  | 
"a + c < b + c ==> a < (b::'a::pordered_ab_semigroup_add_imp_le)"  | 
|
321  | 
apply (rule add_less_imp_less_left [of c])  | 
|
322  | 
apply (simp add: add_commute)  | 
|
323  | 
done  | 
|
324  | 
||
325  | 
lemma add_less_cancel_left [simp]:  | 
|
326  | 
"(c+a < c+b) = (a < (b::'a::pordered_ab_semigroup_add_imp_le))"  | 
|
327  | 
by (blast intro: add_less_imp_less_left add_strict_left_mono)  | 
|
328  | 
||
329  | 
lemma add_less_cancel_right [simp]:  | 
|
330  | 
"(a+c < b+c) = (a < (b::'a::pordered_ab_semigroup_add_imp_le))"  | 
|
331  | 
by (blast intro: add_less_imp_less_right add_strict_right_mono)  | 
|
332  | 
||
333  | 
lemma add_le_cancel_left [simp]:  | 
|
334  | 
"(c+a \<le> c+b) = (a \<le> (b::'a::pordered_ab_semigroup_add_imp_le))"  | 
|
335  | 
by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono)  | 
|
336  | 
||
337  | 
lemma add_le_cancel_right [simp]:  | 
|
338  | 
"(a+c \<le> b+c) = (a \<le> (b::'a::pordered_ab_semigroup_add_imp_le))"  | 
|
339  | 
by (simp add: add_commute[of a c] add_commute[of b c])  | 
|
340  | 
||
341  | 
lemma add_le_imp_le_right:  | 
|
342  | 
"a + c \<le> b + c ==> a \<le> (b::'a::pordered_ab_semigroup_add_imp_le)"  | 
|
343  | 
by simp  | 
|
344  | 
||
| 
15234
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
345  | 
lemma add_increasing:  | 
| 
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
346  | 
  fixes c :: "'a::{pordered_ab_semigroup_add_imp_le, comm_monoid_add}"
 | 
| 
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
347  | 
shows "[|0\<le>a; b\<le>c|] ==> b \<le> a + c"  | 
| 14738 | 348  | 
by (insert add_mono [of 0 a b c], simp)  | 
349  | 
||
| 15539 | 350  | 
lemma add_increasing2:  | 
351  | 
  fixes c :: "'a::{pordered_ab_semigroup_add_imp_le, comm_monoid_add}"
 | 
|
352  | 
shows "[|0\<le>c; b\<le>a|] ==> b \<le> a + c"  | 
|
353  | 
by (simp add:add_increasing add_commute[of a])  | 
|
354  | 
||
| 
15234
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
355  | 
lemma add_strict_increasing:  | 
| 
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
356  | 
  fixes c :: "'a::{pordered_ab_semigroup_add_imp_le, comm_monoid_add}"
 | 
| 
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
357  | 
shows "[|0<a; b\<le>c|] ==> b < a + c"  | 
| 
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
358  | 
by (insert add_less_le_mono [of 0 a b c], simp)  | 
| 
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
359  | 
|
| 
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
360  | 
lemma add_strict_increasing2:  | 
| 
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
361  | 
  fixes c :: "'a::{pordered_ab_semigroup_add_imp_le, comm_monoid_add}"
 | 
| 
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
362  | 
shows "[|0\<le>a; b<c|] ==> b < a + c"  | 
| 
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
363  | 
by (insert add_le_less_mono [of 0 a b c], simp)  | 
| 
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
364  | 
|
| 19527 | 365  | 
lemma max_add_distrib_left:  | 
366  | 
fixes z :: "'a::pordered_ab_semigroup_add_imp_le"  | 
|
367  | 
shows "(max x y) + z = max (x+z) (y+z)"  | 
|
368  | 
by (rule max_of_mono [THEN sym], rule add_le_cancel_right)  | 
|
369  | 
||
370  | 
lemma min_add_distrib_left:  | 
|
371  | 
fixes z :: "'a::pordered_ab_semigroup_add_imp_le"  | 
|
372  | 
shows "(min x y) + z = min (x+z) (y+z)"  | 
|
373  | 
by (rule min_of_mono [THEN sym], rule add_le_cancel_right)  | 
|
374  | 
||
375  | 
lemma max_diff_distrib_left:  | 
|
376  | 
fixes z :: "'a::pordered_ab_group_add"  | 
|
377  | 
shows "(max x y) - z = max (x-z) (y-z)"  | 
|
378  | 
by (simp add: diff_minus, rule max_add_distrib_left)  | 
|
379  | 
||
380  | 
lemma min_diff_distrib_left:  | 
|
381  | 
fixes z :: "'a::pordered_ab_group_add"  | 
|
382  | 
shows "(min x y) - z = min (x-z) (y-z)"  | 
|
383  | 
by (simp add: diff_minus, rule min_add_distrib_left)  | 
|
384  | 
||
| 
15234
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
385  | 
|
| 14738 | 386  | 
subsection {* Ordering Rules for Unary Minus *}
 | 
387  | 
||
388  | 
lemma le_imp_neg_le:  | 
|
| 23389 | 389  | 
  assumes "a \<le> (b::'a::{pordered_ab_semigroup_add_imp_le, ab_group_add})" shows "-b \<le> -a"
 | 
| 14738 | 390  | 
proof -  | 
391  | 
have "-a+a \<le> -a+b"  | 
|
| 23389 | 392  | 
using `a \<le> b` by (rule add_left_mono)  | 
| 14738 | 393  | 
hence "0 \<le> -a+b"  | 
394  | 
by simp  | 
|
395  | 
hence "0 + (-b) \<le> (-a + b) + (-b)"  | 
|
396  | 
by (rule add_right_mono)  | 
|
397  | 
thus ?thesis  | 
|
398  | 
by (simp add: add_assoc)  | 
|
399  | 
qed  | 
|
400  | 
||
401  | 
lemma neg_le_iff_le [simp]: "(-b \<le> -a) = (a \<le> (b::'a::pordered_ab_group_add))"  | 
|
402  | 
proof  | 
|
403  | 
assume "- b \<le> - a"  | 
|
404  | 
hence "- (- a) \<le> - (- b)"  | 
|
405  | 
by (rule le_imp_neg_le)  | 
|
406  | 
thus "a\<le>b" by simp  | 
|
407  | 
next  | 
|
408  | 
assume "a\<le>b"  | 
|
409  | 
thus "-b \<le> -a" by (rule le_imp_neg_le)  | 
|
410  | 
qed  | 
|
411  | 
||
412  | 
lemma neg_le_0_iff_le [simp]: "(-a \<le> 0) = (0 \<le> (a::'a::pordered_ab_group_add))"  | 
|
413  | 
by (subst neg_le_iff_le [symmetric], simp)  | 
|
414  | 
||
415  | 
lemma neg_0_le_iff_le [simp]: "(0 \<le> -a) = (a \<le> (0::'a::pordered_ab_group_add))"  | 
|
416  | 
by (subst neg_le_iff_le [symmetric], simp)  | 
|
417  | 
||
418  | 
lemma neg_less_iff_less [simp]: "(-b < -a) = (a < (b::'a::pordered_ab_group_add))"  | 
|
419  | 
by (force simp add: order_less_le)  | 
|
420  | 
||
421  | 
lemma neg_less_0_iff_less [simp]: "(-a < 0) = (0 < (a::'a::pordered_ab_group_add))"  | 
|
422  | 
by (subst neg_less_iff_less [symmetric], simp)  | 
|
423  | 
||
424  | 
lemma neg_0_less_iff_less [simp]: "(0 < -a) = (a < (0::'a::pordered_ab_group_add))"  | 
|
425  | 
by (subst neg_less_iff_less [symmetric], simp)  | 
|
426  | 
||
427  | 
text{*The next several equations can make the simplifier loop!*}
 | 
|
428  | 
||
429  | 
lemma less_minus_iff: "(a < - b) = (b < - (a::'a::pordered_ab_group_add))"  | 
|
430  | 
proof -  | 
|
431  | 
have "(- (-a) < - b) = (b < - a)" by (rule neg_less_iff_less)  | 
|
432  | 
thus ?thesis by simp  | 
|
433  | 
qed  | 
|
434  | 
||
435  | 
lemma minus_less_iff: "(- a < b) = (- b < (a::'a::pordered_ab_group_add))"  | 
|
436  | 
proof -  | 
|
437  | 
have "(- a < - (-b)) = (- b < a)" by (rule neg_less_iff_less)  | 
|
438  | 
thus ?thesis by simp  | 
|
439  | 
qed  | 
|
440  | 
||
441  | 
lemma le_minus_iff: "(a \<le> - b) = (b \<le> - (a::'a::pordered_ab_group_add))"  | 
|
442  | 
proof -  | 
|
443  | 
have mm: "!! a (b::'a). (-(-a)) < -b \<Longrightarrow> -(-b) < -a" by (simp only: minus_less_iff)  | 
|
444  | 
have "(- (- a) <= -b) = (b <= - a)"  | 
|
445  | 
apply (auto simp only: order_le_less)  | 
|
446  | 
apply (drule mm)  | 
|
447  | 
apply (simp_all)  | 
|
448  | 
apply (drule mm[simplified], assumption)  | 
|
449  | 
done  | 
|
450  | 
then show ?thesis by simp  | 
|
451  | 
qed  | 
|
452  | 
||
453  | 
lemma minus_le_iff: "(- a \<le> b) = (- b \<le> (a::'a::pordered_ab_group_add))"  | 
|
454  | 
by (auto simp add: order_le_less minus_less_iff)  | 
|
455  | 
||
456  | 
lemma add_diff_eq: "a + (b - c) = (a + b) - (c::'a::ab_group_add)"  | 
|
457  | 
by (simp add: diff_minus add_ac)  | 
|
458  | 
||
459  | 
lemma diff_add_eq: "(a - b) + c = (a + c) - (b::'a::ab_group_add)"  | 
|
460  | 
by (simp add: diff_minus add_ac)  | 
|
461  | 
||
462  | 
lemma diff_eq_eq: "(a-b = c) = (a = c + (b::'a::ab_group_add))"  | 
|
463  | 
by (auto simp add: diff_minus add_assoc)  | 
|
464  | 
||
465  | 
lemma eq_diff_eq: "(a = c-b) = (a + (b::'a::ab_group_add) = c)"  | 
|
466  | 
by (auto simp add: diff_minus add_assoc)  | 
|
467  | 
||
468  | 
lemma diff_diff_eq: "(a - b) - c = a - (b + (c::'a::ab_group_add))"  | 
|
469  | 
by (simp add: diff_minus add_ac)  | 
|
470  | 
||
471  | 
lemma diff_diff_eq2: "a - (b - c) = (a + c) - (b::'a::ab_group_add)"  | 
|
472  | 
by (simp add: diff_minus add_ac)  | 
|
473  | 
||
474  | 
lemma diff_add_cancel: "a - b + b = (a::'a::ab_group_add)"  | 
|
475  | 
by (simp add: diff_minus add_ac)  | 
|
476  | 
||
477  | 
lemma add_diff_cancel: "a + b - b = (a::'a::ab_group_add)"  | 
|
478  | 
by (simp add: diff_minus add_ac)  | 
|
479  | 
||
| 
14754
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
480  | 
text{*Further subtraction laws*}
 | 
| 14738 | 481  | 
|
482  | 
lemma less_iff_diff_less_0: "(a < b) = (a - b < (0::'a::pordered_ab_group_add))"  | 
|
483  | 
proof -  | 
|
484  | 
have "(a < b) = (a + (- b) < b + (-b))"  | 
|
485  | 
by (simp only: add_less_cancel_right)  | 
|
486  | 
also have "... = (a - b < 0)" by (simp add: diff_minus)  | 
|
487  | 
finally show ?thesis .  | 
|
488  | 
qed  | 
|
489  | 
||
490  | 
lemma diff_less_eq: "(a-b < c) = (a < c + (b::'a::pordered_ab_group_add))"  | 
|
| 15481 | 491  | 
apply (subst less_iff_diff_less_0 [of a])  | 
| 14738 | 492  | 
apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])  | 
493  | 
apply (simp add: diff_minus add_ac)  | 
|
494  | 
done  | 
|
495  | 
||
496  | 
lemma less_diff_eq: "(a < c-b) = (a + (b::'a::pordered_ab_group_add) < c)"  | 
|
| 15481 | 497  | 
apply (subst less_iff_diff_less_0 [of "a+b"])  | 
498  | 
apply (subst less_iff_diff_less_0 [of a])  | 
|
| 14738 | 499  | 
apply (simp add: diff_minus add_ac)  | 
500  | 
done  | 
|
501  | 
||
502  | 
lemma diff_le_eq: "(a-b \<le> c) = (a \<le> c + (b::'a::pordered_ab_group_add))"  | 
|
503  | 
by (auto simp add: order_le_less diff_less_eq diff_add_cancel add_diff_cancel)  | 
|
504  | 
||
505  | 
lemma le_diff_eq: "(a \<le> c-b) = (a + (b::'a::pordered_ab_group_add) \<le> c)"  | 
|
506  | 
by (auto simp add: order_le_less less_diff_eq diff_add_cancel add_diff_cancel)  | 
|
507  | 
||
508  | 
text{*This list of rewrites simplifies (in)equalities by bringing subtractions
 | 
|
509  | 
to the top and then moving negative terms to the other side.  | 
|
510  | 
  Use with @{text add_ac}*}
 | 
|
511  | 
lemmas compare_rls =  | 
|
512  | 
diff_minus [symmetric]  | 
|
513  | 
add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2  | 
|
514  | 
diff_less_eq less_diff_eq diff_le_eq le_diff_eq  | 
|
515  | 
diff_eq_eq eq_diff_eq  | 
|
516  | 
||
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
517  | 
subsection {* Support for reasoning about signs *}
 | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
518  | 
|
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
519  | 
lemma add_pos_pos: "0 <  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
520  | 
    (x::'a::{comm_monoid_add,pordered_cancel_ab_semigroup_add}) 
 | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
521  | 
==> 0 < y ==> 0 < x + y"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
522  | 
apply (subgoal_tac "0 + 0 < x + y")  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
523  | 
apply simp  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
524  | 
apply (erule add_less_le_mono)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
525  | 
apply (erule order_less_imp_le)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
526  | 
done  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
527  | 
|
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
528  | 
lemma add_pos_nonneg: "0 <  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
529  | 
    (x::'a::{comm_monoid_add,pordered_cancel_ab_semigroup_add}) 
 | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
530  | 
==> 0 <= y ==> 0 < x + y"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
531  | 
apply (subgoal_tac "0 + 0 < x + y")  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
532  | 
apply simp  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
533  | 
apply (erule add_less_le_mono, assumption)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
534  | 
done  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
535  | 
|
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
536  | 
lemma add_nonneg_pos: "0 <=  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
537  | 
    (x::'a::{comm_monoid_add,pordered_cancel_ab_semigroup_add}) 
 | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
538  | 
==> 0 < y ==> 0 < x + y"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
539  | 
apply (subgoal_tac "0 + 0 < x + y")  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
540  | 
apply simp  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
541  | 
apply (erule add_le_less_mono, assumption)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
542  | 
done  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
543  | 
|
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
544  | 
lemma add_nonneg_nonneg: "0 <=  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
545  | 
    (x::'a::{comm_monoid_add,pordered_cancel_ab_semigroup_add}) 
 | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
546  | 
==> 0 <= y ==> 0 <= x + y"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
547  | 
apply (subgoal_tac "0 + 0 <= x + y")  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
548  | 
apply simp  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
549  | 
apply (erule add_mono, assumption)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
550  | 
done  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
551  | 
|
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
552  | 
lemma add_neg_neg: "(x::'a::{comm_monoid_add,pordered_cancel_ab_semigroup_add})
 | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
553  | 
< 0 ==> y < 0 ==> x + y < 0"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
554  | 
apply (subgoal_tac "x + y < 0 + 0")  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
555  | 
apply simp  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
556  | 
apply (erule add_less_le_mono)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
557  | 
apply (erule order_less_imp_le)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
558  | 
done  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
559  | 
|
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
560  | 
lemma add_neg_nonpos:  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
561  | 
    "(x::'a::{comm_monoid_add,pordered_cancel_ab_semigroup_add}) < 0 
 | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
562  | 
==> y <= 0 ==> x + y < 0"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
563  | 
apply (subgoal_tac "x + y < 0 + 0")  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
564  | 
apply simp  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
565  | 
apply (erule add_less_le_mono, assumption)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
566  | 
done  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
567  | 
|
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
568  | 
lemma add_nonpos_neg:  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
569  | 
    "(x::'a::{comm_monoid_add,pordered_cancel_ab_semigroup_add}) <= 0 
 | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
570  | 
==> y < 0 ==> x + y < 0"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
571  | 
apply (subgoal_tac "x + y < 0 + 0")  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
572  | 
apply simp  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
573  | 
apply (erule add_le_less_mono, assumption)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
574  | 
done  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
575  | 
|
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
576  | 
lemma add_nonpos_nonpos:  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
577  | 
    "(x::'a::{comm_monoid_add,pordered_cancel_ab_semigroup_add}) <= 0 
 | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
578  | 
==> y <= 0 ==> x + y <= 0"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
579  | 
apply (subgoal_tac "x + y <= 0 + 0")  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
580  | 
apply simp  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
581  | 
apply (erule add_mono, assumption)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
582  | 
done  | 
| 14738 | 583  | 
|
584  | 
subsection{*Lemmas for the @{text cancel_numerals} simproc*}
 | 
|
585  | 
||
586  | 
lemma eq_iff_diff_eq_0: "(a = b) = (a-b = (0::'a::ab_group_add))"  | 
|
587  | 
by (simp add: compare_rls)  | 
|
588  | 
||
589  | 
lemma le_iff_diff_le_0: "(a \<le> b) = (a-b \<le> (0::'a::pordered_ab_group_add))"  | 
|
590  | 
by (simp add: compare_rls)  | 
|
591  | 
||
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
592  | 
|
| 14738 | 593  | 
subsection {* Lattice Ordered (Abelian) Groups *}
 | 
594  | 
||
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
595  | 
class lordered_ab_group_meet = pordered_ab_group_add + lower_semilattice  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
596  | 
|
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
597  | 
class lordered_ab_group_join = pordered_ab_group_add + upper_semilattice  | 
| 14738 | 598  | 
|
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
599  | 
class lordered_ab_group = pordered_ab_group_add + lattice  | 
| 14738 | 600  | 
|
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
601  | 
instance lordered_ab_group \<subseteq> lordered_ab_group_meet by default  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
602  | 
instance lordered_ab_group \<subseteq> lordered_ab_group_join by default  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
603  | 
|
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
604  | 
lemma add_inf_distrib_left: "a + inf b c = inf (a + b) (a + (c::'a::{pordered_ab_group_add, lower_semilattice}))"
 | 
| 14738 | 605  | 
apply (rule order_antisym)  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
606  | 
apply (simp_all add: le_infI)  | 
| 14738 | 607  | 
apply (rule add_le_imp_le_left [of "-a"])  | 
608  | 
apply (simp only: add_assoc[symmetric], simp)  | 
|
| 21312 | 609  | 
apply rule  | 
610  | 
apply (rule add_le_imp_le_left[of "a"], simp only: add_assoc[symmetric], simp)+  | 
|
| 14738 | 611  | 
done  | 
612  | 
||
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
613  | 
lemma add_sup_distrib_left: "a + sup b c = sup (a + b) (a+ (c::'a::{pordered_ab_group_add, upper_semilattice}))" 
 | 
| 14738 | 614  | 
apply (rule order_antisym)  | 
615  | 
apply (rule add_le_imp_le_left [of "-a"])  | 
|
616  | 
apply (simp only: add_assoc[symmetric], simp)  | 
|
| 21312 | 617  | 
apply rule  | 
618  | 
apply (rule add_le_imp_le_left [of "a"], simp only: add_assoc[symmetric], simp)+  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
619  | 
apply (rule le_supI)  | 
| 21312 | 620  | 
apply (simp_all)  | 
| 14738 | 621  | 
done  | 
622  | 
||
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
623  | 
lemma add_inf_distrib_right: "inf a b + (c::'a::lordered_ab_group) = inf (a+c) (b+c)"  | 
| 14738 | 624  | 
proof -  | 
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
625  | 
have "c + inf a b = inf (c+a) (c+b)" by (simp add: add_inf_distrib_left)  | 
| 14738 | 626  | 
thus ?thesis by (simp add: add_commute)  | 
627  | 
qed  | 
|
628  | 
||
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
629  | 
lemma add_sup_distrib_right: "sup a b + (c::'a::lordered_ab_group) = sup (a+c) (b+c)"  | 
| 14738 | 630  | 
proof -  | 
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
631  | 
have "c + sup a b = sup (c+a) (c+b)" by (simp add: add_sup_distrib_left)  | 
| 14738 | 632  | 
thus ?thesis by (simp add: add_commute)  | 
633  | 
qed  | 
|
634  | 
||
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
635  | 
lemmas add_sup_inf_distribs = add_inf_distrib_right add_inf_distrib_left add_sup_distrib_right add_sup_distrib_left  | 
| 14738 | 636  | 
|
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
637  | 
lemma inf_eq_neg_sup: "inf a (b\<Colon>'a\<Colon>lordered_ab_group) = - sup (-a) (-b)"  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
638  | 
proof (rule inf_unique)  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
639  | 
fix a b :: 'a  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
640  | 
show "- sup (-a) (-b) \<le> a" by (rule add_le_imp_le_right [of _ "sup (-a) (-b)"])  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
641  | 
(simp, simp add: add_sup_distrib_left)  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
642  | 
next  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
643  | 
fix a b :: 'a  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
644  | 
show "- sup (-a) (-b) \<le> b" by (rule add_le_imp_le_right [of _ "sup (-a) (-b)"])  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
645  | 
(simp, simp add: add_sup_distrib_left)  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
646  | 
next  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
647  | 
fix a b c :: 'a  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
648  | 
assume "a \<le> b" "a \<le> c"  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
649  | 
then show "a \<le> - sup (-b) (-c)" by (subst neg_le_iff_le [symmetric])  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
650  | 
(simp add: le_supI)  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
651  | 
qed  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
652  | 
|
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
653  | 
lemma sup_eq_neg_inf: "sup a (b\<Colon>'a\<Colon>lordered_ab_group) = - inf (-a) (-b)"  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
654  | 
proof (rule sup_unique)  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
655  | 
fix a b :: 'a  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
656  | 
show "a \<le> - inf (-a) (-b)" by (rule add_le_imp_le_right [of _ "inf (-a) (-b)"])  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
657  | 
(simp, simp add: add_inf_distrib_left)  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
658  | 
next  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
659  | 
fix a b :: 'a  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
660  | 
show "b \<le> - inf (-a) (-b)" by (rule add_le_imp_le_right [of _ "inf (-a) (-b)"])  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
661  | 
(simp, simp add: add_inf_distrib_left)  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
662  | 
next  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
663  | 
fix a b c :: 'a  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
664  | 
assume "a \<le> c" "b \<le> c"  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
665  | 
then show "- inf (-a) (-b) \<le> c" by (subst neg_le_iff_le [symmetric])  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
666  | 
(simp add: le_infI)  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
667  | 
qed  | 
| 14738 | 668  | 
|
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
669  | 
lemma add_eq_inf_sup: "a + b = sup a b + inf a (b\<Colon>'a\<Colon>lordered_ab_group)"  | 
| 14738 | 670  | 
proof -  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
671  | 
have "0 = - inf 0 (a-b) + inf (a-b) 0" by (simp add: inf_commute)  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
672  | 
hence "0 = sup 0 (b-a) + inf (a-b) 0" by (simp add: inf_eq_neg_sup)  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
673  | 
hence "0 = (-a + sup a b) + (inf a b + (-b))"  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
674  | 
apply (simp add: add_sup_distrib_left add_inf_distrib_right)  | 
| 14738 | 675  | 
by (simp add: diff_minus add_commute)  | 
676  | 
thus ?thesis  | 
|
677  | 
apply (simp add: compare_rls)  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
678  | 
apply (subst add_left_cancel[symmetric, of "a+b" "sup a b + inf a b" "-a"])  | 
| 14738 | 679  | 
apply (simp only: add_assoc, simp add: add_assoc[symmetric])  | 
680  | 
done  | 
|
681  | 
qed  | 
|
682  | 
||
683  | 
subsection {* Positive Part, Negative Part, Absolute Value *}
 | 
|
684  | 
||
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
685  | 
definition  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
686  | 
  nprt :: "'a \<Rightarrow> ('a::lordered_ab_group)" where
 | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
687  | 
"nprt x = inf x 0"  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
688  | 
|
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
689  | 
definition  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
690  | 
  pprt :: "'a \<Rightarrow> ('a::lordered_ab_group)" where
 | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
691  | 
"pprt x = sup x 0"  | 
| 14738 | 692  | 
|
693  | 
lemma prts: "a = pprt a + nprt a"  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
694  | 
by (simp add: pprt_def nprt_def add_eq_inf_sup[symmetric])  | 
| 14738 | 695  | 
|
696  | 
lemma zero_le_pprt[simp]: "0 \<le> pprt a"  | 
|
| 21312 | 697  | 
by (simp add: pprt_def)  | 
| 14738 | 698  | 
|
699  | 
lemma nprt_le_zero[simp]: "nprt a \<le> 0"  | 
|
| 21312 | 700  | 
by (simp add: nprt_def)  | 
| 14738 | 701  | 
|
702  | 
lemma le_eq_neg: "(a \<le> -b) = (a + b \<le> (0::_::lordered_ab_group))" (is "?l = ?r")  | 
|
703  | 
proof -  | 
|
704  | 
have a: "?l \<longrightarrow> ?r"  | 
|
705  | 
apply (auto)  | 
|
706  | 
apply (rule add_le_imp_le_right[of _ "-b" _])  | 
|
707  | 
apply (simp add: add_assoc)  | 
|
708  | 
done  | 
|
709  | 
have b: "?r \<longrightarrow> ?l"  | 
|
710  | 
apply (auto)  | 
|
711  | 
apply (rule add_le_imp_le_right[of _ "b" _])  | 
|
712  | 
apply (simp)  | 
|
713  | 
done  | 
|
714  | 
from a b show ?thesis by blast  | 
|
715  | 
qed  | 
|
716  | 
||
| 15580 | 717  | 
lemma pprt_0[simp]: "pprt 0 = 0" by (simp add: pprt_def)  | 
718  | 
lemma nprt_0[simp]: "nprt 0 = 0" by (simp add: nprt_def)  | 
|
719  | 
||
| 
24286
 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 
paulson 
parents: 
24137 
diff
changeset
 | 
720  | 
lemma pprt_eq_id[simp,noatp]: "0 <= x \<Longrightarrow> pprt x = x"  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
721  | 
by (simp add: pprt_def le_iff_sup sup_aci)  | 
| 15580 | 722  | 
|
| 
24286
 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 
paulson 
parents: 
24137 
diff
changeset
 | 
723  | 
lemma nprt_eq_id[simp,noatp]: "x <= 0 \<Longrightarrow> nprt x = x"  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
724  | 
by (simp add: nprt_def le_iff_inf inf_aci)  | 
| 15580 | 725  | 
|
| 
24286
 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 
paulson 
parents: 
24137 
diff
changeset
 | 
726  | 
lemma pprt_eq_0[simp,noatp]: "x <= 0 \<Longrightarrow> pprt x = 0"  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
727  | 
by (simp add: pprt_def le_iff_sup sup_aci)  | 
| 15580 | 728  | 
|
| 
24286
 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 
paulson 
parents: 
24137 
diff
changeset
 | 
729  | 
lemma nprt_eq_0[simp,noatp]: "0 <= x \<Longrightarrow> nprt x = 0"  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
730  | 
by (simp add: nprt_def le_iff_inf inf_aci)  | 
| 15580 | 731  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
732  | 
lemma sup_0_imp_0: "sup a (-a) = 0 \<Longrightarrow> a = (0::'a::lordered_ab_group)"  | 
| 14738 | 733  | 
proof -  | 
734  | 
  {
 | 
|
735  | 
fix a::'a  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
736  | 
assume hyp: "sup a (-a) = 0"  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
737  | 
hence "sup a (-a) + a = a" by (simp)  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
738  | 
hence "sup (a+a) 0 = a" by (simp add: add_sup_distrib_right)  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
739  | 
hence "sup (a+a) 0 <= a" by (simp)  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
740  | 
hence "0 <= a" by (blast intro: order_trans inf_sup_ord)  | 
| 14738 | 741  | 
}  | 
742  | 
note p = this  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
743  | 
assume hyp:"sup a (-a) = 0"  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
744  | 
hence hyp2:"sup (-a) (-(-a)) = 0" by (simp add: sup_commute)  | 
| 14738 | 745  | 
from p[OF hyp] p[OF hyp2] show "a = 0" by simp  | 
746  | 
qed  | 
|
747  | 
||
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
748  | 
lemma inf_0_imp_0: "inf a (-a) = 0 \<Longrightarrow> a = (0::'a::lordered_ab_group)"  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
749  | 
apply (simp add: inf_eq_neg_sup)  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
750  | 
apply (simp add: sup_commute)  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
751  | 
apply (erule sup_0_imp_0)  | 
| 15481 | 752  | 
done  | 
| 14738 | 753  | 
|
| 
24286
 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 
paulson 
parents: 
24137 
diff
changeset
 | 
754  | 
lemma inf_0_eq_0[simp,noatp]: "(inf a (-a) = 0) = (a = (0::'a::lordered_ab_group))"  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
755  | 
by (auto, erule inf_0_imp_0)  | 
| 14738 | 756  | 
|
| 
24286
 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 
paulson 
parents: 
24137 
diff
changeset
 | 
757  | 
lemma sup_0_eq_0[simp,noatp]: "(sup a (-a) = 0) = (a = (0::'a::lordered_ab_group))"  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
758  | 
by (auto, erule sup_0_imp_0)  | 
| 14738 | 759  | 
|
760  | 
lemma zero_le_double_add_iff_zero_le_single_add[simp]: "(0 \<le> a + a) = (0 \<le> (a::'a::lordered_ab_group))"  | 
|
761  | 
proof  | 
|
762  | 
assume "0 <= a + a"  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
763  | 
hence a:"inf (a+a) 0 = 0" by (simp add: le_iff_inf inf_commute)  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
764  | 
have "(inf a 0)+(inf a 0) = inf (inf (a+a) 0) a" (is "?l=_") by (simp add: add_sup_inf_distribs inf_aci)  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
765  | 
hence "?l = 0 + inf a 0" by (simp add: a, simp add: inf_commute)  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
766  | 
hence "inf a 0 = 0" by (simp only: add_right_cancel)  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
767  | 
then show "0 <= a" by (simp add: le_iff_inf inf_commute)  | 
| 14738 | 768  | 
next  | 
769  | 
assume a: "0 <= a"  | 
|
770  | 
show "0 <= a + a" by (simp add: add_mono[OF a a, simplified])  | 
|
771  | 
qed  | 
|
772  | 
||
773  | 
lemma double_add_le_zero_iff_single_add_le_zero[simp]: "(a + a <= 0) = ((a::'a::lordered_ab_group) <= 0)"  | 
|
774  | 
proof -  | 
|
775  | 
have "(a + a <= 0) = (0 <= -(a+a))" by (subst le_minus_iff, simp)  | 
|
776  | 
moreover have "\<dots> = (a <= 0)" by (simp add: zero_le_double_add_iff_zero_le_single_add)  | 
|
777  | 
ultimately show ?thesis by blast  | 
|
778  | 
qed  | 
|
779  | 
||
780  | 
lemma double_add_less_zero_iff_single_less_zero[simp]: "(a+a<0) = ((a::'a::{pordered_ab_group_add,linorder}) < 0)" (is ?s)
 | 
|
781  | 
proof cases  | 
|
782  | 
assume a: "a < 0"  | 
|
783  | 
thus ?s by (simp add: add_strict_mono[OF a a, simplified])  | 
|
784  | 
next  | 
|
785  | 
assume "~(a < 0)"  | 
|
786  | 
hence a:"0 <= a" by (simp)  | 
|
787  | 
hence "0 <= a+a" by (simp add: add_mono[OF a a, simplified])  | 
|
788  | 
hence "~(a+a < 0)" by simp  | 
|
789  | 
with a show ?thesis by simp  | 
|
790  | 
qed  | 
|
791  | 
||
| 23879 | 792  | 
class lordered_ab_group_abs = lordered_ab_group + abs +  | 
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
793  | 
assumes abs_lattice: "abs x = sup x (uminus x)"  | 
| 14738 | 794  | 
|
795  | 
lemma abs_zero[simp]: "abs 0 = (0::'a::lordered_ab_group_abs)"  | 
|
796  | 
by (simp add: abs_lattice)  | 
|
797  | 
||
798  | 
lemma abs_eq_0[simp]: "(abs a = 0) = (a = (0::'a::lordered_ab_group_abs))"  | 
|
799  | 
by (simp add: abs_lattice)  | 
|
800  | 
||
801  | 
lemma abs_0_eq[simp]: "(0 = abs a) = (a = (0::'a::lordered_ab_group_abs))"  | 
|
802  | 
proof -  | 
|
803  | 
have "(0 = abs a) = (abs a = 0)" by (simp only: eq_ac)  | 
|
804  | 
thus ?thesis by simp  | 
|
805  | 
qed  | 
|
806  | 
||
| 
24286
 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 
paulson 
parents: 
24137 
diff
changeset
 | 
807  | 
declare abs_0_eq [noatp] (*essentially the same as the other one*)  | 
| 
 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 
paulson 
parents: 
24137 
diff
changeset
 | 
808  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
809  | 
lemma neg_inf_eq_sup[simp]: "- inf a (b::_::lordered_ab_group) = sup (-a) (-b)"  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
810  | 
by (simp add: inf_eq_neg_sup)  | 
| 14738 | 811  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
812  | 
lemma neg_sup_eq_inf[simp]: "- sup a (b::_::lordered_ab_group) = inf (-a) (-b)"  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
813  | 
by (simp del: neg_inf_eq_sup add: sup_eq_neg_inf)  | 
| 14738 | 814  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
815  | 
lemma sup_eq_if: "sup a (-a) = (if a < 0 then -a else (a::'a::{lordered_ab_group, linorder}))"
 | 
| 14738 | 816  | 
proof -  | 
817  | 
note b = add_le_cancel_right[of a a "-a",symmetric,simplified]  | 
|
818  | 
have c: "a + a = 0 \<Longrightarrow> -a = a" by (rule add_right_imp_eq[of _ a], simp)  | 
|
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
819  | 
show ?thesis by (auto simp add: max_def b linorder_not_less sup_max)  | 
| 14738 | 820  | 
qed  | 
821  | 
||
822  | 
lemma abs_if_lattice: "\<bar>a\<bar> = (if a < 0 then -a else (a::'a::{lordered_ab_group_abs, linorder}))"
 | 
|
823  | 
proof -  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
824  | 
show ?thesis by (simp add: abs_lattice sup_eq_if)  | 
| 14738 | 825  | 
qed  | 
826  | 
||
827  | 
lemma abs_ge_zero[simp]: "0 \<le> abs (a::'a::lordered_ab_group_abs)"  | 
|
828  | 
proof -  | 
|
| 21312 | 829  | 
have a:"a <= abs a" and b:"-a <= abs a" by (auto simp add: abs_lattice)  | 
| 14738 | 830  | 
show ?thesis by (rule add_mono[OF a b, simplified])  | 
831  | 
qed  | 
|
832  | 
||
833  | 
lemma abs_le_zero_iff [simp]: "(abs a \<le> (0::'a::lordered_ab_group_abs)) = (a = 0)"  | 
|
834  | 
proof  | 
|
835  | 
assume "abs a <= 0"  | 
|
836  | 
hence "abs a = 0" by (auto dest: order_antisym)  | 
|
837  | 
thus "a = 0" by simp  | 
|
838  | 
next  | 
|
839  | 
assume "a = 0"  | 
|
840  | 
thus "abs a <= 0" by simp  | 
|
841  | 
qed  | 
|
842  | 
||
843  | 
lemma zero_less_abs_iff [simp]: "(0 < abs a) = (a \<noteq> (0::'a::lordered_ab_group_abs))"  | 
|
844  | 
by (simp add: order_less_le)  | 
|
845  | 
||
846  | 
lemma abs_not_less_zero [simp]: "~ abs a < (0::'a::lordered_ab_group_abs)"  | 
|
847  | 
proof -  | 
|
848  | 
have a:"!! x (y::_::order). x <= y \<Longrightarrow> ~(y < x)" by auto  | 
|
849  | 
show ?thesis by (simp add: a)  | 
|
850  | 
qed  | 
|
851  | 
||
852  | 
lemma abs_ge_self: "a \<le> abs (a::'a::lordered_ab_group_abs)"  | 
|
| 21312 | 853  | 
by (simp add: abs_lattice)  | 
| 14738 | 854  | 
|
855  | 
lemma abs_ge_minus_self: "-a \<le> abs (a::'a::lordered_ab_group_abs)"  | 
|
| 21312 | 856  | 
by (simp add: abs_lattice)  | 
| 14738 | 857  | 
|
858  | 
lemma abs_prts: "abs (a::_::lordered_ab_group_abs) = pprt a - nprt a"  | 
|
859  | 
apply (simp add: pprt_def nprt_def diff_minus)  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
860  | 
apply (simp add: add_sup_inf_distribs sup_aci abs_lattice[symmetric])  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
861  | 
apply (subst sup_absorb2, auto)  | 
| 14738 | 862  | 
done  | 
863  | 
||
864  | 
lemma abs_minus_cancel [simp]: "abs (-a) = abs(a::'a::lordered_ab_group_abs)"  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
865  | 
by (simp add: abs_lattice sup_commute)  | 
| 14738 | 866  | 
|
867  | 
lemma abs_idempotent [simp]: "abs (abs a) = abs (a::'a::lordered_ab_group_abs)"  | 
|
868  | 
apply (simp add: abs_lattice[of "abs a"])  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
869  | 
apply (subst sup_absorb1)  | 
| 14738 | 870  | 
apply (rule order_trans[of _ 0])  | 
871  | 
by auto  | 
|
872  | 
||
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
15010 
diff
changeset
 | 
873  | 
lemma abs_minus_commute:  | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
15010 
diff
changeset
 | 
874  | 
fixes a :: "'a::lordered_ab_group_abs"  | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
15010 
diff
changeset
 | 
875  | 
shows "abs (a-b) = abs(b-a)"  | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
15010 
diff
changeset
 | 
876  | 
proof -  | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
15010 
diff
changeset
 | 
877  | 
have "abs (a-b) = abs (- (a-b))" by (simp only: abs_minus_cancel)  | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
15010 
diff
changeset
 | 
878  | 
also have "... = abs(b-a)" by simp  | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
15010 
diff
changeset
 | 
879  | 
finally show ?thesis .  | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
15010 
diff
changeset
 | 
880  | 
qed  | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
15010 
diff
changeset
 | 
881  | 
|
| 14738 | 882  | 
lemma zero_le_iff_zero_nprt: "(0 \<le> a) = (nprt a = 0)"  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
883  | 
by (simp add: le_iff_inf nprt_def inf_commute)  | 
| 14738 | 884  | 
|
885  | 
lemma le_zero_iff_zero_pprt: "(a \<le> 0) = (pprt a = 0)"  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
886  | 
by (simp add: le_iff_sup pprt_def sup_commute)  | 
| 14738 | 887  | 
|
888  | 
lemma le_zero_iff_pprt_id: "(0 \<le> a) = (pprt a = a)"  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
889  | 
by (simp add: le_iff_sup pprt_def sup_commute)  | 
| 14738 | 890  | 
|
891  | 
lemma zero_le_iff_nprt_id: "(a \<le> 0) = (nprt a = a)"  | 
|
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
892  | 
by (simp add: le_iff_inf nprt_def inf_commute)  | 
| 14738 | 893  | 
|
| 
24286
 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 
paulson 
parents: 
24137 
diff
changeset
 | 
894  | 
lemma pprt_mono[simp,noatp]: "(a::_::lordered_ab_group) <= b \<Longrightarrow> pprt a <= pprt b"  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
895  | 
by (simp add: le_iff_sup pprt_def sup_aci)  | 
| 15580 | 896  | 
|
| 
24286
 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 
paulson 
parents: 
24137 
diff
changeset
 | 
897  | 
lemma nprt_mono[simp,noatp]: "(a::_::lordered_ab_group) <= b \<Longrightarrow> nprt a <= nprt b"  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
898  | 
by (simp add: le_iff_inf nprt_def inf_aci)  | 
| 15580 | 899  | 
|
| 19404 | 900  | 
lemma pprt_neg: "pprt (-x) = - nprt x"  | 
901  | 
by (simp add: pprt_def nprt_def)  | 
|
902  | 
||
903  | 
lemma nprt_neg: "nprt (-x) = - pprt x"  | 
|
904  | 
by (simp add: pprt_def nprt_def)  | 
|
905  | 
||
| 14738 | 906  | 
lemma iff2imp: "(A=B) \<Longrightarrow> (A \<Longrightarrow> B)"  | 
907  | 
by (simp)  | 
|
908  | 
||
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
909  | 
lemma abs_of_nonneg [simp]: "0 \<le> a \<Longrightarrow> abs a = (a::'a::lordered_ab_group_abs)"  | 
| 14738 | 910  | 
by (simp add: iff2imp[OF zero_le_iff_zero_nprt] iff2imp[OF le_zero_iff_pprt_id] abs_prts)  | 
911  | 
||
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
912  | 
lemma abs_of_pos: "0 < (x::'a::lordered_ab_group_abs) ==> abs x = x";  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
913  | 
by (rule abs_of_nonneg, rule order_less_imp_le);  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
914  | 
|
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
915  | 
lemma abs_of_nonpos [simp]: "a \<le> 0 \<Longrightarrow> abs a = -(a::'a::lordered_ab_group_abs)"  | 
| 14738 | 916  | 
by (simp add: iff2imp[OF le_zero_iff_zero_pprt] iff2imp[OF zero_le_iff_nprt_id] abs_prts)  | 
917  | 
||
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
918  | 
lemma abs_of_neg: "(x::'a::lordered_ab_group_abs) < 0 ==>  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
919  | 
abs x = - x"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
920  | 
by (rule abs_of_nonpos, rule order_less_imp_le)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
921  | 
|
| 14738 | 922  | 
lemma abs_leI: "[|a \<le> b; -a \<le> b|] ==> abs a \<le> (b::'a::lordered_ab_group_abs)"  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
923  | 
by (simp add: abs_lattice le_supI)  | 
| 14738 | 924  | 
|
925  | 
lemma le_minus_self_iff: "(a \<le> -a) = (a \<le> (0::'a::lordered_ab_group))"  | 
|
926  | 
proof -  | 
|
927  | 
from add_le_cancel_left[of "-a" "a+a" "0"] have "(a <= -a) = (a+a <= 0)"  | 
|
928  | 
by (simp add: add_assoc[symmetric])  | 
|
929  | 
thus ?thesis by simp  | 
|
930  | 
qed  | 
|
931  | 
||
932  | 
lemma minus_le_self_iff: "(-a \<le> a) = (0 \<le> (a::'a::lordered_ab_group))"  | 
|
933  | 
proof -  | 
|
934  | 
from add_le_cancel_left[of "-a" "0" "a+a"] have "(-a <= a) = (0 <= a+a)"  | 
|
935  | 
by (simp add: add_assoc[symmetric])  | 
|
936  | 
thus ?thesis by simp  | 
|
937  | 
qed  | 
|
938  | 
||
939  | 
lemma abs_le_D1: "abs a \<le> b ==> a \<le> (b::'a::lordered_ab_group_abs)"  | 
|
940  | 
by (insert abs_ge_self, blast intro: order_trans)  | 
|
941  | 
||
942  | 
lemma abs_le_D2: "abs a \<le> b ==> -a \<le> (b::'a::lordered_ab_group_abs)"  | 
|
943  | 
by (insert abs_le_D1 [of "-a"], simp)  | 
|
944  | 
||
945  | 
lemma abs_le_iff: "(abs a \<le> b) = (a \<le> b & -a \<le> (b::'a::lordered_ab_group_abs))"  | 
|
946  | 
by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2)  | 
|
947  | 
||
| 15539 | 948  | 
lemma abs_triangle_ineq: "abs(a+b) \<le> abs a + abs(b::'a::lordered_ab_group_abs)"  | 
| 14738 | 949  | 
proof -  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
950  | 
have g:"abs a + abs b = sup (a+b) (sup (-a-b) (sup (-a+b) (a + (-b))))" (is "_=sup ?m ?n")  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
951  | 
by (simp add: abs_lattice add_sup_inf_distribs sup_aci diff_minus)  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
952  | 
have a:"a+b <= sup ?m ?n" by (simp)  | 
| 21312 | 953  | 
have b:"-a-b <= ?n" by (simp)  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
954  | 
have c:"?n <= sup ?m ?n" by (simp)  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
955  | 
from b c have d: "-a-b <= sup ?m ?n" by(rule order_trans)  | 
| 14738 | 956  | 
have e:"-a-b = -(a+b)" by (simp add: diff_minus)  | 
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
22390 
diff
changeset
 | 
957  | 
from a d e have "abs(a+b) <= sup ?m ?n"  | 
| 14738 | 958  | 
by (drule_tac abs_leI, auto)  | 
959  | 
with g[symmetric] show ?thesis by simp  | 
|
960  | 
qed  | 
|
961  | 
||
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
962  | 
lemma abs_triangle_ineq2: "abs (a::'a::lordered_ab_group_abs) -  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
963  | 
abs b <= abs (a - b)"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
964  | 
apply (simp add: compare_rls)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
965  | 
apply (subgoal_tac "abs a = abs (a - b + b)")  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
966  | 
apply (erule ssubst)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
967  | 
apply (rule abs_triangle_ineq)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
968  | 
apply (rule arg_cong);back;  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
969  | 
apply (simp add: compare_rls)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
970  | 
done  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
971  | 
|
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
972  | 
lemma abs_triangle_ineq3:  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
973  | 
"abs(abs (a::'a::lordered_ab_group_abs) - abs b) <= abs (a - b)"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
974  | 
apply (subst abs_le_iff)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
975  | 
apply auto  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
976  | 
apply (rule abs_triangle_ineq2)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
977  | 
apply (subst abs_minus_commute)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
978  | 
apply (rule abs_triangle_ineq2)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
979  | 
done  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
980  | 
|
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
981  | 
lemma abs_triangle_ineq4: "abs ((a::'a::lordered_ab_group_abs) - b) <=  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
982  | 
abs a + abs b"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
983  | 
proof -;  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
984  | 
have "abs(a - b) = abs(a + - b)"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
985  | 
by (subst diff_minus, rule refl)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
986  | 
also have "... <= abs a + abs (- b)"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
987  | 
by (rule abs_triangle_ineq)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
988  | 
finally show ?thesis  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
989  | 
by simp  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
990  | 
qed  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
991  | 
|
| 14738 | 992  | 
lemma abs_diff_triangle_ineq:  | 
993  | 
"\<bar>(a::'a::lordered_ab_group_abs) + b - (c+d)\<bar> \<le> \<bar>a-c\<bar> + \<bar>b-d\<bar>"  | 
|
994  | 
proof -  | 
|
995  | 
have "\<bar>a + b - (c+d)\<bar> = \<bar>(a-c) + (b-d)\<bar>" by (simp add: diff_minus add_ac)  | 
|
996  | 
also have "... \<le> \<bar>a-c\<bar> + \<bar>b-d\<bar>" by (rule abs_triangle_ineq)  | 
|
997  | 
finally show ?thesis .  | 
|
998  | 
qed  | 
|
999  | 
||
| 15539 | 1000  | 
lemma abs_add_abs[simp]:  | 
1001  | 
fixes a:: "'a::{lordered_ab_group_abs}"
 | 
|
1002  | 
shows "abs(abs a + abs b) = abs a + abs b" (is "?L = ?R")  | 
|
1003  | 
proof (rule order_antisym)  | 
|
1004  | 
show "?L \<ge> ?R" by(rule abs_ge_self)  | 
|
1005  | 
next  | 
|
1006  | 
have "?L \<le> \<bar>\<bar>a\<bar>\<bar> + \<bar>\<bar>b\<bar>\<bar>" by(rule abs_triangle_ineq)  | 
|
1007  | 
also have "\<dots> = ?R" by simp  | 
|
1008  | 
finally show "?L \<le> ?R" .  | 
|
1009  | 
qed  | 
|
1010  | 
||
| 
14754
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1011  | 
text {* Needed for abelian cancellation simprocs: *}
 | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1012  | 
|
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1013  | 
lemma add_cancel_21: "((x::'a::ab_group_add) + (y + z) = y + u) = (x + z = u)"  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1014  | 
apply (subst add_left_commute)  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1015  | 
apply (subst add_left_cancel)  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1016  | 
apply simp  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1017  | 
done  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1018  | 
|
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1019  | 
lemma add_cancel_end: "(x + (y + z) = y) = (x = - (z::'a::ab_group_add))"  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1020  | 
apply (subst add_cancel_21[of _ _ _ 0, simplified])  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1021  | 
apply (simp add: add_right_cancel[symmetric, of "x" "-z" "z", simplified])  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1022  | 
done  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1023  | 
|
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1024  | 
lemma less_eqI: "(x::'a::pordered_ab_group_add) - y = x' - y' \<Longrightarrow> (x < y) = (x' < y')"  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1025  | 
by (simp add: less_iff_diff_less_0[of x y] less_iff_diff_less_0[of x' y'])  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1026  | 
|
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1027  | 
lemma le_eqI: "(x::'a::pordered_ab_group_add) - y = x' - y' \<Longrightarrow> (y <= x) = (y' <= x')"  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1028  | 
apply (simp add: le_iff_diff_le_0[of y x] le_iff_diff_le_0[of y' x'])  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1029  | 
apply (simp add: neg_le_iff_le[symmetric, of "y-x" 0] neg_le_iff_le[symmetric, of "y'-x'" 0])  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1030  | 
done  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1031  | 
|
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1032  | 
lemma eq_eqI: "(x::'a::ab_group_add) - y = x' - y' \<Longrightarrow> (x = y) = (x' = y')"  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1033  | 
by (simp add: eq_iff_diff_eq_0[of x y] eq_iff_diff_eq_0[of x' y'])  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1034  | 
|
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1035  | 
lemma diff_def: "(x::'a::ab_group_add) - y == x + (-y)"  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1036  | 
by (simp add: diff_minus)  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1037  | 
|
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1038  | 
lemma add_minus_cancel: "(a::'a::ab_group_add) + (-a + b) = b"  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1039  | 
by (simp add: add_assoc[symmetric])  | 
| 
 
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
 
obua 
parents: 
14738 
diff
changeset
 | 
1040  | 
|
| 15178 | 1041  | 
lemma le_add_right_mono:  | 
1042  | 
assumes  | 
|
1043  | 
"a <= b + (c::'a::pordered_ab_group_add)"  | 
|
1044  | 
"c <= d"  | 
|
1045  | 
shows "a <= b + d"  | 
|
1046  | 
apply (rule_tac order_trans[where y = "b+c"])  | 
|
1047  | 
apply (simp_all add: prems)  | 
|
1048  | 
done  | 
|
1049  | 
||
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23389 
diff
changeset
 | 
1050  | 
lemmas group_simps =  | 
| 15178 | 1051  | 
mult_ac  | 
1052  | 
add_ac  | 
|
1053  | 
add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2  | 
|
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23389 
diff
changeset
 | 
1054  | 
diff_eq_eq eq_diff_eq diff_minus[symmetric] uminus_add_conv_diff  | 
| 
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23389 
diff
changeset
 | 
1055  | 
diff_less_eq less_diff_eq diff_le_eq le_diff_eq  | 
| 15178 | 1056  | 
|
1057  | 
lemma estimate_by_abs:  | 
|
| 
24380
 
c215e256beca
moved ordered_ab_semigroup_add to OrderedGroup.thy
 
haftmann 
parents: 
24286 
diff
changeset
 | 
1058  | 
"a + b <= (c::'a::lordered_ab_group_abs) \<Longrightarrow> a <= c + abs b"  | 
| 15178 | 1059  | 
proof -  | 
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23389 
diff
changeset
 | 
1060  | 
assume "a+b <= c"  | 
| 
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23389 
diff
changeset
 | 
1061  | 
hence 2: "a <= c+(-b)" by (simp add: group_simps)  | 
| 15178 | 1062  | 
have 3: "(-b) <= abs b" by (rule abs_ge_minus_self)  | 
1063  | 
show ?thesis by (rule le_add_right_mono[OF 2 3])  | 
|
1064  | 
qed  | 
|
1065  | 
||
| 22482 | 1066  | 
|
1067  | 
subsection {* Tools setup *}
 | 
|
1068  | 
||
| 17085 | 1069  | 
text{*Simplification of @{term "x-y < 0"}, etc.*}
 | 
| 
24380
 
c215e256beca
moved ordered_ab_semigroup_add to OrderedGroup.thy
 
haftmann 
parents: 
24286 
diff
changeset
 | 
1070  | 
lemmas diff_less_0_iff_less [simp] = less_iff_diff_less_0 [symmetric]  | 
| 
 
c215e256beca
moved ordered_ab_semigroup_add to OrderedGroup.thy
 
haftmann 
parents: 
24286 
diff
changeset
 | 
1071  | 
lemmas diff_eq_0_iff_eq [simp, noatp] = eq_iff_diff_eq_0 [symmetric]  | 
| 
 
c215e256beca
moved ordered_ab_semigroup_add to OrderedGroup.thy
 
haftmann 
parents: 
24286 
diff
changeset
 | 
1072  | 
lemmas diff_le_0_iff_le [simp] = le_iff_diff_le_0 [symmetric]  | 
| 17085 | 1073  | 
|
| 22482 | 1074  | 
ML {*
 | 
1075  | 
structure ab_group_add_cancel = Abel_Cancel(  | 
|
1076  | 
struct  | 
|
1077  | 
||
1078  | 
(* term order for abelian groups *)  | 
|
1079  | 
||
1080  | 
fun agrp_ord (Const (a, _)) = find_index (fn a' => a = a')  | 
|
| 22997 | 1081  | 
      [@{const_name HOL.zero}, @{const_name HOL.plus},
 | 
1082  | 
        @{const_name HOL.uminus}, @{const_name HOL.minus}]
 | 
|
| 22482 | 1083  | 
| agrp_ord _ = ~1;  | 
1084  | 
||
1085  | 
fun termless_agrp (a, b) = (Term.term_lpo agrp_ord (a, b) = LESS);  | 
|
1086  | 
||
1087  | 
local  | 
|
1088  | 
  val ac1 = mk_meta_eq @{thm add_assoc};
 | 
|
1089  | 
  val ac2 = mk_meta_eq @{thm add_commute};
 | 
|
1090  | 
  val ac3 = mk_meta_eq @{thm add_left_commute};
 | 
|
| 22997 | 1091  | 
  fun solve_add_ac thy _ (_ $ (Const (@{const_name HOL.plus},_) $ _ $ _) $ _) =
 | 
| 22482 | 1092  | 
SOME ac1  | 
| 22997 | 1093  | 
    | solve_add_ac thy _ (_ $ x $ (Const (@{const_name HOL.plus},_) $ y $ z)) =
 | 
| 22482 | 1094  | 
if termless_agrp (y, x) then SOME ac3 else NONE  | 
1095  | 
| solve_add_ac thy _ (_ $ x $ y) =  | 
|
1096  | 
if termless_agrp (y, x) then SOME ac2 else NONE  | 
|
1097  | 
| solve_add_ac thy _ _ = NONE  | 
|
1098  | 
in  | 
|
1099  | 
  val add_ac_proc = Simplifier.simproc @{theory}
 | 
|
1100  | 
"add_ac_proc" ["x + y::'a::ab_semigroup_add"] solve_add_ac;  | 
|
1101  | 
end;  | 
|
1102  | 
||
1103  | 
val cancel_ss = HOL_basic_ss settermless termless_agrp  | 
|
1104  | 
addsimprocs [add_ac_proc] addsimps  | 
|
| 23085 | 1105  | 
  [@{thm add_0_left}, @{thm add_0_right}, @{thm diff_def},
 | 
| 22482 | 1106  | 
   @{thm minus_add_distrib}, @{thm minus_minus}, @{thm minus_zero},
 | 
1107  | 
   @{thm right_minus}, @{thm left_minus}, @{thm add_minus_cancel},
 | 
|
1108  | 
   @{thm minus_add_cancel}];
 | 
|
1109  | 
||
| 22548 | 1110  | 
val eq_reflection = @{thm eq_reflection};
 | 
| 22482 | 1111  | 
|
| 
24137
 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23879 
diff
changeset
 | 
1112  | 
val thy_ref = Theory.check_thy @{theory};
 | 
| 22482 | 1113  | 
|
| 22548 | 1114  | 
val T = TFree("'a", ["OrderedGroup.ab_group_add"]);
 | 
| 22482 | 1115  | 
|
| 22548 | 1116  | 
val eqI_rules = [@{thm less_eqI}, @{thm le_eqI}, @{thm eq_eqI}];
 | 
| 22482 | 1117  | 
|
1118  | 
val dest_eqI =  | 
|
1119  | 
fst o HOLogic.dest_bin "op =" HOLogic.boolT o HOLogic.dest_Trueprop o concl_of;  | 
|
1120  | 
||
1121  | 
end);  | 
|
1122  | 
*}  | 
|
1123  | 
||
1124  | 
ML_setup {*
 | 
|
1125  | 
Addsimprocs [ab_group_add_cancel.sum_conv, ab_group_add_cancel.rel_conv];  | 
|
1126  | 
*}  | 
|
| 17085 | 1127  | 
|
| 14738 | 1128  | 
end  |