| author | paulson | 
| Tue, 30 Nov 1999 17:53:34 +0100 | |
| changeset 8042 | ecdedff41e67 | 
| parent 8024 | 199721f2eb2d | 
| child 8214 | d612354445b6 | 
| 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: 
5316diff
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: 
5316diff
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: 
5316diff
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: 
5316diff
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: 
5316diff
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"; | |
| 6956 | 20 | AddXIs [monoI]; | 
| 923 | 21 | |
| 5316 | 22 | Goalw [mono_def] "[| mono(f); A <= B |] ==> f(A) <= f(B)"; | 
| 23 | by (Fast_tac 1); | |
| 923 | 24 | qed "monoD"; | 
| 6956 | 25 | AddXDs [monoD]; | 
| 923 | 26 | |
| 2608 | 27 | |
| 28 | section "Orders"; | |
| 29 | ||
| 5538 | 30 | (** Reflexivity **) | 
| 31 | ||
| 6115 | 32 | AddIffs [order_refl]; | 
| 2608 | 33 | |
| 4600 | 34 | (*This form is useful with the classical reasoner*) | 
| 5069 | 35 | Goal "!!x::'a::order. x = y ==> x <= y"; | 
| 4600 | 36 | by (etac ssubst 1); | 
| 37 | by (rtac order_refl 1); | |
| 38 | qed "order_eq_refl"; | |
| 39 | ||
| 5069 | 40 | Goal "~ x < (x::'a::order)"; | 
| 4089 | 41 | by (simp_tac (simpset() addsimps [order_less_le]) 1); | 
| 2608 | 42 | qed "order_less_irrefl"; | 
| 5449 
d853d1ac85a3
Adds order_refl, order_less_irrefl as simps, not as Iffs, to avoid PROOF FAILED
 paulson parents: 
