| author | wenzelm | 
| Wed, 06 Oct 1999 18:50:40 +0200 | |
| changeset 7760 | 43f8d28dbc6e | 
| parent 6176 | 707b6f9859d2 | 
| child 8127 | 68c6159440f1 | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/Ordinal.thy | 
| 435 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 435 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 5321 | 6 | Ordinals in Zermelo-Fraenkel Set Theory | 
| 435 | 7 | *) | 
| 8 | ||
| 9 | (*** Rules for Transset ***) | |
| 10 | ||
| 11 | (** Two neat characterisations of Transset **) | |
| 12 | ||
| 5067 | 13 | Goalw [Transset_def] "Transset(A) <-> A<=Pow(A)"; | 
| 2925 | 14 | by (Blast_tac 1); | 
| 760 | 15 | qed "Transset_iff_Pow"; | 
| 435 | 16 | |
| 5067 | 17 | Goalw [Transset_def] "Transset(A) <-> Union(succ(A)) = A"; | 
| 4091 | 18 | by (blast_tac (claset() addSEs [equalityE]) 1); | 
| 760 | 19 | qed "Transset_iff_Union_succ"; | 
| 435 | 20 | |
| 21 | (** Consequences of downwards closure **) | |
| 22 | ||
| 5067 | 23 | Goalw [Transset_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 24 |     "[| Transset(C); {a,b}: C |] ==> a:C & b: C";
 | 
| 2925 | 25 | by (Blast_tac 1); | 
| 760 | 26 | qed "Transset_doubleton_D"; | 
| 435 | 27 | |
| 28 | val [prem1,prem2] = goalw Ordinal.thy [Pair_def] | |
| 29 | "[| Transset(C); <a,b>: C |] ==> a:C & b: C"; | |
| 1461 | 30 | by (cut_facts_tac [prem2] 1); | 
| 4091 | 31 | by (blast_tac (claset() addSDs [prem1 RS Transset_doubleton_D]) 1); | 
| 760 | 32 | qed "Transset_Pair_D"; | 
| 435 | 33 | |
| 34 | val prem1::prems = goal Ordinal.thy | |
| 35 | "[| Transset(C); A*B <= C; b: B |] ==> A <= C"; | |
| 36 | by (cut_facts_tac prems 1); | |
| 4091 | 37 | by (blast_tac (claset() addSDs [prem1 RS Transset_Pair_D]) 1); | 
| 760 | 38 | qed "Transset_includes_domain"; | 
| 435 | 39 | |
| 40 | val prem1::prems = goal Ordinal.thy | |
| 41 | "[| Transset(C); A*B <= C; a: A |] ==> B <= C"; | |
| 42 | by (cut_facts_tac prems 1); | |
| 4091 | 43 | by (blast_tac (claset() addSDs [prem1 RS Transset_Pair_D]) 1); | 
| 760 | 44 | qed "Transset_includes_range"; | 
| 435 | 45 | |
| 46 | (** Closure properties **) | |
| 47 | ||
| 5067 | 48 | Goalw [Transset_def] "Transset(0)"; | 
| 2925 | 49 | by (Blast_tac 1); | 
| 760 | 50 | qed "Transset_0"; | 
| 435 | 51 | |
| 5067 | 52 | Goalw [Transset_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 53 | "[| Transset(i); Transset(j) |] ==> Transset(i Un j)"; | 
| 2925 | 54 | by (Blast_tac 1); | 
| 760 | 55 | qed "Transset_Un"; | 
| 435 | 56 | |
| 5067 | 57 | Goalw [Transset_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 58 | "[| Transset(i); Transset(j) |] ==> Transset(i Int j)"; | 
| 2925 | 59 | by (Blast_tac 1); | 
| 760 | 60 | qed "Transset_Int"; | 
| 435 | 61 | |
| 5137 | 62 | Goalw [Transset_def] "Transset(i) ==> Transset(succ(i))"; | 
| 2925 | 63 | by (Blast_tac 1); | 
| 760 | 64 | qed "Transset_succ"; | 
| 435 | 65 | |
| 5137 | 66 | Goalw [Transset_def] "Transset(i) ==> Transset(Pow(i))"; | 
| 2925 | 67 | by (Blast_tac 1); | 
| 760 | 68 | qed "Transset_Pow"; | 
| 435 | 69 | |
| 5137 | 70 | Goalw [Transset_def] "Transset(A) ==> Transset(Union(A))"; | 
| 2925 | 71 | by (Blast_tac 1); | 
| 760 | 72 | qed "Transset_Union"; | 
| 435 | 73 | |
| 5321 | 74 | val [Transprem] = Goalw [Transset_def] | 
| 435 | 75 | "[| !!i. i:A ==> Transset(i) |] ==> Transset(Union(A))"; | 
| 4091 | 76 | by (blast_tac (claset() addDs [Transprem RS bspec RS subsetD]) 1); | 
| 760 | 77 | qed "Transset_Union_family"; | 
| 435 | 78 | |
| 5321 | 79 | val [prem,Transprem] = Goalw [Transset_def] | 
| 435 | 80 | "[| j:A; !!i. i:A ==> Transset(i) |] ==> Transset(Inter(A))"; | 
| 81 | by (cut_facts_tac [prem] 1); | |
| 4091 | 82 | by (blast_tac (claset() addDs [Transprem RS bspec RS subsetD]) 1); | 
| 760 | 83 | qed "Transset_Inter_family"; | 
| 435 | 84 | |
| 85 | (*** Natural Deduction rules for Ord ***) | |
| 86 | ||
| 5321 | 87 | val prems = Goalw [Ord_def] | 
| 2717 
b29c45ef3d86
best_tac avoids looping with change to RepFun_eqI in claset
 paulson parents: 
2493diff
changeset | 88 | "[| Transset(i); !!x. x:i ==> Transset(x) |] ==> Ord(i)"; | 
| 435 | 89 | by (REPEAT (ares_tac (prems@[ballI,conjI]) 1)); | 
| 760 | 90 | qed "OrdI"; | 
| 435 | 91 | |
| 5321 | 92 | Goalw [Ord_def] "Ord(i) ==> Transset(i)"; | 
| 93 | by (Blast_tac 1); | |
| 760 | 94 | qed "Ord_is_Transset"; | 
| 435 | 95 | |
| 5321 | 96 | Goalw [Ord_def] | 
| 435 | 97 | "[| Ord(i); j:i |] ==> Transset(j) "; | 
| 5321 | 98 | by (Blast_tac 1); | 
| 760 | 99 | qed "Ord_contains_Transset"; | 
| 435 | 100 | |
| 101 | (*** Lemmas for ordinals ***) | |
| 102 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5137diff
changeset | 103 | Goalw [Ord_def,Transset_def] "[| Ord(i); j:i |] ==> Ord(j)"; | 
| 2925 | 104 | by (Blast_tac 1); | 
| 760 | 105 | qed "Ord_in_Ord"; | 
| 435 | 106 | |
| 107 | (* Ord(succ(j)) ==> Ord(j) *) | |
| 108 | val Ord_succD = succI1 RSN (2, Ord_in_Ord); | |
| 109 | ||
| 3016 | 110 | AddSDs [Ord_succD]; | 
| 111 | ||
| 5137 | 112 | Goal "[| Ord(i); Transset(j); j<=i |] ==> Ord(j)"; | 
| 435 | 113 | by (REPEAT (ares_tac [OrdI] 1 | 
| 114 | ORELSE eresolve_tac [Ord_contains_Transset, subsetD] 1)); | |
| 760 | 115 | qed "Ord_subset_Ord"; | 
| 435 | 116 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5137diff
changeset | 117 | Goalw [Ord_def,Transset_def] "[| j:i; Ord(i) |] ==> j<=i"; | 
| 2925 | 118 | by (Blast_tac 1); | 
| 760 | 119 | qed "OrdmemD"; | 
| 435 | 120 | |
| 5137 | 121 | Goal "[| i:j; j:k; Ord(k) |] ==> i:k"; | 
| 435 | 122 | by (REPEAT (ares_tac [OrdmemD RS subsetD] 1)); | 
| 760 | 123 | qed "Ord_trans"; | 
| 435 | 124 | |
| 5137 | 125 | Goal "[| i:j; Ord(j) |] ==> succ(i) <= j"; | 
| 435 | 126 | by (REPEAT (ares_tac [OrdmemD RSN (2,succ_subsetI)] 1)); | 
| 760 | 127 | qed "Ord_succ_subsetI"; | 
| 435 | 128 | |
| 129 | ||
| 130 | (*** The construction of ordinals: 0, succ, Union ***) | |
| 131 | ||
| 5067 | 132 | Goal "Ord(0)"; | 
| 435 | 133 | by (REPEAT (ares_tac [OrdI,Transset_0] 1 ORELSE etac emptyE 1)); | 
| 760 | 134 | qed "Ord_0"; | 
| 435 | 135 | |
| 5137 | 136 | Goal "Ord(i) ==> Ord(succ(i))"; | 
| 435 | 137 | by (REPEAT (ares_tac [OrdI,Transset_succ] 1 | 
| 138 | ORELSE eresolve_tac [succE,ssubst,Ord_is_Transset, | |
| 1461 | 139 | Ord_contains_Transset] 1)); | 
| 760 | 140 | qed "Ord_succ"; | 
| 435 | 141 | |
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 142 | bind_thm ("Ord_1", Ord_0 RS Ord_succ);
 | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 143 | |
| 5067 | 144 | Goal "Ord(succ(i)) <-> Ord(i)"; | 
| 4091 | 145 | by (blast_tac (claset() addIs [Ord_succ] addDs [Ord_succD]) 1); | 
| 760 | 146 | qed "Ord_succ_iff"; | 
| 435 | 147 | |
| 2469 | 148 | Addsimps [Ord_0, Ord_succ_iff]; | 
| 149 | AddSIs [Ord_0, Ord_succ]; | |
| 6153 | 150 | AddTCs [Ord_0, Ord_succ]; | 
| 2469 | 151 | |
| 5137 | 152 | Goalw [Ord_def] "[| Ord(i); Ord(j) |] ==> Ord(i Un j)"; | 
| 4091 | 153 | by (blast_tac (claset() addSIs [Transset_Un]) 1); | 
| 760 | 154 | qed "Ord_Un"; | 
| 435 | 155 | |
| 5137 | 156 | Goalw [Ord_def] "[| Ord(i); Ord(j) |] ==> Ord(i Int j)"; | 
| 4091 | 157 | by (blast_tac (claset() addSIs [Transset_Int]) 1); | 
| 760 | 158 | qed "Ord_Int"; | 
| 6153 | 159 | AddTCs [Ord_Un, Ord_Int]; | 
| 435 | 160 | |
| 5321 | 161 | val nonempty::prems = Goal | 
| 435 | 162 | "[| j:A; !!i. i:A ==> Ord(i) |] ==> Ord(Inter(A))"; | 
| 163 | by (rtac (nonempty RS Transset_Inter_family RS OrdI) 1); | |
| 164 | by (rtac Ord_is_Transset 1); | |
| 165 | by (REPEAT (ares_tac ([Ord_contains_Transset,nonempty]@prems) 1 | |
| 166 | ORELSE etac InterD 1)); | |
| 760 | 167 | qed "Ord_Inter"; | 
| 435 | 168 | |
| 5321 | 169 | val jmemA::prems = Goal | 
| 435 | 170 | "[| j:A; !!x. x:A ==> Ord(B(x)) |] ==> Ord(INT x:A. B(x))"; | 
| 171 | by (rtac (jmemA RS RepFunI RS Ord_Inter) 1); | |
| 172 | by (etac RepFunE 1); | |
| 173 | by (etac ssubst 1); | |
| 174 | by (eresolve_tac prems 1); | |
| 760 | 175 | qed "Ord_INT"; | 
| 435 | 176 | |
| 177 | (*There is no set of all ordinals, for then it would contain itself*) | |
| 5067 | 178 | Goal "~ (ALL i. i:X <-> Ord(i))"; | 
| 435 | 179 | by (rtac notI 1); | 
| 180 | by (forw_inst_tac [("x", "X")] spec 1);
 | |
| 4091 | 181 | by (safe_tac (claset() addSEs [mem_irrefl])); | 
| 435 | 182 | by (swap_res_tac [Ord_is_Transset RSN (2,OrdI)] 1); | 
| 2925 | 183 | by (Blast_tac 2); | 
| 437 | 184 | by (rewtac Transset_def); | 
| 4152 | 185 | by Safe_tac; | 
| 2469 | 186 | by (Asm_full_simp_tac 1); | 
| 435 | 187 | by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 1)); | 
| 760 | 188 | qed "ON_class"; | 
| 435 | 189 | |
| 190 | (*** < is 'less than' for ordinals ***) | |
| 191 | ||
| 5137 | 192 | Goalw [lt_def] "[| i:j; Ord(j) |] ==> i<j"; | 
| 435 | 193 | by (REPEAT (ares_tac [conjI] 1)); | 
| 760 | 194 | qed "ltI"; | 
| 435 | 195 | |
| 5321 | 196 | val major::prems = Goalw [lt_def] | 
| 435 | 197 | "[| i<j; [| i:j; Ord(i); Ord(j) |] ==> P |] ==> P"; | 
| 198 | by (rtac (major RS conjE) 1); | |
| 199 | by (REPEAT (ares_tac (prems@[Ord_in_Ord]) 1)); | |
| 760 | 200 | qed "ltE"; | 
| 435 | 201 | |
| 5137 | 202 | Goal "i<j ==> i:j"; | 
| 435 | 203 | by (etac ltE 1); | 
| 204 | by (assume_tac 1); | |
| 760 | 205 | qed "ltD"; | 
| 435 | 206 | |
| 5067 | 207 | Goalw [lt_def] "~ i<0"; | 
| 2925 | 208 | by (Blast_tac 1); | 
| 760 | 209 | qed "not_lt0"; | 
| 435 | 210 | |
| 2469 | 211 | Addsimps [not_lt0]; | 
| 212 | ||
| 5137 | 213 | Goal "j<i ==> Ord(j)"; | 
| 1461 | 214 | by (etac ltE 1 THEN assume_tac 1); | 
| 830 
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
 lcp parents: 
782diff
changeset | 215 | qed "lt_Ord"; | 
| 
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
 lcp parents: 
782diff
changeset | 216 | |
| 5137 | 217 | Goal "j<i ==> Ord(i)"; | 
| 1461 | 218 | by (etac ltE 1 THEN assume_tac 1); | 
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 219 | qed "lt_Ord2"; | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 220 | |
| 1034 | 221 | (* "ja le j ==> Ord(j)" *) | 
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 222 | bind_thm ("le_Ord2", lt_Ord2 RS Ord_succD);
 | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 223 | |
| 435 | 224 | (* i<0 ==> R *) | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
772diff
changeset | 225 | bind_thm ("lt0E", not_lt0 RS notE);
 | 
| 435 | 226 | |
| 5137 | 227 | Goal "[| i<j; j<k |] ==> i<k"; | 
| 4091 | 228 | by (blast_tac (claset() addSIs [ltI] addSEs [ltE] addIs [Ord_trans]) 1); | 
| 760 | 229 | qed "lt_trans"; | 
| 435 | 230 | |
| 5465 | 231 | Goalw [lt_def] "i<j ==> ~ (j<i)"; | 
| 232 | by (blast_tac (claset() addEs [mem_asym]) 1); | |
| 233 | qed "lt_not_sym"; | |
| 234 | ||
| 235 | (* [| i<j; ~P ==> j<i |] ==> P *) | |
| 236 | bind_thm ("lt_asym", lt_not_sym RS swap);
 | |
| 435 | 237 | |
| 760 | 238 | qed_goal "lt_irrefl" Ordinal.thy "i<i ==> P" | 
| 437 | 239 | (fn [major]=> [ (rtac (major RS (major RS lt_asym)) 1) ]); | 
| 435 | 240 | |
| 760 | 241 | qed_goal "lt_not_refl" Ordinal.thy "~ i<i" | 
| 437 | 242 | (fn _=> [ (rtac notI 1), (etac lt_irrefl 1) ]); | 
| 435 | 243 | |
| 2469 | 244 | AddSEs [lt_irrefl, lt0E]; | 
| 245 | ||
| 435 | 246 | (** le is less than or equals; recall i le j abbrevs i<succ(j) !! **) | 
| 247 | ||
| 5067 | 248 | Goalw [lt_def] "i le j <-> i<j | (i=j & Ord(j))"; | 
| 4091 | 249 | by (blast_tac (claset() addSIs [Ord_succ] addSDs [Ord_succD]) 1); | 
| 760 | 250 | qed "le_iff"; | 
| 435 | 251 | |
| 772 | 252 | (*Equivalently, i<j ==> i < succ(j)*) | 
| 5137 | 253 | Goal "i<j ==> i le j"; | 
| 4091 | 254 | by (asm_simp_tac (simpset() addsimps [le_iff]) 1); | 
| 760 | 255 | qed "leI"; | 
| 435 | 256 | |
| 5137 | 257 | Goal "[| i=j; Ord(j) |] ==> i le j"; | 
| 4091 | 258 | by (asm_simp_tac (simpset() addsimps [le_iff]) 1); | 
| 760 | 259 | qed "le_eqI"; | 
| 435 | 260 | |
| 261 | val le_refl = refl RS le_eqI; | |
| 262 | ||
| 5321 | 263 | val [prem] = Goal "(~ (i=j & Ord(j)) ==> i<j) ==> i le j"; | 
| 435 | 264 | by (rtac (disjCI RS (le_iff RS iffD2)) 1); | 
| 265 | by (etac prem 1); | |
| 760 | 266 | qed "leCI"; | 
| 435 | 267 | |
| 5321 | 268 | val major::prems = Goal | 
| 435 | 269 | "[| i le j; i<j ==> P; [| i=j; Ord(j) |] ==> P |] ==> P"; | 
| 270 | by (rtac (major RS (le_iff RS iffD1 RS disjE)) 1); | |
| 271 | by (DEPTH_SOLVE (ares_tac prems 1 ORELSE etac conjE 1)); | |
| 760 | 272 | qed "leE"; | 
| 435 | 273 | |
| 5137 | 274 | Goal "[| i le j; j le i |] ==> i=j"; | 
| 4091 | 275 | by (asm_full_simp_tac (simpset() addsimps [le_iff]) 1); | 
| 276 | by (blast_tac (claset() addEs [lt_asym]) 1); | |
| 760 | 277 | qed "le_anti_sym"; | 
| 435 | 278 | |
| 5067 | 279 | Goal "i le 0 <-> i=0"; | 
| 4091 | 280 | by (blast_tac (claset() addSIs [Ord_0 RS le_refl] addSEs [leE]) 1); | 
| 760 | 281 | qed "le0_iff"; | 
| 435 | 282 | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
772diff
changeset | 283 | bind_thm ("le0D", le0_iff RS iffD1);
 | 
| 435 | 284 | |
| 2469 | 285 | AddIs [le_refl]; | 
| 286 | AddSDs [le0D]; | |
| 287 | Addsimps [le0_iff]; | |
| 288 | ||
| 4091 | 289 | val le_cs = claset() addSIs [leCI] addSEs [leE] addEs [lt_asym]; | 
| 435 | 290 | |
| 291 | ||
| 292 | (*** Natural Deduction rules for Memrel ***) | |
| 293 | ||
| 5067 | 294 | Goalw [Memrel_def] "<a,b> : Memrel(A) <-> a:b & a:A & b:A"; | 
| 2925 | 295 | by (Blast_tac 1); | 
| 760 | 296 | qed "Memrel_iff"; | 
| 435 | 297 | |
| 5137 | 298 | Goal "[| a: b; a: A; b: A |] ==> <a,b> : Memrel(A)"; | 
| 2925 | 299 | by (REPEAT (ares_tac [conjI, Memrel_iff RS iffD2] 1)); | 
| 760 | 300 | qed "MemrelI"; | 
| 435 | 301 | |
| 5321 | 302 | val [major,minor] = Goal | 
| 435 | 303 | "[| <a,b> : Memrel(A); \ | 
| 304 | \ [| a: A; b: A; a:b |] ==> P \ | |
| 305 | \ |] ==> P"; | |
| 306 | by (rtac (major RS (Memrel_iff RS iffD1) RS conjE) 1); | |
| 307 | by (etac conjE 1); | |
| 308 | by (rtac minor 1); | |
| 309 | by (REPEAT (assume_tac 1)); | |
| 760 | 310 | qed "MemrelE"; | 
| 435 | 311 | |
| 2925 | 312 | AddSIs [MemrelI]; | 
| 313 | AddSEs [MemrelE]; | |
| 314 | ||
| 5067 | 315 | Goalw [Memrel_def] "Memrel(A) <= A*A"; | 
| 2925 | 316 | by (Blast_tac 1); | 
| 830 
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
 lcp parents: 
