src/ZF/AC/AC16_WO4.ML
author paulson
Fri, 08 Nov 1996 14:02:51 +0100
changeset 2167 5819e85ad261
parent 1932 cc9f1ba8f29a
child 2469 b50b8c0eec01
permissions -rw-r--r--
Adjusting to new version of uresult
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
     1
(*  Title:      ZF/AC/AC16_WO4.ML
1200
d4551b1a6da7 Many small changes to make proofs run faster
lcp
parents: 1196
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
     3
    Author:     Krzysztof Grabczewski
1200
d4551b1a6da7 Many small changes to make proofs run faster
lcp
parents: 1196
diff changeset
     4
d4551b1a6da7 Many small changes to make proofs run faster
lcp
parents: 1196
diff changeset
     5
  The proof of AC16(n, k) ==> WO4(n-k)
d4551b1a6da7 Many small changes to make proofs run faster
lcp
parents: 1196
diff changeset
     6
*)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     7
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     8
open AC16_WO4;
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     9
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    10
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    11
(* The case of finite set                                                 *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    12
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    13
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    14
goalw thy [Finite_def] "!!A. [| Finite(A); 0<m; m:nat |] ==>  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    15
\       EX a f. Ord(a) & domain(f) = a &  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    16
\               (UN b<a. f`b) = A & (ALL b<a. f`b lepoll m)";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
    17
by (etac bexE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    18
by (dresolve_tac [eqpoll_sym RS (eqpoll_def RS def_imp_iff RS iffD1)] 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
    19
by (etac exE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    20
by (res_inst_tac [("x","n")] exI 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    21
by (res_inst_tac [("x","lam i:n. {f`i}")] exI 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    22
by (asm_full_simp_tac AC_ss 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    23
by (rewrite_goals_tac [bij_def, surj_def]);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    24
by (fast_tac (AC_cs addSIs [ltI, nat_into_Ord, lam_funtype RS domain_of_fun,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    25
        equalityI, singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    26
        nat_1_lepoll_iff RS iffD2]
1932
cc9f1ba8f29a Tidying: removing redundant args in classical reasoner calls
paulson
parents: 1924
diff changeset
    27
        addSEs [apply_type, ltE]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    28
val lemma1 = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    29
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    30
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    31
(* The case of infinite set                                               *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    32
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    33
2167
5819e85ad261 Adjusting to new version of uresult
paulson
parents: 1932
diff changeset
    34
(* well_ord(x,r) ==> well_ord({{y,z}. y:x}, Something(x,z))  **)
5819e85ad261 Adjusting to new version of uresult
paulson
parents: 1932
diff changeset
    35
val well_ord_paired = standard (paired_bij RS bij_is_inj RS well_ord_rvimage);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    36
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    37
goal thy "!!A. [| A lepoll B; ~ A lepoll C |] ==> ~ B lepoll C";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    38
by (fast_tac (FOL_cs addEs [notE, lepoll_trans]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    39
val lepoll_trans1 = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    40
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    41
goalw thy [lepoll_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    42
        "!!X.[| Y lepoll X; well_ord(X, R) |] ==> EX S. well_ord(Y, S)";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    43
by (fast_tac (AC_cs addSEs [well_ord_rvimage]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    44
val well_ord_lepoll = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    45
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    46
goal thy "!!X. [| well_ord(X,R); well_ord(Y,S)  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    47
\               |] ==> EX T. well_ord(X Un Y, T)";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    48
by (eresolve_tac [well_ord_radd RS (Un_lepoll_sum RS well_ord_lepoll)] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    49
by (assume_tac 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    50
val well_ord_Un = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    51
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    52
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    53
(* There exists a well ordered set y such that ...                        *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    54
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    55
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    56
goal thy "EX y R. well_ord(y,R) & x Int y = 0 & ~y lepoll z & ~Finite(y)";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    57
by (res_inst_tac [("x","{{a,x}. a:nat Un  Hartog(z)}")] exI 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    58
by (resolve_tac [Ord_nat RS well_ord_Memrel RS (Ord_Hartog RS
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    59
                well_ord_Memrel RSN (2, well_ord_Un)) RS exE] 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    60
by (fast_tac (AC_cs addSIs [Ord_Hartog, well_ord_Memrel, well_ord_paired,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    61
        equals0I, HartogI RSN (2, lepoll_trans1),
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    62
        subset_imp_lepoll RS (paired_eqpoll RS eqpoll_sym RS
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    63
        eqpoll_imp_lepoll RSN (2, lepoll_trans))]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    64
        addSEs [RepFunE, nat_not_Finite RS notE] addEs [mem_asym]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    65
        addSDs [Un_upper1 RS subset_imp_lepoll RS lepoll_Finite,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    66
        paired_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    67
        RS lepoll_Finite]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    68
val lemma2 = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    69
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    70
val [prem] = goal thy "~Finite(B) ==> ~Finite(A Un B)";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    71
by (fast_tac (AC_cs
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    72
        addSIs [subset_imp_lepoll RS (prem RSN (2, lepoll_infinite))]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    73
val infinite_Un = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    74
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    75
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    76
(* There is a v : s_u such that k lepoll x Int y (in our case succ(k))    *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    77
(* The idea of the proof is the following :                               *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    78
(* Suppose not, i.e. every element of s_u has exactly k-1 elements of y   *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    79
(* Thence y is less than or equipollent to {v:Pow(x). v eqpoll n#-k}      *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    80
(*   We have obtained this result in two steps :                          *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    81
(*   1. y is less than or equipollent to {v:s_u. a <= v}                  *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    82
(*      where a is certain k-2 element subset of y                        *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    83
(*   2. {v:s_u. a <= v} is less than or equipollent                       *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    84
(*      to {v:Pow(x). v eqpoll n-k}                                       *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    85
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    86
1710
f50ec5b35937 Simplified KG's proofs
paulson
parents: 1461
diff changeset
    87
(*Proof simplified by LCP*)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    88
goal thy "!!A. [| ~(EX x:A. f`x=y); f : inj(A, B); y:B |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    89
\       ==> (lam a:succ(A). if(a=A, y, f`a)) : inj(succ(A), B)";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    90
by (res_inst_tac [("d","%z. if(z=y, A, converse(f)`z)")] lam_injective 1);
1710
f50ec5b35937 Simplified KG's proofs
paulson
parents: 1461
diff changeset
    91
by (ALLGOALS
f50ec5b35937 Simplified KG's proofs
paulson
parents: 1461
diff changeset
    92
    (asm_simp_tac 
f50ec5b35937 Simplified KG's proofs
paulson
parents: 1461
diff changeset
    93
     (ZF_ss addsimps [inj_is_fun RS apply_type, left_inverse] 
f50ec5b35937 Simplified KG's proofs
paulson
parents: 1461
diff changeset
    94
      setloop (split_tac [expand_if] ORELSE' step_tac ZF_cs))));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    95
val succ_not_lepoll_lemma = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    96
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    97
goalw thy [lepoll_def, eqpoll_def, bij_def, surj_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    98
        "!!A. [| ~A eqpoll B; A lepoll B |] ==> succ(A) lepoll B";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    99
by (fast_tac (AC_cs addSEs [succ_not_lepoll_lemma, inj_is_fun]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   100
val succ_not_lepoll_imp_eqpoll = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   101
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   102
val [prem] = goalw thy [s_u_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   103
        "(ALL v:s_u(u, t_n, k, y). k eqpoll v Int y ==> False)  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   104
\       ==> EX v : s_u(u, t_n, k, y). succ(k) lepoll v Int y";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   105
by (excluded_middle_tac "?P" 1 THEN (assume_tac 2));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   106
by (resolve_tac [prem RS FalseE] 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   107
by (rtac ballI 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   108
by (etac CollectE 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   109
by (etac conjE 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   110
by (etac swap 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   111
by (fast_tac (AC_cs addSEs [succ_not_lepoll_imp_eqpoll]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   112
val suppose_not = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   113
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   114
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   115
(* There is a k-2 element subset of y                                     *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   116
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   117
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   118
goalw thy [lepoll_def, eqpoll_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   119
        "!!X. [| n:nat; nat lepoll X |] ==> EX Y. Y<=X & n eqpoll Y";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   120
by (fast_tac (FOL_cs addSDs [Ord_nat RSN (2, OrdmemD) RSN (2, restrict_inj)]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   121
        addSIs [subset_refl]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   122
        addSEs [restrict_bij, inj_is_fun RS fun_is_rel RS image_subset]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   123
val nat_lepoll_imp_ex_eqpoll_n = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   124
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   125
val ordertype_eqpoll =
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   126
        ordermap_bij RS (exI RS (eqpoll_def RS def_imp_iff RS iffD2));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   127
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   128
goal thy "!!y. [| well_ord(y,R); ~Finite(y); n:nat  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   129
\       |] ==> EX z. z<=y & n eqpoll z";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   130
by (etac nat_lepoll_imp_ex_eqpoll_n 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   131
by (resolve_tac [ordertype_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   132
        RSN (2, lepoll_trans)] 1 THEN (assume_tac 2));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   133
by (fast_tac (AC_cs addSIs [nat_le_infinite_Ord RS le_imp_lepoll]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   134
                addSEs [Ord_ordertype, ordertype_eqpoll RS eqpoll_imp_lepoll
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   135
                        RS lepoll_infinite]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   136
val ex_subset_eqpoll_n = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   137
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   138
goalw thy [lesspoll_def] "!!n. n: nat ==> n lesspoll nat";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   139
by (fast_tac (AC_cs addSEs [Ord_nat RSN (2, ltI) RS leI RS le_imp_lepoll,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   140
        eqpoll_sym RS eqpoll_imp_lepoll]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   141
        addIs [Ord_nat RSN (2, nat_succI RS ltI) RS leI
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   142
        RS le_imp_lepoll RS lepoll_trans RS succ_lepoll_natE]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   143
val n_lesspoll_nat = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   144
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   145
goal thy "!!y. [| well_ord(y,R); ~Finite(y); k eqpoll a; a <= y; k: nat |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   146
\       ==> y - a eqpoll y";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   147
by (fast_tac (empty_cs addIs [lepoll_lesspoll_lesspoll]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   148
        addSIs [Card_cardinal, Diff_lesspoll_eqpoll_Card RS eqpoll_trans,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   149
                Card_cardinal RS Card_is_Ord RS nat_le_infinite_Ord
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   150
                RS le_imp_lepoll]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   151
        addSEs [well_ord_cardinal_eqpoll,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   152
                well_ord_cardinal_eqpoll RS eqpoll_sym,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   153
                eqpoll_sym RS eqpoll_imp_lepoll,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   154
                n_lesspoll_nat RS lesspoll_lepoll_lesspoll,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   155
                well_ord_cardinal_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   156
                RS lepoll_infinite]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   157
val Diff_Finite_eqpoll = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   158
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   159
goal thy "!!x. [| a<=y; b:y-a; u:x |] ==> cons(b, cons(u, a)) : Pow(x Un y)";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   160
by (fast_tac AC_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   161
val cons_cons_subset = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   162
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   163
goal thy "!!x. [| a eqpoll k; a<=y; b:y-a; u:x; x Int y = 0  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   164
\       |] ==> cons(b, cons(u, a)) eqpoll succ(succ(k))";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   165
by (fast_tac (AC_cs addSIs [cons_eqpoll_succ] addEs [equals0D]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   166
val cons_cons_eqpoll = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   167
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   168
goalw thy [s_u_def] "s_u(u, t_n, k, y) <= t_n";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   169
by (fast_tac AC_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   170
val s_u_subset = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   171
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   172
goalw thy [s_u_def, succ_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   173
        "!!w. [| w:t_n; cons(b,cons(u,a)) <= w; a <= y; b : y-a; k eqpoll a  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   174
\       |] ==> w: s_u(u, t_n, succ(k), y)";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   175
by (fast_tac (AC_cs addDs [eqpoll_imp_lepoll RS cons_lepoll_cong]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   176
                addSEs [subset_imp_lepoll RSN (2, lepoll_trans)]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   177
val s_uI = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   178
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   179
goalw thy [s_u_def] "!!v. v : s_u(u, t_n, k, y) ==> u : v";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   180
by (fast_tac AC_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   181
val in_s_u_imp_u_in = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   182
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   183
goal thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   184
        "!!y. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(k))}.  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   185
\       EX! w. w:t_n & z <= w; \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   186
\       k eqpoll a; a <= y; b : y - a; u : x; x Int y = 0 |]  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   187
\       ==> EX! c. c:{v:s_u(u, t_n, succ(k), y). a <= v} & b:c";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   188
by (etac ballE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   189
by (fast_tac (FOL_cs addSIs [CollectI, cons_cons_subset,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   190
                eqpoll_sym RS cons_cons_eqpoll]) 2);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   191
by (etac ex1E 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   192
by (res_inst_tac [("a","w")] ex1I 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   193
by (rtac conjI 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   194
by (rtac CollectI 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   195
by (fast_tac (FOL_cs addSEs [s_uI]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   196
by (fast_tac AC_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   197
by (fast_tac AC_cs 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   198
by (etac allE 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   199
by (etac impE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   200
by (assume_tac 2);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   201
by (fast_tac (AC_cs addSEs [s_u_subset RS subsetD, in_s_u_imp_u_in]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   202
val ex1_superset_a = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   203
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   204
goal thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   205
        "!!A. [| succ(k) eqpoll A; k eqpoll B; B <= A; a : A-B; k:nat  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   206
\       |] ==> A = cons(a, B)";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   207
by (rtac equalityI 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   208
by (fast_tac AC_cs 2);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   209
by (resolve_tac [Diff_eq_0_iff RS iffD1] 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   210
by (rtac equals0I 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   211
by (dresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   212
by (dresolve_tac [eqpoll_sym RS cons_eqpoll_succ] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   213
by (fast_tac AC_cs 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   214
by (dtac cons_eqpoll_succ 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   215
by (fast_tac AC_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   216
by (fast_tac (AC_cs addSIs [nat_succI]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   217
        addSEs [[eqpoll_sym RS eqpoll_imp_lepoll, subset_imp_lepoll] MRS
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   218
        (lepoll_trans RS lepoll_trans) RS succ_lepoll_natE]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   219
val set_eq_cons = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   220
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   221
goal thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   222
        "!!y. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(k))}.  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   223
\       EX! w. w:t_n & z <= w; \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   224
\       ALL v:s_u(u, t_n, succ(k), y). succ(k) eqpoll v Int y;  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   225
\       k eqpoll a; a <= y; b : y - a; u : x; x Int y = 0; k:nat   \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   226
\       |] ==> (THE c. c:{v:s_u(u, t_n, succ(k), y). a <= v} & b:c)  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   227
\       Int y = cons(b, a)";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   228
by (dresolve_tac [ex1_superset_a RS theI] 1 THEN REPEAT (assume_tac 1));
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   229
by (rtac set_eq_cons 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   230
by (REPEAT (fast_tac AC_cs 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   231
val the_eq_cons = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   232
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   233
goal thy "!!a. [| cons(x,a) = cons(y,a); x~: a |] ==> x = y ";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   234
by (fast_tac (AC_cs addSEs [equalityE]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   235
val cons_eqE = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   236
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   237
goal thy "!!A. A = B ==> A Int C = B Int C";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   238
by (asm_simp_tac AC_ss 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   239
val eq_imp_Int_eq = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   240
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   241
goal thy "!!a. [| a=b; a=c; b=d |] ==> c=d";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   242
by (asm_full_simp_tac AC_ss 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   243
val msubst = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   244
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   245
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   246
(*   1. y is less than or equipollent to {v:s_u. a <= v}                  *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   247
(*      where a is certain k-2 element subset of y                        *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   248
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   249
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   250
goal thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   251
        "!!y. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(k))}.  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   252
\       EX! w. w:t_n & z <= w; \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   253
\       ALL v:s_u(u, t_n, succ(k), y). succ(k) eqpoll v Int y;  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   254
\       well_ord(y,R); ~Finite(y); k eqpoll a; a <= y;  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   255
\       k:nat; u:x; x Int y = 0 |]  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   256
\       ==> y lepoll {v:s_u(u, t_n, succ(k), y). a <= v}";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   257
by (resolve_tac [Diff_Finite_eqpoll RS eqpoll_sym RS 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   258
        eqpoll_imp_lepoll RS lepoll_trans] 1
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   259
        THEN REPEAT (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   260
by (res_inst_tac [("f3","lam b:y-a.  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   261
\       THE c. c:{v:s_u(u, t_n, succ(k), y). a <= v} & b:c")]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   262
        (exI RS (lepoll_def RS def_imp_iff RS iffD2)) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   263
by (resolve_tac [inj_def RS def_imp_eq RS ssubst] 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   264
by (rtac CollectI 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   265
by (rtac lam_type 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   266
by (resolve_tac [ex1_superset_a RS theI RS conjunct1] 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   267
        THEN REPEAT (assume_tac 1));
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   268
by (rtac ballI 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   269
by (rtac ballI 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   270
by (resolve_tac [beta RS ssubst] 1 THEN (assume_tac 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   271
by (resolve_tac [beta RS ssubst] 1 THEN (assume_tac 1));
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   272
by (rtac impI 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   273
by (rtac cons_eqE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   274
by (fast_tac AC_cs 2);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   275
by (dres_inst_tac [("A","THE c. ?P(c)"), ("C","y")] eq_imp_Int_eq 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   276
by (eresolve_tac [[asm_rl, the_eq_cons, the_eq_cons] MRS msubst] 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   277
        THEN REPEAT (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   278
val y_lepoll_subset_s_u = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   279
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   280
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   281
(* some arithmetic                                                        *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   282
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   283
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   284
goal thy 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   285
        "!!k. [| k:nat; m:nat |] ==>  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   286
\       ALL A B. A eqpoll k #+ m & k lepoll B & B<=A --> A-B lepoll m";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   287
by (eres_inst_tac [("n","k")] nat_induct 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   288
by (simp_tac (AC_ss addsimps [add_0]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   289
by (fast_tac (AC_cs addIs [eqpoll_imp_lepoll RS
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   290
        (Diff_subset RS subset_imp_lepoll RS lepoll_trans)]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   291
by (REPEAT (resolve_tac [allI,impI] 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   292
by (resolve_tac [succ_lepoll_imp_not_empty RS not_emptyE] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   293
by (fast_tac AC_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   294
by (eres_inst_tac [("x","A - {xa}")] allE 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   295
by (eres_inst_tac [("x","B - {xa}")] allE 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   296
by (etac impE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   297
by (asm_full_simp_tac (AC_ss addsimps [add_succ]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   298
by (fast_tac (AC_cs addSIs [Diff_sing_eqpoll, lepoll_Diff_sing]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   299
by (res_inst_tac [("P","%z. z lepoll m")] subst 1 THEN (assume_tac 2));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   300
by (fast_tac (AC_cs addSIs [equalityI]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   301
val eqpoll_sum_imp_Diff_lepoll_lemma = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   302
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   303
goal thy "!!k. [| A eqpoll succ(k #+ m); B<=A; succ(k) lepoll B;  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   304
\       k:nat; m:nat |]  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   305
\       ==> A-B lepoll m";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   306
by (dresolve_tac [add_succ RS ssubst] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   307
by (dresolve_tac [nat_succI RS eqpoll_sum_imp_Diff_lepoll_lemma] 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   308
        THEN (REPEAT (assume_tac 1)));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   309
by (fast_tac AC_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   310
val eqpoll_sum_imp_Diff_lepoll = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   311
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   312
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   313
(* similar properties for eqpoll                                          *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   314
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   315
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   316
goal thy 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   317
        "!!k. [| k:nat; m:nat |] ==>  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   318
\       ALL A B. A eqpoll k #+ m & k eqpoll B & B<=A --> A-B eqpoll m";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   319
by (eres_inst_tac [("n","k")] nat_induct 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   320
by (fast_tac (AC_cs addSDs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_0_is_0]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   321
        addss (AC_ss addsimps [add_0])) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   322
by (REPEAT (resolve_tac [allI,impI] 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   323
by (resolve_tac [succ_lepoll_imp_not_empty RS not_emptyE] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   324
by (fast_tac (AC_cs addSEs [eqpoll_imp_lepoll]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   325
by (eres_inst_tac [("x","A - {xa}")] allE 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   326
by (eres_inst_tac [("x","B - {xa}")] allE 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   327
by (etac impE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   328
by (fast_tac (AC_cs addSIs [Diff_sing_eqpoll,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   329
        eqpoll_sym RSN (2, Diff_sing_eqpoll) RS eqpoll_sym]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   330
        addss (AC_ss addsimps [add_succ])) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   331
by (res_inst_tac [("P","%z. z eqpoll m")] subst 1 THEN (assume_tac 2));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   332
by (fast_tac (AC_cs addSIs [equalityI]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   333
val eqpoll_sum_imp_Diff_eqpoll_lemma = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   334
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   335
goal thy "!!k. [| A eqpoll succ(k #+ m); B<=A; succ(k) eqpoll B;  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   336
\       k:nat; m:nat |]  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   337
\       ==> A-B eqpoll m";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   338
by (dresolve_tac [add_succ RS ssubst] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   339
by (dresolve_tac [nat_succI RS eqpoll_sum_imp_Diff_eqpoll_lemma] 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   340
        THEN (REPEAT (assume_tac 1)));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   341
by (fast_tac AC_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   342
val eqpoll_sum_imp_Diff_eqpoll = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   343
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   344
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   345
(* back to the second part                                                *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   346
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   347
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   348
goal thy "!!w. [| x Int y = 0; w <= x Un y |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   349
\        ==> w Int (x - {u}) = w - cons(u, w Int y)";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   350
by (fast_tac (AC_cs addSIs [equalityI] addEs [equals0D]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   351
val w_Int_eq_w_Diff = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   352
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   353
goal thy "!!w. [| w:{v:s_u(u, t_n, succ(l), y). a <= v};  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   354
\       l eqpoll a; t_n <= {v:Pow(x Un y). v eqpoll succ(succ(l) #+ m)};  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   355
\       m:nat; l:nat; x Int y = 0; u : x;  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   356
\       ALL v:s_u(u, t_n, succ(l), y). succ(l) eqpoll v Int y  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   357
\       |] ==> w Int (x - {u}) eqpoll m";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   358
by (etac CollectE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   359
by (resolve_tac [w_Int_eq_w_Diff RS ssubst] 1 THEN (assume_tac 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   360
by (fast_tac (AC_cs addSDs [s_u_subset RS subsetD]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   361
by (fast_tac (AC_cs addEs [equals0D] addSDs [bspec]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   362
        addDs [s_u_subset RS subsetD]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   363
        addSEs [eqpoll_sym RS cons_eqpoll_succ RS eqpoll_sym, in_s_u_imp_u_in]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   364
        addSIs [nat_succI, eqpoll_sum_imp_Diff_eqpoll]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   365
val w_Int_eqpoll_m = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   366
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   367
goal thy "!!m. [| 0<m; x eqpoll m; m:nat |] ==> x ~= 0";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   368
by (fast_tac (empty_cs
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   369
        addSEs [mem_irrefl, ltE, eqpoll_succ_imp_not_empty, natE]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   370
val eqpoll_m_not_empty = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   371
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   372
goal thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   373
        "!!z. [| z : xa Int (x - {u}); l eqpoll a; a <= y; x Int y = 0; u:x  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   374
\       |] ==> cons(z, cons(u, a)) : {v: Pow(x Un y). v eqpoll succ(succ(l))}";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   375
by (fast_tac (AC_cs addSIs [cons_eqpoll_succ] addEs [equals0D, eqpoll_sym]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   376
val cons_cons_in = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   377
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   378
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   379
(*   2. {v:s_u. a <= v} is less than or equipollent                       *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   380
(*      to {v:Pow(x). v eqpoll n-k}                                       *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   381
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   382
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   383
goal thy 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   384
        "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(l))}.  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   385
\       EX! w. w:t_n & z <= w; \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   386
\       t_n <= {v:Pow(x Un y). v eqpoll succ(succ(l) #+ m)};  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   387
\       0<m; m:nat; l:nat;  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   388
\       ALL v:s_u(u, t_n, succ(l), y). succ(l) eqpoll v Int y;  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   389
\       a <= y; l eqpoll a; x Int y = 0; u : x  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   390
\       |] ==> {v:s_u(u, t_n, succ(l), y). a <= v}  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   391
\               lepoll {v:Pow(x). v eqpoll m}";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   392
by (res_inst_tac [("f3","lam w:{v:s_u(u, t_n, succ(l), y). a <= v}.  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   393
\       w Int (x - {u})")] 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   394
        (exI RS (lepoll_def RS def_imp_iff RS iffD2)) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   395
by (resolve_tac [inj_def RS def_imp_eq RS ssubst] 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   396
by (rtac CollectI 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   397
by (rtac lam_type 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   398
by (rtac CollectI 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   399
by (fast_tac AC_cs 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   400
by (rtac w_Int_eqpoll_m 1 THEN REPEAT (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   401
by (simp_tac AC_ss 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   402
by (REPEAT (resolve_tac [ballI, impI] 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   403
by (eresolve_tac [w_Int_eqpoll_m RSN (2, eqpoll_m_not_empty) RS not_emptyE] 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   404
        THEN REPEAT (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   405
by (dresolve_tac [equalityD1 RS subsetD] 1 THEN (assume_tac 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   406
by (dresolve_tac [cons_cons_in RSN (2, bspec)] 1 THEN REPEAT (assume_tac 1));
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   407
by (etac ex1_two_eq 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   408
by (fast_tac (AC_cs addSEs [s_u_subset RS subsetD, in_s_u_imp_u_in]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   409
by (fast_tac (AC_cs addSEs [s_u_subset RS subsetD, in_s_u_imp_u_in]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   410
val subset_s_u_lepoll_w = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   411
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   412
goal thy "!!k. [| 0<k; k:nat |] ==> EX l:nat. k = succ(l)";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   413
by (etac natE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   414
by (fast_tac (empty_cs addSEs [ltE, mem_irrefl]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   415
by (fast_tac (empty_cs addSIs [refl, bexI]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   416
val ex_eq_succ = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   417
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   418
goal thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   419
        "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}.  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   420
\       EX! w. w:t_n & z <= w; \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   421
\       well_ord(y,R); ~Finite(y); u:x; x Int y = 0;  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   422
\       t_n <= {v:Pow(x Un y). v eqpoll succ(k #+ m)};  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   423
\       ~ y lepoll {v:Pow(x). v eqpoll m}; 0<k; 0<m; k:nat; m:nat  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   424
\       |] ==> EX v : s_u(u, t_n, k, y). succ(k) lepoll v Int y";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   425
by (rtac suppose_not 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   426
by (eresolve_tac [ex_eq_succ RS bexE] 1 THEN (assume_tac 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   427
by (hyp_subst_tac 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   428
by (res_inst_tac [("n1","xa")] (ex_subset_eqpoll_n RS exE) 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   429
        THEN REPEAT (assume_tac 1));
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   430
by (etac conjE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   431
by (forward_tac [[y_lepoll_subset_s_u, subset_s_u_lepoll_w] MRS lepoll_trans] 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   432
        THEN REPEAT (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   433
by (contr_tac 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   434
val exists_proper_in_s_u = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   435
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   436
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   437
(* LL can be well ordered                                                 *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   438
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   439
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   440
goal thy "{x:Pow(X). x lepoll 0} = {0}";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   441
by (fast_tac (AC_cs addSDs [lepoll_0_is_0]
1924
0f1a583457da Corrected for new classical reasoner: redundant rules
paulson
parents: 1710
diff changeset
   442
                    addSIs [lepoll_refl, equalityI]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   443
val subsets_lepoll_0_eq_unit = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   444
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   445
goal thy "!!X. [| well_ord(X, R); ~Finite(X); n:nat |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   446
\               ==> EX S. well_ord({Y: Pow(X) . Y eqpoll succ(n)}, S)";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   447
by (resolve_tac [well_ord_infinite_subsets_eqpoll_X
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   448
        RS (eqpoll_def RS def_imp_iff RS iffD1) RS exE] 1
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   449
    THEN (REPEAT (assume_tac 1)));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   450
by (fast_tac (ZF_cs addSEs [bij_is_inj RS well_ord_rvimage]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   451
val well_ord_subsets_eqpoll_n = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   452
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   453
goal thy "!!n. n:nat ==> {z:Pow(y). z lepoll succ(n)} =  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   454
\       {z:Pow(y). z lepoll n} Un {z:Pow(y). z eqpoll succ(n)}";
1200
d4551b1a6da7 Many small changes to make proofs run faster
lcp
parents: 1196
diff changeset
   455
by (fast_tac (ZF_cs addIs [le_refl, leI,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   456
                le_imp_lepoll, equalityI]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   457
                addSDs [lepoll_succ_disj]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   458
                addSEs [nat_into_Ord, lepoll_trans, eqpoll_imp_lepoll]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   459
val subsets_lepoll_succ = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   460
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   461
goal thy "!!n. n:nat ==>  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   462
\       {z:Pow(y). z lepoll n} Int {z:Pow(y). z eqpoll succ(n)} = 0";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   463
by (fast_tac (ZF_cs addSEs [eqpoll_sym RS eqpoll_imp_lepoll 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   464
                RS lepoll_trans RS succ_lepoll_natE]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   465
                addSIs [equals0I]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   466
val Int_empty = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   467
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   468
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   469
(* unit set is well-ordered by the empty relation                         *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   470
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   471
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   472
goalw thy [irrefl_def, trans_on_def, part_ord_def, linear_def, tot_ord_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   473
        "tot_ord({a},0)";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   474
by (simp_tac ZF_ss 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   475
val tot_ord_unit = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   476
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   477
goalw thy [wf_on_def, wf_def] "wf[{a}](0)";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   478
by (fast_tac (ZF_cs addSIs [equalityI]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   479
val wf_on_unit = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   480
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   481
goalw thy [well_ord_def] "well_ord({a},0)";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   482
by (simp_tac (ZF_ss addsimps [tot_ord_unit, wf_on_unit]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   483
val well_ord_unit = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   484
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   485
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   486
(* well_ord_subsets_lepoll_n                                              *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   487
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   488
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   489
goal thy "!!y r. [| well_ord(y,r); ~Finite(y); n:nat |] ==>  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   490
\       EX R. well_ord({z:Pow(y). z lepoll n}, R)";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   491
by (etac nat_induct 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   492
by (fast_tac (ZF_cs addSIs [well_ord_unit]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   493
        addss (ZF_ss addsimps [subsets_lepoll_0_eq_unit])) 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   494
by (etac exE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   495
by (eresolve_tac [well_ord_subsets_eqpoll_n RS exE] 1 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   496
        THEN REPEAT (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   497
by (asm_simp_tac (ZF_ss addsimps [subsets_lepoll_succ]) 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   498
by (dtac well_ord_radd 1 THEN (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   499
by (eresolve_tac [Int_empty RS disj_Un_eqpoll_sum RS 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   500
                (eqpoll_def RS def_imp_iff RS iffD1) RS exE] 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   501
by (fast_tac (ZF_cs addSEs [bij_is_inj RS well_ord_rvimage]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   502
val well_ord_subsets_lepoll_n = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   503
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   504
goalw thy [LL_def, MM_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   505
        "!!x. t_n <= {v:Pow(x Un y). v eqpoll n}  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   506
\               ==> LL(t_n, k, y) <= {z:Pow(y). z lepoll n}";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   507
by (fast_tac (AC_cs addSEs [RepFunE]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   508
        addIs [subset_imp_lepoll RS (eqpoll_imp_lepoll
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   509
                RSN (2, lepoll_trans))]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   510
val LL_subset = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   511
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   512
goal thy "!!x. [| t_n <= {v:Pow(x Un y). v eqpoll n}; \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   513
\               well_ord(y, R); ~Finite(y); n:nat  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   514
\               |] ==> EX S. well_ord(LL(t_n, k, y), S)";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   515
by (fast_tac (FOL_cs addIs [exI]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   516
                addSEs [LL_subset RSN (2,  well_ord_subset)]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   517
                addEs [well_ord_subsets_lepoll_n RS exE]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   518
val well_ord_LL = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   519
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   520
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   521
(* every element of LL is a contained in exactly one element of MM        *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   522
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   523
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   524
goalw thy [MM_def, LL_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   525
        "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   526
\       t_n <= {v:Pow(x Un y). v eqpoll n}; \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   527
\       v:LL(t_n, k, y)  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   528
\       |] ==> EX! w. w:MM(t_n, k, y) & v<=w";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   529
by (step_tac (AC_cs addSEs [RepFunE]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   530
by (resolve_tac [lepoll_imp_eqpoll_subset RS exE] 1 THEN (assume_tac 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   531
by (eres_inst_tac [("x","xa")] ballE 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   532
by (fast_tac (AC_cs addSEs [eqpoll_sym]) 2);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   533
by (res_inst_tac [("a","v")] ex1I 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   534
by (fast_tac AC_cs 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   535
by (etac ex1E 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   536
by (res_inst_tac [("x","v")] allE 1 THEN (assume_tac 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   537
by (eres_inst_tac [("x","xb")] allE 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   538
by (fast_tac AC_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   539
val unique_superset_in_MM = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   540
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   541
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   542
(* The function GG satisfies the conditions of WO4                        *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   543
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   544
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   545
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   546
(* The union of appropriate values is the whole x                         *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   547
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   548
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   549
goal thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   550
        "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}.  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   551
\       EX! w. w:t_n & z <= w; \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   552
\       well_ord(y,R); ~Finite(y); u:x; x Int y = 0;  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   553
\       t_n <= {v:Pow(x Un y). v eqpoll succ(k #+ m)};  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   554
\       ~ y lepoll {v:Pow(x). v eqpoll m}; 0<k; 0<m; k:nat; m:nat  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   555
\       |] ==> EX w:MM(t_n, succ(k), y). u:w";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   556
by (eresolve_tac [exists_proper_in_s_u RS bexE] 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   557
        THEN REPEAT (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   558
by (rewrite_goals_tac [MM_def, s_u_def]);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   559
by (fast_tac AC_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   560
val exists_in_MM = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   561
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   562
goalw thy [LL_def] "!!w. w : MM(t_n, k, y) ==> w Int y : LL(t_n, k, y)";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   563
by (fast_tac AC_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   564
val Int_in_LL = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   565
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   566
goalw thy [MM_def] "MM(t_n, k, y) <= t_n";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   567
by (fast_tac AC_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   568
val MM_subset = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   569
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   570
goal thy 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   571
        "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}.  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   572
\       EX! w. w:t_n & z <= w; \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   573
\       well_ord(y,R); ~Finite(y); u:x; x Int y = 0;  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   574
\       t_n <= {v:Pow(x Un y). v eqpoll succ(k #+ m)};  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   575
\       ~ y lepoll {v:Pow(x). v eqpoll m}; 0<k; 0<m; k:nat; m:nat  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   576
\       |] ==> EX w:LL(t_n, succ(k), y). u:GG(t_n, succ(k), y)`w";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   577
by (forward_tac [exists_in_MM] 1 THEN REPEAT (assume_tac 1));
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   578
by (etac bexE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   579
by (res_inst_tac [("x","w Int y")] bexI 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   580
by (etac Int_in_LL 2);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   581
by (rewtac GG_def);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   582
by (asm_full_simp_tac (AC_ss addsimps [Int_in_LL]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   583
by (eresolve_tac [unique_superset_in_MM RS the_equality2 RS ssubst] 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   584
        THEN (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   585
by (REPEAT (fast_tac (AC_cs addEs [equals0D] addSEs [Int_in_LL]) 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   586
val exists_in_LL = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   587
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   588
goalw thy [LL_def] 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   589
        "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   590
\       t_n <= {v:Pow(x Un y). v eqpoll n};  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   591
\       v : LL(t_n, k, y) |]  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   592
\       ==> v = (THE x. x : MM(t_n, k, y) & v <= x) Int y";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   593
by (fast_tac (AC_cs addSEs [Int_in_LL,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   594
                unique_superset_in_MM RS the_equality2 RS ssubst]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   595
val in_LL_eq_Int = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   596
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   597
goal thy  
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   598
        "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   599
\       t_n <= {v:Pow(x Un y). v eqpoll n};  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   600
\       v : LL(t_n, k, y) |]  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   601
\       ==> (THE x. x : MM(t_n, k, y) & v <= x) <= x Un y";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   602
by (fast_tac (AC_cs addSDs [unique_superset_in_MM RS theI RS conjunct1 RS 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   603
        (MM_subset RS subsetD)]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   604
val the_in_MM_subset = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   605
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   606
goalw thy [GG_def] 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   607
        "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   608
\       t_n <= {v:Pow(x Un y). v eqpoll n};  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   609
\       v : LL(t_n, k, y) |]  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   610
\       ==> GG(t_n, k, y) ` v <= x";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   611
by (asm_full_simp_tac AC_ss 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   612
by (forward_tac [the_in_MM_subset] 1 THEN REPEAT (assume_tac 1));
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   613
by (dtac in_LL_eq_Int 1 THEN REPEAT (assume_tac 1));
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   614
by (rtac subsetI 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   615
by (etac DiffE 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   616
by (etac swap 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   617
by (fast_tac (AC_cs addEs [ssubst]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   618
val GG_subset = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   619
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   620
goal thy  
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   621
        "!!x. [| well_ord(LL(t_n, succ(k), y), S);  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   622
\       ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}. EX! w. w:t_n & z <= w; \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   623
\       well_ord(y,R); ~Finite(y); x Int y = 0;  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   624
\       t_n <= {v:Pow(x Un y). v eqpoll succ(k #+ m)};  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   625
\       ~ y lepoll {v:Pow(x). v eqpoll m}; 0<k; 0<m; k:nat; m:nat  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   626
\       |] ==> (UN b<ordertype(LL(t_n, succ(k), y), S).  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   627
\       (GG(t_n, succ(k), y)) `  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   628
\       (converse(ordermap(LL(t_n, succ(k), y), S)) ` b)) = x";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   629
by (rtac equalityI 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   630
by (rtac subsetI 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   631
by (etac OUN_E 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   632
by (eresolve_tac [GG_subset RS subsetD] 1 THEN TRYALL assume_tac);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   633
by (eresolve_tac [ordermap_bij RS bij_converse_bij RS
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   634
                bij_is_fun RS apply_type] 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   635
by (etac ltD 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   636
by (rtac subsetI 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   637
by (eresolve_tac [exists_in_LL RS bexE] 1 THEN REPEAT (assume_tac 1));
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   638
by (rtac OUN_I 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   639
by (resolve_tac [Ord_ordertype RSN (2, ltI)] 1 THEN (assume_tac 2));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   640
by (eresolve_tac [ordermap_type RS apply_type] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   641
by (eresolve_tac [ordermap_bij RS bij_is_inj RS left_inverse RS ssubst] 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   642
        THEN REPEAT (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   643
val OUN_eq_x = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   644
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   645
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   646
(* Every element of the family is less than or equipollent to n-k (m)     *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   647
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   648
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   649
goalw thy [MM_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   650
        "!!w. [| w : MM(t_n, k, y); t_n <= {v:Pow(x Un y). v eqpoll n}  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   651
\       |] ==> w eqpoll n";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   652
by (fast_tac AC_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   653
val in_MM_eqpoll_n = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   654
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   655
goalw thy [LL_def, MM_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   656
        "!!w. w : LL(t_n, k, y) ==> k lepoll w";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   657
by (fast_tac AC_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   658
val in_LL_eqpoll_n = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   659
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   660
goalw thy [GG_def] 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   661
        "!!x. [|  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   662
\       ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}. EX! w. w:t_n & z <= w; \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   663
\       t_n <= {v:Pow(x Un y). v eqpoll succ(k #+ m)};  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   664
\       well_ord(LL(t_n, succ(k), y), S); k:nat; m:nat |]  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   665
\       ==> ALL b<ordertype(LL(t_n, succ(k), y), S).  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   666
\       (GG(t_n, succ(k), y)) `  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   667
\       (converse(ordermap(LL(t_n, succ(k), y), S)) ` b) lepoll m";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   668
by (rtac oallI 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   669
by (asm_full_simp_tac (AC_ss addsimps [ltD,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   670
        ordermap_bij RS bij_converse_bij RS bij_is_fun RS apply_type]) 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   671
by (rtac eqpoll_sum_imp_Diff_lepoll 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   672
by (REPEAT (fast_tac (FOL_cs addSDs [ltD]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   673
        addSIs [eqpoll_sum_imp_Diff_lepoll, in_LL_eqpoll_n]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   674
        addEs [unique_superset_in_MM RS theI RS conjunct1 RS in_MM_eqpoll_n,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   675
        in_LL_eq_Int RS equalityD1 RS (Int_lower1 RSN (2, subset_trans)),
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   676
        ordermap_bij RS bij_converse_bij RS bij_is_fun RS apply_type]) 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   677
val all_in_lepoll_m = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   678
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   679
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   680
(* The main theorem AC16(n, k) ==> WO4(n-k)                               *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   681
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   682
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   683
goalw thy [AC16_def,WO4_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   684
        "!!n k. [| AC16(k #+ m, k); 0 < k; 0 < m; k:nat; m:nat |] ==> WO4(m)";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   685
by (rtac allI 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   686
by (excluded_middle_tac "Finite(A)" 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1200
diff changeset
   687
by (rtac lemma1 2 THEN REPEAT (assume_tac 2));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   688
by (resolve_tac [lemma2 RS revcut_rl] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   689
by (REPEAT (eresolve_tac [exE, conjE] 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   690
by (eres_inst_tac [("x","A Un y")] allE 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   691
by (forward_tac [infinite_Un] 1 THEN (mp_tac 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   692
by (REPEAT (eresolve_tac [exE, conjE] 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   693
by (resolve_tac [well_ord_LL RS exE] 1 THEN REPEAT (assume_tac 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   694
by (fast_tac (AC_cs addSIs [nat_succI, add_type]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   695
by (res_inst_tac [("x","ordertype(LL(T, succ(k), y), x)")] exI 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   696
by (res_inst_tac [("x","lam b:ordertype(LL(T, succ(k), y), x).  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   697
\       (GG(T, succ(k), y)) `  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   698
\       (converse(ordermap(LL(T, succ(k), y), x)) ` b)")] exI 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   699
by (simp_tac AC_ss 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   700
by (fast_tac (empty_cs addSIs [conjI, lam_funtype RS domain_of_fun]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   701
        addSEs [Ord_ordertype, all_in_lepoll_m, OUN_eq_x]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   702
qed "AC16_WO4";