src/ZF/AC/Cardinal_aux.ML
author wenzelm
Fri, 16 Jul 1999 22:25:07 +0200
changeset 7025 afbd8241797b
parent 6153 bff90585cce5
child 8123 a71686059be0
permissions -rw-r--r--
tuned dest_lexicon;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
     1
(*  Title:      ZF/AC/Cardinal_aux.ML
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
     3
    Author:     Krzysztof Grabczewski
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     4
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     5
Auxiliary lemmas concerning cardinalities
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     6
*)
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
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
     9
(* Lemmas involving ordinals and cardinalities used in the proofs         *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    10
(* concerning AC16 and DC                                                 *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    11
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    12
1616
6d7278c3f686 Moved some proofs to Cardinal.ML; simplified others
paulson
parents: 1461
diff changeset
    13
(* j=|A| *)
6d7278c3f686 Moved some proofs to Cardinal.ML; simplified others
paulson
parents: 1461
diff changeset
    14
goal Cardinal.thy
6d7278c3f686 Moved some proofs to Cardinal.ML; simplified others
paulson
parents: 1461
diff changeset
    15
    "!!A. [| A lepoll i; Ord(i) |] ==> EX j. j le i & A eqpoll j";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    16
by (fast_tac (claset() addIs [lepoll_cardinal_le, well_ord_Memrel,
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1932
diff changeset
    17
                            well_ord_cardinal_eqpoll RS eqpoll_sym]
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1932
diff changeset
    18
                    addDs [lepoll_well_ord]) 1);
1616
6d7278c3f686 Moved some proofs to Cardinal.ML; simplified others
paulson
parents: 1461
diff changeset
    19
qed "lepoll_imp_ex_le_eqpoll";
6d7278c3f686 Moved some proofs to Cardinal.ML; simplified others
paulson
parents: 1461
diff changeset
    20
6d7278c3f686 Moved some proofs to Cardinal.ML; simplified others
paulson
parents: 1461
diff changeset
    21
(* j=|A| *)
6d7278c3f686 Moved some proofs to Cardinal.ML; simplified others
paulson
parents: 1461
diff changeset
    22
goalw Cardinal.thy [lesspoll_def]
6d7278c3f686 Moved some proofs to Cardinal.ML; simplified others
paulson
parents: 1461
diff changeset
    23
    "!!A a. [| A lesspoll i; Ord(i) |] ==> EX j. j<i & A eqpoll j";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    24
by (fast_tac (claset() addSDs [lepoll_imp_ex_le_eqpoll] addSEs [leE]) 1);
1616
6d7278c3f686 Moved some proofs to Cardinal.ML; simplified others
paulson
parents: 1461
diff changeset
    25
qed "lesspoll_imp_ex_lt_eqpoll";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    26
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    27
Goalw [InfCard_def] "[| ~Finite(i); Ord(i) |] ==> InfCard(|i|)";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
    28
by (rtac conjI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
    29
by (rtac Card_cardinal 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    30
by (resolve_tac [Card_nat RS (Card_def RS def_imp_iff RS iffD1 RS ssubst)] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    31
by (resolve_tac [nat_le_infinite_Ord RS le_imp_lepoll
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    32
        RSN (2, well_ord_Memrel RS well_ord_lepoll_imp_Card_le)] 1 
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    33
    THEN REPEAT (assume_tac 1));
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
    34
qed "Inf_Ord_imp_InfCard_cardinal";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    35
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    36
Goal "[| A eqpoll i; B eqpoll i; ~Finite(i); Ord(i) |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    37
\               ==> A Un B eqpoll i";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
    38
by (rtac eqpollI 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    39
by (eresolve_tac [subset_imp_lepoll RSN (2, eqpoll_sym RS eqpoll_imp_lepoll
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    40
        RS  lepoll_trans)] 2);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1932
diff changeset
    41
by (Fast_tac 2);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    42
by (resolve_tac [Un_lepoll_sum RS lepoll_trans] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    43
by (resolve_tac [lepoll_imp_sum_lepoll_prod RS lepoll_trans] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    44
by (eresolve_tac [eqpoll_sym RSN (2, eqpoll_trans) RS eqpoll_imp_lepoll] 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    45
        THEN (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    46
by (resolve_tac [nat_le_infinite_Ord RS le_imp_lepoll RS 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    47
        (Ord_nat RS (nat_2I RS OrdmemD) RS subset_imp_lepoll RS lepoll_trans)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    48
        RS (eqpoll_sym RS eqpoll_imp_lepoll RSN (2, lepoll_trans))] 1 
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    49
    THEN (REPEAT (assume_tac 1)));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    50
by (eresolve_tac [prod_eqpoll_cong RS eqpoll_imp_lepoll RS lepoll_trans] 1 
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    51
    THEN (assume_tac 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    52
by (resolve_tac [Inf_Ord_imp_InfCard_cardinal RSN (2, well_ord_Memrel RS 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    53
        well_ord_InfCard_square_eq) RS eqpoll_imp_lepoll] 1 
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    54
    THEN REPEAT (assume_tac 1));
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
    55
qed "Un_eqpoll_Inf_Ord";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    56
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    57
5315
c9ad6bbf3a34 tidying
paulson
parents: 5284
diff changeset
    58
Goal "?f : bij({{y,z}. y:x}, x)";
5284
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5241
diff changeset
    59
by (rtac RepFun_bijective 1);
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5241
diff changeset
    60
by (simp_tac (simpset() addsimps [doubleton_eq_iff]) 1);
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5241
diff changeset
    61
by (Blast_tac 1);
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
    62
qed "paired_bij";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    63
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
    64
Goalw [eqpoll_def] "{{y,z}. y:x} eqpoll x";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    65
by (fast_tac (claset() addSIs [paired_bij]) 1);
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
    66
qed "paired_eqpoll";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    67
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    68
Goal "EX B. B eqpoll A & B Int C = 0";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    69
by (fast_tac (claset() addSIs [paired_eqpoll, equals0I] addEs [mem_asym]) 1);
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
    70
qed "ex_eqpoll_disjoint";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    71
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    72
Goal "[| A lepoll i; B lepoll i; ~Finite(i); Ord(i) |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    73
\               ==> A Un B lepoll i";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    74
by (res_inst_tac [("A1","i"), ("C1","i")] (ex_eqpoll_disjoint RS exE) 1);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
    75
by (etac conjE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    76
by (dresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RSN (2, lepoll_trans)] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    77
by (assume_tac 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    78
by (resolve_tac [Un_lepoll_Un RS lepoll_trans] 1 THEN (REPEAT (assume_tac 1)));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    79
by (eresolve_tac [eqpoll_refl RSN (2, Un_eqpoll_Inf_Ord) RS
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    80
        eqpoll_imp_lepoll] 1
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    81
        THEN (REPEAT (assume_tac 1)));
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
    82
qed "Un_lepoll_Inf_Ord";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    83
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    84
Goal "[| P(i); i:j; Ord(j) |] ==> (LEAST i. P(i)) : j";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    85
by (eresolve_tac [Least_le RS leE] 1);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
    86
by (etac Ord_in_Ord 1 THEN (assume_tac 1));
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
    87
by (etac ltE 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    88
by (fast_tac (claset() addDs [OrdmemD]) 1);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
    89
by (etac subst_elem 1 THEN (assume_tac 1));
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
    90
qed "Least_in_Ord";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    91
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    92
Goal "[| well_ord(x,r); y<=x; y lepoll succ(n); n:nat |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    93
\       ==> y-{THE b. first(b,y,r)} lepoll n";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    94
by (res_inst_tac [("Q","y=0")] (excluded_middle RS disjE) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    95
by (fast_tac (claset() addSIs [Diff_sing_lepoll, the_first_in]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    96
by (res_inst_tac [("b","y-{THE b. first(b, y, r)}")] subst 1);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
    97
by (rtac empty_lepollI 2);
2496
40efb87985b5 Removal of needless "addIs [equality]", etc.
paulson
parents: 2493
diff changeset
    98
by (Fast_tac 1);
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
    99
qed "Diff_first_lepoll";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   100
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
   101
Goal "(UN x:X. P(x)) <= (UN x:X. P(x)-Q(x)) Un (UN x:X. Q(x))";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1932
diff changeset
   102
by (Fast_tac 1);
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
   103
qed "UN_subset_split";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   104
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   105
Goalw [lepoll_def] "Ord(a) ==> (UN x:a. {P(x)}) lepoll a";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   106
by (res_inst_tac [("x","lam z:(UN x:a. {P(x)}). (LEAST i. P(i)=z)")] exI 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   107
by (res_inst_tac [("d","%z. P(z)")] lam_injective 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   108
by (fast_tac (claset() addSIs [Least_in_Ord]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   109
by (fast_tac (claset() addIs [LeastI] addSEs [Ord_in_Ord]) 1);
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
   110
qed "UN_sing_lepoll";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   111
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   112
Goal "[| well_ord(T, R); ~Finite(a); Ord(a); n:nat |] ==>  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   113
\       ALL f. (ALL b:a. f`b lepoll n & f`b <= T) --> (UN b:a. f`b) lepoll a";
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 5315
diff changeset
   114
by (induct_tac "n" 1);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   115
by (rtac allI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   116
by (rtac impI 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   117
by (res_inst_tac [("b","UN b:a. f`b")] subst 1);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   118
by (rtac empty_lepollI 2);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   119
by (resolve_tac [equals0I RS sym] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   120
by (REPEAT (eresolve_tac [UN_E, allE] 1));
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   121
by (fast_tac (claset() addDs [lepoll_0_is_0 RS subst]) 1);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   122
by (rtac allI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   123
by (rtac impI 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   124
by (eres_inst_tac [("x","lam x:a. f`x - {THE b. first(b,f`x,R)}")] allE 1);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   125
by (etac impE 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1932
diff changeset
   126
by (Asm_full_simp_tac 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   127
by (fast_tac (claset() addSIs [Diff_first_lepoll]) 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1932
diff changeset
   128
by (Asm_full_simp_tac 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   129
by (resolve_tac [UN_subset_split RS subset_imp_lepoll RS lepoll_trans] 1);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   130
by (rtac Un_lepoll_Inf_Ord 1 THEN (REPEAT_FIRST assume_tac));
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   131
by (etac UN_sing_lepoll 1);
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
   132
qed "UN_fun_lepoll_lemma";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   133
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   134
Goal "[| ALL b:a. f`b lepoll n & f`b <= T; well_ord(T, R);  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   135
\       ~Finite(a); Ord(a); n:nat |] ==> (UN b:a. f`b) lepoll a";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   136
by (eresolve_tac [UN_fun_lepoll_lemma RS allE] 1 THEN (REPEAT (assume_tac 1)));
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1932
diff changeset
   137
by (Fast_tac 1);
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
   138
qed "UN_fun_lepoll";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   139
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   140
Goal "[| ALL b:a. F(b) lepoll n & F(b) <= T; well_ord(T, R);  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   141
\       ~Finite(a); Ord(a); n:nat |] ==> (UN b:a. F(b)) lepoll a";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   142
by (rtac impE 1 THEN (assume_tac 3));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   143
by (res_inst_tac [("f","lam b:a. F(b)")] (UN_fun_lepoll) 2 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   144
        THEN (TRYALL assume_tac));
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6070
diff changeset
   145
by Auto_tac;
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
   146
qed "UN_lepoll";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   147
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   148
Goal "Ord(a) ==> (UN b:a. F(b)) = (UN b:a. F(b) - (UN c:b. F(c)))";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   149
by (rtac equalityI 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1932
diff changeset
   150
by (Fast_tac 2);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   151
by (rtac subsetI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   152
by (etac UN_E 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   153
by (rtac UN_I 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   154
by (res_inst_tac [("P","%z. x:F(z)")] Least_in_Ord 1 THEN (REPEAT (assume_tac 1)));
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   155
by (rtac DiffI 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   156
by (resolve_tac [Ord_in_Ord RSN (2, LeastI)] 1 THEN (REPEAT (assume_tac 1)));
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   157
by (rtac notI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   158
by (etac UN_E 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   159
by (eres_inst_tac [("P","%z. x:F(z)"),("i","c")] less_LeastE 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   160
by (eresolve_tac [Ord_Least RSN (2, ltI)] 1);
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
   161
qed "UN_eq_UN_Diffs";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   162
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
   163
Goalw [lepoll_def, eqpoll_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   164
        "a lepoll X ==> EX Y. Y<=X & a eqpoll Y";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   165
by (etac exE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   166
by (forward_tac [subset_refl RSN (2, restrict_bij)] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   167
by (res_inst_tac [("x","f``a")] exI 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   168
by (fast_tac (claset() addSEs [inj_is_fun RS fun_is_rel RS image_subset]) 1);
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
   169
qed "lepoll_imp_eqpoll_subset";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   170
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   171
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   172
(* Diff_lesspoll_eqpoll_Card                                              *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   173
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   174
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   175
Goal "[| A eqpoll a; ~Finite(a); Card(a); B lesspoll a;  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   176
\               A-B lesspoll a |] ==> P";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   177
by (REPEAT (eresolve_tac [lesspoll_imp_ex_lt_eqpoll RS exE,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   178
        Card_is_Ord, conjE] 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   179
by (forw_inst_tac [("j","xa")] (lt_Ord RS (lt_Ord RSN (2, Un_upper1_le))) 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   180
        THEN (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   181
by (forw_inst_tac [("j","xa")] (lt_Ord RS (lt_Ord RSN (2, Un_upper2_le))) 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   182
        THEN (assume_tac 1));
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   183
by (dtac Un_least_lt 1 THEN (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   184
by (dresolve_tac [le_imp_lepoll RSN
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   185
        (2, eqpoll_imp_lepoll RS lepoll_trans)] 1
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   186
        THEN (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   187
by (dresolve_tac [le_imp_lepoll RSN
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   188
        (2, eqpoll_imp_lepoll RS lepoll_trans)] 1
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   189
        THEN (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   190
by (res_inst_tac [("Q","Finite(x Un xa)")] (excluded_middle RS disjE) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   191
by (dresolve_tac [[lepoll_Finite, lepoll_Finite] MRS Finite_Un] 2
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   192
        THEN (REPEAT (assume_tac 2)));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   193
by (dresolve_tac [subset_Un_Diff RS subset_imp_lepoll RS lepoll_Finite] 2);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   194
by (fast_tac (claset()
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   195
        addDs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_Finite]) 2);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   196
by (dresolve_tac [ Un_lepoll_Inf_Ord] 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   197
        THEN (REPEAT (assume_tac 1)));
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   198
by (fast_tac (claset() addSEs [ltE, Ord_in_Ord]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   199
by (dresolve_tac [subset_Un_Diff RS subset_imp_lepoll RS lepoll_trans RSN
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   200
        (3, lt_Card_imp_lesspoll RS lepoll_lesspoll_lesspoll)] 1
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   201
        THEN (TRYALL assume_tac));
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   202
by (fast_tac (claset() addSDs [lesspoll_def RS def_imp_iff RS iffD1]) 1);
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
   203
qed "Diff_lesspoll_eqpoll_Card_lemma";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   204
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   205
Goal "[| A eqpoll a; ~Finite(a); Card(a); B lesspoll a |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   206
\       ==> A - B eqpoll a";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1932
diff changeset
   207
by (rtac swap 1 THEN (Fast_tac 1));
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1200
diff changeset
   208
by (rtac Diff_lesspoll_eqpoll_Card_lemma 1 THEN (REPEAT (assume_tac 1)));
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   209
by (fast_tac (claset() addSIs [lesspoll_def RS def_imp_iff RS iffD2,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   210
        subset_imp_lepoll RS (eqpoll_imp_lepoll RSN (2, lepoll_trans))]) 1);
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2873
diff changeset
   211
qed "Diff_lesspoll_eqpoll_Card";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   212