| author | wenzelm | 
| Sun, 25 Oct 2009 19:19:41 +0100 | |
| changeset 33170 | dd6d8d1f70d2 | 
| parent 24893 | b8ef7afe3a6b | 
| child 35762 | af3ff2ba4c54 | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: ZF/Ordinal.thy | 
| 435 | 2 | ID: $Id$ | 
| 1478 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 435 | 4 | Copyright 1994 University of Cambridge | 
| 5 | ||
| 6 | *) | |
| 7 | ||
| 13356 | 8 | header{*Transitive Sets and Ordinals*}
 | 
| 9 | ||
| 16417 | 10 | theory Ordinal imports WF Bool equalities begin | 
| 13155 | 11 | |
| 24893 | 12 | definition | 
| 13 | Memrel :: "i=>i" where | |
| 13155 | 14 |     "Memrel(A)   == {z: A*A . EX x y. z=<x,y> & x:y }"
 | 
| 15 | ||
| 24893 | 16 | definition | 
| 17 | Transset :: "i=>o" where | |
| 13155 | 18 | "Transset(i) == ALL x:i. x<=i" | 
| 19 | ||
| 24893 | 20 | definition | 
| 21 | Ord :: "i=>o" where | |
| 13155 | 22 | "Ord(i) == Transset(i) & (ALL x:i. Transset(x))" | 
| 23 | ||
| 24893 | 24 | definition | 
| 25 | lt :: "[i,i] => o" (infixl "<" 50) (*less-than on ordinals*) where | |
| 13155 | 26 | "i<j == i:j & Ord(j)" | 
| 27 | ||
| 24893 | 28 | definition | 
| 29 | Limit :: "i=>o" where | |
| 13155 | 30 | "Limit(i) == Ord(i) & 0<i & (ALL y. y<i --> succ(y)<i)" | 
| 2539 | 31 | |
| 22808 | 32 | abbreviation | 
| 33 | le (infixl "le" 50) where | |
| 34 | "x le y == x < succ(y)" | |
| 435 | 35 | |
| 22808 | 36 | notation (xsymbols) | 
| 37 | le (infixl "\<le>" 50) | |
| 435 | 38 | |
| 22808 | 39 | notation (HTML output) | 
| 40 | le (infixl "\<le>" 50) | |
| 13155 | 41 | |
| 42 | ||
| 13356 | 43 | subsection{*Rules for Transset*}
 | 
| 13155 | 44 | |
| 13356 | 45 | subsubsection{*Three Neat Characterisations of Transset*}
 | 
| 13155 | 46 | |
| 47 | lemma Transset_iff_Pow: "Transset(A) <-> A<=Pow(A)" | |
| 48 | by (unfold Transset_def, blast) | |
| 49 | ||
| 50 | lemma Transset_iff_Union_succ: "Transset(A) <-> Union(succ(A)) = A" | |
| 51 | apply (unfold Transset_def) | |
| 52 | apply (blast elim!: equalityE) | |
| 53 | done | |
| 54 | ||
| 55 | lemma Transset_iff_Union_subset: "Transset(A) <-> Union(A) <= A" | |
| 56 | by (unfold Transset_def, blast) | |
| 57 | ||
| 13356 | 58 | subsubsection{*Consequences of Downwards Closure*}
 | 
| 13155 | 59 | |
| 60 | lemma Transset_doubleton_D: | |
| 61 |     "[| Transset(C); {a,b}: C |] ==> a:C & b: C"
 | |
| 62 | by (unfold Transset_def, blast) | |
| 63 | ||
| 64 | lemma Transset_Pair_D: | |
| 65 | "[| Transset(C); <a,b>: C |] ==> a:C & b: C" | |
| 66 | apply (simp add: Pair_def) | |
| 67 | apply (blast dest: Transset_doubleton_D) | |
| 68 | done | |
| 69 | ||
| 70 | lemma Transset_includes_domain: | |
| 71 | "[| Transset(C); A*B <= C; b: B |] ==> A <= C" | |
| 72 | by (blast dest: Transset_Pair_D) | |
| 73 | ||
| 74 | lemma Transset_includes_range: | |
| 75 | "[| Transset(C); A*B <= C; a: A |] ==> B <= C" | |
| 76 | by (blast dest: Transset_Pair_D) | |
| 77 | ||
| 13356 | 78 | subsubsection{*Closure Properties*}
 | 
| 13155 | 79 | |
| 80 | lemma Transset_0: "Transset(0)" | |
| 81 | by (unfold Transset_def, blast) | |
| 82 | ||
| 83 | lemma Transset_Un: | |
| 84 | "[| Transset(i); Transset(j) |] ==> Transset(i Un j)" | |
| 85 | by (unfold Transset_def, blast) | |
| 86 | ||
| 87 | lemma Transset_Int: | |
| 88 | "[| Transset(i); Transset(j) |] ==> Transset(i Int j)" | |
| 89 | by (unfold Transset_def, blast) | |
| 90 | ||
| 91 | lemma Transset_succ: "Transset(i) ==> Transset(succ(i))" | |
| 92 | by (unfold Transset_def, blast) | |
| 93 | ||
| 94 | lemma Transset_Pow: "Transset(i) ==> Transset(Pow(i))" | |
| 95 | by (unfold Transset_def, blast) | |
| 96 | ||
| 97 | lemma Transset_Union: "Transset(A) ==> Transset(Union(A))" | |
| 98 | by (unfold Transset_def, blast) | |
| 99 | ||
| 100 | lemma Transset_Union_family: | |
| 101 | "[| !!i. i:A ==> Transset(i) |] ==> Transset(Union(A))" | |
| 102 | by (unfold Transset_def, blast) | |
| 103 | ||
| 104 | lemma Transset_Inter_family: | |
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 105 | "[| !!i. i:A ==> Transset(i) |] ==> Transset(Inter(A))" | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 106 | by (unfold Inter_def Transset_def, blast) | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 107 | |
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 108 | lemma Transset_UN: | 
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13544diff
changeset | 109 | "(!!x. x \<in> A ==> Transset(B(x))) ==> Transset (\<Union>x\<in>A. B(x))" | 
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 110 | by (rule Transset_Union_family, auto) | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 111 | |
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 112 | lemma Transset_INT: | 
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13544diff
changeset | 113 | "(!!x. x \<in> A ==> Transset(B(x))) ==> Transset (\<Inter>x\<in>A. B(x))" | 
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 114 | by (rule Transset_Inter_family, auto) | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 115 | |
| 13155 | 116 | |
| 13356 | 117 | subsection{*Lemmas for Ordinals*}
 | 
