| author | nipkow | 
| Wed, 29 Mar 2000 15:09:51 +0200 | |
| changeset 8604 | c99e0024050c | 
| parent 8128 | 3a5864b465e2 | 
| child 8948 | b797cfa3548d | 
| permissions | -rw-r--r-- | 
| 4776 | 1 | (* Title: HOL/LessThan/LessThan | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5 | ||
| 6 | lessThan, greaterThan, atLeast, atMost | |
| 7 | *) | |
| 8 | ||
| 9 | ||
| 7878 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 10 | (*** Restrict [MOVE TO RELATION.THY??] ***) | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 11 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 12 | Goalw [Restrict_def] "((x,y): Restrict A r) = ((x,y): r & x: A)"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 13 | by (Blast_tac 1); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 14 | qed "Restrict_iff"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 15 | AddIffs [Restrict_iff]; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 16 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 17 | Goal "Restrict UNIV = id"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 18 | by (rtac ext 1); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 19 | by (auto_tac (claset(), simpset() addsimps [Restrict_def])); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 20 | qed "Restrict_UNIV"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 21 | Addsimps [Restrict_UNIV]; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 22 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 23 | Goal "Restrict {} r = {}";
 | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 24 | by (auto_tac (claset(), simpset() addsimps [Restrict_def])); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 25 | qed "Restrict_empty"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 26 | Addsimps [Restrict_empty]; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 27 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 28 | Goalw [Restrict_def] "Restrict A (Restrict B r) = Restrict (A Int B) r"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 29 | by (Blast_tac 1); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 30 | qed "Restrict_Int"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 31 | Addsimps [Restrict_Int]; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 32 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 33 | Goalw [Restrict_def] "Domain r <= A ==> Restrict A r = r"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 34 | by Auto_tac; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 35 | qed "Restrict_triv"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 36 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 37 | Goalw [Restrict_def] "Restrict A r <= r"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 38 | by Auto_tac; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 39 | qed "Restrict_subset"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 40 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 41 | Goalw [Restrict_def] | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 42 | "[| A <= B; Restrict B r = Restrict B s |] \ | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 43 | \ ==> Restrict A r = Restrict A s"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 44 | by (blast_tac (claset() addSEs [equalityE]) 1); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 45 | qed "Restrict_eq_mono"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 46 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 47 | Goalw [Restrict_def, image_def] | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 48 | "[| s : RR; Restrict A r = Restrict A s |] \ | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 49 | \ ==> Restrict A r : Restrict A `` RR"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 50 | by Auto_tac; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 51 | qed "Restrict_imageI"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 52 | |
| 7915 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 53 | Goal "Domain (Restrict A r) = A Int Domain r"; | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 54 | by (Blast_tac 1); | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 55 | qed "Domain_Restrict"; | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 56 | |
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 57 | Goal "(Restrict A r) ^^ B = r ^^ (A Int B)"; | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 58 | by (Blast_tac 1); | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 59 | qed "Image_Restrict"; | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 60 | |
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 61 | Addsimps [Domain_Restrict, Image_Restrict]; | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 62 | |
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 63 | |
| 7878 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7521diff
changeset | 64 | |
| 4776 | 65 | (*** lessThan ***) | 
| 66 | ||
| 5069 | 67 | Goalw [lessThan_def] "(i: lessThan k) = (i<k)"; | 
| 4776 | 68 | by (Blast_tac 1); | 
| 69 | qed "lessThan_iff"; | |
| 70 | AddIffs [lessThan_iff]; | |
| 71 | ||
| 5069 | 72 | Goalw [lessThan_def] "lessThan 0 = {}";
 | 
