| author | paulson | 
| Tue, 17 May 2005 17:01:35 +0200 | |
| changeset 15977 | aa6744dd998e | 
| parent 15950 | 5c067c956a20 | 
| child 16121 | a80aa66d2271 | 
| 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  | 
|
12  | 
files ("antisym_setup.ML")
 | 
|
13  | 
begin  | 
|
14  | 
||
15  | 
subsection {* Order signatures and orders *}
 | 
|
16  | 
||
17  | 
axclass  | 
|
18  | 
ord < type  | 
|
19  | 
||
20  | 
syntax  | 
|
21  | 
  "op <"        :: "['a::ord, 'a] => bool"             ("op <")
 | 
|
22  | 
  "op <="       :: "['a::ord, 'a] => bool"             ("op <=")
 | 
|
23  | 
||
24  | 
global  | 
|
25  | 
||
26  | 
consts  | 
|
27  | 
  "op <"        :: "['a::ord, 'a] => bool"             ("(_/ < _)"  [50, 51] 50)
 | 
|
28  | 
  "op <="       :: "['a::ord, 'a] => bool"             ("(_/ <= _)" [50, 51] 50)
 | 
|
29  | 
||
30  | 
local  | 
|
31  | 
||
32  | 
syntax (xsymbols)  | 
|
33  | 
  "op <="       :: "['a::ord, 'a] => bool"             ("op \<le>")
 | 
|
34  | 
  "op <="       :: "['a::ord, 'a] => bool"             ("(_/ \<le> _)"  [50, 51] 50)
 | 
|
35  | 
||
36  | 
syntax (HTML output)  | 
|
37  | 
  "op <="       :: "['a::ord, 'a] => bool"             ("op \<le>")
 | 
|
38  | 
  "op <="       :: "['a::ord, 'a] => bool"             ("(_/ \<le> _)"  [50, 51] 50)
 | 
|
39  | 
||
40  | 
text{* Syntactic sugar: *}
 | 
|
41  | 
||
| 15822 | 42  | 
syntax  | 
| 15524 | 43  | 
"_gt" :: "'a::ord => 'a => bool" (infixl ">" 50)  | 
44  | 
"_ge" :: "'a::ord => 'a => bool" (infixl ">=" 50)  | 
|
45  | 
translations  | 
|
46  | 
"x > y" => "y < x"  | 
|
47  | 
"x >= y" => "y <= x"  | 
|
48  | 
||
49  | 
syntax (xsymbols)  | 
|
50  | 
"_ge" :: "'a::ord => 'a => bool" (infixl "\<ge>" 50)  | 
|
51  | 
||
52  | 
syntax (HTML output)  | 
|
53  | 
"_ge" :: "['a::ord, 'a] => bool" (infixl "\<ge>" 50)  | 
|
54  | 
||
55  | 
||
56  | 
subsection {* Monotonicity *}
 | 
|
57  | 
||
58  | 
locale mono =  | 
|
59  | 
fixes f  | 
|
60  | 
assumes mono: "A <= B ==> f A <= f B"  | 
|
61  | 
||
62  | 
lemmas monoI [intro?] = mono.intro  | 
|
63  | 
and monoD [dest?] = mono.mono  | 
|
64  | 
||
65  | 
constdefs  | 
|
66  | 
min :: "['a::ord, 'a] => 'a"  | 
|
67  | 
"min a b == (if a <= b then a else b)"  | 
|
68  | 
max :: "['a::ord, 'a] => 'a"  | 
|
69  | 
"max a b == (if a <= b then b else a)"  | 
|
70  | 
||
71  | 
lemma min_leastL: "(!!x. least <= x) ==> min least x = least"  | 
|
72  | 
by (simp add: min_def)  | 
|
73  | 
||
74  | 
lemma min_of_mono:  | 
|
75  | 
"ALL x y. (f x <= f y) = (x <= y) ==> min (f m) (f n) = f (min m n)"  | 
|
76  | 
by (simp add: min_def)  | 
|
77  | 
||
78  | 
lemma max_leastL: "(!!x. least <= x) ==> max least x = x"  | 
|
79  | 
by (simp add: max_def)  | 
|
80  | 
||
81  | 
lemma max_of_mono:  | 
|
82  | 
"ALL x y. (f x <= f y) = (x <= y) ==> max (f m) (f n) = f (max m n)"  | 
|
83  | 
by (simp add: max_def)  | 
|
84  | 
||
85  | 
||
86  | 
subsection "Orders"  | 
|
87  | 
||
88  | 
axclass order < ord  | 
|
89  | 
order_refl [iff]: "x <= x"  | 
|
90  | 
order_trans: "x <= y ==> y <= z ==> x <= z"  | 
|
91  | 
order_antisym: "x <= y ==> y <= x ==> x = y"  | 
|
92  | 
order_less_le: "(x < y) = (x <= y & x ~= y)"  | 
|
93  | 
||
94  | 
text{* Connection to locale: *}
 | 
