| author | wenzelm | 
| Fri, 14 Mar 2025 23:03:58 +0100 | |
| changeset 82276 | d22e9c5b5dc6 | 
| parent 81125 | ec121999a9cb | 
| permissions | -rw-r--r-- | 
| 63588 | 1 | (* Title: HOL/Groups.thy | 
| 2 | Author: Gertrud Bauer | |
| 3 | Author: Steven Obua | |
| 4 | Author: Lawrence C Paulson | |
| 5 | Author: Markus Wenzel | |
| 6 | Author: Jeremy Avigad | |
| 14738 | 7 | *) | 
| 8 | ||
| 60758 | 9 | section \<open>Groups, also combined with orderings\<close> | 
| 14738 | 10 | |
| 35050 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
35036diff
changeset | 11 | theory Groups | 
| 63588 | 12 | imports Orderings | 
| 15131 | 13 | begin | 
| 14738 | 14 | |
| 60758 | 15 | subsection \<open>Dynamic facts\<close> | 
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 16 | |
| 57950 | 17 | named_theorems ac_simps "associativity and commutativity simplification rules" | 
| 70347 | 18 | and algebra_simps "algebra simplification rules for rings" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 19 | and algebra_split_simps "algebra simplification rules for rings, with potential goal splitting" | 
| 63588 | 20 | and field_simps "algebra simplification rules for fields" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 21 | and field_split_simps "algebra simplification rules for fields, with potential goal splitting" | 
| 36343 | 22 | |
| 63325 | 23 | text \<open> | 
| 63588 | 24 | The rewrites accumulated in \<open>algebra_simps\<close> deal with the classical | 
| 25 | algebraic structures of groups, rings and family. They simplify terms by | |
| 26 | multiplying everything out (in case of a ring) and bringing sums and | |
| 27 | products into a canonical form (by ordered rewriting). As a result it | |
| 28 | decides group and ring equalities but also helps with inequalities. | |
| 36348 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 haftmann parents: 
36343diff
changeset | 29 | |
| 63588 | 30 | Of course it also works for fields, but it knows nothing about | 
| 31 | multiplicative inverses or division. This is catered for by \<open>field_simps\<close>. | |
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 32 | |
| 63588 | 33 | Facts in \<open>field_simps\<close> multiply with denominators in (in)equations if they | 
| 34 | can be proved to be non-zero (for equations) or positive/negative (for | |
| 35 | inequalities). Can be too aggressive and is therefore separate from the more | |
| 36 | benign \<open>algebra_simps\<close>. | |
| 70347 | 37 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 38 | Collections \<open>algebra_split_simps\<close> and \<open>field_split_simps\<close> | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 39 | correspond to \<open>algebra_simps\<close> and \<open>field_simps\<close> | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 40 | but contain more aggresive rules that may lead to goal splitting. | 
| 63325 | 41 | \<close> | 
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 42 | |
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 43 | |
| 60758 | 44 | subsection \<open>Abstract structures\<close> | 
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 45 | |
| 60758 | 46 | text \<open> | 
| 63588 | 47 | These locales provide basic structures for interpretation into bigger | 
| 48 | structures; extensions require careful thinking, otherwise undesired effects | |
| 49 | may occur due to interpretation. | |
| 60758 | 50 | \<close> | 
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 51 | |
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 52 | locale semigroup = | 
| 80932 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 wenzelm parents: 
80097diff
changeset | 53 | fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>\<^bold>*\<close> 70) | 
| 63290 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 haftmann parents: 
63145diff
changeset | 54 | assumes assoc [ac_simps]: "a \<^bold>* b \<^bold>* c = a \<^bold>* (b \<^bold>* c)" | 
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 55 | |
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 56 | locale abel_semigroup = semigroup + | 
| 63290 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 haftmann parents: 
63145diff
changeset | 57 | assumes commute [ac_simps]: "a \<^bold>* b = b \<^bold>* a" | 
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 58 | begin | 
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 59 | |
| 63325 | 60 | lemma left_commute [ac_simps]: "b \<^bold>* (a \<^bold>* c) = a \<^bold>* (b \<^bold>* c)" | 
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 61 | proof - | 
| 63290 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 haftmann parents: 
63145diff
changeset | 62 | have "(b \<^bold>* a) \<^bold>* c = (a \<^bold>* b) \<^bold>* c" | 
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 63 | by (simp only: commute) | 
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 64 | then show ?thesis | 
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 65 | by (simp only: assoc) | 
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 66 | qed | 
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 67 | |
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 68 | end | 
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 69 | |
| 35720 | 70 | locale monoid = semigroup + | 
| 80932 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 wenzelm parents: 
80097diff
changeset | 71 | fixes z :: 'a (\<open>\<^bold>1\<close>) | 
| 63290 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 haftmann parents: 
63145diff
changeset | 72 | assumes left_neutral [simp]: "\<^bold>1 \<^bold>* a = a" | 
| 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 haftmann parents: 
63145diff
changeset | 73 | assumes right_neutral [simp]: "a \<^bold>* \<^bold>1 = a" | 
| 35720 | 74 | |
| 75 | locale comm_monoid = abel_semigroup + | |
| 80932 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 wenzelm parents: 
80097diff
changeset | 76 | fixes z :: 'a (\<open>\<^bold>1\<close>) | 
| 63290 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 haftmann parents: 
63145diff
changeset | 77 | assumes comm_neutral: "a \<^bold>* \<^bold>1 = a" | 
| 54868 | 78 | begin | 
| 35720 | 79 | |
| 54868 | 80 | sublocale monoid | 
| 61169 | 81 | by standard (simp_all add: commute comm_neutral) | 
| 35720 | 82 | |
| 54868 | 83 | end | 
| 84 | ||
| 63364 | 85 | locale group = semigroup + | 
| 80932 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 wenzelm parents: 
80097diff
changeset | 86 | fixes z :: 'a (\<open>\<^bold>1\<close>) | 
| 63364 | 87 | fixes inverse :: "'a \<Rightarrow> 'a" | 
| 88 | assumes group_left_neutral: "\<^bold>1 \<^bold>* a = a" | |
| 89 | assumes left_inverse [simp]: "inverse a \<^bold>* a = \<^bold>1" | |
| 90 | begin | |
| 91 | ||
| 92 | lemma left_cancel: "a \<^bold>* b = a \<^bold>* c \<longleftrightarrow> b = c" | |
| 93 | proof | |
| 94 | assume "a \<^bold>* b = a \<^bold>* c" | |
| 95 | then have "inverse a \<^bold>* (a \<^bold>* b) = inverse a \<^bold>* (a \<^bold>* c)" by simp | |
| 96 | then have "(inverse a \<^bold>* a) \<^bold>* b = (inverse a \<^bold>* a) \<^bold>* c" | |
| 97 | by (simp only: assoc) | |
| 98 | then show "b = c" by (simp add: group_left_neutral) | |
| 99 | qed simp | |
| 100 | ||
| 101 | sublocale monoid | |
| 102 | proof | |
| 103 | fix a | |
| 104 | have "inverse a \<^bold>* a = \<^bold>1" by simp | |
| 105 | then have "inverse a \<^bold>* (a \<^bold>* \<^bold>1) = inverse a \<^bold>* a" | |
| 106 | by (simp add: group_left_neutral assoc [symmetric]) | |
| 107 | with left_cancel show "a \<^bold>* \<^bold>1 = a" | |
| 108 | by (simp only: left_cancel) | |
| 109 | qed (fact group_left_neutral) | |
| 110 | ||
| 111 | lemma inverse_unique: | |
| 112 | assumes "a \<^bold>* b = \<^bold>1" | |
| 113 | shows "inverse a = b" | |
| 114 | proof - | |
| 115 | from assms have "inverse a \<^bold>* (a \<^bold>* b) = inverse a" | |
| 116 | by simp | |
| 117 | then show ?thesis | |
| 118 | by (simp add: assoc [symmetric]) | |
| 119 | qed | |
| 120 | ||
| 63588 | 121 | lemma inverse_neutral [simp]: "inverse \<^bold>1 = \<^bold>1" | 
| 63364 | 122 | by (rule inverse_unique) simp | 
| 123 | ||
| 63588 | 124 | lemma inverse_inverse [simp]: "inverse (inverse a) = a" | 
| 63364 | 125 | by (rule inverse_unique) simp | 
| 126 | ||
| 63588 | 127 | lemma right_inverse [simp]: "a \<^bold>* inverse a = \<^bold>1" | 
| 63364 | 128 | proof - | 
| 129 | have "a \<^bold>* inverse a = inverse (inverse a) \<^bold>* inverse a" | |
| 130 | by simp | |
| 131 | also have "\<dots> = \<^bold>1" | |
| 132 | by (rule left_inverse) | |
| 133 | then show ?thesis by simp | |
| 134 | qed | |
| 135 | ||
| 63588 | 136 | lemma inverse_distrib_swap: "inverse (a \<^bold>* b) = inverse b \<^bold>* inverse a" | 
| 63364 | 137 | proof (rule inverse_unique) | 
| 138 | have "a \<^bold>* b \<^bold>* (inverse b \<^bold>* inverse a) = | |
| 139 | a \<^bold>* (b \<^bold>* inverse b) \<^bold>* inverse a" | |
| 140 | by (simp only: assoc) | |
| 141 | also have "\<dots> = \<^bold>1" | |
| 142 | by simp | |
| 143 | finally show "a \<^bold>* b \<^bold>* (inverse b \<^bold>* inverse a) = \<^bold>1" . | |
| 144 | qed | |
| 145 | ||
| 63588 | 146 | lemma right_cancel: "b \<^bold>* a = c \<^bold>* a \<longleftrightarrow> b = c" | 
| 63364 | 147 | proof | 
| 148 | assume "b \<^bold>* a = c \<^bold>* a" | |
| 149 | then have "b \<^bold>* a \<^bold>* inverse a= c \<^bold>* a \<^bold>* inverse a" | |
| 150 | by simp | |
| 151 | then show "b = c" | |
| 152 | by (simp add: assoc) | |
| 153 | qed simp | |
| 154 | ||
| 155 | end | |
| 156 | ||
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 157 | |
| 60758 | 158 | subsection \<open>Generic operations\<close> | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 159 | |
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 160 | class zero = | 
| 80932 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 wenzelm parents: 
80097diff
changeset | 161 | fixes zero :: 'a (\<open>0\<close>) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 162 | |
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 163 | class one = | 
| 80932 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 wenzelm parents: 
80097diff
changeset | 164 | fixes one :: 'a (\<open>1\<close>) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 165 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35828diff
changeset | 166 | hide_const (open) zero one | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 167 | |
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 168 | lemma Let_0 [simp]: "Let 0 f = f 0" | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 169 | unfolding Let_def .. | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 170 | |
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 171 | lemma Let_1 [simp]: "Let 1 f = f 1" | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 172 | unfolding Let_def .. | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 173 | |
| 60758 | 174 | setup \<open> | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 175 | Reorient_Proc.add | 
| 69593 | 176 | (fn Const(\<^const_name>\<open>Groups.zero\<close>, _) => true | 
| 177 | | Const(\<^const_name>\<open>Groups.one\<close>, _) => true | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 178 | | _ => false) | 
| 60758 | 179 | \<close> | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 180 | |
| 78099 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 wenzelm parents: 
76054diff
changeset | 181 | simproc_setup reorient_zero ("0 = x") = \<open>K Reorient_Proc.proc\<close>
 | 
| 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 wenzelm parents: 
76054diff
changeset | 182 | simproc_setup reorient_one ("1 = x") = \<open>K Reorient_Proc.proc\<close>
 | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 183 | |
| 60758 | 184 | typed_print_translation \<open> | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 185 | let | 
| 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 186 | fun tr' c = (c, fn ctxt => fn T => fn ts => | 
| 52210 
0226035df99d
more explicit Printer.type_emphasis, depending on show_type_emphasis;
 wenzelm parents: 
52143diff
changeset | 187 | if null ts andalso Printer.type_emphasis ctxt T then | 
| 69593 | 188 | Syntax.const \<^syntax_const>\<open>_constrain\<close> $ Syntax.const c $ | 
| 52210 
0226035df99d
more explicit Printer.type_emphasis, depending on show_type_emphasis;
 wenzelm parents: 
52143diff
changeset | 189 | Syntax_Phases.term_of_typ ctxt T | 
| 
0226035df99d
more explicit Printer.type_emphasis, depending on show_type_emphasis;
 wenzelm parents: 