| 13155 | 118 | |
| 119 | lemma OrdI: | |
| 120 | "[| Transset(i); !!x. x:i ==> Transset(x) |] ==> Ord(i)" | |
| 121 | by (simp add: Ord_def) | |
| 122 | ||
| 123 | lemma Ord_is_Transset: "Ord(i) ==> Transset(i)" | |
| 124 | by (simp add: Ord_def) | |
| 125 | ||
| 126 | lemma Ord_contains_Transset: | |
| 127 | "[| Ord(i); j:i |] ==> Transset(j) " | |
| 128 | by (unfold Ord_def, blast) | |
| 129 | ||
| 130 | ||
| 131 | lemma Ord_in_Ord: "[| Ord(i); j:i |] ==> Ord(j)" | |
| 132 | by (unfold Ord_def Transset_def, blast) | |
| 133 | ||
| 13243 | 134 | (*suitable for rewriting PROVIDED i has been fixed*) | 
| 135 | lemma Ord_in_Ord': "[| j:i; Ord(i) |] ==> Ord(j)" | |
| 136 | by (blast intro: Ord_in_Ord) | |
| 137 | ||
| 13155 | 138 | (* Ord(succ(j)) ==> Ord(j) *) | 
| 139 | lemmas Ord_succD = Ord_in_Ord [OF _ succI1] | |
| 140 | ||
| 141 | lemma Ord_subset_Ord: "[| Ord(i); Transset(j); j<=i |] ==> Ord(j)" | |
| 142 | by (simp add: Ord_def Transset_def, blast) | |
| 143 | ||
| 144 | lemma OrdmemD: "[| j:i; Ord(i) |] ==> j<=i" | |
| 145 | by (unfold Ord_def Transset_def, blast) | |
| 146 | ||
| 147 | lemma Ord_trans: "[| i:j; j:k; Ord(k) |] ==> i:k" | |
| 148 | by (blast dest: OrdmemD) | |
| 149 | ||
| 150 | lemma Ord_succ_subsetI: "[| i:j; Ord(j) |] ==> succ(i) <= j" | |
| 151 | by (blast dest: OrdmemD) | |
| 152 | ||
| 153 | ||
| 13356 | 154 | subsection{*The Construction of Ordinals: 0, succ, Union*}
 | 
| 13155 | 155 | |
| 156 | lemma Ord_0 [iff,TC]: "Ord(0)" | |
| 157 | by (blast intro: OrdI Transset_0) | |
| 158 | ||
| 159 | lemma Ord_succ [TC]: "Ord(i) ==> Ord(succ(i))" | |
| 160 | by (blast intro: OrdI Transset_succ Ord_is_Transset Ord_contains_Transset) | |
| 161 | ||
| 162 | lemmas Ord_1 = Ord_0 [THEN Ord_succ] | |
| 163 | ||
| 164 | lemma Ord_succ_iff [iff]: "Ord(succ(i)) <-> Ord(i)" | |
| 165 | by (blast intro: Ord_succ dest!: Ord_succD) | |
| 166 | ||
| 13172 | 167 | lemma Ord_Un [intro,simp,TC]: "[| Ord(i); Ord(j) |] ==> Ord(i Un j)" | 
| 13155 | 168 | apply (unfold Ord_def) | 
| 169 | apply (blast intro!: Transset_Un) | |
| 170 | done | |
| 171 | ||
| 172 | lemma Ord_Int [TC]: "[| Ord(i); Ord(j) |] ==> Ord(i Int j)" | |
| 173 | apply (unfold Ord_def) | |
| 174 | apply (blast intro!: Transset_Int) | |
| 175 | done | |
| 176 | ||
| 177 | (*There is no set of all ordinals, for then it would contain itself*) | |
| 178 | lemma ON_class: "~ (ALL i. i:X <-> Ord(i))" | |
| 179 | apply (rule notI) | |
| 13784 | 180 | apply (frule_tac x = X in spec) | 
| 13155 | 181 | apply (safe elim!: mem_irrefl) | 
| 182 | apply (erule swap, rule OrdI [OF _ Ord_is_Transset]) | |
| 183 | apply (simp add: Transset_def) | |
| 184 | apply (blast intro: Ord_in_Ord)+ | |
| 185 | done | |
| 186 | ||
| 13356 | 187 | subsection{*< is 'less Than' for Ordinals*}
 | 