|
95  | 
||
| 15837 | 96  | 
interpretation order:  | 
| 15780 | 97  | 
partial_order["op \<le> :: 'a::order \<Rightarrow> 'a \<Rightarrow> bool"]  | 
| 15524 | 98  | 
apply(rule partial_order.intro)  | 
99  | 
apply(rule order_refl, erule (1) order_trans, erule (1) order_antisym)  | 
|
100  | 
done  | 
|
101  | 
||
102  | 
text {* Reflexivity. *}
 | 
|
103  | 
||
104  | 
lemma order_eq_refl: "!!x::'a::order. x = y ==> x <= y"  | 
|
105  | 
    -- {* This form is useful with the classical reasoner. *}
 | 
|
106  | 
apply (erule ssubst)  | 
|
107  | 
apply (rule order_refl)  | 
|
108  | 
done  | 
|
109  | 
||
110  | 
lemma order_less_irrefl [iff]: "~ x < (x::'a::order)"  | 
|
111  | 
by (simp add: order_less_le)  | 
|
112  | 
||
113  | 
lemma order_le_less: "((x::'a::order) <= y) = (x < y | x = y)"  | 
|
114  | 
    -- {* NOT suitable for iff, since it can cause PROOF FAILED. *}
 | 
|
115  | 
apply (simp add: order_less_le, blast)  | 
|
116  | 
done  | 
|
117  | 
||
118  | 
lemmas order_le_imp_less_or_eq = order_le_less [THEN iffD1, standard]  | 
|
119  | 
||
120  | 
lemma order_less_imp_le: "!!x::'a::order. x < y ==> x <= y"  | 
|
121  | 
by (simp add: order_less_le)  | 
|
122  | 
||
123  | 
||
124  | 
text {* Asymmetry. *}
 | 
|
125  | 
||
126  | 
lemma order_less_not_sym: "(x::'a::order) < y ==> ~ (y < x)"  | 
|
127  | 
by (simp add: order_less_le order_antisym)  | 
|
128  | 
||
129  | 
lemma order_less_asym: "x < (y::'a::order) ==> (~P ==> y < x) ==> P"  | 
|
130  | 
apply (drule order_less_not_sym)  | 
|
131  | 
apply (erule contrapos_np, simp)  | 
|
132  | 
done  | 
|
133  | 
||
134  | 
lemma order_eq_iff: "!!x::'a::order. (x = y) = (x \<le> y & y \<le> x)"  | 
|
135  | 
by (blast intro: order_antisym)  | 
|
136  | 
||
137  | 
lemma order_antisym_conv: "(y::'a::order) <= x ==> (x <= y) = (x = y)"  | 
|
138  | 
by(blast intro:order_antisym)  | 
|
139  | 
||
140  | 
text {* Transitivity. *}
 | 
|
141  | 
||
142  | 
lemma order_less_trans: "!!x::'a::order. [| x < y; y < z |] ==> x < z"  | 
|
143  | 
apply (simp add: order_less_le)  | 
|
144  | 
apply (blast intro: order_trans order_antisym)  | 
|
145  | 
done  | 
|
146  | 
||
147  | 
lemma order_le_less_trans: "!!x::'a::order. [| x <= y; y < z |] ==> x < z"  | 
|
148  | 
apply (simp add: order_less_le)  | 
|
149  | 
apply (blast intro: order_trans order_antisym)  | 
|
150  | 
done  | 
|
151  | 
||
152  | 
lemma order_less_le_trans: "!!x::'a::order. [| x < y; y <= z |] ==> x < z"  | 
|
153  | 
apply (simp add: order_less_le)  | 
|
154  | 
apply (blast intro: order_trans order_antisym)  | 
|
155  | 
done  | 
|
156  | 
||
157  | 
||
158  | 
text {* Useful for simplification, but too risky to include by default. *}
 | 