782diff
changeset | 317 | qed "Memrel_type"; | 
| 
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
 lcp parents: 
782diff
changeset | 318 | |
| 5137 | 319 | Goalw [Memrel_def] "A<=B ==> Memrel(A) <= Memrel(B)"; | 
| 2925 | 320 | by (Blast_tac 1); | 
| 830 
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
 lcp parents: 
782diff
changeset | 321 | qed "Memrel_mono"; | 
| 
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
 lcp parents: 
782diff
changeset | 322 | |
| 5067 | 323 | Goalw [Memrel_def] "Memrel(0) = 0"; | 
| 2925 | 324 | by (Blast_tac 1); | 
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 325 | qed "Memrel_0"; | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 326 | |
| 5067 | 327 | Goalw [Memrel_def] "Memrel(1) = 0"; | 
| 2925 | 328 | by (Blast_tac 1); | 
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 329 | qed "Memrel_1"; | 
| 830 
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
 lcp parents: 
782diff
changeset | 330 | |
| 2469 | 331 | Addsimps [Memrel_0, Memrel_1]; | 
| 332 | ||
| 435 | 333 | (*The membership relation (as a set) is well-founded. | 
| 334 | Proof idea: show A<=B by applying the foundation axiom to A-B *) | |
| 5067 | 335 | Goalw [wf_def] "wf(Memrel(A))"; | 
| 435 | 336 | by (EVERY1 [rtac (foundation RS disjE RS allI), | 
| 1461 | 337 | etac disjI1, | 
| 338 | etac bexE, | |
| 339 | rtac (impI RS allI RS bexI RS disjI2), | |
| 340 | etac MemrelE, | |
| 341 | etac bspec, | |
| 342 | REPEAT o assume_tac]); | |
| 760 | 343 | qed "wf_Memrel"; | 
| 435 | 344 | |
| 345 | (*Transset(i) does not suffice, though ALL j:i.Transset(j) does*) | |
| 5067 | 346 | Goalw [Ord_def, Transset_def, trans_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 347 | "Ord(i) ==> trans(Memrel(i))"; | 
| 2925 | 348 | by (Blast_tac 1); | 
| 760 | 349 | qed "trans_Memrel"; | 
| 435 | 350 | |
| 351 | (*If Transset(A) then Memrel(A) internalizes the membership relation below A*) | |
| 5067 | 352 | Goalw [Transset_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 353 | "Transset(A) ==> <a,b> : Memrel(A) <-> a:b & b:A"; | 
| 2925 | 354 | by (Blast_tac 1); | 
| 760 | 355 | qed "Transset_Memrel_iff"; | 
| 435 | 356 | |
| 357 | ||
| 358 | (*** Transfinite induction ***) | |
| 359 | ||
| 360 | (*Epsilon induction over a transitive set*) | |
| 5321 | 361 | val major::prems = Goalw [Transset_def] | 
| 435 | 362 | "[| i: k; Transset(k); \ | 
| 363 | \ !!x.[| x: k; ALL y:x. P(y) |] ==> P(x) \ | |
| 364 | \ |] ==> P(i)"; | |
| 365 | by (rtac (major RS (wf_Memrel RS wf_induct2)) 1); | |
| 2925 | 366 | by (Blast_tac 1); | 
| 435 | 367 | by (resolve_tac prems 1); | 
| 368 | by (assume_tac 1); | |
| 369 | by (cut_facts_tac prems 1); | |
| 2925 | 370 | by (Blast_tac 1); | 
| 760 | 371 | qed "Transset_induct"; | 
| 435 | 372 | |
| 373 | (*Induction over an ordinal*) | |
| 374 | val Ord_induct = Ord_is_Transset RSN (2, Transset_induct); | |
| 375 | ||
| 376 | (*Induction over the class of ordinals -- a useful corollary of Ord_induct*) | |
| 5321 | 377 | val [major,indhyp] = Goal | 
| 435 | 378 | "[| Ord(i); \ | 
| 379 | \ !!x.[| Ord(x); ALL y:x. P(y) |] ==> P(x) \ | |
| 380 | \ |] ==> P(i)"; | |
| 381 | by (rtac (major RS Ord_succ RS (succI1 RS Ord_induct)) 1); | |
| 382 | by (rtac indhyp 1); | |
| 383 | by (rtac (major RS Ord_succ RS Ord_in_Ord) 1); | |
| 384 | by (REPEAT (assume_tac 1)); | |
| 760 | 385 | qed "trans_induct"; | 
| 435 | 386 | |
| 387 | (*Perform induction on i, then prove the Ord(i) subgoal using prems. *) | |
| 388 | fun trans_ind_tac a prems i = | |
| 389 |     EVERY [res_inst_tac [("i",a)] trans_induct i,
 | |
| 1461 | 390 | rename_last_tac a ["1"] (i+1), | 
| 391 | ares_tac prems i]; | |
| 435 | 392 | |
| 393 | ||
| 394 | (*** Fundamental properties of the epsilon ordering (< on ordinals) ***) | |
| 395 | ||
| 396 | (*Finds contradictions for the following proof*) | |
| 397 | val Ord_trans_tac = EVERY' [etac notE, etac Ord_trans, REPEAT o atac]; | |
| 398 | ||
| 399 | (** Proving that < is a linear ordering on the ordinals **) | |
| 400 | ||
| 5321 | 401 | Goal "Ord(i) ==> (ALL j. Ord(j) --> i:j | i=j | j:i)"; | 
| 402 | by (etac trans_induct 1); | |
| 435 | 403 | by (rtac (impI RS allI) 1); | 
| 404 | by (trans_ind_tac "j" [] 1); | |
| 2493 | 405 | by (DEPTH_SOLVE (Step_tac 1 ORELSE Ord_trans_tac 1)); | 
| 3736 
39ee3d31cfbc
Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
 paulson parents: 