| 13155 | 188 | |
| 189 | lemma ltI: "[| i:j; Ord(j) |] ==> i<j" | |
| 190 | by (unfold lt_def, blast) | |
| 191 | ||
| 192 | lemma ltE: | |
| 193 | "[| i<j; [| i:j; Ord(i); Ord(j) |] ==> P |] ==> P" | |
| 194 | apply (unfold lt_def) | |
| 195 | apply (blast intro: Ord_in_Ord) | |
| 196 | done | |
| 197 | ||
| 198 | lemma ltD: "i<j ==> i:j" | |
| 199 | by (erule ltE, assumption) | |
| 200 | ||
| 201 | lemma not_lt0 [simp]: "~ i<0" | |
| 202 | by (unfold lt_def, blast) | |
| 203 | ||
| 204 | lemma lt_Ord: "j<i ==> Ord(j)" | |
| 205 | by (erule ltE, assumption) | |
| 206 | ||
| 207 | lemma lt_Ord2: "j<i ==> Ord(i)" | |
| 208 | by (erule ltE, assumption) | |
| 209 | ||
| 210 | (* "ja le j ==> Ord(j)" *) | |
| 211 | lemmas le_Ord2 = lt_Ord2 [THEN Ord_succD] | |
| 212 | ||
| 213 | (* i<0 ==> R *) | |
| 214 | lemmas lt0E = not_lt0 [THEN notE, elim!] | |
| 215 | ||
| 216 | lemma lt_trans: "[| i<j; j<k |] ==> i<k" | |
| 217 | by (blast intro!: ltI elim!: ltE intro: Ord_trans) | |
| 218 | ||
| 219 | lemma lt_not_sym: "i<j ==> ~ (j<i)" | |
| 220 | apply (unfold lt_def) | |
| 221 | apply (blast elim: mem_asym) | |
| 222 | done | |
| 223 | ||
| 224 | (* [| i<j; ~P ==> j<i |] ==> P *) | |
| 225 | lemmas lt_asym = lt_not_sym [THEN swap] | |
| 226 | ||
| 227 | lemma lt_irrefl [elim!]: "i<i ==> P" | |
| 228 | by (blast intro: lt_asym) | |
| 229 | ||
| 230 | lemma lt_not_refl: "~ i<i" | |
| 231 | apply (rule notI) | |
| 232 | apply (erule lt_irrefl) | |
| 233 | done | |
| 234 | ||
| 235 | ||
| 236 | (** le is less than or equals; recall i le j abbrevs i<succ(j) !! **) | |
| 237 | ||
| 238 | lemma le_iff: "i le j <-> i<j | (i=j & Ord(j))" | |
| 239 | by (unfold lt_def, blast) | |
| 240 | ||
| 241 | (*Equivalently, i<j ==> i < succ(j)*) | |
| 242 | lemma leI: "i<j ==> i le j" | |
| 243 | by (simp (no_asm_simp) add: le_iff) | |
| 244 | ||
| 245 | lemma le_eqI: "[| i=j; Ord(j) |] ==> i le j" | |
| 246 | by (simp (no_asm_simp) add: le_iff) | |
| 247 | ||
| 248 | lemmas le_refl = refl [THEN le_eqI] | |
| 249 | ||
| 250 | lemma le_refl_iff [iff]: "i le i <-> Ord(i)" | |
| 251 | by (simp (no_asm_simp) add: lt_not_refl le_iff) | |
| 252 | ||
| 253 | lemma leCI: "(~ (i=j & Ord(j)) ==> i<j) ==> i le j" | |
| 254 | by (simp add: le_iff, blast) | |
| 255 | ||
| 256 | lemma leE: | |
| 257 | "[| i le j; i<j ==> P; [| i=j; Ord(j) |] ==> P |] ==> P" | |
| 258 | by (simp add: le_iff, blast) | |
| 259 | ||
| 260 | lemma le_anti_sym: "[| i le j; j le i |] ==> i=j" | |
| 261 | apply (simp add: le_iff) | |
| 262 | apply (blast elim: lt_asym) | |
| 263 | done | |
| 264 | ||
| 265 | lemma le0_iff [simp]: "i le 0 <-> i=0" | |
| 266 | by (blast elim!: leE) | |
| 267 | ||
| 268 | lemmas le0D = le0_iff [THEN iffD1, dest!] | |
| 269 | ||
| 13356 | 270 | subsection{*Natural Deduction Rules for Memrel*}
 | 
| 13155 | 271 | |
| 272 | (*The lemmas MemrelI/E give better speed than [iff] here*) | |
| 273 | lemma Memrel_iff [simp]: "<a,b> : Memrel(A) <-> a:b & a:A & b:A" | |
| 274 | by (unfold Memrel_def, blast) | |
| 275 | ||
| 276 | lemma MemrelI [intro!]: "[| a: b; a: A; b: A |] ==> <a,b> : Memrel(A)" | |
| 277 | by auto | |
| 278 | ||
| 279 | lemma MemrelE [elim!]: | |
| 280 | "[| <a,b> : Memrel(A); | |
| 281 | [| a: A; b: A; a:b |] ==> P |] | |
| 282 | ==> P" | |
| 283 | by auto | |
| 284 | ||
| 285 | lemma Memrel_type: "Memrel(A) <= A*A" | |
| 286 | by (unfold Memrel_def, blast) | |
| 287 | ||
| 288 | lemma Memrel_mono: "A<=B ==> Memrel(A) <= Memrel(B)" | |
| 289 | by (unfold Memrel_def, blast) | |
| 290 | ||
| 291 | lemma Memrel_0 [simp]: "Memrel(0) = 0" | |
| 292 | by (unfold Memrel_def, blast) | |
| 293 | ||
| 294 | lemma Memrel_1 [simp]: "Memrel(1) = 0" | |
| 295 | by (unfold Memrel_def, blast) | |
| 296 | ||
| 13269 | 297 | lemma relation_Memrel: "relation(Memrel(A))" | 
| 14864 | 298 | by (simp add: relation_def Memrel_def) | 
| 13269 | 299 | |
| 13155 | 300 | (*The membership relation (as a set) is well-founded. | 
| 301 | Proof idea: show A<=B by applying the foundation axiom to A-B *) | |
| 302 | lemma wf_Memrel: "wf(Memrel(A))" | |
| 303 | apply (unfold wf_def) | |
| 304 | apply (rule foundation [THEN disjE, THEN allI], erule disjI1, blast) | |
| 305 | done | |
| 306 | ||
| 13396 | 307 | text{*The premise @{term "Ord(i)"} does not suffice.*}
 | 