52143diff
changeset | 190 | else raise Match); | 
| 69593 | 191 | in map tr' [\<^const_syntax>\<open>Groups.one\<close>, \<^const_syntax>\<open>Groups.zero\<close>] end | 
| 61799 | 192 | \<close> \<comment> \<open>show types that are presumably too general\<close> | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 193 | |
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 194 | class plus = | 
| 80932 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 wenzelm parents: 
80097diff
changeset | 195 | fixes plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>+\<close> 65) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 196 | |
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 197 | class minus = | 
| 80932 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 wenzelm parents: 
80097diff
changeset | 198 | fixes minus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>-\<close> 65) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 199 | |
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 200 | class uminus = | 
| 81125 | 201 | fixes uminus :: "'a \<Rightarrow> 'a" (\<open>(\<open>open_block notation=\<open>prefix -\<close>\<close>- _)\<close> [81] 80) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 202 | |
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 203 | class times = | 
| 80932 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 wenzelm parents: 
80097diff
changeset | 204 | fixes times :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>*\<close> 70) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35216diff
changeset | 205 | |
| 81125 | 206 | bundle uminus_syntax | 
| 207 | begin | |
| 208 | notation uminus (\<open>(\<open>open_block notation=\<open>prefix -\<close>\<close>- _)\<close> [81] 80) | |
| 209 | end | |
| 210 | ||
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35050diff
changeset | 211 | |
| 60758 | 212 | subsection \<open>Semigroups and Monoids\<close> | 
| 14738 | 213 | |
| 22390 | 214 | class semigroup_add = plus + | 
| 80097 
5ed992c47cdc
prefer canonical theorem name for fact collection declarations
 haftmann parents: 
78099diff
changeset | 215 | assumes add_assoc: "(a + b) + c = a + (b + c)" | 
| 54868 | 216 | begin | 
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34147diff
changeset | 217 | |
| 61605 | 218 | sublocale add: semigroup plus | 
| 61169 | 219 | by standard (fact add_assoc) | 
| 22390 | 220 | |
| 80097 
5ed992c47cdc
prefer canonical theorem name for fact collection declarations
 haftmann parents: 
78099diff
changeset | 221 | declare add.assoc [algebra_simps, algebra_split_simps, field_simps, field_split_simps] | 
| 
5ed992c47cdc
prefer canonical theorem name for fact collection declarations
 haftmann parents: 
78099diff
changeset | 222 | |
| 54868 | 223 | end | 
| 224 | ||
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 225 | hide_fact add_assoc | 
| 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 226 | |
| 22390 | 227 | class ab_semigroup_add = semigroup_add + | 
| 80097 
5ed992c47cdc
prefer canonical theorem name for fact collection declarations
 haftmann parents: 
78099diff
changeset | 228 | assumes add_commute: "a + b = b + a" | 
| 54868 | 229 | begin | 
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34147diff
changeset | 230 | |
| 61605 | 231 | sublocale add: abel_semigroup plus | 
| 61169 | 232 | by standard (fact add_commute) | 
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34147diff
changeset | 233 | |
| 80097 
5ed992c47cdc
prefer canonical theorem name for fact collection declarations
 haftmann parents: 
78099diff
changeset | 234 | declare add.commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps] | 
| 
5ed992c47cdc
prefer canonical theorem name for fact collection declarations
 haftmann parents: 
78099diff
changeset | 235 | add.left_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps] | 
| 25062 | 236 | |
| 61337 | 237 | lemmas add_ac = add.assoc add.commute add.left_commute | 
| 57571 
d38a98f496dd
reinstated popular add_ac and mult_ac to avoid needless incompatibilities in user space
 nipkow parents: 
57514diff
changeset | 238 | |
| 25062 | 239 | end | 
| 14738 | 240 | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 241 | hide_fact add_commute | 
| 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 242 | |
| 61337 | 243 | lemmas add_ac = add.assoc add.commute add.left_commute | 
| 57571 
d38a98f496dd
reinstated popular add_ac and mult_ac to avoid needless incompatibilities in user space
 nipkow parents: 
57514diff
changeset | 244 | |
| 22390 | 245 | class semigroup_mult = times + | 
| 80097 
5ed992c47cdc
prefer canonical theorem name for fact collection declarations
 haftmann parents: 
78099diff
changeset | 246 | assumes mult_assoc: "(a * b) * c = a * (b * c)" | 
| 54868 | 247 | begin | 
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34147diff
changeset | 248 | |
| 61605 | 249 | sublocale mult: semigroup times | 
| 61169 | 250 | by standard (fact mult_assoc) | 
| 14738 | 251 | |
| 80097 
5ed992c47cdc
prefer canonical theorem name for fact collection declarations
 haftmann parents: 
78099diff
changeset | 252 | declare mult.assoc [algebra_simps, algebra_split_simps, field_simps, field_split_simps] | 
| 
5ed992c47cdc
prefer canonical theorem name for fact collection declarations
 haftmann parents: 
78099diff
changeset | 253 | |
| 54868 | 254 | end | 
| 255 | ||
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 256 | hide_fact mult_assoc | 
| 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 257 | |
| 22390 | 258 | class ab_semigroup_mult = semigroup_mult + | 
| 80097 
5ed992c47cdc
prefer canonical theorem name for fact collection declarations
 haftmann parents: 
78099diff
changeset | 259 | assumes mult_commute: "a * b = b * a" | 
| 54868 | 260 | begin | 
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34147diff
changeset | 261 | |
| 61605 | 262 | sublocale mult: abel_semigroup times | 
| 61169 | 263 | by standard (fact mult_commute) | 
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34147diff
changeset | 264 | |
| 80097 
5ed992c47cdc
prefer canonical theorem name for fact collection declarations
 haftmann parents: 
78099diff
changeset | 265 | declare mult.commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps] | 
| 
5ed992c47cdc
prefer canonical theorem name for fact collection declarations
 haftmann parents: 
78099diff
changeset | 266 | mult.left_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps] | 
| 25062 | 267 | |
| 61337 | 268 | lemmas mult_ac = mult.assoc mult.commute mult.left_commute | 
| 57571 
d38a98f496dd
reinstated popular add_ac and mult_ac to avoid needless incompatibilities in user space
 nipkow parents: 
57514diff
changeset | 269 | |
| 23181 | 270 | end | 
| 14738 | 271 | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 272 | hide_fact mult_commute | 
| 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 273 | |
| 61337 | 274 | lemmas mult_ac = mult.assoc mult.commute mult.left_commute | 
| 57571 
d38a98f496dd
reinstated popular add_ac and mult_ac to avoid needless incompatibilities in user space
 nipkow parents: 
57514diff
changeset | 275 | |
| 23085 | 276 | class monoid_add = zero + semigroup_add + | 
| 35720 | 277 | assumes add_0_left: "0 + a = a" | 
| 278 | and add_0_right: "a + 0 = a" | |
| 54868 | 279 | begin | 
| 35720 | 280 | |
| 61605 | 281 | sublocale add: monoid plus 0 | 
| 61169 | 282 | by standard (fact add_0_left add_0_right)+ | 
| 23085 | 283 | |
| 54868 | 284 | end | 
| 285 | ||
| 26071 | 286 | lemma zero_reorient: "0 = x \<longleftrightarrow> x = 0" | 
| 54868 | 287 | by (fact eq_commute) | 
| 26071 | 288 | |
| 22390 | 289 | class comm_monoid_add = zero + ab_semigroup_add + | 
| 25062 | 290 | assumes add_0: "0 + a = a" | 
| 54868 | 291 | begin | 
| 23085 | 292 | |
| 59559 
35da1bbf234e
more canonical order of subscriptions avoids superfluous facts
 haftmann parents: 
59557diff
changeset | 293 | subclass monoid_add | 
| 61169 | 294 | by standard (simp_all add: add_0 add.commute [of _ 0]) | 
| 25062 | 295 | |
| 61605 | 296 | sublocale add: comm_monoid plus 0 | 
| 61169 | 297 | by standard (simp add: ac_simps) | 
| 14738 | 298 | |
| 54868 | 299 | end | 
| 300 | ||
| 22390 | 301 | class monoid_mult = one + semigroup_mult + | 
| 35720 | 302 | assumes mult_1_left: "1 * a = a" | 
| 303 | and mult_1_right: "a * 1 = a" | |
| 54868 | 304 | begin | 
| 35720 | 305 | |
| 61605 | 306 | sublocale mult: monoid times 1 | 
| 61169 | 307 | by standard (fact mult_1_left mult_1_right)+ | 
| 14738 | 308 | |
| 54868 | 309 | end | 
| 310 | ||
| 26071 | 311 | lemma one_reorient: "1 = x \<longleftrightarrow> x = 1" | 
| 54868 | 312 | by (fact eq_commute) | 
| 26071 | 313 | |
| 22390 | 314 | class comm_monoid_mult = one + ab_semigroup_mult + | 
| 25062 | 315 | assumes mult_1: "1 * a = a" | 
| 54868 | 316 | begin | 
| 14738 | 317 | |
| 59559 
35da1bbf234e
more canonical order of subscriptions avoids superfluous facts
 haftmann parents: 
59557diff
changeset | 318 | subclass monoid_mult | 
| 61169 | 319 | by standard (simp_all add: mult_1 mult.commute [of _ 1]) | 
| 25062 | 320 | |
| 61605 | 321 | sublocale mult: comm_monoid times 1 | 
| 61169 | 322 | by standard (simp add: ac_simps) | 
| 14738 | 323 | |
| 54868 | 324 | end | 
| 325 | ||
| 22390 | 326 | class cancel_semigroup_add = semigroup_add + | 
| 25062 | 327 | assumes add_left_imp_eq: "a + b = a + c \<Longrightarrow> b = c" | 
| 328 | 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: 
27250diff
changeset | 329 | begin | 
| 
a89d755b029d
move proofs of add_left_cancel and add_right_cancel into the correct locale
 huffman parents: 
27250diff
changeset | 330 | |
| 63325 | 331 | lemma add_left_cancel [simp]: "a + b = a + c \<longleftrightarrow> b = c" | 
| 332 | 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: 
27250diff
changeset | 333 | |
| 63325 | 334 | lemma add_right_cancel [simp]: "b + a = c + a \<longleftrightarrow> b = c" | 
| 335 | 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: 
27250diff
changeset | 336 | |
| 
a89d755b029d
move proofs of add_left_cancel and add_right_cancel into the correct locale
 huffman parents: 