|
159  | 
||
160  | 
lemma order_less_imp_not_less: "(x::'a::order) < y ==> (~ y < x) = True"  | 
|
161  | 
by (blast elim: order_less_asym)  | 
|
162  | 
||
163  | 
lemma order_less_imp_triv: "(x::'a::order) < y ==> (y < x --> P) = True"  | 
|
164  | 
by (blast elim: order_less_asym)  | 
|
165  | 
||
166  | 
lemma order_less_imp_not_eq: "(x::'a::order) < y ==> (x = y) = False"  | 
|
167  | 
by auto  | 
|
168  | 
||
169  | 
lemma order_less_imp_not_eq2: "(x::'a::order) < y ==> (y = x) = False"  | 
|
170  | 
by auto  | 
|
171  | 
||
172  | 
||
173  | 
text {* Other operators. *}
 | 
|
174  | 
||
175  | 
lemma min_leastR: "(!!x::'a::order. least <= x) ==> min x least = least"  | 
|
176  | 
apply (simp add: min_def)  | 
|
177  | 
apply (blast intro: order_antisym)  | 
|
178  | 
done  | 
|
179  | 
||
180  | 
lemma max_leastR: "(!!x::'a::order. least <= x) ==> max x least = x"  | 
|
181  | 
apply (simp add: max_def)  | 
|
182  | 
apply (blast intro: order_antisym)  | 
|
183  | 
done  | 
|
184  | 
||
185  | 
||
186  | 
subsection {* Transitivity rules for calculational reasoning *}
 | 
|
187  | 
||
188  | 
||
189  | 
lemma order_neq_le_trans: "a ~= b ==> (a::'a::order) <= b ==> a < b"  | 
|
190  | 
by (simp add: order_less_le)  | 
|
191  | 
||
192  | 
lemma order_le_neq_trans: "(a::'a::order) <= b ==> a ~= b ==> a < b"  | 
|
193  | 
by (simp add: order_less_le)  | 
|
194  | 
||
195  | 
lemma order_less_asym': "(a::'a::order) < b ==> b < a ==> P"  | 
|
196  | 
by (rule order_less_asym)  | 
|
197  | 
||
198  | 
||
199  | 
subsection {* Least value operator *}
 | 
|
200  | 
||
201  | 
constdefs  | 
|
202  | 
  Least :: "('a::ord => bool) => 'a"               (binder "LEAST " 10)
 | 
|
203  | 
"Least P == THE x. P x & (ALL y. P y --> x <= y)"  | 
|
204  | 
    -- {* We can no longer use LeastM because the latter requires Hilbert-AC. *}
 | 
