| author | blanchet | 
| Thu, 31 Jan 2013 17:54:05 +0100 | |
| changeset 51003 | 198cb05fb35b | 
| parent 49690 | a6814de45b69 | 
| child 51546 | 2e26df807dc7 | 
| permissions | -rw-r--r-- | 
| 
35050
 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 
haftmann 
parents: 
35036 
diff
changeset
 | 
1  | 
(* Title: HOL/Groups.thy  | 
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
28823 
diff
changeset
 | 
2  | 
Author: Gertrud Bauer, Steven Obua, Lawrence C Paulson, Markus Wenzel, Jeremy Avigad  | 
| 14738 | 3  | 
*)  | 
4  | 
||
| 
35050
 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 
haftmann 
parents: 
35036 
diff
changeset
 | 
5  | 
header {* Groups, also combined with orderings *}
 | 
| 14738 | 6  | 
|
| 
35050
 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 
haftmann 
parents: 
35036 
diff
changeset
 | 
7  | 
theory Groups  | 
| 
35092
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
8  | 
imports Orderings  | 
| 15131 | 9  | 
begin  | 
| 14738 | 10  | 
|
| 
35301
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
11  | 
subsection {* Fact collections *}
 | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
12  | 
|
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
13  | 
ML {*
 | 
| 45294 | 14  | 
structure Ac_Simps = Named_Thms  | 
15  | 
(  | 
|
16  | 
  val name = @{binding ac_simps}
 | 
|
| 36343 | 17  | 
val description = "associativity and commutativity simplification rules"  | 
18  | 
)  | 
|
19  | 
*}  | 
|
20  | 
||
21  | 
setup Ac_Simps.setup  | 
|
22  | 
||
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
23  | 
text{* The rewrites accumulated in @{text algebra_simps} deal with the
 | 
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
24  | 
classical algebraic structures of groups, rings and family. They simplify  | 
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
25  | 
terms by multiplying everything out (in case of a ring) and bringing sums and  | 
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
26  | 
products into a canonical form (by ordered rewriting). As a result it decides  | 
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
27  | 
group and ring equalities but also helps with inequalities.  | 
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
28  | 
|
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
29  | 
Of course it also works for fields, but it knows nothing about multiplicative  | 
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
30  | 
inverses or division. This is catered for by @{text field_simps}. *}
 | 
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
31  | 
|
| 36343 | 32  | 
ML {*
 | 
| 45294 | 33  | 
structure Algebra_Simps = Named_Thms  | 
34  | 
(  | 
|
35  | 
  val name = @{binding algebra_simps}
 | 
|
| 
35301
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
36  | 
val description = "algebra simplification rules"  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
37  | 
)  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
38  | 
*}  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
39  | 
|
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
40  | 
setup Algebra_Simps.setup  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
41  | 
|
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
42  | 
text{* Lemmas @{text field_simps} multiply with denominators in (in)equations
 | 
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
43  | 
if they can be proved to be non-zero (for equations) or positive/negative  | 
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
44  | 
(for inequations). Can be too aggressive and is therefore separate from the  | 
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
45  | 
more benign @{text algebra_simps}. *}
 | 
| 
35301
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
46  | 
|
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
47  | 
ML {*
 | 
| 45294 | 48  | 
structure Field_Simps = Named_Thms  | 
49  | 
(  | 
|
50  | 
  val name = @{binding field_simps}
 | 
|
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
51  | 
val description = "algebra simplification rules for fields"  | 
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
52  | 
)  | 
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
53  | 
*}  | 
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
54  | 
|
| 
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
55  | 
setup Field_Simps.setup  | 
| 
35301
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
56  | 
|
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
57  | 
|
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
58  | 
subsection {* Abstract structures *}
 | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
59  | 
|
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
60  | 
text {*
 | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
61  | 
These locales provide basic structures for interpretation into  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
62  | 
bigger structures; extensions require careful thinking, otherwise  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
63  | 
undesired effects may occur due to interpretation.  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
64  | 
*}  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
65  | 
|
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
66  | 
locale semigroup =  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
67  | 
fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "*" 70)  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
68  | 
assumes assoc [ac_simps]: "a * b * c = a * (b * c)"  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
69  | 
|
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
70  | 
locale abel_semigroup = semigroup +  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
71  | 
assumes commute [ac_simps]: "a * b = b * a"  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
72  | 
begin  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
73  | 
|
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
74  | 
lemma left_commute [ac_simps]:  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
75  | 
"b * (a * c) = a * (b * c)"  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
76  | 
proof -  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
77  | 
have "(b * a) * c = (a * b) * c"  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
78  | 
by (simp only: commute)  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
79  | 
then show ?thesis  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
80  | 
by (simp only: assoc)  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
81  | 
qed  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
82  | 
|
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
83  | 
end  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
84  | 
|
| 35720 | 85  | 
locale monoid = semigroup +  | 
| 
35723
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
86  | 
  fixes z :: 'a ("1")
 | 
| 
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
87  | 
assumes left_neutral [simp]: "1 * a = a"  | 
| 
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
88  | 
assumes right_neutral [simp]: "a * 1 = a"  | 
| 35720 | 89  | 
|
90  | 
locale comm_monoid = abel_semigroup +  | 
|
| 
35723
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
91  | 
  fixes z :: 'a ("1")
 | 
| 
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
92  | 
assumes comm_neutral: "a * 1 = a"  | 
| 35720 | 93  | 
|
94  | 
sublocale comm_monoid < monoid proof  | 
|
95  | 
qed (simp_all add: commute comm_neutral)  | 
|
96  | 
||
| 
35301
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
97  | 
|
| 
35267
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
98  | 
subsection {* Generic operations *}
 | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
99  | 
|
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
100  | 
class zero =  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
101  | 
  fixes zero :: 'a  ("0")
 | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
102  | 
|
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
103  | 
class one =  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
104  | 
  fixes one  :: 'a  ("1")
 | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
105  | 
|
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
35828 
diff
changeset
 | 
106  | 
hide_const (open) zero one  | 
| 
35267
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
107  | 
|
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
108  | 
lemma Let_0 [simp]: "Let 0 f = f 0"  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
109  | 
unfolding Let_def ..  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
110  | 
|
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
111  | 
lemma Let_1 [simp]: "Let 1 f = f 1"  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
112  | 
unfolding Let_def ..  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
113  | 
|
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
114  | 
setup {*
 | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
115  | 
Reorient_Proc.add  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
116  | 
    (fn Const(@{const_name Groups.zero}, _) => true
 | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
117  | 
      | Const(@{const_name Groups.one}, _) => true
 | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
118  | 
| _ => false)  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
119  | 
*}  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
120  | 
|
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
121  | 
simproc_setup reorient_zero ("0 = x") = Reorient_Proc.proc
 | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
122  | 
simproc_setup reorient_one ("1 = x") = Reorient_Proc.proc
 | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
123  | 
|
| 
39134
 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 
wenzelm 
parents: 
37986 
diff
changeset
 | 
124  | 
typed_print_translation (advanced) {*
 | 
| 
42247
 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 
wenzelm 
parents: 
42245 
diff
changeset
 | 
125  | 
let  | 
| 
 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 
wenzelm 
parents: 
42245 
diff
changeset
 | 
126  | 
fun tr' c = (c, fn ctxt => fn T => fn ts =>  | 
| 
49690
 
a6814de45b69
more explicit show_type_constraint, show_sort_constraint;
 
wenzelm 
parents: 
49388 
diff
changeset
 | 
127  | 
if not (null ts) orelse T = dummyT orelse  | 
| 
 
a6814de45b69
more explicit show_type_constraint, show_sort_constraint;
 
wenzelm 
parents: 
49388 
diff
changeset
 | 
128  | 
not (Printer.show_type_constraint ctxt) andalso can Term.dest_Type T  | 
| 
42247
 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 
wenzelm 
parents: 
42245 
diff
changeset
 | 
129  | 
then raise Match  | 
| 42248 | 130  | 
else  | 
131  | 
        Syntax.const @{syntax_const "_constrain"} $ Syntax.const c $
 | 
|
132  | 
Syntax_Phases.term_of_typ ctxt T);  | 
|
| 
42247
 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 
wenzelm 
parents: 
42245 
diff
changeset
 | 
133  | 
  in map tr' [@{const_syntax Groups.one}, @{const_syntax Groups.zero}] end;
 | 
| 
35267
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
134  | 
*} -- {* show types that are presumably too general *}
 | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
135  | 
|
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
136  | 
class plus =  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
137  | 
fixes plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "+" 65)  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
138  | 
|
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
139  | 
class minus =  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
140  | 
fixes minus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "-" 65)  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
141  | 
|
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
142  | 
class uminus =  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
143  | 
  fixes uminus :: "'a \<Rightarrow> 'a"  ("- _" [81] 80)
 | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
144  | 
|
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
145  | 
class times =  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
146  | 
fixes times :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "*" 70)  | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35216 
diff
changeset
 | 
147  | 
|
| 
35092
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
148  | 
|
| 23085 | 149  | 
subsection {* Semigroups and Monoids *}
 | 
| 14738 | 150  | 
|
| 22390 | 151  | 
class semigroup_add = plus +  | 
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
152  | 
assumes add_assoc [algebra_simps, field_simps]: "(a + b) + c = a + (b + c)"  | 
| 
34973
 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 
haftmann 
parents: 
34147 
diff
changeset
 | 
153  | 
|
| 
35723
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
154  | 
sublocale semigroup_add < add!: semigroup plus proof  | 
| 
34973
 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 
haftmann 
parents: 
34147 
diff
changeset
 | 
155  | 
qed (fact add_assoc)  | 
| 22390 | 156  | 
|
157  | 
class ab_semigroup_add = semigroup_add +  | 
|
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
158  | 
assumes add_commute [algebra_simps, field_simps]: "a + b = b + a"  | 
| 
34973
 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 
haftmann 
parents: 
34147 
diff
changeset
 | 
159  | 
|
| 
35723
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
160  | 
sublocale ab_semigroup_add < add!: abel_semigroup plus proof  | 
| 
34973
 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 
haftmann 
parents: 
34147 
diff
changeset
 | 
161  | 
qed (fact add_commute)  | 
| 
 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 
haftmann 
parents: 
34147 
diff
changeset
 | 
162  | 
|
| 
 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 
haftmann 
parents: 
34147 
diff
changeset
 | 
163  | 
context ab_semigroup_add  | 
| 25062 | 164  | 
begin  | 
| 14738 | 165  | 
|
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
166  | 
lemmas add_left_commute [algebra_simps, field_simps] = add.left_commute  | 
| 25062 | 167  | 
|
168  | 
theorems add_ac = add_assoc add_commute add_left_commute  | 
|
169  | 
||
170  | 
end  | 
|
| 14738 | 171  | 
|
172  | 
theorems add_ac = add_assoc add_commute add_left_commute  | 
|
173  | 
||
| 22390 | 174  | 
class semigroup_mult = times +  | 
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
175  | 
assumes mult_assoc [algebra_simps, field_simps]: "(a * b) * c = a * (b * c)"  | 
| 
34973
 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 
haftmann 
parents: 
34147 
diff
changeset
 | 
176  | 
|
| 
35723
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
177  | 
sublocale semigroup_mult < mult!: semigroup times proof  | 
| 
34973
 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 
haftmann 
parents: 
34147 
diff
changeset
 | 