3016diff
changeset | 406 | qed_spec_mp "Ord_linear"; | 
| 435 | 407 | |
| 408 | (*The trichotomy law for ordinals!*) | |
| 5321 | 409 | val ordi::ordj::prems = Goalw [lt_def] | 
| 435 | 410 | "[| Ord(i); Ord(j); i<j ==> P; i=j ==> P; j<i ==> P |] ==> P"; | 
| 411 | by (rtac ([ordi,ordj] MRS Ord_linear RS disjE) 1); | |
| 412 | by (etac disjE 2); | |
| 413 | by (DEPTH_SOLVE (ares_tac ([ordi,ordj,conjI] @ prems) 1)); | |
| 760 | 414 | qed "Ord_linear_lt"; | 
| 435 | 415 | |
| 5321 | 416 | val prems = Goal | 
| 435 | 417 | "[| Ord(i); Ord(j); i<j ==> P; j le i ==> P |] ==> P"; | 
| 418 | by (res_inst_tac [("i","i"),("j","j")] Ord_linear_lt 1);
 | |
| 419 | by (DEPTH_SOLVE (ares_tac ([leI, sym RS le_eqI] @ prems) 1)); | |
| 760 | 420 | qed "Ord_linear2"; | 
| 435 | 421 | |
| 5321 | 422 | val prems = Goal | 
| 435 | 423 | "[| Ord(i); Ord(j); i le j ==> P; j le i ==> P |] ==> P"; | 
| 424 | by (res_inst_tac [("i","i"),("j","j")] Ord_linear_lt 1);
 | |
