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