| 4776 | 73 | by (Simp_tac 1); | 
| 74 | qed "lessThan_0"; | |
| 75 | Addsimps [lessThan_0]; | |
| 76 | ||
| 5069 | 77 | Goalw [lessThan_def] "lessThan (Suc k) = insert k (lessThan k)"; | 
| 4776 | 78 | by (simp_tac (simpset() addsimps [less_Suc_eq]) 1); | 
| 79 | by (Blast_tac 1); | |
| 80 | qed "lessThan_Suc"; | |
| 81 | ||
| 5648 | 82 | Goalw [lessThan_def, atMost_def] "lessThan (Suc k) = atMost k"; | 
| 83 | by (simp_tac (simpset() addsimps [less_Suc_eq_le]) 1); | |
| 84 | qed "lessThan_Suc_atMost"; | |
| 85 | ||
| 6825 | 86 | Goal "finite (lessThan k)"; | 
| 87 | by (induct_tac "k" 1); | |
| 88 | by (simp_tac (simpset() addsimps [lessThan_Suc]) 2); | |
| 89 | by Auto_tac; | |
| 90 | qed "finite_lessThan"; | |
| 91 | ||
| 5069 | 92 | Goal "(UN m. lessThan m) = UNIV"; | 
| 4776 | 93 | by (Blast_tac 1); | 
| 94 | qed "UN_lessThan_UNIV"; | |
| 95 | ||
| 5069 | 96 | Goalw [lessThan_def, atLeast_def, le_def] | 
| 5490 | 97 | "-lessThan k = atLeast k"; | 
| 4776 | 98 | by (Blast_tac 1); | 
| 99 | qed "Compl_lessThan"; | |
| 100 | ||
| 7521 | 101 | Goal "{k} - lessThan k = {k}";
 | 
| 102 | by (Blast_tac 1); | |
| 103 | qed "single_Diff_lessThan"; | |
| 104 | Addsimps [single_Diff_lessThan]; | |
| 4776 | 105 | |
| 106 | (*** greaterThan ***) | |
| 107 | ||
| 5069 | 108 | Goalw [greaterThan_def] "(i: greaterThan k) = (k<i)"; | 
| 4776 | 109 | by (Blast_tac 1); | 
| 110 | qed "greaterThan_iff"; | |
| 111 | AddIffs [greaterThan_iff]; | |
| 112 | ||
| 5069 | 113 | Goalw [greaterThan_def] "greaterThan 0 = range Suc"; | 
| 4776 | 114 | by (blast_tac (claset() addIs [Suc_pred RS sym]) 1); | 
| 115 | qed "greaterThan_0"; | |
| 116 | Addsimps [greaterThan_0]; | |
| 117 | ||
| 5069 | 118 | Goalw [greaterThan_def] "greaterThan (Suc k) = greaterThan k - {Suc k}";
 | 
| 5625 | 119 | by (auto_tac (claset() addEs [linorder_neqE], simpset())); | 
| 4776 | 120 | qed "greaterThan_Suc"; | 
| 121 | ||
| 5069 | 122 | Goal "(INT m. greaterThan m) = {}";
 | 
| 4776 | 123 | by (Blast_tac 1); | 
| 124 | qed "INT_greaterThan_UNIV"; | |
| 125 | ||
| 5069 | 126 | Goalw [greaterThan_def, atMost_def, le_def] | 
| 5490 | 127 | "-greaterThan k = atMost k"; | 
| 4776 | 128 | by (Blast_tac 1); | 
| 129 | qed "Compl_greaterThan"; | |
| 130 | ||
| 5069 | 131 | Goalw [greaterThan_def, atMost_def, le_def] | 
| 5490 | 132 | "-atMost k = greaterThan k"; | 
| 4776 | 133 | by (Blast_tac 1); | 
| 134 | qed "Compl_atMost"; | |
| 135 | ||
| 5069 | 136 | Goal "less_than ^^ {k} = greaterThan k";
 | 