178  | 
qed (fact mult_assoc)  | 
| 14738 | 179  | 
|
| 22390 | 180  | 
class ab_semigroup_mult = semigroup_mult +  | 
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
181  | 
assumes mult_commute [algebra_simps, field_simps]: "a * b = b * a"  | 
| 
34973
 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 
haftmann 
parents: 
34147 
diff
changeset
 | 
182  | 
|
| 
35723
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
183  | 
sublocale ab_semigroup_mult < mult!: abel_semigroup times proof  | 
| 
34973
 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 
haftmann 
parents: 
34147 
diff
changeset
 | 
184  | 
qed (fact mult_commute)  | 
| 
 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 
haftmann 
parents: 
34147 
diff
changeset
 | 
185  | 
|
| 
 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 
haftmann 
parents: 
34147 
diff
changeset
 | 
186  | 
context ab_semigroup_mult  | 
| 23181 | 187  | 
begin  | 
| 14738 | 188  | 
|
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
189  | 
lemmas mult_left_commute [algebra_simps, field_simps] = mult.left_commute  | 
| 25062 | 190  | 
|
191  | 
theorems mult_ac = mult_assoc mult_commute mult_left_commute  | 
|
| 23181 | 192  | 
|
193  | 
end  | 
|
| 14738 | 194  | 
|
195  | 
theorems mult_ac = mult_assoc mult_commute mult_left_commute  | 
|
196  | 
||
| 23085 | 197  | 
class monoid_add = zero + semigroup_add +  | 
| 35720 | 198  | 
assumes add_0_left: "0 + a = a"  | 
199  | 
and add_0_right: "a + 0 = a"  | 
|
200  | 
||
| 
35723
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
201  | 
sublocale monoid_add < add!: monoid plus 0 proof  | 
| 35720 | 202  | 
qed (fact add_0_left add_0_right)+  | 
| 23085 | 203  | 
|
| 26071 | 204  | 
lemma zero_reorient: "0 = x \<longleftrightarrow> x = 0"  | 
| 29667 | 205  | 
by (rule eq_commute)  | 
| 26071 | 206  | 
|
| 22390 | 207  | 
class comm_monoid_add = zero + ab_semigroup_add +  | 
| 25062 | 208  | 
assumes add_0: "0 + a = a"  | 
| 23085 | 209  | 
|
| 
35723
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
210  | 
sublocale comm_monoid_add < add!: comm_monoid plus 0 proof  | 
| 35720 | 211  | 
qed (insert add_0, simp add: ac_simps)  | 
| 25062 | 212  | 
|
| 35720 | 213  | 
subclass (in comm_monoid_add) monoid_add proof  | 
| 
35723
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
214  | 
qed (fact add.left_neutral add.right_neutral)+  | 
| 14738 | 215  | 
|
| 49388 | 216  | 
class comm_monoid_diff = comm_monoid_add + minus +  | 
217  | 
assumes diff_zero [simp]: "a - 0 = a"  | 
|
218  | 
and zero_diff [simp]: "0 - a = 0"  | 
|
219  | 
and add_diff_cancel_left [simp]: "(c + a) - (c + b) = a - b"  | 
|
220  | 
and diff_diff_add: "a - b - c = a - (b + c)"  | 
|
221  | 
begin  | 
|
222  | 
||
223  | 
lemma add_diff_cancel_right [simp]:  | 
|
224  | 
"(a + c) - (b + c) = a - b"  | 
|
225  | 
using add_diff_cancel_left [symmetric] by (simp add: add.commute)  | 
|
226  | 
||
227  | 
lemma add_diff_cancel_left' [simp]:  | 
|
228  | 
"(b + a) - b = a"  | 
|
229  | 
proof -  | 
|
230  | 
have "(b + a) - (b + 0) = a" by (simp only: add_diff_cancel_left diff_zero)  | 
|
231  | 
then show ?thesis by simp  | 
|
232  | 
qed  | 
|
233  | 
||
234  | 
lemma add_diff_cancel_right' [simp]:  | 
|
235  | 
"(a + b) - b = a"  | 
|
236  | 
using add_diff_cancel_left' [symmetric] by (simp add: add.commute)  | 
|
237  | 
||
238  | 
lemma diff_add_zero [simp]:  | 
|
239  | 
"a - (a + b) = 0"  | 
|
240  | 
proof -  | 
|
241  | 
have "a - (a + b) = (a + 0) - (a + b)" by simp  | 
|
242  | 
also have "\<dots> = 0" by (simp only: add_diff_cancel_left zero_diff)  | 
|
243  | 
finally show ?thesis .  | 
|
244  | 
qed  | 
|
245  | 
||
246  | 
lemma diff_cancel [simp]:  | 
|
247  | 
"a - a = 0"  | 
|
248  | 
proof -  | 
|
249  | 
have "(a + 0) - (a + 0) = 0" by (simp only: add_diff_cancel_left diff_zero)  | 
|
250  | 
then show ?thesis by simp  | 
|
251  | 
qed  | 
|
252  | 
||
253  | 
lemma diff_right_commute:  | 
|
254  | 
"a - c - b = a - b - c"  | 
|
255  | 
by (simp add: diff_diff_add add.commute)  | 
|
256  | 
||
257  | 
lemma add_implies_diff:  | 
|
258  | 
assumes "c + b = a"  | 
|
259  | 
shows "c = a - b"  | 
|
260  | 
proof -  | 
|
261  | 
from assms have "(b + c) - (b + 0) = a - b" by (simp add: add.commute)  | 
|
262  | 
then show "c = a - b" by simp  | 
|
263  | 
qed  | 
|
264  | 
||
265  | 
end  | 
|
266  | 
||
| 22390 | 267  | 
class monoid_mult = one + semigroup_mult +  | 
| 35720 | 268  | 
assumes mult_1_left: "1 * a = a"  | 
269  | 
and mult_1_right: "a * 1 = a"  | 
|
270  | 
||
| 
35723
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
271  | 
sublocale monoid_mult < mult!: monoid times 1 proof  | 
| 35720 | 272  | 
qed (fact mult_1_left mult_1_right)+  | 
| 14738 | 273  | 
|
| 26071 | 274  | 
lemma one_reorient: "1 = x \<longleftrightarrow> x = 1"  | 
| 29667 | 275  | 
by (rule eq_commute)  | 
| 26071 | 276  | 
|
| 22390 | 277  | 
class comm_monoid_mult = one + ab_semigroup_mult +  | 
| 25062 | 278  | 
assumes mult_1: "1 * a = a"  | 
| 14738 | 279  | 
|
| 
35723
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
280  | 
sublocale comm_monoid_mult < mult!: comm_monoid times 1 proof  | 
| 35720 | 281  | 
qed (insert mult_1, simp add: ac_simps)  | 
| 25062 | 282  | 
|
| 35720 | 283  | 
subclass (in comm_monoid_mult) monoid_mult proof  | 
| 
35723
 
b6cf98f25c3f
tuned monoid locales and prefix of sublocale interpretations
 
haftmann 
parents: 
35720 
diff
changeset
 | 
284  | 
qed (fact mult.left_neutral mult.right_neutral)+  | 
| 14738 | 285  | 
|
| 22390 | 286  | 
class cancel_semigroup_add = semigroup_add +  | 
| 25062 | 287  | 
assumes add_left_imp_eq: "a + b = a + c \<Longrightarrow> b = c"  | 
288  | 
assumes add_right_imp_eq: "b + a = c + a \<Longrightarrow> b = c"  | 
|
| 
27474
 
a89d755b029d
move proofs of add_left_cancel and add_right_cancel into the correct locale
 
huffman 
parents: 
27250 
diff
changeset
 | 
289  | 
begin  | 
| 
 
a89d755b029d
move proofs of add_left_cancel and add_right_cancel into the correct locale
 
huffman 
parents: 
27250 
diff
changeset
 | 
290  | 
|
| 
 
a89d755b029d
move proofs of add_left_cancel and add_right_cancel into the correct locale
 
huffman 
parents: 
27250 
diff
changeset
 | 
291  | 
lemma add_left_cancel [simp]:  | 
| 
 
a89d755b029d
move proofs of add_left_cancel and add_right_cancel into the correct locale
 
huffman 
parents: 
27250 
diff
changeset
 | 
292  | 
"a + b = a + c \<longleftrightarrow> b = c"  | 
| 29667 | 293  | 
by (blast dest: add_left_imp_eq)  | 
| 
27474
 
a89d755b029d
move proofs of add_left_cancel and add_right_cancel into the correct locale
 
huffman 
parents: 
27250 
diff
changeset
 | 
294  | 
|
| 
 
a89d755b029d
move proofs of add_left_cancel and add_right_cancel into the correct locale
 
huffman 
parents: 
27250 
diff
changeset
 | 
295  | 
lemma add_right_cancel [simp]:  | 
| 
 
a89d755b029d
move proofs of add_left_cancel and add_right_cancel into the correct locale
 
huffman 
parents: 
27250 
diff
changeset
 | 
296  | 
"b + a = c + a \<longleftrightarrow> b = c"  | 
| 29667 | 297  | 
by (blast dest: add_right_imp_eq)  | 
| 
27474
 
a89d755b029d
move proofs of add_left_cancel and add_right_cancel into the correct locale
 
huffman 
parents: 
27250 
diff
changeset
 | 
298  | 
|
| 
 
a89d755b029d
move proofs of add_left_cancel and add_right_cancel into the correct locale
 
huffman 
parents: 
27250 
diff
changeset
 | 
299  | 
end  | 
| 14738 | 300  | 
|
| 22390 | 301  | 
class cancel_ab_semigroup_add = ab_semigroup_add +  | 
| 25062 | 302  | 
assumes add_imp_eq: "a + b = a + c \<Longrightarrow> b = c"  | 
| 25267 | 303  | 
begin  | 
| 14738 | 304  | 
|
| 25267 | 305  | 
subclass cancel_semigroup_add  | 
| 28823 | 306  | 
proof  | 
| 22390 | 307  | 
fix a b c :: 'a  | 
308  | 
assume "a + b = a + c"  | 
|
309  | 
then show "b = c" by (rule add_imp_eq)  | 
|
310  | 
next  | 
|
| 14738 | 311  | 
fix a b c :: 'a  | 
312  | 
assume "b + a = c + a"  | 
|
| 22390 | 313  | 
then have "a + b = a + c" by (simp only: add_commute)  | 
314  | 
then show "b = c" by (rule add_imp_eq)  | 
|
| 14738 | 315  | 
qed  | 
316  | 
||
| 25267 | 317  | 
end  | 
318  | 
||
| 29904 | 319  | 
class cancel_comm_monoid_add = cancel_ab_semigroup_add + comm_monoid_add  | 
320  | 
||
321  | 
||
| 23085 | 322  | 
subsection {* Groups *}
 | 
