src/HOL/Real/Hyperreal/Zorn.ML
author wenzelm
Tue, 07 Sep 1999 10:40:58 +0200
changeset 7499 23e090051cb8
parent 7322 d16d7ddcc842
child 8856 435187ffc64e
permissions -rw-r--r--
isatool expandshort;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5979
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     1
(*  Title       : Zorn.ML
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 5979
diff changeset
     2
    ID          : $Id$
5979
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     3
    Author      : Jacques D. Fleuriot
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     4
    Copyright   : 1998  University of Cambridge
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     5
    Description : Zorn's Lemma -- adapted proofs from lcp's ZF/Zorn.ML
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     6
*) 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     7
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     8
(*---------------------------------------------------------------
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     9
      Section 1.  Mathematical Preamble 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    10
 ---------------------------------------------------------------*)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    11
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    12
Goal "(ALL x:C. x<=A | B<=x) ==> Union(C)<=A | B<=Union(C)";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    13
by (Blast_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    14
qed "Union_lemma0";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    15
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    16
(*-- similar to subset_cs in ZF/subset.thy --*)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    17
val thissubset_SIs =
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    18
    [subset_refl,Union_least, UN_least, Un_least, 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    19
     Inter_greatest, Int_greatest,
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    20
     Un_upper1, Un_upper2, Int_lower1, Int_lower2];
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    21
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    22
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    23
(*A claset for subset reasoning*)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    24
val thissubset_cs = claset() 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    25
    delrules [subsetI, subsetCE]
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    26
    addSIs thissubset_SIs
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    27
    addIs  [Union_upper, Inter_lower];
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    28
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    29
(* increasingD2 of ZF/Zorn.ML *) 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    30
Goalw [succ_def] "x <= succ S x";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    31
by (rtac (expand_if RS iffD2) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    32
by (auto_tac (claset(),simpset() addsimps [super_def,
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    33
               maxchain_def,psubset_def]));
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    34
by (rtac swap 1 THEN assume_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    35
by (rtac selectI2 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    36
by (ALLGOALS(Blast_tac));
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    37
qed "Abrial_axiom1";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    38
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    39
val [TFin_succI, Pow_TFin_UnionI] = TFin.intrs;
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    40
val TFin_UnionI = PowI RS Pow_TFin_UnionI;
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    41
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    42
val major::prems = Goal  
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    43
          "[| n : TFin S; \
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    44
\            !!x. [| x: TFin S; P(x) |] ==> P(succ S x); \
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    45
\            !!Y. [| Y <= TFin S; Ball Y P |] ==> P(Union Y) |] \
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    46
\         ==> P(n)";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    47
by (rtac (major RS TFin.induct) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    48
by (ALLGOALS (fast_tac (claset() addIs prems)));
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    49
qed "TFin_induct";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    50
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    51
(*Perform induction on n, then prove the major premise using prems. *)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    52
fun TFin_ind_tac a prems i = 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    53
    EVERY [res_inst_tac [("n",a)] TFin_induct i,
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    54
           rename_last_tac a ["1"] (i+1),
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    55
           rename_last_tac a ["2"] (i+2),
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    56
           ares_tac prems i];
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    57
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    58
Goal "x <= y ==> x <= succ S y";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    59
by (etac (Abrial_axiom1 RSN (2,subset_trans)) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    60
qed "succ_trans";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    61
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    62
(*Lemma 1 of section 3.1*)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    63
Goal "[| n: TFin S;  m: TFin S;  \
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    64
\        ALL x: TFin S. x <= m --> x = m | succ S x <= m \
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    65
\     |] ==> n <= m | succ S m <= n";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    66
by (etac TFin_induct 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    67
by (etac Union_lemma0 2);               (*or just Blast_tac*)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    68
by (blast_tac (thissubset_cs addIs [succ_trans]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    69
qed "TFin_linear_lemma1";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    70
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    71
(* Lemma 2 of section 3.2 *)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    72
Goal "m: TFin S ==> ALL n: TFin S. n<=m --> n=m | succ S n<=m";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    73
by (etac TFin_induct 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    74
by (rtac (impI RS ballI) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    75
(*case split using TFin_linear_lemma1*)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    76
by (res_inst_tac [("n1","n"), ("m1","x")] 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    77
    (TFin_linear_lemma1 RS disjE) 1  THEN  REPEAT (assume_tac 1));
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    78
by (dres_inst_tac [("x","n")] bspec 1 THEN assume_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    79
by (blast_tac (thissubset_cs addIs [succ_trans]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    80
by (REPEAT (ares_tac [disjI1,equalityI] 1));
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    81
(*second induction step*)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    82
by (rtac (impI RS ballI) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    83
by (rtac (Union_lemma0 RS disjE) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    84
by (rtac disjI2 3);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    85
by (REPEAT (ares_tac [disjI1,equalityI] 2));
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    86
by (rtac ballI 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    87
by (ball_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    88
by (set_mp_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    89
by (res_inst_tac [("n1","n"), ("m1","x")] 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    90
    (TFin_linear_lemma1 RS disjE) 1  THEN  REPEAT (assume_tac 1));
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    91
by (blast_tac thissubset_cs 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    92
by (rtac (Abrial_axiom1 RS subset_trans RS disjI1) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    93
by (assume_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    94
qed "TFin_linear_lemma2";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    95
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    96
(*a more convenient form for Lemma 2*)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    97
Goal "[| n<=m;  m: TFin S;  n: TFin S |] ==> n=m | succ S n<=m";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    98
by (rtac (TFin_linear_lemma2 RS bspec RS mp) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    99
by (REPEAT (assume_tac 1));
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   100
qed "TFin_subsetD";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   101
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   102
(*Consequences from section 3.3 -- Property 3.2, the ordering is total*)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   103
Goal "[| m: TFin S;  n: TFin S|] ==> n<=m | m<=n";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   104
by (rtac (TFin_linear_lemma2 RSN (3,TFin_linear_lemma1) RS disjE) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   105
by (REPEAT (assume_tac 1) THEN etac disjI2 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   106
by (blast_tac (thissubset_cs addIs [Abrial_axiom1 RS subset_trans]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   107
qed "TFin_subset_linear";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   108
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   109
(*Lemma 3 of section 3.3*)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   110
Goal "[| n: TFin S;  m: TFin S;  m = succ S m |] ==> n<=m";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   111
by (etac TFin_induct 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   112
by (dtac TFin_subsetD 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   113
by (REPEAT (assume_tac 1));
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   114
by (fast_tac (claset() addEs [ssubst]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   115
by (blast_tac (thissubset_cs) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   116
qed "eq_succ_upper";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   117
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   118
(*Property 3.3 of section 3.3*)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   119
Goal "m: TFin S ==> (m = succ S m) = (m = Union(TFin S))";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   120
by (rtac iffI 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   121
by (rtac (Union_upper RS equalityI) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   122
by (rtac (eq_succ_upper RS Union_least) 2);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   123
by (REPEAT (assume_tac 1));
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   124
by (etac ssubst 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   125
by (rtac (Abrial_axiom1 RS equalityI) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   126
by (blast_tac (thissubset_cs addIs [TFin_UnionI, TFin_succI]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   127
qed "equal_succ_Union";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   128
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   129
(*-------------------------------------------------------------------------
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   130
    Section 4.  Hausdorff's Theorem: every set contains a maximal chain 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   131
    NB: We assume the partial ordering is <=, the subset relation! 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   132
 -------------------------------------------------------------------------*)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   133
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   134
Goalw [chain_def] "({} :: 'a set set) : chain S";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   135
by (Auto_tac);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   136
qed "empty_set_mem_chain";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   137
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   138
Goalw [super_def] "super S c <= chain S";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   139
by (Fast_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   140
qed "super_subset_chain";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   141
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   142
Goalw [maxchain_def] "maxchain S <= chain S";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   143
by (Fast_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   144
qed "maxchain_subset_chain";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   145
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   146
Goalw [succ_def] "c ~: chain S ==> succ S c = c";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   147
by (fast_tac (claset() addSIs [if_P]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   148
qed "succI1";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   149
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   150
Goalw [succ_def] "c: maxchain S ==> succ S c = c";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   151
by (fast_tac (claset() addSIs [if_P]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   152
qed "succI2";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   153
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   154
Goalw [succ_def] "c: chain S - maxchain S ==> \
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   155
\                         succ S c = (@c'. c': super S c)";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   156
by (fast_tac (claset() addSIs [if_not_P]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   157
qed "succI3";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   158
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   159
Goal "c: chain S - maxchain S ==> ? d. d: super S c";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   160
by (rewrite_goals_tac [super_def,maxchain_def]);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   161
by (Auto_tac);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   162
qed "mem_super_Ex";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   163
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   164
Goal "c: chain S - maxchain S ==> \
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   165
\                         (@c'. c': super S c): super S c";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   166
by (etac (mem_super_Ex RS exE) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   167
by (rtac selectI2 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   168
by (Auto_tac);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   169
qed "select_super";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   170
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   171
Goal "c: chain S - maxchain S ==> \
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   172
\                         (@c'. c': super S c) ~= c";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   173
by (rtac notI 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   174
by (dtac select_super 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   175
by (asm_full_simp_tac (simpset() addsimps [super_def,psubset_def]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   176
qed "select_not_equals";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   177
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   178
Goal "c: chain S - maxchain S ==> \
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   179
\                         succ S c ~= c";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7322
diff changeset
   180
by (ftac succI3 1);
5979
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   181
by (Asm_simp_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   182
by (rtac select_not_equals 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   183
by (assume_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   184
qed "succ_not_equals";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   185
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   186
Goal "c: TFin S ==> (c :: 'a set set): chain S";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   187
by (etac TFin_induct 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   188
by (asm_simp_tac (simpset() addsimps [succ_def,
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   189
    select_super RS (super_subset_chain RS subsetD)]
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   190
                   setloop split_tac [expand_if]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   191
by (rewtac chain_def);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   192
by (rtac CollectI 1);
7322
d16d7ddcc842 isatool expandshort;
wenzelm
parents: 7219
diff changeset
   193
by Safe_tac;
5979
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   194
by (dtac bspec 1 THEN assume_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   195
by (res_inst_tac  [("m1","Xa"), ("n1","X")] (TFin_subset_linear RS disjE) 2);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   196
by (ALLGOALS(Blast_tac));
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   197
qed "TFin_chain_lemm4";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   198
 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   199
Goal "EX c. (c :: 'a set set): maxchain S";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   200
by (res_inst_tac [("x", "Union(TFin S)")] exI 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   201
by (rtac classical 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   202
by (subgoal_tac "succ S (Union(TFin S)) = Union(TFin S)" 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   203
by (resolve_tac [equal_succ_Union RS iffD2 RS sym] 2);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   204
by (resolve_tac [subset_refl RS TFin_UnionI] 2);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   205
by (rtac refl 2);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   206
by (cut_facts_tac [subset_refl RS TFin_UnionI RS TFin_chain_lemm4] 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   207
by (dtac (DiffI RS succ_not_equals) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   208
by (ALLGOALS(Blast_tac));
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   209
qed "Hausdorff";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   210
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   211
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   212
(*---------------------------------------------------------------
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   213
  Section 5.  Zorn's Lemma: if all chains have upper bounds 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   214
                               there is  a maximal element 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   215
 ----------------------------------------------------------------*)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   216
Goalw [chain_def]
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   217
    "[| c: chain S; z: S; \
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   218
\             ALL x:c. x<=(z:: 'a set) |] ==> {z} Un c : chain S";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   219
by (Blast_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   220
qed "chain_extend";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   221
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   222
Goalw [chain_def] "[| c: chain S; x: c |] ==> x <= Union(c)";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   223
by (Auto_tac);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   224
qed "chain_Union_upper";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   225
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   226
Goalw [chain_def] "c: chain S ==> ! x: c. x <= Union(c)";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   227
by (Auto_tac);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   228
qed "chain_ball_Union_upper";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   229
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   230
Goal "[| c: maxchain S; u: S; Union(c) <= u |] ==> Union(c) = u";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   231
by (rtac ccontr 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   232
by (asm_full_simp_tac (simpset() addsimps [maxchain_def]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   233
by (etac conjE 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   234
by (subgoal_tac "({u} Un c): super S c" 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   235
by (Asm_full_simp_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   236
by (rewrite_tac [super_def,psubset_def]);
7322
d16d7ddcc842 isatool expandshort;
wenzelm
parents: 7219
diff changeset
   237
by Safe_tac;
5979
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   238
by (fast_tac (claset() addEs [chain_extend]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   239
by (subgoal_tac "u ~: c" 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   240
by (blast_tac (claset() addEs [equalityE]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   241
by (blast_tac (claset() addDs [chain_Union_upper]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   242
qed "maxchain_Zorn";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   243
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   244
Goal "ALL c: chain S. Union(c): S ==> \
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   245
\     EX y: S. ALL z: S. y <= z --> y = z";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   246
by (cut_facts_tac [Hausdorff,maxchain_subset_chain] 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   247
by (etac exE 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   248
by (dtac subsetD 1 THEN assume_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   249
by (dtac bspec 1 THEN assume_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   250
by (res_inst_tac [("x","Union(c)")] bexI 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   251
by (rtac ballI 1 THEN rtac impI 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   252
by (blast_tac (claset() addSDs [maxchain_Zorn]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   253
by (assume_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   254
qed "Zorn_Lemma";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   255
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   256
(*-------------------------------------------------------------
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   257
             Alternative version of Zorn's Lemma
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   258
 --------------------------------------------------------------*)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   259
Goal "ALL (c:: 'a set set): chain S. EX y : S. ALL x : c. x <= y ==> \
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   260
\     EX y : S. ALL x : S. (y :: 'a set) <= x --> y = x";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   261
by (cut_facts_tac [Hausdorff,maxchain_subset_chain] 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   262
by (EVERY1[etac exE, dtac subsetD, assume_tac]);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   263
by (EVERY1[dtac bspec, assume_tac, etac bexE]);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   264
by (res_inst_tac [("x","y")] bexI 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   265
by (assume_tac 2);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   266
by (EVERY1[rtac ballI, rtac impI, rtac ccontr]);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   267
by (forw_inst_tac [("z","x")]  chain_extend 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   268
by (assume_tac 1 THEN Blast_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   269
by (rewrite_tac [maxchain_def,super_def,psubset_def]);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   270
by (Step_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   271
by (eres_inst_tac [("c","{x} Un c")] equalityCE 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   272
by (Step_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   273
by (subgoal_tac "x ~: c" 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   274
by (blast_tac (claset() addEs [equalityE]) 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   275
by (Blast_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   276
qed "Zorn_Lemma2";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   277
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   278
(** misc. lemmas **)
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   279
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   280
Goalw [chain_def] "[| c : chain S; x: c; y: c |] ==> x <= y | y <= x";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   281
by (Blast_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   282
qed "chainD";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   283
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   284
Goalw [chain_def] "!!(c :: 'a set set). c: chain S ==> c <= S";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   285
by (Blast_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   286
qed "chainD2";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   287
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   288
(* proved elsewhere? *) 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   289
Goal "x : Union(c) ==> EX m:c. x:m";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   290
by (Blast_tac 1);
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   291
qed "mem_UnionD";
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
   292