| author | webertj | 
| Tue, 01 Aug 2006 14:58:43 +0200 | |
| changeset 20276 | d94dc40673b1 | 
| parent 19984 | 29bb4659f80a | 
| child 20588 | c847c56edf0c | 
| permissions | -rw-r--r-- | 
| 15524 | 1 | (* Title: HOL/Orderings.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson | |
| 4 | ||
| 5 | FIXME: derive more of the min/max laws generically via semilattices | |
| 6 | *) | |
| 7 | ||
| 8 | header {* Type classes for $\le$ *}
 | |
| 9 | ||
| 10 | theory Orderings | |
| 11 | imports Lattice_Locales | |
| 16417 | 12 | uses ("antisym_setup.ML")
 | 
| 15524 | 13 | begin | 
| 14 | ||
| 15 | subsection {* Order signatures and orders *}
 | |
| 16 | ||
| 17 | axclass | |
| 18 | ord < type | |
| 19 | ||
| 20 | consts | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 21 | less :: "['a::ord, 'a] => bool" | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 22 | less_eq :: "['a::ord, 'a] => bool" | 
| 15524 | 23 | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 24 | const_syntax | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 25 |   less  ("op <")
 | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 26 |   less  ("(_/ < _)"  [50, 51] 50)
 | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 27 |   less_eq  ("op <=")
 | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 28 |   less_eq  ("(_/ <= _)" [50, 51] 50)
 | 
| 15524 | 29 | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 30 | const_syntax (xsymbols) | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 31 |   less_eq  ("op \<le>")
 | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 32 |   less_eq  ("(_/ \<le> _)"  [50, 51] 50)
 | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 33 | |
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 34 | const_syntax (HTML output) | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 35 |   less_eq  ("op \<le>")
 | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 36 |   less_eq  ("(_/ \<le> _)"  [50, 51] 50)
 | 
| 15524 | 37 | |
| 19536 | 38 | abbreviation (input) | 
| 39 | greater (infixl ">" 50) | |
| 40 | "x > y == y < x" | |
| 41 | greater_eq (infixl ">=" 50) | |
| 42 | "x >= y == y <= x" | |
| 15524 | 43 | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 44 | const_syntax (xsymbols) | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 45 | greater_eq (infixl "\<ge>" 50) | 
| 15524 | 46 | |
| 47 | ||
| 48 | subsection {* Monotonicity *}
 | |
| 49 | ||
| 50 | locale mono = | |
| 51 | fixes f | |
| 52 | assumes mono: "A <= B ==> f A <= f B" | |
| 53 | ||
| 54 | lemmas monoI [intro?] = mono.intro | |
| 55 | and monoD [dest?] = mono.mono | |
| 56 | ||
| 57 | constdefs | |
| 58 | min :: "['a::ord, 'a] => 'a" | |
| 59 | "min a b == (if a <= b then a else b)" | |
| 60 | max :: "['a::ord, 'a] => 'a" | |
| 61 | "max a b == (if a <= b then b else a)" | |
| 62 | ||
| 63 | lemma min_leastL: "(!!x. least <= x) ==> min least x = least" | |
| 64 | by (simp add: min_def) | |
| 65 | ||
| 66 | lemma min_of_mono: | |
| 19527 | 67 | "(!!x y. (f x <= f y) = (x <= y)) ==> min (f m) (f n) = f (min m n)" | 
| 15524 | 68 | by (simp add: min_def) | 
| 69 | ||
| 70 | lemma max_leastL: "(!!x. least <= x) ==> max least x = x" | |
| 71 | by (simp add: max_def) | |
| 72 | ||
| 73 | lemma max_of_mono: | |
| 19527 | 74 | "(!!x y. (f x <= f y) = (x <= y)) ==> max (f m) (f n) = f (max m n)" | 
| 15524 | 75 | by (simp add: max_def) | 
| 76 | ||
| 77 | ||
| 78 | subsection "Orders" | |
| 79 | ||
| 80 | axclass order < ord | |
| 81 | order_refl [iff]: "x <= x" | |
| 82 | order_trans: "x <= y ==> y <= z ==> x <= z" | |
| 83 | order_antisym: "x <= y ==> y <= x ==> x = y" | |
| 84 | order_less_le: "(x < y) = (x <= y & x ~= y)" | |
| 85 | ||
| 86 | text{* Connection to locale: *}
 | |
| 87 | ||
| 15837 | 88 | interpretation order: | 
| 15780 | 89 | partial_order["op \<le> :: 'a::order \<Rightarrow> 'a \<Rightarrow> bool"] | 
| 15524 | 90 | apply(rule partial_order.intro) | 
| 91 | apply(rule order_refl, erule (1) order_trans, erule (1) order_antisym) | |
| 92 | done | |
| 93 | ||
| 94 | text {* Reflexivity. *}
 | |
| 95 | ||
| 96 | lemma order_eq_refl: "!!x::'a::order. x = y ==> x <= y" | |
| 97 |     -- {* This form is useful with the classical reasoner. *}
 | |