5316diff
changeset | 43 | Addsimps [order_less_irrefl]; | 
| 2608 | 44 | |
| 5069 | 45 | Goal "(x::'a::order) <= y = (x < y | x = y)"; | 
| 4089 | 46 | 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: 
5316diff
changeset | 47 | (*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: 
5316diff
changeset | 48 | by (blast_tac (claset() addSIs [order_refl]) 1); | 
| 2608 | 49 | qed "order_le_less"; | 
| 50 | ||
| 5538 | 51 | (** Asymmetry **) | 
| 52 | ||
| 53 | Goal "(x::'a::order) < y ==> ~ (y<x)"; | |
| 54 | by (asm_full_simp_tac (simpset() addsimps [order_less_le, order_antisym]) 1); | |
| 55 | qed "order_less_not_sym"; | |
| 56 | ||
| 57 | (* [| n<m; ~P ==> m<n |] ==> P *) | |
| 58 | bind_thm ("order_less_asym", order_less_not_sym RS swap);
 | |
| 59 | ||
| 6073 | 60 | (* Transitivity *) | 
| 61 | ||
| 62 | Goal "!!x::'a::order. [| x < y; y < z |] ==> x < z"; | |
| 63 | by (asm_full_simp_tac (simpset() addsimps [order_less_le]) 1); | |
| 64 | by (blast_tac (claset() addIs [order_trans,order_antisym]) 1); | |
| 65 | qed "order_less_trans"; | |
| 66 | ||
| 6780 | 67 | Goal "!!x::'a::order. [| x <= y; y < z |] ==> x < z"; | 
| 68 | by (asm_full_simp_tac (simpset() addsimps [order_less_le]) 1); | |
| 69 | by (blast_tac (claset() addIs [order_trans,order_antisym]) 1); | |
| 70 | qed "order_le_less_trans"; | |
| 71 | ||
| 72 | Goal "!!x::'a::order. [| x < y; y <= z |] ==> x < z"; | |
| 73 | by (asm_full_simp_tac (simpset() addsimps [order_less_le]) 1); | |
| 74 | by (blast_tac (claset() addIs [order_trans,order_antisym]) 1); | |
| 75 | qed "order_less_le_trans"; | |
| 76 | ||
| 5538 | 77 | |
| 78 | (** Useful for simplification, but too risky to include by default. **) | |
| 79 | ||
| 80 | Goal "(x::'a::order) < y ==> (~ y < x) = True"; | |
| 81 | by (blast_tac (claset() addEs [order_less_asym]) 1); | |
| 82 | qed "order_less_imp_not_less"; | |
| 83 | ||
| 84 | Goal "(x::'a::order) < y ==> (y < x --> P) = True"; | |
| 85 | by (blast_tac (claset() addEs [order_less_asym]) 1); | |
| 86 | qed "order_less_imp_triv"; | |
| 87 | ||
| 88 | Goal "(x::'a::order) < y ==> (x = y) = False"; | |
| 89 | by Auto_tac; | |
| 90 | qed "order_less_imp_not_eq"; | |
| 91 | ||
| 92 | Goal "(x::'a::order) < y ==> (y = x) = False"; | |
| 93 | by Auto_tac; | |
| 94 | qed "order_less_imp_not_eq2"; | |
| 95 | ||
| 96 | ||
| 2608 | 97 | (** min **) | 
| 98 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 99 | val prems = Goalw [min_def] "(!!x. least <= x) ==> min least x = least"; | 
| 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 100 | by (simp_tac (simpset() addsimps prems) 1); | 
| 2608 | 101 | qed "min_leastL"; | 
| 102 | ||
| 5316 | 103 | val prems = Goalw [min_def] | 
| 2608 | 104 | "(!!x::'a::order. least <= x) ==> min x least = least"; | 
| 2935 | 105 | by (cut_facts_tac prems 1); | 
| 106 | by (Asm_simp_tac 1); | |
| 4089 | 107 | by (blast_tac (claset() addIs [order_antisym]) 1); | 
| 2608 | 108 | qed "min_leastR"; | 
| 4640 | 109 | |
| 110 | ||
| 111 | section "Linear/Total Orders"; | |
| 112 | ||
| 5069 | 113 | Goal "!!x::'a::linorder. x<y | x=y | y<x"; | 
| 4640 | 114 | by (simp_tac (simpset() addsimps [order_less_le]) 1); | 
| 5132 | 115 | by (cut_facts_tac [linorder_linear] 1); | 
| 4640 | 116 | by (Blast_tac 1); | 
| 117 | qed "linorder_less_linear"; | |
| 118 | ||
| 8024 | 119 | val prems = goal thy | 
| 120 | "[| (x::'a::linorder)<y ==> P; x=y ==> P; y<x ==> P |] ==> P"; | |
| 121 | by(cut_facts_tac [linorder_less_linear] 1); | |
| 122 | by(REPEAT(eresolve_tac (prems@[disjE]) 1)); | |
| 123 | qed "linorder_less_split"; | |
| 124 | ||
| 6128 | 125 | Goal "!!x::'a::linorder. (~ x < y) = (y <= x)"; | 
| 126 | by (simp_tac (simpset() addsimps [order_less_le]) 1); | |
| 127 | by (cut_facts_tac [linorder_linear] 1); | |
| 128 | by (blast_tac (claset() addIs [order_antisym]) 1); | |
| 129 | qed "linorder_not_less"; | |
| 130 | ||
| 131 | Goal "!!x::'a::linorder. (~ x <= y) = (y < x)"; | |
| 132 | by (simp_tac (simpset() addsimps [order_less_le]) 1); | |
| 133 | by (cut_facts_tac [linorder_linear] 1); | |
| 134 | by (blast_tac (claset() addIs [order_antisym]) 1); | |
| 135 | qed "linorder_not_le"; | |
| 136 | ||
| 137 | Goal "!!x::'a::linorder. (x ~= y) = (x<y | y<x)"; | |
| 138 | by (cut_inst_tac [("x","x"),("y","y")] linorder_less_linear 1);
 | |
| 139 | by Auto_tac; | |
| 140 | qed "linorder_neq_iff"; | |
| 141 | ||
| 142 | (* eliminates ~= in premises *) | |
| 143 | bind_thm("linorder_neqE", linorder_neq_iff RS iffD1 RS disjE);
 | |
| 144 | ||
| 145 | (** min & max **) | |
| 146 | ||
| 6433 | 147 | Goalw [min_def] "min (x::'a::order) x = x"; | 
| 6814 | 148 | by (Simp_tac 1); | 
| 6433 | 149 | qed "min_same"; | 
| 150 | Addsimps [min_same]; | |
| 151 | ||
| 152 | Goalw [max_def] "max (x::'a::order) x = x"; | |
| 6814 | 153 | by (Simp_tac 1); | 
| 6433 | 154 | qed "max_same"; | 
| 155 | Addsimps [max_same]; | |
| 156 | ||
| 5069 | 157 | Goalw [max_def] "!!z::'a::linorder. (z <= max x y) = (z <= x | z <= y)"; | 
| 4686 | 158 | by (Simp_tac 1); | 
| 5132 | 159 | by (cut_facts_tac [linorder_linear] 1); | 
| 4640 | 160 | by (blast_tac (claset() addIs [order_trans]) 1); | 
| 161 | qed "le_max_iff_disj"; | |
| 162 | ||
| 7494 
45905028bb1d
added theorems le_maxI1 and le_maxI2, also in claset
 oheimb parents: 
6956diff
changeset | 163 | qed_goal "le_maxI1" Ord.thy "(x::'a::linorder) <= max x y" | 
| 
45905028bb1d
added theorems le_maxI1 and le_maxI2, also in claset
 oheimb parents: 
6956diff
changeset | 164 | (K [rtac (le_max_iff_disj RS iffD2) 1, rtac (order_refl RS disjI1) 1]); | 
| 
45905028bb1d
added theorems le_maxI1 and le_maxI2, also in claset
 oheimb parents: 
6956diff
changeset | 165 | qed_goal "le_maxI2" Ord.thy "(y::'a::linorder) <= max x y" | 
| 
45905028bb1d
added theorems le_maxI1 and le_maxI2, also in claset
 oheimb parents: 
6956diff
changeset | 166 | (K [rtac (le_max_iff_disj RS iffD2) 1, rtac (order_refl RS disjI2) 1]); | 
| 7617 
e783adccf39e
removed order-sorted theorems from the default claset
 paulson parents: 
7494diff
changeset | 167 | (*CANNOT use with AddSIs because blast_tac will give PROOF FAILED.*) | 
| 7494 
45905028bb1d
added theorems le_maxI1 and le_maxI2, also in claset
 oheimb parents: 
6956diff
changeset | 168 | |
| 6073 | 169 | Goalw [max_def] "!!z::'a::linorder. (z < max x y) = (z < x | z < y)"; | 
| 170 | by (simp_tac (simpset() addsimps [order_le_less]) 1); | |
| 171 | by (cut_facts_tac [linorder_less_linear] 1); | |
| 172 | by (blast_tac (claset() addIs [order_less_trans]) 1); | |
| 173 | qed "less_max_iff_disj"; | |
| 174 | ||
| 5069 | 175 | Goalw [max_def] "!!z::'a::linorder. (max x y <= z) = (x <= z & y <= z)"; | 
| 4686 | 176 | by (Simp_tac 1); | 
| 5132 | 177 | by (cut_facts_tac [linorder_linear] 1); | 
| 4640 | 178 | by (blast_tac (claset() addIs [order_trans]) 1); | 
| 179 | qed "max_le_iff_conj"; | |
| 5673 | 180 | Addsimps [max_le_iff_conj]; | 
| 4640 | 181 | |
| 6433 | 182 | Goalw [max_def] "!!z::'a::linorder. (max x y < z) = (x < z & y < z)"; | 
| 183 | by (simp_tac (simpset() addsimps [order_le_less]) 1); | |
| 184 | by (cut_facts_tac [linorder_less_linear] 1); | |
| 185 | by (blast_tac (claset() addIs [order_less_trans]) 1); | |
| 186 | qed "max_less_iff_conj"; | |
| 187 | Addsimps [max_less_iff_conj]; | |
| 188 | ||
| 5069 | 189 | Goalw [min_def] "!!z::'a::linorder. (z <= min x y) = (z <= x & z <= y)"; | 
| 4686 | 190 | by (Simp_tac 1); | 
| 5132 | 191 | by (cut_facts_tac [linorder_linear] 1); | 
| 4640 | 192 | by (blast_tac (claset() addIs [order_trans]) 1); | 
| 193 | qed "le_min_iff_conj"; | |
| 5673 | 194 | Addsimps [le_min_iff_conj]; | 
| 195 | (* AddIffs screws up a blast_tac in MiniML *) | |
| 4640 | 196 | |
| 6433 | 197 | Goalw [min_def] "!!z::'a::linorder. (z < min x y) = (z < x & z < y)"; | 
| 198 | by (simp_tac (simpset() addsimps [order_le_less]) 1); | |
| 199 | by (cut_facts_tac [linorder_less_linear] 1); | |
| 200 | by (blast_tac (claset() addIs [order_less_trans]) 1); | |
| 201 | qed "min_less_iff_conj"; | |
| 202 | Addsimps [min_less_iff_conj]; | |
| 203 | ||
| 5069 | 204 | Goalw [min_def] "!!z::'a::linorder. (min x y <= z) = (x <= z | y <= z)"; | 
| 4686 | 205 | by (Simp_tac 1); | 
| 5132 | 206 | by (cut_facts_tac [linorder_linear] 1); | 
| 4640 | 207 | by (blast_tac (claset() addIs [order_trans]) 1); | 
| 208 | qed "min_le_iff_disj"; | |
| 6157 | 209 | |
| 210 | Goalw [min_def] | |
| 211 | "P(min (i::'a::linorder) j) = ((i <= j --> P(i)) & (~ i <= j --> P(j)))"; | |
| 6301 | 212 | by (Simp_tac 1); | 
| 6157 | 213 | qed "split_min"; | 
| 214 | ||
| 215 | Goalw [max_def] | |
| 216 | "P(max (i::'a::linorder) j) = ((i <= j --> P(j)) & (~ i <= j --> P(i)))"; | |
| 6301 | 217 | by (Simp_tac 1); | 
| 6157 | 218 | qed "split_max"; |