323  | 
||
| 25762 | 324  | 
class group_add = minus + uminus + monoid_add +  | 
| 25062 | 325  | 
assumes left_minus [simp]: "- a + a = 0"  | 
326  | 
assumes diff_minus: "a - b = a + (- b)"  | 
|
327  | 
begin  | 
|
| 23085 | 328  | 
|
| 
34147
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
329  | 
lemma minus_unique:  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
330  | 
assumes "a + b = 0" shows "- a = b"  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
331  | 
proof -  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
332  | 
have "- a = - a + (a + b)" using assms by simp  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
333  | 
also have "\<dots> = b" by (simp add: add_assoc [symmetric])  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
334  | 
finally show ?thesis .  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
335  | 
qed  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
336  | 
|
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
337  | 
lemmas equals_zero_I = minus_unique (* legacy name *)  | 
| 14738 | 338  | 
|
| 25062 | 339  | 
lemma minus_zero [simp]: "- 0 = 0"  | 
| 14738 | 340  | 
proof -  | 
| 
34147
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
341  | 
have "0 + 0 = 0" by (rule add_0_right)  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
342  | 
thus "- 0 = 0" by (rule minus_unique)  | 
| 14738 | 343  | 
qed  | 
344  | 
||
| 25062 | 345  | 
lemma minus_minus [simp]: "- (- a) = a"  | 
| 23085 | 346  | 
proof -  | 
| 
34147
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
347  | 
have "- a + a = 0" by (rule left_minus)  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
348  | 
thus "- (- a) = a" by (rule minus_unique)  | 
| 23085 | 349  | 
qed  | 
| 14738 | 350  | 
|
| 25062 | 351  | 
lemma right_minus [simp]: "a + - a = 0"  | 
| 14738 | 352  | 
proof -  | 
| 25062 | 353  | 
have "a + - a = - (- a) + - a" by simp  | 
354  | 
also have "\<dots> = 0" by (rule left_minus)  | 
|
| 14738 | 355  | 
finally show ?thesis .  | 
356  | 
qed  | 
|
357  | 
||
| 
40368
 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 
haftmann 
parents: 
39134 
diff
changeset
 | 
358  | 
subclass cancel_semigroup_add  | 
| 
 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 
haftmann 
parents: 
39134 
diff
changeset
 | 
359  | 
proof  | 
| 
 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 
haftmann 
parents: 
39134 
diff
changeset
 | 
360  | 
fix a b c :: 'a  | 
| 
 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 
haftmann 
parents: 
39134 
diff
changeset
 | 
361  | 
assume "a + b = a + c"  | 
| 
 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 
haftmann 
parents: 
39134 
diff
changeset
 | 
362  | 
then have "- a + a + b = - a + a + c"  | 
| 
 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 
haftmann 
parents: 
39134 
diff
changeset
 | 
363  | 
unfolding add_assoc by simp  | 
| 
 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 
haftmann 
parents: 
39134 
diff
changeset
 | 
364  | 
then show "b = c" by simp  | 
| 
 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 
haftmann 
parents: 
39134 
diff
changeset
 | 
365  | 
next  | 
| 
 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 
haftmann 
parents: 
39134 
diff
changeset
 | 
366  | 
fix a b c :: 'a  | 
| 
 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 
haftmann 
parents: 
39134 
diff
changeset
 | 
367  | 
assume "b + a = c + a"  | 
| 
 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 
haftmann 
parents: 
39134 
diff
changeset
 | 
368  | 
then have "b + a + - a = c + a + - a" by simp  | 
| 
 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 
haftmann 
parents: 
39134 
diff
changeset
 | 
369  | 
then show "b = c" unfolding add_assoc by simp  | 
| 
 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 
haftmann 
parents: 
39134 
diff
changeset
 | 
370  | 
qed  | 
| 
 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 
haftmann 
parents: 
39134 
diff
changeset
 | 
371  | 
|
| 
34147
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
372  | 
lemma minus_add_cancel: "- a + (a + b) = b"  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
373  | 
by (simp add: add_assoc [symmetric])  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
374  | 
|
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
375  | 
lemma add_minus_cancel: "a + (- a + b) = b"  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
376  | 
by (simp add: add_assoc [symmetric])  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
377  | 
|
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
378  | 
lemma minus_add: "- (a + b) = - b + - a"  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
379  | 
proof -  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
380  | 
have "(a + b) + (- b + - a) = 0"  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
381  | 
by (simp add: add_assoc add_minus_cancel)  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
382  | 
thus "- (a + b) = - b + - a"  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
383  | 
by (rule minus_unique)  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
384  | 
qed  | 
| 
 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 
huffman 
parents: 
34146 
diff
changeset
 | 
385  | 
|
| 25062 | 386  | 
lemma right_minus_eq: "a - b = 0 \<longleftrightarrow> a = b"  | 
| 14738 | 387  | 
proof  | 
| 23085 | 388  | 
assume "a - b = 0"  | 
389  | 
have "a = (a - b) + b" by (simp add:diff_minus add_assoc)  | 
|
390  | 
also have "\<dots> = b" using `a - b = 0` by simp  | 
|
391  | 
finally show "a = b" .  | 
|
| 14738 | 392  | 
next  | 
| 23085 | 393  | 
assume "a = b" thus "a - b = 0" by (simp add: diff_minus)  | 
| 14738 | 394  | 
qed  | 
395  | 
||
| 25062 | 396  | 
lemma diff_self [simp]: "a - a = 0"  | 
| 29667 | 397  | 
by (simp add: diff_minus)  | 
| 14738 | 398  | 
|
| 25062 | 399  | 
lemma diff_0 [simp]: "0 - a = - a"  | 
| 29667 | 400  | 
by (simp add: diff_minus)  | 
| 14738 | 401  | 
|
| 25062 | 402  | 
lemma diff_0_right [simp]: "a - 0 = a"  | 
| 29667 | 403  | 
by (simp add: diff_minus)  | 
| 14738 | 404  | 
|
| 25062 | 405  | 
lemma diff_minus_eq_add [simp]: "a - - b = a + b"  | 
| 29667 | 406  | 
by (simp add: diff_minus)  | 
| 14738 | 407  | 
|
| 25062 | 408  | 
lemma neg_equal_iff_equal [simp]:  | 
409  | 
"- a = - b \<longleftrightarrow> a = b"  | 
|
| 14738 | 410  | 
proof  | 
411  | 
assume "- a = - b"  | 
|
| 29667 | 412  | 
hence "- (- a) = - (- b)" by simp  | 
| 25062 | 413  | 
thus "a = b" by simp  | 
| 14738 | 414  | 
next  | 
| 25062 | 415  | 
assume "a = b"  | 
416  | 
thus "- a = - b" by simp  | 
|
| 14738 | 417  | 
qed  | 
418  | 
||
| 25062 | 419  | 
lemma neg_equal_0_iff_equal [simp]:  | 
420  | 
"- a = 0 \<longleftrightarrow> a = 0"  | 
|
| 29667 | 421  | 
by (subst neg_equal_iff_equal [symmetric], simp)  | 
| 14738 | 422  | 
|
| 25062 | 423  | 
lemma neg_0_equal_iff_equal [simp]:  | 
424  | 
"0 = - a \<longleftrightarrow> 0 = a"  | 
|
| 29667 | 425  | 
by (subst neg_equal_iff_equal [symmetric], simp)  | 
| 14738 | 426  | 
|
427  | 
text{*The next two equations can make the simplifier loop!*}
 | 
|
428  | 
||
| 25062 | 429  | 
lemma equation_minus_iff:  | 
430  | 
"a = - b \<longleftrightarrow> b = - a"  | 
|
| 14738 | 431  | 
proof -  | 
| 25062 | 432  | 
have "- (- a) = - b \<longleftrightarrow> - a = b" by (rule neg_equal_iff_equal)  | 
433  | 
thus ?thesis by (simp add: eq_commute)  | 
|
434  | 
qed  | 
|
435  | 
||
436  | 
lemma minus_equation_iff:  | 
|
437  | 
"- a = b \<longleftrightarrow> - b = a"  | 
|
438  | 
proof -  | 
|
439  | 
have "- a = - (- b) \<longleftrightarrow> a = -b" by (rule neg_equal_iff_equal)  | 
|
| 14738 | 440  | 
thus ?thesis by (simp add: eq_commute)  | 
441  | 
qed  | 
|
442  | 
||
| 
28130
 
32b4185bfdc7
move diff_add_cancel, add_diff_cancel from class ab_group_add to group_add
 
huffman 
parents: 
27516 
diff
changeset
 | 
443  | 
lemma diff_add_cancel: "a - b + b = a"  | 
| 29667 | 444  | 
by (simp add: diff_minus add_assoc)  | 
| 
28130
 
32b4185bfdc7
move diff_add_cancel, add_diff_cancel from class ab_group_add to group_add
 
huffman 
parents: 
27516 
diff
changeset
 | 
445  | 
|
| 
 
32b4185bfdc7
move diff_add_cancel, add_diff_cancel from class ab_group_add to group_add
 
huffman 
parents: 
27516 
diff
changeset
 | 
446  | 
lemma add_diff_cancel: "a + b - b = a"  | 
| 29667 | 447  | 
by (simp add: diff_minus add_assoc)  | 
448  | 
||
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
449  | 
declare diff_minus[symmetric, algebra_simps, field_simps]  | 
| 
28130
 
32b4185bfdc7
move diff_add_cancel, add_diff_cancel from class ab_group_add to group_add
 
huffman 
parents: 
27516 
diff
changeset
 | 
450  | 
|
| 
29914
 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 
huffman 
parents: 
29904 
diff
changeset
 | 
451  | 
lemma eq_neg_iff_add_eq_0: "a = - b \<longleftrightarrow> a + b = 0"  | 
| 
 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 
huffman 
parents: 
29904 
diff
changeset
 | 
452  | 
proof  | 
| 
 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 
huffman 
parents: 
29904 
diff
changeset
 | 
453  | 
assume "a = - b" then show "a + b = 0" by simp  | 
| 
 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 
huffman 
parents: 
29904 
diff
changeset
 | 
454  | 
next  | 
| 
 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 
huffman 
parents: 
29904 
diff
changeset
 | 
455  | 
assume "a + b = 0"  | 
| 
 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 
huffman 
parents: 
29904 
diff
changeset
 | 
456  | 
moreover have "a + (b + - b) = (a + b) + - b"  | 
| 
 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 
huffman 
parents: 
29904 
diff
changeset
 | 
457  | 
by (simp only: add_assoc)  | 
| 
 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 
huffman 
parents: 
29904 
diff
changeset
 | 
458  | 
ultimately show "a = - b" by simp  | 
| 
 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 
huffman 
parents: 
29904 
diff
changeset
 | 
459  | 
qed  | 
| 
 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 
huffman 
parents: 
29904 
diff
changeset
 | 
460  | 
|
| 44348 | 461  | 
lemma add_eq_0_iff: "x + y = 0 \<longleftrightarrow> y = - x"  | 
462  | 
unfolding eq_neg_iff_add_eq_0 [symmetric]  | 
|
463  | 
by (rule equation_minus_iff)  | 
|
464  | 
||
| 
45548
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
465  | 
lemma minus_diff_eq [simp]: "- (a - b) = b - a"  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
466  | 
by (simp add: diff_minus minus_add)  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
467  | 
|
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
468  | 
lemma add_diff_eq[algebra_simps, field_simps]: "a + (b - c) = (a + b) - c"  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
469  | 
by (simp add: diff_minus add_assoc)  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
470  | 
|
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
471  | 
lemma diff_eq_eq[algebra_simps, field_simps]: "a - b = c \<longleftrightarrow> a = c + b"  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
472  | 
by (auto simp add: diff_minus add_assoc)  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
473  | 
|
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
474  | 
lemma eq_diff_eq[algebra_simps, field_simps]: "a = c - b \<longleftrightarrow> a + b = c"  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
475  | 
by (auto simp add: diff_minus add_assoc)  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
476  | 
|
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
477  | 
lemma diff_diff_eq2[algebra_simps, field_simps]: "a - (b - c) = (a + c) - b"  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
478  | 
by (simp add: diff_minus minus_add add_assoc)  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
479  | 
|
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
480  | 
lemma eq_iff_diff_eq_0: "a = b \<longleftrightarrow> a - b = 0"  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
481  | 
by (fact right_minus_eq [symmetric])  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
482  | 
|
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
483  | 
lemma diff_eq_diff_eq:  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
484  | 
"a - b = c - d \<Longrightarrow> a = b \<longleftrightarrow> c = d"  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
485  | 
by (simp add: eq_iff_diff_eq_0 [of a b] eq_iff_diff_eq_0 [of c d])  | 
| 
 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 