| 98 | apply (erule ssubst) | |
| 99 | apply (rule order_refl) | |
| 100 | done | |
| 101 | ||
| 102 | lemma order_less_irrefl [iff]: "~ x < (x::'a::order)" | |
| 103 | by (simp add: order_less_le) | |
| 104 | ||
| 105 | lemma order_le_less: "((x::'a::order) <= y) = (x < y | x = y)" | |
| 106 |     -- {* NOT suitable for iff, since it can cause PROOF FAILED. *}
 | |
| 107 | apply (simp add: order_less_le, blast) | |
| 108 | done | |
| 109 | ||
| 110 | lemmas order_le_imp_less_or_eq = order_le_less [THEN iffD1, standard] | |
| 111 | ||
| 112 | lemma order_less_imp_le: "!!x::'a::order. x < y ==> x <= y" | |
| 113 | by (simp add: order_less_le) | |
| 114 | ||
| 115 | ||
| 116 | text {* Asymmetry. *}
 | |
| 117 | ||
| 118 | lemma order_less_not_sym: "(x::'a::order) < y ==> ~ (y < x)" | |
| 119 | by (simp add: order_less_le order_antisym) | |
| 120 | ||
| 121 | lemma order_less_asym: "x < (y::'a::order) ==> (~P ==> y < x) ==> P" | |
| 122 | apply (drule order_less_not_sym) | |
| 123 | apply (erule contrapos_np, simp) | |
| 124 | done | |
| 125 | ||
| 126 | lemma order_eq_iff: "!!x::'a::order. (x = y) = (x \<le> y & y \<le> x)" | |
| 127 | by (blast intro: order_antisym) | |
| 128 | ||
| 129 | lemma order_antisym_conv: "(y::'a::order) <= x ==> (x <= y) = (x = y)" | |
| 130 | by(blast intro:order_antisym) | |
| 131 | ||
| 132 | text {* Transitivity. *}
 | |
| 133 | ||
| 134 | lemma order_less_trans: "!!x::'a::order. [| x < y; y < z |] ==> x < z" | |
| 135 | apply (simp add: order_less_le) | |
| 136 | apply (blast intro: order_trans order_antisym) | |
| 137 | done | |
| 138 | ||
| 139 | lemma order_le_less_trans: "!!x::'a::order. [| x <= y; y < z |] ==> x < z" | |
| 140 | apply (simp add: order_less_le) | |
| 141 | apply (blast intro: order_trans order_antisym) | |
| 142 | done | |
| 143 | ||
| 144 | lemma order_less_le_trans: "!!x::'a::order. [| x < y; y <= z |] ==> x < z" | |
| 145 | apply (simp add: order_less_le) | |
| 146 | apply (blast intro: order_trans order_antisym) | |
| 147 | done | |
| 148 | ||
| 149 | ||
| 150 | text {* Useful for simplification, but too risky to include by default. *}
 | |
| 151 | ||
| 152 | lemma order_less_imp_not_less: "(x::'a::order) < y ==> (~ y < x) = True" | |
| 153 | by (blast elim: order_less_asym) | |
| 154 | ||
| 155 | lemma order_less_imp_triv: "(x::'a::order) < y ==> (y < x --> P) = True" | |
| 156 | by (blast elim: order_less_asym) | |
| 157 | ||
| 158 | lemma order_less_imp_not_eq: "(x::'a::order) < y ==> (x = y) = False" | |
| 159 | by auto | |
| 160 | ||
| 161 | lemma order_less_imp_not_eq2: "(x::'a::order) < y ==> (y = x) = False" | |
| 162 | by auto | |
| 163 | ||
| 164 | ||
| 165 | text {* Other operators. *}
 | |
| 166 | ||
| 167 | lemma min_leastR: "(!!x::'a::order. least <= x) ==> min x least = least" | |
| 168 | apply (simp add: min_def) | |
| 169 | apply (blast intro: order_antisym) | |
| 170 | done | |
| 171 | ||
| 172 | lemma max_leastR: "(!!x::'a::order. least <= x) ==> max x least = x" | |
| 173 | apply (simp add: max_def) | |
| 174 | apply (blast intro: order_antisym) | |
| 175 | done | |
| 176 | ||
| 177 | ||
| 178 | subsection {* Transitivity rules for calculational reasoning *}
 | |
| 179 | ||
| 180 | ||
| 181 | lemma order_neq_le_trans: "a ~= b ==> (a::'a::order) <= b ==> a < b" | |
| 182 | by (simp add: order_less_le) | |
| 183 | ||
| 184 | lemma order_le_neq_trans: "(a::'a::order) <= b ==> a ~= b ==> a < b" | |
| 185 | by (simp add: order_less_le) | |
| 186 | ||
| 187 | lemma order_less_asym': "(a::'a::order) < b ==> b < a ==> P" | |
| 188 | by (rule order_less_asym) | |
| 189 | ||
| 190 | ||
| 191 | subsection {* Least value operator *}
 | |
| 192 | ||
| 193 | constdefs | |
| 194 |   Least :: "('a::ord => bool) => 'a"               (binder "LEAST " 10)
 | |
| 195 | "Least P == THE x. P x & (ALL y. P y --> x <= y)" | |
| 196 |     -- {* We can no longer use LeastM because the latter requires Hilbert-AC. *}
 | |