| 425 | by (DEPTH_SOLVE (ares_tac ([leI,le_eqI] @ prems) 1)); | |
| 760 | 426 | qed "Ord_linear_le"; | 
| 435 | 427 | |
| 5137 | 428 | Goal "j le i ==> ~ i<j"; | 
| 2925 | 429 | by (blast_tac le_cs 1); | 
| 760 | 430 | qed "le_imp_not_lt"; | 
| 435 | 431 | |
| 5137 | 432 | Goal "[| ~ i<j; Ord(i); Ord(j) |] ==> j le i"; | 
| 435 | 433 | by (res_inst_tac [("i","i"),("j","j")] Ord_linear2 1);
 | 
| 434 | by (REPEAT (SOMEGOAL assume_tac)); | |
| 2925 | 435 | by (blast_tac le_cs 1); | 
| 760 | 436 | qed "not_lt_imp_le"; | 
| 435 | 437 | |
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 438 | (** Some rewrite rules for <, le **) | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 439 | |
| 5137 | 440 | Goalw [lt_def] "Ord(j) ==> i:j <-> i<j"; | 
| 2925 | 441 | by (Blast_tac 1); | 
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 442 | qed "Ord_mem_iff_lt"; | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 443 | |
| 5137 | 444 | Goal "[| Ord(i); Ord(j) |] ==> ~ i<j <-> j le i"; | 
| 435 | 445 | by (REPEAT (ares_tac [iffI, le_imp_not_lt, not_lt_imp_le] 1)); | 
| 760 | 446 | qed "not_lt_iff_le"; | 
| 435 | 447 | |
| 5137 | 448 | Goal "[| Ord(i); Ord(j) |] ==> ~ i le j <-> j<i"; | 
| 4091 | 449 | by (asm_simp_tac (simpset() addsimps [not_lt_iff_le RS iff_sym]) 1); | 
| 760 | 450 | qed "not_le_iff_lt"; | 
| 435 | 451 | |
| 1610 | 452 | (*This is identical to 0<succ(i) *) | 
| 5137 | 453 | Goal "Ord(i) ==> 0 le i"; | 
| 435 | 454 | by (etac (not_lt_iff_le RS iffD1) 1); | 
| 455 | by (REPEAT (resolve_tac [Ord_0, not_lt0] 1)); | |
| 760 | 456 | qed "Ord_0_le"; | 
| 435 | 457 | |
| 5137 | 458 | Goal "[| Ord(i); i~=0 |] ==> 0<i"; | 
| 435 | 459 | by (etac (not_le_iff_lt RS iffD1) 1); | 
| 460 | by (rtac Ord_0 1); | |
| 2925 | 461 | by (Blast_tac 1); | 
| 760 | 462 | qed "Ord_0_lt"; | 
| 435 | 463 | |
| 464 | (*** Results about less-than or equals ***) | |
| 465 | ||
| 466 | (** For ordinals, j<=i (subset) implies j le i (less-than or equals) **) | |
| 467 | ||
| 5137 | 468 | Goal "[| j<=i; Ord(i); Ord(j) |] ==> j le i"; | 
| 435 | 469 | by (rtac (not_lt_iff_le RS iffD1) 1); | 
| 470 | by (assume_tac 1); | |
| 471 | by (assume_tac 1); | |
| 4091 | 472 | by (blast_tac (claset() addEs [ltE, mem_irrefl]) 1); | 
| 760 | 473 | qed "subset_imp_le"; | 
| 435 | 474 | |
| 5137 | 475 | Goal "i le j ==> i<=j"; | 
| 435 | 476 | by (etac leE 1); | 
| 2925 | 477 | by (Blast_tac 2); | 
| 478 | by (blast_tac (subset_cs addIs [OrdmemD] addEs [ltE]) 1); | |
| 760 | 479 | qed "le_imp_subset"; | 
| 435 | 480 | |
| 5067 | 481 | Goal "j le i <-> j<=i & Ord(i) & Ord(j)"; | 
| 4091 | 482 | by (blast_tac (claset() addDs [Ord_succD, subset_imp_le, le_imp_subset] | 
| 3016 | 483 | addEs [ltE]) 1); | 
| 760 | 484 | qed "le_subset_iff"; | 
| 435 | 485 | |
| 5067 | 486 | Goal "i le succ(j) <-> i le j | i=succ(j) & Ord(i)"; | 
| 4091 | 487 | by (simp_tac (simpset() addsimps [le_iff]) 1); | 
| 488 | by (blast_tac (claset() addIs [Ord_succ] addDs [Ord_succD]) 1); | |
| 760 | 489 | qed "le_succ_iff"; | 
| 435 | 490 | |
| 491 | (*Just a variant of subset_imp_le*) | |
| 5321 | 492 | val [ordi,ordj,minor] = Goal | 
| 435 | 493 | "[| Ord(i); Ord(j); !!x. x<j ==> x<i |] ==> j le i"; | 
| 494 | by (REPEAT_FIRST (ares_tac [notI RS not_lt_imp_le, ordi, ordj])); | |
| 437 | 495 | by (etac (minor RS lt_irrefl) 1); | 
| 760 | 496 | qed "all_lt_imp_le"; | 
| 435 | 497 | |
| 498 | (** Transitive laws **) | |
| 499 | ||
| 5137 | 500 | Goal "[| i le j; j<k |] ==> i<k"; | 
| 4091 | 501 | by (blast_tac (claset() addSEs [leE] addIs [lt_trans]) 1); | 
| 760 | 502 | qed "lt_trans1"; | 
| 435 | 503 | |
| 5137 | 504 | Goal "[| i<j; j le k |] ==> i<k"; | 
| 4091 | 505 | by (blast_tac (claset() addSEs [leE] addIs [lt_trans]) 1); | 
| 760 | 506 | qed "lt_trans2"; | 
| 435 | 507 | |
| 5137 | 508 | Goal "[| i le j; j le k |] ==> i le k"; | 
| 435 | 509 | by (REPEAT (ares_tac [lt_trans1] 1)); | 
| 760 | 510 | qed "le_trans"; | 
| 435 | 511 | |
| 5137 | 512 | Goal "i<j ==> succ(i) le j"; | 
| 435 | 513 | by (rtac (not_lt_iff_le RS iffD1) 1); | 
| 2925 | 514 | by (blast_tac le_cs 3); | 
| 4091 | 515 | by (ALLGOALS (blast_tac (claset() addEs [ltE]))); | 
| 760 | 516 | qed "succ_leI"; | 
| 435 | 517 | |
| 830 
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
 lcp parents: 