|
205  | 
||
| 15950 | 206  | 
lemma LeastI2_order:  | 
| 15524 | 207  | 
"[| P (x::'a::order);  | 
208  | 
!!y. P y ==> x <= y;  | 
|
209  | 
!!x. [| P x; ALL y. P y --> x \<le> y |] ==> Q x |]  | 
|
210  | 
==> Q (Least P)"  | 
|
211  | 
apply (unfold Least_def)  | 
|
212  | 
apply (rule theI2)  | 
|
213  | 
apply (blast intro: order_antisym)+  | 
|
214  | 
done  | 
|
215  | 
||
216  | 
lemma Least_equality:  | 
|
217  | 
"[| P (k::'a::order); !!x. P x ==> k <= x |] ==> (LEAST x. P x) = k"  | 
|
218  | 
apply (simp add: Least_def)  | 
|
219  | 
apply (rule the_equality)  | 
|
220  | 
apply (auto intro!: order_antisym)  | 
|
221  | 
done  | 
|
222  | 
||
223  | 
||
224  | 
subsection "Linear / total orders"  | 
|
225  | 
||
226  | 
axclass linorder < order  | 
|
227  | 
linorder_linear: "x <= y | y <= x"  | 
|
228  | 
||
229  | 
lemma linorder_less_linear: "!!x::'a::linorder. x<y | x=y | y<x"  | 
|
230  | 
apply (simp add: order_less_le)  | 
|
231  | 
apply (insert linorder_linear, blast)  | 
|
232  | 
done  | 
|
233  | 
||
234  | 
lemma linorder_le_less_linear: "!!x::'a::linorder. x\<le>y | y<x"  | 
|
235  | 
by (simp add: order_le_less linorder_less_linear)  | 
|
236  | 
||
237  | 
lemma linorder_le_cases [case_names le ge]:  | 
|
238  | 
"((x::'a::linorder) \<le> y ==> P) ==> (y \<le> x ==> P) ==> P"  | 
|
239  | 
by (insert linorder_linear, blast)  | 
|
240  | 
||
241  | 
lemma linorder_cases [case_names less equal greater]:  | 
|
242  | 
"((x::'a::linorder) < y ==> P) ==> (x = y ==> P) ==> (y < x ==> P) ==> P"  | 
|
243  | 
by (insert linorder_less_linear, blast)  | 
|
244  | 
||
245  | 
lemma linorder_not_less: "!!x::'a::linorder. (~ x < y) = (y <= x)"  | 
|
246  | 
apply (simp add: order_less_le)  | 
|
247  | 
apply (insert linorder_linear)  | 
|
248  | 
apply (blast intro: order_antisym)  | 
|
249  | 
done  | 
|
250  | 
||
251  | 
lemma linorder_not_le: "!!x::'a::linorder. (~ x <= y) = (y < x)"  | 
|
252  | 
apply (simp add: order_less_le)  | 
|
253  | 
apply (insert linorder_linear)  | 
|
254  | 
apply (blast intro: order_antisym)  | 
|
255  | 
done  | 
|
256  | 
||
257  | 
lemma linorder_neq_iff: "!!x::'a::linorder. (x ~= y) = (x<y | y<x)"  | 
|
258  | 
by (cut_tac x = x and y = y in linorder_less_linear, auto)  | 
|
259  | 
||
260  | 
lemma linorder_neqE: "x ~= (y::'a::linorder) ==> (x < y ==> R) ==> (y < x ==> R) ==> R"  | 
|
261  | 
by (simp add: linorder_neq_iff, blast)  | 
|
262  | 
||
263  | 
lemma linorder_antisym_conv1: "~ (x::'a::linorder) < y ==> (x <= y) = (x = y)"  | 
|
264  | 
by(blast intro:order_antisym dest:linorder_not_less[THEN iffD1])  | 
|
265  | 
||
266  | 
lemma linorder_antisym_conv2: "(x::'a::linorder) <= y ==> (~ x < y) = (x = y)"  | 
|
267  | 
by(blast intro:order_antisym dest:linorder_not_less[THEN iffD1])  | 
|
268  | 
||
269  | 
lemma linorder_antisym_conv3: "~ (y::'a::linorder) < x ==> (~ x < y) = (x = y)"  | 
|
270  | 
by(blast intro:order_antisym dest:linorder_not_less[THEN iffD1])  | 
|
271  | 
||
272  | 
use "antisym_setup.ML";  | 
|
273  | 
setup antisym_setup  | 
|
274  | 
||
275  | 
subsection {* Setup of transitivity reasoner as Solver *}
 | 