27250diff
changeset | 337 | end | 
| 14738 | 338 | |
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 339 | class cancel_ab_semigroup_add = ab_semigroup_add + minus + | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 340 | assumes add_diff_cancel_left' [simp]: "(a + b) - a = b" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 341 | assumes diff_diff_add [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 342 | "a - b - c = a - (b + c)" | 
| 25267 | 343 | begin | 
| 14738 | 344 | |
| 63325 | 345 | lemma add_diff_cancel_right' [simp]: "(a + b) - b = a" | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 346 | using add_diff_cancel_left' [of b a] by (simp add: ac_simps) | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 347 | |
| 25267 | 348 | subclass cancel_semigroup_add | 
| 28823 | 349 | proof | 
| 22390 | 350 | fix a b c :: 'a | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 351 | assume "a + b = a + c" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 352 | then have "a + b - a = a + c - a" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 353 | by simp | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 354 | then show "b = c" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 355 | by simp | 
| 22390 | 356 | next | 
| 14738 | 357 | fix a b c :: 'a | 
| 358 | assume "b + a = c + a" | |
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 359 | then have "b + a - a = c + a - a" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 360 | by simp | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 361 | then show "b = c" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 362 | by simp | 
| 14738 | 363 | qed | 
| 364 | ||
| 63325 | 365 | lemma add_diff_cancel_left [simp]: "(c + a) - (c + b) = a - b" | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 366 | unfolding diff_diff_add [symmetric] by simp | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 367 | |
| 63325 | 368 | lemma add_diff_cancel_right [simp]: "(a + c) - (b + c) = a - b" | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 369 | using add_diff_cancel_left [symmetric] by (simp add: ac_simps) | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 370 | |
| 63325 | 371 | lemma diff_right_commute: "a - c - b = a - b - c" | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 372 | by (simp add: diff_diff_add add.commute) | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 373 | |
| 25267 | 374 | end | 
| 375 | ||
| 29904 | 376 | class cancel_comm_monoid_add = cancel_ab_semigroup_add + comm_monoid_add | 
| 59322 | 377 | begin | 
| 378 | ||
| 63325 | 379 | lemma diff_zero [simp]: "a - 0 = a" | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 380 | using add_diff_cancel_right' [of a 0] by simp | 
| 59322 | 381 | |
| 63325 | 382 | lemma diff_cancel [simp]: "a - a = 0" | 
| 59322 | 383 | proof - | 
| 63325 | 384 | have "(a + 0) - (a + 0) = 0" | 
| 385 | by (simp only: add_diff_cancel_left diff_zero) | |
| 59322 | 386 | then show ?thesis by simp | 
| 387 | qed | |
| 388 | ||
| 389 | lemma add_implies_diff: | |
| 390 | assumes "c + b = a" | |
| 391 | shows "c = a - b" | |
| 392 | proof - | |
| 63325 | 393 | from assms have "(b + c) - (b + 0) = a - b" | 
| 394 | by (simp add: add.commute) | |
| 59322 | 395 | then show "c = a - b" by simp | 
| 396 | qed | |
| 397 | ||
| 63325 | 398 | lemma add_cancel_right_right [simp]: "a = a + b \<longleftrightarrow> b = 0" | 
| 399 | (is "?P \<longleftrightarrow> ?Q") | |
| 62608 | 400 | proof | 
| 63325 | 401 | assume ?Q | 
| 402 | then show ?P by simp | |
| 62608 | 403 | next | 
| 63325 | 404 | assume ?P | 
| 405 | then have "a - a = a + b - a" by simp | |
| 62608 | 406 | then show ?Q by simp | 
| 407 | qed | |
| 408 | ||
| 63325 | 409 | lemma add_cancel_right_left [simp]: "a = b + a \<longleftrightarrow> b = 0" | 
| 62608 | 410 | using add_cancel_right_right [of a b] by (simp add: ac_simps) | 
| 411 | ||
| 63325 | 412 | lemma add_cancel_left_right [simp]: "a + b = a \<longleftrightarrow> b = 0" | 
| 62608 | 413 | by (auto dest: sym) | 
| 414 | ||
| 63325 | 415 | lemma add_cancel_left_left [simp]: "b + a = a \<longleftrightarrow> b = 0" | 
| 62608 | 416 | by (auto dest: sym) | 
| 417 | ||
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 418 | end | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 419 | |
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 420 | class comm_monoid_diff = cancel_comm_monoid_add + | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 421 | assumes zero_diff [simp]: "0 - a = 0" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 422 | begin | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 423 | |
| 63325 | 424 | lemma diff_add_zero [simp]: "a - (a + b) = 0" | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 425 | proof - | 
| 63325 | 426 | have "a - (a + b) = (a + 0) - (a + b)" | 
| 427 | by simp | |
| 428 | also have "\<dots> = 0" | |
| 429 | by (simp only: add_diff_cancel_left zero_diff) | |
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 430 | finally show ?thesis . | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 431 | qed | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 432 | |
| 59322 | 433 | end | 
| 434 | ||
| 29904 | 435 | |
| 60758 | 436 | subsection \<open>Groups\<close> | 
| 23085 | 437 | |
| 25762 | 438 | class group_add = minus + uminus + monoid_add + | 
| 63364 | 439 | assumes left_minus: "- a + a = 0" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 440 | assumes add_uminus_conv_diff [simp]: "a + (- b) = a - b" | 
| 25062 | 441 | begin | 
| 23085 | 442 | |
| 63325 | 443 | lemma diff_conv_add_uminus: "a - b = a + (- b)" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 444 | by simp | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 445 | |
| 63364 | 446 | sublocale add: group plus 0 uminus | 
| 447 | by standard (simp_all add: left_minus) | |
| 448 | ||
| 63588 | 449 | lemma minus_unique: "a + b = 0 \<Longrightarrow> - a = b" | 
| 450 | by (fact add.inverse_unique) | |
| 34147 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 huffman parents: 
34146diff
changeset | 451 | |
| 63364 | 452 | lemma minus_zero: "- 0 = 0" | 
| 453 | by (fact add.inverse_neutral) | |
| 14738 | 454 | |
| 63364 | 455 | lemma minus_minus: "- (- a) = a" | 
| 456 | by (fact add.inverse_inverse) | |
| 14738 | 457 | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 458 | lemma right_minus: "a + - a = 0" | 
| 63364 | 459 | by (fact add.right_inverse) | 
| 14738 | 460 | |
| 63325 | 461 | lemma diff_self [simp]: "a - a = 0" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 462 | using right_minus [of a] by simp | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 463 | |
| 40368 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 haftmann parents: 
39134diff
changeset | 464 | subclass cancel_semigroup_add | 
| 63364 | 465 | by standard (simp_all add: add.left_cancel add.right_cancel) | 
| 40368 
47c186c8577d
added class relation group_add < cancel_semigroup_add
 haftmann parents: 
39134diff
changeset | 466 | |
| 63325 | 467 | lemma minus_add_cancel [simp]: "- a + (a + b) = b" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 468 | by (simp add: add.assoc [symmetric]) | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 469 | |
| 63325 | 470 | lemma add_minus_cancel [simp]: "a + (- a + b) = b" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 471 | by (simp add: add.assoc [symmetric]) | 
| 34147 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 huffman parents: 
34146diff
changeset | 472 | |
| 63325 | 473 | lemma diff_add_cancel [simp]: "a - b + b = a" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 474 | by (simp only: diff_conv_add_uminus add.assoc) simp | 
| 34147 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 huffman parents: 
34146diff
changeset | 475 | |
| 63325 | 476 | lemma add_diff_cancel [simp]: "a + b - b = a" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 477 | by (simp only: diff_conv_add_uminus add.assoc) simp | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 478 | |
| 63325 | 479 | lemma minus_add: "- (a + b) = - b + - a" | 
| 63364 | 480 | by (fact add.inverse_distrib_swap) | 
| 34147 
319616f4eecf
generalize lemma add_minus_cancel, add lemma minus_add, simplify some proofs
 huffman parents: 
34146diff
changeset | 481 | |
| 63325 | 482 | lemma right_minus_eq [simp]: "a - b = 0 \<longleftrightarrow> a = b" | 
| 14738 | 483 | proof | 
| 23085 | 484 | assume "a - b = 0" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 485 | have "a = (a - b) + b" by (simp add: add.assoc) | 
| 60758 | 486 | also have "\<dots> = b" using \<open>a - b = 0\<close> by simp | 
| 23085 | 487 | finally show "a = b" . | 
| 14738 | 488 | next | 
| 63325 | 489 | assume "a = b" | 
| 490 | then show "a - b = 0" by simp | |
| 14738 | 491 | qed | 
| 492 | ||
| 63325 | 493 | lemma eq_iff_diff_eq_0: "a = b \<longleftrightarrow> a - b = 0" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 494 | by (fact right_minus_eq [symmetric]) | 
| 14738 | 495 | |
| 63325 | 496 | lemma diff_0 [simp]: "0 - a = - a" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 497 | by (simp only: diff_conv_add_uminus add_0_left) | 
| 14738 | 498 | |
| 63325 | 499 | lemma diff_0_right [simp]: "a - 0 = a" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 500 | by (simp only: diff_conv_add_uminus minus_zero add_0_right) | 
| 14738 | 501 | |
| 63325 | 502 | lemma diff_minus_eq_add [simp]: "a - - b = a + b" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 503 | by (simp only: diff_conv_add_uminus minus_minus) | 
| 14738 | 504 | |
| 63325 | 505 | lemma neg_equal_iff_equal [simp]: "- a = - b \<longleftrightarrow> a = b" | 
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 506 | proof | 
| 14738 | 507 | assume "- a = - b" | 
| 63325 | 508 | then have "- (- a) = - (- b)" by simp | 
| 509 | then show "a = b" by simp | |
| 14738 | 510 | next | 
| 25062 | 511 | assume "a = b" | 
| 63325 | 512 | then show "- a = - b" by simp | 
| 14738 | 513 | qed | 
| 514 | ||
| 63325 | 515 | lemma neg_equal_0_iff_equal [simp]: "- a = 0 \<longleftrightarrow> a = 0" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 516 | by (subst neg_equal_iff_equal [symmetric]) simp | 
| 14738 | 517 | |
| 63325 | 518 | lemma neg_0_equal_iff_equal [simp]: "0 = - a \<longleftrightarrow> 0 = a" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 519 | by (subst neg_equal_iff_equal [symmetric]) simp | 
| 14738 | 520 | |
| 63325 | 521 | text \<open>The next two equations can make the simplifier loop!\<close> | 
| 14738 | 522 | |
| 63325 | 523 | lemma equation_minus_iff: "a = - b \<longleftrightarrow> b = - a" | 
| 14738 | 524 | proof - | 
| 63325 | 525 | have "- (- a) = - b \<longleftrightarrow> - a = b" | 
| 526 | by (rule neg_equal_iff_equal) | |
| 527 | then show ?thesis | |
| 528 | by (simp add: eq_commute) | |
| 25062 | 529 | qed | 
| 530 | ||
| 63325 | 531 | lemma minus_equation_iff: "- a = b \<longleftrightarrow> - b = a" | 
| 25062 | 532 | proof - | 
| 63325 | 533 | have "- a = - (- b) \<longleftrightarrow> a = -b" | 
| 534 | by (rule neg_equal_iff_equal) | |
| 535 | then show ?thesis | |
| 536 | by (simp add: eq_commute) | |
| 14738 | 537 | qed | 
| 538 | ||
| 63325 | 539 | lemma eq_neg_iff_add_eq_0: "a = - b \<longleftrightarrow> a + b = 0" | 
| 29914 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 huffman parents: 
29904diff
changeset | 540 | proof | 
| 63325 | 541 | assume "a = - b" | 
| 542 | then show "a + b = 0" by simp | |
| 29914 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 huffman parents: 
29904diff
changeset | 543 | next | 
| 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 huffman parents: 
29904diff
changeset | 544 | assume "a + b = 0" | 
| 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 huffman parents: 
29904diff
changeset | 545 | moreover have "a + (b + - b) = (a + b) + - b" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 546 | by (simp only: add.assoc) | 
| 63325 | 547 | ultimately show "a = - b" | 
| 548 | by simp | |
| 29914 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 huffman parents: 
29904diff
changeset | 549 | qed | 
| 
c9ced4f54e82
generalize lemma eq_neg_iff_add_eq_0, and move to OrderedGroup
 huffman parents: 