782diff
changeset | 518 | (*Identical to succ(i) < succ(j) ==> i<j *) | 
| 5137 | 519 | Goal "succ(i) le j ==> i<j"; | 
| 435 | 520 | by (rtac (not_le_iff_lt RS iffD1) 1); | 
| 4475 | 521 | by (blast_tac le_cs 3); | 
| 4091 | 522 | by (ALLGOALS (blast_tac (claset() addEs [ltE, make_elim Ord_succD]))); | 
| 760 | 523 | qed "succ_leE"; | 
| 435 | 524 | |
| 5067 | 525 | Goal "succ(i) le j <-> i<j"; | 
| 435 | 526 | by (REPEAT (ares_tac [iffI,succ_leI,succ_leE] 1)); | 
| 760 | 527 | qed "succ_le_iff"; | 
| 435 | 528 | |
| 2469 | 529 | Addsimps [succ_le_iff]; | 
| 530 | ||
| 5137 | 531 | Goal "succ(i) le succ(j) ==> i le j"; | 
| 4091 | 532 | by (blast_tac (claset() addSDs [succ_leE]) 1); | 
| 830 
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
 lcp parents: 
782diff
changeset | 533 | qed "succ_le_imp_le"; | 
| 
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
 lcp parents: 
782diff
changeset | 534 | |
| 6176 
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
 paulson parents: 