huffman 
parents: 
45294 
diff
changeset
 | 
486  | 
|
| 25062 | 487  | 
end  | 
488  | 
||
| 25762 | 489  | 
class ab_group_add = minus + uminus + comm_monoid_add +  | 
| 25062 | 490  | 
assumes ab_left_minus: "- a + a = 0"  | 
491  | 
assumes ab_diff_minus: "a - b = a + (- b)"  | 
|
| 25267 | 492  | 
begin  | 
| 25062 | 493  | 
|
| 25267 | 494  | 
subclass group_add  | 
| 28823 | 495  | 
proof qed (simp_all add: ab_left_minus ab_diff_minus)  | 
| 25062 | 496  | 
|
| 29904 | 497  | 
subclass cancel_comm_monoid_add  | 
| 28823 | 498  | 
proof  | 
| 25062 | 499  | 
fix a b c :: 'a  | 
500  | 
assume "a + b = a + c"  | 
|
501  | 
then have "- a + a + b = - a + a + c"  | 
|
502  | 
unfolding add_assoc by simp  | 
|
503  | 
then show "b = c" by simp  | 
|
504  | 
qed  | 
|
505  | 
||
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
506  | 
lemma uminus_add_conv_diff[algebra_simps, field_simps]:  | 
| 25062 | 507  | 
"- a + b = b - a"  | 
| 29667 | 508  | 
by (simp add:diff_minus add_commute)  | 
| 25062 | 509  | 
|
510  | 
lemma minus_add_distrib [simp]:  | 
|
511  | 
"- (a + b) = - a + - b"  | 
|
| 
34146
 
14595e0c27e8
rename equals_zero_I to minus_unique (keep old name too)
 
huffman 
parents: 
33364 
diff
changeset
 | 
512  | 
by (rule minus_unique) (simp add: add_ac)  | 
| 25062 | 513  | 
|
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
514  | 
lemma diff_add_eq[algebra_simps, field_simps]: "(a - b) + c = (a + c) - b"  | 
| 29667 | 515  | 
by (simp add: diff_minus add_ac)  | 
| 25077 | 516  | 
|
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
517  | 
lemma diff_diff_eq[algebra_simps, field_simps]: "(a - b) - c = a - (b + c)"  | 
| 29667 | 518  | 
by (simp add: diff_minus add_ac)  | 
| 25077 | 519  | 
|
| 35216 | 520  | 
(* FIXME: duplicates right_minus_eq from class group_add *)  | 
521  | 
(* but only this one is declared as a simp rule. *)  | 
|
| 
35828
 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 
blanchet 
parents: 
35723 
diff
changeset
 | 
522  | 
lemma diff_eq_0_iff_eq [simp, no_atp]: "a - b = 0 \<longleftrightarrow> a = b"  | 
| 44348 | 523  | 
by (rule right_minus_eq)  | 
| 30629 | 524  | 
|
| 
48556
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
525  | 
lemma add_diff_cancel_left: "(c + a) - (c + b) = a - b"  | 
| 
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
526  | 
by (simp add: diff_minus add_ac)  | 
| 
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
527  | 
|
| 25062 | 528  | 
end  | 
| 14738 | 529  | 
|
| 
37884
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
530  | 
|
| 14738 | 531  | 
subsection {* (Partially) Ordered Groups *} 
 | 
532  | 
||
| 
35301
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
533  | 
text {*
 | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
534  | 
The theory of partially ordered groups is taken from the books:  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
535  | 
  \begin{itemize}
 | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
536  | 
  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
 | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
537  | 
  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
 | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
538  | 
  \end{itemize}
 | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
539  | 
Most of the used notions can also be looked up in  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
540  | 
  \begin{itemize}
 | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
541  | 
  \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
 | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
542  | 
  \item \emph{Algebra I} by van der Waerden, Springer.
 | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
543  | 
  \end{itemize}
 | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
544  | 
*}  | 
| 
 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 
haftmann 
parents: 
35267 
diff
changeset
 | 
545  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
546  | 
class ordered_ab_semigroup_add = order + ab_semigroup_add +  | 
| 25062 | 547  | 
assumes add_left_mono: "a \<le> b \<Longrightarrow> c + a \<le> c + b"  | 
548  | 
begin  | 
|
| 
24380
 
c215e256beca
moved ordered_ab_semigroup_add to OrderedGroup.thy
 
haftmann 
parents: 
24286 
diff
changeset
 | 
549  | 
|
| 25062 | 550  | 
lemma add_right_mono:  | 
551  | 
"a \<le> b \<Longrightarrow> a + c \<le> b + c"  | 
|
| 29667 | 552  | 
by (simp add: add_commute [of _ c] add_left_mono)  | 
| 14738 | 553  | 
|
554  | 
text {* non-strict, in both arguments *}
 | 
|
555  | 
lemma add_mono:  | 
|
| 25062 | 556  | 
"a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> a + c \<le> b + d"  | 
| 14738 | 557  | 
apply (erule add_right_mono [THEN order_trans])  | 
558  | 
apply (simp add: add_commute add_left_mono)  | 
|
559  | 
done  | 
|
560  | 
||
| 25062 | 561  | 
end  | 
562  | 
||
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
563  | 
class ordered_cancel_ab_semigroup_add =  | 
| 
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
564  | 
ordered_ab_semigroup_add + cancel_ab_semigroup_add  | 
| 25062 | 565  | 
begin  | 
566  | 
||
| 14738 | 567  | 
lemma add_strict_left_mono:  | 
| 25062 | 568  | 
"a < b \<Longrightarrow> c + a < c + b"  | 
| 29667 | 569  | 
by (auto simp add: less_le add_left_mono)  | 
| 14738 | 570  | 
|
571  | 
lemma add_strict_right_mono:  | 
|
| 25062 | 572  | 
"a < b \<Longrightarrow> a + c < b + c"  | 
| 29667 | 573  | 
by (simp add: add_commute [of _ c] add_strict_left_mono)  | 
| 14738 | 574  | 
|
575  | 
text{*Strict monotonicity in both arguments*}
 | 
|
| 25062 | 576  | 
lemma add_strict_mono:  | 
577  | 
"a < b \<Longrightarrow> c < d \<Longrightarrow> a + c < b + d"  | 
|
578  | 
apply (erule add_strict_right_mono [THEN less_trans])  | 
|
| 14738 | 579  | 
apply (erule add_strict_left_mono)  | 
580  | 
done  | 
|
581  | 
||
582  | 
lemma add_less_le_mono:  | 
|
| 25062 | 583  | 
"a < b \<Longrightarrow> c \<le> d \<Longrightarrow> a + c < b + d"  | 
584  | 
apply (erule add_strict_right_mono [THEN less_le_trans])  | 
|
585  | 
apply (erule add_left_mono)  | 
|
| 14738 | 586  | 
done  | 
587  | 
||
588  | 
lemma add_le_less_mono:  | 
|
| 25062 | 589  | 
"a \<le> b \<Longrightarrow> c < d \<Longrightarrow> a + c < b + d"  | 
590  | 
apply (erule add_right_mono [THEN le_less_trans])  | 
|
| 14738 | 591  | 
apply (erule add_strict_left_mono)  | 
592  | 
done  | 
|
593  | 
||
| 25062 | 594  | 
end  | 
595  | 
||
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
596  | 
class ordered_ab_semigroup_add_imp_le =  | 
| 
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
597  | 
ordered_cancel_ab_semigroup_add +  | 
| 25062 | 598  | 
assumes add_le_imp_le_left: "c + a \<le> c + b \<Longrightarrow> a \<le> b"  | 
599  | 
begin  | 
|
600  | 
||
| 14738 | 601  | 
lemma add_less_imp_less_left:  | 
| 29667 | 602  | 
assumes less: "c + a < c + b" shows "a < b"  | 
| 14738 | 603  | 
proof -  | 
604  | 
from less have le: "c + a <= c + b" by (simp add: order_le_less)  | 
|
605  | 
have "a <= b"  | 
|
606  | 
apply (insert le)  | 
|
607  | 
apply (drule add_le_imp_le_left)  | 
|
608  | 
by (insert le, drule add_le_imp_le_left, assumption)  | 
|
609  | 
moreover have "a \<noteq> b"  | 
|
610  | 
proof (rule ccontr)  | 
|
611  | 
assume "~(a \<noteq> b)"  | 
|
612  | 
then have "a = b" by simp  | 
|
613  | 
then have "c + a = c + b" by simp  | 
|
614  | 
with less show "False"by simp  | 
|
615  | 
qed  | 
|
616  | 
ultimately show "a < b" by (simp add: order_le_less)  | 
|
617  | 
qed  | 
|
618  | 
||
619  | 
lemma add_less_imp_less_right:  | 
|
| 25062 | 620  | 
"a + c < b + c \<Longrightarrow> a < b"  | 
| 14738 | 621  | 
apply (rule add_less_imp_less_left [of c])  | 
622  | 
apply (simp add: add_commute)  | 
|
623  | 
done  | 
|
624  | 
||
625  | 
lemma add_less_cancel_left [simp]:  | 
|
| 25062 | 626  | 
"c + a < c + b \<longleftrightarrow> a < b"  | 
| 29667 | 627  | 
by (blast intro: add_less_imp_less_left add_strict_left_mono)  | 
| 14738 | 628  | 
|
629  | 
lemma add_less_cancel_right [simp]:  | 
|
| 25062 | 630  | 
"a + c < b + c \<longleftrightarrow> a < b"  | 
| 29667 | 631  | 
by (blast intro: add_less_imp_less_right add_strict_right_mono)  | 
| 14738 | 632  | 
|
633  | 
lemma add_le_cancel_left [simp]:  | 
|
| 25062 | 634  | 
"c + a \<le> c + b \<longleftrightarrow> a \<le> b"  | 
| 29667 | 635  | 
by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono)  | 
| 14738 | 636  | 
|
637  | 
lemma add_le_cancel_right [simp]:  | 
|
| 25062 | 638  | 
"a + c \<le> b + c \<longleftrightarrow> a \<le> b"  | 
| 29667 | 639  | 
by (simp add: add_commute [of a c] add_commute [of b c])  | 
| 14738 | 640  | 
|
641  | 
lemma add_le_imp_le_right:  | 
|
| 25062 | 642  | 
"a + c \<le> b + c \<Longrightarrow> a \<le> b"  | 
| 29667 | 643  | 
by simp  | 
| 25062 | 644  | 
|
| 25077 | 645  | 
lemma max_add_distrib_left:  | 
646  | 
"max x y + z = max (x + z) (y + z)"  | 
|
647  | 
unfolding max_def by auto  | 
|
648  | 
||
649  | 
lemma min_add_distrib_left:  | 
|
650  | 
"min x y + z = min (x + z) (y + z)"  | 
|
651  | 
unfolding min_def by auto  | 
|
652  | 
||
| 
44848
 