29904diff
changeset | 550 | |
| 63325 | 551 | lemma add_eq_0_iff2: "a + b = 0 \<longleftrightarrow> a = - b" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 552 | by (fact eq_neg_iff_add_eq_0 [symmetric]) | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 553 | |
| 63325 | 554 | lemma neg_eq_iff_add_eq_0: "- a = b \<longleftrightarrow> a + b = 0" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 555 | by (auto simp add: add_eq_0_iff2) | 
| 44348 | 556 | |
| 63325 | 557 | lemma add_eq_0_iff: "a + b = 0 \<longleftrightarrow> b = - a" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 558 | by (auto simp add: neg_eq_iff_add_eq_0 [symmetric]) | 
| 45548 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 huffman parents: 
45294diff
changeset | 559 | |
| 63325 | 560 | lemma minus_diff_eq [simp]: "- (a - b) = b - a" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 561 | by (simp only: neg_eq_iff_add_eq_0 diff_conv_add_uminus add.assoc minus_add_cancel) simp | 
| 45548 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 huffman parents: 
45294diff
changeset | 562 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 563 | lemma add_diff_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 564 | "a + (b - c) = (a + b) - c" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 565 | by (simp only: diff_conv_add_uminus add.assoc) | 
| 45548 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 huffman parents: 
45294diff
changeset | 566 | |
| 63325 | 567 | lemma diff_add_eq_diff_diff_swap: "a - (b + c) = a - c - b" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 568 | by (simp only: diff_conv_add_uminus add.assoc minus_add) | 
| 45548 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 huffman parents: 
45294diff
changeset | 569 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 570 | lemma diff_eq_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 571 | "a - b = c \<longleftrightarrow> a = c + b" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 572 | by auto | 
| 45548 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 huffman parents: 
45294diff
changeset | 573 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 574 | lemma eq_diff_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 575 | "a = c - b \<longleftrightarrow> a + b = c" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 576 | by auto | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 577 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 578 | lemma diff_diff_eq2 [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 579 | "a - (b - c) = (a + c) - b" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 580 | by (simp only: diff_conv_add_uminus add.assoc) simp | 
| 45548 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 huffman parents: 
45294diff
changeset | 581 | |
| 63325 | 582 | lemma diff_eq_diff_eq: "a - b = c - d \<Longrightarrow> a = b \<longleftrightarrow> c = d" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 583 | by (simp only: eq_iff_diff_eq_0 [of a b] eq_iff_diff_eq_0 [of c d]) | 
| 45548 
3e2722d66169
Groups.thy: generalize several lemmas from class ab_group_add to class group_add
 huffman parents: 
45294diff
changeset | 584 | |
| 25062 | 585 | end | 
| 586 | ||
| 25762 | 587 | class ab_group_add = minus + uminus + comm_monoid_add + | 
| 25062 | 588 | assumes ab_left_minus: "- a + a = 0" | 
| 59557 | 589 | assumes ab_diff_conv_add_uminus: "a - b = a + (- b)" | 
| 25267 | 590 | begin | 
| 25062 | 591 | |
| 25267 | 592 | subclass group_add | 
| 63325 | 593 | by standard (simp_all add: ab_left_minus ab_diff_conv_add_uminus) | 
| 25062 | 594 | |
| 29904 | 595 | subclass cancel_comm_monoid_add | 
| 28823 | 596 | proof | 
| 25062 | 597 | fix a b c :: 'a | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 598 | have "b + a - a = b" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 599 | by simp | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 600 | then show "a + b - a = b" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 601 | by (simp add: ac_simps) | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 602 | show "a - b - c = a - (b + c)" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 603 | by (simp add: algebra_simps) | 
| 25062 | 604 | qed | 
| 605 | ||
| 63325 | 606 | lemma uminus_add_conv_diff [simp]: "- a + b = b - a" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 607 | by (simp add: add.commute) | 
| 25062 | 608 | |
| 63325 | 609 | lemma minus_add_distrib [simp]: "- (a + b) = - a + - b" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 610 | by (simp add: algebra_simps) | 
| 25062 | 611 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 612 | lemma diff_add_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 613 | "(a - b) + c = (a + c) - b" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 614 | by (simp add: algebra_simps) | 
| 25077 | 615 | |
| 71093 | 616 | lemma minus_diff_commute: | 
| 617 | "- b - a = - a - b" | |
| 618 | by (simp only: diff_conv_add_uminus add.commute) | |
| 619 | ||
| 25062 | 620 | end | 
| 14738 | 621 | |
| 37884 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
36977diff
changeset | 622 | |
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 623 | subsection \<open>(Partially) Ordered Groups\<close> | 
| 14738 | 624 | |
| 60758 | 625 | text \<open> | 
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 626 | The theory of partially ordered groups is taken from the books: | 
| 63325 | 627 | |
| 628 | \<^item> \<^emph>\<open>Lattice Theory\<close> by Garret Birkhoff, American Mathematical Society, 1979 | |
| 629 | \<^item> \<^emph>\<open>Partially Ordered Algebraic Systems\<close>, Pergamon Press, 1963 | |
| 630 | ||
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 631 | Most of the used notions can also be looked up in | 
| 63680 | 632 | \<^item> \<^url>\<open>http://www.mathworld.com\<close> by Eric Weisstein et. al. | 
| 63325 | 633 | \<^item> \<^emph>\<open>Algebra I\<close> by van der Waerden, Springer | 
| 60758 | 634 | \<close> | 
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35267diff
changeset | 635 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 636 | class ordered_ab_semigroup_add = order + ab_semigroup_add + | 
| 25062 | 637 | assumes add_left_mono: "a \<le> b \<Longrightarrow> c + a \<le> c + b" | 
| 638 | begin | |
| 24380 
c215e256beca
moved ordered_ab_semigroup_add to OrderedGroup.thy
 haftmann parents: 
24286diff
changeset | 639 | |
| 63325 | 640 | lemma add_right_mono: "a \<le> b \<Longrightarrow> a + c \<le> b + c" | 
| 641 | by (simp add: add.commute [of _ c] add_left_mono) | |
| 14738 | 642 | |
| 60758 | 643 | text \<open>non-strict, in both arguments\<close> | 
| 63325 | 644 | lemma add_mono: "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> a + c \<le> b + d" | 
| 71743 | 645 | by (simp add: add.commute add_left_mono add_right_mono [THEN order_trans]) | 
| 14738 | 646 | |
| 25062 | 647 | end | 
| 648 | ||
| 63325 | 649 | text \<open>Strict monotonicity in both arguments\<close> | 
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 650 | class strict_ordered_ab_semigroup_add = ordered_ab_semigroup_add + | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 651 | assumes add_strict_mono: "a < b \<Longrightarrow> c < d \<Longrightarrow> a + c < b + d" | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 652 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 653 | class ordered_cancel_ab_semigroup_add = | 
| 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 654 | ordered_ab_semigroup_add + cancel_ab_semigroup_add | 
| 25062 | 655 | begin | 
| 656 | ||
| 63325 | 657 | lemma add_strict_left_mono: "a < b \<Longrightarrow> c + a < c + b" | 
| 658 | by (auto simp add: less_le add_left_mono) | |
| 14738 | 659 | |
| 63325 | 660 | lemma add_strict_right_mono: "a < b \<Longrightarrow> a + c < b + c" | 
| 661 | by (simp add: add.commute [of _ c] add_strict_left_mono) | |
| 14738 | 662 | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 663 | subclass strict_ordered_ab_semigroup_add | 
| 71743 | 664 | proof | 
| 665 | show "\<And>a b c d. \<lbrakk>a < b; c < d\<rbrakk> \<Longrightarrow> a + c < b + d" | |
| 666 | by (iprover intro: add_strict_left_mono add_strict_right_mono less_trans) | |
| 667 | qed | |
| 14738 | 668 | |
| 63325 | 669 | lemma add_less_le_mono: "a < b \<Longrightarrow> c \<le> d \<Longrightarrow> a + c < b + d" | 
| 71743 | 670 | by (iprover intro: add_left_mono add_strict_right_mono less_le_trans) | 
| 14738 | 671 | |
| 63325 | 672 | lemma add_le_less_mono: "a \<le> b \<Longrightarrow> c < d \<Longrightarrow> a + c < b + d" | 
| 71743 | 673 | by (iprover intro: add_strict_left_mono add_right_mono less_le_trans) | 
| 14738 | 674 | |
| 25062 | 675 | end | 
| 676 | ||
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 677 | class ordered_ab_semigroup_add_imp_le = ordered_cancel_ab_semigroup_add + | 
| 25062 | 678 | assumes add_le_imp_le_left: "c + a \<le> c + b \<Longrightarrow> a \<le> b" | 
| 679 | begin | |
| 680 | ||
| 14738 | 681 | lemma add_less_imp_less_left: | 
| 63325 | 682 | assumes less: "c + a < c + b" | 
| 683 | shows "a < b" | |
| 14738 | 684 | proof - | 
| 63325 | 685 | from less have le: "c + a \<le> c + b" | 
| 686 | by (simp add: order_le_less) | |
| 687 | have "a \<le> b" | |
| 71743 | 688 | using add_le_imp_le_left [OF le] . | 
| 14738 | 689 | moreover have "a \<noteq> b" | 
| 690 | proof (rule ccontr) | |
| 63325 | 691 | assume "\<not> ?thesis" | 
| 14738 | 692 | then have "a = b" by simp | 
| 693 | then have "c + a = c + b" by simp | |
| 63325 | 694 | with less show "False" by simp | 
| 14738 | 695 | qed | 
| 63325 | 696 | ultimately show "a < b" | 
| 697 | by (simp add: order_le_less) | |
| 14738 | 698 | qed | 
| 699 | ||
| 63325 | 700 | lemma add_less_imp_less_right: "a + c < b + c \<Longrightarrow> a < b" | 
| 701 | by (rule add_less_imp_less_left [of c]) (simp add: add.commute) | |
| 14738 | 702 | |
| 63325 | 703 | lemma add_less_cancel_left [simp]: "c + a < c + b \<longleftrightarrow> a < b" | 
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 704 | by (blast intro: add_less_imp_less_left add_strict_left_mono) | 
| 14738 | 705 | |
| 63325 | 706 | lemma add_less_cancel_right [simp]: "a + c < b + c \<longleftrightarrow> a < b" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 707 | by (blast intro: add_less_imp_less_right add_strict_right_mono) | 
| 14738 | 708 | |
| 63325 | 709 | lemma add_le_cancel_left [simp]: "c + a \<le> c + b \<longleftrightarrow> a \<le> b" | 
| 71743 | 710 | by (auto simp: dest: add_le_imp_le_left add_left_mono) | 
| 14738 | 711 | |
| 63325 | 712 | lemma add_le_cancel_right [simp]: "a + c \<le> b + c \<longleftrightarrow> a \<le> b" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56950diff
changeset | 713 | by (simp add: add.commute [of a c] add.commute [of b c]) | 
| 14738 | 714 | |
| 63325 | 715 | lemma add_le_imp_le_right: "a + c \<le> b + c \<Longrightarrow> a \<le> b" | 
| 716 | by simp | |
| 25062 | 717 | |
| 63325 | 718 | lemma max_add_distrib_left: "max x y + z = max (x + z) (y + z)" | 
| 25077 | 719 | unfolding max_def by auto | 
| 720 | ||
| 63325 | 721 | lemma min_add_distrib_left: "min x y + z = min (x + z) (y + z)" | 
| 25077 | 722 | unfolding min_def by auto | 
| 723 | ||
| 63325 | 724 | lemma max_add_distrib_right: "x + max y z = max (x + y) (x + z)" | 
| 44848 
f4d0b060c7ca
remove lemmas nat_add_min_{left,right} in favor of generic lemmas min_add_distrib_{left,right}
 huffman parents: 
44433diff
changeset | 725 | 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: 
44433diff
changeset | 726 | |
| 63325 | 727 | lemma min_add_distrib_right: "x + min y z = min (x + y) (x + z)" | 
| 44848 
f4d0b060c7ca
remove lemmas nat_add_min_{left,right} in favor of generic lemmas min_add_distrib_{left,right}
 huffman parents: 
44433diff
changeset | 728 | 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: 
44433diff
changeset | 729 | |
| 25062 | 730 | end | 
| 731 | ||
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 732 | subsection \<open>Support for reasoning about signs\<close> | 
| 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 733 | |
| 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 734 | class ordered_comm_monoid_add = comm_monoid_add + ordered_ab_semigroup_add | 
| 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 735 | begin | 
| 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 736 | |
| 63325 | 737 | lemma add_nonneg_nonneg [simp]: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a + b" | 
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 738 | using add_mono[of 0 a 0 b] by simp | 
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 739 | |
| 63325 | 740 | lemma add_nonpos_nonpos: "a \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> a + b \<le> 0" | 
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 741 | using add_mono[of a 0 b 0] by simp | 
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 742 | |
| 63325 | 743 | lemma add_nonneg_eq_0_iff: "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> x + y = 0 \<longleftrightarrow> x = 0 \<and> y = 0" | 
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 744 | using add_left_mono[of 0 y x] add_right_mono[of 0 x y] by auto | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 745 | |
| 63325 | 746 | lemma add_nonpos_eq_0_iff: "x \<le> 0 \<Longrightarrow> y \<le> 0 \<Longrightarrow> x + y = 0 \<longleftrightarrow> x = 0 \<and> y = 0" | 
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 747 | using add_left_mono[of y 0 x] add_right_mono[of x 0 y] by auto | 
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 748 | |
| 63325 | 749 | lemma add_increasing: "0 \<le> a \<Longrightarrow> b \<le> c \<Longrightarrow> b \<le> a + c" | 
| 750 | using add_mono [of 0 a b c] by simp | |
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 751 | |
| 63325 | 752 | lemma add_increasing2: "0 \<le> c \<Longrightarrow> b \<le> a \<Longrightarrow> b \<le> a + c" | 
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 753 | by (simp add: add_increasing add.commute [of a]) | 
| 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 754 | |
| 63325 | 755 | lemma add_decreasing: "a \<le> 0 \<Longrightarrow> c \<le> b \<Longrightarrow> a + c \<le> b" | 
| 756 | using add_mono [of a 0 c b] by simp | |
| 52289 | 757 | |
| 63325 | 758 | lemma add_decreasing2: "c \<le> 0 \<Longrightarrow> a \<le> b \<Longrightarrow> a + c \<le> b" | 
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 759 | using add_mono[of a b c 0] by simp | 
| 52289 | 760 | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 761 | lemma add_pos_nonneg: "0 < a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 < a + b" | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 762 | using less_le_trans[of 0 a "a + b"] by (simp add: add_increasing2) | 
| 52289 | 763 | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 764 | lemma add_pos_pos: "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> 0 < a + b" | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 765 | by (intro add_pos_nonneg less_imp_le) | 
| 52289 | 766 | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 767 | lemma add_nonneg_pos: "0 \<le> a \<Longrightarrow> 0 < b \<Longrightarrow> 0 < a + b" | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 768 | using add_pos_nonneg[of b a] by (simp add: add_commute) | 
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 769 | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 770 | lemma add_neg_nonpos: "a < 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> a + b < 0" | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 771 | using le_less_trans[of "a + b" a 0] by (simp add: add_decreasing2) | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 772 | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 773 | lemma add_neg_neg: "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> a + b < 0" | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 774 | by (intro add_neg_nonpos less_imp_le) | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 775 | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 776 | lemma add_nonpos_neg: "a \<le> 0 \<Longrightarrow> b < 0 \<Longrightarrow> a + b < 0" | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 777 | using add_neg_nonpos[of b a] by (simp add: add_commute) | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 778 | |
| 30691 | 779 | lemmas add_sign_intros = | 
| 780 | add_pos_nonneg add_pos_pos add_nonneg_pos add_nonneg_nonneg | |
| 781 | add_neg_nonpos add_neg_neg add_nonpos_neg add_nonpos_nonpos | |
| 782 | ||
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 783 | end | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 784 | |
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 785 | class strict_ordered_comm_monoid_add = comm_monoid_add + strict_ordered_ab_semigroup_add | 
| 62378 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 786 | begin | 
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 787 | |
| 63325 | 788 | lemma pos_add_strict: "0 < a \<Longrightarrow> b < c \<Longrightarrow> b < a + c" | 
| 62378 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 789 | using add_strict_mono [of 0 a b c] by simp | 
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 790 | |
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 791 | end | 
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 792 | |
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 793 | class ordered_cancel_comm_monoid_add = ordered_comm_monoid_add + cancel_ab_semigroup_add | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 794 | begin | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 795 | |
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 796 | subclass ordered_cancel_ab_semigroup_add .. | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 797 | subclass strict_ordered_comm_monoid_add .. | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 798 | |
| 63325 | 799 | lemma add_strict_increasing: "0 < a \<Longrightarrow> b \<le> c \<Longrightarrow> b < a + c" | 
| 800 | using add_less_le_mono [of 0 a b c] by simp | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 801 | |
| 63325 | 802 | lemma add_strict_increasing2: "0 \<le> a \<Longrightarrow> b < c \<Longrightarrow> b < a + c" | 
| 803 | using add_le_less_mono [of 0 a b c] by simp | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 804 | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 805 | end | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 806 | |
| 63456 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 807 | class ordered_ab_semigroup_monoid_add_imp_le = monoid_add + ordered_ab_semigroup_add_imp_le | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 808 | begin | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 809 | |
| 63588 | 810 | lemma add_less_same_cancel1 [simp]: "b + a < b \<longleftrightarrow> a < 0" | 
| 63456 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 811 | using add_less_cancel_left [of _ _ 0] by simp | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 812 | |
| 63588 | 813 | lemma add_less_same_cancel2 [simp]: "a + b < b \<longleftrightarrow> a < 0" | 
| 63456 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 814 | using add_less_cancel_right [of _ _ 0] by simp | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 815 | |
| 63588 | 816 | lemma less_add_same_cancel1 [simp]: "a < a + b \<longleftrightarrow> 0 < b" | 
| 63456 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 817 | using add_less_cancel_left [of _ 0] by simp | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 818 | |
| 63588 | 819 | lemma less_add_same_cancel2 [simp]: "a < b + a \<longleftrightarrow> 0 < b" | 
| 63456 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 820 | using add_less_cancel_right [of 0] by simp | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 821 | |
| 63588 | 822 | lemma add_le_same_cancel1 [simp]: "b + a \<le> b \<longleftrightarrow> a \<le> 0" | 
| 63456 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 823 | using add_le_cancel_left [of _ _ 0] by simp | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 824 | |
| 63588 | 825 | lemma add_le_same_cancel2 [simp]: "a + b \<le> b \<longleftrightarrow> a \<le> 0" | 
| 63456 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 826 | using add_le_cancel_right [of _ _ 0] by simp | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 827 | |
| 63588 | 828 | lemma le_add_same_cancel1 [simp]: "a \<le> a + b \<longleftrightarrow> 0 \<le> b" | 
| 63456 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 829 | using add_le_cancel_left [of _ 0] by simp | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 830 | |
| 63588 | 831 | lemma le_add_same_cancel2 [simp]: "a \<le> b + a \<longleftrightarrow> 0 \<le> b" | 
| 63456 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 832 | using add_le_cancel_right [of 0] by simp | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 833 | |
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 834 | subclass cancel_comm_monoid_add | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 835 | by standard auto | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 836 | |
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 837 | subclass ordered_cancel_comm_monoid_add | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 838 | by standard | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 839 | |
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 840 | end | 
| 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 841 | |
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 842 | class ordered_ab_group_add = ab_group_add + ordered_ab_semigroup_add | 
| 25062 | 843 | begin | 
| 844 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 845 | subclass ordered_cancel_ab_semigroup_add .. | 
| 25062 | 846 | |
| 63456 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63364diff
changeset | 847 | subclass ordered_ab_semigroup_monoid_add_imp_le | 
| 28823 | 848 | proof | 
| 25062 | 849 | fix a b c :: 'a | 
| 850 | assume "c + a \<le> c + b" | |
| 63325 | 851 | then have "(-c) + (c + a) \<le> (-c) + (c + b)" | 
| 852 | by (rule add_left_mono) | |
| 853 | then have "((-c) + c) + a \<le> ((-c) + c) + b" | |
| 854 | by (simp only: add.assoc) | |
| 855 | then show "a \<le> b" by simp | |
| 25062 | 856 | qed | 
| 857 | ||
| 63325 | 858 | lemma max_diff_distrib_left: "max x y - z = max (x - z) (y - z)" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 859 | using max_add_distrib_left [of x y "- z"] by simp | 
| 25077 | 860 | |
| 63325 | 861 | lemma min_diff_distrib_left: "min x y - z = min (x - z) (y - z)" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 862 | using min_add_distrib_left [of x y "- z"] by simp | 
| 25077 | 863 | |
| 864 | lemma le_imp_neg_le: | |
| 63325 | 865 | assumes "a \<le> b" | 
| 866 | shows "- b \<le> - a" | |
| 25077 | 867 | proof - | 
| 63325 | 868 | from assms have "- a + a \<le> - a + b" | 
| 869 | by (rule add_left_mono) | |
| 870 | then have "0 \<le> - a + b" | |
| 871 | by simp | |
| 872 | then have "0 + (- b) \<le> (- a + b) + (- b)" | |
| 873 | by (rule add_right_mono) | |
| 874 | then show ?thesis | |
| 875 | by (simp add: algebra_simps) | |
| 25077 | 876 | qed | 
| 877 | ||
| 878 | lemma neg_le_iff_le [simp]: "- b \<le> - a \<longleftrightarrow> a \<le> b" | |
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 879 | proof | 
| 25077 | 880 | assume "- b \<le> - a" | 
| 63325 | 881 | then have "- (- a) \<le> - (- b)" | 
| 882 | by (rule le_imp_neg_le) | |
| 883 | then show "a \<le> b" | |
| 884 | by simp | |
| 25077 | 885 | next | 
| 63325 | 886 | assume "a \<le> b" | 
| 887 | then show "- b \<le> - a" | |
| 888 | by (rule le_imp_neg_le) | |
| 25077 | 889 | qed | 
| 890 | ||
| 891 | lemma neg_le_0_iff_le [simp]: "- a \<le> 0 \<longleftrightarrow> 0 \<le> a" | |
| 63325 | 892 | by (subst neg_le_iff_le [symmetric]) simp | 
| 25077 | 893 | |
| 894 | lemma neg_0_le_iff_le [simp]: "0 \<le> - a \<longleftrightarrow> a \<le> 0" | |
| 63325 | 895 | by (subst neg_le_iff_le [symmetric]) simp | 
| 25077 | 896 | |
| 897 | lemma neg_less_iff_less [simp]: "- b < - a \<longleftrightarrow> a < b" | |
| 63325 | 898 | by (auto simp add: less_le) | 
| 25077 | 899 | |
| 900 | lemma neg_less_0_iff_less [simp]: "- a < 0 \<longleftrightarrow> 0 < a" | |
| 63325 | 901 | by (subst neg_less_iff_less [symmetric]) simp | 
| 25077 | 902 | |
| 903 | lemma neg_0_less_iff_less [simp]: "0 < - a \<longleftrightarrow> a < 0" | |
| 63325 | 904 | by (subst neg_less_iff_less [symmetric]) simp | 
| 25077 | 905 | |
| 63325 | 906 | text \<open>The next several equations can make the simplifier loop!\<close> | 
| 25077 | 907 | |
| 908 | lemma less_minus_iff: "a < - b \<longleftrightarrow> b < - a" | |
| 909 | proof - | |
| 63588 | 910 | have "- (- a) < - b \<longleftrightarrow> b < - a" | 
| 63325 | 911 | by (rule neg_less_iff_less) | 
| 912 | then show ?thesis by simp | |
| 25077 | 913 | qed | 
| 914 | ||
| 915 | lemma minus_less_iff: "- a < b \<longleftrightarrow> - b < a" | |
| 916 | proof - | |
| 63325 | 917 | have "- a < - (- b) \<longleftrightarrow> - b < a" | 
| 918 | by (rule neg_less_iff_less) | |
| 919 | then show ?thesis by simp | |
| 25077 | 920 | qed | 
| 921 | ||
| 922 | lemma le_minus_iff: "a \<le> - b \<longleftrightarrow> b \<le> - a" | |
| 71743 | 923 | by (auto simp: order.order_iff_strict less_minus_iff) | 
| 25077 | 924 | |
| 925 | lemma minus_le_iff: "- a \<le> b \<longleftrightarrow> - b \<le> a" | |
| 63325 | 926 | by (auto simp add: le_less minus_less_iff) | 
| 25077 | 927 | |
| 63325 | 928 | lemma diff_less_0_iff_less [simp]: "a - b < 0 \<longleftrightarrow> a < b" | 
| 25077 | 929 | proof - | 
| 63325 | 930 | have "a - b < 0 \<longleftrightarrow> a + (- b) < b + (- b)" | 
| 931 | by simp | |
| 932 | also have "\<dots> \<longleftrightarrow> a < b" | |
| 933 | by (simp only: add_less_cancel_right) | |
| 25077 | 934 | finally show ?thesis . | 
| 935 | qed | |
| 936 | ||
| 37884 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
36977diff
changeset | 937 | 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: 
36977diff
changeset | 938 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 939 | lemma diff_less_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 940 | "a - b < c \<longleftrightarrow> a < c + b" | 
| 71743 | 941 | proof (subst less_iff_diff_less_0 [of a]) | 
| 942 | show "(a - b < c) = (a - (c + b) < 0)" | |
| 943 | by (simp add: algebra_simps less_iff_diff_less_0 [of _ c]) | |
| 944 | qed | |
| 25077 | 945 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 946 | lemma less_diff_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 947 | "a < c - b \<longleftrightarrow> a + b < c" | 
| 71743 | 948 | proof (subst less_iff_diff_less_0 [of "a + b"]) | 
| 949 | show "(a < c - b) = (a + b - c < 0)" | |
| 950 | by (simp add: algebra_simps less_iff_diff_less_0 [of a]) | |
| 951 | qed | |
| 25077 | 952 | |
| 63325 | 953 | lemma diff_gt_0_iff_gt [simp]: "a - b > 0 \<longleftrightarrow> a > b" | 
| 62348 | 954 | by (simp add: less_diff_eq) | 
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61605diff
changeset | 955 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 956 | lemma diff_le_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 957 | "a - b \<le> c \<longleftrightarrow> a \<le> c + b" | 
| 62348 | 958 | by (auto simp add: le_less diff_less_eq ) | 
| 25077 | 959 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 960 | lemma le_diff_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70490diff
changeset | 961 | "a \<le> c - b \<longleftrightarrow> a + b \<le> c" | 
| 62348 | 962 | by (auto simp add: le_less less_diff_eq) | 
| 25077 | 963 | |
| 63325 | 964 | lemma diff_le_0_iff_le [simp]: "a - b \<le> 0 \<longleftrightarrow> a \<le> b" | 
| 37884 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
36977diff
changeset | 965 | by (simp add: algebra_simps) | 
| 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
36977diff
changeset | 966 | |
| 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
36977diff
changeset | 967 | 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: 
36977diff
changeset | 968 | |
| 63325 | 969 | lemma diff_ge_0_iff_ge [simp]: "a - b \<ge> 0 \<longleftrightarrow> a \<ge> b" | 
| 62348 | 970 | by (simp add: le_diff_eq) | 
| 971 | ||
| 63325 | 972 | lemma diff_eq_diff_less: "a - b = c - d \<Longrightarrow> a < b \<longleftrightarrow> c < d" | 
| 37884 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
36977diff
changeset | 973 | 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: 
36977diff
changeset | 974 | |
| 63325 | 975 | lemma diff_eq_diff_less_eq: "a - b = c - d \<Longrightarrow> a \<le> b \<longleftrightarrow> c \<le> d" | 
| 37889 
0d8058e0c270
keep explicit diff_def as legacy theorem; modernized abel_cancel simproc setup
 haftmann parents: 
37884diff
changeset | 976 | by (auto simp only: le_iff_diff_le_0 [of a b] le_iff_diff_le_0 [of c d]) | 
| 25077 | 977 | |
| 56950 | 978 | lemma diff_mono: "a \<le> b \<Longrightarrow> d \<le> c \<Longrightarrow> a - c \<le> b - d" | 
| 979 | by (simp add: field_simps add_mono) | |
| 980 | ||
| 981 | lemma diff_left_mono: "b \<le> a \<Longrightarrow> c - a \<le> c - b" | |
| 982 | by (simp add: field_simps) | |
| 983 | ||
| 984 | lemma diff_right_mono: "a \<le> b \<Longrightarrow> a - c \<le> b - c" | |
| 985 | by (simp add: field_simps) | |
| 986 | ||
| 987 | lemma diff_strict_mono: "a < b \<Longrightarrow> d < c \<Longrightarrow> a - c < b - d" | |
| 988 | by (simp add: field_simps add_strict_mono) | |
| 989 | ||
| 990 | lemma diff_strict_left_mono: "b < a \<Longrightarrow> c - a < c - b" | |
| 991 | by (simp add: field_simps) | |
| 992 | ||
| 993 | lemma diff_strict_right_mono: "a < b \<Longrightarrow> a - c < b - c" | |
| 994 | by (simp add: field_simps) | |
| 995 | ||
| 25077 | 996 | end | 
| 997 | ||
| 70490 | 998 | locale group_cancel | 
| 999 | begin | |
| 1000 | ||
| 1001 | lemma add1: "(A::'a::comm_monoid_add) \<equiv> k + a \<Longrightarrow> A + b \<equiv> k + (a + b)" | |
| 1002 | by (simp only: ac_simps) | |
| 1003 | ||
| 1004 | lemma add2: "(B::'a::comm_monoid_add) \<equiv> k + b \<Longrightarrow> a + B \<equiv> k + (a + b)" | |
| 1005 | by (simp only: ac_simps) | |
| 1006 | ||
| 1007 | lemma sub1: "(A::'a::ab_group_add) \<equiv> k + a \<Longrightarrow> A - b \<equiv> k + (a - b)" | |
| 1008 | by (simp only: add_diff_eq) | |
| 1009 | ||
| 1010 | lemma sub2: "(B::'a::ab_group_add) \<equiv> k + b \<Longrightarrow> a - B \<equiv> - k + (a - b)" | |
| 1011 | by (simp only: minus_add diff_conv_add_uminus ac_simps) | |
| 1012 | ||
| 1013 | lemma neg1: "(A::'a::ab_group_add) \<equiv> k + a \<Longrightarrow> - A \<equiv> - k + - a" | |
| 1014 | by (simp only: minus_add_distrib) | |
| 1015 | ||
| 1016 | lemma rule0: "(a::'a::comm_monoid_add) \<equiv> a + 0" | |
| 1017 | by (simp only: add_0_right) | |
| 1018 | ||
| 1019 | end | |
| 1020 | ||
| 69605 | 1021 | ML_file \<open>Tools/group_cancel.ML\<close> | 
| 48556 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: 
45548diff
changeset | 1022 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: 
45548diff
changeset | 1023 | simproc_setup group_cancel_add ("a + b::'a::ab_group_add") =
 | 
| 60758 | 1024 | \<open>fn phi => fn ss => try Group_Cancel.cancel_add_conv\<close> | 
| 48556 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: 
45548diff
changeset | 1025 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: 
45548diff
changeset | 1026 | simproc_setup group_cancel_diff ("a - b::'a::ab_group_add") =
 | 
| 60758 | 1027 | \<open>fn phi => fn ss => try Group_Cancel.cancel_diff_conv\<close> | 
| 37884 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
36977diff
changeset | 1028 | |
| 48556 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: 
45548diff
changeset | 1029 | simproc_setup group_cancel_eq ("a = (b::'a::ab_group_add)") =
 | 
| 60758 | 1030 | \<open>fn phi => fn ss => try Group_Cancel.cancel_eq_conv\<close> | 
| 37889 
0d8058e0c270
keep explicit diff_def as legacy theorem; modernized abel_cancel simproc setup
 haftmann parents: 
37884diff
changeset | 1031 | |
| 48556 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: 
45548diff
changeset | 1032 | simproc_setup group_cancel_le ("a \<le> (b::'a::ordered_ab_group_add)") =
 | 
| 60758 | 1033 | \<open>fn phi => fn ss => try Group_Cancel.cancel_le_conv\<close> | 
| 48556 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: 
45548diff
changeset | 1034 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: 
45548diff
changeset | 1035 | simproc_setup group_cancel_less ("a < (b::'a::ordered_ab_group_add)") =
 | 
| 60758 | 1036 | \<open>fn phi => fn ss => try Group_Cancel.cancel_less_conv\<close> | 
| 37884 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
36977diff
changeset | 1037 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 1038 | class linordered_ab_semigroup_add = | 
| 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 1039 | linorder + ordered_ab_semigroup_add | 
| 25062 | 1040 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 1041 | class linordered_cancel_ab_semigroup_add = | 
| 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 1042 | linorder + ordered_cancel_ab_semigroup_add | 
| 25267 | 1043 | begin | 
| 25062 | 1044 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 1045 | subclass linordered_ab_semigroup_add .. | 
| 25062 | 1046 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 1047 | subclass ordered_ab_semigroup_add_imp_le | 
| 28823 | 1048 | proof | 
| 25062 | 1049 | fix a b c :: 'a | 
| 63325 | 1050 | assume le1: "c + a \<le> c + b" | 
| 1051 | show "a \<le> b" | |
| 25062 | 1052 | proof (rule ccontr) | 
| 63325 | 1053 | assume *: "\<not> ?thesis" | 
| 1054 | then have "b \<le> a" by (simp add: linorder_not_le) | |
| 63588 | 1055 | then have "c + b \<le> c + a" by (rule add_left_mono) | 
| 71743 | 1056 | then have "c + a = c + b" | 
| 73411 | 1057 | using le1 by (iprover intro: order.antisym) | 
| 71743 | 1058 | then have "a = b" | 
| 1059 | by simp | |
| 63325 | 1060 | with * show False | 
| 25062 | 1061 | by (simp add: linorder_not_le [symmetric]) | 
| 1062 | qed | |
| 1063 | qed | |
| 1064 | ||
| 25267 | 1065 | end | 
| 1066 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 1067 | class linordered_ab_group_add = linorder + ordered_ab_group_add | 
| 25267 | 1068 | begin | 
| 25230 | 1069 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 1070 | subclass linordered_cancel_ab_semigroup_add .. | 
| 25230 | 1071 | |
| 63325 | 1072 | lemma equal_neg_zero [simp]: "a = - a \<longleftrightarrow> a = 0" | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1073 | proof | 
| 63325 | 1074 | assume "a = 0" | 
| 1075 | then show "a = - a" by simp | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1076 | next | 
| 63325 | 1077 | assume A: "a = - a" | 
| 1078 | show "a = 0" | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1079 | proof (cases "0 \<le> a") | 
| 63325 | 1080 | case True | 
| 1081 | with A have "0 \<le> - a" by auto | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1082 | with le_minus_iff have "a \<le> 0" by simp | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1083 | with True show ?thesis by (auto intro: order_trans) | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1084 | next | 
| 63325 | 1085 | case False | 
| 1086 | then have B: "a \<le> 0" by auto | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1087 | with A have "- a \<le> 0" by auto | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1088 | with B show ?thesis by (auto intro: order_trans) | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1089 | qed | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1090 | qed | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1091 | |
| 63325 | 1092 | lemma neg_equal_zero [simp]: "- a = a \<longleftrightarrow> a = 0" | 
| 35036 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1093 | by (auto dest: sym) | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1094 | |
| 63325 | 1095 | lemma neg_less_eq_nonneg [simp]: "- a \<le> a \<longleftrightarrow> 0 \<le> a" | 
| 54250 | 1096 | proof | 
| 63325 | 1097 | assume *: "- a \<le> a" | 
| 1098 | show "0 \<le> a" | |
| 54250 | 1099 | proof (rule classical) | 
| 63325 | 1100 | assume "\<not> ?thesis" | 
| 54250 | 1101 | then have "a < 0" by auto | 
| 63325 | 1102 | with * have "- a < 0" by (rule le_less_trans) | 
| 54250 | 1103 | then show ?thesis by auto | 
| 1104 | qed | |
| 1105 | next | |
| 63325 | 1106 | assume *: "0 \<le> a" | 
| 1107 | then have "- a \<le> 0" by (simp add: minus_le_iff) | |
| 1108 | from this * show "- a \<le> a" by (rule order_trans) | |
| 54250 | 1109 | qed | 
| 1110 | ||
| 63325 | 1111 | lemma neg_less_pos [simp]: "- a < a \<longleftrightarrow> 0 < a" | 
| 54250 | 1112 | by (auto simp add: less_le) | 
| 1113 | ||
| 63325 | 1114 | lemma less_eq_neg_nonpos [simp]: "a \<le> - a \<longleftrightarrow> a \<le> 0" | 
| 54250 | 1115 | using neg_less_eq_nonneg [of "- a"] by simp | 
| 1116 | ||
| 63325 | 1117 | lemma less_neg_neg [simp]: "a < - a \<longleftrightarrow> a < 0" | 
| 54250 | 1118 | using neg_less_pos [of "- a"] by simp | 
| 1119 | ||
| 63325 | 1120 | lemma double_zero [simp]: "a + a = 0 \<longleftrightarrow> a = 0" | 
| 35036 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1121 | proof | 
| 63325 | 1122 | assume "a + a = 0" | 
| 35036 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1123 | then have a: "- a = a" by (rule minus_unique) | 
| 35216 | 1124 | then show "a = 0" by (simp only: neg_equal_zero) | 
| 63325 | 1125 | next | 
| 1126 | assume "a = 0" | |
| 1127 | then show "a + a = 0" by simp | |
| 1128 | qed | |
| 35036 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1129 | |
| 63325 | 1130 | lemma double_zero_sym [simp]: "0 = a + a \<longleftrightarrow> a = 0" | 
| 71743 | 1131 | using double_zero [of a] by (simp only: eq_commute) | 
| 35036 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1132 | |
| 63325 | 1133 | lemma zero_less_double_add_iff_zero_less_single_add [simp]: "0 < a + a \<longleftrightarrow> 0 < a" | 
| 35036 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1134 | proof | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1135 | assume "0 < a + a" | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1136 | 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: 
35028diff
changeset | 1137 | then have "- a < a" by simp | 
| 54250 | 1138 | then show "0 < a" by simp | 
| 35036 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1139 | next | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1140 | assume "0 < a" | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1141 | 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: 
35028diff
changeset | 1142 | by (rule add_strict_mono) | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1143 | 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: 
35028diff
changeset | 1144 | qed | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1145 | |
| 63325 | 1146 | lemma zero_le_double_add_iff_zero_le_single_add [simp]: "0 \<le> a + a \<longleftrightarrow> 0 \<le> a" | 
| 35036 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1147 | by (auto simp add: le_less) | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1148 | |
| 63325 | 1149 | lemma double_add_less_zero_iff_single_add_less_zero [simp]: "a + a < 0 \<longleftrightarrow> a < 0" | 
| 35036 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1150 | proof - | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1151 | 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: 
35028diff
changeset | 1152 | by (simp add: not_less) | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1153 | then show ?thesis by simp | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1154 | qed | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1155 | |
| 63325 | 1156 | lemma double_add_le_zero_iff_single_add_le_zero [simp]: "a + a \<le> 0 \<longleftrightarrow> a \<le> 0" | 
| 35036 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1157 | proof - | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1158 | 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: 
35028diff
changeset | 1159 | by (simp add: not_le) | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1160 | then show ?thesis by simp | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1161 | qed | 
| 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1162 | |
| 63325 | 1163 | lemma minus_max_eq_min: "- max x y = min (- x) (- y)" | 
| 35036 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1164 | 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: 
35028diff
changeset | 1165 | |
| 63325 | 1166 | lemma minus_min_eq_max: "- min x y = max (- x) (- y)" | 
| 35036 
b8c8d01cc20d
separate library theory for type classes combining lattices with various algebraic structures; more simp rules
 haftmann parents: 
35028diff
changeset | 1167 | by (auto simp add: max_def min_def) | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1168 | |
| 25267 | 1169 | end | 
| 1170 | ||
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35050diff
changeset | 1171 | class abs = | 
| 81125 | 1172 | fixes abs :: "'a \<Rightarrow> 'a" (\<open>(\<open>open_block notation=\<open>mixfix abs\<close>\<close>\<bar>_\<bar>)\<close>) | 
| 1173 | ||
| 1174 | bundle abs_syntax | |
| 1175 | begin | |
| 1176 | notation abs (\<open>(\<open>open_block notation=\<open>mixfix abs\<close>\<close>\<bar>_\<bar>)\<close>) | |
| 1177 | end | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35050diff
changeset | 1178 | |
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35050diff
changeset | 1179 | class sgn = | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35050diff
changeset | 1180 | fixes sgn :: "'a \<Rightarrow> 'a" | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35050diff
changeset | 1181 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 1182 | class ordered_ab_group_add_abs = ordered_ab_group_add + abs + | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1183 | assumes abs_ge_zero [simp]: "\<bar>a\<bar> \<ge> 0" | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1184 | and abs_ge_self: "a \<le> \<bar>a\<bar>" | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1185 | 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: 
25267diff
changeset | 1186 | and abs_minus_cancel [simp]: "\<bar>-a\<bar> = \<bar>a\<bar>" | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1187 | 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: 
25267diff
changeset | 1188 | begin | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1189 | |
| 25307 | 1190 | lemma abs_minus_le_zero: "- \<bar>a\<bar> \<le> 0" | 
| 1191 | unfolding neg_le_0_iff_le by simp | |
| 1192 | ||
| 1193 | lemma abs_of_nonneg [simp]: | |
| 63325 | 1194 | assumes nonneg: "0 \<le> a" | 
| 1195 | shows "\<bar>a\<bar> = a" | |
| 73411 | 1196 | proof (rule order.antisym) | 
| 63325 | 1197 | show "a \<le> \<bar>a\<bar>" by (rule abs_ge_self) | 
| 25307 | 1198 | from nonneg le_imp_neg_le have "- a \<le> 0" by simp | 
| 1199 | from this nonneg have "- a \<le> a" by (rule order_trans) | |
| 1200 | then show "\<bar>a\<bar> \<le> a" by (auto intro: abs_leI) | |
| 63325 | 1201 | qed | 
| 25307 | 1202 | |
| 1203 | lemma abs_idempotent [simp]: "\<bar>\<bar>a\<bar>\<bar> = \<bar>a\<bar>" | |
| 73411 | 1204 | by (rule order.antisym) (auto intro!: abs_ge_self abs_leI order_trans [of "- \<bar>a\<bar>" 0 "\<bar>a\<bar>"]) | 
| 25307 | 1205 | |
| 1206 | lemma abs_eq_0 [simp]: "\<bar>a\<bar> = 0 \<longleftrightarrow> a = 0" | |
| 1207 | proof - | |
| 1208 | have "\<bar>a\<bar> = 0 \<Longrightarrow> a = 0" | |
| 73411 | 1209 | proof (rule order.antisym) | 
| 25307 | 1210 | assume zero: "\<bar>a\<bar> = 0" | 
| 1211 | with abs_ge_self show "a \<le> 0" by auto | |
| 1212 | from zero have "\<bar>-a\<bar> = 0" by simp | |
| 36302 | 1213 | with abs_ge_self [of "- a"] have "- a \<le> 0" by auto | 
| 25307 | 1214 | with neg_le_0_iff_le show "0 \<le> a" by auto | 
| 1215 | qed | |
| 1216 | then show ?thesis by auto | |
| 1217 | qed | |
| 1218 | ||
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1219 | lemma abs_zero [simp]: "\<bar>0\<bar> = 0" | 
| 63325 | 1220 | by simp | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16417diff
changeset | 1221 | |
| 54148 | 1222 | lemma abs_0_eq [simp]: "0 = \<bar>a\<bar> \<longleftrightarrow> a = 0" | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1223 | proof - | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1224 | have "0 = \<bar>a\<bar> \<longleftrightarrow> \<bar>a\<bar> = 0" by (simp only: eq_ac) | 
| 63325 | 1225 | then show ?thesis by simp | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1226 | qed | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1227 | |
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 1228 | lemma abs_le_zero_iff [simp]: "\<bar>a\<bar> \<le> 0 \<longleftrightarrow> a = 0" | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1229 | proof | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1230 | assume "\<bar>a\<bar> \<le> 0" | 
| 73411 | 1231 | then have "\<bar>a\<bar> = 0" by (rule order.antisym) simp | 
| 63325 | 1232 | then show "a = 0" by simp | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1233 | next | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1234 | assume "a = 0" | 
| 63325 | 1235 | then show "\<bar>a\<bar> \<le> 0" by simp | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1236 | qed | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1237 | |
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62378diff
changeset | 1238 | lemma abs_le_self_iff [simp]: "\<bar>a\<bar> \<le> a \<longleftrightarrow> 0 \<le> a" | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62378diff
changeset | 1239 | proof - | 
| 63325 | 1240 | have "0 \<le> \<bar>a\<bar>" | 
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62378diff
changeset | 1241 | using abs_ge_zero by blast | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62378diff
changeset | 1242 | then have "\<bar>a\<bar> \<le> a \<Longrightarrow> 0 \<le> a" | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62378diff
changeset | 1243 | using order.trans by blast | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62378diff
changeset | 1244 | then show ?thesis | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62378diff
changeset | 1245 | using abs_of_nonneg eq_refl by blast | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62378diff
changeset | 1246 | qed | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62378diff
changeset | 1247 | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1248 | lemma zero_less_abs_iff [simp]: "0 < \<bar>a\<bar> \<longleftrightarrow> a \<noteq> 0" | 
| 63325 | 1249 | by (simp add: less_le) | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1250 | |
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1251 | lemma abs_not_less_zero [simp]: "\<not> \<bar>a\<bar> < 0" | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1252 | proof - | 
| 63325 | 1253 | have "x \<le> y \<Longrightarrow> \<not> y < x" for x y by auto | 
| 1254 | then show ?thesis by simp | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1255 | qed | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16417diff
changeset | 1256 | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1257 | lemma abs_ge_minus_self: "- a \<le> \<bar>a\<bar>" | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1258 | proof - | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1259 | have "- a \<le> \<bar>-a\<bar>" by (rule abs_ge_self) | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1260 | then show ?thesis by simp | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1261 | qed | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1262 | |
| 63325 | 1263 | lemma abs_minus_commute: "\<bar>a - b\<bar> = \<bar>b - a\<bar>" | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1264 | proof - | 
| 63325 | 1265 | have "\<bar>a - b\<bar> = \<bar>- (a - b)\<bar>" | 
| 1266 | by (simp only: abs_minus_cancel) | |
| 1267 | also have "\<dots> = \<bar>b - a\<bar>" by simp | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1268 | finally show ?thesis . | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1269 | qed | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1270 | |
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1271 | lemma abs_of_pos: "0 < a \<Longrightarrow> \<bar>a\<bar> = a" | 
| 63325 | 1272 | 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: 
16417diff
changeset | 1273 | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1274 | lemma abs_of_nonpos [simp]: | 
| 63325 | 1275 | assumes "a \<le> 0" | 
| 1276 | shows "\<bar>a\<bar> = - a" | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1277 | proof - | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1278 | let ?b = "- a" | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1279 | have "- ?b \<le> 0 \<Longrightarrow> \<bar>- ?b\<bar> = - (- ?b)" | 
| 63325 | 1280 | unfolding abs_minus_cancel [of ?b] | 
| 1281 | unfolding neg_le_0_iff_le [of ?b] | |
| 1282 | unfolding minus_minus by (erule abs_of_nonneg) | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1283 | then show ?thesis using assms by auto | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1284 | qed | 
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62348diff
changeset | 1285 | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1286 | lemma abs_of_neg: "a < 0 \<Longrightarrow> \<bar>a\<bar> = - a" | 
| 63325 | 1287 | by (rule abs_of_nonpos) (rule less_imp_le) | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1288 | |
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1289 | lemma abs_le_D1: "\<bar>a\<bar> \<le> b \<Longrightarrow> a \<le> b" | 
| 63325 | 1290 | using abs_ge_self by (blast intro: order_trans) | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1291 | |
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1292 | lemma abs_le_D2: "\<bar>a\<bar> \<le> b \<Longrightarrow> - a \<le> b" | 
| 63325 | 1293 | using abs_le_D1 [of "- a"] by simp | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1294 | |
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1295 | lemma abs_le_iff: "\<bar>a\<bar> \<le> b \<longleftrightarrow> a \<le> b \<and> - a \<le> b" | 
| 63325 | 1296 | by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2) | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1297 | |
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1298 | lemma abs_triangle_ineq2: "\<bar>a\<bar> - \<bar>b\<bar> \<le> \<bar>a - b\<bar>" | 
| 36302 | 1299 | proof - | 
| 1300 | have "\<bar>a\<bar> = \<bar>b + (a - b)\<bar>" | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54148diff
changeset | 1301 | by (simp add: algebra_simps) | 
| 36302 | 1302 | then have "\<bar>a\<bar> \<le> \<bar>b\<bar> + \<bar>a - b\<bar>" | 
| 1303 | by (simp add: abs_triangle_ineq) | |
| 1304 | then show ?thesis | |
| 1305 | by (simp add: algebra_simps) | |
| 1306 | qed | |
| 1307 | ||
| 1308 | lemma abs_triangle_ineq2_sym: "\<bar>a\<bar> - \<bar>b\<bar> \<le> \<bar>b - a\<bar>" | |
| 1309 | 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: 
16417diff
changeset | 1310 | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1311 | lemma abs_triangle_ineq3: "\<bar>\<bar>a\<bar> - \<bar>b\<bar>\<bar> \<le> \<bar>a - b\<bar>" | 
| 36302 | 1312 | 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: 
16417diff
changeset | 1313 | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1314 | 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: 
25267diff
changeset | 1315 | proof - | 
| 63325 | 1316 | have "\<bar>a - b\<bar> = \<bar>a + - b\<bar>" | 
| 1317 | by (simp add: algebra_simps) | |
| 1318 | also have "\<dots> \<le> \<bar>a\<bar> + \<bar>- b\<bar>" | |
| 1319 | by (rule abs_triangle_ineq) | |
| 29667 | 1320 | finally show ?thesis by simp | 
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1321 | qed | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16417diff
changeset | 1322 | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1323 | 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: 
25267diff
changeset | 1324 | proof - | 
| 63325 | 1325 | have "\<bar>a + b - (c + d)\<bar> = \<bar>(a - c) + (b - d)\<bar>" | 
| 1326 | by (simp add: algebra_simps) | |
| 1327 | also have "\<dots> \<le> \<bar>a - c\<bar> + \<bar>b - d\<bar>" | |
| 1328 | by (rule abs_triangle_ineq) | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1329 | finally show ?thesis . | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1330 | qed | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16417diff
changeset | 1331 | |
| 63325 | 1332 | lemma abs_add_abs [simp]: "\<bar>\<bar>a\<bar> + \<bar>b\<bar>\<bar> = \<bar>a\<bar> + \<bar>b\<bar>" | 
| 1333 | (is "?L = ?R") | |
| 73411 | 1334 | proof (rule order.antisym) | 
| 63325 | 1335 | show "?L \<ge> ?R" by (rule abs_ge_self) | 
| 1336 | have "?L \<le> \<bar>\<bar>a\<bar>\<bar> + \<bar>\<bar>b\<bar>\<bar>" by (rule abs_triangle_ineq) | |
| 25303 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1337 | also have "\<dots> = ?R" by simp | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1338 | finally show "?L \<le> ?R" . | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1339 | qed | 
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1340 | |
| 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 haftmann parents: 
25267diff
changeset | 1341 | end | 
| 14738 | 1342 | |
| 60762 | 1343 | lemma dense_eq0_I: | 
| 1344 |   fixes x::"'a::{dense_linorder,ordered_ab_group_add_abs}"
 | |