6153diff
changeset | 535 | Goal "[| i <= j; j<k; Ord(i) |] ==> i<k"; | 
| 
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
 paulson parents: 
6153diff
changeset | 536 | by (resolve_tac [subset_imp_le RS lt_trans1] 1); | 
| 
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
 paulson parents: 
6153diff
changeset | 537 | by (REPEAT (eresolve_tac [asm_rl, ltE] 1)); | 
| 
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
 paulson parents: 
6153diff
changeset | 538 | qed "lt_subset_trans"; | 
| 
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
 paulson parents: 
6153diff
changeset | 539 | |
| 435 | 540 | (** Union and Intersection **) | 
| 541 | ||
| 5137 | 542 | Goal "[| Ord(i); Ord(j) |] ==> i le i Un j"; | 
| 435 | 543 | by (rtac (Un_upper1 RS subset_imp_le) 1); | 
| 544 | by (REPEAT (ares_tac [Ord_Un] 1)); | |
| 760 | 545 | qed "Un_upper1_le"; | 
| 435 | 546 | |
| 5137 | 547 | Goal "[| Ord(i); Ord(j) |] ==> j le i Un j"; | 
| 435 | 548 | by (rtac (Un_upper2 RS subset_imp_le) 1); | 
| 549 | by (REPEAT (ares_tac [Ord_Un] 1)); | |
| 760 | 550 | qed "Un_upper2_le"; | 
| 435 | 551 | |
| 552 | (*Replacing k by succ(k') yields the similar rule for le!*) | |
| 5137 | 553 | Goal "[| i<k; j<k |] ==> i Un j < k"; | 
| 435 | 554 | by (res_inst_tac [("i","i"),("j","j")] Ord_linear_le 1);
 | 
| 2033 | 555 | by (stac Un_commute 4); | 
| 4091 | 556 | by (asm_full_simp_tac (simpset() addsimps [le_subset_iff, subset_Un_iff]) 4); | 
| 557 | by (asm_full_simp_tac (simpset() addsimps [le_subset_iff, subset_Un_iff]) 3); | |
| 435 | 558 | by (REPEAT (eresolve_tac [asm_rl, ltE] 1)); | 
| 760 | 559 | qed "Un_least_lt"; | 
| 435 | 560 | |
| 5137 | 561 | Goal "[| Ord(i); Ord(j) |] ==> i Un j < k <-> i<k & j<k"; | 
| 4091 | 562 | by (safe_tac (claset() addSIs [Un_least_lt])); | 
| 437 | 563 | by (rtac (Un_upper2_le RS lt_trans1) 2); | 
| 564 | by (rtac (Un_upper1_le RS lt_trans1) 1); | |
| 435 | 565 | by (REPEAT_SOME assume_tac); | 
| 760 | 566 | qed "Un_least_lt_iff"; | 
| 435 | 567 | |
| 568 | val [ordi,ordj,ordk] = goal Ordinal.thy | |
| 569 | "[| Ord(i); Ord(j); Ord(k) |] ==> i Un j : k <-> i:k & j:k"; | |
| 570 | by (cut_facts_tac [[ordi,ordj] MRS | |
| 1461 | 571 |                    read_instantiate [("k","k")] Un_least_lt_iff] 1);
 | 
| 4091 | 572 | by (asm_full_simp_tac (simpset() addsimps [lt_def,ordi,ordj,ordk]) 1); | 
| 760 | 573 | qed "Un_least_mem_iff"; | 
| 435 | 574 | |
| 575 | (*Replacing k by succ(k') yields the similar rule for le!*) | |
| 5137 | 576 | Goal "[| i<k; j<k |] ==> i Int j < k"; | 
| 435 | 577 | by (res_inst_tac [("i","i"),("j","j")] Ord_linear_le 1);
 | 
| 2033 | 578 | by (stac Int_commute 4); | 
| 4091 | 579 | by (asm_full_simp_tac (simpset() addsimps [le_subset_iff, subset_Int_iff]) 4); | 
| 580 | by (asm_full_simp_tac (simpset() addsimps [le_subset_iff, subset_Int_iff]) 3); | |
| 435 | 581 | by (REPEAT (eresolve_tac [asm_rl, ltE] 1)); | 
| 760 | 582 | qed "Int_greatest_lt"; | 
| 435 | 583 | |
| 584 | (*FIXME: the Intersection duals are missing!*) | |
| 585 | ||
| 586 | ||
| 587 | (*** Results about limits ***) | |
| 588 | ||
| 5321 | 589 | val prems = Goal "[| !!i. i:A ==> Ord(i) |] ==> Ord(Union(A))"; | 
| 435 | 590 | by (rtac (Ord_is_Transset RS Transset_Union_family RS OrdI) 1); | 
| 591 | by (REPEAT (etac UnionE 1 ORELSE ares_tac ([Ord_contains_Transset]@prems) 1)); | |
| 760 | 592 | qed "Ord_Union"; | 
| 435 | 593 | |
| 5321 | 594 | val prems = Goal | 
| 435 | 595 | "[| !!x. x:A ==> Ord(B(x)) |] ==> Ord(UN x:A. B(x))"; | 
| 596 | by (rtac Ord_Union 1); | |
| 597 | by (etac RepFunE 1); | |
| 598 | by (etac ssubst 1); | |
| 599 | by (eresolve_tac prems 1); | |
| 760 | 600 | qed "Ord_UN"; | 
| 435 | 601 | |
| 602 | (* No < version; consider (UN i:nat.i)=nat *) | |
| 5321 | 603 | val [ordi,limit] = Goal | 
| 435 | 604 | "[| Ord(i); !!x. x:A ==> b(x) le i |] ==> (UN x:A. b(x)) le i"; | 
| 605 | by (rtac (le_imp_subset RS UN_least RS subset_imp_le) 1); | |
| 606 | by (REPEAT (ares_tac [ordi, Ord_UN, limit] 1 ORELSE etac (limit RS ltE) 1)); | |
| 760 | 607 | qed "UN_least_le"; | 
| 435 | 608 | |
| 5321 | 609 | val [jlti,limit] = Goal | 
| 435 | 610 | "[| j<i; !!x. x:A ==> b(x)<j |] ==> (UN x:A. succ(b(x))) < i"; | 
| 611 | by (rtac (jlti RS ltE) 1); | |
| 612 | by (rtac (UN_least_le RS lt_trans2) 1); | |
| 613 | by (REPEAT (ares_tac [jlti, succ_leI, limit] 1)); | |
| 760 | 614 | qed "UN_succ_least_lt"; | 
| 435 | 615 | |
| 5321 | 616 | val prems = Goal | 
| 435 | 617 | "[| a: A; i le b(a); !!x. x:A ==> Ord(b(x)) |] ==> i le (UN x:A. b(x))"; | 
| 618 | by (resolve_tac (prems RL [ltE]) 1); | |
| 619 | by (rtac (le_imp_subset RS subset_trans RS subset_imp_le) 1); | |
| 620 | by (REPEAT (ares_tac (prems @ [UN_upper, Ord_UN]) 1)); | |
| 760 | 621 | qed "UN_upper_le"; | 
| 435 | 622 | |
| 5321 | 623 | val [leprem] = Goal | 
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 624 | "[| !!x. x:A ==> c(x) le d(x) |] ==> (UN x:A. c(x)) le (UN x:A. d(x))"; | 
| 1461 | 625 | by (rtac UN_least_le 1); | 
| 626 | by (rtac UN_upper_le 2); | |
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 627 | by (REPEAT (ares_tac [leprem] 2)); | 
| 1461 | 628 | by (rtac Ord_UN 1); | 
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 629 | by (REPEAT (eresolve_tac [asm_rl, leprem RS ltE] 1 | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 630 | ORELSE dtac Ord_succD 1)); | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 631 | qed "le_implies_UN_le_UN"; | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 632 | |
| 5137 | 633 | Goal "Ord(i) ==> (UN y:i. succ(y)) = i"; | 
| 4091 | 634 | by (blast_tac (claset() addIs [Ord_trans]) 1); | 
| 760 | 635 | qed "Ord_equality"; | 
| 435 | 636 | |
| 637 | (*Holds for all transitive sets, not just ordinals*) | |
| 5137 | 638 | Goal "Ord(i) ==> Union(i) <= i"; | 
| 4091 | 639 | by (blast_tac (claset() addIs [Ord_trans]) 1); | 
| 760 | 640 | qed "Ord_Union_subset"; | 
| 435 | 641 | |
| 642 | ||
| 643 | (*** Limit ordinals -- general properties ***) | |
| 644 | ||
| 5137 | 645 | Goalw [Limit_def] "Limit(i) ==> Union(i) = i"; | 
| 4091 | 646 | by (fast_tac (claset() addSIs [ltI] addSEs [ltE] addEs [Ord_trans]) 1); | 
| 760 | 647 | qed "Limit_Union_eq"; | 
| 435 | 648 | |
| 5137 | 649 | Goalw [Limit_def] "Limit(i) ==> Ord(i)"; | 
| 435 | 650 | by (etac conjunct1 1); | 
| 760 | 651 | qed "Limit_is_Ord"; | 
| 435 | 652 | |
| 5137 | 653 | Goalw [Limit_def] "Limit(i) ==> 0 < i"; | 
| 435 | 654 | by (etac (conjunct2 RS conjunct1) 1); | 
| 760 | 655 | qed "Limit_has_0"; | 
| 435 | 656 | |
| 5137 | 657 | Goalw [Limit_def] "[| Limit(i); j<i |] ==> succ(j) < i"; | 
| 2925 | 658 | by (Blast_tac 1); | 
| 760 | 659 | qed "Limit_has_succ"; | 
| 435 | 660 | |
| 5067 | 661 | Goalw [Limit_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 662 | "[| 0<i; ALL y. succ(y) ~= i |] ==> Limit(i)"; | 
| 435 | 663 | by (safe_tac subset_cs); | 
| 664 | by (rtac (not_le_iff_lt RS iffD1) 2); | |
| 2925 | 665 | by (blast_tac le_cs 4); | 
| 435 | 666 | by (REPEAT (eresolve_tac [asm_rl, ltE, Ord_succ] 1)); | 
| 760 | 667 | qed "non_succ_LimitI"; | 
| 435 | 668 | |
| 5137 | 669 | Goal "Limit(succ(i)) ==> P"; | 
| 437 | 670 | by (rtac lt_irrefl 1); | 
| 671 | by (rtac Limit_has_succ 1); | |
| 672 | by (assume_tac 1); | |
| 673 | by (etac (Limit_is_Ord RS Ord_succD RS le_refl) 1); | |
| 760 | 674 | qed "succ_LimitE"; | 
| 435 | 675 | |
| 5137 | 676 | Goal "[| Limit(i); i le succ(j) |] ==> i le j"; | 
| 4091 | 677 | by (safe_tac (claset() addSEs [succ_LimitE, leE])); | 
| 760 | 678 | qed "Limit_le_succD"; | 
| 435 | 679 | |
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 680 | (** Traditional 3-way case analysis on ordinals **) | 
| 435 | 681 | |
| 5137 | 682 | Goal "Ord(i) ==> i=0 | (EX j. Ord(j) & i=succ(j)) | Limit(i)"; | 
| 4091 | 683 | by (blast_tac (claset() addSIs [non_succ_LimitI, Ord_0_lt] | 
| 2493 | 684 | addSDs [Ord_succD]) 1); | 
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 685 | qed "Ord_cases_disj"; | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 686 | |
| 5321 | 687 | val major::prems = Goal | 
| 1461 | 688 | "[| Ord(i); \ | 
| 689 | \ i=0 ==> P; \ | |
| 690 | \ !!j. [| Ord(j); i=succ(j) |] ==> P; \ | |
| 691 | \ Limit(i) ==> P \ | |
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 692 | \ |] ==> P"; | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 693 | by (cut_facts_tac [major RS Ord_cases_disj] 1); | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 694 | by (REPEAT (eresolve_tac (prems@[asm_rl, disjE, exE, conjE]) 1)); | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 695 | qed "Ord_cases"; | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 696 | |
| 5321 | 697 | val major::prems = Goal | 
| 1461 | 698 | "[| Ord(i); \ | 
| 699 | \ P(0); \ | |
| 700 | \ !!x. [| Ord(x); P(x) |] ==> P(succ(x)); \ | |
| 701 | \ !!x. [| Limit(x); ALL y:x. P(y) |] ==> P(x) \ | |
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 702 | \ |] ==> P(i)"; | 
| 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 703 | by (resolve_tac [major RS trans_induct] 1); | 
| 1461 | 704 | by (etac Ord_cases 1); | 
| 4091 | 705 | by (ALLGOALS (blast_tac (claset() addIs prems))); | 
| 851 
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
 lcp parents: 
830diff
changeset | 706 | qed "trans_induct3"; |