| 4776 | 137 | by (Blast_tac 1); | 
| 138 | qed "Image_less_than"; | |
| 139 | ||
| 140 | Addsimps [Compl_greaterThan, Compl_atMost, Image_less_than]; | |
| 141 | ||
| 142 | (*** atLeast ***) | |
| 143 | ||
| 5069 | 144 | Goalw [atLeast_def] "(i: atLeast k) = (k<=i)"; | 
| 4776 | 145 | by (Blast_tac 1); | 
| 146 | qed "atLeast_iff"; | |
| 147 | AddIffs [atLeast_iff]; | |
| 148 | ||
| 5069 | 149 | Goalw [atLeast_def, UNIV_def] "atLeast 0 = UNIV"; | 
| 4776 | 150 | by (Simp_tac 1); | 
| 151 | qed "atLeast_0"; | |
| 152 | Addsimps [atLeast_0]; | |
| 153 | ||
| 5069 | 154 | Goalw [atLeast_def] "atLeast (Suc k) = atLeast k - {k}";
 | 
| 4776 | 155 | by (simp_tac (simpset() addsimps [Suc_le_eq]) 1); | 
| 5596 | 156 | by (simp_tac (simpset() addsimps [order_le_less]) 1); | 
| 4776 | 157 | by (Blast_tac 1); | 
| 158 | qed "atLeast_Suc"; | |
| 159 | ||
| 5069 | 160 | Goal "(UN m. atLeast m) = UNIV"; | 
| 4776 | 161 | by (Blast_tac 1); | 
| 162 | qed "UN_atLeast_UNIV"; | |
| 163 | ||
| 5069 | 164 | Goalw [lessThan_def, atLeast_def, le_def] | 
| 5490 | 165 | "-atLeast k = lessThan k"; | 
| 4776 | 166 | by (Blast_tac 1); | 
| 167 | qed "Compl_atLeast"; | |
| 168 | ||
| 5069 | 169 | Goal "less_than^-1 ^^ {k} = lessThan k";
 | 
| 4776 | 170 | by (Blast_tac 1); | 
| 171 | qed "Image_inverse_less_than"; | |
| 172 | ||
| 173 | Addsimps [Compl_lessThan, Compl_atLeast, Image_inverse_less_than]; | |
| 174 | ||
| 175 | (*** atMost ***) | |
| 176 | ||
| 5069 | 177 | Goalw [atMost_def] "(i: atMost k) = (i<=k)"; | 
| 4776 | 178 | by (Blast_tac 1); | 
| 179 | qed "atMost_iff"; | |
| 180 | AddIffs [atMost_iff]; | |
| 181 | ||
| 5069 | 182 | Goalw [atMost_def] "atMost 0 = {0}";
 | 
| 4776 | 183 | by (Simp_tac 1); | 
| 184 | qed "atMost_0"; | |
| 185 | Addsimps [atMost_0]; | |
| 186 | ||
| 5069 | 187 | Goalw [atMost_def] "atMost (Suc k) = insert (Suc k) (atMost k)"; | 
| 5596 | 188 | by (simp_tac (simpset() addsimps [less_Suc_eq, order_le_less]) 1); | 
| 4776 | 189 | by (Blast_tac 1); | 
| 190 | qed "atMost_Suc"; | |
| 191 | ||
| 6825 | 192 | Goal "finite (atMost k)"; | 
| 193 | by (induct_tac "k" 1); | |
| 194 | by (simp_tac (simpset() addsimps [atMost_Suc]) 2); | |
| 195 | by Auto_tac; | |
| 196 | qed "finite_atMost"; | |
| 197 | ||
| 5069 | 198 | Goal "(UN m. atMost m) = UNIV"; | 
| 4776 | 199 | by (Blast_tac 1); | 
| 200 | qed "UN_atMost_UNIV"; | |
| 201 | ||
| 5069 | 202 | Goalw [atMost_def, le_def] | 
| 5490 | 203 | "-atMost k = greaterThan k"; | 
| 4776 | 204 | by (Blast_tac 1); | 
| 205 | qed "Compl_atMost"; | |
| 206 | Addsimps [Compl_atMost]; | |
| 207 | ||
| 208 | ||
| 209 | (*** Combined properties ***) | |
| 210 | ||
| 5069 | 211 | Goal "atMost n Int atLeast n = {n}";
 | 