| 71743 | 1345 | assumes "\<And>e. 0 < e \<Longrightarrow> \<bar>x\<bar> \<le> e" | 
| 1346 | shows "x = 0" | |
| 1347 | proof (cases "\<bar>x\<bar> = 0") | |
| 1348 | case False | |
| 1349 | then have "\<bar>x\<bar> > 0" | |
| 1350 | by simp | |
| 1351 | then obtain z where "0 < z" "z < \<bar>x\<bar>" | |
| 1352 | using dense by force | |
| 1353 | then show ?thesis | |
| 1354 | using assms by (simp flip: not_less) | |
| 1355 | qed auto | |
| 60762 | 1356 | |
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 1357 | hide_fact (open) ab_diff_conv_add_uminus add_0 mult_1 ab_left_minus | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 1358 | |
| 63325 | 1359 | lemmas add_0 = add_0_left (* FIXME duplicate *) | 
| 1360 | lemmas mult_1 = mult_1_left (* FIXME duplicate *) | |
| 1361 | lemmas ab_left_minus = left_minus (* FIXME duplicate *) | |
| 1362 | lemmas diff_diff_eq = diff_diff_add (* FIXME duplicate *) | |
| 1363 | ||
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59559diff
changeset | 1364 | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1365 | subsection \<open>Canonically ordered monoids\<close> | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1366 | |
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1367 | text \<open>Canonically ordered monoids are never groups.\<close> | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1368 | |
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1369 | class canonically_ordered_monoid_add = comm_monoid_add + order + | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1370 | assumes le_iff_add: "a \<le> b \<longleftrightarrow> (\<exists>c. b = a + c)" | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1371 | begin | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1372 | |
| 62378 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1373 | lemma zero_le[simp]: "0 \<le> x" | 
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1374 | by (auto simp: le_iff_add) | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1375 | |
| 62378 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1376 | lemma le_zero_eq[simp]: "n \<le> 0 \<longleftrightarrow> n = 0" | 
| 73411 | 1377 | by (auto intro: order.antisym) | 
| 62378 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1378 | |
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1379 | lemma not_less_zero[simp]: "\<not> n < 0" | 
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1380 | by (auto simp: less_le) | 
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1381 | |
| 63325 | 1382 | lemma zero_less_iff_neq_zero: "0 < n \<longleftrightarrow> n \<noteq> 0" | 
| 62378 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1383 | by (auto simp: less_le) | 
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1384 | |
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1385 | text \<open>This theorem is useful with \<open>blast\<close>\<close> | 
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1386 | lemma gr_zeroI: "(n = 0 \<Longrightarrow> False) \<Longrightarrow> 0 < n" | 
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1387 | by (rule zero_less_iff_neq_zero[THEN iffD2]) iprover | 
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1388 | |
| 63325 | 1389 | lemma not_gr_zero[simp]: "\<not> 0 < n \<longleftrightarrow> n = 0" | 
| 62378 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1390 | by (simp add: zero_less_iff_neq_zero) | 
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1391 | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1392 | subclass ordered_comm_monoid_add | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1393 | proof qed (auto simp: le_iff_add add_ac) | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1394 | |
| 63878 
e26c7f58d78e
add add_eq_0_iff_both_eq_0 and zero_eq_add_iff_both_eq_0 to simp set
 hoelzl parents: 
