author | wenzelm |
Wed, 07 Jun 2006 00:57:14 +0200 | |
changeset 19801 | b2af2549efd1 |
parent 19656 | 09be06943252 |
child 19931 | fb32b43e7f80 |
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:
19637
diff
changeset
|
21 |
less :: "['a::ord, 'a] => bool" |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
changeset
|
22 |
less_eq :: "['a::ord, 'a] => bool" |
15524 | 23 |
|
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
changeset
|
24 |
const_syntax |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
changeset
|
25 |
less ("op <") |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
changeset
|
26 |
less ("(_/ < _)" [50, 51] 50) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
changeset
|
27 |
less_eq ("op <=") |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
changeset
|
28 |
less_eq ("(_/ <= _)" [50, 51] 50) |
15524 | 29 |
|
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
changeset
|
30 |
const_syntax (xsymbols) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
changeset
|
31 |
less_eq ("op \<le>") |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
changeset
|
32 |
less_eq ("(_/ \<le> _)" [50, 51] 50) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
changeset
|
33 |
|
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
changeset
|
34 |
const_syntax (HTML output) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
changeset
|
35 |
less_eq ("op \<le>") |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
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:
19637
diff
changeset
|
44 |
const_syntax (xsymbols) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
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:
15531
diff
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:
15531
diff
changeset
|
294 |
(* exclude numeric types: linear arithmetic subsumes transitivity *) |
4723248c982b
Transitivity reasoner ignores types amenable to linear arithmetic.
ballarin
parents:
15531
diff
changeset
|
295 |
T <> HOLogic.natT andalso T <> HOLogic.intT andalso |
4723248c982b
Transitivity reasoner ignores types amenable to linear arithmetic.
ballarin
parents:
15531
diff
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:
16417
diff
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"] |
15524 | 391 |
apply(rule lower_semilattice_axioms.intro) |
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"] |
399 |
apply - |
|
15524 | 400 |
apply(rule upper_semilattice_axioms.intro) |
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 |
apply(simp add: max_def linorder_not_le order_less_imp_le) |
|
404 |
done |
|
405 |
||
15837 | 406 |
interpretation min_max: |
15780 | 407 |
lattice["op \<le>" "min :: 'a::linorder \<Rightarrow> 'a \<Rightarrow> 'a" "max"] |
408 |
. |
|
15524 | 409 |
|
15837 | 410 |
interpretation min_max: |
15780 | 411 |
distrib_lattice["op \<le>" "min :: 'a::linorder \<Rightarrow> 'a \<Rightarrow> 'a" "max"] |
15524 | 412 |
apply(rule distrib_lattice_axioms.intro) |
413 |
apply(rule_tac x=x and y=y in linorder_le_cases) |
|
414 |
apply(rule_tac x=x and y=z in linorder_le_cases) |
|
415 |
apply(rule_tac x=y and y=z in linorder_le_cases) |
|
416 |
apply(simp add:min_def max_def) |
|
417 |
apply(simp add:min_def max_def) |
|
418 |
apply(rule_tac x=y and y=z in linorder_le_cases) |
|
419 |
apply(simp add:min_def max_def) |
|
420 |
apply(simp add:min_def max_def) |
|
421 |
apply(rule_tac x=x and y=z in linorder_le_cases) |
|
422 |
apply(rule_tac x=y and y=z in linorder_le_cases) |
|
423 |
apply(simp add:min_def max_def) |
|
424 |
apply(simp add:min_def max_def) |
|
425 |
apply(rule_tac x=y and y=z in linorder_le_cases) |
|
426 |
apply(simp add:min_def max_def) |
|
427 |
apply(simp add:min_def max_def) |
|
428 |
done |
|
429 |
||
430 |
lemma le_max_iff_disj: "!!z::'a::linorder. (z <= max x y) = (z <= x | z <= y)" |
|
431 |
apply(simp add:max_def) |
|
432 |
apply (insert linorder_linear) |
|
433 |
apply (blast intro: order_trans) |
|
434 |
done |
|
435 |
||
15780 | 436 |
lemmas le_maxI1 = min_max.sup_ge1 |
437 |
lemmas le_maxI2 = min_max.sup_ge2 |
|
15524 | 438 |
|
439 |
lemma less_max_iff_disj: "!!z::'a::linorder. (z < max x y) = (z < x | z < y)" |
|
440 |
apply (simp add: max_def order_le_less) |
|
441 |
apply (insert linorder_less_linear) |
|
442 |
apply (blast intro: order_less_trans) |
|
443 |
done |
|
444 |
||
445 |
lemma max_less_iff_conj [simp]: |
|
446 |
"!!z::'a::linorder. (max x y < z) = (x < z & y < z)" |
|
447 |
apply (simp add: order_le_less max_def) |
|
448 |
apply (insert linorder_less_linear) |
|
449 |
apply (blast intro: order_less_trans) |
|
450 |
done |
|
15791 | 451 |
|
15524 | 452 |
lemma min_less_iff_conj [simp]: |
453 |
"!!z::'a::linorder. (z < min x y) = (z < x & z < y)" |
|
454 |
apply (simp add: order_le_less min_def) |
|
455 |
apply (insert linorder_less_linear) |
|
456 |
apply (blast intro: order_less_trans) |
|
457 |
done |
|
458 |
||
459 |
lemma min_le_iff_disj: "!!z::'a::linorder. (min x y <= z) = (x <= z | y <= z)" |
|
460 |
apply (simp add: min_def) |
|
461 |
apply (insert linorder_linear) |
|
462 |
apply (blast intro: order_trans) |
|
463 |
done |
|
464 |
||
465 |
lemma min_less_iff_disj: "!!z::'a::linorder. (min x y < z) = (x < z | y < z)" |
|
466 |
apply (simp add: min_def order_le_less) |
|
467 |
apply (insert linorder_less_linear) |
|
468 |
apply (blast intro: order_less_trans) |
|
469 |
done |
|
470 |
||
15780 | 471 |
lemmas max_ac = min_max.sup_assoc min_max.sup_commute |
472 |
mk_left_commute[of max,OF min_max.sup_assoc min_max.sup_commute] |
|
15524 | 473 |
|
15780 | 474 |
lemmas min_ac = min_max.inf_assoc min_max.inf_commute |
475 |
mk_left_commute[of min,OF min_max.inf_assoc min_max.inf_commute] |
|
15524 | 476 |
|
477 |
lemma split_min: |
|
478 |
"P (min (i::'a::linorder) j) = ((i <= j --> P(i)) & (~ i <= j --> P(j)))" |
|
479 |
by (simp add: min_def) |
|
480 |
||
481 |
lemma split_max: |
|
482 |
"P (max (i::'a::linorder) j) = ((i <= j --> P(j)) & (~ i <= j --> P(i)))" |
|
483 |
by (simp add: max_def) |
|
484 |
||
485 |
||
486 |
subsection "Bounded quantifiers" |
|
487 |
||
488 |
syntax |
|
489 |
"_lessAll" :: "[idt, 'a, bool] => bool" ("(3ALL _<_./ _)" [0, 0, 10] 10) |
|
490 |
"_lessEx" :: "[idt, 'a, bool] => bool" ("(3EX _<_./ _)" [0, 0, 10] 10) |
|
491 |
"_leAll" :: "[idt, 'a, bool] => bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10) |
|
492 |
"_leEx" :: "[idt, 'a, bool] => bool" ("(3EX _<=_./ _)" [0, 0, 10] 10) |
|
493 |
||
494 |
"_gtAll" :: "[idt, 'a, bool] => bool" ("(3ALL _>_./ _)" [0, 0, 10] 10) |
|
495 |
"_gtEx" :: "[idt, 'a, bool] => bool" ("(3EX _>_./ _)" [0, 0, 10] 10) |
|
496 |
"_geAll" :: "[idt, 'a, bool] => bool" ("(3ALL _>=_./ _)" [0, 0, 10] 10) |
|
497 |
"_geEx" :: "[idt, 'a, bool] => bool" ("(3EX _>=_./ _)" [0, 0, 10] 10) |
|
498 |
||
499 |
syntax (xsymbols) |
|
500 |
"_lessAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_<_./ _)" [0, 0, 10] 10) |
|
501 |
"_lessEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_<_./ _)" [0, 0, 10] 10) |
|
502 |
"_leAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10) |
|
503 |
"_leEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10) |
|
504 |
||
505 |
"_gtAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_>_./ _)" [0, 0, 10] 10) |
|
506 |
"_gtEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_>_./ _)" [0, 0, 10] 10) |
|
507 |
"_geAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_\<ge>_./ _)" [0, 0, 10] 10) |
|
508 |
"_geEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_\<ge>_./ _)" [0, 0, 10] 10) |
|
509 |
||
510 |
syntax (HOL) |
|
511 |
"_lessAll" :: "[idt, 'a, bool] => bool" ("(3! _<_./ _)" [0, 0, 10] 10) |
|
512 |
"_lessEx" :: "[idt, 'a, bool] => bool" ("(3? _<_./ _)" [0, 0, 10] 10) |
|
513 |
"_leAll" :: "[idt, 'a, bool] => bool" ("(3! _<=_./ _)" [0, 0, 10] 10) |
|
514 |
"_leEx" :: "[idt, 'a, bool] => bool" ("(3? _<=_./ _)" [0, 0, 10] 10) |
|
515 |
||
516 |
syntax (HTML output) |
|
517 |
"_lessAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_<_./ _)" [0, 0, 10] 10) |
|
518 |
"_lessEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_<_./ _)" [0, 0, 10] 10) |
|
519 |
"_leAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10) |
|
520 |
"_leEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10) |
|
521 |
||
522 |
"_gtAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_>_./ _)" [0, 0, 10] 10) |
|
523 |
"_gtEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_>_./ _)" [0, 0, 10] 10) |
|
524 |
"_geAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_\<ge>_./ _)" [0, 0, 10] 10) |
|
525 |
"_geEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_\<ge>_./ _)" [0, 0, 10] 10) |
|
526 |
||
527 |
translations |
|
528 |
"ALL x<y. P" => "ALL x. x < y --> P" |
|
529 |
"EX x<y. P" => "EX x. x < y & P" |
|
530 |
"ALL x<=y. P" => "ALL x. x <= y --> P" |
|
531 |
"EX x<=y. P" => "EX x. x <= y & P" |
|
532 |
"ALL x>y. P" => "ALL x. x > y --> P" |
|
533 |
"EX x>y. P" => "EX x. x > y & P" |
|
534 |
"ALL x>=y. P" => "ALL x. x >= y --> P" |
|
535 |
"EX x>=y. P" => "EX x. x >= y & P" |
|
536 |
||
537 |
print_translation {* |
|
538 |
let |
|
539 |
fun mk v v' q n P = |
|
16861 | 540 |
if v=v' andalso not (v mem (map fst (Term.add_frees n []))) |
15524 | 541 |
then Syntax.const q $ Syntax.mark_bound v' $ n $ P else raise Match; |
542 |
fun all_tr' [Const ("_bound",_) $ Free (v,_), |
|
19637 | 543 |
Const("op -->",_) $ (Const ("less",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] = |
15524 | 544 |
mk v v' "_lessAll" n P |
545 |
||
546 |
| all_tr' [Const ("_bound",_) $ Free (v,_), |
|
19637 | 547 |
Const("op -->",_) $ (Const ("less_eq",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] = |
15524 | 548 |
mk v v' "_leAll" n P |
549 |
||
550 |
| all_tr' [Const ("_bound",_) $ Free (v,_), |
|
19637 | 551 |
Const("op -->",_) $ (Const ("less",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] = |
15524 | 552 |
mk v v' "_gtAll" n P |
553 |
||
554 |
| all_tr' [Const ("_bound",_) $ Free (v,_), |
|
19637 | 555 |
Const("op -->",_) $ (Const ("less_eq",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] = |
15524 | 556 |
mk v v' "_geAll" n P; |
557 |
||
558 |
fun ex_tr' [Const ("_bound",_) $ Free (v,_), |
|
19637 | 559 |
Const("op &",_) $ (Const ("less",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] = |
15524 | 560 |
mk v v' "_lessEx" n P |
561 |
||
562 |
| ex_tr' [Const ("_bound",_) $ Free (v,_), |
|
19637 | 563 |
Const("op &",_) $ (Const ("less_eq",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] = |
15524 | 564 |
mk v v' "_leEx" n P |
565 |
||
566 |
| ex_tr' [Const ("_bound",_) $ Free (v,_), |
|
19637 | 567 |
Const("op &",_) $ (Const ("less",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] = |
15524 | 568 |
mk v v' "_gtEx" n P |
569 |
||
570 |
| ex_tr' [Const ("_bound",_) $ Free (v,_), |
|
19637 | 571 |
Const("op &",_) $ (Const ("less_eq",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] = |
15524 | 572 |
mk v v' "_geEx" n P |
573 |
in |
|
574 |
[("ALL ", all_tr'), ("EX ", ex_tr')] |
|
575 |
end |
|
576 |
*} |
|
577 |
||
17012 | 578 |
subsection {* Extra transitivity rules *} |
579 |
||
580 |
text {* These support proving chains of decreasing inequalities |
|
581 |
a >= b >= c ... in Isar proofs. *} |
|
582 |
||
583 |
lemma xt1: "a = b ==> b > c ==> a > c" |
|
584 |
by simp |
|
585 |
||
586 |
lemma xt2: "a > b ==> b = c ==> a > c" |
|
587 |
by simp |
|
588 |
||
589 |
lemma xt3: "a = b ==> b >= c ==> a >= c" |
|
590 |
by simp |
|
591 |
||
592 |
lemma xt4: "a >= b ==> b = c ==> a >= c" |
|
593 |
by simp |
|
594 |
||
595 |
lemma xt5: "(x::'a::order) >= y ==> y >= x ==> x = y" |
|
596 |
by simp |
|
597 |
||
598 |
lemma xt6: "(x::'a::order) >= y ==> y >= z ==> x >= z" |
|
599 |
by simp |
|
600 |
||
601 |
lemma xt7: "(x::'a::order) > y ==> y >= z ==> x > z" |
|
602 |
by simp |
|
603 |
||
604 |
lemma xt8: "(x::'a::order) >= y ==> y > z ==> x > z" |
|
605 |
by simp |
|
606 |
||
607 |
lemma xt9: "(a::'a::order) > b ==> b > a ==> ?P" |
|
608 |
by simp |
|
609 |
||
610 |
lemma xt10: "(x::'a::order) > y ==> y > z ==> x > z" |
|
611 |
by simp |
|
612 |
||
613 |
lemma xt11: "(a::'a::order) >= b ==> a ~= b ==> a > b" |
|
614 |
by simp |
|
615 |
||
616 |
lemma xt12: "(a::'a::order) ~= b ==> a >= b ==> a > b" |
|
617 |
by simp |
|
618 |
||
619 |
lemma xt13: "a = f b ==> b > c ==> (!!x y. x > y ==> f x > f y) ==> |
|
620 |
a > f c" |
|
621 |
by simp |
|
622 |
||
623 |
lemma xt14: "a > b ==> f b = c ==> (!!x y. x > y ==> f x > f y) ==> |
|
624 |
f a > c" |
|
625 |
by auto |
|
626 |
||
627 |
lemma xt15: "a = f b ==> b >= c ==> (!!x y. x >= y ==> f x >= f y) ==> |
|
628 |
a >= f c" |
|
629 |
by simp |
|
630 |
||
631 |
lemma xt16: "a >= b ==> f b = c ==> (!! x y. x >= y ==> f x >= f y) ==> |
|
632 |
f a >= c" |
|
633 |
by auto |
|
634 |
||
635 |
lemma xt17: "(a::'a::order) >= f b ==> b >= c ==> |
|
636 |
(!!x y. x >= y ==> f x >= f y) ==> a >= f c" |
|
637 |
by (subgoal_tac "f b >= f c", force, force) |
|
638 |
||
639 |
lemma xt18: "(a::'a::order) >= b ==> (f b::'b::order) >= c ==> |
|
640 |
(!!x y. x >= y ==> f x >= f y) ==> f a >= c" |
|
641 |
by (subgoal_tac "f a >= f b", force, force) |
|
642 |
||
643 |
lemma xt19: "(a::'a::order) > f b ==> (b::'b::order) >= c ==> |
|
644 |
(!!x y. x >= y ==> f x >= f y) ==> a > f c" |
|
645 |
by (subgoal_tac "f b >= f c", force, force) |
|
646 |
||
647 |
lemma xt20: "(a::'a::order) > b ==> (f b::'b::order) >= c==> |
|
648 |
(!!x y. x > y ==> f x > f y) ==> f a > c" |
|
649 |
by (subgoal_tac "f a > f b", force, force) |
|
650 |
||
651 |
lemma xt21: "(a::'a::order) >= f b ==> b > c ==> |
|
652 |
(!!x y. x > y ==> f x > f y) ==> a > f c" |
|
653 |
by (subgoal_tac "f b > f c", force, force) |
|
654 |
||
655 |
lemma xt22: "(a::'a::order) >= b ==> (f b::'b::order) > c ==> |
|
656 |
(!!x y. x >= y ==> f x >= f y) ==> f a > c" |
|
657 |
by (subgoal_tac "f a >= f b", force, force) |
|
658 |
||
659 |
lemma xt23: "(a::'a::order) > f b ==> (b::'b::order) > c ==> |
|
660 |
(!!x y. x > y ==> f x > f y) ==> a > f c" |
|
661 |
by (subgoal_tac "f b > f c", force, force) |
|
662 |
||
663 |
lemma xt24: "(a::'a::order) > b ==> (f b::'b::order) > c ==> |
|
664 |
(!!x y. x > y ==> f x > f y) ==> f a > c" |
|
665 |
by (subgoal_tac "f a > f b", force, force) |
|
666 |
||
667 |
||
668 |
lemmas xtrans = xt1 xt2 xt3 xt4 xt5 xt6 xt7 xt8 xt9 xt10 xt11 xt12 |
|
669 |
xt13 xt14 xt15 xt15 xt17 xt18 xt19 xt20 xt21 xt22 xt23 xt24 |
|
670 |
||
671 |
(* |
|
672 |
Since "a >= b" abbreviates "b <= a", the abbreviation "..." stands |
|
673 |
for the wrong thing in an Isar proof. |
|
674 |
||
675 |
The extra transitivity rules can be used as follows: |
|
676 |
||
677 |
lemma "(a::'a::order) > z" |
|
678 |
proof - |
|
679 |
have "a >= b" (is "_ >= ?rhs") |
|
680 |
sorry |
|
681 |
also have "?rhs >= c" (is "_ >= ?rhs") |
|
682 |
sorry |
|
683 |
also (xtrans) have "?rhs = d" (is "_ = ?rhs") |
|
684 |
sorry |
|
685 |
also (xtrans) have "?rhs >= e" (is "_ >= ?rhs") |
|
686 |
sorry |
|
687 |
also (xtrans) have "?rhs > f" (is "_ > ?rhs") |
|
688 |
sorry |
|
689 |
also (xtrans) have "?rhs > z" |
|
690 |
sorry |
|
691 |
finally (xtrans) show ?thesis . |
|
692 |
qed |
|
693 |
||
694 |
Alternatively, one can use "declare xtrans [trans]" and then |
|
695 |
leave out the "(xtrans)" above. |
|
696 |
*) |
|
697 |
||
15524 | 698 |
end |