|
276  | 
||
277  | 
lemma less_imp_neq: "[| (x::'a::order) < y |] ==> x ~= y"  | 
|
278  | 
by (erule contrapos_pn, erule subst, rule order_less_irrefl)  | 
|
279  | 
||
280  | 
lemma eq_neq_eq_imp_neq: "[| x = a ; a ~= b; b = y |] ==> x ~= y"  | 
|
281  | 
by (erule subst, erule ssubst, assumption)  | 
|
282  | 
||
283  | 
ML_setup {*
 | 
|
284  | 
||
285  | 
(* The setting up of Quasi_Tac serves as a demo. Since there is no  | 
|
286  | 
class for quasi orders, the tactics Quasi_Tac.trans_tac and  | 
|
287  | 
Quasi_Tac.quasi_tac are not of much use. *)  | 
|
288  | 
||
289  | 
fun decomp_gen sort sign (Trueprop $ t) =  | 
|
| 
15622
 
4723248c982b
Transitivity reasoner ignores types amenable to linear arithmetic.
 
ballarin 
parents: 
15531 
diff
changeset
 | 
290  | 
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
 | 
291  | 
(* exclude numeric types: linear arithmetic subsumes transitivity *)  | 
| 
 
4723248c982b
Transitivity reasoner ignores types amenable to linear arithmetic.
 
ballarin 
parents: 
15531 
diff
changeset
 | 
292  | 
T <> HOLogic.natT andalso T <> HOLogic.intT andalso  | 
| 
 
4723248c982b
Transitivity reasoner ignores types amenable to linear arithmetic.
 
ballarin 
parents: 
15531 
diff
changeset
 | 
293  | 
T <> HOLogic.realT andalso Sign.of_sort sign (T, sort) end  | 
| 15524 | 294  | 
  fun dec (Const ("Not", _) $ t) = (
 | 
295  | 
case dec t of  | 
|
| 15531 | 296  | 
NONE => NONE  | 
297  | 
| SOME (t1, rel, t2) => SOME (t1, "~" ^ rel, t2))  | 
|
| 15524 | 298  | 
	| dec (Const ("op =",  _) $ t1 $ t2) =
 | 
299  | 
if of_sort t1  | 
|
| 15531 | 300  | 
then SOME (t1, "=", t2)  | 
301  | 
else NONE  | 
|
| 15524 | 302  | 
	| dec (Const ("op <=",  _) $ t1 $ t2) =
 | 
303  | 
if of_sort t1  | 
|
| 15531 | 304  | 
then SOME (t1, "<=", t2)  | 
305  | 
else NONE  | 
|
| 15524 | 306  | 
	| dec (Const ("op <",  _) $ t1 $ t2) =
 | 
307  | 
if of_sort t1  | 
|
| 15531 | 308  | 
then SOME (t1, "<", t2)  | 
309  | 
else NONE  | 
|
310  | 
| dec _ = NONE  | 
|
| 15524 | 311  | 
in dec t end;  | 
312  | 
||
313  | 
structure Quasi_Tac = Quasi_Tac_Fun (  | 
|
314  | 
struct  | 
|
315  | 
val le_trans = thm "order_trans";  | 
|
316  | 
val le_refl = thm "order_refl";  | 
|
317  | 
val eqD1 = thm "order_eq_refl";  | 
|
318  | 
val eqD2 = thm "sym" RS thm "order_eq_refl";  | 
|
319  | 
val less_reflE = thm "order_less_irrefl" RS thm "notE";  | 
|
320  | 
val less_imp_le = thm "order_less_imp_le";  | 
|
321  | 
val le_neq_trans = thm "order_le_neq_trans";  | 
|
322  | 
val neq_le_trans = thm "order_neq_le_trans";  | 
|
323  | 
val less_imp_neq = thm "less_imp_neq";  | 
|
324  | 
val decomp_trans = decomp_gen ["Orderings.order"];  | 
|
325  | 
val decomp_quasi = decomp_gen ["Orderings.order"];  | 
|
326  | 
||
327  | 
end); (* struct *)  | 
|
328  | 
||
329  | 
structure Order_Tac = Order_Tac_Fun (  | 
|
330  | 
struct  | 
|
331  | 
val less_reflE = thm "order_less_irrefl" RS thm "notE";  | 
|
332  | 
val le_refl = thm "order_refl";  | 
|
333  | 
val less_imp_le = thm "order_less_imp_le";  | 
|
334  | 
val not_lessI = thm "linorder_not_less" RS thm "iffD2";  | 
|
335  | 
val not_leI = thm "linorder_not_le" RS thm "iffD2";  | 
|
336  | 
val not_lessD = thm "linorder_not_less" RS thm "iffD1";  | 
|
337  | 
val not_leD = thm "linorder_not_le" RS thm "iffD1";  | 
|
338  | 
val eqI = thm "order_antisym";  | 
|
339  | 
val eqD1 = thm "order_eq_refl";  | 
|
340  | 
val eqD2 = thm "sym" RS thm "order_eq_refl";  | 
|
341  | 
val less_trans = thm "order_less_trans";  | 
|
342  | 
val less_le_trans = thm "order_less_le_trans";  | 
|
343  | 
val le_less_trans = thm "order_le_less_trans";  | 
|
344  | 
val le_trans = thm "order_trans";  | 
|
345  | 
val le_neq_trans = thm "order_le_neq_trans";  | 
|
346  | 
val neq_le_trans = thm "order_neq_le_trans";  | 
|
347  | 
val less_imp_neq = thm "less_imp_neq";  | 
|
348  | 
val eq_neq_eq_imp_neq = thm "eq_neq_eq_imp_neq";  | 
|
349  | 
val decomp_part = decomp_gen ["Orderings.order"];  | 
|
350  | 
val decomp_lin = decomp_gen ["Orderings.linorder"];  | 
|
351  | 
||
352  | 
end); (* struct *)  | 
|
353  | 
||
354  | 
simpset_ref() := simpset ()  | 
|
355  | 
addSolver (mk_solver "Trans_linear" (fn _ => Order_Tac.linear_tac))  | 
|
356  | 
addSolver (mk_solver "Trans_partial" (fn _ => Order_Tac.partial_tac));  | 
|
357  | 
(* Adding the transitivity reasoners also as safe solvers showed a slight  | 
|
358  | 
speed up, but the reasoning strength appears to be not higher (at least  | 
|
359  | 
no breaking of additional proofs in the entire HOL distribution, as  | 
|
360  | 
of 5 March 2004, was observed). *)  | 
|
361  | 
*}  | 
|
362  | 
||
363  | 
(* Optional setup of methods *)  | 
|
364  | 
||
365  | 
(*  | 
|
366  | 
method_setup trans_partial =  | 
|
367  | 
  {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (Order_Tac.partial_tac)) *}
 | 