| 6707 | 212 | by (blast_tac (claset() addIs [order_antisym]) 1); | 
| 4776 | 213 | qed "atMost_Int_atLeast"; | 
| 214 | ||
| 215 | ||
| 5232 | 216 | |
| 217 | ||
| 218 | (*** Finally, a few set-theoretic laws... | |
| 219 | CAN BOOLEAN SIMPLIFICATION BE AUTOMATED? ***) | |
| 220 | ||
| 221 | context Set.thy; | |
| 222 | ||
| 5490 | 223 | (** Rewrite rules to eliminate A. Conditions can be satisfied by letting B | 
| 5232 | 224 | be any set including A Int C and contained in A Un C, such as B=A or B=C. | 
| 225 | **) | |
| 226 | ||
| 227 | Goal "[| A Int C <= B; B <= A Un C |] \ | |
| 5490 | 228 | \ ==> (A Int B) Un (-A Int C) = B Un C"; | 
| 5232 | 229 | by (Blast_tac 1); | 
| 5490 | 230 | qed "set_cancel1"; | 
| 5232 | 231 | |
| 232 | Goal "[| A Int C <= B; B <= A Un C |] \ | |
| 5490 | 233 | \ ==> (A Un B) Int (-A Un C) = B Int C"; | 
| 5232 | 234 | by (Blast_tac 1); | 
| 5490 | 235 | qed "set_cancel2"; | 
| 5232 | 236 | |
| 237 | (*The base B=A*) | |
| 5490 | 238 | Goal "A Un (-A Int C) = A Un C"; | 
| 5232 | 239 | by (Blast_tac 1); | 
| 5490 | 240 | qed "set_cancel3"; | 
| 5232 | 241 | |
| 5490 | 242 | Goal "A Int (-A Un C) = A Int C"; | 
| 5232 | 243 | by (Blast_tac 1); | 
| 5490 | 244 | qed "set_cancel4"; | 
| 5232 | 245 | |
| 246 | (*The base B=C*) | |
| 5490 | 247 | Goal "(A Int C) Un (-A Int C) = C"; | 
| 5232 | 248 | by (Blast_tac 1); | 
| 5490 | 249 | qed "set_cancel5"; | 
| 5232 | 250 | |
| 5490 | 251 | Goal "(A Un C) Int (-A Un C) = C"; | 
| 5232 | 252 | by (Blast_tac 1); | 
| 5490 | 253 | qed "set_cancel6"; | 
| 254 | ||
| 255 | Addsimps [set_cancel1, set_cancel2, set_cancel3, | |
| 256 | set_cancel4, set_cancel5, set_cancel6]; | |
| 5232 | 257 | |
| 258 | ||
| 259 | (** More ad-hoc rules **) | |
| 260 | ||
| 261 | Goal "A Un B - (A - B) = B"; | |
| 262 | by (Blast_tac 1); | |
| 263 | qed "Un_Diff_Diff"; | |
| 5490 | 264 | Addsimps [Un_Diff_Diff]; | 
| 5232 | 265 | |
| 266 | Goal "A Int (B - C) Un C = A Int B Un C"; | |
| 267 | by (Blast_tac 1); | |
| 268 | qed "Int_Diff_Un"; | |
| 269 | ||
| 270 | Goal "Union(B) Int A = (UN C:B. C Int A)"; | |
| 271 | by (Blast_tac 1); | |
| 272 | qed "Int_Union2"; | |
| 273 | ||
| 274 | Goal "Union(B) Int A = Union((%C. C Int A)``B)"; | |
| 275 | by (Blast_tac 1); | |
| 276 | qed "Int_Union_Union"; | |
| 277 |