f4d0b060c7ca
remove lemmas nat_add_min_{left,right} in favor of generic lemmas min_add_distrib_{left,right}
 
huffman 
parents: 
44433 
diff
changeset
 | 
653  | 
lemma max_add_distrib_right:  | 
| 
 
f4d0b060c7ca
remove lemmas nat_add_min_{left,right} in favor of generic lemmas min_add_distrib_{left,right}
 
huffman 
parents: 
44433 
diff
changeset
 | 
654  | 
"x + max y z = max (x + y) (x + z)"  | 
| 
 
f4d0b060c7ca
remove lemmas nat_add_min_{left,right} in favor of generic lemmas min_add_distrib_{left,right}
 
huffman 
parents: 
44433 
diff
changeset
 | 
655  | 
unfolding max_def by auto  | 
| 
 
f4d0b060c7ca
remove lemmas nat_add_min_{left,right} in favor of generic lemmas min_add_distrib_{left,right}
 
huffman 
parents: 
44433 
diff
changeset
 | 
656  | 
|
| 
 
f4d0b060c7ca
remove lemmas nat_add_min_{left,right} in favor of generic lemmas min_add_distrib_{left,right}
 
huffman 
parents: 
44433 
diff
changeset
 | 
657  | 
lemma min_add_distrib_right:  | 
| 
 
f4d0b060c7ca
remove lemmas nat_add_min_{left,right} in favor of generic lemmas min_add_distrib_{left,right}
 
huffman 
parents: 
44433 
diff
changeset
 | 
658  | 
"x + min y z = min (x + y) (x + z)"  | 
| 
 
f4d0b060c7ca
remove lemmas nat_add_min_{left,right} in favor of generic lemmas min_add_distrib_{left,right}
 
huffman 
parents: 
44433 
diff
changeset
 | 
659  | 
unfolding min_def by auto  | 
| 
 
f4d0b060c7ca
remove lemmas nat_add_min_{left,right} in favor of generic lemmas min_add_distrib_{left,right}
 
huffman 
parents: 
44433 
diff
changeset
 | 
660  | 
|
| 25062 | 661  | 
end  | 
662  | 
||
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
663  | 
subsection {* Support for reasoning about signs *}
 | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
664  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
665  | 
class ordered_comm_monoid_add =  | 
| 
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
666  | 
ordered_cancel_ab_semigroup_add + comm_monoid_add  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
667  | 
begin  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
668  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
669  | 
lemma add_pos_nonneg:  | 
| 29667 | 670  | 
assumes "0 < a" and "0 \<le> b" shows "0 < a + b"  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
671  | 
proof -  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
672  | 
have "0 + 0 < a + b"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
673  | 
using assms by (rule add_less_le_mono)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
674  | 
then show ?thesis by simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
675  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
676  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
677  | 
lemma add_pos_pos:  | 
| 29667 | 678  | 
assumes "0 < a" and "0 < b" shows "0 < a + b"  | 
679  | 
by (rule add_pos_nonneg) (insert assms, auto)  | 
|
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
680  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
681  | 
lemma add_nonneg_pos:  | 
| 29667 | 682  | 
assumes "0 \<le> a" and "0 < b" shows "0 < a + b"  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
683  | 
proof -  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
684  | 
have "0 + 0 < a + b"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
685  | 
using assms by (rule add_le_less_mono)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
686  | 
then show ?thesis by simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
687  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
688  | 
|
| 
36977
 
71c8973a604b
declare add_nonneg_nonneg [simp]; remove now-redundant lemmas realpow_two_le_order(2)
 
huffman 
parents: 
36348 
diff
changeset
 | 
689  | 
lemma add_nonneg_nonneg [simp]:  | 
| 29667 | 690  | 
assumes "0 \<le> a" and "0 \<le> b" shows "0 \<le> a + b"  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
691  | 
proof -  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
692  | 
have "0 + 0 \<le> a + b"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
693  | 
using assms by (rule add_mono)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
694  | 
then show ?thesis by simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
695  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
696  | 
|
| 30691 | 697  | 
lemma add_neg_nonpos:  | 
| 29667 | 698  | 
assumes "a < 0" and "b \<le> 0" shows "a + b < 0"  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
699  | 
proof -  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
700  | 
have "a + b < 0 + 0"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
701  | 
using assms by (rule add_less_le_mono)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
702  | 
then show ?thesis by simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
703  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
704  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
705  | 
lemma add_neg_neg:  | 
| 29667 | 706  | 
assumes "a < 0" and "b < 0" shows "a + b < 0"  | 
707  | 
by (rule add_neg_nonpos) (insert assms, auto)  | 
|
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
708  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
709  | 
lemma add_nonpos_neg:  | 
| 29667 | 710  | 
assumes "a \<le> 0" and "b < 0" shows "a + b < 0"  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
711  | 
proof -  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
712  | 
have "a + b < 0 + 0"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
713  | 
using assms by (rule add_le_less_mono)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
714  | 
then show ?thesis by simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
715  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
716  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
717  | 
lemma add_nonpos_nonpos:  | 
| 29667 | 718  | 
assumes "a \<le> 0" and "b \<le> 0" shows "a + b \<le> 0"  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
719  | 
proof -  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
720  | 
have "a + b \<le> 0 + 0"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
721  | 
using assms by (rule add_mono)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
722  | 
then show ?thesis by simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
723  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
724  | 
|
| 30691 | 725  | 
lemmas add_sign_intros =  | 
726  | 
add_pos_nonneg add_pos_pos add_nonneg_pos add_nonneg_nonneg  | 
|
727  | 
add_neg_nonpos add_neg_neg add_nonpos_neg add_nonpos_nonpos  | 
|
728  | 
||
| 29886 | 729  | 
lemma add_nonneg_eq_0_iff:  | 
730  | 
assumes x: "0 \<le> x" and y: "0 \<le> y"  | 
|
731  | 
shows "x + y = 0 \<longleftrightarrow> x = 0 \<and> y = 0"  | 
|
732  | 
proof (intro iffI conjI)  | 
|
733  | 
have "x = x + 0" by simp  | 
|
734  | 
also have "x + 0 \<le> x + y" using y by (rule add_left_mono)  | 
|
735  | 
also assume "x + y = 0"  | 
|
736  | 
also have "0 \<le> x" using x .  | 
|
737  | 
finally show "x = 0" .  | 
|
738  | 
next  | 
|
739  | 
have "y = 0 + y" by simp  | 
|
740  | 
also have "0 + y \<le> x + y" using x by (rule add_right_mono)  | 
|
741  | 
also assume "x + y = 0"  | 
|
742  | 
also have "0 \<le> y" using y .  | 
|
743  | 
finally show "y = 0" .  | 
|
744  | 
next  | 
|
745  | 
assume "x = 0 \<and> y = 0"  | 
|
746  | 
then show "x + y = 0" by simp  | 
|
747  | 
qed  | 
|
748  | 
||
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
749  | 
end  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
750  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
751  | 
class ordered_ab_group_add =  | 
| 
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
752  | 
ab_group_add + ordered_ab_semigroup_add  | 
| 25062 | 753  | 
begin  | 
754  | 
||
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
755  | 
subclass ordered_cancel_ab_semigroup_add ..  | 
| 25062 | 756  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
757  | 
subclass ordered_ab_semigroup_add_imp_le  | 
| 28823 | 758  | 
proof  | 
| 25062 | 759  | 
fix a b c :: 'a  | 
760  | 
assume "c + a \<le> c + b"  | 
|
761  | 
hence "(-c) + (c + a) \<le> (-c) + (c + b)" by (rule add_left_mono)  | 
|
762  | 
hence "((-c) + c) + a \<le> ((-c) + c) + b" by (simp only: add_assoc)  | 
|
763  | 
thus "a \<le> b" by simp  | 
|
764  | 
qed  | 
|
765  | 
||
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
766  | 
subclass ordered_comm_monoid_add ..  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
767  | 
|
| 25077 | 768  | 
lemma max_diff_distrib_left:  | 
769  | 
shows "max x y - z = max (x - z) (y - z)"  | 
|
| 29667 | 770  | 
by (simp add: diff_minus, rule max_add_distrib_left)  | 
| 25077 | 771  | 
|
772  | 
lemma min_diff_distrib_left:  | 
|
773  | 
shows "min x y - z = min (x - z) (y - z)"  | 
|
| 29667 | 774  | 
by (simp add: diff_minus, rule min_add_distrib_left)  | 
| 25077 | 775  | 
|
776  | 
lemma le_imp_neg_le:  | 
|
| 29667 | 777  | 
assumes "a \<le> b" shows "-b \<le> -a"  | 
| 25077 | 778  | 
proof -  | 
| 29667 | 779  | 
have "-a+a \<le> -a+b" using `a \<le> b` by (rule add_left_mono)  | 
780  | 
hence "0 \<le> -a+b" by simp  | 
|
781  | 
hence "0 + (-b) \<le> (-a + b) + (-b)" by (rule add_right_mono)  | 
|
782  | 
thus ?thesis by (simp add: add_assoc)  | 
|
| 25077 | 783  | 
qed  | 
784  | 
||
785  | 
lemma neg_le_iff_le [simp]: "- b \<le> - a \<longleftrightarrow> a \<le> b"  | 
|
786  | 
proof  | 
|
787  | 
assume "- b \<le> - a"  | 
|
| 29667 | 788  | 
hence "- (- a) \<le> - (- b)" by (rule le_imp_neg_le)  | 
| 25077 | 789  | 
thus "a\<le>b" by simp  | 
790  | 
next  | 
|
791  | 
assume "a\<le>b"  | 
|
792  | 
thus "-b \<le> -a" by (rule le_imp_neg_le)  | 
|
793  | 
qed  | 
|
794  | 
||
795  | 
lemma neg_le_0_iff_le [simp]: "- a \<le> 0 \<longleftrightarrow> 0 \<le> a"  | 
|
| 29667 | 796  | 
by (subst neg_le_iff_le [symmetric], simp)  | 
| 25077 | 797  | 
|
798  | 
lemma neg_0_le_iff_le [simp]: "0 \<le> - a \<longleftrightarrow> a \<le> 0"  | 
|
| 29667 | 799  | 
by (subst neg_le_iff_le [symmetric], simp)  | 
| 25077 | 800  | 
|
801  | 
lemma neg_less_iff_less [simp]: "- b < - a \<longleftrightarrow> a < b"  | 
|
| 29667 | 802  | 
by (force simp add: less_le)  | 
| 25077 | 803  | 
|
804  | 
lemma neg_less_0_iff_less [simp]: "- a < 0 \<longleftrightarrow> 0 < a"  | 
|
| 29667 | 805  | 
by (subst neg_less_iff_less [symmetric], simp)  | 
| 25077 | 806  | 
|
807  | 
lemma neg_0_less_iff_less [simp]: "0 < - a \<longleftrightarrow> a < 0"  | 
|
| 29667 | 808  | 
by (subst neg_less_iff_less [symmetric], simp)  | 
| 25077 | 809  | 
|
810  | 
text{*The next several equations can make the simplifier loop!*}
 | 