| 13155 | 308 | lemma trans_Memrel: | 
| 309 | "Ord(i) ==> trans(Memrel(i))" | |
| 310 | by (unfold Ord_def Transset_def trans_def, blast) | |
| 311 | ||
| 13396 | 312 | text{*However, the following premise is strong enough.*}
 | 
| 313 | lemma Transset_trans_Memrel: | |
| 314 | "\<forall>j\<in>i. Transset(j) ==> trans(Memrel(i))" | |
| 315 | by (unfold Transset_def trans_def, blast) | |
| 316 | ||
| 13155 | 317 | (*If Transset(A) then Memrel(A) internalizes the membership relation below A*) | 
| 318 | lemma Transset_Memrel_iff: | |
| 319 | "Transset(A) ==> <a,b> : Memrel(A) <-> a:b & b:A" | |
| 320 | by (unfold Transset_def, blast) | |
| 321 | ||
| 322 | ||
| 13356 | 323 | subsection{*Transfinite Induction*}
 | 
| 13155 | 324 | |
| 325 | (*Epsilon induction over a transitive set*) | |
| 326 | lemma Transset_induct: | |
| 327 | "[| i: k; Transset(k); | |
| 328 | !!x.[| x: k; ALL y:x. P(y) |] ==> P(x) |] | |
| 329 | ==> P(i)" | |
| 330 | apply (simp add: Transset_def) | |
| 13269 | 331 | apply (erule wf_Memrel [THEN wf_induct2], blast+) | 
| 13155 | 332 | done | 
| 333 | ||
| 334 | (*Induction over an ordinal*) | |
| 13534 | 335 | lemmas Ord_induct [consumes 2] = Transset_induct [OF _ Ord_is_Transset] | 
| 336 | lemmas Ord_induct_rule = Ord_induct [rule_format, consumes 2] | |
| 13155 | 337 | |
| 338 | (*Induction over the class of ordinals -- a useful corollary of Ord_induct*) | |
| 339 | ||
| 13534 | 340 | lemma trans_induct [consumes 1]: | 
| 13155 | 341 | "[| Ord(i); | 
| 342 | !!x.[| Ord(x); ALL y:x. P(y) |] ==> P(x) |] | |
| 343 | ==> P(i)" | |
| 344 | apply (rule Ord_succ [THEN succI1 [THEN Ord_induct]], assumption) | |
| 345 | apply (blast intro: Ord_succ [THEN Ord_in_Ord]) | |
| 346 | done | |
| 347 | ||
| 13534 | 348 | lemmas trans_induct_rule = trans_induct [rule_format, consumes 1] | 
| 349 | ||
| 13155 | 350 | |
| 351 | (*** Fundamental properties of the epsilon ordering (< on ordinals) ***) | |
| 352 | ||
| 353 | ||
| 13356 | 354 | subsubsection{*Proving That < is a Linear Ordering on the Ordinals*}
 | 
| 13155 | 355 | |
| 356 | lemma Ord_linear [rule_format]: | |
| 357 | "Ord(i) ==> (ALL j. Ord(j) --> i:j | i=j | j:i)" | |
| 358 | apply (erule trans_induct) | |
| 359 | apply (rule impI [THEN allI]) | |
| 360 | apply (erule_tac i=j in trans_induct) | |
| 361 | apply (blast dest: Ord_trans) | |
| 362 | done | |
| 363 | ||
| 364 | (*The trichotomy law for ordinals!*) | |
| 365 | lemma Ord_linear_lt: | |
| 366 | "[| Ord(i); Ord(j); i<j ==> P; i=j ==> P; j<i ==> P |] ==> P" | |
| 367 | apply (simp add: lt_def) | |
| 368 | apply (rule_tac i1=i and j1=j in Ord_linear [THEN disjE], blast+) | |
| 369 | done | |
| 370 | ||
| 371 | lemma Ord_linear2: | |
| 372 | "[| Ord(i); Ord(j); i<j ==> P; j le i ==> P |] ==> P" | |
| 13784 | 373 | apply (rule_tac i = i and j = j in Ord_linear_lt) | 
| 13155 | 374 | apply (blast intro: leI le_eqI sym ) + | 
| 375 | done | |
| 376 | ||
| 377 | lemma Ord_linear_le: | |
| 378 | "[| Ord(i); Ord(j); i le j ==> P; j le i ==> P |] ==> P" | |
| 13784 | 379 | apply (rule_tac i = i and j = j in Ord_linear_lt) | 
| 13155 | 380 | apply (blast intro: leI le_eqI ) + | 
| 381 | done | |
| 382 | ||
| 383 | lemma le_imp_not_lt: "j le i ==> ~ i<j" | |
| 384 | by (blast elim!: leE elim: lt_asym) | |
| 385 | ||
| 386 | lemma not_lt_imp_le: "[| ~ i<j; Ord(i); Ord(j) |] ==> j le i" | |
| 13784 | 387 | by (rule_tac i = i and j = j in Ord_linear2, auto) | 
| 13155 | 388 | |
| 13356 | 389 | subsubsection{*Some Rewrite Rules for <, le*}
 | 