63680diff
changeset | 1395 | lemma gr_implies_not_zero: "m < n \<Longrightarrow> n \<noteq> 0" | 
| 
e26c7f58d78e
add add_eq_0_iff_both_eq_0 and zero_eq_add_iff_both_eq_0 to simp set
 hoelzl parents: 
63680diff
changeset | 1396 | by auto | 
| 
e26c7f58d78e
add add_eq_0_iff_both_eq_0 and zero_eq_add_iff_both_eq_0 to simp set
 hoelzl parents: 
63680diff
changeset | 1397 | |
| 
e26c7f58d78e
add add_eq_0_iff_both_eq_0 and zero_eq_add_iff_both_eq_0 to simp set
 hoelzl parents: 
63680diff
changeset | 1398 | lemma add_eq_0_iff_both_eq_0[simp]: "x + y = 0 \<longleftrightarrow> x = 0 \<and> y = 0" | 
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1399 | by (intro add_nonneg_eq_0_iff zero_le) | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1400 | |
| 63878 
e26c7f58d78e
add add_eq_0_iff_both_eq_0 and zero_eq_add_iff_both_eq_0 to simp set
 hoelzl parents: 
63680diff
changeset | 1401 | lemma zero_eq_add_iff_both_eq_0[simp]: "0 = x + y \<longleftrightarrow> x = 0 \<and> y = 0" | 
| 
e26c7f58d78e
add add_eq_0_iff_both_eq_0 and zero_eq_add_iff_both_eq_0 to simp set
 hoelzl parents: 
