src/ZF/Zorn.thy
author paulson
Sun, 14 Jul 2002 15:14:43 +0200
changeset 13356 c9cfe1638bf2
parent 13269 3ba9be497c33
child 13558 18acbbd4d225
permissions -rw-r--r--
improved presentation markup
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/Zorn.thy
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
     4
    Copyright   1994  University of Cambridge
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
     5
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
     6
*)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
     7
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
     8
header{*Zorn's Lemma*}
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
     9
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    10
theory Zorn = OrderArith + AC + Inductive:
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
    11
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    12
text{*Based upon the unpublished article ``Towards the Mechanization of the
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    13
Proofs of Some Classical Theorems of Set Theory,'' by Abrial and Laffitte.*}
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    14
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    15
constdefs
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    16
  Subset_rel :: "i=>i"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    17
   "Subset_rel(A) == {z: A*A . EX x y. z=<x,y> & x<=y & x~=y}"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    18
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    19
  chain      :: "i=>i"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    20
   "chain(A)      == {F: Pow(A). ALL X:F. ALL Y:F. X<=Y | Y<=X}"
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
    21
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    22
  maxchain   :: "i=>i"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    23
   "maxchain(A)   == {c: chain(A). super(A,c)=0}"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    24
  
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    25
  super      :: "[i,i]=>i"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    26
   "super(A,c)    == {d: chain(A). c<=d & c~=d}"
485
5e00a676a211 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    27
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
    28
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    29
constdefs
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    30
  increasing :: "i=>i"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    31
    "increasing(A) == {f: Pow(A)->Pow(A). ALL x. x<=A --> x<=f`x}"
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
    32
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    33
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    34
(*Lemma for the inductive definition below*)
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    35
lemma Union_in_Pow: "Y : Pow(Pow(A)) ==> Union(Y) : Pow(A)"
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    36
by blast
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    37
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    38
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    39
text{*We could make the inductive definition conditional on 
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    40
    @{term "next \<in> increasing(S)"}
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
    41
    but instead we make this a side-condition of an introduction rule.  Thus
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
    42
    the induction rule lets us assume that condition!  Many inductive proofs
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    43
    are therefore unconditional.*}
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
    44
consts
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    45
  "TFin" :: "[i,i]=>i"
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
    46
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
    47
inductive
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
    48
  domains       "TFin(S,next)" <= "Pow(S)"
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    49
  intros
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    50
    nextI:       "[| x : TFin(S,next);  next: increasing(S) |]
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    51
                  ==> next`x : TFin(S,next)"
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
    52
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    53
    Pow_UnionI: "Y : Pow(TFin(S,next)) ==> Union(Y) : TFin(S,next)"
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
    54
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1478
diff changeset
    55
  monos         Pow_mono
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1478
diff changeset
    56
  con_defs      increasing_def
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    57
  type_intros   CollectD1 [THEN apply_funtype] Union_in_Pow
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    58
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    59
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    60
subsection{*Mathematical Preamble *}
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    61
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    62
lemma Union_lemma0: "(ALL x:C. x<=A | B<=x) ==> Union(C)<=A | B<=Union(C)"
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
    63
by blast
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    64
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    65
lemma Inter_lemma0:
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    66
     "[| c:C; ALL x:C. A<=x | x<=B |] ==> A<=Inter(C) | Inter(C)<=B"
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
    67
by blast
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    68
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    69
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    70
subsection{*The Transfinite Construction *}
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    71
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    72
lemma increasingD1: "f: increasing(A) ==> f: Pow(A)->Pow(A)"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    73
apply (unfold increasing_def)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    74
apply (erule CollectD1)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    75
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    76
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    77
lemma increasingD2: "[| f: increasing(A); x<=A |] ==> x <= f`x"
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
    78
by (unfold increasing_def, blast)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    79
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    80
lemmas TFin_UnionI = PowI [THEN TFin.Pow_UnionI, standard]
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    81
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    82
lemmas TFin_is_subset = TFin.dom_subset [THEN subsetD, THEN PowD, standard]
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    83
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    84
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    85
(** Structural induction on TFin(S,next) **)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    86
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    87
lemma TFin_induct:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    88
  "[| n: TFin(S,next);   
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    89
      !!x. [| x : TFin(S,next);  P(x);  next: increasing(S) |] ==> P(next`x);  
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    90
      !!Y. [| Y <= TFin(S,next);  ALL y:Y. P(y) |] ==> P(Union(Y))  
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    91
   |] ==> P(n)"
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    92
by (erule TFin.induct, blast+)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    93
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    94
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
    95
subsection{*Some Properties of the Transfinite Construction *}
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    96
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    97
lemmas increasing_trans = subset_trans [OF _ increasingD2, 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    98
                                        OF _ _ TFin_is_subset]
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    99
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   100
(*Lemma 1 of section 3.1*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   101
lemma TFin_linear_lemma1:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   102
     "[| n: TFin(S,next);  m: TFin(S,next);   
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   103
         ALL x: TFin(S,next) . x<=m --> x=m | next`x<=m |] 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   104
      ==> n<=m | next`m<=n"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   105
apply (erule TFin_induct)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   106
apply (erule_tac [2] Union_lemma0) (*or just Blast_tac*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   107
(*downgrade subsetI from intro! to intro*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   108
apply (blast dest: increasing_trans)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   109
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   110
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   111
(*Lemma 2 of section 3.2.  Interesting in its own right!
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   112
  Requires next: increasing(S) in the second induction step. *)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   113
lemma TFin_linear_lemma2:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   114
    "[| m: TFin(S,next);  next: increasing(S) |] 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   115
     ==> ALL n: TFin(S,next) . n<=m --> n=m | next`n<=m"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   116
apply (erule TFin_induct)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   117
apply (rule impI [THEN ballI])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   118
(*case split using TFin_linear_lemma1*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   119
apply (rule_tac n1 = "n" and m1 = "x" in TFin_linear_lemma1 [THEN disjE],
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   120
       assumption+)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   121
apply (blast del: subsetI
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   122
	     intro: increasing_trans subsetI, blast) 
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   123
(*second induction step*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   124
apply (rule impI [THEN ballI])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   125
apply (rule Union_lemma0 [THEN disjE])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   126
apply (erule_tac [3] disjI2)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   127
prefer 2 apply blast 
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   128
apply (rule ballI)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   129
apply (drule bspec, assumption) 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   130
apply (drule subsetD, assumption) 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   131
apply (rule_tac n1 = "n" and m1 = "x" in TFin_linear_lemma1 [THEN disjE],
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   132
       assumption+, blast) 
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   133
apply (erule increasingD2 [THEN subset_trans, THEN disjI1])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   134
apply (blast dest: TFin_is_subset)+
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   135
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   136
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   137
(*a more convenient form for Lemma 2*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   138
lemma TFin_subsetD:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   139
     "[| n<=m;  m: TFin(S,next);  n: TFin(S,next);  next: increasing(S) |]  
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   140
      ==> n=m | next`n<=m"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   141
by (blast dest: TFin_linear_lemma2 [rule_format]) 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   142
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   143
(*Consequences from section 3.3 -- Property 3.2, the ordering is total*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   144
lemma TFin_subset_linear:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   145
     "[| m: TFin(S,next);  n: TFin(S,next);  next: increasing(S) |]  
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   146
      ==> n<=m | m<=n"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   147
apply (rule disjE) 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   148
apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   149
apply (assumption+, erule disjI2)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   150
apply (blast del: subsetI 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   151
             intro: subsetI increasingD2 [THEN subset_trans] TFin_is_subset)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   152
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   153
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   154
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   155
(*Lemma 3 of section 3.3*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   156
lemma equal_next_upper:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   157
     "[| n: TFin(S,next);  m: TFin(S,next);  m = next`m |] ==> n<=m"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   158
apply (erule TFin_induct)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   159
apply (drule TFin_subsetD)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   160
apply (assumption+, force) 
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   161
apply blast
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   162
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   163
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   164
(*Property 3.3 of section 3.3*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   165
lemma equal_next_Union: "[| m: TFin(S,next);  next: increasing(S) |]   
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   166
      ==> m = next`m <-> m = Union(TFin(S,next))"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   167
apply (rule iffI)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   168
apply (rule Union_upper [THEN equalityI])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   169
apply (rule_tac [2] equal_next_upper [THEN Union_least])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   170
apply (assumption+)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   171
apply (erule ssubst)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   172
apply (rule increasingD2 [THEN equalityI], assumption)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   173
apply (blast del: subsetI
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   174
	     intro: subsetI TFin_UnionI TFin.nextI TFin_is_subset)+
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   175
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   176
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   177
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
   178
subsection{*Hausdorff's Theorem: Every Set Contains a Maximal Chain*}
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
   179
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
   180
text{*NOTE: We assume the partial ordering is @{text "\<subseteq>"}, the subset
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
   181
relation!*}
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   182
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   183
(** Defining the "next" operation for Hausdorff's Theorem **)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   184
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   185
lemma chain_subset_Pow: "chain(A) <= Pow(A)"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   186
apply (unfold chain_def)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   187
apply (rule Collect_subset)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   188
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   189
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   190
lemma super_subset_chain: "super(A,c) <= chain(A)"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   191
apply (unfold super_def)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   192
apply (rule Collect_subset)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   193
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   194
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   195
lemma maxchain_subset_chain: "maxchain(A) <= chain(A)"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   196
apply (unfold maxchain_def)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   197
apply (rule Collect_subset)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   198
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   199
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   200
lemma choice_super: "[| ch : (PROD X:Pow(chain(S)) - {0}. X);   
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   201
         X : chain(S);  X ~: maxchain(S) |]      
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   202
      ==> ch ` super(S,X) : super(S,X)"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   203
apply (erule apply_type)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   204
apply (unfold super_def maxchain_def, blast)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   205
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   206
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   207
lemma choice_not_equals:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   208
     "[| ch : (PROD X:Pow(chain(S)) - {0}. X);       
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   209
         X : chain(S);  X ~: maxchain(S) |]      
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   210
      ==> ch ` super(S,X) ~= X"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   211
apply (rule notI)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   212
apply (drule choice_super, assumption)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   213
apply assumption
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   214
apply (simp add: super_def)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   215
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   216
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   217
(*This justifies Definition 4.4*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   218
lemma Hausdorff_next_exists:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   219
     "ch: (PROD X: Pow(chain(S))-{0}. X) ==>         
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   220
      EX next: increasing(S). ALL X: Pow(S).        
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   221
                   next`X = if(X: chain(S)-maxchain(S), ch`super(S,X), X)"
13175
81082cfa5618 new definition of "apply" and new simprule "beta_if"
paulson
parents: 13134
diff changeset
   222
apply (rule_tac x="\<lambda>X\<in>Pow(S). 
81082cfa5618 new definition of "apply" and new simprule "beta_if"
paulson
parents: 13134
diff changeset
   223
                   if X \<in> chain(S) - maxchain(S) then ch ` super(S, X) else X" 
81082cfa5618 new definition of "apply" and new simprule "beta_if"
paulson
parents: 13134
diff changeset
   224
       in bexI)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   225
apply force 
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   226
apply (unfold increasing_def)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   227
apply (rule CollectI)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   228
apply (rule lam_type)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   229
apply (simp (no_asm_simp))
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   230
apply (blast dest: super_subset_chain [THEN subsetD] chain_subset_Pow [THEN subsetD] choice_super)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   231
(*Now, verify that it increases*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   232
apply (simp (no_asm_simp) add: Pow_iff subset_refl)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   233
apply safe
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   234
apply (drule choice_super)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   235
apply (assumption+)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   236
apply (simp add: super_def, blast)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   237
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   238
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   239
(*Lemma 4*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   240
lemma TFin_chain_lemma4:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   241
     "[| c: TFin(S,next);                               
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   242
         ch: (PROD X: Pow(chain(S))-{0}. X);            
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   243
         next: increasing(S);                           
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   244
         ALL X: Pow(S). next`X =        
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   245
                          if(X: chain(S)-maxchain(S), ch`super(S,X), X) |] 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   246
     ==> c: chain(S)"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   247
apply (erule TFin_induct)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   248
apply (simp (no_asm_simp) add: chain_subset_Pow [THEN subsetD, THEN PowD] 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   249
            choice_super [THEN super_subset_chain [THEN subsetD]])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   250
apply (unfold chain_def)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   251
apply (rule CollectI, blast, safe)
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
   252
apply (rule_tac m1=B and n1=Ba in TFin_subset_linear [THEN disjE], fast+) 
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
   253
      (*Blast_tac's slow*)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   254
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   255
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
   256
theorem Hausdorff: "EX c. c : maxchain(S)"
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   257
apply (rule AC_Pi_Pow [THEN exE])
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   258
apply (rule Hausdorff_next_exists [THEN bexE], assumption)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   259
apply (rename_tac ch "next")
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   260
apply (subgoal_tac "Union (TFin (S,next)) : chain (S) ")
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   261
prefer 2
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   262
 apply (blast intro!: TFin_chain_lemma4 subset_refl [THEN TFin_UnionI])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   263
apply (rule_tac x = "Union (TFin (S,next))" in exI)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   264
apply (rule classical)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   265
apply (subgoal_tac "next ` Union (TFin (S,next)) = Union (TFin (S,next))")
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   266
apply (rule_tac [2] equal_next_Union [THEN iffD2, symmetric])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   267
apply (rule_tac [2] subset_refl [THEN TFin_UnionI])
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   268
prefer 2 apply assumption
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   269
apply (rule_tac [2] refl)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   270
apply (simp add: subset_refl [THEN TFin_UnionI, 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   271
                              THEN TFin.dom_subset [THEN subsetD, THEN PowD]])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   272
apply (erule choice_not_equals [THEN notE])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   273
apply (assumption+)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   274
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   275
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   276
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
   277
subsection{*Zorn's Lemma*}
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
   278
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
   279
text{*If all chains in S have upper bounds in S,
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
   280
       then S contains a maximal element *}
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   281
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   282
(*Used in the proof of Zorn's Lemma*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   283
lemma chain_extend: 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   284
    "[| c: chain(A);  z: A;  ALL x:c. x<=z |] ==> cons(z,c) : chain(A)"
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
   285
by (unfold chain_def, blast)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   286
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   287
lemma Zorn: "ALL c: chain(S). Union(c) : S ==> EX y:S. ALL z:S. y<=z --> y=z"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   288
apply (rule Hausdorff [THEN exE])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   289
apply (simp add: maxchain_def)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   290
apply (rename_tac c)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   291
apply (rule_tac x = "Union (c)" in bexI)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   292
prefer 2 apply blast
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   293
apply safe
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   294
apply (rename_tac z)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   295
apply (rule classical)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   296
apply (subgoal_tac "cons (z,c) : super (S,c) ")
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   297
apply (blast elim: equalityE)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   298
apply (unfold super_def, safe)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   299
apply (fast elim: chain_extend)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   300
apply (fast elim: equalityE)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   301
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   302
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   303
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
   304
subsection{*Zermelo's Theorem: Every Set can be Well-Ordered*}
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   305
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   306
(*Lemma 5*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   307
lemma TFin_well_lemma5:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   308
     "[| n: TFin(S,next);  Z <= TFin(S,next);  z:Z;  ~ Inter(Z) : Z |]   
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   309
      ==> ALL m:Z. n<=m"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   310
apply (erule TFin_induct)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   311
prefer 2 apply blast (*second induction step is easy*)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   312
apply (rule ballI)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   313
apply (rule bspec [THEN TFin_subsetD, THEN disjE], auto) 
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   314
apply (subgoal_tac "m = Inter (Z) ")
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   315
apply blast+
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   316
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   317
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   318
(*Well-ordering of TFin(S,next)*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   319
lemma well_ord_TFin_lemma: "[| Z <= TFin(S,next);  z:Z |] ==> Inter(Z) : Z"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   320
apply (rule classical)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   321
apply (subgoal_tac "Z = {Union (TFin (S,next))}")
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   322
apply (simp (no_asm_simp) add: Inter_singleton)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   323
apply (erule equal_singleton)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   324
apply (rule Union_upper [THEN equalityI])
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   325
apply (rule_tac [2] subset_refl [THEN TFin_UnionI, THEN TFin_well_lemma5, THEN bspec], blast+)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   326
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   327
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   328
(*This theorem just packages the previous result*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   329
lemma well_ord_TFin:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   330
     "next: increasing(S) ==> well_ord(TFin(S,next), Subset_rel(TFin(S,next)))"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   331
apply (rule well_ordI)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   332
apply (unfold Subset_rel_def linear_def)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   333
(*Prove the well-foundedness goal*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   334
apply (rule wf_onI)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   335
apply (frule well_ord_TFin_lemma, assumption)
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   336
apply (drule_tac x = "Inter (Z) " in bspec, assumption)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   337
apply blast
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   338
(*Now prove the linearity goal*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   339
apply (intro ballI)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   340
apply (case_tac "x=y")
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   341
 apply blast
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   342
(*The x~=y case remains*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   343
apply (rule_tac n1=x and m1=y in TFin_subset_linear [THEN disjE],
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   344
       assumption+, blast+)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   345
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   346
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   347
(** Defining the "next" operation for Zermelo's Theorem **)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   348
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   349
lemma choice_Diff:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   350
     "[| ch \<in> (\<Pi>X \<in> Pow(S) - {0}. X);  X \<subseteq> S;  X\<noteq>S |] ==> ch ` (S-X) \<in> S-X"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   351
apply (erule apply_type)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   352
apply (blast elim!: equalityE)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   353
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   354
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   355
(*This justifies Definition 6.1*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   356
lemma Zermelo_next_exists:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   357
     "ch: (PROD X: Pow(S)-{0}. X) ==>                
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   358
           EX next: increasing(S). ALL X: Pow(S).        
13175
81082cfa5618 new definition of "apply" and new simprule "beta_if"
paulson
parents: 13134
diff changeset
   359
                      next`X = (if X=S then S else cons(ch`(S-X), X))"
81082cfa5618 new definition of "apply" and new simprule "beta_if"
paulson
parents: 13134
diff changeset
   360
apply (rule_tac x="\<lambda>X\<in>Pow(S). if X=S then S else cons(ch`(S-X), X)"
81082cfa5618 new definition of "apply" and new simprule "beta_if"
paulson
parents: 13134
diff changeset
   361
       in bexI)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   362
apply force  
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   363
apply (unfold increasing_def)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   364
apply (rule CollectI)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   365
apply (rule lam_type)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   366
(*Type checking is surprisingly hard!*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   367
apply (simp (no_asm_simp) add: Pow_iff cons_subset_iff subset_refl)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   368
apply (blast intro!: choice_Diff [THEN DiffD1])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   369
(*Verify that it increases*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   370
apply (intro allI impI) 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   371
apply (simp add: Pow_iff subset_consI subset_refl)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   372
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   373
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   374
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   375
(*The construction of the injection*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   376
lemma choice_imp_injection:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   377
     "[| ch: (PROD X: Pow(S)-{0}. X);                  
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   378
         next: increasing(S);                          
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   379
         ALL X: Pow(S). next`X = if(X=S, S, cons(ch`(S-X), X)) |]  
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   380
      ==> (lam x:S. Union({y: TFin(S,next). x~: y}))        
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   381
               : inj(S, TFin(S,next) - {S})"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   382
apply (rule_tac d = "%y. ch` (S-y) " in lam_injective)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   383
apply (rule DiffI)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   384
apply (rule Collect_subset [THEN TFin_UnionI])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   385
apply (blast intro!: Collect_subset [THEN TFin_UnionI] elim: equalityE)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   386
apply (subgoal_tac "x ~: Union ({y: TFin (S,next) . x~: y}) ")
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   387
prefer 2 apply (blast elim: equalityE)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   388
apply (subgoal_tac "Union ({y: TFin (S,next) . x~: y}) ~= S")
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   389
prefer 2 apply (blast elim: equalityE)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   390
(*For proving x : next`Union(...)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   391
  Abrial & Laffitte's justification appears to be faulty.*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   392
apply (subgoal_tac "~ next ` Union ({y: TFin (S,next) . x~: y}) <= Union ({y: TFin (S,next) . x~: y}) ")
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   393
prefer 2
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   394
apply (simp del: Union_iff 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   395
            add: Collect_subset [THEN TFin_UnionI, THEN TFin_is_subset] 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   396
            Pow_iff cons_subset_iff subset_refl choice_Diff [THEN DiffD2])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   397
apply (subgoal_tac "x : next ` Union ({y: TFin (S,next) . x~: y}) ")
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   398
prefer 2
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   399
apply (blast intro!: Collect_subset [THEN TFin_UnionI] TFin.nextI)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   400
(*End of the lemmas!*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   401
apply (simp add: Collect_subset [THEN TFin_UnionI, THEN TFin_is_subset])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   402
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   403
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   404
(*The wellordering theorem*)
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
   405
theorem AC_well_ord: "EX r. well_ord(S,r)"
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   406
apply (rule AC_Pi_Pow [THEN exE])
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   407
apply (rule Zermelo_next_exists [THEN bexE], assumption)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   408
apply (rule exI)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   409
apply (rule well_ord_rvimage)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   410
apply (erule_tac [2] well_ord_TFin)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13175
diff changeset
   411
apply (rule choice_imp_injection [THEN inj_weaken_type], blast+)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
   412
done
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
   413
  
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 485
diff changeset
   414
end