|
811  | 
||
812  | 
lemma less_minus_iff: "a < - b \<longleftrightarrow> b < - a"  | 
|
813  | 
proof -  | 
|
814  | 
have "(- (-a) < - b) = (b < - a)" by (rule neg_less_iff_less)  | 
|
815  | 
thus ?thesis by simp  | 
|
816  | 
qed  | 
|
817  | 
||
818  | 
lemma minus_less_iff: "- a < b \<longleftrightarrow> - b < a"  | 
|
819  | 
proof -  | 
|
820  | 
have "(- a < - (-b)) = (- b < a)" by (rule neg_less_iff_less)  | 
|
821  | 
thus ?thesis by simp  | 
|
822  | 
qed  | 
|
823  | 
||
824  | 
lemma le_minus_iff: "a \<le> - b \<longleftrightarrow> b \<le> - a"  | 
|
825  | 
proof -  | 
|
826  | 
have mm: "!! a (b::'a). (-(-a)) < -b \<Longrightarrow> -(-b) < -a" by (simp only: minus_less_iff)  | 
|
827  | 
have "(- (- a) <= -b) = (b <= - a)"  | 
|
828  | 
apply (auto simp only: le_less)  | 
|
829  | 
apply (drule mm)  | 
|
830  | 
apply (simp_all)  | 
|
831  | 
apply (drule mm[simplified], assumption)  | 
|
832  | 
done  | 
|
833  | 
then show ?thesis by simp  | 
|
834  | 
qed  | 
|
835  | 
||
836  | 
lemma minus_le_iff: "- a \<le> b \<longleftrightarrow> - b \<le> a"  | 
|
| 29667 | 837  | 
by (auto simp add: le_less minus_less_iff)  | 
| 25077 | 838  | 
|
| 
37884
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
839  | 
lemma diff_less_0_iff_less [simp, no_atp]:  | 
| 
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
840  | 
"a - b < 0 \<longleftrightarrow> a < b"  | 
| 25077 | 841  | 
proof -  | 
| 
37884
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
842  | 
have "a - b < 0 \<longleftrightarrow> a + (- b) < b + (- b)" by (simp add: diff_minus)  | 
| 
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
843  | 
also have "... \<longleftrightarrow> a < b" by (simp only: add_less_cancel_right)  | 
| 25077 | 844  | 
finally show ?thesis .  | 
845  | 
qed  | 
|
846  | 
||
| 
37884
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
847  | 
lemmas less_iff_diff_less_0 = diff_less_0_iff_less [symmetric]  | 
| 
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
848  | 
|
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
849  | 
lemma diff_less_eq[algebra_simps, field_simps]: "a - b < c \<longleftrightarrow> a < c + b"  | 
| 25077 | 850  | 
apply (subst less_iff_diff_less_0 [of a])  | 
851  | 
apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])  | 
|
852  | 
apply (simp add: diff_minus add_ac)  | 
|
853  | 
done  | 
|
854  | 
||
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
855  | 
lemma less_diff_eq[algebra_simps, field_simps]: "a < c - b \<longleftrightarrow> a + b < c"  | 
| 36302 | 856  | 
apply (subst less_iff_diff_less_0 [of "a + b"])  | 
| 25077 | 857  | 
apply (subst less_iff_diff_less_0 [of a])  | 
858  | 
apply (simp add: diff_minus add_ac)  | 
|
859  | 
done  | 
|
860  | 
||
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
861  | 
lemma diff_le_eq[algebra_simps, field_simps]: "a - b \<le> c \<longleftrightarrow> a \<le> c + b"  | 
| 29667 | 862  | 
by (auto simp add: le_less diff_less_eq diff_add_cancel add_diff_cancel)  | 
| 25077 | 863  | 
|
| 
36348
 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 
haftmann 
parents: 
36343 
diff
changeset
 | 
864  | 
lemma le_diff_eq[algebra_simps, field_simps]: "a \<le> c - b \<longleftrightarrow> a + b \<le> c"  | 
| 29667 | 865  | 
by (auto simp add: le_less less_diff_eq diff_add_cancel add_diff_cancel)  | 
| 25077 | 866  | 
|
| 
37884
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
867  | 
lemma diff_le_0_iff_le [simp, no_atp]:  | 
| 
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
868  | 
"a - b \<le> 0 \<longleftrightarrow> a \<le> b"  | 
| 
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
869  | 
by (simp add: algebra_simps)  | 
| 
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
870  | 
|
| 
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
871  | 
lemmas le_iff_diff_le_0 = diff_le_0_iff_le [symmetric]  | 
| 
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
872  | 
|
| 
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
873  | 
lemma diff_eq_diff_less:  | 
| 
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
874  | 
"a - b = c - d \<Longrightarrow> a < b \<longleftrightarrow> c < d"  | 
| 
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
875  | 
by (auto simp only: less_iff_diff_less_0 [of a b] less_iff_diff_less_0 [of c d])  | 
| 
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
876  | 
|
| 
37889
 
0d8058e0c270
keep explicit diff_def as legacy theorem; modernized abel_cancel simproc setup
 
haftmann 
parents: 
37884 
diff
changeset
 | 
877  | 
lemma diff_eq_diff_less_eq:  | 
| 
 
0d8058e0c270
keep explicit diff_def as legacy theorem; modernized abel_cancel simproc setup
 
haftmann 
parents: 
37884 
diff
changeset
 | 
878  | 
"a - b = c - d \<Longrightarrow> a \<le> b \<longleftrightarrow> c \<le> d"  | 
| 
 
0d8058e0c270
keep explicit diff_def as legacy theorem; modernized abel_cancel simproc setup
 
haftmann 
parents: 
37884 
diff
changeset
 | 
879  | 
by (auto simp only: le_iff_diff_le_0 [of a b] le_iff_diff_le_0 [of c d])  | 
| 25077 | 880  | 
|
881  | 
end  | 
|
882  | 
||
| 48891 | 883  | 
ML_file "Tools/group_cancel.ML"  | 
| 
48556
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
884  | 
|
| 
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
885  | 
simproc_setup group_cancel_add ("a + b::'a::ab_group_add") =
 | 
| 
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
886  | 
  {* fn phi => fn ss => try Group_Cancel.cancel_add_conv *}
 | 
| 
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
887  | 
|
| 
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
888  | 
simproc_setup group_cancel_diff ("a - b::'a::ab_group_add") =
 | 
| 
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
889  | 
  {* fn phi => fn ss => try Group_Cancel.cancel_diff_conv *}
 | 
| 
37884
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
890  | 
|
| 
48556
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
891  | 
simproc_setup group_cancel_eq ("a = (b::'a::ab_group_add)") =
 | 
| 
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
892  | 
  {* fn phi => fn ss => try Group_Cancel.cancel_eq_conv *}
 | 
| 
37889
 
0d8058e0c270
keep explicit diff_def as legacy theorem; modernized abel_cancel simproc setup
 
haftmann 
parents: 
37884 
diff
changeset
 | 
893  | 
|
| 
48556
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
894  | 
simproc_setup group_cancel_le ("a \<le> (b::'a::ordered_ab_group_add)") =
 | 
| 
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
895  | 
  {* fn phi => fn ss => try Group_Cancel.cancel_le_conv *}
 | 
| 
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
896  | 
|
| 
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
897  | 
simproc_setup group_cancel_less ("a < (b::'a::ordered_ab_group_add)") =
 | 
| 
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
45548 
diff
changeset
 | 
898  | 
  {* fn phi => fn ss => try Group_Cancel.cancel_less_conv *}
 | 
| 
37884
 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 
haftmann 
parents: 
36977 
diff
changeset
 | 