|
368  | 
  {* transitivity reasoner for partial orders *}	
 | 
|
369  | 
method_setup trans_linear =  | 
|
370  | 
  {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (Order_Tac.linear_tac)) *}
 | 
|
371  | 
  {* transitivity reasoner for linear orders *}
 | 
|
372  | 
*)  | 
|
373  | 
||
374  | 
(*  | 
|
375  | 
declare order.order_refl [simp del] order_less_irrefl [simp del]  | 
|
376  | 
||
377  | 
can currently not be removed, abel_cancel relies on it.  | 
|
378  | 
*)  | 
|
379  | 
||
380  | 
||
381  | 
subsection "Min and max on (linear) orders"  | 
|
382  | 
||
383  | 
text{* Instantiate locales: *}
 | 
|
384  | 
||
| 15837 | 385  | 
interpretation min_max:  | 
| 15780 | 386  | 
lower_semilattice["op \<le>" "min :: 'a::linorder \<Rightarrow> 'a \<Rightarrow> 'a"]  | 
387  | 
apply -  | 
|
| 15524 | 388  | 
apply(rule lower_semilattice_axioms.intro)  | 
389  | 
apply(simp add:min_def linorder_not_le order_less_imp_le)  | 
|
390  | 
apply(simp add:min_def linorder_not_le order_less_imp_le)  | 
|
391  | 
apply(simp add:min_def linorder_not_le order_less_imp_le)  | 
|
392  | 
done  | 
|
393  | 
||
| 15837 | 394  | 
interpretation min_max:  | 
| 15780 | 395  | 
upper_semilattice["op \<le>" "max :: 'a::linorder \<Rightarrow> 'a \<Rightarrow> 'a"]  | 
396  | 
apply -  | 
|
| 15524 | 397  | 
apply(rule upper_semilattice_axioms.intro)  | 
398  | 
apply(simp add: max_def linorder_not_le order_less_imp_le)  | 
|
399  | 
apply(simp add: max_def linorder_not_le order_less_imp_le)  | 
|
400  | 
apply(simp add: max_def linorder_not_le order_less_imp_le)  | 
|
401  | 
done  | 
|
402  | 
||
| 15837 | 403  | 
interpretation min_max:  | 
| 15780 | 404  | 
lattice["op \<le>" "min :: 'a::linorder \<Rightarrow> 'a \<Rightarrow> 'a" "max"]  | 
405  | 
.  | 
|
| 15524 | 406  | 
|
| 15837 | 407  | 
interpretation min_max:  | 
| 15780 | 408  | 
distrib_lattice["op \<le>" "min :: 'a::linorder \<Rightarrow> 'a \<Rightarrow> 'a" "max"]  | 
| 15524 | 409  | 
apply(rule distrib_lattice_axioms.intro)  | 
410  | 
apply(rule_tac x=x and y=y in linorder_le_cases)  | 
|
411  | 
apply(rule_tac x=x and y=z in linorder_le_cases)  | 
|
412  | 
apply(rule_tac x=y and y=z in linorder_le_cases)  | 
|
413  | 
apply(simp add:min_def max_def)  | 
|
414  | 
apply(simp add:min_def max_def)  | 
|
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=x and y=z in linorder_le_cases)  | 
|
419  | 
apply(rule_tac x=y and y=z in linorder_le_cases)  | 
|
420  | 
apply(simp add:min_def max_def)  | 
|
421  | 
apply(simp add:min_def max_def)  | 
|
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  | 
done  | 
|
426  | 
||
427  | 
lemma le_max_iff_disj: "!!z::'a::linorder. (z <= max x y) = (z <= x | z <= y)"  | 
|
428  | 
apply(simp add:max_def)  | 
|
429  | 
apply (insert linorder_linear)  | 
|
430  | 
apply (blast intro: order_trans)  | 
|
431  | 
done  | 
|
432  | 
||
| 15780 | 433  | 
lemmas le_maxI1 = min_max.sup_ge1  | 
434  | 
lemmas le_maxI2 = min_max.sup_ge2  | 
|
| 15524 | 435  | 
|
436  | 
lemma less_max_iff_disj: "!!z::'a::linorder. (z < max x y) = (z < x | z < y)"  | 
|
437  | 
apply (simp add: max_def order_le_less)  | 
|
438  | 
apply (insert linorder_less_linear)  | 
|
439  | 
apply (blast intro: order_less_trans)  | 
|
440  | 
done  | 
|
441  | 
||
442  | 
lemma max_less_iff_conj [simp]:  | 
|
443  | 
"!!z::'a::linorder. (max x y < z) = (x < z & y < z)"  | 
|
444  | 
apply (simp add: order_le_less max_def)  | 
|
445  | 
apply (insert linorder_less_linear)  | 
|
446  | 
apply (blast intro: order_less_trans)  | 
|
447  | 
done  | 
|
| 15791 | 448  | 
|
| 15524 | 449  | 
lemma min_less_iff_conj [simp]:  | 
450  | 
"!!z::'a::linorder. (z < min x y) = (z < x & z < y)"  | 
|
451  | 
apply (simp add: order_le_less min_def)  | 
|
452  | 
apply (insert linorder_less_linear)  | 
|
453  | 
apply (blast intro: order_less_trans)  | 
|
454  | 
done  | 
|
455  | 
||
456  | 
lemma min_le_iff_disj: "!!z::'a::linorder. (min x y <= z) = (x <= z | y <= z)"  | 
|
457  | 
apply (simp add: min_def)  | 
|
458  | 
apply (insert linorder_linear)  | 
|
459  | 
apply (blast intro: order_trans)  | 
|
460  | 
done  | 
|
461  | 
||
462  | 
lemma min_less_iff_disj: "!!z::'a::linorder. (min x y < z) = (x < z | y < z)"  | 
|
463  | 
apply (simp add: min_def order_le_less)  | 
|
464  | 
apply (insert linorder_less_linear)  | 
|
465  | 
apply (blast intro: order_less_trans)  | 
|
466  | 
done  | 
|
467  | 
||
| 15780 | 468  | 
lemmas max_ac = min_max.sup_assoc min_max.sup_commute  | 
469  | 
mk_left_commute[of max,OF min_max.sup_assoc min_max.sup_commute]  | 
|
| 15524 | 470  | 
|
| 15780 | 471  | 
lemmas min_ac = min_max.inf_assoc min_max.inf_commute  | 
472  | 
mk_left_commute[of min,OF min_max.inf_assoc min_max.inf_commute]  | 
|
| 15524 | 473  | 
|
474  | 
lemma split_min:  | 
|
475  | 
"P (min (i::'a::linorder) j) = ((i <= j --> P(i)) & (~ i <= j --> P(j)))"  | 
|
476  | 
by (simp add: min_def)  | 
|
477  | 
||
478  | 
lemma split_max:  | 
|
479  | 
"P (max (i::'a::linorder) j) = ((i <= j --> P(j)) & (~ i <= j --> P(i)))"  | 
|
480  | 
by (simp add: max_def)  | 
|
481  | 
||
482  | 
||
483  | 
subsection "Bounded quantifiers"  | 
|
484  | 
||
485  | 
syntax  | 
|
486  | 
  "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3ALL _<_./ _)"  [0, 0, 10] 10)
 | 