| 197 | ||
| 15950 | 198 | lemma LeastI2_order: | 
| 15524 | 199 | "[| P (x::'a::order); | 
| 200 | !!y. P y ==> x <= y; | |
| 201 | !!x. [| P x; ALL y. P y --> x \<le> y |] ==> Q x |] | |
| 202 | ==> Q (Least P)" | |
| 203 | apply (unfold Least_def) | |
| 204 | apply (rule theI2) | |
| 205 | apply (blast intro: order_antisym)+ | |
| 206 | done | |
| 207 | ||
| 208 | lemma Least_equality: | |
| 209 | "[| P (k::'a::order); !!x. P x ==> k <= x |] ==> (LEAST x. P x) = k" | |
| 210 | apply (simp add: Least_def) | |
| 211 | apply (rule the_equality) | |
| 212 | apply (auto intro!: order_antisym) | |
| 213 | done | |
| 214 | ||
| 215 | ||
| 216 | subsection "Linear / total orders" | |
| 217 | ||
| 218 | axclass linorder < order | |
| 219 | linorder_linear: "x <= y | y <= x" | |
| 220 | ||
| 221 | lemma linorder_less_linear: "!!x::'a::linorder. x<y | x=y | y<x" | |
| 222 | apply (simp add: order_less_le) | |
| 223 | apply (insert linorder_linear, blast) | |
| 224 | done | |
| 225 | ||
| 226 | lemma linorder_le_less_linear: "!!x::'a::linorder. x\<le>y | y<x" | |
| 227 | by (simp add: order_le_less linorder_less_linear) | |
| 228 | ||
| 229 | lemma linorder_le_cases [case_names le ge]: | |
| 230 | "((x::'a::linorder) \<le> y ==> P) ==> (y \<le> x ==> P) ==> P" | |
| 231 | by (insert linorder_linear, blast) | |
| 232 | ||
| 233 | lemma linorder_cases [case_names less equal greater]: | |
| 234 | "((x::'a::linorder) < y ==> P) ==> (x = y ==> P) ==> (y < x ==> P) ==> P" | |
| 235 | by (insert linorder_less_linear, blast) | |
| 236 | ||
| 237 | lemma linorder_not_less: "!!x::'a::linorder. (~ x < y) = (y <= x)" | |
| 238 | apply (simp add: order_less_le) | |
| 239 | apply (insert linorder_linear) | |
| 240 | apply (blast intro: order_antisym) | |
| 241 | done | |
| 242 | ||
| 243 | lemma linorder_not_le: "!!x::'a::linorder. (~ x <= y) = (y < x)" | |
| 244 | apply (simp add: order_less_le) | |
| 245 | apply (insert linorder_linear) | |
| 246 | apply (blast intro: order_antisym) | |
| 247 | done | |
| 248 | ||
| 249 | lemma linorder_neq_iff: "!!x::'a::linorder. (x ~= y) = (x<y | y<x)" | |
| 250 | by (cut_tac x = x and y = y in linorder_less_linear, auto) | |
| 251 | ||
| 252 | lemma linorder_neqE: "x ~= (y::'a::linorder) ==> (x < y ==> R) ==> (y < x ==> R) ==> R" | |
| 253 | by (simp add: linorder_neq_iff, blast) | |
| 254 | ||
| 255 | lemma linorder_antisym_conv1: "~ (x::'a::linorder) < y ==> (x <= y) = (x = y)" | |
| 256 | by(blast intro:order_antisym dest:linorder_not_less[THEN iffD1]) | |
| 257 | ||
| 258 | lemma linorder_antisym_conv2: "(x::'a::linorder) <= y ==> (~ x < y) = (x = y)" | |
| 259 | by(blast intro:order_antisym dest:linorder_not_less[THEN iffD1]) | |
| 260 | ||
| 261 | lemma linorder_antisym_conv3: "~ (y::'a::linorder) < x ==> (~ x < y) = (x = y)" | |
| 262 | by(blast intro:order_antisym dest:linorder_not_less[THEN iffD1]) | |
| 263 | ||
| 16796 | 264 | text{*Replacing the old Nat.leI*}
 | 
| 265 | lemma leI: "~ x < y ==> y <= (x::'a::linorder)" | |
| 266 | by (simp only: linorder_not_less) | |
| 267 | ||
| 268 | lemma leD: "y <= (x::'a::linorder) ==> ~ x < y" | |
| 269 | by (simp only: linorder_not_less) | |
| 270 | ||
| 271 | (*FIXME inappropriate name (or delete altogether)*) | |
| 272 | lemma not_leE: "~ y <= (x::'a::linorder) ==> x < y" | |
| 273 | by (simp only: linorder_not_le) | |
| 274 | ||
| 15524 | 275 | use "antisym_setup.ML"; | 
| 276 | setup antisym_setup | |
| 277 | ||
| 278 | subsection {* Setup of transitivity reasoner as Solver *}
 | |