899  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
900  | 
class linordered_ab_semigroup_add =  | 
| 
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
901  | 
linorder + ordered_ab_semigroup_add  | 
| 25062 | 902  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
903  | 
class linordered_cancel_ab_semigroup_add =  | 
| 
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
904  | 
linorder + ordered_cancel_ab_semigroup_add  | 
| 25267 | 905  | 
begin  | 
| 25062 | 906  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
907  | 
subclass linordered_ab_semigroup_add ..  | 
| 25062 | 908  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
909  | 
subclass ordered_ab_semigroup_add_imp_le  | 
| 28823 | 910  | 
proof  | 
| 25062 | 911  | 
fix a b c :: 'a  | 
912  | 
assume le: "c + a <= c + b"  | 
|
913  | 
show "a <= b"  | 
|
914  | 
proof (rule ccontr)  | 
|
915  | 
assume w: "~ a \<le> b"  | 
|
916  | 
hence "b <= a" by (simp add: linorder_not_le)  | 
|
917  | 
hence le2: "c + b <= c + a" by (rule add_left_mono)  | 
|
918  | 
have "a = b"  | 
|
919  | 
apply (insert le)  | 
|
920  | 
apply (insert le2)  | 
|
921  | 
apply (drule antisym, simp_all)  | 
|
922  | 
done  | 
|
923  | 
with w show False  | 
|
924  | 
by (simp add: linorder_not_le [symmetric])  | 
|
925  | 
qed  | 
|
926  | 
qed  | 
|
927  | 
||
| 25267 | 928  | 
end  | 
929  | 
||
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
930  | 
class linordered_ab_group_add = linorder + ordered_ab_group_add  | 
| 25267 | 931  | 
begin  | 
| 25230 | 932  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
933  | 
subclass linordered_cancel_ab_semigroup_add ..  | 
| 25230 | 934  | 
|
| 
35036
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
935  | 
lemma neg_less_eq_nonneg [simp]:  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
936  | 
"- a \<le> a \<longleftrightarrow> 0 \<le> a"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
937  | 
proof  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
938  | 
assume A: "- a \<le> a" show "0 \<le> a"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
939  | 
proof (rule classical)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
940  | 
assume "\<not> 0 \<le> a"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
941  | 
then have "a < 0" by auto  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
942  | 
with A have "- a < 0" by (rule le_less_trans)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
943  | 
then show ?thesis by auto  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
944  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
945  | 
next  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
946  | 
assume A: "0 \<le> a" show "- a \<le> a"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
947  | 
proof (rule order_trans)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
948  | 
show "- a \<le> 0" using A by (simp add: minus_le_iff)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
949  | 
next  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
950  | 
show "0 \<le> a" using A .  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
951  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
952  | 
qed  | 
| 
35036
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
953  | 
|
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
954  | 
lemma neg_less_nonneg [simp]:  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
955  | 
"- a < a \<longleftrightarrow> 0 < a"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
956  | 
proof  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
957  | 
assume A: "- a < a" show "0 < a"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
958  | 
proof (rule classical)  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
959  | 
assume "\<not> 0 < a"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
960  | 
then have "a \<le> 0" by auto  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
961  | 
with A have "- a < 0" by (rule less_le_trans)  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
962  | 
then show ?thesis by auto  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
963  | 
qed  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
964  | 
next  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
965  | 
assume A: "0 < a" show "- a < a"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
966  | 
proof (rule less_trans)  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
967  | 
show "- a < 0" using A by (simp add: minus_le_iff)  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
968  | 
next  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
969  | 
show "0 < a" using A .  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
970  | 
qed  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
971  | 
qed  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
972  | 
|
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
973  | 
lemma less_eq_neg_nonpos [simp]:  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
974  | 
"a \<le> - a \<longleftrightarrow> a \<le> 0"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
975  | 
proof  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
976  | 
assume A: "a \<le> - a" show "a \<le> 0"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
977  | 
proof (rule classical)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
978  | 
assume "\<not> a \<le> 0"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
979  | 
then have "0 < a" by auto  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
980  | 
then have "0 < - a" using A by (rule less_le_trans)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
981  | 
then show ?thesis by auto  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
982  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
983  | 
next  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
984  | 
assume A: "a \<le> 0" show "a \<le> - a"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
985  | 
proof (rule order_trans)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
986  | 
show "0 \<le> - a" using A by (simp add: minus_le_iff)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
987  | 
next  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
988  | 
show "a \<le> 0" using A .  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
989  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
990  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
991  | 
|
| 
35036
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
992  | 
lemma equal_neg_zero [simp]:  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
993  | 
"a = - a \<longleftrightarrow> a = 0"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
994  | 
proof  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
995  | 
assume "a = 0" then show "a = - a" by simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
996  | 
next  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
997  | 
assume A: "a = - a" show "a = 0"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
998  | 
proof (cases "0 \<le> a")  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
999  | 
case True with A have "0 \<le> - a" by auto  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1000  | 
with le_minus_iff have "a \<le> 0" by simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1001  | 
with True show ?thesis by (auto intro: order_trans)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1002  | 
next  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1003  | 
case False then have B: "a \<le> 0" by auto  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1004  | 
with A have "- a \<le> 0" by auto  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1005  | 
with B show ?thesis by (auto intro: order_trans)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1006  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1007  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1008  | 
|
| 
35036
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1009  | 
lemma neg_equal_zero [simp]:  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1010  | 
"- a = a \<longleftrightarrow> a = 0"  | 
| 
35036
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1011  | 
by (auto dest: sym)  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1012  | 
|
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1013  | 
lemma double_zero [simp]:  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1014  | 
"a + a = 0 \<longleftrightarrow> a = 0"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1015  | 
proof  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1016  | 
assume assm: "a + a = 0"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1017  | 
then have a: "- a = a" by (rule minus_unique)  | 
| 35216 | 1018  | 
then show "a = 0" by (simp only: neg_equal_zero)  | 
| 
35036
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1019  | 
qed simp  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1020  | 
|
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1021  | 
lemma double_zero_sym [simp]:  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1022  | 
"0 = a + a \<longleftrightarrow> a = 0"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1023  | 
by (rule, drule sym) simp_all  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1024  | 
|
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1025  | 
lemma zero_less_double_add_iff_zero_less_single_add [simp]:  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1026  | 
"0 < a + a \<longleftrightarrow> 0 < a"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1027  | 
proof  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1028  | 
assume "0 < a + a"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1029  | 
then have "0 - a < a" by (simp only: diff_less_eq)  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1030  | 
then have "- a < a" by simp  | 
| 35216 | 1031  | 
then show "0 < a" by (simp only: neg_less_nonneg)  | 
| 
35036
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1032  | 
next  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1033  | 
assume "0 < a"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1034  | 
with this have "0 + 0 < a + a"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1035  | 
by (rule add_strict_mono)  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1036  | 
then show "0 < a + a" by simp  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1037  | 
qed  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1038  | 
|
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1039  | 
lemma zero_le_double_add_iff_zero_le_single_add [simp]:  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1040  | 
"0 \<le> a + a \<longleftrightarrow> 0 \<le> a"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1041  | 
by (auto simp add: le_less)  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1042  | 
|
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1043  | 
lemma double_add_less_zero_iff_single_add_less_zero [simp]:  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1044  | 
"a + a < 0 \<longleftrightarrow> a < 0"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1045  | 
proof -  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1046  | 
have "\<not> a + a < 0 \<longleftrightarrow> \<not> a < 0"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1047  | 
by (simp add: not_less)  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1048  | 
then show ?thesis by simp  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1049  | 
qed  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1050  | 
|
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1051  | 
lemma double_add_le_zero_iff_single_add_le_zero [simp]:  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1052  | 
"a + a \<le> 0 \<longleftrightarrow> a \<le> 0"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1053  | 
proof -  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1054  | 
have "\<not> a + a \<le> 0 \<longleftrightarrow> \<not> a \<le> 0"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1055  | 
by (simp add: not_le)  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1056  | 
then show ?thesis by simp  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1057  | 
qed  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1058  | 
|
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1059  | 
lemma le_minus_self_iff:  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1060  | 
"a \<le> - a \<longleftrightarrow> a \<le> 0"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1061  | 
proof -  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1062  | 
from add_le_cancel_left [of "- a" "a + a" 0]  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1063  | 
have "a \<le> - a \<longleftrightarrow> a + a \<le> 0"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1064  | 
by (simp add: add_assoc [symmetric])  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1065  | 
thus ?thesis by simp  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1066  | 
qed  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1067  | 
|
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1068  | 
lemma minus_le_self_iff:  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1069  | 
"- a \<le> a \<longleftrightarrow> 0 \<le> a"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1070  | 
proof -  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1071  | 
from add_le_cancel_left [of "- a" 0 "a + a"]  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1072  | 
have "- a \<le> a \<longleftrightarrow> 0 \<le> a + a"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1073  | 
by (simp add: add_assoc [symmetric])  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1074  | 
thus ?thesis by simp  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1075  | 
qed  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1076  | 
|
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1077  | 
lemma minus_max_eq_min:  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1078  | 
"- max x y = min (-x) (-y)"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1079  | 
by (auto simp add: max_def min_def)  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1080  | 
|
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1081  | 
lemma minus_min_eq_max:  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1082  | 
"- min x y = max (-x) (-y)"  | 
| 
 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1083  | 
by (auto simp add: max_def min_def)  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1084  | 
|
| 25267 | 1085  | 
end  | 
1086  | 
||
| 36302 | 1087  | 
context ordered_comm_monoid_add  | 
1088  | 
begin  | 
|
| 14738 | 1089  | 
|
| 
15234
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
1090  | 
lemma add_increasing:  | 
| 36302 | 1091  | 
"0 \<le> a \<Longrightarrow> b \<le> c \<Longrightarrow> b \<le> a + c"  | 
1092  | 
by (insert add_mono [of 0 a b c], simp)  | 
|
| 14738 | 1093  | 
|
| 15539 | 1094  | 
lemma add_increasing2:  | 
| 36302 | 1095  | 
"0 \<le> c \<Longrightarrow> b \<le> a \<Longrightarrow> b \<le> a + c"  | 
1096  | 
by (simp add: add_increasing add_commute [of a])  | 
|
| 15539 | 1097  | 
|
| 
15234
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
1098  | 
lemma add_strict_increasing:  | 
| 36302 | 1099  | 
"0 < a \<Longrightarrow> b \<le> c \<Longrightarrow> b < a + c"  | 
1100  | 
by (insert add_less_le_mono [of 0 a b c], simp)  | 
|
| 
15234
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
1101  | 
|
| 
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
1102  | 
lemma add_strict_increasing2:  | 
| 36302 | 1103  | 
"0 \<le> a \<Longrightarrow> b < c \<Longrightarrow> b < a + c"  | 
1104  | 
by (insert add_le_less_mono [of 0 a b c], simp)  | 
|
1105  | 
||
1106  | 
end  | 
|
| 
15234
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
1107  | 
|
| 
35092
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1108  | 
class abs =  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1109  | 
fixes abs :: "'a \<Rightarrow> 'a"  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1110  | 
begin  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1111  | 
|
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1112  | 
notation (xsymbols)  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1113  | 
  abs  ("\<bar>_\<bar>")
 | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1114  | 
|
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1115  | 
notation (HTML output)  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1116  | 
  abs  ("\<bar>_\<bar>")
 | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1117  | 
|
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1118  | 
end  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1119  | 
|
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1120  | 
class sgn =  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1121  | 
fixes sgn :: "'a \<Rightarrow> 'a"  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1122  | 
|
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1123  | 
class abs_if = minus + uminus + ord + zero + abs +  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1124  | 
assumes abs_if: "\<bar>a\<bar> = (if a < 0 then - a else a)"  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1125  | 
|
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1126  | 
class sgn_if = minus + uminus + zero + one + ord + sgn +  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1127  | 
assumes sgn_if: "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)"  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1128  | 
begin  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1129  | 
|
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1130  | 
lemma sgn0 [simp]: "sgn 0 = 0"  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1131  | 
by (simp add:sgn_if)  | 
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1132  | 
|
| 
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35050 
diff
changeset
 | 
1133  | 
end  | 
| 14738 | 1134  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
1135  | 
class ordered_ab_group_add_abs = ordered_ab_group_add + abs +  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1136  | 
assumes abs_ge_zero [simp]: "\<bar>a\<bar> \<ge> 0"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1137  | 
and abs_ge_self: "a \<le> \<bar>a\<bar>"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1138  | 
and abs_leI: "a \<le> b \<Longrightarrow> - a \<le> b \<Longrightarrow> \<bar>a\<bar> \<le> b"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1139  | 
and abs_minus_cancel [simp]: "\<bar>-a\<bar> = \<bar>a\<bar>"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1140  | 
and abs_triangle_ineq: "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1141  | 
begin  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1142  | 
|
| 25307 | 1143  | 
lemma abs_minus_le_zero: "- \<bar>a\<bar> \<le> 0"  | 
1144  | 
unfolding neg_le_0_iff_le by simp  | 
|
1145  | 
||
1146  | 
lemma abs_of_nonneg [simp]:  | 
|
| 29667 | 1147  | 
assumes nonneg: "0 \<le> a" shows "\<bar>a\<bar> = a"  | 
| 25307 | 1148  | 
proof (rule antisym)  | 
1149  | 
from nonneg le_imp_neg_le have "- a \<le> 0" by simp  | 
|
1150  | 
from this nonneg have "- a \<le> a" by (rule order_trans)  | 
|
1151  | 
then show "\<bar>a\<bar> \<le> a" by (auto intro: abs_leI)  | 
|
1152  | 
qed (rule abs_ge_self)  | 
|
1153  | 
||
1154  | 
lemma abs_idempotent [simp]: "\<bar>\<bar>a\<bar>\<bar> = \<bar>a\<bar>"  | 
|
| 29667 | 1155  | 
by (rule antisym)  | 
| 36302 | 1156  | 
(auto intro!: abs_ge_self abs_leI order_trans [of "- \<bar>a\<bar>" 0 "\<bar>a\<bar>"])  | 
| 25307 | 1157  | 
|
1158  | 
lemma abs_eq_0 [simp]: "\<bar>a\<bar> = 0 \<longleftrightarrow> a = 0"  | 
|
1159  | 
proof -  | 
|
1160  | 
have "\<bar>a\<bar> = 0 \<Longrightarrow> a = 0"  | 
|
1161  | 
proof (rule antisym)  | 
|
1162  | 
assume zero: "\<bar>a\<bar> = 0"  | 
|
1163  | 
with abs_ge_self show "a \<le> 0" by auto  | 
|
1164  | 
from zero have "\<bar>-a\<bar> = 0" by simp  | 
|
| 36302 | 1165  | 
with abs_ge_self [of "- a"] have "- a \<le> 0" by auto  | 
| 25307 | 1166  | 
with neg_le_0_iff_le show "0 \<le> a" by auto  | 
1167  | 
qed  | 
|
1168  | 
then show ?thesis by auto  | 
|
1169  | 
qed  | 
|
1170  | 
||
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1171  | 
lemma abs_zero [simp]: "\<bar>0\<bar> = 0"  | 
| 29667 | 1172  | 
by simp  | 
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
1173  | 
|
| 
35828
 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 
blanchet 
parents: 
35723 
diff
changeset
 | 