63680diff
changeset | 1402 | using add_eq_0_iff_both_eq_0[of x y] unfolding eq_commute[of 0] . | 
| 62378 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1403 | |
| 71425 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1404 | lemma less_eqE: | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1405 | assumes \<open>a \<le> b\<close> | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1406 | obtains c where \<open>b = a + c\<close> | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1407 | using assms by (auto simp add: le_iff_add) | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1408 | |
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1409 | lemma lessE: | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1410 | assumes \<open>a < b\<close> | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1411 | obtains c where \<open>b = a + c\<close> and \<open>c \<noteq> 0\<close> | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1412 | proof - | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1413 | from assms have \<open>a \<le> b\<close> \<open>a \<noteq> b\<close> | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1414 | by simp_all | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1415 | from \<open>a \<le> b\<close> obtain c where \<open>b = a + c\<close> | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1416 | by (rule less_eqE) | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1417 | moreover have \<open>c \<noteq> 0\<close> using \<open>a \<noteq> b\<close> \<open>b = a + c\<close> | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1418 | by auto | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1419 | ultimately show ?thesis | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1420 | by (rule that) | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1421 | qed | 
| 
f2da99316b86
more rules for natural deduction from inequalities
 haftmann parents: 
71093diff
changeset | 1422 | |
| 62378 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1423 | lemmas zero_order = zero_le le_zero_eq not_less_zero zero_less_iff_neq_zero not_gr_zero | 
| 63145 | 1424 | \<comment> \<open>This should be attributed with \<open>[iff]\<close>, but then \<open>blast\<close> fails in \<open>Set\<close>.\<close> | 
| 62378 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62377diff
changeset | 1425 | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1426 | end | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1427 | |
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1428 | class ordered_cancel_comm_monoid_diff = | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1429 | canonically_ordered_monoid_add + comm_monoid_diff + ordered_ab_semigroup_add_imp_le | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1430 | begin | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1431 | |
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1432 | context | 
| 63588 | 1433 | fixes a b :: 'a | 
| 63325 | 1434 | assumes le: "a \<le> b" | 
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1435 | begin | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1436 | |
| 63325 | 1437 | lemma add_diff_inverse: "a + (b - a) = b" | 
| 1438 | using le by (auto simp add: le_iff_add) | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1439 | |
| 63325 | 1440 | lemma add_diff_assoc: "c + (b - a) = c + b - a" | 
| 1441 | using le by (auto simp add: le_iff_add add.left_commute [of c]) | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1442 | |
| 63325 | 1443 | lemma add_diff_assoc2: "b - a + c = b + c - a" | 
| 1444 | using le by (auto simp add: le_iff_add add.assoc) | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1445 | |
| 63325 | 1446 | lemma diff_add_assoc: "c + b - a = c + (b - a)" | 
| 1447 | using le by (simp add: add.commute add_diff_assoc) | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1448 | |
| 63325 | 1449 | lemma diff_add_assoc2: "b + c - a = b - a + c" | 
| 1450 | using le by (simp add: add.commute add_diff_assoc) | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1451 | |
| 63325 | 1452 | lemma diff_diff_right: "c - (b - a) = c + a - b" | 
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1453 | by (simp add: add_diff_inverse add_diff_cancel_left [of a c "b - a", symmetric] add.commute) | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1454 | |
| 63325 | 1455 | lemma diff_add: "b - a + a = b" | 
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1456 | by (simp add: add.commute add_diff_inverse) | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1457 | |
| 63325 | 1458 | lemma le_add_diff: "c \<le> b + c - a" | 
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1459 | by (auto simp add: add.commute diff_add_assoc2 le_iff_add) | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1460 | |
| 63325 | 1461 | lemma le_imp_diff_is_add: "a \<le> b \<Longrightarrow> b - a = c \<longleftrightarrow> b = c + a" | 
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1462 | by (auto simp add: add.commute add_diff_inverse) | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1463 | |
| 63325 | 1464 | lemma le_diff_conv2: "c \<le> b - a \<longleftrightarrow> c + a \<le> b" | 
| 1465 | (is "?P \<longleftrightarrow> ?Q") | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1466 | proof | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1467 | assume ?P | 
| 63325 | 1468 | then have "c + a \<le> b - a + a" | 
| 1469 | by (rule add_right_mono) | |
| 1470 | then show ?Q | |
| 1471 | by (simp add: add_diff_inverse add.commute) | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1472 | next | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1473 | assume ?Q | 
| 63325 | 1474 | then have "a + c \<le> a + (b - a)" | 
| 1475 | by (simp add: add_diff_inverse add.commute) | |
| 62377 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1476 | then show ?P by simp | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1477 | qed | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1478 | |
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1479 | end | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1480 | |
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1481 | end | 
| 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 hoelzl parents: 
62376diff
changeset | 1482 | |
| 63325 | 1483 | |
| 60758 | 1484 | subsection \<open>Tools setup\<close> | 
| 25090 | 1485 | |
| 54147 
97a8ff4e4ac9
killed most "no_atp", to make Sledgehammer more complete
 blanchet parents: 