| 279 | ||
| 280 | lemma less_imp_neq: "[| (x::'a::order) < y |] ==> x ~= y" | |
| 281 | by (erule contrapos_pn, erule subst, rule order_less_irrefl) | |
| 282 | ||
| 283 | lemma eq_neq_eq_imp_neq: "[| x = a ; a ~= b; b = y |] ==> x ~= y" | |
| 284 | by (erule subst, erule ssubst, assumption) | |
| 285 | ||
| 286 | ML_setup {*
 | |
| 287 | ||
| 288 | (* The setting up of Quasi_Tac serves as a demo. Since there is no | |
| 289 | class for quasi orders, the tactics Quasi_Tac.trans_tac and | |
| 290 | Quasi_Tac.quasi_tac are not of much use. *) | |
| 291 | ||
| 292 | fun decomp_gen sort sign (Trueprop $ t) = | |
| 15622 
4723248c982b
Transitivity reasoner ignores types amenable to linear arithmetic.
 ballarin parents: 
15531diff
changeset | 293 | let fun of_sort t = let val T = type_of t in | 
| 
4723248c982b
Transitivity reasoner ignores types amenable to linear arithmetic.
 ballarin parents: 
15531diff
changeset | 294 | (* exclude numeric types: linear arithmetic subsumes transitivity *) | 
| 
4723248c982b
Transitivity reasoner ignores types amenable to linear arithmetic.
 ballarin parents: 
15531diff
changeset | 295 | T <> HOLogic.natT andalso T <> HOLogic.intT andalso | 
| 
4723248c982b
Transitivity reasoner ignores types amenable to linear arithmetic.
 ballarin parents: 
15531diff
changeset | 296 | T <> HOLogic.realT andalso Sign.of_sort sign (T, sort) end | 
| 15524 | 297 |   fun dec (Const ("Not", _) $ t) = (
 | 
| 298 | case dec t of | |
| 15531 | 299 | NONE => NONE | 
| 300 | | SOME (t1, rel, t2) => SOME (t1, "~" ^ rel, t2)) | |
| 15524 | 301 | 	| dec (Const ("op =",  _) $ t1 $ t2) =
 | 
| 302 | if of_sort t1 | |
| 15531 | 303 | then SOME (t1, "=", t2) | 
| 304 | else NONE | |
| 19277 | 305 | 	| dec (Const ("Orderings.less_eq",  _) $ t1 $ t2) =
 | 
| 15524 | 306 | if of_sort t1 | 
| 15531 | 307 | then SOME (t1, "<=", t2) | 
| 308 | else NONE | |
| 19277 | 309 | 	| dec (Const ("Orderings.less",  _) $ t1 $ t2) =
 | 
| 15524 | 310 | if of_sort t1 | 
| 15531 | 311 | then SOME (t1, "<", t2) | 
| 312 | else NONE | |
| 313 | | dec _ = NONE | |
| 15524 | 314 | in dec t end; | 
| 315 | ||
| 316 | structure Quasi_Tac = Quasi_Tac_Fun ( | |
| 317 | struct | |
| 318 | val le_trans = thm "order_trans"; | |
| 319 | val le_refl = thm "order_refl"; | |
| 320 | val eqD1 = thm "order_eq_refl"; | |
| 321 | val eqD2 = thm "sym" RS thm "order_eq_refl"; | |
| 322 | val less_reflE = thm "order_less_irrefl" RS thm "notE"; | |
| 323 | val less_imp_le = thm "order_less_imp_le"; | |
| 324 | val le_neq_trans = thm "order_le_neq_trans"; | |
| 325 | val neq_le_trans = thm "order_neq_le_trans"; | |
| 326 | val less_imp_neq = thm "less_imp_neq"; | |
| 327 | val decomp_trans = decomp_gen ["Orderings.order"]; | |
| 328 | val decomp_quasi = decomp_gen ["Orderings.order"]; | |
| 329 | ||
| 330 | end); (* struct *) | |
| 331 | ||
| 332 | structure Order_Tac = Order_Tac_Fun ( | |
| 333 | struct | |
| 334 | val less_reflE = thm "order_less_irrefl" RS thm "notE"; | |
| 335 | val le_refl = thm "order_refl"; | |
| 336 | val less_imp_le = thm "order_less_imp_le"; | |
| 337 | val not_lessI = thm "linorder_not_less" RS thm "iffD2"; | |
| 338 | val not_leI = thm "linorder_not_le" RS thm "iffD2"; | |
| 339 | val not_lessD = thm "linorder_not_less" RS thm "iffD1"; | |
| 340 | val not_leD = thm "linorder_not_le" RS thm "iffD1"; | |
| 341 | val eqI = thm "order_antisym"; | |
| 342 | val eqD1 = thm "order_eq_refl"; | |
| 343 | val eqD2 = thm "sym" RS thm "order_eq_refl"; | |
| 344 | val less_trans = thm "order_less_trans"; | |
| 345 | val less_le_trans = thm "order_less_le_trans"; | |
| 346 | val le_less_trans = thm "order_le_less_trans"; | |
| 347 | val le_trans = thm "order_trans"; | |
| 348 | val le_neq_trans = thm "order_le_neq_trans"; | |
| 349 | val neq_le_trans = thm "order_neq_le_trans"; | |
| 350 | val less_imp_neq = thm "less_imp_neq"; | |
| 351 | val eq_neq_eq_imp_neq = thm "eq_neq_eq_imp_neq"; | |
| 16743 
21dbff595bf6
1) all theorems in Orderings can now be given as a parameter
 obua parents: 
16417diff
changeset | 352 | val not_sym = thm "not_sym"; | 
| 15524 | 353 | val decomp_part = decomp_gen ["Orderings.order"]; | 
| 354 | val decomp_lin = decomp_gen ["Orderings.linorder"]; | |
| 355 | ||
| 356 | end); (* struct *) | |
| 357 | ||
| 17876 | 358 | change_simpset (fn ss => ss | 
| 15524 | 359 | addSolver (mk_solver "Trans_linear" (fn _ => Order_Tac.linear_tac)) | 
| 17876 | 360 | addSolver (mk_solver "Trans_partial" (fn _ => Order_Tac.partial_tac))); | 
| 15524 | 361 | (* Adding the transitivity reasoners also as safe solvers showed a slight | 
| 362 | speed up, but the reasoning strength appears to be not higher (at least | |
| 363 | no breaking of additional proofs in the entire HOL distribution, as | |
| 364 | of 5 March 2004, was observed). *) | |
| 365 | *} | |
| 366 | ||
| 367 | (* Optional setup of methods *) | |
| 368 | ||
| 369 | (* | |
| 370 | method_setup trans_partial = | |
| 371 |   {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (Order_Tac.partial_tac)) *}
 | |