|
487  | 
  "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3EX _<_./ _)"  [0, 0, 10] 10)
 | 
|
488  | 
  "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3ALL _<=_./ _)" [0, 0, 10] 10)
 | 
|
489  | 
  "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3EX _<=_./ _)" [0, 0, 10] 10)
 | 
|
490  | 
||
491  | 
  "_gtAll" :: "[idt, 'a, bool] => bool"   ("(3ALL _>_./ _)"  [0, 0, 10] 10)
 | 
|
492  | 
  "_gtEx"  :: "[idt, 'a, bool] => bool"   ("(3EX _>_./ _)"  [0, 0, 10] 10)
 | 
|
493  | 
  "_geAll"   :: "[idt, 'a, bool] => bool"   ("(3ALL _>=_./ _)" [0, 0, 10] 10)
 | 
|
494  | 
  "_geEx"    :: "[idt, 'a, bool] => bool"   ("(3EX _>=_./ _)" [0, 0, 10] 10)
 | 
|
495  | 
||
496  | 
syntax (xsymbols)  | 
|
497  | 
  "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_<_./ _)"  [0, 0, 10] 10)
 | 
|
498  | 
  "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_<_./ _)"  [0, 0, 10] 10)
 | 
|
499  | 
  "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10)
 | 
|
500  | 
  "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10)
 | 