| 13155 | 390 | |
| 391 | lemma Ord_mem_iff_lt: "Ord(j) ==> i:j <-> i<j" | |
| 392 | by (unfold lt_def, blast) | |
| 393 | ||
| 394 | lemma not_lt_iff_le: "[| Ord(i); Ord(j) |] ==> ~ i<j <-> j le i" | |
| 395 | by (blast dest: le_imp_not_lt not_lt_imp_le) | |
| 2540 | 396 | |
| 13155 | 397 | lemma not_le_iff_lt: "[| Ord(i); Ord(j) |] ==> ~ i le j <-> j<i" | 
| 398 | by (simp (no_asm_simp) add: not_lt_iff_le [THEN iff_sym]) | |
| 399 | ||
| 400 | (*This is identical to 0<succ(i) *) | |
| 401 | lemma Ord_0_le: "Ord(i) ==> 0 le i" | |
| 402 | by (erule not_lt_iff_le [THEN iffD1], auto) | |
| 403 | ||
| 404 | lemma Ord_0_lt: "[| Ord(i); i~=0 |] ==> 0<i" | |
| 405 | apply (erule not_le_iff_lt [THEN iffD1]) | |
| 406 | apply (rule Ord_0, blast) | |
| 407 | done | |
| 408 | ||
| 409 | lemma Ord_0_lt_iff: "Ord(i) ==> i~=0 <-> 0<i" | |
| 410 | by (blast intro: Ord_0_lt) | |
| 411 | ||
| 412 | ||
| 13356 | 413 | subsection{*Results about Less-Than or Equals*}
 | 
| 13155 | 414 | |
| 415 | (** For ordinals, j<=i (subset) implies j le i (less-than or equals) **) | |
| 416 | ||
| 417 | lemma zero_le_succ_iff [iff]: "0 le succ(x) <-> Ord(x)" | |
| 418 | by (blast intro: Ord_0_le elim: ltE) | |
| 419 | ||
| 420 | lemma subset_imp_le: "[| j<=i; Ord(i); Ord(j) |] ==> j le i" | |
| 13269 | 421 | apply (rule not_lt_iff_le [THEN iffD1], assumption+) | 
| 13155 | 422 | apply (blast elim: ltE mem_irrefl) | 
| 423 | done | |
| 424 | ||
| 425 | lemma le_imp_subset: "i le j ==> i<=j" | |
| 426 | by (blast dest: OrdmemD elim: ltE leE) | |
| 427 | ||
| 428 | lemma le_subset_iff: "j le i <-> j<=i & Ord(i) & Ord(j)" | |
| 429 | by (blast dest: subset_imp_le le_imp_subset elim: ltE) | |
| 430 | ||
| 431 | lemma le_succ_iff: "i le succ(j) <-> i le j | i=succ(j) & Ord(i)" | |
| 432 | apply (simp (no_asm) add: le_iff) | |
| 433 | apply blast | |
| 434 | done | |
| 435 | ||
| 436 | (*Just a variant of subset_imp_le*) | |
| 437 | lemma all_lt_imp_le: "[| Ord(i); Ord(j); !!x. x<j ==> x<i |] ==> j le i" | |
| 438 | by (blast intro: not_lt_imp_le dest: lt_irrefl) | |
| 439 | ||
| 13356 | 440 | subsubsection{*Transitivity Laws*}
 | 
| 13155 | 441 | |
| 442 | lemma lt_trans1: "[| i le j; j<k |] ==> i<k" | |
| 443 | by (blast elim!: leE intro: lt_trans) | |
| 444 | ||
| 445 | lemma lt_trans2: "[| i<j; j le k |] ==> i<k" | |
| 446 | by (blast elim!: leE intro: lt_trans) | |
| 447 | ||
| 448 | lemma le_trans: "[| i le j; j le k |] ==> i le k" | |
| 449 | by (blast intro: lt_trans1) | |
| 450 | ||
| 451 | lemma succ_leI: "i<j ==> succ(i) le j" | |
| 452 | apply (rule not_lt_iff_le [THEN iffD1]) | |
| 453 | apply (blast elim: ltE leE lt_asym)+ | |
| 454 | done | |
| 455 | ||
| 456 | (*Identical to succ(i) < succ(j) ==> i<j *) | |
| 457 | lemma succ_leE: "succ(i) le j ==> i<j" | |
| 458 | apply (rule not_le_iff_lt [THEN iffD1]) | |
| 459 | apply (blast elim: ltE leE lt_asym)+ | |
| 460 | done | |
| 461 | ||
| 462 | lemma succ_le_iff [iff]: "succ(i) le j <-> i<j" | |
| 463 | by (blast intro: succ_leI succ_leE) | |
| 464 | ||
| 465 | lemma succ_le_imp_le: "succ(i) le succ(j) ==> i le j" | |
| 466 | by (blast dest!: succ_leE) | |
| 467 | ||
| 468 | lemma lt_subset_trans: "[| i <= j; j<k; Ord(i) |] ==> i<k" | |
| 469 | apply (rule subset_imp_le [THEN lt_trans1]) | |
| 470 | apply (blast intro: elim: ltE) + | |
| 471 | done | |
| 472 | ||
| 13172 | 473 | lemma lt_imp_0_lt: "j<i ==> 0<i" | 
| 474 | by (blast intro: lt_trans1 Ord_0_le [OF lt_Ord]) | |
| 475 | ||
| 13243 | 476 | lemma succ_lt_iff: "succ(i) < j <-> i<j & succ(i) \<noteq> j" | 
| 13162 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 477 | apply auto | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 478 | apply (blast intro: lt_trans le_refl dest: lt_Ord) | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 479 | apply (frule lt_Ord) | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 480 | apply (rule not_le_iff_lt [THEN iffD1]) | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 481 | apply (blast intro: lt_Ord2) | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 482 | apply blast | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 483 | apply (simp add: lt_Ord lt_Ord2 le_iff) | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 484 | apply (blast dest: lt_asym) | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 485 | done | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 486 | |
| 13243 | 487 | lemma Ord_succ_mem_iff: "Ord(j) ==> succ(i) \<in> succ(j) <-> i\<in>j" | 
| 488 | apply (insert succ_le_iff [of i j]) | |
| 489 | apply (simp add: lt_def) | |
| 490 | done | |
| 491 | ||
| 13356 | 492 | subsubsection{*Union and Intersection*}
 | 