| 372 |   {* transitivity reasoner for partial orders *}	
 | |
| 373 | method_setup trans_linear = | |
| 374 |   {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (Order_Tac.linear_tac)) *}
 | |
| 375 |   {* transitivity reasoner for linear orders *}
 | |
| 376 | *) | |
| 377 | ||
| 378 | (* | |
| 379 | declare order.order_refl [simp del] order_less_irrefl [simp del] | |
| 380 | ||
| 381 | can currently not be removed, abel_cancel relies on it. | |
| 382 | *) | |
| 383 | ||
| 384 | ||
| 385 | subsection "Min and max on (linear) orders" | |
| 386 | ||
| 387 | text{* Instantiate locales: *}
 | |
| 388 | ||
| 15837 | 389 | interpretation min_max: | 
| 15780 | 390 | lower_semilattice["op \<le>" "min :: 'a::linorder \<Rightarrow> 'a \<Rightarrow> 'a"] | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 391 | apply unfold_locales | 
| 15524 | 392 | apply(simp add:min_def linorder_not_le order_less_imp_le) | 
| 393 | apply(simp add:min_def linorder_not_le order_less_imp_le) | |
| 394 | apply(simp add:min_def linorder_not_le order_less_imp_le) | |
| 395 | done | |
| 396 | ||
| 15837 | 397 | interpretation min_max: | 
| 15780 | 398 | upper_semilattice["op \<le>" "max :: 'a::linorder \<Rightarrow> 'a \<Rightarrow> 'a"] | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 399 | apply unfold_locales | 
| 15524 | 400 | apply(simp add: max_def linorder_not_le order_less_imp_le) | 
| 401 | apply(simp add: max_def linorder_not_le order_less_imp_le) | |
| 402 | apply(simp add: max_def linorder_not_le order_less_imp_le) | |
| 403 | done | |
| 404 | ||
| 15837 | 405 | interpretation min_max: | 
| 15780 | 406 | lattice["op \<le>" "min :: 'a::linorder \<Rightarrow> 'a \<Rightarrow> 'a" "max"] | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 407 | by unfold_locales | 
| 15524 | 408 | |
| 15837 | 409 | interpretation min_max: | 
| 15780 | 410 | distrib_lattice["op \<le>" "min :: 'a::linorder \<Rightarrow> 'a \<Rightarrow> 'a" "max"] | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 411 | apply unfold_locales | 
| 15524 | 412 | apply(rule_tac x=x and y=y in linorder_le_cases) | 
| 413 | apply(rule_tac x=x and y=z in linorder_le_cases) | |
| 414 | apply(rule_tac x=y and y=z in linorder_le_cases) | |
| 415 | apply(simp add:min_def max_def) | |
| 416 | apply(simp add:min_def max_def) | |
| 417 | apply(rule_tac x=y and y=z in linorder_le_cases) | |
| 418 | apply(simp add:min_def max_def) | |
| 419 | apply(simp add:min_def max_def) | |
| 420 | apply(rule_tac x=x and y=z in linorder_le_cases) | |
| 421 | apply(rule_tac x=y and y=z in linorder_le_cases) | |
| 422 | apply(simp add:min_def max_def) | |
| 423 | apply(simp add:min_def max_def) | |
| 424 | apply(rule_tac x=y and y=z in linorder_le_cases) | |
| 425 | apply(simp add:min_def max_def) | |
| 426 | apply(simp add:min_def max_def) | |
| 427 | done | |
| 428 | ||
| 429 | lemma le_max_iff_disj: "!!z::'a::linorder. (z <= max x y) = (z <= x | z <= y)" | |
| 430 | apply(simp add:max_def) | |
| 431 | apply (insert linorder_linear) | |
| 432 | apply (blast intro: order_trans) | |
| 433 | done | |
| 434 | ||
| 15780 | 435 | lemmas le_maxI1 = min_max.sup_ge1 | 
| 436 | lemmas le_maxI2 = min_max.sup_ge2 | |
| 15524 | 437 | |
| 438 | lemma less_max_iff_disj: "!!z::'a::linorder. (z < max x y) = (z < x | z < y)" | |
| 439 | apply (simp add: max_def order_le_less) | |
| 440 | apply (insert linorder_less_linear) | |
| 441 | apply (blast intro: order_less_trans) | |
| 442 | done | |
| 443 | ||
| 444 | lemma max_less_iff_conj [simp]: | |
| 445 | "!!z::'a::linorder. (max x y < z) = (x < z & y < z)" | |
| 446 | apply (simp add: order_le_less max_def) | |
| 447 | apply (insert linorder_less_linear) | |
| 448 | apply (blast intro: order_less_trans) | |
| 449 | done | |
| 15791 | 450 | |
| 15524 | 451 | lemma min_less_iff_conj [simp]: | 
| 452 | "!!z::'a::linorder. (z < min x y) = (z < x & z < y)" | |
| 453 | apply (simp add: order_le_less min_def) | |
| 454 | apply (insert linorder_less_linear) | |
| 455 | apply (blast intro: order_less_trans) | |
| 456 | done | |
| 457 | ||
| 458 | lemma min_le_iff_disj: "!!z::'a::linorder. (min x y <= z) = (x <= z | y <= z)" | |
| 459 | apply (simp add: min_def) | |
| 460 | apply (insert linorder_linear) | |
| 461 | apply (blast intro: order_trans) | |
| 462 | done | |
| 463 | ||
| 464 | lemma min_less_iff_disj: "!!z::'a::linorder. (min x y < z) = (x < z | y < z)" | |
| 465 | apply (simp add: min_def order_le_less) | |
| 466 | apply (insert linorder_less_linear) | |
| 467 | apply (blast intro: order_less_trans) | |
| 468 | done | |
| 469 | ||
| 15780 | 470 | lemmas max_ac = min_max.sup_assoc min_max.sup_commute | 
| 471 | mk_left_commute[of max,OF min_max.sup_assoc min_max.sup_commute] | |
| 15524 | 472 | |
| 15780 | 473 | lemmas min_ac = min_max.inf_assoc min_max.inf_commute | 
| 474 | mk_left_commute[of min,OF min_max.inf_assoc min_max.inf_commute] | |
| 15524 | 475 | |
| 476 | lemma split_min: | |
| 477 | "P (min (i::'a::linorder) j) = ((i <= j --> P(i)) & (~ i <= j --> P(j)))" | |
| 478 | by (simp add: min_def) | |
| 479 | ||
| 480 | lemma split_max: | |
| 481 | "P (max (i::'a::linorder) j) = ((i <= j --> P(j)) & (~ i <= j --> P(i)))" | |
| 482 | by (simp add: max_def) | |
| 483 | ||
| 484 | ||
| 485 | subsection "Bounded quantifiers" | |
| 486 | ||
| 487 | syntax | |
| 488 |   "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3ALL _<_./ _)"  [0, 0, 10] 10)
 | |
