author | oheimb |
Mon, 21 Sep 1998 23:25:27 +0200 | |
changeset 5526 | e7617b57a3e6 |
parent 5449 | d853d1ac85a3 |
child 5538 | c55bf0487abe |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: HOL/Ord.ML |
923 | 2 |
ID: $Id$ |
1465 | 3 |
Author: Tobias Nipkow, Cambridge University Computer Laboratory |
923 | 4 |
Copyright 1993 University of Cambridge |
5 |
||
6 |
The type class for ordered types |
|
7 |
*) |
|
8 |
||
5449
d853d1ac85a3
Adds order_refl, order_less_irrefl as simps, not as Iffs, to avoid PROOF FAILED
paulson
parents:
5316
diff
changeset
|
9 |
(*Tell Blast_tac about overloading of < and <= to reduce the risk of |
d853d1ac85a3
Adds order_refl, order_less_irrefl as simps, not as Iffs, to avoid PROOF FAILED
paulson
parents:
5316
diff
changeset
|
10 |
its applying a rule for the wrong type*) |
d853d1ac85a3
Adds order_refl, order_less_irrefl as simps, not as Iffs, to avoid PROOF FAILED
paulson
parents:
5316
diff
changeset
|
11 |
Blast.overloaded ("op <", domain_type); |
d853d1ac85a3
Adds order_refl, order_less_irrefl as simps, not as Iffs, to avoid PROOF FAILED
paulson
parents:
5316
diff
changeset
|
12 |
Blast.overloaded ("op <=", domain_type); |
d853d1ac85a3
Adds order_refl, order_less_irrefl as simps, not as Iffs, to avoid PROOF FAILED
paulson
parents:
5316
diff
changeset
|
13 |
|
2608 | 14 |
(** mono **) |
923 | 15 |
|
5316 | 16 |
val [prem] = Goalw [mono_def] |
923 | 17 |
"[| !!A B. A <= B ==> f(A) <= f(B) |] ==> mono(f)"; |
18 |
by (REPEAT (ares_tac [allI, impI, prem] 1)); |
|
19 |
qed "monoI"; |
|
20 |
||
5316 | 21 |
Goalw [mono_def] "[| mono(f); A <= B |] ==> f(A) <= f(B)"; |
22 |
by (Fast_tac 1); |
|
923 | 23 |
qed "monoD"; |
24 |
||
2608 | 25 |
|
26 |
section "Orders"; |
|
27 |
||
5449
d853d1ac85a3
Adds order_refl, order_less_irrefl as simps, not as Iffs, to avoid PROOF FAILED
paulson
parents:
5316
diff
changeset
|
28 |
Addsimps [order_refl]; |
2608 | 29 |
|
4600 | 30 |
(*This form is useful with the classical reasoner*) |
5069 | 31 |
Goal "!!x::'a::order. x = y ==> x <= y"; |
4600 | 32 |
by (etac ssubst 1); |
33 |
by (rtac order_refl 1); |
|
34 |
qed "order_eq_refl"; |
|
35 |
||
5069 | 36 |
Goal "~ x < (x::'a::order)"; |
4089 | 37 |
by (simp_tac (simpset() addsimps [order_less_le]) 1); |
2608 | 38 |
qed "order_less_irrefl"; |
5449
d853d1ac85a3
Adds order_refl, order_less_irrefl as simps, not as Iffs, to avoid PROOF FAILED
paulson
parents:
5316
diff
changeset
|
39 |
Addsimps [order_less_irrefl]; |
2608 | 40 |
|
5069 | 41 |
Goal "(x::'a::order) <= y = (x < y | x = y)"; |
4089 | 42 |
by (simp_tac (simpset() addsimps [order_less_le]) 1); |
5449
d853d1ac85a3
Adds order_refl, order_less_irrefl as simps, not as Iffs, to avoid PROOF FAILED
paulson
parents:
5316
diff
changeset
|
43 |
(*NOT suitable for AddIffs, since it can cause PROOF FAILED*) |
d853d1ac85a3
Adds order_refl, order_less_irrefl as simps, not as Iffs, to avoid PROOF FAILED
paulson
parents:
5316
diff
changeset
|
44 |
by (blast_tac (claset() addSIs [order_refl]) 1); |
2608 | 45 |
qed "order_le_less"; |
46 |
||
47 |
(** min **) |
|
48 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5132
diff
changeset
|
49 |
val prems = Goalw [min_def] "(!!x. least <= x) ==> min least x = least"; |
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5132
diff
changeset
|
50 |
by (simp_tac (simpset() addsimps prems) 1); |
2608 | 51 |
qed "min_leastL"; |
52 |
||
5316 | 53 |
val prems = Goalw [min_def] |
2608 | 54 |
"(!!x::'a::order. least <= x) ==> min x least = least"; |
2935 | 55 |
by (cut_facts_tac prems 1); |
56 |
by (Asm_simp_tac 1); |
|
4089 | 57 |
by (blast_tac (claset() addIs [order_antisym]) 1); |
2608 | 58 |
qed "min_leastR"; |
4640 | 59 |
|
60 |
||
61 |
section "Linear/Total Orders"; |
|
62 |
||
5069 | 63 |
Goal "!!x::'a::linorder. x<y | x=y | y<x"; |
4640 | 64 |
by (simp_tac (simpset() addsimps [order_less_le]) 1); |
5132 | 65 |
by (cut_facts_tac [linorder_linear] 1); |
4640 | 66 |
by (Blast_tac 1); |
67 |
qed "linorder_less_linear"; |
|
68 |
||
5069 | 69 |
Goalw [max_def] "!!z::'a::linorder. (z <= max x y) = (z <= x | z <= y)"; |
4686 | 70 |
by (Simp_tac 1); |
5132 | 71 |
by (cut_facts_tac [linorder_linear] 1); |
4640 | 72 |
by (blast_tac (claset() addIs [order_trans]) 1); |
73 |
qed "le_max_iff_disj"; |
|
74 |
||
5069 | 75 |
Goalw [max_def] "!!z::'a::linorder. (max x y <= z) = (x <= z & y <= z)"; |
4686 | 76 |
by (Simp_tac 1); |
5132 | 77 |
by (cut_facts_tac [linorder_linear] 1); |
4640 | 78 |
by (blast_tac (claset() addIs [order_trans]) 1); |
79 |
qed "max_le_iff_conj"; |
|
80 |
||
5069 | 81 |
Goalw [min_def] "!!z::'a::linorder. (z <= min x y) = (z <= x & z <= y)"; |
4686 | 82 |
by (Simp_tac 1); |
5132 | 83 |
by (cut_facts_tac [linorder_linear] 1); |
4640 | 84 |
by (blast_tac (claset() addIs [order_trans]) 1); |
85 |
qed "le_min_iff_conj"; |
|
86 |
||
5069 | 87 |
Goalw [min_def] "!!z::'a::linorder. (min x y <= z) = (x <= z | y <= z)"; |
4686 | 88 |
by (Simp_tac 1); |
5132 | 89 |
by (cut_facts_tac [linorder_linear] 1); |
4640 | 90 |
by (blast_tac (claset() addIs [order_trans]) 1); |
91 |
qed "min_le_iff_disj"; |
|
92 |
||
93 |