|
501  | 
||
502  | 
  "_gtAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_>_./ _)"  [0, 0, 10] 10)
 | 
|
503  | 
  "_gtEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_>_./ _)"  [0, 0, 10] 10)
 | 
|
504  | 
  "_geAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<ge>_./ _)" [0, 0, 10] 10)
 | 
|
505  | 
  "_geEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<ge>_./ _)" [0, 0, 10] 10)
 | 
|
506  | 
||
507  | 
syntax (HOL)  | 
|
508  | 
  "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3! _<_./ _)"  [0, 0, 10] 10)
 | 
|
509  | 
  "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3? _<_./ _)"  [0, 0, 10] 10)
 | 
|
510  | 
  "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3! _<=_./ _)" [0, 0, 10] 10)
 | 
|
511  | 
  "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3? _<=_./ _)" [0, 0, 10] 10)
 | 
|
512  | 
||
513  | 
syntax (HTML output)  | 
|
514  | 
  "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_<_./ _)"  [0, 0, 10] 10)
 | 
|
515  | 
  "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_<_./ _)"  [0, 0, 10] 10)
 | 
|
516  | 
  "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10)
 | 
|
517  | 
  "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10)
 | 
|
518  | 
||
519  | 
  "_gtAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_>_./ _)"  [0, 0, 10] 10)
 | 
|
520  | 
  "_gtEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_>_./ _)"  [0, 0, 10] 10)
 | 
|
521  | 
  "_geAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<ge>_./ _)" [0, 0, 10] 10)
 | 
|
522  | 
  "_geEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<ge>_./ _)" [0, 0, 10] 10)
 | 
|
523  | 
||
524  | 
translations  | 
|
525  | 
"ALL x<y. P" => "ALL x. x < y --> P"  | 
|
526  | 
"EX x<y. P" => "EX x. x < y & P"  | 
|
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  | 
||
534  | 
print_translation {*
 | 
|
535  | 
let  | 
|
536  | 
fun mk v v' q n P =  | 
|
537  | 
if v=v' andalso not(v mem (map fst (Term.add_frees([],n))))  | 
|
538  | 
then Syntax.const q $ Syntax.mark_bound v' $ n $ P else raise Match;  | 
|
539  | 
  fun all_tr' [Const ("_bound",_) $ Free (v,_),
 | 
|
540  | 
               Const("op -->",_) $ (Const ("op <",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] =
 | 
|
541  | 
mk v v' "_lessAll" n P  | 
|
542  | 
||
543  | 
  | all_tr' [Const ("_bound",_) $ Free (v,_),
 | 
|
544  | 
               Const("op -->",_) $ (Const ("op <=",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] =
 | 
|
545  | 
mk v v' "_leAll" n P  | 
|
546  | 
||
547  | 
  | all_tr' [Const ("_bound",_) $ Free (v,_),
 | 
|
548  | 
               Const("op -->",_) $ (Const ("op <",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] =
 | 
|
549  | 
mk v v' "_gtAll" n P  | 
|
550  | 
||
551  | 
  | all_tr' [Const ("_bound",_) $ Free (v,_),
 | 
|
552  | 
               Const("op -->",_) $ (Const ("op <=",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] =
 | 
|
553  | 
mk v v' "_geAll" n P;  | 
|
554  | 
||
555  | 
  fun ex_tr' [Const ("_bound",_) $ Free (v,_),
 | 
|
556  | 
               Const("op &",_) $ (Const ("op <",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] =
 | 
|
557  | 
mk v v' "_lessEx" n P  | 
|
558  | 
||
559  | 
  | ex_tr' [Const ("_bound",_) $ Free (v,_),
 | 
|
560  | 
               Const("op &",_) $ (Const ("op <=",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] =
 | 
|
561  | 
mk v v' "_leEx" n P  | 
|
562  | 
||
563  | 
  | ex_tr' [Const ("_bound",_) $ Free (v,_),
 | 
|
564  | 
               Const("op &",_) $ (Const ("op <",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] =
 | 
|
565  | 
mk v v' "_gtEx" n P  | 
|
566  | 
||
567  | 
  | ex_tr' [Const ("_bound",_) $ Free (v,_),
 | 
|
568  | 
               Const("op &",_) $ (Const ("op <=",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] =
 | 
|
569  | 
mk v v' "_geEx" n P  | 
|
570  | 
in  | 
|
571  | 
[("ALL ", all_tr'), ("EX ", ex_tr')]
 | 
|
572  | 
end  | 
|
573  | 
*}  | 
|
574  | 
||
575  | 
end  |