| 489 |   "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3EX _<_./ _)"  [0, 0, 10] 10)
 | |
| 490 |   "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3ALL _<=_./ _)" [0, 0, 10] 10)
 | |
| 491 |   "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3EX _<=_./ _)" [0, 0, 10] 10)
 | |
| 492 | ||
| 493 |   "_gtAll" :: "[idt, 'a, bool] => bool"   ("(3ALL _>_./ _)"  [0, 0, 10] 10)
 | |
| 494 |   "_gtEx"  :: "[idt, 'a, bool] => bool"   ("(3EX _>_./ _)"  [0, 0, 10] 10)
 | |
| 495 |   "_geAll"   :: "[idt, 'a, bool] => bool"   ("(3ALL _>=_./ _)" [0, 0, 10] 10)
 | |
| 496 |   "_geEx"    :: "[idt, 'a, bool] => bool"   ("(3EX _>=_./ _)" [0, 0, 10] 10)
 | |
| 497 | ||
| 498 | syntax (xsymbols) | |
| 499 |   "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_<_./ _)"  [0, 0, 10] 10)
 | |
| 500 |   "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_<_./ _)"  [0, 0, 10] 10)
 | |
| 501 |   "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10)
 | |
| 502 |   "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10)
 | |
| 503 | ||
| 504 |   "_gtAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_>_./ _)"  [0, 0, 10] 10)
 | |
| 505 |   "_gtEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_>_./ _)"  [0, 0, 10] 10)
 | |
| 506 |   "_geAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<ge>_./ _)" [0, 0, 10] 10)
 | |
| 507 |   "_geEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<ge>_./ _)" [0, 0, 10] 10)
 | |
| 508 | ||
| 509 | syntax (HOL) | |
| 510 |   "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3! _<_./ _)"  [0, 0, 10] 10)
 | |
| 511 |   "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3? _<_./ _)"  [0, 0, 10] 10)
 | |
| 512 |   "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3! _<=_./ _)" [0, 0, 10] 10)
 | |
| 513 |   "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3? _<=_./ _)" [0, 0, 10] 10)
 | |
| 514 | ||
| 515 | syntax (HTML output) | |
| 516 |   "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_<_./ _)"  [0, 0, 10] 10)
 | |
| 517 |   "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_<_./ _)"  [0, 0, 10] 10)
 | |
| 518 |   "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10)
 | |
| 519 |   "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10)
 | |
| 520 | ||
| 521 |   "_gtAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_>_./ _)"  [0, 0, 10] 10)
 | |