52435diff
changeset | 1486 | lemma add_mono_thms_linordered_semiring: | 
| 61076 | 1487 | fixes i j k :: "'a::ordered_ab_semigroup_add" | 
| 25077 | 1488 | shows "i \<le> j \<and> k \<le> l \<Longrightarrow> i + k \<le> j + l" | 
| 1489 | and "i = j \<and> k \<le> l \<Longrightarrow> i + k \<le> j + l" | |
| 1490 | and "i \<le> j \<and> k = l \<Longrightarrow> i + k \<le> j + l" | |
| 1491 | and "i = j \<and> k = l \<Longrightarrow> i + k = j + l" | |
| 63325 | 1492 | by (rule add_mono, clarify+)+ | 
| 25077 | 1493 | |
| 54147 
97a8ff4e4ac9
killed most "no_atp", to make Sledgehammer more complete
 blanchet parents: 
52435diff
changeset | 1494 | lemma add_mono_thms_linordered_field: | 
| 61076 | 1495 | fixes i j k :: "'a::ordered_cancel_ab_semigroup_add" | 
| 25077 | 1496 | shows "i < j \<and> k = l \<Longrightarrow> i + k < j + l" | 
| 1497 | and "i = j \<and> k < l \<Longrightarrow> i + k < j + l" | |
| 1498 | and "i < j \<and> k \<le> l \<Longrightarrow> i + k < j + l" | |
| 1499 | and "i \<le> j \<and> k < l \<Longrightarrow> i + k < j + l" | |
| 1500 | and "i < j \<and> k < l \<Longrightarrow> i + k < j + l" | |
| 63325 | 1501 | by (auto intro: add_strict_right_mono add_strict_left_mono | 
| 1502 | add_less_le_mono add_le_less_mono add_strict_mono) | |
| 25077 | 1503 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52289diff
changeset | 1504 | code_identifier | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52289diff
changeset | 1505 | code_module Groups \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith | 
| 33364 | 1506 | |
| 14738 | 1507 | end |