1174  | 
lemma abs_0_eq [simp, no_atp]: "0 = \<bar>a\<bar> \<longleftrightarrow> a = 0"  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1175  | 
proof -  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1176  | 
have "0 = \<bar>a\<bar> \<longleftrightarrow> \<bar>a\<bar> = 0" by (simp only: eq_ac)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1177  | 
thus ?thesis by simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1178  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1179  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1180  | 
lemma abs_le_zero_iff [simp]: "\<bar>a\<bar> \<le> 0 \<longleftrightarrow> a = 0"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1181  | 
proof  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1182  | 
assume "\<bar>a\<bar> \<le> 0"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1183  | 
then have "\<bar>a\<bar> = 0" by (rule antisym) simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1184  | 
thus "a = 0" by simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1185  | 
next  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1186  | 
assume "a = 0"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1187  | 
thus "\<bar>a\<bar> \<le> 0" by simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1188  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1189  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1190  | 
lemma zero_less_abs_iff [simp]: "0 < \<bar>a\<bar> \<longleftrightarrow> a \<noteq> 0"  | 
| 29667 | 1191  | 
by (simp add: less_le)  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1192  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1193  | 
lemma abs_not_less_zero [simp]: "\<not> \<bar>a\<bar> < 0"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1194  | 
proof -  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1195  | 
have a: "\<And>x y. x \<le> y \<Longrightarrow> \<not> y < x" by auto  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1196  | 
show ?thesis by (simp add: a)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1197  | 
qed  | 
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
1198  | 
|
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1199  | 
lemma abs_ge_minus_self: "- a \<le> \<bar>a\<bar>"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1200  | 
proof -  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1201  | 
have "- a \<le> \<bar>-a\<bar>" by (rule abs_ge_self)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1202  | 
then show ?thesis by simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1203  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1204  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1205  | 
lemma abs_minus_commute:  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1206  | 
"\<bar>a - b\<bar> = \<bar>b - a\<bar>"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1207  | 
proof -  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1208  | 
have "\<bar>a - b\<bar> = \<bar>- (a - b)\<bar>" by (simp only: abs_minus_cancel)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1209  | 
also have "... = \<bar>b - a\<bar>" by simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1210  | 
finally show ?thesis .  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1211  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1212  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1213  | 
lemma abs_of_pos: "0 < a \<Longrightarrow> \<bar>a\<bar> = a"  | 
| 29667 | 1214  | 
by (rule abs_of_nonneg, rule less_imp_le)  | 
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
1215  | 
|
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1216  | 
lemma abs_of_nonpos [simp]:  | 
| 29667 | 1217  | 
assumes "a \<le> 0" shows "\<bar>a\<bar> = - a"  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1218  | 
proof -  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1219  | 
let ?b = "- a"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1220  | 
have "- ?b \<le> 0 \<Longrightarrow> \<bar>- ?b\<bar> = - (- ?b)"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1221  | 
unfolding abs_minus_cancel [of "?b"]  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1222  | 
unfolding neg_le_0_iff_le [of "?b"]  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1223  | 
unfolding minus_minus by (erule abs_of_nonneg)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1224  | 
then show ?thesis using assms by auto  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1225  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1226  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1227  | 
lemma abs_of_neg: "a < 0 \<Longrightarrow> \<bar>a\<bar> = - a"  | 
| 29667 | 1228  | 
by (rule abs_of_nonpos, rule less_imp_le)  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1229  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1230  | 
lemma abs_le_D1: "\<bar>a\<bar> \<le> b \<Longrightarrow> a \<le> b"  | 
| 29667 | 1231  | 
by (insert abs_ge_self, blast intro: order_trans)  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1232  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1233  | 
lemma abs_le_D2: "\<bar>a\<bar> \<le> b \<Longrightarrow> - a \<le> b"  | 
| 36302 | 1234  | 
by (insert abs_le_D1 [of "- a"], simp)  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1235  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1236  | 
lemma abs_le_iff: "\<bar>a\<bar> \<le> b \<longleftrightarrow> a \<le> b \<and> - a \<le> b"  | 
| 29667 | 1237  | 
by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2)  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1238  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1239  | 
lemma abs_triangle_ineq2: "\<bar>a\<bar> - \<bar>b\<bar> \<le> \<bar>a - b\<bar>"  | 
| 36302 | 1240  | 
proof -  | 
1241  | 
have "\<bar>a\<bar> = \<bar>b + (a - b)\<bar>"  | 
|
1242  | 
by (simp add: algebra_simps add_diff_cancel)  | 
|
1243  | 
then have "\<bar>a\<bar> \<le> \<bar>b\<bar> + \<bar>a - b\<bar>"  | 
|
1244  | 
by (simp add: abs_triangle_ineq)  | 
|
1245  | 
then show ?thesis  | 
|
1246  | 
by (simp add: algebra_simps)  | 
|
1247  | 
qed  | 
|
1248  | 
||
1249  | 
lemma abs_triangle_ineq2_sym: "\<bar>a\<bar> - \<bar>b\<bar> \<le> \<bar>b - a\<bar>"  | 
|
1250  | 
by (simp only: abs_minus_commute [of b] abs_triangle_ineq2)  | 
|
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
1251  | 
|
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1252  | 
lemma abs_triangle_ineq3: "\<bar>\<bar>a\<bar> - \<bar>b\<bar>\<bar> \<le> \<bar>a - b\<bar>"  | 
| 36302 | 1253  | 
by (simp add: abs_le_iff abs_triangle_ineq2 abs_triangle_ineq2_sym)  | 
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
1254  | 
|
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1255  | 
lemma abs_triangle_ineq4: "\<bar>a - b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1256  | 
proof -  | 
| 36302 | 1257  | 
have "\<bar>a - b\<bar> = \<bar>a + - b\<bar>" by (subst diff_minus, rule refl)  | 
1258  | 
also have "... \<le> \<bar>a\<bar> + \<bar>- b\<bar>" by (rule abs_triangle_ineq)  | 
|
| 29667 | 1259  | 
finally show ?thesis by simp  | 
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1260  | 
qed  | 
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
1261  | 
|
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1262  | 
lemma abs_diff_triangle_ineq: "\<bar>a + b - (c + d)\<bar> \<le> \<bar>a - c\<bar> + \<bar>b - d\<bar>"  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1263  | 
proof -  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1264  | 
have "\<bar>a + b - (c+d)\<bar> = \<bar>(a-c) + (b-d)\<bar>" by (simp add: diff_minus add_ac)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1265  | 
also have "... \<le> \<bar>a-c\<bar> + \<bar>b-d\<bar>" by (rule abs_triangle_ineq)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1266  | 
finally show ?thesis .  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1267  | 
qed  | 
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16417 
diff
changeset
 | 
1268  | 
|
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1269  | 
lemma abs_add_abs [simp]:  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1270  | 
"\<bar>\<bar>a\<bar> + \<bar>b\<bar>\<bar> = \<bar>a\<bar> + \<bar>b\<bar>" (is "?L = ?R")  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1271  | 
proof (rule antisym)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1272  | 
show "?L \<ge> ?R" by(rule abs_ge_self)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1273  | 
next  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1274  | 
have "?L \<le> \<bar>\<bar>a\<bar>\<bar> + \<bar>\<bar>b\<bar>\<bar>" by(rule abs_triangle_ineq)  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1275  | 
also have "\<dots> = ?R" by simp  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1276  | 
finally show "?L \<le> ?R" .  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1277  | 
qed  | 
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1278  | 
|
| 
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
25267 
diff
changeset
 | 
1279  | 
end  | 
| 14738 | 1280  | 
|
| 15178 | 1281  | 
|
| 25090 | 1282  | 
subsection {* Tools setup *}
 | 
1283  | 
||
| 
35828
 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 
blanchet 
parents: 
35723 
diff
changeset
 | 
1284  | 
lemma add_mono_thms_linordered_semiring [no_atp]:  | 
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
1285  | 
fixes i j k :: "'a\<Colon>ordered_ab_semigroup_add"  | 
| 25077 | 1286  | 
shows "i \<le> j \<and> k \<le> l \<Longrightarrow> i + k \<le> j + l"  | 
1287  | 
and "i = j \<and> k \<le> l \<Longrightarrow> i + k \<le> j + l"  | 
|
1288  | 
and "i \<le> j \<and> k = l \<Longrightarrow> i + k \<le> j + l"  | 
|
1289  | 
and "i = j \<and> k = l \<Longrightarrow> i + k = j + l"  | 
|
1290  | 
by (rule add_mono, clarify+)+  | 
|
1291  | 
||
| 
35828
 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 
blanchet 
parents: 
35723 
diff
changeset
 | 
1292  | 
lemma add_mono_thms_linordered_field [no_atp]:  | 
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34973 
diff
changeset
 | 
1293  | 
fixes i j k :: "'a\<Colon>ordered_cancel_ab_semigroup_add"  | 
| 25077 | 1294  | 
shows "i < j \<and> k = l \<Longrightarrow> i + k < j + l"  | 
1295  | 
and "i = j \<and> k < l \<Longrightarrow> i + k < j + l"  | 
|
1296  | 
and "i < j \<and> k \<le> l \<Longrightarrow> i + k < j + l"  | 
|
1297  | 
and "i \<le> j \<and> k < l \<Longrightarrow> i + k < j + l"  | 
|
1298  | 
and "i < j \<and> k < l \<Longrightarrow> i + k < j + l"  | 
|
1299  | 
by (auto intro: add_strict_right_mono add_strict_left_mono  | 
|
1300  | 
add_less_le_mono add_le_less_mono add_strict_mono)  | 
|
1301  | 
||
| 33364 | 1302  | 
code_modulename SML  | 
| 
35050
 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 
haftmann 
parents: 
35036 
diff
changeset
 | 
1303  | 
Groups Arith  | 
| 33364 | 1304  | 
|
1305  | 
code_modulename OCaml  | 
|
| 
35050
 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 
haftmann 
parents: 
35036 
diff
changeset
 | 
1306  | 
Groups Arith  | 
| 33364 | 1307  | 
|
1308  | 
code_modulename Haskell  | 
|
| 
35050
 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 
haftmann 
parents: 
35036 
diff
changeset
 | 
1309  | 
Groups Arith  | 
| 33364 | 1310  | 
|
| 
37889
 
0d8058e0c270
keep explicit diff_def as legacy theorem; modernized abel_cancel simproc setup
 
haftmann 
parents: 
37884 
diff
changeset
 | 
1311  | 
|
| 
 
0d8058e0c270
keep explicit diff_def as legacy theorem; modernized abel_cancel simproc setup
 
haftmann 
parents: 
37884 
diff
changeset
 | 
1312  | 
text {* Legacy *}
 | 
| 
 
0d8058e0c270
keep explicit diff_def as legacy theorem; modernized abel_cancel simproc setup
 
haftmann 
parents: 
37884 
diff
changeset
 | 
1313  | 
|
| 
 
0d8058e0c270
keep explicit diff_def as legacy theorem; modernized abel_cancel simproc setup
 
haftmann 
parents: 
37884 
diff
changeset
 | 
1314  | 
lemmas diff_def = diff_minus  | 
| 
 
0d8058e0c270
keep explicit diff_def as legacy theorem; modernized abel_cancel simproc setup
 
haftmann 
parents: 
37884 
diff
changeset
 | 
1315  | 
|
| 14738 | 1316  | 
end  | 
| 49388 | 1317  |