| 522 |   "_gtEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_>_./ _)"  [0, 0, 10] 10)
 | |
| 523 |   "_geAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<ge>_./ _)" [0, 0, 10] 10)
 | |
| 524 |   "_geEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<ge>_./ _)" [0, 0, 10] 10)
 | |
| 525 | ||
| 526 | translations | |
| 527 | "ALL x<y. P" => "ALL x. x < y --> P" | |
| 528 | "EX x<y. P" => "EX x. x < y & P" | |
| 529 | "ALL x<=y. P" => "ALL x. x <= y --> P" | |
| 530 | "EX x<=y. P" => "EX x. x <= y & P" | |
| 531 | "ALL x>y. P" => "ALL x. x > y --> P" | |
| 532 | "EX x>y. P" => "EX x. x > y & P" | |
| 533 | "ALL x>=y. P" => "ALL x. x >= y --> P" | |
| 534 | "EX x>=y. P" => "EX x. x >= y & P" | |
| 535 | ||
| 536 | print_translation {*
 | |
| 537 | let | |
| 538 | fun mk v v' q n P = | |
| 16861 | 539 | if v=v' andalso not (v mem (map fst (Term.add_frees n []))) | 
| 15524 | 540 | then Syntax.const q $ Syntax.mark_bound v' $ n $ P else raise Match; | 
| 541 |   fun all_tr' [Const ("_bound",_) $ Free (v,_),
 | |
| 19637 | 542 |                Const("op -->",_) $ (Const ("less",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] =
 | 
| 15524 | 543 | mk v v' "_lessAll" n P | 
| 544 | ||
| 545 |   | all_tr' [Const ("_bound",_) $ Free (v,_),
 | |
| 19637 | 546 |                Const("op -->",_) $ (Const ("less_eq",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] =
 | 
| 15524 | 547 | mk v v' "_leAll" n P | 
| 548 | ||
| 549 |   | all_tr' [Const ("_bound",_) $ Free (v,_),
 | |
| 19637 | 550 |                Const("op -->",_) $ (Const ("less",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] =
 | 
| 15524 | 551 | mk v v' "_gtAll" n P | 
| 552 | ||
| 553 |   | all_tr' [Const ("_bound",_) $ Free (v,_),
 | |
| 19637 | 554 |                Const("op -->",_) $ (Const ("less_eq",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] =
 | 
| 15524 | 555 | mk v v' "_geAll" n P; | 
| 556 | ||
| 557 |   fun ex_tr' [Const ("_bound",_) $ Free (v,_),
 | |
| 19637 | 558 |                Const("op &",_) $ (Const ("less",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] =
 | 
| 15524 | 559 | mk v v' "_lessEx" n P | 
| 560 | ||
| 561 |   | ex_tr' [Const ("_bound",_) $ Free (v,_),
 | |
| 19637 | 562 |                Const("op &",_) $ (Const ("less_eq",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] =
 | 
| 15524 | 563 | mk v v' "_leEx" n P | 
| 564 | ||
| 565 |   | ex_tr' [Const ("_bound",_) $ Free (v,_),
 | |
| 19637 | 566 |                Const("op &",_) $ (Const ("less",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] =
 | 
| 15524 | 567 | mk v v' "_gtEx" n P | 
| 568 | ||
| 569 |   | ex_tr' [Const ("_bound",_) $ Free (v,_),
 | |
| 19637 | 570 |                Const("op &",_) $ (Const ("less_eq",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] =
 | 
| 15524 | 571 | mk v v' "_geEx" n P | 
| 572 | in | |
| 573 | [("ALL ", all_tr'), ("EX ", ex_tr')]
 | |
| 574 | end | |
| 575 | *} | |
| 576 | ||
| 17012 | 577 | subsection {* Extra transitivity rules *}
 | 
| 578 | ||
| 579 | text {* These support proving chains of decreasing inequalities
 | |
| 580 | a >= b >= c ... in Isar proofs. *} | |
| 581 | ||
| 582 | lemma xt1: "a = b ==> b > c ==> a > c" | |
| 583 | by simp | |
| 584 | ||
| 585 | lemma xt2: "a > b ==> b = c ==> a > c" | |
| 586 | by simp | |
| 587 | ||
| 588 | lemma xt3: "a = b ==> b >= c ==> a >= c" | |
| 589 | by simp | |
| 590 | ||
| 591 | lemma xt4: "a >= b ==> b = c ==> a >= c" | |
| 592 | by simp | |
| 593 | ||
| 594 | lemma xt5: "(x::'a::order) >= y ==> y >= x ==> x = y" | |
| 595 | by simp | |
| 596 | ||
| 597 | lemma xt6: "(x::'a::order) >= y ==> y >= z ==> x >= z" | |
| 598 | by simp | |
| 599 | ||
| 600 | lemma xt7: "(x::'a::order) > y ==> y >= z ==> x > z" | |
| 601 | by simp | |
| 602 | ||
| 603 | lemma xt8: "(x::'a::order) >= y ==> y > z ==> x > z" | |
| 604 | by simp | |
| 605 | ||
| 606 | lemma xt9: "(a::'a::order) > b ==> b > a ==> ?P" | |
| 607 | by simp | |
| 608 | ||
| 609 | lemma xt10: "(x::'a::order) > y ==> y > z ==> x > z" | |
| 610 | by simp | |
| 611 | ||
| 612 | lemma xt11: "(a::'a::order) >= b ==> a ~= b ==> a > b" | |
| 613 | by simp | |
| 614 | ||
| 615 | lemma xt12: "(a::'a::order) ~= b ==> a >= b ==> a > b" | |
| 616 | by simp | |
| 617 | ||
| 618 | lemma xt13: "a = f b ==> b > c ==> (!!x y. x > y ==> f x > f y) ==> | |
| 619 | a > f c" | |
| 620 | by simp | |
| 621 | ||
| 622 | lemma xt14: "a > b ==> f b = c ==> (!!x y. x > y ==> f x > f y) ==> | |
| 623 | f a > c" | |
| 624 | by auto | |
| 625 | ||
| 626 | lemma xt15: "a = f b ==> b >= c ==> (!!x y. x >= y ==> f x >= f y) ==> | |
| 627 | a >= f c" | |
| 628 | by simp | |
| 629 | ||
| 630 | lemma xt16: "a >= b ==> f b = c ==> (!! x y. x >= y ==> f x >= f y) ==> | |
| 631 | f a >= c" | |
| 632 | by auto | |
| 633 | ||
| 634 | lemma xt17: "(a::'a::order) >= f b ==> b >= c ==> | |
| 635 | (!!x y. x >= y ==> f x >= f y) ==> a >= f c" | |
| 636 | by (subgoal_tac "f b >= f c", force, force) | |
| 637 | ||
| 638 | lemma xt18: "(a::'a::order) >= b ==> (f b::'b::order) >= c ==> | |
| 639 | (!!x y. x >= y ==> f x >= f y) ==> f a >= c" | |
| 640 | by (subgoal_tac "f a >= f b", force, force) | |
| 641 | ||
| 642 | lemma xt19: "(a::'a::order) > f b ==> (b::'b::order) >= c ==> | |
| 643 | (!!x y. x >= y ==> f x >= f y) ==> a > f c" | |
| 644 | by (subgoal_tac "f b >= f c", force, force) | |
| 645 | ||
| 646 | lemma xt20: "(a::'a::order) > b ==> (f b::'b::order) >= c==> | |
| 647 | (!!x y. x > y ==> f x > f y) ==> f a > c" | |
| 648 | by (subgoal_tac "f a > f b", force, force) | |
| 649 | ||
| 650 | lemma xt21: "(a::'a::order) >= f b ==> b > c ==> | |
| 651 | (!!x y. x > y ==> f x > f y) ==> a > f c" | |
| 652 | by (subgoal_tac "f b > f c", force, force) | |
| 653 | ||
| 654 | lemma xt22: "(a::'a::order) >= b ==> (f b::'b::order) > c ==> | |
| 655 | (!!x y. x >= y ==> f x >= f y) ==> f a > c" | |
| 656 | by (subgoal_tac "f a >= f b", force, force) | |
| 657 | ||
| 658 | lemma xt23: "(a::'a::order) > f b ==> (b::'b::order) > c ==> | |
| 659 | (!!x y. x > y ==> f x > f y) ==> a > f c" | |
| 660 | by (subgoal_tac "f b > f c", force, force) | |
| 661 | ||
| 662 | lemma xt24: "(a::'a::order) > b ==> (f b::'b::order) > c ==> | |
| 663 | (!!x y. x > y ==> f x > f y) ==> f a > c" | |
| 664 | by (subgoal_tac "f a > f b", force, force) | |
| 665 | ||
| 666 | ||
| 667 | lemmas xtrans = xt1 xt2 xt3 xt4 xt5 xt6 xt7 xt8 xt9 xt10 xt11 xt12 | |
| 668 | xt13 xt14 xt15 xt15 xt17 xt18 xt19 xt20 xt21 xt22 xt23 xt24 | |
| 669 | ||
| 670 | (* | |
| 671 | Since "a >= b" abbreviates "b <= a", the abbreviation "..." stands | |
| 672 | for the wrong thing in an Isar proof. | |
| 673 | ||
| 674 | The extra transitivity rules can be used as follows: | |
| 675 | ||
| 676 | lemma "(a::'a::order) > z" | |
| 677 | proof - | |
| 678 | have "a >= b" (is "_ >= ?rhs") | |
| 679 | sorry | |
| 680 | also have "?rhs >= c" (is "_ >= ?rhs") | |
| 681 | sorry | |
| 682 | also (xtrans) have "?rhs = d" (is "_ = ?rhs") | |
| 683 | sorry | |
| 684 | also (xtrans) have "?rhs >= e" (is "_ >= ?rhs") | |
| 685 | sorry | |
| 686 | also (xtrans) have "?rhs > f" (is "_ > ?rhs") | |
| 687 | sorry | |
| 688 | also (xtrans) have "?rhs > z" | |
| 689 | sorry | |
| 690 | finally (xtrans) show ?thesis . | |
| 691 | qed | |
| 692 | ||
| 693 | Alternatively, one can use "declare xtrans [trans]" and then | |
| 694 | leave out the "(xtrans)" above. | |
| 695 | *) | |
| 696 | ||
| 15524 | 697 | end |