| 13155 | 493 | |
| 494 | lemma Un_upper1_le: "[| Ord(i); Ord(j) |] ==> i le i Un j" | |
| 495 | by (rule Un_upper1 [THEN subset_imp_le], auto) | |
| 496 | ||
| 497 | lemma Un_upper2_le: "[| Ord(i); Ord(j) |] ==> j le i Un j" | |
| 498 | by (rule Un_upper2 [THEN subset_imp_le], auto) | |
| 499 | ||
| 500 | (*Replacing k by succ(k') yields the similar rule for le!*) | |
| 501 | lemma Un_least_lt: "[| i<k; j<k |] ==> i Un j < k" | |
| 13784 | 502 | apply (rule_tac i = i and j = j in Ord_linear_le) | 
| 13155 | 503 | apply (auto simp add: Un_commute le_subset_iff subset_Un_iff lt_Ord) | 
| 504 | done | |
| 505 | ||
| 506 | lemma Un_least_lt_iff: "[| Ord(i); Ord(j) |] ==> i Un j < k <-> i<k & j<k" | |
| 507 | apply (safe intro!: Un_least_lt) | |
| 508 | apply (rule_tac [2] Un_upper2_le [THEN lt_trans1]) | |
| 509 | apply (rule Un_upper1_le [THEN lt_trans1], auto) | |
| 510 | done | |
| 511 | ||
| 512 | lemma Un_least_mem_iff: | |
| 513 | "[| Ord(i); Ord(j); Ord(k) |] ==> i Un j : k <-> i:k & j:k" | |
| 514 | apply (insert Un_least_lt_iff [of i j k]) | |
| 515 | apply (simp add: lt_def) | |
| 516 | done | |
| 517 | ||
| 518 | (*Replacing k by succ(k') yields the similar rule for le!*) | |
| 519 | lemma Int_greatest_lt: "[| i<k; j<k |] ==> i Int j < k" | |
| 13784 | 520 | apply (rule_tac i = i and j = j in Ord_linear_le) | 
| 13155 | 521 | apply (auto simp add: Int_commute le_subset_iff subset_Int_iff lt_Ord) | 
| 522 | done | |
| 523 | ||
| 13162 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 524 | lemma Ord_Un_if: | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 525 | "[| Ord(i); Ord(j) |] ==> i \<union> j = (if j<i then i else j)" | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 526 | by (simp add: not_lt_iff_le le_imp_subset leI | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 527 | subset_Un_iff [symmetric] subset_Un_iff2 [symmetric]) | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 528 | |
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 529 | lemma succ_Un_distrib: | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 530 | "[| Ord(i); Ord(j) |] ==> succ(i \<union> j) = succ(i) \<union> succ(j)" | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 531 | by (simp add: Ord_Un_if lt_Ord le_Ord2) | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 532 | |
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 533 | lemma lt_Un_iff: | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 534 | "[| Ord(i); Ord(j) |] ==> k < i \<union> j <-> k < i | k < j"; | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 535 | apply (simp add: Ord_Un_if not_lt_iff_le) | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 536 | apply (blast intro: leI lt_trans2)+ | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 537 | done | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 538 | |
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 539 | lemma le_Un_iff: | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 540 | "[| Ord(i); Ord(j) |] ==> k \<le> i \<union> j <-> k \<le> i | k \<le> j"; | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 541 | by (simp add: succ_Un_distrib lt_Un_iff [symmetric]) | 
| 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 542 | |
| 13172 | 543 | lemma Un_upper1_lt: "[|k < i; Ord(j)|] ==> k < i Un j" | 
| 544 | by (simp add: lt_Un_iff lt_Ord2) | |
| 545 | ||
| 546 | lemma Un_upper2_lt: "[|k < j; Ord(i)|] ==> k < i Un j" | |
| 547 | by (simp add: lt_Un_iff lt_Ord2) | |
| 548 | ||
| 549 | (*See also Transset_iff_Union_succ*) | |
| 550 | lemma Ord_Union_succ_eq: "Ord(i) ==> \<Union>(succ(i)) = i" | |
| 551 | by (blast intro: Ord_trans) | |
| 552 | ||
| 13162 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 paulson parents: 
13155diff
changeset | 553 | |
| 13356 | 554 | subsection{*Results about Limits*}
 | 
| 13155 | 555 | |
| 13172 | 556 | lemma Ord_Union [intro,simp,TC]: "[| !!i. i:A ==> Ord(i) |] ==> Ord(Union(A))" | 
| 13155 | 557 | apply (rule Ord_is_Transset [THEN Transset_Union_family, THEN OrdI]) | 
| 558 | apply (blast intro: Ord_contains_Transset)+ | |
| 559 | done | |
| 560 | ||
| 13172 | 561 | lemma Ord_UN [intro,simp,TC]: | 
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13544diff
changeset | 562 | "[| !!x. x:A ==> Ord(B(x)) |] ==> Ord(\<Union>x\<in>A. B(x))" | 
| 13155 | 563 | by (rule Ord_Union, blast) | 
| 564 | ||
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 565 | lemma Ord_Inter [intro,simp,TC]: | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 566 | "[| !!i. i:A ==> Ord(i) |] ==> Ord(Inter(A))" | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 567 | apply (rule Transset_Inter_family [THEN OrdI]) | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 568 | apply (blast intro: Ord_is_Transset) | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 569 | apply (simp add: Inter_def) | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 570 | apply (blast intro: Ord_contains_Transset) | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 571 | done | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 572 | |
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 573 | lemma Ord_INT [intro,simp,TC]: | 
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13544diff
changeset | 574 | "[| !!x. x:A ==> Ord(B(x)) |] ==> Ord(\<Inter>x\<in>A. B(x))" | 
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 575 | by (rule Ord_Inter, blast) | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 576 | |
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13172diff
changeset | 577 | |
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13544diff
changeset | 578 | (* No < version; consider (\<Union>i\<in>nat.i)=nat *) | 
| 13155 | 579 | lemma UN_least_le: | 
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13544diff
changeset | 580 | "[| Ord(i); !!x. x:A ==> b(x) le i |] ==> (\<Union>x\<in>A. b(x)) le i" | 
| 13155 | 581 | apply (rule le_imp_subset [THEN UN_least, THEN subset_imp_le]) | 
| 582 | apply (blast intro: Ord_UN elim: ltE)+ | |
| 583 | done | |
| 584 | ||
| 585 | lemma UN_succ_least_lt: | |
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13544diff
changeset | 586 | "[| j<i; !!x. x:A ==> b(x)<j |] ==> (\<Union>x\<in>A. succ(b(x))) < i" | 
| 13155 | 587 | apply (rule ltE, assumption) | 
| 588 | apply (rule UN_least_le [THEN lt_trans2]) | |
| 589 | apply (blast intro: succ_leI)+ | |
| 590 | done | |
| 591 | ||
| 13172 | 592 | lemma UN_upper_lt: | 
| 593 | "[| a\<in>A; i < b(a); Ord(\<Union>x\<in>A. b(x)) |] ==> i < (\<Union>x\<in>A. b(x))" | |
| 594 | by (unfold lt_def, blast) | |
| 595 | ||
| 13155 | 596 | lemma UN_upper_le: | 
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13544diff
changeset | 597 | "[| a: A; i le b(a); Ord(\<Union>x\<in>A. b(x)) |] ==> i le (\<Union>x\<in>A. b(x))" | 
| 13155 | 598 | apply (frule ltD) | 
| 599 | apply (rule le_imp_subset [THEN subset_trans, THEN subset_imp_le]) | |
| 600 | apply (blast intro: lt_Ord UN_upper)+ | |
| 601 | done | |
| 602 | ||
| 13172 | 603 | lemma lt_Union_iff: "\<forall>i\<in>A. Ord(i) ==> (j < \<Union>(A)) <-> (\<exists>i\<in>A. j<i)" | 
| 604 | by (auto simp: lt_def Ord_Union) | |
| 605 | ||
| 606 | lemma Union_upper_le: | |
| 607 | "[| j: J; i\<le>j; Ord(\<Union>(J)) |] ==> i \<le> \<Union>J" | |
| 608 | apply (subst Union_eq_UN) | |
| 609 | apply (rule UN_upper_le, auto) | |
| 610 | done | |
| 611 | ||
| 13155 | 612 | lemma le_implies_UN_le_UN: | 
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13544diff
changeset | 613 | "[| !!x. x:A ==> c(x) le d(x) |] ==> (\<Union>x\<in>A. c(x)) le (\<Union>x\<in>A. d(x))" | 
| 13155 | 614 | apply (rule UN_least_le) | 
| 615 | apply (rule_tac [2] UN_upper_le) | |
| 616 | apply (blast intro: Ord_UN le_Ord2)+ | |
| 617 | done | |
| 618 | ||
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13544diff
changeset | 619 | lemma Ord_equality: "Ord(i) ==> (\<Union>y\<in>i. succ(y)) = i" | 
| 13155 | 620 | by (blast intro: Ord_trans) | 
| 621 | ||
| 622 | (*Holds for all transitive sets, not just ordinals*) | |
| 623 | lemma Ord_Union_subset: "Ord(i) ==> Union(i) <= i" | |
| 624 | by (blast intro: Ord_trans) | |
| 625 | ||
| 626 | ||
| 13356 | 627 | subsection{*Limit Ordinals -- General Properties*}
 | 
| 13155 | 628 | |
| 629 | lemma Limit_Union_eq: "Limit(i) ==> Union(i) = i" | |
| 630 | apply (unfold Limit_def) | |
| 631 | apply (fast intro!: ltI elim!: ltE elim: Ord_trans) | |
| 632 | done | |
| 633 | ||
| 634 | lemma Limit_is_Ord: "Limit(i) ==> Ord(i)" | |
| 635 | apply (unfold Limit_def) | |
| 636 | apply (erule conjunct1) | |
| 637 | done | |
| 638 | ||
| 639 | lemma Limit_has_0: "Limit(i) ==> 0 < i" | |
| 640 | apply (unfold Limit_def) | |
| 641 | apply (erule conjunct2 [THEN conjunct1]) | |
| 642 | done | |
| 643 | ||
| 13544 | 644 | lemma Limit_nonzero: "Limit(i) ==> i ~= 0" | 
| 645 | by (drule Limit_has_0, blast) | |
| 646 | ||
| 13155 | 647 | lemma Limit_has_succ: "[| Limit(i); j<i |] ==> succ(j) < i" | 
| 648 | by (unfold Limit_def, blast) | |
| 649 | ||
| 13544 | 650 | lemma Limit_succ_lt_iff [simp]: "Limit(i) ==> succ(j) < i <-> (j<i)" | 
| 651 | apply (safe intro!: Limit_has_succ) | |
| 652 | apply (frule lt_Ord) | |
| 653 | apply (blast intro: lt_trans) | |
| 654 | done | |
| 655 | ||
| 13172 | 656 | lemma zero_not_Limit [iff]: "~ Limit(0)" | 
| 657 | by (simp add: Limit_def) | |
| 658 | ||
| 659 | lemma Limit_has_1: "Limit(i) ==> 1 < i" | |
| 660 | by (blast intro: Limit_has_0 Limit_has_succ) | |
| 661 | ||
| 662 | lemma increasing_LimitI: "[| 0<l; \<forall>x\<in>l. \<exists>y\<in>l. x<y |] ==> Limit(l)" | |
| 13544 | 663 | apply (unfold Limit_def, simp add: lt_Ord2, clarify) | 
| 13172 | 664 | apply (drule_tac i=y in ltD) | 
| 665 | apply (blast intro: lt_trans1 [OF _ ltI] lt_Ord2) | |
| 666 | done | |
| 667 | ||
| 13155 | 668 | lemma non_succ_LimitI: | 
| 669 | "[| 0<i; ALL y. succ(y) ~= i |] ==> Limit(i)" | |
| 670 | apply (unfold Limit_def) | |
| 671 | apply (safe del: subsetI) | |
| 672 | apply (rule_tac [2] not_le_iff_lt [THEN iffD1]) | |
| 673 | apply (simp_all add: lt_Ord lt_Ord2) | |
| 674 | apply (blast elim: leE lt_asym) | |
| 675 | done | |
| 676 | ||
| 677 | lemma succ_LimitE [elim!]: "Limit(succ(i)) ==> P" | |
| 678 | apply (rule lt_irrefl) | |
| 679 | apply (rule Limit_has_succ, assumption) | |
| 680 | apply (erule Limit_is_Ord [THEN Ord_succD, THEN le_refl]) | |
| 681 | done | |
| 682 | ||
| 683 | lemma not_succ_Limit [simp]: "~ Limit(succ(i))" | |
| 684 | by blast | |
| 685 | ||
| 686 | lemma Limit_le_succD: "[| Limit(i); i le succ(j) |] ==> i le j" | |
| 687 | by (blast elim!: leE) | |
| 688 | ||
| 13172 | 689 | |
| 13356 | 690 | subsubsection{*Traditional 3-Way Case Analysis on Ordinals*}
 | 
| 13155 | 691 | |
| 692 | lemma Ord_cases_disj: "Ord(i) ==> i=0 | (EX j. Ord(j) & i=succ(j)) | Limit(i)" | |
| 693 | by (blast intro!: non_succ_LimitI Ord_0_lt) | |
| 694 | ||
| 695 | lemma Ord_cases: | |
| 696 | "[| Ord(i); | |
| 697 | i=0 ==> P; | |
| 698 | !!j. [| Ord(j); i=succ(j) |] ==> P; | |
| 699 | Limit(i) ==> P | |
| 700 | |] ==> P" | |
| 701 | by (drule Ord_cases_disj, blast) | |
| 702 | ||
| 13534 | 703 | lemma trans_induct3 [case_names 0 succ limit, consumes 1]: | 
| 13155 | 704 | "[| Ord(i); | 
| 705 | P(0); | |
| 706 | !!x. [| Ord(x); P(x) |] ==> P(succ(x)); | |
| 707 | !!x. [| Limit(x); ALL y:x. P(y) |] ==> P(x) | |
| 708 | |] ==> P(i)" | |
| 709 | apply (erule trans_induct) | |
| 710 | apply (erule Ord_cases, blast+) | |
| 711 | done | |
| 712 | ||
| 13534 | 713 | lemmas trans_induct3_rule = trans_induct3 [rule_format, case_names 0 succ limit, consumes 1] | 
| 714 | ||
| 13172 | 715 | text{*A set of ordinals is either empty, contains its own union, or its
 | 
| 716 | union is a limit ordinal.*} | |
| 717 | lemma Ord_set_cases: | |
| 718 | "\<forall>i\<in>I. Ord(i) ==> I=0 \<or> \<Union>(I) \<in> I \<or> (\<Union>(I) \<notin> I \<and> Limit(\<Union>(I)))" | |
| 719 | apply (clarify elim!: not_emptyE) | |
| 720 | apply (cases "\<Union>(I)" rule: Ord_cases) | |
| 721 | apply (blast intro: Ord_Union) | |
| 722 | apply (blast intro: subst_elem) | |
| 723 | apply auto | |
| 724 | apply (clarify elim!: equalityE succ_subsetE) | |
| 725 | apply (simp add: Union_subset_iff) | |
| 726 | apply (subgoal_tac "B = succ(j)", blast) | |
| 727 | apply (rule le_anti_sym) | |
| 728 | apply (simp add: le_subset_iff) | |
| 729 | apply (simp add: ltI) | |
| 730 | done | |
| 731 | ||
| 732 | text{*If the union of a set of ordinals is a successor, then it is
 | |
| 733 | an element of that set.*} | |
| 734 | lemma Ord_Union_eq_succD: "[|\<forall>x\<in>X. Ord(x); \<Union>X = succ(j)|] ==> succ(j) \<in> X" | |
| 735 | by (drule Ord_set_cases, auto) | |
| 736 | ||
| 737 | lemma Limit_Union [rule_format]: "[| I \<noteq> 0; \<forall>i\<in>I. Limit(i) |] ==> Limit(\<Union>I)" | |
| 738 | apply (simp add: Limit_def lt_def) | |
| 739 | apply (blast intro!: equalityI) | |
| 740 | done | |
| 741 | ||
| 435 | 742 | end |