# HG changeset patch # User paulson # Date 1011199926 -3600 # Node ID 249600a63ba9b78ac6e7235eb344d8033ff2057a # Parent 1748c16c2df38df24427358a62c21d4df35ce319 Isar version of AC diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC15_WO6.ML --- a/src/ZF/AC/AC15_WO6.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,292 +0,0 @@ -(* Title: ZF/AC/AC15_WO6.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -The proofs needed to state that AC10, ..., AC15 are equivalent to the rest. -We need the following: - -WO1 ==> AC10(n) ==> AC11 ==> AC12 ==> AC15 ==> WO6 - -In order to add the formulations AC13 and AC14 we need: - -AC10(succ(n)) ==> AC13(n) ==> AC14 ==> AC15 - -or - -AC1 ==> AC13(1); AC13(m) ==> AC13(n) ==> AC14 ==> AC15 (m le n) - -So we don't have to prove all implications of both cases. -Moreover we don't need to prove AC13(1) ==> AC1 and AC11 ==> AC14 as -Rubin & Rubin do. -*) - -(* ********************************************************************** *) -(* Lemmas used in the proofs in which the conclusion is AC13, AC14 *) -(* or AC15 *) -(* - cons_times_nat_not_Finite *) -(* - ex_fun_AC13_AC15 *) -(* ********************************************************************** *) - -Goalw [lepoll_def] "A\\0 ==> B lepoll A*B"; -by (etac not_emptyE 1); -by (res_inst_tac [("x","\\z \\ B. ")] exI 1); -by (fast_tac (claset() addSIs [snd_conv, lam_injective]) 1); -qed "lepoll_Sigma"; - -Goal "0\\A ==> \\B \\ {cons(0,x*nat). x \\ A}. ~Finite(B)"; -by (rtac ballI 1); -by (etac RepFunE 1); -by (hyp_subst_tac 1); -by (rtac notI 1); -by (dresolve_tac [subset_consI RS subset_imp_lepoll RS lepoll_Finite] 1); -by (resolve_tac [lepoll_Sigma RS lepoll_Finite RS (nat_not_Finite RS notE)] 1 - THEN (assume_tac 2)); -by (Fast_tac 1); -qed "cons_times_nat_not_Finite"; - -Goal "[| Union(C)=A; a \\ A |] ==> \\B \\ C. a \\ B & B \\ A"; -by (Fast_tac 1); -val lemma1 = result(); - -Goalw [pairwise_disjoint_def] - "[| pairwise_disjoint(A); B \\ A; C \\ A; a \\ B; a \\ C |] ==> B=C"; -by (dtac IntI 1 THEN (assume_tac 1)); -by (dres_inst_tac [("A","B Int C")] not_emptyI 1); -by (Fast_tac 1); -val lemma2 = result(); - -Goalw [sets_of_size_between_def] - "\\B \\ {cons(0, x*nat). x \\ A}. pairwise_disjoint(f`B) & \ -\ sets_of_size_between(f`B, 2, n) & Union(f`B)=B \ -\ ==> \\B \\ A. \\! u. u \\ f`cons(0, B*nat) & u \\ cons(0, B*nat) & \ -\ 0 \\ u & 2 lepoll u & u lepoll n"; -by (rtac ballI 1); -by (etac ballE 1); -by (Fast_tac 2); -by (REPEAT (etac conjE 1)); -by (dresolve_tac [consI1 RSN (2, lemma1)] 1); -by (etac bexE 1); -by (rtac ex1I 1); -by (Fast_tac 1); -by (REPEAT (etac conjE 1)); -by (rtac lemma2 1 THEN (REPEAT (assume_tac 1))); -val lemma3 = result(); - -Goalw [lepoll_def] "[| A lepoll i; Ord(i) |] ==> {P(a). a \\ A} lepoll i"; -by (etac exE 1); -by (res_inst_tac - [("x", "\\x \\ RepFun(A, P). LEAST j. \\a \\ A. x=P(a) & f`a=j")] exI 1); -by (res_inst_tac [("d", "%y. P(converse(f)`y)")] lam_injective 1); -by (etac RepFunE 1); -by (forward_tac [inj_is_fun RS apply_type] 1 THEN (assume_tac 1)); -by (fast_tac (claset() addIs [LeastI2] - addSEs [Ord_in_Ord, inj_is_fun RS apply_type]) 1); -by (etac RepFunE 1); -by (rtac LeastI2 1); -by (Fast_tac 1); -by (fast_tac (claset() addSEs [Ord_in_Ord, inj_is_fun RS apply_type]) 1); -by (fast_tac (claset() addEs [sym, left_inverse RS ssubst]) 1); -val lemma4 = result(); - -Goal "[| n \\ nat; B \\ A; u(B) \\ cons(0, B*nat); 0 \\ u(B); 2 lepoll u(B); \ -\ u(B) lepoll succ(n) |] \ -\ ==> (\\x \\ A. {fst(x). x \\ u(x)-{0}})`B \\ 0 & \ -\ (\\x \\ A. {fst(x). x \\ u(x)-{0}})`B \\ B & \ -\ (\\x \\ A. {fst(x). x \\ u(x)-{0}})`B lepoll n"; -by (Asm_simp_tac 1); -by (rtac conjI 1); -by (fast_tac (empty_cs addSDs [RepFun_eq_0_iff RS iffD1] - addDs [lepoll_Diff_sing] - addEs [lepoll_trans RS succ_lepoll_natE, ssubst] - addSIs [notI, lepoll_refl, nat_0I]) 1); -by (rtac conjI 1); -by (fast_tac (claset() addSIs [fst_type] addSEs [consE]) 1); -by (fast_tac (claset() addSEs [equalityE, - Diff_lepoll RS (nat_into_Ord RSN (2, lemma4))]) 1); -val lemma5 = result(); - -Goal "[| \\f. \\B \\ {cons(0, x*nat). x \\ A}. \ -\ pairwise_disjoint(f`B) & \ -\ sets_of_size_between(f`B, 2, succ(n)) & \ -\ Union(f`B)=B; n \\ nat |] \ -\ ==> \\f. \\B \\ A. f`B \\ 0 & f`B \\ B & f`B lepoll n"; -by (fast_tac (empty_cs addSDs [lemma3, theI] addDs [bspec] - addSEs [exE, conjE] - addIs [exI, ballI, lemma5]) 1); -qed "ex_fun_AC13_AC15"; - -(* ********************************************************************** *) -(* The target proofs *) -(* ********************************************************************** *) - -(* ********************************************************************** *) -(* AC10(n) ==> AC11 *) -(* ********************************************************************** *) - -Goalw AC_defs "[| n \\ nat; 1 le n; AC10(n) |] ==> AC11"; -by (rtac bexI 1 THEN (assume_tac 2)); -by (Fast_tac 1); -qed "AC10_AC11"; - -(* ********************************************************************** *) -(* AC11 ==> AC12 *) -(* ********************************************************************** *) - -Goalw AC_defs "AC11 ==> AC12"; -by (fast_tac (FOL_cs addSEs [bexE] addIs [bexI]) 1); -qed "AC11_AC12"; - -(* ********************************************************************** *) -(* AC12 ==> AC15 *) -(* ********************************************************************** *) - -Goalw AC_defs "AC12 ==> AC15"; -by Safe_tac; -by (etac allE 1); -by (etac impE 1); -by (etac cons_times_nat_not_Finite 1); -by (fast_tac (claset() addSIs [ex_fun_AC13_AC15]) 1); -qed "AC12_AC15"; - -(* ********************************************************************** *) -(* AC15 ==> WO6 *) -(* ********************************************************************** *) - -Goal "Ord(x) ==> (\\aa \\ x. F(a))"; -by (fast_tac (claset() addSIs [ltI] addSDs [ltD]) 1); -qed "OUN_eq_UN"; - -val [prem] = goal thy "\\x \\ Pow(A)-{0}. f`x\\0 & f`x \\ x & f`x lepoll m ==> \ -\ (\\ix \\ Pow(A)-{0}. f`x\\0 & f`x \\ x & f`x lepoll m ==> \ -\ \\x WO6"; -by (rtac allI 1); -by (eres_inst_tac [("x","Pow(A)-{0}")] allE 1); -by (etac impE 1); -by (Fast_tac 1); -by (REPEAT (eresolve_tac [bexE,conjE,exE] 1)); -by (rtac bexI 1 THEN (assume_tac 2)); -by (rtac conjI 1 THEN (assume_tac 1)); -by (res_inst_tac [("x","LEAST i. HH(f,A,i)={A}")] exI 1); -by (res_inst_tac [("x","\\j \\ (LEAST i. HH(f,A,i)={A}). HH(f,A,j)")] exI 1); -by (Asm_full_simp_tac 1); -by (fast_tac (claset() addSIs [Ord_Least, lam_type RS domain_of_fun] - addSEs [less_Least_subset_x, lemma1, lemma2]) 1); -qed "AC15_WO6"; - - -(* ********************************************************************** *) -(* The proof needed in the first case, not in the second *) -(* ********************************************************************** *) - -(* ********************************************************************** *) -(* AC10(n) ==> AC13(n-1) if 2 le n *) -(* *) -(* Because of the change to the formal definition of AC10(n) we prove *) -(* the following obviously equivalent theorem \\ *) -(* AC10(n) implies AC13(n) for (1 le n) *) -(* ********************************************************************** *) - -Goalw AC_defs "[| n \\ nat; 1 le n; AC10(n) |] ==> AC13(n)"; -by Safe_tac; -by (fast_tac (empty_cs addSEs [allE, cons_times_nat_not_Finite RSN (2, impE), - ex_fun_AC13_AC15]) 1); -qed "AC10_AC13"; - -(* ********************************************************************** *) -(* The proofs needed in the second case, not in the first *) -(* ********************************************************************** *) - -(* ********************************************************************** *) -(* AC1 ==> AC13(1) *) -(* ********************************************************************** *) - -Goalw AC_defs "AC1 ==> AC13(1)"; -by (rtac allI 1); -by (etac allE 1); -by (rtac impI 1); -by (mp_tac 1); -by (etac exE 1); -by (res_inst_tac [("x","\\x \\ A. {f`x}")] exI 1); -by (asm_simp_tac (simpset() addsimps - [singleton_eqpoll_1 RS eqpoll_imp_lepoll, - singletonI RS not_emptyI]) 1); -qed "AC1_AC13"; - -(* ********************************************************************** *) -(* AC13(m) ==> AC13(n) for m \\ n *) -(* ********************************************************************** *) - -Goalw AC_defs "[| m le n; AC13(m) |] ==> AC13(n)"; -by (dtac le_imp_lepoll 1); -by (fast_tac (claset() addSEs [lepoll_trans]) 1); -qed "AC13_mono"; - -(* ********************************************************************** *) -(* The proofs necessary for both cases *) -(* ********************************************************************** *) - -(* ********************************************************************** *) -(* AC13(n) ==> AC14 if 1 \\ n *) -(* ********************************************************************** *) - -Goalw AC_defs "[| n \\ nat; 1 le n; AC13(n) |] ==> AC14"; -by (fast_tac (FOL_cs addIs [bexI]) 1); -qed "AC13_AC14"; - -(* ********************************************************************** *) -(* AC14 ==> AC15 *) -(* ********************************************************************** *) - -Goalw AC_defs "AC14 ==> AC15"; -by (Fast_tac 1); -qed "AC14_AC15"; - -(* ********************************************************************** *) -(* The redundant proofs; however cited by Rubin & Rubin *) -(* ********************************************************************** *) - -(* ********************************************************************** *) -(* AC13(1) ==> AC1 *) -(* ********************************************************************** *) - -Goal "[| A\\0; A lepoll 1 |] ==> \\a. A={a}"; -by (fast_tac (claset() addSEs [not_emptyE, lepoll_1_is_sing]) 1); -qed "lemma_aux"; - -Goal "\\B \\ A. f(B)\\0 & f(B)<=B & f(B) lepoll 1 \ -\ ==> (\\x \\ A. THE y. f(x)={y}) \\ (\\X \\ A. X)"; -by (rtac lam_type 1); -by (dtac bspec 1 THEN (assume_tac 1)); -by (REPEAT (etac conjE 1)); -by (eresolve_tac [lemma_aux RS exE] 1 THEN (assume_tac 1)); -by (asm_full_simp_tac (simpset() addsimps [the_element]) 1); -val lemma = result(); - -Goalw AC_defs "AC13(1) ==> AC1"; -by (fast_tac (claset() addSEs [lemma]) 1); -qed "AC13_AC1"; - -(* ********************************************************************** *) -(* AC11 ==> AC14 *) -(* ********************************************************************** *) - -Goalw [AC11_def, AC14_def] "AC11 ==> AC14"; -by (fast_tac (claset() addSIs [AC10_AC13]) 1); -qed "AC11_AC14"; diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC15_WO6.thy --- a/src/ZF/AC/AC15_WO6.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/AC15_WO6.thy Wed Jan 16 17:52:06 2002 +0100 @@ -1,3 +1,290 @@ -(*Dummy theory to document dependencies *) +(* Title: ZF/AC/AC15_WO6.thy + ID: $Id$ + Author: Krzysztof Grabczewski + +The proofs needed to state that AC10, ..., AC15 are equivalent to the rest. +We need the following: + +WO1 ==> AC10(n) ==> AC11 ==> AC12 ==> AC15 ==> WO6 + +In order to add the formulations AC13 and AC14 we need: + +AC10(succ(n)) ==> AC13(n) ==> AC14 ==> AC15 + +or + +AC1 ==> AC13(1); AC13(m) ==> AC13(n) ==> AC14 ==> AC15 (m\n) + +So we don't have to prove all implications of both cases. +Moreover we don't need to prove AC13(1) ==> AC1 and AC11 ==> AC14 as +Rubin & Rubin do. +*) + +theory AC15_WO6 = HH + Cardinal_aux: + + +(* ********************************************************************** *) +(* Lemmas used in the proofs in which the conclusion is AC13, AC14 *) +(* or AC15 *) +(* - cons_times_nat_not_Finite *) +(* - ex_fun_AC13_AC15 *) +(* ********************************************************************** *) + +lemma lepoll_Sigma: "A\0 ==> B \ A*B" +apply (unfold lepoll_def) +apply (erule not_emptyE) +apply (rule_tac x = "\z \ B. " in exI) +apply (fast intro!: snd_conv lam_injective) +done + +lemma cons_times_nat_not_Finite: + "0\A ==> \B \ {cons(0,x*nat). x \ A}. ~Finite(B)" +apply clarify +apply (drule subset_consI [THEN subset_imp_lepoll, THEN lepoll_Finite]) +apply (rule nat_not_Finite [THEN notE] ) +apply (subgoal_tac "x ~= 0") +apply (blast intro: lepoll_Sigma [THEN lepoll_Finite] , blast) +done + +lemma lemma1: "[| Union(C)=A; a \ A |] ==> \B \ C. a \ B & B \ A" +by fast + +lemma lemma2: + "[| pairwise_disjoint(A); B \ A; C \ A; a \ B; a \ C |] ==> B=C" +by (unfold pairwise_disjoint_def, blast) + +lemma lemma3: + "\B \ {cons(0, x*nat). x \ A}. pairwise_disjoint(f`B) & + sets_of_size_between(f`B, 2, n) & Union(f`B)=B + ==> \B \ A. \! u. u \ f`cons(0, B*nat) & u \ cons(0, B*nat) & + 0 \ u & 2 \ u & u \ n" +apply (unfold sets_of_size_between_def) +apply (rule ballI) +apply (erule ballE) +prefer 2 apply blast +apply (blast dest: lemma1 intro!: lemma2) +done + +lemma lemma4: "[| A \ i; Ord(i) |] ==> {P(a). a \ A} \ i" +apply (unfold lepoll_def) +apply (erule exE) +apply (rule_tac x = "\x \ RepFun(A,P). LEAST j. \a\A. x=P(a) & f`a=j" + in exI) +apply (rule_tac d = "%y. P (converse (f) `y) " in lam_injective) +apply (erule RepFunE) +apply (frule inj_is_fun [THEN apply_type], assumption) +apply (fast intro: LeastI2 elim!: Ord_in_Ord inj_is_fun [THEN apply_type]) +apply (erule RepFunE) +apply (rule LeastI2) + apply fast + apply (fast elim!: Ord_in_Ord inj_is_fun [THEN apply_type]) +apply (fast elim: sym left_inverse [THEN ssubst]) +done + +lemma lemma5_1: + "[| B \ A; 2 \ u(B) |] ==> (\x \ A. {fst(x). x \ u(x)-{0}})`B \ 0" +apply simp +apply (fast dest: lepoll_Diff_sing + elim: lepoll_trans [THEN succ_lepoll_natE] ssubst + intro!: lepoll_refl) +done + +lemma lemma5_2: + "[| B \ A; u(B) \ cons(0, B*nat) |] + ==> (\x \ A. {fst(x). x \ u(x)-{0}})`B \ B" +apply auto +done + +lemma lemma5_3: + "[| n \ nat; B \ A; 0 \ u(B); u(B) \ succ(n) |] + ==> (\x \ A. {fst(x). x \ u(x)-{0}})`B \ n" +apply simp +apply (fast elim!: Diff_lepoll [THEN lemma4 [OF _ nat_into_Ord]]) +done + +lemma ex_fun_AC13_AC15: + "[| \B \ {cons(0, x*nat). x \ A}. + pairwise_disjoint(f`B) & + sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B; + n \ nat |] + ==> \f. \B \ A. f`B \ 0 & f`B \ B & f`B \ n" +by (fast del: subsetI notI + dest!: lemma3 theI intro!: lemma5_1 lemma5_2 lemma5_3) + + +(* ********************************************************************** *) +(* The target proofs *) +(* ********************************************************************** *) + +(* ********************************************************************** *) +(* AC10(n) ==> AC11 *) +(* ********************************************************************** *) + +lemma AC10_AC11: "[| n \ nat; 1\n; AC10(n) |] ==> AC11" +by (unfold AC10_def AC11_def, blast) + +(* ********************************************************************** *) +(* AC11 ==> AC12 *) +(* ********************************************************************** *) + +lemma AC11_AC12: "AC11 ==> AC12" +by (unfold AC10_def AC11_def AC11_def AC12_def, blast) + +(* ********************************************************************** *) +(* AC12 ==> AC15 *) +(* ********************************************************************** *) + +lemma AC12_AC15: "AC12 ==> AC15" +apply (unfold AC12_def AC15_def) +apply (blast del: ballI + intro!: cons_times_nat_not_Finite ex_fun_AC13_AC15) +done -AC15_WO6 = HH +(* ********************************************************************** *) +(* AC15 ==> WO6 *) +(* ********************************************************************** *) + +lemma OUN_eq_UN: "Ord(x) ==> (\aa \ x. F(a))" +by (fast intro!: ltI dest!: ltD) + +lemma lemma1: + "\x \ Pow(A)-{0}. f`x\0 & f`x \ x & f`x \ m + ==> (\ix \ Pow(A)-{0}. f`x\0 & f`x \ x & f`x \ m + ==> \x < (LEAST x. HH(f,A,x)={A}). HH(f,A,x) \ m" +apply (rule oallI) +apply (drule ltD [THEN less_Least_subset_x]) +apply (frule HH_subset_imp_eq) +apply (erule ssubst) +apply (blast dest!: HH_subset_x_imp_subset_Diff_UN [THEN not_emptyI2]) + (*but can't use del: DiffE despite the obvious conflictc*) +done + +lemma AC15_WO6: "AC15 ==> WO6" +apply (unfold AC15_def WO6_def) +apply (rule allI) +apply (erule_tac x = "Pow (A) -{0}" in allE) +apply (erule impE, fast) +apply (elim bexE conjE exE) +apply (rule bexI) +apply (rule conjI, assumption) +apply (rule_tac x = "LEAST i. HH (f,A,i) ={A}" in exI) +apply (rule_tac x = "\j \ (LEAST i. HH (f,A,i) ={A}) . HH (f,A,j) " in exI) +apply simp +apply (fast intro!: Ord_Least lam_type [THEN domain_of_fun] + elim!: less_Least_subset_x lemma1 lemma2, assumption); +done + + +(* ********************************************************************** *) +(* The proof needed in the first case, not in the second *) +(* ********************************************************************** *) + +(* ********************************************************************** *) +(* AC10(n) ==> AC13(n-1) if 2\n *) +(* *) +(* Because of the change to the formal definition of AC10(n) we prove *) +(* the following obviously equivalent theorem \ *) +(* AC10(n) implies AC13(n) for (1\n) *) +(* ********************************************************************** *) + +lemma AC10_AC13: "[| n \ nat; 1\n; AC10(n) |] ==> AC13(n)" +apply (unfold AC10_def AC13_def, safe) +apply (erule allE) +apply (erule impE [OF _ cons_times_nat_not_Finite], assumption); +apply (fast elim!: impE [OF _ cons_times_nat_not_Finite] + dest!: ex_fun_AC13_AC15) +done + +(* ********************************************************************** *) +(* The proofs needed in the second case, not in the first *) +(* ********************************************************************** *) + +(* ********************************************************************** *) +(* AC1 ==> AC13(1) *) +(* ********************************************************************** *) + +lemma AC1_AC13: "AC1 ==> AC13(1)" +apply (unfold AC1_def AC13_def) +apply (rule allI) +apply (erule allE) +apply (rule impI) +apply (drule mp, assumption) +apply (elim exE) +apply (rule_tac x = "\x \ A. {f`x}" in exI) +apply (simp add: singleton_eqpoll_1 [THEN eqpoll_imp_lepoll]) +done + +(* ********************************************************************** *) +(* AC13(m) ==> AC13(n) for m \ n *) +(* ********************************************************************** *) + +lemma AC13_mono: "[| m\n; AC13(m) |] ==> AC13(n)" +apply (unfold AC13_def) +apply (drule le_imp_lepoll) +apply (fast elim!: lepoll_trans) +done + +(* ********************************************************************** *) +(* The proofs necessary for both cases *) +(* ********************************************************************** *) + +(* ********************************************************************** *) +(* AC13(n) ==> AC14 if 1 \ n *) +(* ********************************************************************** *) + +lemma AC13_AC14: "[| n \ nat; 1\n; AC13(n) |] ==> AC14" +by (unfold AC13_def AC14_def, auto) + +(* ********************************************************************** *) +(* AC14 ==> AC15 *) +(* ********************************************************************** *) + +lemma AC14_AC15: "AC14 ==> AC15" +by (unfold AC13_def AC14_def AC15_def, fast) + +(* ********************************************************************** *) +(* The redundant proofs; however cited by Rubin & Rubin *) +(* ********************************************************************** *) + +(* ********************************************************************** *) +(* AC13(1) ==> AC1 *) +(* ********************************************************************** *) + +lemma lemma_aux: "[| A\0; A \ 1 |] ==> \a. A={a}" +by (fast elim!: not_emptyE lepoll_1_is_sing) + +lemma AC13_AC1_lemma: + "\B \ A. f(B)\0 & f(B)<=B & f(B) \ 1 + ==> (\x \ A. THE y. f(x)={y}) \ (\X \ A. X)" +apply (rule lam_type) +apply (drule bspec, assumption) +apply (elim conjE) +apply (erule lemma_aux [THEN exE], assumption) +apply (simp add: the_element) +done + +lemma AC13_AC1: "AC13(1) ==> AC1" +apply (unfold AC13_def AC1_def) +apply (fast elim!: AC13_AC1_lemma) +done + +(* ********************************************************************** *) +(* AC11 ==> AC14 *) +(* ********************************************************************** *) + +lemma AC11_AC14: "AC11 ==> AC14" +apply (unfold AC11_def AC14_def) +apply (fast intro!: AC10_AC13) +done + +end + diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC16_WO4.ML --- a/src/ZF/AC/AC16_WO4.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,611 +0,0 @@ -(* Title: ZF/AC/AC16_WO4.ML - ID: $Id$ - Author: Krzysztof Grabczewski - - The proof of AC16(n, k) ==> WO4(n-k) -*) - -(* ********************************************************************** *) -(* The case of finite set *) -(* ********************************************************************** *) - -Goalw [Finite_def] "[| Finite(A); 0 nat |] ==> \ -\ \\a f. Ord(a) & domain(f) = a & \ -\ (\\bbi \\ n. {f`i}")] exI 1); -by (Asm_full_simp_tac 1); -by (rewrite_goals_tac [bij_def, surj_def]); -by (fast_tac (claset() addSIs [ltI, nat_into_Ord, lam_funtype RS domain_of_fun, - equalityI, singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans, - nat_1_lepoll_iff RS iffD2] - addSEs [apply_type, ltE]) 1); -val lemma1 = result(); - -(* ********************************************************************** *) -(* The case of infinite set *) -(* ********************************************************************** *) - -(* well_ord(x,r) ==> well_ord({{y,z}. y \\ x}, Something(x,z)) **) -bind_thm ("well_ord_paired", (paired_bij RS bij_is_inj RS well_ord_rvimage)); - -Goal "[| A lepoll B; ~ A lepoll C |] ==> ~ B lepoll C"; -by (fast_tac (claset() addEs [notE, lepoll_trans]) 1); -qed "lepoll_trans1"; - -(* ********************************************************************** *) -(* There exists a well ordered set y such that ... *) -(* ********************************************************************** *) - -val lepoll_paired = paired_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll; - -Goal "\\y R. well_ord(y,R) & x Int y = 0 & ~y lepoll z & ~Finite(y)"; -by (res_inst_tac [("x","{{a,x}. a \\ nat Un Hartog(z)}")] exI 1); -by (resolve_tac [transfer thy Ord_nat RS well_ord_Memrel RS - (Ord_Hartog RS - well_ord_Memrel RSN (2, well_ord_Un)) RS exE] 1); -by (fast_tac - (claset() addSIs [Ord_Hartog, well_ord_Memrel, well_ord_paired, - HartogI RSN (2, lepoll_trans1), - subset_imp_lepoll RS (lepoll_paired RSN (2, lepoll_trans))] - addSEs [nat_not_Finite RS notE] addEs [mem_asym] - addSDs [Un_upper1 RS subset_imp_lepoll RS lepoll_Finite, - lepoll_paired RS lepoll_Finite]) 1); -val lemma2 = result(); - -Goal "~Finite(B) ==> ~Finite(A Un B)"; -by (blast_tac (claset() addIs [subset_Finite]) 1); -qed "infinite_Un"; - -(* ********************************************************************** *) -(* There is a v \\ s(u) such that k lepoll x Int y (in our case succ(k)) *) -(* The idea of the proof is the following \\ *) -(* Suppose not, i.e. every element of s(u) has exactly k-1 elements of y *) -(* Thence y is less than or equipollent to {v \\ Pow(x). v eqpoll n#-k} *) -(* We have obtained this result in two steps \\ *) -(* 1. y is less than or equipollent to {v \\ s(u). a \\ v} *) -(* where a is certain k-2 element subset of y *) -(* 2. {v \\ s(u). a \\ v} is less than or equipollent *) -(* to {v \\ Pow(x). v eqpoll n-k} *) -(* ********************************************************************** *) - -(*Proof simplified by LCP*) -Goal "[| ~(\\x \\ A. f`x=y); f \\ inj(A, B); y \\ B |] \ -\ ==> (\\a \\ succ(A). if(a=A, y, f`a)) \\ inj(succ(A), B)"; -by (res_inst_tac [("d","%z. if(z=y, A, converse(f)`z)")] lam_injective 1); -by (force_tac (claset(), simpset() addsimps [inj_is_fun RS apply_type]) 1); -(*this preliminary simplification prevents looping somehow*) -by (Asm_simp_tac 1); -by (force_tac (claset(), simpset() addsimps []) 1); -qed "succ_not_lepoll_lemma"; - -Goalw [lepoll_def, eqpoll_def, bij_def, surj_def] - "[| ~A eqpoll B; A lepoll B |] ==> succ(A) lepoll B"; -by (fast_tac (claset() addSEs [succ_not_lepoll_lemma, inj_is_fun]) 1); -qed "succ_not_lepoll_imp_eqpoll"; - - -(* ********************************************************************** *) -(* There is a k-2 element subset of y *) -(* ********************************************************************** *) - -val ordertype_eqpoll = - ordermap_bij RS (exI RS (eqpoll_def RS def_imp_iff RS iffD2)); - -Goal "[| a \\ y; b \\ y-a; u \\ x |] ==> cons(b, cons(u, a)) \\ Pow(x Un y)"; -by (Fast_tac 1); -qed "cons_cons_subset"; - -Goal "[| a eqpoll k; a \\ y; b \\ y-a; u \\ x; x Int y = 0 |] \ -\ ==> cons(b, cons(u, a)) eqpoll succ(succ(k))"; -by (fast_tac (claset() addSIs [cons_eqpoll_succ]) 1); -qed "cons_cons_eqpoll"; - -Goal "[| succ(k) eqpoll A; k eqpoll B; B \\ A; a \\ A-B; k \\ nat |] \ -\ ==> A = cons(a, B)"; -by (rtac equalityI 1); -by (Fast_tac 2); -by (resolve_tac [Diff_eq_0_iff RS iffD1] 1); -by (rtac equals0I 1); -by (dresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll] 1); -by (dresolve_tac [eqpoll_sym RS cons_eqpoll_succ] 1); -by (Fast_tac 1); -by (dtac cons_eqpoll_succ 1); -by (Fast_tac 1); -by (fast_tac - (claset() - addSEs [[eqpoll_sym RS eqpoll_imp_lepoll, subset_imp_lepoll] MRS - (lepoll_trans RS lepoll_trans) RS succ_lepoll_natE]) 1); -qed "set_eq_cons"; - -Goal "[| cons(x,a) = cons(y,a); x\\ a |] ==> x = y "; -by (fast_tac (claset() addSEs [equalityE]) 1); -qed "cons_eqE"; - -Goal "A = B ==> A Int C = B Int C"; -by (Asm_simp_tac 1); -qed "eq_imp_Int_eq"; - -(* ********************************************************************** *) -(* some arithmetic *) -(* ********************************************************************** *) - -Goal "[| k \\ nat; m \\ nat |] ==> \ -\ \\A B. A eqpoll k #+ m & k lepoll B & B \\ A --> A-B lepoll m"; -by (induct_tac "k" 1); -by (asm_simp_tac (simpset() addsimps [add_0]) 1); -by (fast_tac (claset() addIs [eqpoll_imp_lepoll RS - (Diff_subset RS subset_imp_lepoll RS lepoll_trans)]) 1); -by (REPEAT (resolve_tac [allI,impI] 1)); -by (resolve_tac [succ_lepoll_imp_not_empty RS not_emptyE] 1); -by (Fast_tac 1); -by (eres_inst_tac [("x","A - {xa}")] allE 1); -by (eres_inst_tac [("x","B - {xa}")] allE 1); -by (etac impE 1); -by (asm_full_simp_tac (simpset() addsimps [add_succ]) 1); -by (fast_tac (claset() addSIs [Diff_sing_eqpoll, lepoll_Diff_sing]) 1); -by (res_inst_tac [("P","%z. z lepoll m")] subst 1 THEN (assume_tac 2)); -by (Fast_tac 1); -qed "eqpoll_sum_imp_Diff_lepoll_lemma"; - -Goal "[| A eqpoll succ(k #+ m); B \\ A; succ(k) lepoll B; k \\ nat; m \\ nat |] \ -\ ==> A-B lepoll m"; -by (dresolve_tac [add_succ RS ssubst] 1); -by (dresolve_tac [nat_succI RS eqpoll_sum_imp_Diff_lepoll_lemma] 1 - THEN (REPEAT (assume_tac 1))); -by (Fast_tac 1); -qed "eqpoll_sum_imp_Diff_lepoll"; - -(* ********************************************************************** *) -(* similar properties for eqpoll *) -(* ********************************************************************** *) - -Goal "[| k \\ nat; m \\ nat |] ==> \ -\ \\A B. A eqpoll k #+ m & k eqpoll B & B \\ A --> A-B eqpoll m"; -by (induct_tac "k" 1); -by (fast_tac (claset() addSDs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_0_is_0] - addss (simpset() addsimps [add_0])) 1); -by (REPEAT (resolve_tac [allI,impI] 1)); -by (resolve_tac [succ_lepoll_imp_not_empty RS not_emptyE] 1); -by (fast_tac (claset() addSEs [eqpoll_imp_lepoll]) 1); -by (eres_inst_tac [("x","A - {xa}")] allE 1); -by (eres_inst_tac [("x","B - {xa}")] allE 1); -by (etac impE 1); -by (fast_tac (claset() addSIs [Diff_sing_eqpoll, - eqpoll_sym RSN (2, Diff_sing_eqpoll) RS eqpoll_sym] - addss (simpset() addsimps [add_succ])) 1); -by (res_inst_tac [("P","%z. z eqpoll m")] subst 1 THEN (assume_tac 2)); -by (Fast_tac 1); -qed "eqpoll_sum_imp_Diff_eqpoll_lemma"; - -Goal "[| A eqpoll succ(k #+ m); B \\ A; succ(k) eqpoll B; k \\ nat; m \\ nat |] \ -\ ==> A-B eqpoll m"; -by (dresolve_tac [add_succ RS ssubst] 1); -by (dresolve_tac [nat_succI RS eqpoll_sum_imp_Diff_eqpoll_lemma] 1 - THEN (REPEAT (assume_tac 1))); -by (Fast_tac 1); -qed "eqpoll_sum_imp_Diff_eqpoll"; - - -(* ********************************************************************** *) -(* LL can be well ordered *) -(* ********************************************************************** *) - -Goal "{x \\ Pow(X). x lepoll 0} = {0}"; -by (fast_tac (claset() addSDs [lepoll_0_is_0] addSIs [lepoll_refl]) 1); -qed "subsets_lepoll_0_eq_unit"; - -Goal "n \\ nat ==> {z \\ Pow(y). z lepoll succ(n)} = \ -\ {z \\ Pow(y). z lepoll n} Un {z \\ Pow(y). z eqpoll succ(n)}"; -by (fast_tac (claset() addIs [le_refl, leI, le_imp_lepoll] - addSDs [lepoll_succ_disj] - addSEs [nat_into_Ord, lepoll_trans, eqpoll_imp_lepoll]) 1); -qed "subsets_lepoll_succ"; - -Goal "n \\ nat ==> {z \\ Pow(y). z lepoll n} Int {z \\ Pow(y). z eqpoll succ(n)} = 0"; -by (fast_tac (claset() addSEs [eqpoll_sym RS eqpoll_imp_lepoll - RS lepoll_trans RS succ_lepoll_natE] - addSIs [equals0I]) 1); -qed "Int_empty"; - - -Open_locale "AC16"; - -val all_ex = thm "all_ex"; -val disjoint = thm "disjoint"; -val includes = thm "includes"; -val WO_R = thm "WO_R"; -val k_def = thm "k_def"; -val lnat = thm "lnat"; -val mnat = thm "mnat"; -val mpos = thm "mpos"; -val Infinite = thm "Infinite"; -val noLepoll = thm "noLepoll"; - -val LL_def = thm "LL_def"; -val MM_def = thm "MM_def"; -val GG_def = thm "GG_def"; -val s_def = thm "s_def"; - -Addsimps [disjoint, WO_R, lnat, mnat, mpos, Infinite]; -AddSIs [disjoint, WO_R, lnat, mnat, mpos]; - -Goalw [k_def] "k \\ nat"; -by Auto_tac; -qed "knat"; -Addsimps [knat]; AddSIs [knat]; - -AddSIs [Infinite]; (*if notI is removed!*) -AddSEs [Infinite RS notE]; - -AddEs [[disjoint, IntI] MRS (equals0D RS notE)]; - -(*use k = succ(l) *) -val includes_l = simplify (FOL_ss addsimps [k_def]) includes; -val all_ex_l = simplify (FOL_ss addsimps [k_def]) all_ex; - -(* ********************************************************************** *) -(* 1. y is less than or equipollent to {v \\ s(u). a \\ v} *) -(* where a is certain k-2 element subset of y *) -(* ********************************************************************** *) - -Goal "[| l eqpoll a; a \\ y |] ==> y - a eqpoll y"; -by (cut_facts_tac [WO_R, Infinite, lnat] 1); -by (fast_tac (empty_cs addIs [lesspoll_trans1] - addSIs [Card_cardinal, Diff_lesspoll_eqpoll_Card RS eqpoll_trans, - Card_cardinal RS Card_is_Ord RS nat_le_infinite_Ord - RS le_imp_lepoll] - addSEs [well_ord_cardinal_eqpoll, - well_ord_cardinal_eqpoll RS eqpoll_sym, - eqpoll_sym RS eqpoll_imp_lepoll, - n_lesspoll_nat RS lesspoll_trans2, - well_ord_cardinal_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll - RS lepoll_infinite]) 1); -qed "Diff_Finite_eqpoll"; - -Goalw [s_def] "s(u) \\ t_n"; -by (Fast_tac 1); -qed "s_subset"; - -Goalw [s_def, succ_def, k_def] - "[| w \\ t_n; cons(b,cons(u,a)) \\ w; a \\ y; b \\ y-a; l eqpoll a \ -\ |] ==> w \\ s(u)"; -by (fast_tac (claset() addDs [eqpoll_imp_lepoll RS cons_lepoll_cong] - addSEs [subset_imp_lepoll RSN (2, lepoll_trans)]) 1); -qed "sI"; - -Goalw [s_def] "v \\ s(u) ==> u \\ v"; -by (Fast_tac 1); -qed "in_s_imp_u_in"; - - -Goal "[| l eqpoll a; a \\ y; b \\ y - a; u \\ x |] \ -\ ==> \\! c. c \\ s(u) & a \\ c & b \\ c"; -by (rtac (all_ex_l RS ballE) 1); -by (blast_tac (claset() delrules [PowI] - addSIs [cons_cons_subset, - eqpoll_sym RS cons_cons_eqpoll]) 2); -by (etac ex1E 1); -by (res_inst_tac [("a","w")] ex1I 1); -by (blast_tac (claset() addIs [sI]) 1); -by (etac allE 1); -by (etac impE 1); -by (assume_tac 2); -by (fast_tac (claset() addSEs [s_subset RS subsetD, in_s_imp_u_in]) 1); -qed "ex1_superset_a"; - -Goal "[| \\v \\ s(u). succ(l) eqpoll v Int y; \ -\ l eqpoll a; a \\ y; b \\ y - a; u \\ x |] \ -\ ==> (THE c. c \\ s(u) & a \\ c & b \\ c) \ -\ Int y = cons(b, a)"; -by (forward_tac [ex1_superset_a RS theI] 1 THEN REPEAT (assume_tac 1)); -by (rtac set_eq_cons 1); -by (REPEAT (Fast_tac 1)); -qed "the_eq_cons"; - -Goal "[| \\v \\ s(u). succ(l) eqpoll v Int y; \ -\ l eqpoll a; a \\ y; u \\ x |] \ -\ ==> y lepoll {v \\ s(u). a \\ v}"; -by (resolve_tac [Diff_Finite_eqpoll RS eqpoll_sym RS - eqpoll_imp_lepoll RS lepoll_trans] 1 - THEN REPEAT (Fast_tac 1)); -by (res_inst_tac - [("f3", "\\b \\ y-a. THE c. c \\ s(u) & a \\ c & b \\ c")] - (exI RS (lepoll_def RS def_imp_iff RS iffD2)) 1); -by (simp_tac (simpset() addsimps [inj_def]) 1); -by (rtac conjI 1); -by (rtac lam_type 1); -by (forward_tac [ex1_superset_a RS theI] 1 THEN REPEAT (Fast_tac 1)); -by (Asm_simp_tac 1); -by (Clarify_tac 1); -by (rtac cons_eqE 1); -by (Fast_tac 2); -by (dres_inst_tac [("A","THE c. ?P(c)"), ("C","y")] eq_imp_Int_eq 1); -by (asm_full_simp_tac (simpset() addsimps [the_eq_cons]) 1); -qed "y_lepoll_subset_s"; - - -(* ********************************************************************** *) -(* back to the second part *) -(* ********************************************************************** *) - -Goal "w \\ x Un y ==> w Int (x - {u}) = w - cons(u, w Int y)"; -by (Fast_tac 1); -qed "w_Int_eq_w_Diff"; - -Goal "[| w \\ {v \\ s(u). a \\ v}; \ -\ l eqpoll a; u \\ x; \ -\ \\v \\ s(u). succ(l) eqpoll v Int y \ -\ |] ==> w Int (x - {u}) eqpoll m"; -by (etac CollectE 1); -by (stac w_Int_eq_w_Diff 1); -by (fast_tac (claset() addSDs [s_subset RS subsetD, - includes_l RS subsetD]) 1); -by (fast_tac (claset() addSDs [bspec] - addDs [s_subset RS subsetD, includes_l RS subsetD] - addSEs [eqpoll_sym RS cons_eqpoll_succ RS eqpoll_sym, in_s_imp_u_in] - addSIs [eqpoll_sum_imp_Diff_eqpoll]) 1); -qed "w_Int_eqpoll_m"; - -(* ********************************************************************** *) -(* 2. {v \\ s(u). a \\ v} is less than or equipollent *) -(* to {v \\ Pow(x). v eqpoll n-k} *) -(* ********************************************************************** *) - -Goal "x eqpoll m ==> x \\ 0"; -by (cut_facts_tac [mpos] 1); -by (fast_tac (claset() addSEs [zero_lt_natE] - addSDs [eqpoll_succ_imp_not_empty]) 1); -qed "eqpoll_m_not_empty"; - -Goal "[| z \\ xa Int (x - {u}); l eqpoll a; a \\ y; u \\ x |] \ -\ ==> \\! w. w \\ t_n & cons(z, cons(u, a)) \\ w"; -by (rtac (all_ex RS bspec) 1); -by (rewtac k_def); -by (fast_tac (claset() addSIs [cons_eqpoll_succ] addEs [eqpoll_sym]) 1); -qed "cons_cons_in"; - -Goal "[| \\v \\ s(u). succ(l) eqpoll v Int y; \ -\ a \\ y; l eqpoll a; u \\ x |] \ -\ ==> {v \\ s(u). a \\ v} lepoll {v \\ Pow(x). v eqpoll m}"; -by (res_inst_tac [("f3","\\w \\ {v \\ s(u). a \\ v}. w Int (x - {u})")] - (exI RS (lepoll_def RS def_imp_iff RS iffD2)) 1); -by (simp_tac (simpset() addsimps [inj_def]) 1); -by (rtac conjI 1); -by (rtac lam_type 1); -by (rtac CollectI 1); -by (Fast_tac 1); -by (rtac w_Int_eqpoll_m 1 THEN REPEAT (assume_tac 1)); -by (REPEAT (resolve_tac [ballI, impI] 1)); -(** LEVEL 8 **) -by (resolve_tac [w_Int_eqpoll_m RS eqpoll_m_not_empty RS not_emptyE] 1); -by (EVERY (map Blast_tac [4,3,2,1])); -by (dresolve_tac [equalityD1 RS subsetD] 1 THEN (assume_tac 1)); -by (ftac cons_cons_in 1 THEN REPEAT (assume_tac 1)); -by (etac ex1_two_eq 1); -by (REPEAT (blast_tac - (claset() addDs [s_subset RS subsetD, in_s_imp_u_in]) 1)); -qed "subset_s_lepoll_w"; - - -(* ********************************************************************** *) -(* well_ord_subsets_lepoll_n *) -(* ********************************************************************** *) - -Goal "n \\ nat ==> \\S. well_ord({z \\ Pow(y) . z eqpoll succ(n)}, S)"; -by (resolve_tac [WO_R RS well_ord_infinite_subsets_eqpoll_X - RS (eqpoll_def RS def_imp_iff RS iffD1) RS exE] 1); -by (REPEAT (fast_tac (claset() addIs [bij_is_inj RS well_ord_rvimage]) 1)); -qed "well_ord_subsets_eqpoll_n"; - -Goal "n \\ nat ==> \\R. well_ord({z \\ Pow(y). z lepoll n}, R)"; -by (induct_tac "n" 1); -by (force_tac (claset() addSIs [well_ord_unit], - simpset() addsimps [subsets_lepoll_0_eq_unit]) 1); -by (etac exE 1); -by (resolve_tac [well_ord_subsets_eqpoll_n RS exE] 1 THEN assume_tac 1); -by (asm_simp_tac (simpset() addsimps [subsets_lepoll_succ]) 1); -by (dtac well_ord_radd 1 THEN (assume_tac 1)); -by (eresolve_tac [Int_empty RS disj_Un_eqpoll_sum RS - (eqpoll_def RS def_imp_iff RS iffD1) RS exE] 1); -by (fast_tac (claset() addSEs [bij_is_inj RS well_ord_rvimage]) 1); -qed "well_ord_subsets_lepoll_n"; - - -Goalw [LL_def, MM_def] "LL \\ {z \\ Pow(y). z lepoll succ(k #+ m)}"; -by (cut_facts_tac [includes] 1); -by (fast_tac (claset() addIs [subset_imp_lepoll - RS (eqpoll_imp_lepoll - RSN (2, lepoll_trans))]) 1); -qed "LL_subset"; - -Goal "\\S. well_ord(LL,S)"; -by (rtac (well_ord_subsets_lepoll_n RS exE) 1); -by (blast_tac (claset() addIs [LL_subset RSN (2, well_ord_subset)]) 2); -by Auto_tac; -qed "well_ord_LL"; - -(* ********************************************************************** *) -(* every element of LL is a contained in exactly one element of MM *) -(* ********************************************************************** *) - -Goalw [MM_def, LL_def] "v \\ LL ==> \\! w. w \\ MM & v \\ w"; -by Safe_tac; -by (Fast_tac 1); -by (resolve_tac [lepoll_imp_eqpoll_subset RS exE] 1 THEN (assume_tac 1)); -by (res_inst_tac [("x","x")] (all_ex RS ballE) 1); -by (fast_tac (claset() addSEs [eqpoll_sym]) 2); -by (Blast_tac 1); -qed "unique_superset_in_MM"; - -val unique_superset1 = unique_superset_in_MM RS theI RS conjunct1; -val unique_superset2 = unique_superset_in_MM RS the_equality2; - - -(* ********************************************************************** *) -(* The function GG satisfies the conditions of WO4 *) -(* ********************************************************************** *) - -(* ********************************************************************** *) -(* The union of appropriate values is the whole x *) -(* ********************************************************************** *) - -Goalw [LL_def] "w \\ MM ==> w Int y \\ LL"; -by (Fast_tac 1); -qed "Int_in_LL"; - -Goalw [LL_def] - "v \\ LL ==> v = (THE x. x \\ MM & v \\ x) Int y"; -by (Clarify_tac 1); -by (stac unique_superset2 1); -by (auto_tac (claset(), simpset() addsimps [Int_in_LL])); -qed "in_LL_eq_Int"; - -Goal "v \\ LL ==> (THE x. x \\ MM & v \\ x) \\ x Un y"; -by (dtac unique_superset1 1); -by (rewtac MM_def); -by (fast_tac (claset() addSDs [unique_superset1, includes RS subsetD]) 1); -qed "the_in_MM_subset"; - -Goalw [GG_def] "v \\ LL ==> GG ` v \\ x"; -by (ftac the_in_MM_subset 1); -by (ftac in_LL_eq_Int 1); -by (force_tac (claset() addEs [equalityE], simpset()) 1); -qed "GG_subset"; - - -Goal "n \\ nat ==> \\z. z \\ y & n eqpoll z"; -by (etac nat_lepoll_imp_ex_eqpoll_n 1); -by (resolve_tac [ordertype_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll - RSN (2, lepoll_trans)] 1); -by (rtac WO_R 2); -by (fast_tac - (claset() delrules [notI] - addSIs [nat_le_infinite_Ord RS le_imp_lepoll] - addIs [Ord_ordertype, - ordertype_eqpoll RS eqpoll_imp_lepoll - RS lepoll_infinite]) 1); -qed "ex_subset_eqpoll_n"; - - -Goal "u \\ x ==> \\v \\ s(u). succ(k) lepoll v Int y"; -by (rtac ccontr 1); -by (subgoal_tac "\\v \\ s(u). k eqpoll v Int y" 1); -by (full_simp_tac (simpset() addsimps [s_def]) 2); -by (blast_tac (claset() addIs [succ_not_lepoll_imp_eqpoll]) 2); -by (rewtac k_def); -by (cut_facts_tac [all_ex, includes, lnat] 1); -by (rtac (ex_subset_eqpoll_n RS exE) 1 THEN assume_tac 1); -by (rtac (noLepoll RS notE) 1); -by (blast_tac (claset() addIs - [[y_lepoll_subset_s, subset_s_lepoll_w] MRS lepoll_trans]) 1); -qed "exists_proper_in_s"; - -Goal "u \\ x ==> \\w \\ MM. u \\ w"; -by (eresolve_tac [exists_proper_in_s RS bexE] 1); -by (rewrite_goals_tac [MM_def, s_def]); -by (Fast_tac 1); -qed "exists_in_MM"; - -Goal "u \\ x ==> \\w \\ LL. u \\ GG`w"; -by (rtac (exists_in_MM RS bexE) 1); -by (assume_tac 1); -by (rtac bexI 1); -by (etac Int_in_LL 2); -by (rewtac GG_def); -by (asm_simp_tac (simpset() addsimps [Int_in_LL]) 1); -by (stac unique_superset2 1); -by (REPEAT (fast_tac (claset() addSEs [Int_in_LL]) 1)); -qed "exists_in_LL"; - - -Goal "well_ord(LL,S) ==> \ -\ (\\b MM ==> w eqpoll succ(k #+ m)"; -by (fast_tac (claset() addDs [includes RS subsetD]) 1); -qed "in_MM_eqpoll_n"; - -Goalw [LL_def, MM_def] "w \\ LL ==> succ(k) lepoll w"; -by (Fast_tac 1); -qed "in_LL_eqpoll_n"; - -val in_LL = in_LL_eq_Int RS equalityD1 RS (Int_lower1 RSN (2, subset_trans)); - -Goalw [GG_def] - "well_ord(LL,S) ==> \ -\ \\ba f. Ord(a) & domain(f) = a & \ -\ (\\bbb \\ ordertype(LL,S). GG ` (converse(ordermap(LL,S)) ` b)")] - exI 1); -by (Simp_tac 1); -by (REPEAT (ares_tac [conjI, lam_funtype RS domain_of_fun, - Ord_ordertype, - all_in_lepoll_m, OUN_eq_x] 1)); -qed "conclusion"; - -Close_locale "AC16"; - - - -(* ********************************************************************** *) -(* The main theorem AC16(n, k) ==> WO4(n-k) *) -(* ********************************************************************** *) - -Goalw [AC16_def,WO4_def] - "[| AC16(k #+ m, k); 0 < k; 0 < m; k \\ nat; m \\ nat |] ==> WO4(m)"; -by (rtac allI 1); -by (case_tac "Finite(A)" 1); -by (rtac lemma1 1 THEN REPEAT (assume_tac 1)); -by (cut_facts_tac [lemma2] 1); -by (REPEAT (eresolve_tac [exE, conjE] 1)); -by (eres_inst_tac [("x","A Un y")] allE 1); -by (ftac infinite_Un 1 THEN (mp_tac 1)); -by (etac zero_lt_natE 1); -by (assume_tac 1); -by (Clarify_tac 1); -by (DEPTH_SOLVE (ares_tac [export conclusion] 1)); -qed "AC16_WO4"; - diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC16_WO4.thy --- a/src/ZF/AC/AC16_WO4.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/AC16_WO4.thy Wed Jan 16 17:52:06 2002 +0100 @@ -2,40 +2,575 @@ ID: $Id$ Author: Krzysztof Grabczewski -Tidied using locales by LCP +The proof of AC16(n, k) ==> WO4(n-k) + +Tidied (using locales) by LCP *) -AC16_WO4 = OrderType + AC16_lemmas + Cardinal_aux + +theory AC16_WO4 = AC16_lemmas: + +(* ********************************************************************** *) +(* The case of finite set *) +(* ********************************************************************** *) + +lemma lemma1: + "[| Finite(A); 0 nat |] + ==> \a f. Ord(a) & domain(f) = a & + (\bb m)" +apply (unfold Finite_def) +apply (erule bexE) +apply (drule eqpoll_sym [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]]) +apply (erule exE) +apply (rule_tac x = "n" in exI) +apply (rule_tac x = "\i \ n. {f`i}" in exI, simp) +apply (unfold bij_def surj_def) +apply (fast intro!: ltI nat_into_Ord lam_funtype [THEN domain_of_fun] + singleton_eqpoll_1 [THEN eqpoll_imp_lepoll, THEN lepoll_trans] + nat_1_lepoll_iff [THEN iffD2] + elim!: apply_type ltE) +done + +(* ********************************************************************** *) +(* The case of infinite set *) +(* ********************************************************************** *) + +(* well_ord(x,r) ==> well_ord({{y,z}. y \ x}, Something(x,z)) **) +lemmas well_ord_paired = paired_bij [THEN bij_is_inj, THEN well_ord_rvimage] + +lemma lepoll_trans1: "[| A \ B; ~ A \ C |] ==> ~ B \ C" +by (blast intro: lepoll_trans) + +(* ********************************************************************** *) +(* There exists a well ordered set y such that ... *) +(* ********************************************************************** *) + +lemmas lepoll_paired = paired_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll]; + +lemma lemma2: "\y R. well_ord(y,R) & x Int y = 0 & ~y \ z & ~Finite(y)" +apply (rule_tac x = "{{a,x}. a \ nat Un Hartog (z) }" in exI) +apply (rule well_ord_Un [OF Ord_nat [THEN well_ord_Memrel] + Ord_Hartog [THEN well_ord_Memrel], THEN exE]) +apply (blast intro!: Ord_Hartog well_ord_Memrel well_ord_paired + lepoll_trans1 [OF _ not_Hartog_lepoll_self] + lepoll_trans [OF subset_imp_lepoll lepoll_paired] + elim!: nat_not_Finite [THEN notE] + elim: mem_asym + dest!: Un_upper1 [THEN subset_imp_lepoll, THEN lepoll_Finite] + lepoll_paired [THEN lepoll_Finite]) +done + +lemma infinite_Un: "~Finite(B) ==> ~Finite(A Un B)" +by (blast intro: subset_Finite) + +(* ********************************************************************** *) +(* There is a v \ s(u) such that k \ x Int y (in our case succ(k)) *) +(* The idea of the proof is the following \ *) +(* Suppose not, i.e. every element of s(u) has exactly k-1 elements of y *) +(* Thence y is less than or equipollent to {v \ Pow(x). v \ n#-k} *) +(* We have obtained this result in two steps \ *) +(* 1. y is less than or equipollent to {v \ s(u). a \ v} *) +(* where a is certain k-2 element subset of y *) +(* 2. {v \ s(u). a \ v} is less than or equipollent *) +(* to {v \ Pow(x). v \ n-k} *) +(* ********************************************************************** *) + +(*Proof simplified by LCP*) +lemma succ_not_lepoll_lemma: + "[| ~(\x \ A. f`x=y); f \ inj(A, B); y \ B |] + ==> (\a \ succ(A). if(a=A, y, f`a)) \ inj(succ(A), B)" +apply (rule_tac d = "%z. if (z=y, A, converse (f) `z) " in lam_injective) +apply (force simp add: inj_is_fun [THEN apply_type]) +(*this preliminary simplification prevents looping somehow*) +apply (simp (no_asm_simp)) +apply force +done + +lemma succ_not_lepoll_imp_eqpoll: "[| ~A \ B; A \ B |] ==> succ(A) \ B" +apply (unfold lepoll_def eqpoll_def bij_def surj_def) +apply (fast elim!: succ_not_lepoll_lemma inj_is_fun) +done + + +(* ********************************************************************** *) +(* There is a k-2 element subset of y *) +(* ********************************************************************** *) + +lemmas ordertype_eqpoll = + ordermap_bij [THEN exI [THEN eqpoll_def [THEN def_imp_iff, THEN iffD2]]] + +lemma cons_cons_subset: + "[| a \ y; b \ y-a; u \ x |] ==> cons(b, cons(u, a)) \ Pow(x Un y)" +by fast + +lemma cons_cons_eqpoll: + "[| a \ k; a \ y; b \ y-a; u \ x; x Int y = 0 |] + ==> cons(b, cons(u, a)) \ succ(succ(k))" +by (fast intro!: cons_eqpoll_succ) + +lemma set_eq_cons: + "[| succ(k) \ A; k \ B; B \ A; a \ A-B; k \ nat |] ==> A = cons(a, B)" +apply (rule equalityI) +prefer 2 apply fast +apply (rule Diff_eq_0_iff [THEN iffD1]) +apply (rule equals0I) +apply (drule eqpoll_sym [THEN eqpoll_imp_lepoll]) +apply (drule eqpoll_sym [THEN cons_eqpoll_succ], fast) +apply (drule cons_eqpoll_succ, fast) +apply (fast elim!: lepoll_trans [THEN lepoll_trans, THEN succ_lepoll_natE, + OF eqpoll_sym [THEN eqpoll_imp_lepoll] subset_imp_lepoll]) +done + +lemma cons_eqE: "[| cons(x,a) = cons(y,a); x \ a |] ==> x = y " +by (fast elim!: equalityE) + +lemma eq_imp_Int_eq: "A = B ==> A Int C = B Int C" +by blast + +(* ********************************************************************** *) +(* some arithmetic *) +(* ********************************************************************** *) + +lemma eqpoll_sum_imp_Diff_lepoll_lemma [rule_format]: + "[| k \ nat; m \ nat |] + ==> \A B. A \ k #+ m & k \ B & B \ A --> A-B \ m" +apply (induct_tac "k") +apply (simp add: add_0) +apply (blast intro: eqpoll_imp_lepoll lepoll_trans + Diff_subset [THEN subset_imp_lepoll]) +apply (intro allI impI) +apply (rule succ_lepoll_imp_not_empty [THEN not_emptyE], fast) +apply (erule_tac x = "A - {xa}" in allE) +apply (erule_tac x = "B - {xa}" in allE) +apply (erule impE) +apply (simp add: add_succ) +apply (fast intro!: Diff_sing_eqpoll lepoll_Diff_sing) +apply (subgoal_tac "A - {xa} - (B - {xa}) = A - B", simp); +apply blast +done + +lemma eqpoll_sum_imp_Diff_lepoll: + "[| A \ succ(k #+ m); B \ A; succ(k) \ B; k \ nat; m \ nat |] + ==> A-B \ m" +apply (simp only: add_succ [symmetric]) +apply (blast intro: eqpoll_sum_imp_Diff_lepoll_lemma) +done + +(* ********************************************************************** *) +(* similar properties for \ *) +(* ********************************************************************** *) + +lemma eqpoll_sum_imp_Diff_eqpoll_lemma [rule_format]: + "[| k \ nat; m \ nat |] + ==> \A B. A \ k #+ m & k \ B & B \ A --> A-B \ m" +apply (induct_tac "k") +apply (force dest!: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_0_is_0]) +apply (intro allI impI) +apply (rule succ_lepoll_imp_not_empty [THEN not_emptyE]) +apply (fast elim!: eqpoll_imp_lepoll) +apply (erule_tac x = "A - {xa}" in allE) +apply (erule_tac x = "B - {xa}" in allE) +apply (erule impE) +apply (force intro: eqpoll_sym intro!: Diff_sing_eqpoll) +apply (subgoal_tac "A - {xa} - (B - {xa}) = A - B", simp); +apply blast +done + +lemma eqpoll_sum_imp_Diff_eqpoll: + "[| A \ succ(k #+ m); B \ A; succ(k) \ B; k \ nat; m \ nat |] + ==> A-B \ m" +apply (simp only: add_succ [symmetric]) +apply (blast intro: eqpoll_sum_imp_Diff_eqpoll_lemma) +done + + +(* ********************************************************************** *) +(* LL can be well ordered *) +(* ********************************************************************** *) + +lemma subsets_lepoll_0_eq_unit: "{x \ Pow(X). x \ 0} = {0}" +by (fast dest!: lepoll_0_is_0 intro!: lepoll_refl) + +lemma subsets_lepoll_succ: + "n \ nat ==> {z \ Pow(y). z \ succ(n)} = + {z \ Pow(y). z \ n} Un {z \ Pow(y). z \ succ(n)}" +by (blast intro: leI le_imp_lepoll nat_into_Ord + lepoll_trans eqpoll_imp_lepoll + dest!: lepoll_succ_disj) + +lemma Int_empty: + "n \ nat ==> {z \ Pow(y). z \ n} Int {z \ Pow(y). z \ succ(n)} = 0" +by (blast intro: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_trans] + succ_lepoll_natE) + locale AC16 = - fixes - x :: i - y :: i - k :: i - l :: i - m :: i - t_n :: i - R :: i - MM :: i - LL :: i - GG :: i - s :: i=>i - assumes - all_ex "\\z \\ {z \\ Pow(x Un y) . z eqpoll succ(k)}. - \\! w. w \\ t_n & z \\ w " - disjoint "x Int y = 0" - includes "t_n \\ {v \\ Pow(x Un y). v eqpoll succ(k #+ m)}" - WO_R "well_ord(y,R)" - lnat "l \\ nat" - mnat "m \\ nat" - mpos "0 Pow(x). v eqpoll m}" - defines - k_def "k == succ(l)" - MM_def "MM == {v \\ t_n. succ(k) lepoll v Int y}" - LL_def "LL == {v Int y. v \\ MM}" - GG_def "GG == \\v \\ LL. (THE w. w \\ MM & v \\ w) - v" - s_def "s(u) == {v \\ t_n. u \\ v & k lepoll v Int y}" + fixes x and y and k and l and m and t_n and R and MM and LL and GG and s + defines k_def: "k == succ(l)" + and MM_def: "MM == {v \ t_n. succ(k) \ v Int y}" + and LL_def: "LL == {v Int y. v \ MM}" + and GG_def: "GG == \v \ LL. (THE w. w \ MM & v \ w) - v" + and s_def: "s(u) == {v \ t_n. u \ v & k \ v Int y}" + assumes all_ex: "\z \ {z \ Pow(x Un y) . z \ succ(k)}. + \! w. w \ t_n & z \ w " + and disjoint[iff]: "x Int y = 0" + and includes: "t_n \ {v \ Pow(x Un y). v \ succ(k #+ m)}" + and WO_R[iff]: "well_ord(y,R)" + and lnat[iff]: "l \ nat" + and mnat[iff]: "m \ nat" + and mpos[iff]: "0 {v \ Pow(x). v \ m}" + +lemma (in AC16) knat [iff]: "k \ nat" +by (simp add: k_def) + + +(* ********************************************************************** *) +(* 1. y is less than or equipollent to {v \ s(u). a \ v} *) +(* where a is certain k-2 element subset of y *) +(* ********************************************************************** *) + +lemma (in AC16) Diff_Finite_eqpoll: "[| l \ a; a \ y |] ==> y - a \ y" +apply (insert WO_R Infinite lnat) +apply (rule eqpoll_trans) +apply (rule Diff_lesspoll_eqpoll_Card) +apply (erule well_ord_cardinal_eqpoll [THEN eqpoll_sym]) +apply (blast intro: lesspoll_trans1 + intro!: Card_cardinal + Card_cardinal [THEN Card_is_Ord, THEN nat_le_infinite_Ord, + THEN le_imp_lepoll] + dest: well_ord_cardinal_eqpoll + eqpoll_sym eqpoll_imp_lepoll + n_lesspoll_nat [THEN lesspoll_trans2] + well_ord_cardinal_eqpoll [THEN eqpoll_sym, + THEN eqpoll_imp_lepoll, THEN lepoll_infinite])+ +done + + +lemma (in AC16) s_subset: "s(u) \ t_n" +by (unfold s_def, blast) + +lemma (in AC16) sI: + "[| w \ t_n; cons(b,cons(u,a)) \ w; a \ y; b \ y-a; l \ a |] + ==> w \ s(u)" +apply (unfold s_def succ_def k_def) +apply (blast intro!: eqpoll_imp_lepoll [THEN cons_lepoll_cong] + intro: subset_imp_lepoll lepoll_trans) +done + +lemma (in AC16) in_s_imp_u_in: "v \ s(u) ==> u \ v" +by (unfold s_def, blast) + + +lemma (in AC16) ex1_superset_a: + "[| l \ a; a \ y; b \ y - a; u \ x |] + ==> \! c. c \ s(u) & a \ c & b \ c" +apply (rule all_ex [simplified k_def, THEN ballE]) + apply (erule ex1E) + apply (rule_tac a = "w" in ex1I, blast intro: sI) + apply (blast dest: s_subset [THEN subsetD] in_s_imp_u_in) +apply (blast del: PowI + intro!: cons_cons_subset eqpoll_sym [THEN cons_cons_eqpoll]) +done + +lemma (in AC16) the_eq_cons: + "[| \v \ s(u). succ(l) \ v Int y; + l \ a; a \ y; b \ y - a; u \ x |] + ==> (THE c. c \ s(u) & a \ c & b \ c) Int y = cons(b, a)" +apply (frule ex1_superset_a [THEN theI], assumption+) +apply (rule set_eq_cons) +apply (fast+) +done + +lemma (in AC16) y_lepoll_subset_s: + "[| \v \ s(u). succ(l) \ v Int y; + l \ a; a \ y; u \ x |] + ==> y \ {v \ s(u). a \ v}" +apply (rule Diff_Finite_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll, + THEN lepoll_trans], fast+) +apply (rule_tac f3 = "\b \ y-a. THE c. c \ s (u) & a \ c & b \ c" + in exI [THEN lepoll_def [THEN def_imp_iff, THEN iffD2]]) +apply (simp add: inj_def) +apply (rule conjI) +apply (rule lam_type) +apply (frule ex1_superset_a [THEN theI], fast+, clarify) +apply (rule cons_eqE [of _ a]) +apply (drule_tac A = "THE c. ?P (c) " and C = "y" in eq_imp_Int_eq) +apply (simp_all add: the_eq_cons) +done + + +(* ********************************************************************** *) +(* back to the second part *) +(* ********************************************************************** *) + + +(*relies on the disjointness of x, y*) +lemma (in AC16) x_imp_not_y [dest]: "a \ x ==> a \ y" +by (blast dest: disjoint [THEN equalityD1, THEN subsetD, OF IntI]) + +lemma (in AC16) w_Int_eq_w_Diff: + "w \ x Un y ==> w Int (x - {u}) = w - cons(u, w Int y)" +by blast + +lemma (in AC16) w_Int_eqpoll_m: + "[| w \ {v \ s(u). a \ v}; + l \ a; u \ x; + \v \ s(u). succ(l) \ v Int y |] + ==> w Int (x - {u}) \ m" +apply (erule CollectE) +apply (subst w_Int_eq_w_Diff) +apply (fast dest!: s_subset [THEN subsetD] + includes [simplified k_def, THEN subsetD]) +apply (blast dest: s_subset [THEN subsetD] + includes [simplified k_def, THEN subsetD] + dest: eqpoll_sym [THEN cons_eqpoll_succ, THEN eqpoll_sym] + in_s_imp_u_in + intro!: eqpoll_sum_imp_Diff_eqpoll) +done + + +(* ********************************************************************** *) +(* 2. {v \ s(u). a \ v} is less than or equipollent *) +(* to {v \ Pow(x). v \ n-k} *) +(* ********************************************************************** *) + +lemma (in AC16) eqpoll_m_not_empty: "a \ m ==> a \ 0" +apply (insert mpos) +apply (fast elim!: zero_lt_natE dest!: eqpoll_succ_imp_not_empty) +done + +lemma (in AC16) cons_cons_in: + "[| z \ xa Int (x - {u}); l \ a; a \ y; u \ x |] + ==> \! w. w \ t_n & cons(z, cons(u, a)) \ w" +apply (rule all_ex [THEN bspec]) +apply (unfold k_def) +apply (fast intro!: cons_eqpoll_succ elim: eqpoll_sym) +done + +lemma (in AC16) subset_s_lepoll_w: + "[| \v \ s(u). succ(l) \ v Int y; a \ y; l \ a; u \ x |] + ==> {v \ s(u). a \ v} \ {v \ Pow(x). v \ m}" +apply (rule_tac f3 = "\w \ {v \ s (u) . a \ v}. w Int (x - {u})" + in exI [THEN lepoll_def [THEN def_imp_iff, THEN iffD2]]) +apply (simp add: inj_def) +apply (intro conjI lam_type CollectI) + apply fast + apply (blast intro: w_Int_eqpoll_m) +apply (intro ballI impI) +(** LEVEL 8 **) +apply (rule w_Int_eqpoll_m [THEN eqpoll_m_not_empty, THEN not_emptyE]) +apply (blast, assumption+) +apply (drule equalityD1 [THEN subsetD], (assumption)) +apply (frule cons_cons_in, assumption+) +apply (blast dest: ex1_two_eq intro: s_subset [THEN subsetD] in_s_imp_u_in)+ +done + + +(* ********************************************************************** *) +(* well_ord_subsets_lepoll_n *) +(* ********************************************************************** *) + +lemma (in AC16) well_ord_subsets_eqpoll_n: + "n \ nat ==> \S. well_ord({z \ Pow(y) . z \ succ(n)}, S)" +apply (rule WO_R [THEN well_ord_infinite_subsets_eqpoll_X, + THEN eqpoll_def [THEN def_imp_iff, THEN iffD1], THEN exE]) +apply (fast intro: bij_is_inj [THEN well_ord_rvimage])+ +done + +lemma (in AC16) well_ord_subsets_lepoll_n: + "n \ nat ==> \R. well_ord({z \ Pow(y). z \ n}, R)" +apply (induct_tac "n") +apply (force intro!: well_ord_unit simp add: subsets_lepoll_0_eq_unit) +apply (erule exE) +apply (rule well_ord_subsets_eqpoll_n [THEN exE], assumption) +apply (simp add: subsets_lepoll_succ) +apply (drule well_ord_radd, (assumption)) +apply (erule Int_empty [THEN disj_Un_eqpoll_sum, + THEN eqpoll_def [THEN def_imp_iff, THEN iffD1], THEN exE]) +apply (fast elim!: bij_is_inj [THEN well_ord_rvimage]) +done + + +lemma (in AC16) LL_subset: "LL \ {z \ Pow(y). z \ succ(k #+ m)}" +apply (unfold LL_def MM_def) +apply (insert includes) +apply (blast intro: subset_imp_lepoll eqpoll_imp_lepoll lepoll_trans) +done + +lemma (in AC16) well_ord_LL: "\S. well_ord(LL,S)" +apply (rule well_ord_subsets_lepoll_n [THEN exE, of "succ(k#+m)"]) +apply simp +apply (blast intro: well_ord_subset [OF _ LL_subset]) +done + +(* ********************************************************************** *) +(* every element of LL is a contained in exactly one element of MM *) +(* ********************************************************************** *) + +lemma (in AC16) unique_superset_in_MM: + "v \ LL ==> \! w. w \ MM & v \ w" +apply (unfold MM_def LL_def, safe) +apply fast +apply (rule lepoll_imp_eqpoll_subset [THEN exE], (assumption)) +apply (rule_tac x = "x" in all_ex [THEN ballE]) +apply (blast intro: eqpoll_sym)+ +done + + +(* ********************************************************************** *) +(* The function GG satisfies the conditions of WO4 *) +(* ********************************************************************** *) + +(* ********************************************************************** *) +(* The union of appropriate values is the whole x *) +(* ********************************************************************** *) + +lemma (in AC16) Int_in_LL: "w \ MM ==> w Int y \ LL" +by (unfold LL_def, fast) + +lemma (in AC16) in_LL_eq_Int: + "v \ LL ==> v = (THE x. x \ MM & v \ x) Int y" +apply (unfold LL_def, clarify) +apply (subst unique_superset_in_MM [THEN the_equality2]) +apply (auto simp add: Int_in_LL) +done + +lemma (in AC16) unique_superset1: "a \ LL \ (THE x. x \ MM \ a \ x) \ MM" +by (erule unique_superset_in_MM [THEN theI, THEN conjunct1]); + +lemma (in AC16) the_in_MM_subset: + "v \ LL ==> (THE x. x \ MM & v \ x) \ x Un y" +apply (drule unique_superset1) +apply (unfold MM_def) +apply (fast dest!: unique_superset1 includes [THEN subsetD]) +done + +lemma (in AC16) GG_subset: "v \ LL ==> GG ` v \ x" +apply (unfold GG_def) +apply (frule the_in_MM_subset) +apply (frule in_LL_eq_Int) +apply (force elim: equalityE) +done + +lemma (in AC16) nat_lepoll_ordertype: "nat \ ordertype(y, R)" +apply (rule nat_le_infinite_Ord [THEN le_imp_lepoll]) + apply (rule Ord_ordertype [OF WO_R]) +apply (rule ordertype_eqpoll [THEN eqpoll_imp_lepoll, THEN lepoll_infinite]) + apply (rule WO_R) +apply (rule Infinite) +done + +lemma (in AC16) ex_subset_eqpoll_n: "n \ nat ==> \z. z \ y & n \ z" +apply (erule nat_lepoll_imp_ex_eqpoll_n) +apply (rule lepoll_trans [OF nat_lepoll_ordertype]) +apply (rule ordertype_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll]) +apply (rule WO_R) +done + + +lemma (in AC16) exists_proper_in_s: "u \ x ==> \v \ s(u). succ(k) \ v Int y" +apply (rule ccontr) +apply (subgoal_tac "\v \ s (u) . k \ v Int y") +prefer 2 apply (simp add: s_def, blast intro: succ_not_lepoll_imp_eqpoll) +apply (unfold k_def) +apply (insert all_ex includes lnat) +apply (rule ex_subset_eqpoll_n [THEN exE], assumption) +apply (rule noLepoll [THEN notE]) +apply (blast intro: lepoll_trans [OF y_lepoll_subset_s subset_s_lepoll_w]) +done + +lemma (in AC16) exists_in_MM: "u \ x ==> \w \ MM. u \ w" +apply (erule exists_proper_in_s [THEN bexE]) +apply (unfold MM_def s_def, fast) +done + +lemma (in AC16) exists_in_LL: "u \ x ==> \w \ LL. u \ GG`w" +apply (rule exists_in_MM [THEN bexE], assumption) +apply (rule bexI) +apply (erule_tac [2] Int_in_LL) +apply (unfold GG_def) +apply (simp add: Int_in_LL) +apply (subst unique_superset_in_MM [THEN the_equality2]) +apply (fast elim!: Int_in_LL)+ +done + +lemma (in AC16) OUN_eq_x: "well_ord(LL,S) ==> + (\b MM ==> w \ succ(k #+ m)" +apply (unfold MM_def) +apply (fast dest: includes [THEN subsetD]) +done + +lemma (in AC16) in_LL_eqpoll_n: "w \ LL ==> succ(k) \ w" +by (unfold LL_def MM_def, fast) + +lemma (in AC16) in_LL: "w \ LL ==> w \ (THE x. x \ MM \ w \ x)" +by (erule subset_trans [OF in_LL_eq_Int [THEN equalityD1] Int_lower1]) + +lemma (in AC16) all_in_lepoll_m: + "well_ord(LL,S) ==> + \b m" +apply (unfold GG_def) +apply (rule oallI) +apply (simp add: ltD ordermap_bij [THEN bij_converse_bij, THEN bij_is_fun, THEN apply_type]) +apply (insert includes) +apply (rule eqpoll_sum_imp_Diff_lepoll) +apply (blast del: subsetI + dest!: ltD + intro!: eqpoll_sum_imp_Diff_lepoll in_LL_eqpoll_n + intro: in_LL unique_superset1 [THEN in_MM_eqpoll_n] + ordermap_bij [THEN bij_converse_bij, THEN bij_is_fun, + THEN apply_type])+ +done + +lemma (in AC16) conclusion: + "\a f. Ord(a) & domain(f) = a & (\bb m)" +apply (rule well_ord_LL [THEN exE]) +apply (rename_tac S) +apply (rule_tac x = "ordertype (LL,S)" in exI) +apply (rule_tac x = "\b \ ordertype(LL,S). + GG ` (converse (ordermap (LL,S)) ` b)" in exI) +apply simp +apply (blast intro: lam_funtype [THEN domain_of_fun] + Ord_ordertype OUN_eq_x all_in_lepoll_m [THEN ospec]) +done + + +(* ********************************************************************** *) +(* The main theorem AC16(n, k) ==> WO4(n-k) *) +(* ********************************************************************** *) + +theorem AC16_WO4: + "[| AC16(k #+ m, k); 0 < k; 0 < m; k \ nat; m \ nat |] ==> WO4(m)" +apply (unfold AC16_def WO4_def) +apply (rule allI) +apply (case_tac "Finite (A)") +apply (rule lemma1, assumption+) +apply (cut_tac lemma2) +apply (elim exE conjE) +apply (erule_tac x = "A Un y" in allE) +apply (frule infinite_Un, drule mp, assumption) +apply (erule zero_lt_natE, assumption, clarify) +apply (blast intro: AC16.conclusion) +done end diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC16_lemmas.ML --- a/src/ZF/AC/AC16_lemmas.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,275 +0,0 @@ -(* Title: ZF/AC/AC16_lemmas.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -Lemmas used in the proofs concerning AC16 -*) - -Goal "a\\A ==> cons(a,A)-{a}=A"; -by (Fast_tac 1); -qed "cons_Diff_eq"; - -Goalw [lepoll_def] "1 lepoll X <-> (\\x. x \\ X)"; -by (rtac iffI 1); -by (fast_tac (claset() addIs [inj_is_fun RS apply_type]) 1); -by (etac exE 1); -by (res_inst_tac [("x","\\a \\ 1. x")] exI 1); -by (fast_tac (claset() addSIs [lam_injective]) 1); -qed "nat_1_lepoll_iff"; - -Goal "X eqpoll 1 <-> (\\x. X={x})"; -by (rtac iffI 1); -by (etac eqpollE 1); -by (dresolve_tac [nat_1_lepoll_iff RS iffD1] 1); -by (fast_tac (claset() addSIs [lepoll_1_is_sing]) 1); -by (fast_tac (claset() addSIs [singleton_eqpoll_1]) 1); -qed "eqpoll_1_iff_singleton"; - -Goalw [succ_def] "[| x eqpoll n; y\\x |] ==> cons(y,x) eqpoll succ(n)"; -by (fast_tac (claset() addSEs [cons_eqpoll_cong, mem_irrefl]) 1); -qed "cons_eqpoll_succ"; - -Goal "{Y \\ Pow(X). Y eqpoll 1} = {{x}. x \\ X}"; -by (rtac equalityI 1); -by (rtac subsetI 1); -by (etac CollectE 1); -by (dresolve_tac [eqpoll_1_iff_singleton RS iffD1] 1); -by (fast_tac (claset() addSIs [RepFunI]) 1); -by (rtac subsetI 1); -by (etac RepFunE 1); -by (rtac CollectI 1); -by (Fast_tac 1); -by (fast_tac (claset() addSIs [singleton_eqpoll_1]) 1); -qed "subsets_eqpoll_1_eq"; - -Goalw [eqpoll_def, bij_def] "X eqpoll {{x}. x \\ X}"; -by (res_inst_tac [("x","\\x \\ X. {x}")] exI 1); -by (rtac IntI 1); -by (rewrite_goals_tac [inj_def, surj_def]); -by (Asm_full_simp_tac 1); -by (fast_tac (claset() addSIs [lam_type, RepFunI] - addIs [singleton_eq_iff RS iffD1]) 1); -by (Asm_full_simp_tac 1); -by (fast_tac (claset() addSIs [lam_type]) 1); -qed "eqpoll_RepFun_sing"; - -Goal "{Y \\ Pow(X). Y eqpoll 1} eqpoll X"; -by (resolve_tac [subsets_eqpoll_1_eq RS ssubst] 1); -by (resolve_tac [eqpoll_RepFun_sing RS eqpoll_sym] 1); -qed "subsets_eqpoll_1_eqpoll"; - -Goal "[| InfCard(x); y \\ Pow(x); y eqpoll succ(z) |] \ -\ ==> (LEAST i. i \\ y) \\ y"; -by (eresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS - succ_lepoll_imp_not_empty RS not_emptyE] 1); -by (fast_tac (claset() addIs [LeastI] - addSDs [InfCard_is_Card RS Card_is_Ord, PowD RS subsetD] - addEs [Ord_in_Ord]) 1); -qed "InfCard_Least_in"; - -Goalw [lepoll_def] "[| InfCard(x); n \\ nat |] ==> \ -\ {y \\ Pow(x). y eqpoll succ(succ(n))} lepoll \ -\ x*{y \\ Pow(x). y eqpoll succ(n)}"; -by (res_inst_tac [("x","\\y \\ {y \\ Pow(x). y eqpoll succ(succ(n))}. \ -\ y, y-{LEAST i. i \\ y}>")] exI 1); -by (res_inst_tac [("d","%z. cons(fst(z),snd(z))")] lam_injective 1); -by (rtac SigmaI 1); -by (etac CollectE 1); -by (Asm_full_simp_tac 3); -by (rtac equalityI 3); -by (Fast_tac 4); -by (rtac subsetI 3); -by (etac consE 3); -by (Fast_tac 4); -by (rtac CollectI 2); -by (Fast_tac 2); -by (resolve_tac [PowD RS subsetD] 1 THEN (assume_tac 1)); -by (REPEAT (fast_tac (claset() addSIs [Diff_sing_eqpoll] - addIs [InfCard_Least_in]) 1)); -qed "subsets_lepoll_lemma1"; - -val prems = goal thy "(!!y. y \\ z ==> Ord(y)) ==> z \\ succ(Union(z))"; -by (rtac subsetI 1); -by (res_inst_tac [("Q","\\y \\ z. y \\ x")] (excluded_middle RS disjE) 1); -by (Fast_tac 2); -by (etac swap 1); -by (rtac ballI 1); -by (rtac Ord_linear_le 1); -by (dtac le_imp_subset 3 THEN (assume_tac 3)); -by (fast_tac (claset() addDs prems) 1); -by (fast_tac (claset() addDs prems) 1); -by (fast_tac (claset() addSEs [leE,ltE]) 1); -qed "set_of_Ord_succ_Union"; - -Goal "j \\ i ==> i \\ j"; -by (fast_tac (claset() addSEs [mem_irrefl]) 1); -qed "subset_not_mem"; - -val prems = goal thy "(!!y. y \\ z ==> Ord(y)) ==> succ(Union(z)) \\ z"; -by (resolve_tac [set_of_Ord_succ_Union RS subset_not_mem] 1); -by (eresolve_tac prems 1); -qed "succ_Union_not_mem"; - -Goal "Union(cons(succ(Union(z)),z)) = succ(Union(z))"; -by (Fast_tac 1); -qed "Union_cons_eq_succ_Union"; - -Goal "[| Ord(i); Ord(j) |] ==> i Un j = i | i Un j = j"; -by (fast_tac (claset() addSDs [le_imp_subset] addEs [Ord_linear_le]) 1); -qed "Un_Ord_disj"; - -Goal "x \\ X ==> Union(X) = x Un Union(X-{x})"; -by (Fast_tac 1); -qed "Union_eq_Un"; - -Goal "n \\ nat ==> \ -\ \\z. (\\y \\ z. Ord(y)) & z eqpoll n & z\\0 --> Union(z) \\ z"; -by (induct_tac "n" 1); -by (fast_tac (claset() addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]) 1); -by (REPEAT (resolve_tac [allI, impI] 1)); -by (etac natE 1); -by (fast_tac (claset() addSDs [eqpoll_1_iff_singleton RS iffD1] - addSIs [Union_singleton]) 1); -by (hyp_subst_tac 1); -by (REPEAT (eresolve_tac [conjE, not_emptyE] 1)); -by (eres_inst_tac [("x","z-{xb}")] allE 1); -by (etac impE 1); -by (fast_tac (claset() addSEs [Diff_sing_eqpoll, - Diff_sing_eqpoll RS eqpoll_succ_imp_not_empty]) 1); -by (resolve_tac [Union_eq_Un RSN (2, subst_elem)] 1 THEN (assume_tac 2)); -by (ftac bspec 1 THEN (assume_tac 1)); -by (dresolve_tac [Diff_subset RS subsetD RSN (2, bspec)] 1 THEN (assume_tac 1)); -by (dtac Un_Ord_disj 1 THEN (assume_tac 1)); -by (etac DiffE 1); -by (etac disjE 1); -by (etac subst_elem 1 THEN (assume_tac 1)); -by (rtac subst_elem 1 THEN (TRYALL assume_tac)); -qed "Union_in_lemma"; - -Goal "[| \\x \\ z. Ord(x); z eqpoll n; z\\0; n \\ nat |] \ -\ ==> Union(z) \\ z"; -by (dtac Union_in_lemma 1); -by (fast_tac FOL_cs 1); -qed "Union_in"; - -Goal "[| InfCard(x); z \\ Pow(x); z eqpoll n; n \\ nat |] \ -\ ==> succ(Union(z)) \\ x"; -by (resolve_tac [Limit_has_succ RS ltE] 1 THEN (assume_tac 3)); -by (etac InfCard_is_Limit 1); -by (excluded_middle_tac "z=0" 1); -by (fast_tac (claset() addSIs [InfCard_is_Limit RS Limit_has_0] - addss (simpset())) 2); -by (resolve_tac - [PowD RS subsetD RS (InfCard_is_Card RS Card_is_Ord RSN (2, ltI))] 1 - THEN (TRYALL assume_tac)); -by (fast_tac (claset() addSIs [Union_in] - addSEs [PowD RS subsetD RSN - (2, InfCard_is_Card RS Card_is_Ord RS Ord_in_Ord)]) 1); -qed "succ_Union_in_x"; - -Goalw [lepoll_def] "[| InfCard(x); n \\ nat |] ==> \ -\ {y \\ Pow(x). y eqpoll succ(n)} lepoll \ -\ {y \\ Pow(x). y eqpoll succ(succ(n))}"; -by (res_inst_tac [("x","\\z \\ {y \\ Pow(x). y eqpoll succ(n)}. \ -\ cons(succ(Union(z)), z)")] exI 1); -by (res_inst_tac [("d","%z. z-{Union(z)}")] lam_injective 1); -by (resolve_tac [Union_cons_eq_succ_Union RS ssubst] 2); -by (rtac cons_Diff_eq 2); -by (fast_tac (claset() addSDs [InfCard_is_Card RS Card_is_Ord] - addEs [Ord_in_Ord] addSIs [succ_Union_not_mem]) 2); -by (rtac CollectI 1); -by (fast_tac (claset() addSEs [cons_eqpoll_succ] - addSIs [succ_Union_not_mem] - addSDs [InfCard_is_Card RS Card_is_Ord] - addEs [Ord_in_Ord]) 2); -by (fast_tac (claset() addSIs [succ_Union_in_x]) 1); -qed "succ_lepoll_succ_succ"; - -Goal "[| InfCard(X); n \\ nat |] \ -\ ==> {Y \\ Pow(X). Y eqpoll succ(n)} eqpoll X"; -by (induct_tac "n" 1); -by (rtac subsets_eqpoll_1_eqpoll 1); -by (rtac eqpollI 1); -by (resolve_tac [subsets_lepoll_lemma1 RS lepoll_trans] 1 - THEN (REPEAT (assume_tac 1))); -by (resolve_tac [InfCard_is_Card RS Card_is_Ord RS well_ord_Memrel RS - well_ord_InfCard_square_eq RS eqpoll_imp_lepoll - RSN (2, lepoll_trans)] 1 THEN (assume_tac 2)); -by (resolve_tac [InfCard_is_Card RS Card_cardinal_eq RS ssubst] 2 - THEN (REPEAT (assume_tac 2))); -by (eresolve_tac [eqpoll_refl RS prod_eqpoll_cong RS eqpoll_imp_lepoll] 1); -by (fast_tac (claset() addEs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_trans] - addSIs [succ_lepoll_succ_succ]) 1); -qed "subsets_eqpoll_X"; - -Goalw [surj_def] "[| f \\ surj(A,B); y \\ B |] \ -\ ==> f``(converse(f)``y) = y"; -by (fast_tac (claset() addDs [apply_equality2] - addEs [apply_iff RS iffD2]) 1); -qed "image_vimage_eq"; - -Goal "[| f \\ inj(A,B); y \\ A |] ==> converse(f)``(f``y) = y"; -by (fast_tac (claset() addSEs [inj_is_fun RS apply_Pair] - addDs [inj_equality]) 1); -qed "vimage_image_eq"; - -Goalw [eqpoll_def] "A eqpoll B \ -\ ==> {Y \\ Pow(A). Y eqpoll n} eqpoll {Y \\ Pow(B). Y eqpoll n}"; -by (etac exE 1); -by (res_inst_tac [("x","\\X \\ {Y \\ Pow(A). \\f. f \\ bij(Y, n)}. f``X")] exI 1); -by (res_inst_tac [("d","%Z. converse(f)``Z")] lam_bijective 1); -by (fast_tac (claset() - addSIs [bij_is_inj RS restrict_bij RS bij_converse_bij RS comp_bij] - addSEs [bij_is_fun RS fun_is_rel RS image_subset RS PowI]) 1); -by (fast_tac (claset() addSIs [bij_converse_bij RS bij_is_inj RS restrict_bij - RS bij_converse_bij RS comp_bij] - addSEs [bij_converse_bij RS bij_is_fun RS fun_is_rel - RS image_subset RS PowI]) 1); -by (fast_tac (claset() addSEs [bij_is_inj RS vimage_image_eq]) 1); -by (fast_tac (claset() addSEs [bij_is_surj RS image_vimage_eq]) 1); -qed "subsets_eqpoll"; - -Goalw [WO2_def] "WO2 ==> \\a. Card(a) & X eqpoll a"; -by (REPEAT (eresolve_tac [allE,exE,conjE] 1)); -by (fast_tac (claset() addSEs [well_ord_Memrel RS well_ord_cardinal_eqpoll RS - (eqpoll_sym RSN (2, eqpoll_trans)) RS eqpoll_sym] - addSIs [Card_cardinal]) 1); -qed "WO2_imp_ex_Card"; - -Goal "[| X lepoll Y; ~Finite(X) |] ==> ~Finite(Y)"; -by (fast_tac (empty_cs addEs [notE, lepoll_Finite] addSIs [notI]) 1); -qed "lepoll_infinite"; - -Goalw [InfCard_def] "[| ~Finite(X); Card(X) |] ==> InfCard(X)"; -by (fast_tac (claset() addSEs [Card_is_Ord RS nat_le_infinite_Ord]) 1); -qed "infinite_Card_is_InfCard"; - -Goal "[| WO2; n \\ nat; ~Finite(X) |] \ -\ ==> {Y \\ Pow(X). Y eqpoll succ(n)} eqpoll X"; -by (dtac WO2_imp_ex_Card 1); -by (REPEAT (eresolve_tac [allE,exE,conjE] 1)); -by (forward_tac [eqpoll_imp_lepoll RS lepoll_infinite] 1 THEN (assume_tac 1)); -by (dtac infinite_Card_is_InfCard 1 THEN (assume_tac 1)); -by (resolve_tac [eqpoll_trans RS eqpoll_trans] 1); -by (etac subsets_eqpoll 1); -by (etac subsets_eqpoll_X 1 THEN (assume_tac 1)); -by (etac eqpoll_sym 1); -qed "WO2_infinite_subsets_eqpoll_X"; - -Goal "well_ord(X,R) ==> \\a. Card(a) & X eqpoll a"; -by (fast_tac (claset() addSEs [well_ord_cardinal_eqpoll RS eqpoll_sym] - addSIs [Card_cardinal]) 1); -qed "well_ord_imp_ex_Card"; - -Goal "[| well_ord(X,R); n \\ nat; ~Finite(X) |] \ -\ ==> {Y \\ Pow(X). Y eqpoll succ(n)} eqpoll X"; -by (dtac well_ord_imp_ex_Card 1); -by (REPEAT (eresolve_tac [allE,exE,conjE] 1)); -by (forward_tac [eqpoll_imp_lepoll RS lepoll_infinite] 1 THEN (assume_tac 1)); -by (dtac infinite_Card_is_InfCard 1 THEN (assume_tac 1)); -by (resolve_tac [eqpoll_trans RS eqpoll_trans] 1); -by (etac subsets_eqpoll 1); -by (etac subsets_eqpoll_X 1 THEN (assume_tac 1)); -by (etac eqpoll_sym 1); -qed "well_ord_infinite_subsets_eqpoll_X"; - diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC16_lemmas.thy --- a/src/ZF/AC/AC16_lemmas.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/AC16_lemmas.thy Wed Jan 16 17:52:06 2002 +0100 @@ -1,3 +1,244 @@ -(*Dummy theory to document dependencies *) +(* Title: ZF/AC/AC16_lemmas.thy + ID: $Id$ + Author: Krzysztof Grabczewski + +Lemmas used in the proofs concerning AC16 +*) + +theory AC16_lemmas = AC_Equiv + Hartog + Cardinal_aux: + +lemma cons_Diff_eq: "a\A ==> cons(a,A)-{a}=A" +by fast + +lemma nat_1_lepoll_iff: "1\X <-> (\x. x \ X)" +apply (unfold lepoll_def) +apply (rule iffI) +apply (fast intro: inj_is_fun [THEN apply_type]) +apply (erule exE) +apply (rule_tac x = "\a \ 1. x" in exI) +apply (fast intro!: lam_injective) +done + +lemma eqpoll_1_iff_singleton: "X\1 <-> (\x. X={x})" +apply (rule iffI) +apply (erule eqpollE) +apply (drule nat_1_lepoll_iff [THEN iffD1]) +apply (fast intro!: lepoll_1_is_sing) +apply (fast intro!: singleton_eqpoll_1) +done + +lemma cons_eqpoll_succ: "[| x\n; y\x |] ==> cons(y,x)\succ(n)" +apply (unfold succ_def) +apply (fast elim!: cons_eqpoll_cong mem_irrefl) +done + +lemma subsets_eqpoll_1_eq: "{Y \ Pow(X). Y\1} = {{x}. x \ X}" +apply (rule equalityI) +apply (rule subsetI) +apply (erule CollectE) +apply (drule eqpoll_1_iff_singleton [THEN iffD1]) +apply (fast intro!: RepFunI) +apply (rule subsetI) +apply (erule RepFunE) +apply (rule CollectI) +apply fast +apply (fast intro!: singleton_eqpoll_1) +done + +lemma eqpoll_RepFun_sing: "X\{{x}. x \ X}" +apply (unfold eqpoll_def bij_def) +apply (rule_tac x = "\x \ X. {x}" in exI) +apply (rule IntI) +apply (unfold inj_def surj_def) +apply simp +apply (fast intro!: lam_type RepFunI intro: singleton_eq_iff [THEN iffD1]) +apply simp +apply (fast intro!: lam_type) +done + +lemma subsets_eqpoll_1_eqpoll: "{Y \ Pow(X). Y\1}\X" +apply (rule subsets_eqpoll_1_eq [THEN ssubst]) +apply (rule eqpoll_RepFun_sing [THEN eqpoll_sym]) +done + +lemma InfCard_Least_in: + "[| InfCard(x); y \ x; y \ succ(z) |] ==> (LEAST i. i \ y) \ y" +apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll, + THEN succ_lepoll_imp_not_empty, THEN not_emptyE]) +apply (fast intro: LeastI + dest!: InfCard_is_Card [THEN Card_is_Ord] + elim: Ord_in_Ord) +done + +lemma subsets_lepoll_lemma1: + "[| InfCard(x); n \ nat |] + ==> {y \ Pow(x). y\succ(succ(n))} \ x*{y \ Pow(x). y\succ(n)}" +apply (unfold lepoll_def) +apply (rule_tac x = "\y \ {y \ Pow(x) . y\succ (succ (n))}. + y, y-{LEAST i. i \ y}>" in exI) +apply (rule_tac d = "%z. cons (fst(z), snd(z))" in lam_injective); + apply (blast intro!: Diff_sing_eqpoll intro: InfCard_Least_in); +apply (simp, blast intro: InfCard_Least_in); +done + +lemma set_of_Ord_succ_Union: "(\y \ z. Ord(y)) ==> z \ succ(Union(z))" +apply (rule subsetI) +apply (case_tac "\y \ z. y \ x", blast ); +apply (simp, erule bexE); +apply (rule_tac i=xa and j=x in Ord_linear_le) +apply (blast dest: le_imp_subset elim: leE ltE)+ +done + +lemma subset_not_mem: "j \ i ==> i \ j" +by (fast elim!: mem_irrefl) + +lemma succ_Union_not_mem: + "(!!y. y \ z ==> Ord(y)) ==> succ(Union(z)) \ z" +apply (rule set_of_Ord_succ_Union [THEN subset_not_mem]); +apply blast +done + +lemma Union_cons_eq_succ_Union: + "Union(cons(succ(Union(z)),z)) = succ(Union(z))" +by fast + +lemma Un_Ord_disj: "[| Ord(i); Ord(j) |] ==> i Un j = i | i Un j = j" +by (fast dest!: le_imp_subset elim: Ord_linear_le) + +lemma Union_eq_Un: "x \ X ==> Union(X) = x Un Union(X-{x})" +by fast -AC16_lemmas = AC_Equiv + Hartog +lemma Union_in_lemma [rule_format]: + "n \ nat ==> \z. (\y \ z. Ord(y)) & z\n & z\0 --> Union(z) \ z" +apply (induct_tac "n") +apply (fast dest!: eqpoll_imp_lepoll [THEN lepoll_0_is_0]) +apply (intro allI impI) +apply (erule natE) +apply (fast dest!: eqpoll_1_iff_singleton [THEN iffD1] + intro!: Union_singleton) +apply (clarify ); +apply (elim not_emptyE) +apply (erule_tac x = "z-{xb}" in allE) +apply (erule impE) +apply (fast elim!: Diff_sing_eqpoll + Diff_sing_eqpoll [THEN eqpoll_succ_imp_not_empty]) +apply (subgoal_tac "xb \ \(z - {xb}) \ z"); +apply (simp add: Union_eq_Un [symmetric]); +apply (frule bspec, assumption) +apply (drule bspec); +apply (erule Diff_subset [THEN subsetD]); +apply (drule Un_Ord_disj, assumption) +apply (auto ); +done + +lemma Union_in: "[| \x \ z. Ord(x); z\n; z\0; n \ nat |] ==> Union(z) \ z" +apply (blast intro: Union_in_lemma); +done + +lemma succ_Union_in_x: + "[| InfCard(x); z \ Pow(x); z\n; n \ nat |] ==> succ(Union(z)) \ x" +apply (rule Limit_has_succ [THEN ltE]); +prefer 3 apply assumption +apply (erule InfCard_is_Limit) +apply (case_tac "z=0"); +apply (simp, fast intro!: InfCard_is_Limit [THEN Limit_has_0]); +apply (rule ltI [OF PowD [THEN subsetD] InfCard_is_Card [THEN Card_is_Ord]]); +apply assumption; +apply (blast intro: Union_in + InfCard_is_Card [THEN Card_is_Ord, THEN Ord_in_Ord])+ +done + +lemma succ_lepoll_succ_succ: + "[| InfCard(x); n \ nat |] + ==> {y \ Pow(x). y\succ(n)} \ {y \ Pow(x). y\succ(succ(n))}" +apply (unfold lepoll_def); +apply (rule_tac x = "\z \ {y\Pow(x). y\succ(n)}. cons(succ(Union(z)), z)" + in exI) +apply (rule_tac d = "%z. z-{Union (z) }" in lam_injective) +apply (blast intro!: succ_Union_in_x succ_Union_not_mem + intro: cons_eqpoll_succ Ord_in_Ord + dest!: InfCard_is_Card [THEN Card_is_Ord]) +apply (simp only: Union_cons_eq_succ_Union); +apply (rule cons_Diff_eq); +apply (fast dest!: InfCard_is_Card [THEN Card_is_Ord] + elim: Ord_in_Ord + intro!: succ_Union_not_mem); +done + +lemma subsets_eqpoll_X: + "[| InfCard(X); n \ nat |] ==> {Y \ Pow(X). Y\succ(n)} \ X" +apply (induct_tac "n") +apply (rule subsets_eqpoll_1_eqpoll) +apply (rule eqpollI) +apply (rule subsets_lepoll_lemma1 [THEN lepoll_trans], assumption+); +apply (rule eqpoll_trans [THEN eqpoll_imp_lepoll]); + apply (erule eqpoll_refl [THEN prod_eqpoll_cong]); +apply (erule InfCard_square_eqpoll) +apply (fast elim: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_trans] + intro!: succ_lepoll_succ_succ) +done + +lemma image_vimage_eq: + "[| f \ surj(A,B); y \ B |] ==> f``(converse(f)``y) = y" +apply (unfold surj_def) +apply (fast dest: apply_equality2 elim: apply_iff [THEN iffD2]) +done + +lemma vimage_image_eq: "[| f \ inj(A,B); y \ A |] ==> converse(f)``(f``y) = y" +apply (fast elim!: inj_is_fun [THEN apply_Pair] dest: inj_equality) +done + +lemma subsets_eqpoll: + "A\B ==> {Y \ Pow(A). Y\n}\{Y \ Pow(B). Y\n}" +apply (unfold eqpoll_def) +apply (erule exE) +apply (rule_tac x = "\X \ {Y \ Pow (A) . \f. f \ bij (Y, n) }. f``X" in exI) +apply (rule_tac d = "%Z. converse (f) ``Z" in lam_bijective) +apply (fast intro!: bij_is_inj [THEN restrict_bij, THEN bij_converse_bij, + THEN comp_bij] + elim!: bij_is_fun [THEN fun_is_rel, THEN image_subset]) +apply (blast intro!: bij_is_inj [THEN restrict_bij] + comp_bij bij_converse_bij + bij_is_fun [THEN fun_is_rel, THEN image_subset]) +apply (fast elim!: bij_is_inj [THEN vimage_image_eq]) +apply (fast elim!: bij_is_surj [THEN image_vimage_eq]) +done + +lemma WO2_imp_ex_Card: "WO2 ==> \a. Card(a) & X\a" +apply (unfold WO2_def) +apply (drule spec [of _ X]) +apply (blast intro: Card_cardinal eqpoll_trans + well_ord_Memrel [THEN well_ord_cardinal_eqpoll, THEN eqpoll_sym]) +done + +lemma lepoll_infinite: "[| X\Y; ~Finite(X) |] ==> ~Finite(Y)" +by (blast intro: lepoll_Finite) + +lemma infinite_Card_is_InfCard: "[| ~Finite(X); Card(X) |] ==> InfCard(X)" +apply (unfold InfCard_def) +apply (fast elim!: Card_is_Ord [THEN nat_le_infinite_Ord]) +done + +lemma WO2_infinite_subsets_eqpoll_X: "[| WO2; n \ nat; ~Finite(X) |] + ==> {Y \ Pow(X). Y\succ(n)}\X" +apply (drule WO2_imp_ex_Card) +apply (elim allE exE conjE); +apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption) +apply (drule infinite_Card_is_InfCard, assumption) +apply (blast intro: subsets_eqpoll subsets_eqpoll_X eqpoll_sym eqpoll_trans); +done + +lemma well_ord_imp_ex_Card: "well_ord(X,R) ==> \a. Card(a) & X\a" +by (fast elim!: well_ord_cardinal_eqpoll [THEN eqpoll_sym] + intro!: Card_cardinal) + +lemma well_ord_infinite_subsets_eqpoll_X: + "[| well_ord(X,R); n \ nat; ~Finite(X) |] ==> {Y \ Pow(X). Y\succ(n)}\X" +apply (drule well_ord_imp_ex_Card) +apply (elim allE exE conjE) +apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption) +apply (drule infinite_Card_is_InfCard, assumption) +apply (blast intro: subsets_eqpoll subsets_eqpoll_X eqpoll_sym eqpoll_trans); +done + +end diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC17_AC1.ML --- a/src/ZF/AC/AC17_AC1.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,310 +0,0 @@ -(* Title: ZF/AC/AC1_AC17.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -The equivalence of AC0, AC1 and AC17 - -Also, the proofs needed to show that each of AC2, AC3, ..., AC6 is equivalent -to AC0 and AC1. -*) - - -(** AC0 is equivalent to AC1. - AC0 comes from Suppes, AC1 from Rubin & Rubin **) - -Goal "[| f:(\\X \\ A. X); D \\ A |] ==> \\g. g:(\\X \\ D. X)"; -by (fast_tac (claset() addSIs [restrict_type, apply_type]) 1); -val lemma1 = result(); - -Goalw AC_defs "AC0 ==> AC1"; -by (blast_tac (claset() addIs [lemma1]) 1); -qed "AC0_AC1"; - -Goalw AC_defs "AC1 ==> AC0"; -by (Blast_tac 1); -qed "AC1_AC0"; - - -(**** The proof of AC1 ==> AC17 ****) - -Goal "f \\ (\\X \\ Pow(A) - {0}. X) ==> f \\ (Pow(A) - {0} -> A)"; -by (rtac Pi_type 1 THEN (assume_tac 1)); -by (dtac apply_type 1 THEN (assume_tac 1)); -by (Fast_tac 1); -val lemma1 = result(); - -Goalw AC_defs "AC1 ==> AC17"; -by (rtac allI 1); -by (rtac ballI 1); -by (eres_inst_tac [("x","Pow(A)-{0}")] allE 1); -by (etac impE 1); -by (Fast_tac 1); -by (etac exE 1); -by (rtac bexI 1); -by (etac lemma1 2); -by (rtac apply_type 1 THEN (assume_tac 1)); -by (fast_tac (claset() addSDs [lemma1] addSEs [apply_type]) 1); -qed "AC1_AC17"; - - -(**** The proof of AC17 ==> AC1 ****) - -(* *********************************************************************** *) -(* more properties of HH *) -(* *********************************************************************** *) - -Goal "[| x - (\\j \\ LEAST i. HH(\\X \\ Pow(x)-{0}. {f`X}, x, i) = {x}. \ -\ HH(\\X \\ Pow(x)-{0}. {f`X}, x, j)) = 0; \ -\ f \\ Pow(x)-{0} -> x |] \ -\ ==> \\r. well_ord(x,r)"; -by (rtac exI 1); -by (eresolve_tac [[bij_Least_HH_x RS bij_converse_bij RS bij_is_inj, - Ord_Least RS well_ord_Memrel] MRS well_ord_rvimage] 1); -by (assume_tac 1); -qed "UN_eq_imp_well_ord"; - -(* *********************************************************************** *) -(* theorems closer to the proof *) -(* *********************************************************************** *) - -Goalw AC_defs "~AC1 ==> \ -\ \\A. \\f \\ Pow(A)-{0} -> A. \\u \\ Pow(A)-{0}. f`u \\ u"; -by (etac swap 1); -by (rtac allI 1); -by (etac swap 1); -by (res_inst_tac [("x","Union(A)")] exI 1); -by (rtac ballI 1); -by (etac swap 1); -by (rtac impI 1); -by (fast_tac (claset() addSIs [restrict_type]) 1); -qed "not_AC1_imp_ex"; - -Goal "[| \\f \\ Pow(x) - {0} -> x. \\u \\ Pow(x) - {0}. f`u\\u; \ -\ \\f \\ Pow(x)-{0}->x. \ -\ x - (\\a \\ (LEAST i. HH(\\X \\ Pow(x)-{0}. {f`X},x,i)={x}). \ -\ HH(\\X \\ Pow(x)-{0}. {f`X},x,a)) = 0 |] \ -\ ==> P"; -by (etac bexE 1); -by (eresolve_tac [UN_eq_imp_well_ord RS exE] 1 THEN (assume_tac 1)); -by (eresolve_tac [ex_choice_fun_Pow RS exE] 1); -by (etac ballE 1); -by (fast_tac (FOL_cs addEs [bexE, notE, apply_type]) 1); -by (etac notE 1); -by (rtac Pi_type 1 THEN (assume_tac 1)); -by (resolve_tac [apply_type RSN (2, subsetD)] 1 THEN TRYALL assume_tac); -by (Fast_tac 1); -val lemma1 = result(); - -Goal "~ (\\f \\ Pow(x)-{0}->x. x - F(f) = 0) \ -\ ==> (\\f \\ Pow(x)-{0}->x. x - F(f)) \ -\ \\ (Pow(x) -{0} -> x) -> Pow(x) - {0}"; -by (fast_tac (claset() addSIs [lam_type] addSDs [Diff_eq_0_iff RS iffD1]) 1); -val lemma2 = result(); - -Goal "[| f`Z \\ Z; Z \\ Pow(x)-{0} |] ==> \ -\ (\\X \\ Pow(x)-{0}. {f`X})`Z \\ Pow(Z)-{0}"; -by Auto_tac; -val lemma3 = result(); - -Goal "\\f \\ F. f`((\\f \\ F. Q(f))`f) \\ (\\f \\ F. Q(f))`f \ -\ ==> \\f \\ F. f`Q(f) \\ Q(f)"; -by (Asm_full_simp_tac 1); -val lemma4 = result(); - -Goalw [AC17_def] "AC17 ==> AC1"; -by (rtac classical 1); -by (eresolve_tac [not_AC1_imp_ex RS exE] 1); -by (excluded_middle_tac - "\\f \\ Pow(x)-{0}->x. \ -\ x - (\\a \\ (LEAST i. HH(\\X \\ Pow(x)-{0}. {f`X},x,i)={x}). \ -\ HH(\\X \\ Pow(x)-{0}. {f`X},x,a)) = 0" 1); -by (etac lemma1 2 THEN (assume_tac 2)); -by (dtac lemma2 1); -by (etac allE 1); -by (dtac bspec 1 THEN (assume_tac 1)); -by (dtac lemma4 1); -by (etac bexE 1); -by (dtac apply_type 1 THEN (assume_tac 1)); -by (dresolve_tac [beta RS sym RSN (2, subst_elem)] 1); -by (assume_tac 1); -by (dtac lemma3 1 THEN (assume_tac 1)); -by (fast_tac (claset() addSDs [HH_Least_eq_x RS sym RSN (2, subst_elem), - f_subset_imp_HH_subset] addSEs [mem_irrefl]) 1); -qed "AC17_AC1"; - - -(* ********************************************************************** - AC1 ==> AC2 ==> AC1 - AC1 ==> AC4 ==> AC3 ==> AC1 - AC4 ==> AC5 ==> AC4 - AC1 <-> AC6 -************************************************************************* *) - -(* ********************************************************************** *) -(* AC1 ==> AC2 *) -(* ********************************************************************** *) - -Goal "[| f:(\\X \\ A. X); B \\ A; 0\\A |] ==> {f`B} \\ B Int {f`C. C \\ A}"; -by (fast_tac (claset() addSEs [apply_type]) 1); -val lemma1 = result(); - -Goalw [pairwise_disjoint_def] - "[| pairwise_disjoint(A); B \\ A; C \\ A; D \\ B; D \\ C |] ==> f`B = f`C"; -by (Fast_tac 1); -val lemma2 = result(); - -Goalw AC_defs "AC1 ==> AC2"; -by (rtac allI 1); -by (rtac impI 1); -by (REPEAT (eresolve_tac [asm_rl,conjE,allE,exE,impE] 1)); -by (REPEAT (resolve_tac [exI,ballI,equalityI] 1)); -by (rtac lemma1 2 THEN (REPEAT (assume_tac 2))); -by (fast_tac (claset() addSEs [lemma2] addEs [apply_type]) 1); -qed "AC1_AC2"; - - -(* ********************************************************************** *) -(* AC2 ==> AC1 *) -(* ********************************************************************** *) - -Goal "0\\A ==> 0 \\ {B*{B}. B \\ A}"; -by (fast_tac (claset() addSDs [sym RS (Sigma_empty_iff RS iffD1)]) 1); -val lemma1 = result(); - -Goal "[| X*{X} Int C = {y}; X \\ A |] \ -\ ==> (THE y. X*{X} Int C = {y}): X*A"; -by (rtac subst_elem 1); -by (fast_tac (claset() addSIs [the_equality] - addSEs [sym RS trans RS (singleton_eq_iff RS iffD1)]) 2); -by (blast_tac (claset() addSEs [equalityE]) 1); -val lemma2 = result(); - -Goal "\\D \\ {E*{E}. E \\ A}. \\y. D Int C = {y} \ -\ ==> (\\x \\ A. fst(THE z. (x*{x} Int C = {z}))) \\ (\\X \\ A. X)"; -by (fast_tac (claset() addSEs [lemma2] - addSIs [lam_type, RepFunI, fst_type]) 1); -val lemma3 = result(); - -Goalw (AC_defs@AC_aux_defs) "AC2 ==> AC1"; -by (REPEAT (resolve_tac [allI, impI] 1)); -by (REPEAT (eresolve_tac [allE, impE] 1)); -by (fast_tac (claset() addSEs [lemma3]) 2); -by (fast_tac (claset() addSIs [lemma1, equals0I]) 1); -qed "AC2_AC1"; - - -(* ********************************************************************** *) -(* AC1 ==> AC4 *) -(* ********************************************************************** *) - -Goal "0 \\ {R``{x}. x \\ domain(R)}"; -by (Blast_tac 1); -val lemma = result(); - -Goalw AC_defs "AC1 ==> AC4"; -by (REPEAT (resolve_tac [allI, impI] 1)); -by (REPEAT (eresolve_tac [allE, lemma RSN (2, impE), exE] 1)); -by (best_tac (claset() addSIs [lam_type] addSEs [apply_type]) 1); -qed "AC1_AC4"; - - -(* ********************************************************************** *) -(* AC4 ==> AC3 *) -(* ********************************************************************** *) - -Goal "f \\ A->B ==> (\\z \\ A. {z}*f`z) \\ A*Union(B)"; -by (fast_tac (claset() addSDs [apply_type]) 1); -val lemma1 = result(); - -Goal "domain(\\z \\ A. {z}*f(z)) = {a \\ A. f(a)\\0}"; -by (Blast_tac 1); -val lemma2 = result(); - -Goal "x \\ A ==> (\\z \\ A. {z}*f(z))``{x} = f(x)"; -by (Fast_tac 1); -val lemma3 = result(); - -Goalw AC_defs "AC4 ==> AC3"; -by (REPEAT (resolve_tac [allI,ballI] 1)); -by (REPEAT (eresolve_tac [allE,impE] 1)); -by (etac lemma1 1); -by (asm_full_simp_tac (simpset() addsimps [lemma2, lemma3] - addcongs [Pi_cong]) 1); -qed "AC4_AC3"; - -(* ********************************************************************** *) -(* AC3 ==> AC1 *) -(* ********************************************************************** *) - -Goal "b\\A ==> (\\x \\ {a \\ A. id(A)`a\\b}. id(A)`x) = (\\x \\ A. x)"; -by (asm_full_simp_tac (simpset() addsimps [id_def] addcongs [Pi_cong]) 1); -by (res_inst_tac [("b","A")] subst_context 1); -by (Fast_tac 1); -val lemma = result(); - -Goalw AC_defs "AC3 ==> AC1"; -by (fast_tac (claset() addSIs [id_type] addEs [lemma RS subst]) 1); -qed "AC3_AC1"; - -(* ********************************************************************** *) -(* AC4 ==> AC5 *) -(* ********************************************************************** *) - -Goalw (range_def::AC_defs) "AC4 ==> AC5"; -by (REPEAT (resolve_tac [allI,ballI] 1)); -by (REPEAT (eresolve_tac [allE,impE] 1)); -by (eresolve_tac [fun_is_rel RS converse_type] 1); -by (etac exE 1); -by (rtac bexI 1); -by (rtac Pi_type 2 THEN (assume_tac 2)); -by (fast_tac (claset() addSDs [apply_type] - addSEs [fun_is_rel RS converse_type RS subsetD RS SigmaD2]) 2); -by (rtac ballI 1); -by (rtac apply_equality 1 THEN (assume_tac 2)); -by (etac domainE 1); -by (ftac range_type 1 THEN (assume_tac 1)); -by (fast_tac (claset() addDs [apply_equality]) 1); -qed "AC4_AC5"; - - -(* ********************************************************************** *) -(* AC5 ==> AC4, Rubin & Rubin, p. 11 *) -(* ********************************************************************** *) - -Goal "R \\ A*B ==> (\\x \\ R. fst(x)) \\ R -> A"; -by (fast_tac (claset() addSIs [lam_type, fst_type]) 1); -val lemma1 = result(); - -Goalw [range_def] "R \\ A*B ==> range(\\x \\ R. fst(x)) = domain(R)"; -by (force_tac (claset() addIs [lamI RS subst_elem] addSEs [lamE], - simpset()) 1); -val lemma2 = result(); - -Goal "[| \\f \\ A->C. P(f,domain(f)); A=B |] ==> \\f \\ B->C. P(f,B)"; -by (etac bexE 1); -by (ftac domain_of_fun 1); -by (Fast_tac 1); -val lemma3 = result(); - -Goal "[| R \\ A*B; g \\ C->R; \\x \\ C. (\\z \\ R. fst(z))` (g`x) = x |] \ -\ ==> (\\x \\ C. snd(g`x)): (\\x \\ C. R``{x})"; -by (rtac lam_type 1); -by (force_tac (claset() addDs [apply_type], simpset()) 1); -val lemma4 = result(); - -Goalw AC_defs "AC5 ==> AC4"; -by (Clarify_tac 1); -by (REPEAT (eresolve_tac [allE,ballE] 1)); -by (eresolve_tac [lemma1 RSN (2, notE)] 2 THEN (assume_tac 2)); -by (dresolve_tac [lemma2 RSN (2, lemma3)] 1 THEN (assume_tac 1)); -by (fast_tac (claset() addSEs [lemma4]) 1); -qed "AC5_AC4"; - - -(* ********************************************************************** *) -(* AC1 <-> AC6 *) -(* ********************************************************************** *) - -Goalw AC_defs "AC1 <-> AC6"; -by (Blast_tac 1); -qed "AC1_iff_AC6"; diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC17_AC1.thy --- a/src/ZF/AC/AC17_AC1.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/AC17_AC1.thy Wed Jan 16 17:52:06 2002 +0100 @@ -1,3 +1,300 @@ -(*Dummy theory to document dependencies *) +(* Title: ZF/AC/AC1_AC17.thy + ID: $Id$ + Author: Krzysztof Grabczewski + +The equivalence of AC0, AC1 and AC17 + +Also, the proofs needed to show that each of AC2, AC3, ..., AC6 is equivalent +to AC0 and AC1. +*) + +theory AC17_AC1 = HH: + + +(** AC0 is equivalent to AC1. + AC0 comes from Suppes, AC1 from Rubin & Rubin **) + +lemma AC0_AC1_lemma: "[| f:(\X \ A. X); D \ A |] ==> \g. g:(\X \ D. X)" +by (fast intro!: restrict_type apply_type) + +lemma AC0_AC1: "AC0 ==> AC1" +apply (unfold AC0_def AC1_def) +apply (blast intro: AC0_AC1_lemma) +done + +lemma AC1_AC0: "AC1 ==> AC0" +by (unfold AC0_def AC1_def, blast) + + +(**** The proof of AC1 ==> AC17 ****) + +lemma AC1_AC17_lemma: "f \ (\X \ Pow(A) - {0}. X) ==> f \ (Pow(A) - {0} -> A)" +apply (rule Pi_type, assumption) +apply (drule apply_type, assumption, fast) +done + +lemma AC1_AC17: "AC1 ==> AC17" +apply (unfold AC1_def AC17_def) +apply (rule allI) +apply (rule ballI) +apply (erule_tac x = "Pow (A) -{0}" in allE) +apply (erule impE, fast) +apply (erule exE) +apply (rule bexI) +apply (erule_tac [2] AC1_AC17_lemma) +apply (rule apply_type, assumption) +apply (fast dest!: AC1_AC17_lemma elim!: apply_type) +done + + +(**** The proof of AC17 ==> AC1 ****) + +(* *********************************************************************** *) +(* more properties of HH *) +(* *********************************************************************** *) + +lemma UN_eq_imp_well_ord: + "[| x - (\j \ LEAST i. HH(\X \ Pow(x)-{0}. {f`X}, x, i) = {x}. + HH(\X \ Pow(x)-{0}. {f`X}, x, j)) = 0; + f \ Pow(x)-{0} -> x |] + ==> \r. well_ord(x,r)" +apply (rule exI) +apply (erule well_ord_rvimage + [OF bij_Least_HH_x [THEN bij_converse_bij, THEN bij_is_inj] + Ord_Least [THEN well_ord_Memrel]], assumption) +done + +(* *********************************************************************** *) +(* theorems closer to the proof *) +(* *********************************************************************** *) + +lemma not_AC1_imp_ex: + "~AC1 ==> \A. \f \ Pow(A)-{0} -> A. \u \ Pow(A)-{0}. f`u \ u" +apply (unfold AC1_def) +apply (erule swap) +apply (rule allI) +apply (erule swap) +apply (rule_tac x = "Union (A)" in exI) +apply (blast intro: restrict_type) +done + +lemma lemma1: + "[| \f \ Pow(x) - {0} -> x. \u \ Pow(x) - {0}. f`u\u; + \f \ Pow(x)-{0}->x. + x - (\a \ (LEAST i. HH(\X \ Pow(x)-{0}. {f`X},x,i)={x}). + HH(\X \ Pow(x)-{0}. {f`X},x,a)) = 0 |] + ==> P" +apply (erule bexE) +apply (erule UN_eq_imp_well_ord [THEN exE], assumption) +apply (erule ex_choice_fun_Pow [THEN exE]) +apply (erule ballE) +apply (fast intro: apply_type del: DiffE) +apply (erule notE) +apply (rule Pi_type, assumption) +apply (blast dest: apply_type) +done + +lemma lemma2: + "~ (\f \ Pow(x)-{0}->x. x - F(f) = 0) + ==> (\f \ Pow(x)-{0}->x . x - F(f)) + \ (Pow(x) -{0} -> x) -> Pow(x) - {0}" +by (fast intro!: lam_type dest!: Diff_eq_0_iff [THEN iffD1]) + +lemma lemma3: + "[| f`Z \ Z; Z \ Pow(x)-{0} |] + ==> (\X \ Pow(x)-{0}. {f`X})`Z \ Pow(Z)-{0}" +by auto + +lemma lemma4: + "\f \ F. f`((\f \ F. Q(f))`f) \ (\f \ F. Q(f))`f + ==> \f \ F. f`Q(f) \ Q(f)" +by simp + +lemma AC17_AC1: "AC17 ==> AC1" +apply (unfold AC17_def) +apply (rule classical) +apply (erule not_AC1_imp_ex [THEN exE]) +apply (case_tac + "\f \ Pow(x)-{0} -> x. + x - (\a \ (LEAST i. HH (\X \ Pow (x) -{0}. {f`X},x,i) ={x}) . HH (\X \ Pow (x) -{0}. {f`X},x,a)) = 0") +apply (erule lemma1, assumption) +apply (drule lemma2) +apply (erule allE) +apply (drule bspec, assumption) +apply (drule lemma4) +apply (erule bexE) +apply (drule apply_type, assumption) +apply (simp add: HH_Least_eq_x del: Diff_iff ) +apply (drule lemma3, assumption) +apply (fast dest!: subst_elem [OF _ HH_Least_eq_x [symmetric]] + f_subset_imp_HH_subset elim!: mem_irrefl) +done + + +(* ********************************************************************** + AC1 ==> AC2 ==> AC1 + AC1 ==> AC4 ==> AC3 ==> AC1 + AC4 ==> AC5 ==> AC4 + AC1 <-> AC6 +************************************************************************* *) + +(* ********************************************************************** *) +(* AC1 ==> AC2 *) +(* ********************************************************************** *) + +lemma lemma1: + "[| f:(\X \ A. X); B \ A; 0\A |] ==> {f`B} \ B Int {f`C. C \ A}" +by (fast elim!: apply_type) -AC17_AC1 = HH +lemma lemma2: + "[| pairwise_disjoint(A); B \ A; C \ A; D \ B; D \ C |] ==> f`B = f`C" +by (unfold pairwise_disjoint_def, fast) + +lemma AC1_AC2: "AC1 ==> AC2" +apply (unfold AC1_def AC2_def) +apply (rule allI) +apply (rule impI) +apply (elim asm_rl conjE allE exE impE, assumption) +apply (intro exI ballI equalityI) +prefer 2 apply (rule lemma1, assumption+) +apply (fast elim!: lemma2 elim: apply_type) +done + + +(* ********************************************************************** *) +(* AC2 ==> AC1 *) +(* ********************************************************************** *) + +lemma lemma1: "0\A ==> 0 \ {B*{B}. B \ A}" +by (fast dest!: sym [THEN Sigma_empty_iff [THEN iffD1]]) + +lemma lemma2: "[| X*{X} Int C = {y}; X \ A |] + ==> (THE y. X*{X} Int C = {y}): X*A" +apply (rule subst_elem [of y]) +apply (blast elim!: equalityE) +apply (auto simp add: singleton_eq_iff) +done + +lemma lemma3: + "\D \ {E*{E}. E \ A}. \y. D Int C = {y} + ==> (\x \ A. fst(THE z. (x*{x} Int C = {z}))) \ (\X \ A. X)" +apply (rule lam_type) +apply (drule bspec, blast) +apply (blast intro: lemma2 fst_type) +done + +lemma AC2_AC1: "AC2 ==> AC1" +apply (unfold AC1_def AC2_def pairwise_disjoint_def) +apply (intro allI impI) +apply (elim allE impE) +prefer 2 apply (fast elim!: lemma3) +apply (blast intro!: lemma1) +done + + +(* ********************************************************************** *) +(* AC1 ==> AC4 *) +(* ********************************************************************** *) + +lemma empty_notin_images: "0 \ {R``{x}. x \ domain(R)}" +by blast + +lemma AC1_AC4: "AC1 ==> AC4" +apply (unfold AC1_def AC4_def) +apply (intro allI impI) +apply (drule spec, drule mp [OF _ empty_notin_images]) +apply (best intro!: lam_type elim!: apply_type) +done + + +(* ********************************************************************** *) +(* AC4 ==> AC3 *) +(* ********************************************************************** *) + +lemma lemma1: "f \ A->B ==> (\z \ A. {z}*f`z) \ A*Union(B)" +by (fast dest!: apply_type) + +lemma lemma2: "domain(\z \ A. {z}*f(z)) = {a \ A. f(a)\0}" +by blast + +lemma lemma3: "x \ A ==> (\z \ A. {z}*f(z))``{x} = f(x)" +by fast + +lemma AC4_AC3: "AC4 ==> AC3" +apply (unfold AC3_def AC4_def) +apply (intro allI ballI) +apply (elim allE impE) +apply (erule lemma1) +apply (simp add: lemma2 lemma3 cong add: Pi_cong) +done + +(* ********************************************************************** *) +(* AC3 ==> AC1 *) +(* ********************************************************************** *) + +lemma AC3_AC1_lemma: + "b\A ==> (\x \ {a \ A. id(A)`a\b}. id(A)`x) = (\x \ A. x)" +apply (simp add: id_def cong add: Pi_cong) +apply (rule_tac b = "A" in subst_context, fast) +done + +lemma AC3_AC1: "AC3 ==> AC1" +apply (unfold AC1_def AC3_def) +apply (fast intro!: id_type elim: AC3_AC1_lemma [THEN subst]) +done + +(* ********************************************************************** *) +(* AC4 ==> AC5 *) +(* ********************************************************************** *) + +lemma AC4_AC5: "AC4 ==> AC5" +apply (unfold range_def AC4_def AC5_def) +apply (intro allI ballI) +apply (elim allE impE) +apply (erule fun_is_rel [THEN converse_type]) +apply (erule exE) +apply (rename_tac g) +apply (rule_tac x=g in bexI) +apply (blast dest: apply_equality range_type) +apply (blast intro: Pi_type dest: apply_type fun_is_rel) +done + + +(* ********************************************************************** *) +(* AC5 ==> AC4, Rubin & Rubin, p. 11 *) +(* ********************************************************************** *) + +lemma lemma1: "R \ A*B ==> (\x \ R. fst(x)) \ R -> A" +by (fast intro!: lam_type fst_type) + +lemma lemma2: "R \ A*B ==> range(\x \ R. fst(x)) = domain(R)" +by (unfold lam_def, force) + +lemma lemma3: "[| \f \ A->C. P(f,domain(f)); A=B |] ==> \f \ B->C. P(f,B)" +apply (erule bexE) +apply (frule domain_of_fun, fast) +done + +lemma lemma4: "[| R \ A*B; g \ C->R; \x \ C. (\z \ R. fst(z))` (g`x) = x |] + ==> (\x \ C. snd(g`x)): (\x \ C. R``{x})" +apply (rule lam_type) +apply (force dest: apply_type) +done + +lemma AC5_AC4: "AC5 ==> AC4" +apply (unfold AC4_def AC5_def, clarify) +apply (elim allE ballE) +apply (drule lemma3 [OF _ lemma2], assumption) +apply (fast elim!: lemma4) +apply (blast intro: lemma1) +done + + +(* ********************************************************************** *) +(* AC1 <-> AC6 *) +(* ********************************************************************** *) + +lemma AC1_iff_AC6: "AC1 <-> AC6" +by (unfold AC1_def AC6_def, blast) + +end diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC18_AC19.ML --- a/src/ZF/AC/AC18_AC19.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,115 +0,0 @@ -(* Title: ZF/AC/AC18_AC19.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -The proof of AC1 ==> AC18 ==> AC19 ==> AC1 -*) - -(* ********************************************************************** *) -(* AC1 ==> AC18 *) -(* ********************************************************************** *) - -Goal "[| f \\ (\\b \\ {P(a). a \\ A}. b); \\a \\ A. P(a)<=Q(a) |] \ -\ ==> (\\a \\ A. f`P(a)) \\ (\\a \\ A. Q(a))"; -by (rtac lam_type 1); -by (dtac apply_type 1); -by Auto_tac; -qed "PROD_subsets"; - -Goal "[| \\A. 0 \\ A --> (\\f. f \\ (\\X \\ A. X)); A \\ 0 |] ==> \ -\ (\\a \\ A. \\b \\ B(a). X(a, b)) \\ (\\f \\ \\a \\ A. B(a). \\a \\ A. X(a, f`a))"; -by (rtac subsetI 1); -by (eres_inst_tac [("x","{{b \\ B(a). x \\ X(a,b)}. a \\ A}")] allE 1); -by (etac impE 1); -by (Fast_tac 1); -by (etac exE 1); -by (rtac UN_I 1); -by (fast_tac (claset() addSEs [PROD_subsets]) 1); -by (Simp_tac 1); -by (fast_tac (claset() addSEs [not_emptyE] - addDs [RepFunI RSN (2, apply_type)]) 1); -qed "lemma_AC18"; - -val [prem] = goalw thy (AC18_def::AC_defs) "AC1 ==> AC18"; -by (resolve_tac [prem RS revcut_rl] 1); -by (fast_tac (claset() addSEs [lemma_AC18, not_emptyE, apply_type] - addSIs [equalityI, INT_I, UN_I]) 1); -qed "AC1_AC18"; - -(* ********************************************************************** *) -(* AC18 ==> AC19 *) -(* ********************************************************************** *) - -val [prem] = goalw thy [AC18_def, AC19_def] "AC18 ==> AC19"; -by (rtac allI 1); -by (res_inst_tac [("B1","%x. x")] (forall_elim_vars 0 prem RS revcut_rl) 1); -by (Fast_tac 1); -qed "AC18_AC19"; - -(* ********************************************************************** *) -(* AC19 ==> AC1 *) -(* ********************************************************************** *) - -Goalw [u_def] - "[| A \\ 0; 0 \\ A |] ==> {u_(a). a \\ A} \\ 0 & 0 \\ {u_(a). a \\ A}"; -by (fast_tac (claset() addSIs [not_emptyI] - addSEs [not_emptyE] - addSDs [sym RS (RepFun_eq_0_iff RS iffD1)]) 1); -qed "RepRep_conj"; - -Goal "[|c \\ a; x = c Un {0}; x \\ a |] ==> x - {0} \\ a"; -by (hyp_subst_tac 1); -by (rtac subst_elem 1 THEN (assume_tac 1)); -by (rtac equalityI 1); -by (Fast_tac 1); -by (rtac subsetI 1); -by (excluded_middle_tac "x=0" 1); -by (Fast_tac 1); -by (fast_tac (claset() addEs [notE, subst_elem]) 1); -val lemma1_1 = result(); - -Goalw [u_def] - "[| f`(u_(a)) \\ a; f \\ (\\B \\ {u_(a). a \\ A}. B); a \\ A |] \ -\ ==> f`(u_(a))-{0} \\ a"; -by (fast_tac (claset() addSEs [lemma1_1] addSDs [apply_type]) 1); -val lemma1_2 = result(); - -Goal "\\f. f \\ (\\B \\ {u_(a). a \\ A}. B) ==> \\f. f \\ (\\B \\ A. B)"; -by (etac exE 1); -by (res_inst_tac - [("x","\\a \\ A. if(f`(u_(a)) \\ a, f`(u_(a)), f`(u_(a))-{0})")] exI 1); -by (rtac lam_type 1); -by (split_tac [split_if] 1); -by (rtac conjI 1); -by (Fast_tac 1); -by (fast_tac (claset() addSEs [lemma1_2]) 1); -val lemma1 = result(); - -Goalw [u_def] "a\\0 ==> 0 \\ (\\b \\ u_(a). b)"; -by (fast_tac (claset() addSEs [not_emptyE] addSIs [UN_I, RepFunI]) 1); -val lemma2_1 = result(); - -Goal "[| A\\0; 0\\A |] ==> (\\x \\ {u_(a). a \\ A}. \\b \\ x. b) \\ 0"; -by (etac not_emptyE 1); -by (res_inst_tac [("a","0")] not_emptyI 1); -by (fast_tac (claset() addSIs [lemma2_1]) 1); -val lemma2 = result(); - -Goal "(\\f \\ F. P(f)) \\ 0 ==> F \\ 0"; -by (fast_tac (claset() addSEs [not_emptyE]) 1); -val lemma3 = result(); - -Goalw AC_defs "AC19 ==> AC1"; -by (Clarify_tac 1); -by (case_tac "A=0" 1); -by (Force_tac 1); -by (eres_inst_tac [("x","{u_(a). a \\ A}")] allE 1); -by (etac impE 1); -by (etac RepRep_conj 1 THEN (assume_tac 1)); -by (rtac lemma1 1); -by (dtac lemma2 1 THEN (assume_tac 1)); -by (dres_inst_tac [("P","%x. x\\0")] subst 1 THEN (assume_tac 1)); -by (fast_tac (claset() addSEs [lemma3 RS not_emptyE]) 1); -qed "AC19_AC1"; - - diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC18_AC19.thy --- a/src/ZF/AC/AC18_AC19.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/AC18_AC19.thy Wed Jan 16 17:52:06 2002 +0100 @@ -2,17 +2,106 @@ ID: $Id$ Author: Krzysztof Grabczewski -Additional definition used in the proof AC19 ==> AC1 which is a part -of the chain AC1 ==> AC18 ==> AC19 ==> AC1 +The proof of AC1 ==> AC18 ==> AC19 ==> AC1 *) -AC18_AC19 = AC_Equiv + +theory AC18_AC19 = AC_Equiv: + +constdefs + uu :: "i => i" + "uu(a) == {c Un {0}. c \ a}" + + +(* ********************************************************************** *) +(* AC1 ==> AC18 *) +(* ********************************************************************** *) + +lemma PROD_subsets: + "[| f \ (\b \ {P(a). a \ A}. b); \a \ A. P(a)<=Q(a) |] + ==> (\a \ A. f`P(a)) \ (\a \ A. Q(a))" +by (rule lam_type, drule apply_type, auto) + +lemma lemma_AC18: + "[| \A. 0 \ A --> (\f. f \ (\X \ A. X)); A \ 0 |] + ==> (\a \ A. \b \ B(a). X(a, b)) \ + (\f \ \a \ A. B(a). \a \ A. X(a, f`a))" +apply (rule subsetI) +apply (erule_tac x = "{{b \ B (a) . x \ X (a,b) }. a \ A}" in allE) +apply (erule impE, fast) +apply (erule exE) +apply (rule UN_I) + apply (fast elim!: PROD_subsets) +apply (simp, fast elim!: not_emptyE dest: apply_type [OF _ RepFunI]) +done + +lemma AC1_AC18: "AC1 ==> AC18" +apply (unfold AC1_def AC18_def) +apply (fast elim!: lemma_AC18 apply_type intro!: equalityI INT_I UN_I) +done + +(* ********************************************************************** *) +(* AC18 ==> AC19 *) +(* ********************************************************************** *) + +text{*Hard to express because of the need for meta-quantifiers in AC18*} +lemma "AC18 ==> AC19" +proof - + assume ac18 [unfolded AC18_def, norm_hhf]: AC18 + show AC19 + apply (unfold AC18_def AC19_def) + apply (intro allI impI) + apply (rule ac18 [of _ "%x. x", THEN mp], blast) + done +qed -consts - u_ :: i => i - -defs +(* ********************************************************************** *) +(* AC19 ==> AC1 *) +(* ********************************************************************** *) + +lemma RepRep_conj: + "[| A \ 0; 0 \ A |] ==> {uu(a). a \ A} \ 0 & 0 \ {uu(a). a \ A}" +apply (unfold uu_def, auto) +apply (blast dest!: sym [THEN RepFun_eq_0_iff [THEN iffD1]]) +done + +lemma lemma1_1: "[|c \ a; x = c Un {0}; x \ a |] ==> x - {0} \ a" +apply clarify +apply (rule subst_elem , (assumption)) +apply (fast elim: notE subst_elem) +done + +lemma lemma1_2: + "[| f`(uu(a)) \ a; f \ (\B \ {uu(a). a \ A}. B); a \ A |] + ==> f`(uu(a))-{0} \ a" +apply (unfold uu_def, fast elim!: lemma1_1 dest!: apply_type) +done - u_def "u_(a) == {c Un {0}. c \\ a}" +lemma lemma1: "\f. f \ (\B \ {uu(a). a \ A}. B) ==> \f. f \ (\B \ A. B)" +apply (erule exE) +apply (rule_tac x = "\a\A. if (f` (uu(a)) \ a, f` (uu(a)), f` (uu(a))-{0})" + in exI) +apply (rule lam_type) +apply (simp add: lemma1_2) +done + +lemma lemma2_1: "a\0 ==> 0 \ (\b \ uu(a). b)" +by (unfold uu_def, auto) + +lemma lemma2: "[| A\0; 0\A |] ==> (\x \ {uu(a). a \ A}. \b \ x. b) \ 0" +apply (erule not_emptyE) +apply (rule_tac a = "0" in not_emptyI) +apply (fast intro!: lemma2_1) +done + +lemma AC19_AC1: "AC19 ==> AC1" +apply (unfold AC19_def AC1_def, clarify) +apply (case_tac "A=0", force) +apply (erule_tac x = "{uu (a) . a \ A}" in allE) +apply (erule impE) +apply (erule RepRep_conj , (assumption)) +apply (rule lemma1) +apply (drule lemma2 , (assumption)) +apply auto +done end diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC1_WO2.ML --- a/src/ZF/AC/AC1_WO2.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,22 +0,0 @@ -(* Title: ZF/AC/AC1_WO2.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -The proof of AC1 ==> WO2 -*) - -(*Establishing the existence of a bijection -- hence the need for uresult*) -val [prem] = goal thy "f \\ (\\X \\ Pow(x) - {0}. X) ==> \ -\ ?g(f) \\ bij(x, LEAST i. HH(\\X \\ Pow(x)-{0}. {f`X}, x, i) = {x})"; -by (resolve_tac [bij_Least_HH_x RS bij_converse_bij] 1); -by (rtac f_subsets_imp_UN_HH_eq_x 1); -by (resolve_tac [lam_type RS apply_type] 1 THEN (assume_tac 2)); -by (fast_tac (claset() addSDs [prem RS apply_type]) 1); -by (fast_tac (claset() addSIs [prem RS Pi_weaken_type]) 1); -val lemma1 = uresult() |> standard; - -Goalw [AC1_def, WO2_def, eqpoll_def] "AC1 ==> WO2"; -by (rtac allI 1); -by (eres_inst_tac [("x","Pow(A)-{0}")] allE 1); -by (fast_tac (claset() addSDs [lemma1] addSIs [Ord_Least]) 1); -qed "AC1_WO2"; diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC1_WO2.thy --- a/src/ZF/AC/AC1_WO2.thy Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,3 +0,0 @@ -(*Dummy theory to document dependencies *) - -AC1_WO2 = HH diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC7_AC9.ML --- a/src/ZF/AC/AC7_AC9.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,188 +0,0 @@ -(* Title: ZF/AC/AC7-AC9.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -The proofs needed to state that AC7, AC8 and AC9 are equivalent to the previous -instances of AC. -*) - -(* ********************************************************************** *) -(* Lemmas used in the proofs AC7 ==> AC6 and AC9 ==> AC1 *) -(* - Sigma_fun_space_not0 *) -(* - Sigma_fun_space_eqpoll *) -(* ********************************************************************** *) - -Goal "[| 0\\A; B \\ A |] ==> (nat->Union(A)) * B \\ 0"; -by (blast_tac (claset() addSDs [Sigma_empty_iff RS iffD1, - Union_empty_iff RS iffD1]) 1); -qed "Sigma_fun_space_not0"; - -Goalw [inj_def] - "C \\ A ==> (\\g \\ (nat->Union(A))*C. \ -\ (\\n \\ nat. if(n=0, snd(g), fst(g)`(n #- 1)))) \ -\ \\ inj((nat->Union(A))*C, (nat->Union(A)) ) "; -by (rtac CollectI 1); -by (fast_tac (claset() addSIs [lam_type,RepFunI,if_type,snd_type,apply_type, - fst_type,diff_type,nat_succI,nat_0I]) 1); -by (REPEAT (resolve_tac [ballI, impI] 1)); -by (Asm_full_simp_tac 1); -by (REPEAT (etac SigmaE 1)); -by (REPEAT (hyp_subst_tac 1)); -by (Asm_full_simp_tac 1); -by (rtac conjI 1); -by (dresolve_tac [nat_0I RSN (2, lam_eqE)] 2); -by (Asm_full_simp_tac 2); -by (rtac fun_extension 1 THEN REPEAT (assume_tac 1)); -by (dresolve_tac [nat_succI RSN (2, lam_eqE)] 1 THEN (assume_tac 1)); -by (asm_full_simp_tac (simpset() addsimps [succ_not_0 RS if_not_P]) 1); -val lemma = result(); - -Goal "[| C \\ A; 0\\A |] ==> (nat->Union(A)) * C eqpoll (nat->Union(A))"; -by (rtac eqpollI 1); -by (fast_tac (claset() addSEs [prod_lepoll_self, not_sym RS not_emptyE, - subst_elem] addEs [swap]) 2); -by (rewtac lepoll_def); -by (fast_tac (claset() addSIs [lemma]) 1); -qed "Sigma_fun_space_eqpoll"; - - -(* ********************************************************************** *) -(* AC6 ==> AC7 *) -(* ********************************************************************** *) - -Goalw AC_defs "AC6 ==> AC7"; -by (Blast_tac 1); -qed "AC6_AC7"; - -(* ********************************************************************** *) -(* AC7 ==> AC6, Rubin & Rubin p. 12, Theorem 2.8 *) -(* The case of the empty family of sets added in order to complete *) -(* the proof. *) -(* ********************************************************************** *) - -Goal "y \\ (\\B \\ A. Y*B) ==> (\\B \\ A. snd(y`B)): (\\B \\ A. B)"; -by (fast_tac (claset() addSIs [lam_type, snd_type, apply_type]) 1); -val lemma1_1 = result(); - -Goal "y \\ (\\B \\ {Y*C. C \\ A}. B) ==> (\\B \\ A. y`(Y*B)): (\\B \\ A. Y*B)"; -by (fast_tac (claset() addSIs [lam_type, apply_type]) 1); -val lemma1_2 = result(); - -Goal "(\\B \\ {(nat->Union(A))*C. C \\ A}. B) \\ 0 ==> (\\B \\ A. B) \\ 0"; -by (fast_tac (claset() addSIs [equals0I,lemma1_1, lemma1_2]) 1); -val lemma1 = result(); - -Goal "0 \\ A ==> 0 \\ {(nat -> Union(A)) * C. C \\ A}"; -by (fast_tac (claset() addEs [Sigma_fun_space_not0 RS not_sym RS notE]) 1); -val lemma2 = result(); - -Goalw AC_defs "AC7 ==> AC6"; -by (rtac allI 1); -by (rtac impI 1); -by (case_tac "A=0" 1); -by (Asm_simp_tac 1); -by (rtac lemma1 1); -by (etac allE 1); -by (etac impE 1 THEN (assume_tac 2)); -by (blast_tac (claset() addSIs [lemma2] - addIs [eqpoll_sym, eqpoll_trans, Sigma_fun_space_eqpoll]) 1); -qed "AC7_AC6"; - - -(* ********************************************************************** *) -(* AC1 ==> AC8 *) -(* ********************************************************************** *) - -Goalw [eqpoll_def] - "\\B \\ A. \\B1 B2. B= & B1 eqpoll B2 \ -\ ==> 0 \\ { bij(fst(B),snd(B)). B \\ A }"; -by Auto_tac; -val lemma1 = result(); - -Goal "[| f \\ (\\X \\ RepFun(A,p). X); D \\ A |] ==> (\\x \\ A. f`p(x))`D \\ p(D)"; -by (resolve_tac [beta RS ssubst] 1 THEN (assume_tac 1)); -by (fast_tac (claset() addSEs [apply_type]) 1); -val lemma2 = result(); - -Goalw AC_defs "AC1 ==> AC8"; -by (Clarify_tac 1); -by (dtac lemma1 1); -by (fast_tac (claset() addSEs [lemma2]) 1); -qed "AC1_AC8"; - - -(* ********************************************************************** *) -(* AC8 ==> AC9 *) -(* - this proof replaces the following two from Rubin & Rubin: *) -(* AC8 ==> AC1 and AC1 ==> AC9 *) -(* ********************************************************************** *) - -Goal "\\B1 \\ A. \\B2 \\ A. B1 eqpoll B2 \ -\ ==> \\B \\ A*A. \\B1 B2. B= & B1 eqpoll B2"; -by (Fast_tac 1); -val lemma1 = result(); - -Goal "f \\ bij(fst(),snd()) ==> f \\ bij(a,b)"; -by (Asm_full_simp_tac 1); -val lemma2 = result(); - -Goalw AC_defs "AC8 ==> AC9"; -by (rtac allI 1); -by (rtac impI 1); -by (etac allE 1); -by (etac impE 1); -by (etac lemma1 1); -by (fast_tac (claset() addSEs [lemma2]) 1); -qed "AC8_AC9"; - - -(* ********************************************************************** *) -(* AC9 ==> AC1 *) -(* The idea of this proof comes from "Equivalents of the Axiom of Choice" *) -(* by Rubin & Rubin. But (x * y) is not necessarily equipollent to *) -(* (x * y) Un {0} when y is a set of total functions acting from nat to *) -(* Union(A) -- therefore we have used the set (y * nat) instead of y. *) -(* ********************************************************************** *) - -(* Rules nedded to prove lemma1 *) -val snd_lepoll_SigmaI = prod_lepoll_self RS - ((prod_commute_eqpoll RS eqpoll_imp_lepoll) RSN (2,lepoll_trans)); - - -Goal "[|0 \\ A; B \\ A|] ==> nat \\ ((nat \\ Union(A)) \\ B) \\ nat"; -by (blast_tac (claset() addDs [Sigma_fun_space_not0] - addIs [snd_lepoll_SigmaI]) 1); -qed "nat_lepoll_lemma"; - - -Goal "[| 0\\A; A\\0; \ -\ C = {((nat->Union(A))*B)*nat. B \\ A} Un \ -\ {cons(0,((nat->Union(A))*B)*nat). B \\ A}; \ -\ B1: C; B2: C |] \ -\ ==> B1 eqpoll B2"; -by (blast_tac - (claset() delrules [eqpoll_refl] - addSIs [nat_lepoll_lemma, nat_cons_eqpoll RS eqpoll_trans, - eqpoll_refl RSN (2, prod_eqpoll_cong)] - addIs [eqpoll_trans, eqpoll_sym, Sigma_fun_space_eqpoll]) 1); -val lemma1 = result(); - -Goal "\\B1 \\ {(F*B)*N. B \\ A} Un {cons(0,(F*B)*N). B \\ A}. \ -\ \\B2 \\ {(F*B)*N. B \\ A} Un {cons(0,(F*B)*N). B \\ A}. \ -\ f` \\ bij(B1, B2) \ -\ ==> (\\B \\ A. snd(fst((f`)`0))) \\ (\\X \\ A. X)"; -by (rtac lam_type 1); -by (rtac snd_type 1); -by (rtac fst_type 1); -by (resolve_tac [consI1 RSN (2, apply_type)] 1); -by (fast_tac (claset() addSIs [fun_weaken_type, bij_is_fun]) 1); -val lemma2 = result(); - -Goalw AC_defs "AC9 ==> AC1"; -by (rtac allI 1); -by (rtac impI 1); -by (etac allE 1); -by (case_tac "A=0" 1); -by (blast_tac (claset() addDs [lemma1,lemma2]) 2); -by Auto_tac; -qed "AC9_AC1"; diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC_Equiv.ML --- a/src/ZF/AC/AC_Equiv.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,126 +0,0 @@ -(* Title: ZF/AC/AC_Equiv.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -*) - -val WO_defs = [WO1_def, WO2_def, WO3_def, WO4_def, WO5_def, WO6_def, WO8_def]; - -val AC_defs = [AC0_def, AC1_def, AC2_def, AC3_def, AC4_def, AC5_def, - AC6_def, AC7_def, AC8_def, AC9_def, AC10_def, AC11_def, - AC12_def, AC13_def, AC14_def, AC15_def, AC16_def, - AC17_def, AC18_def, AC19_def]; - -val AC_aux_defs = [pairwise_disjoint_def, sets_of_size_between_def]; - - -(* ********************************************************************** *) -(* lemmas concerning FOL and pure ZF theory *) -(* ********************************************************************** *) - -(* used only in WO1_DC.ML *) -(*Note simpler proof*) -Goal "[| \\x \\ A. f`x=g`x; f \\ Df->Cf; g \\ Dg->Cg; A \\ Df; A \\ Dg |] ==> f``A=g``A"; -by (asm_simp_tac (simpset() addsimps [image_fun]) 1); -qed "images_eq"; - -(* used in \\ AC10-AC15.ML AC16WO4.ML WO6WO1.ML *) -(*I don't know where to put this one.*) -Goal - "!!m A B. [| A lepoll succ(m); B \\ A; B\\0 |] ==> A-B lepoll m"; -by (rtac not_emptyE 1 THEN (assume_tac 1)); -by (ftac singleton_subsetI 1); -by (ftac subsetD 1 THEN (assume_tac 1)); -by (res_inst_tac [("A2","A")] - (Diff_sing_lepoll RSN (2, subset_imp_lepoll RS lepoll_trans)) 1 - THEN (REPEAT (assume_tac 2))); -by (Fast_tac 1); -qed "Diff_lepoll"; - -(* ********************************************************************** *) -(* lemmas concerning lepoll and eqpoll relations *) -(* ********************************************************************** *) - -(* ********************************************************************** *) -(* Theorems concerning ordinals *) -(* ********************************************************************** *) - -(* lemma for ordertype_Int *) -goalw Cardinal.thy [rvimage_def] "rvimage(A,id(A),r) = r Int A*A"; -by (rtac equalityI 1); -by Safe_tac; -by (dres_inst_tac [("P","%a. :r")] (id_conv RS subst) 1 - THEN (assume_tac 1)); -by (dres_inst_tac [("P","%a. :r")] (id_conv RS subst) 1 - THEN (REPEAT (assume_tac 1))); -by (fast_tac (claset() addIs [id_conv RS ssubst]) 1); -qed "rvimage_id"; - -(* used only in Hartog.ML *) -goal Cardinal.thy - "!!A r. well_ord(A,r) ==> ordertype(A, r Int A*A) = ordertype(A,r)"; -by (res_inst_tac [("P","%a. ordertype(A,a)=ordertype(A,r)")] - (rvimage_id RS subst) 1); -by (eresolve_tac [id_bij RS bij_ordertype_vimage] 1); -qed "ordertype_Int"; - -(* used only in AC16_lemmas.ML *) -Goalw [InfCard_def] - "!!i. [| ~Finite(i); Card(i) |] ==> InfCard(i)"; -by (asm_simp_tac (simpset() addsimps [Card_is_Ord RS nat_le_infinite_Ord]) 1); -qed "Inf_Card_is_InfCard"; - -Goal "(THE z. {x}={z}) = x"; -by (fast_tac (claset() addSEs [singleton_eq_iff RS iffD1 RS sym]) 1); -qed "the_element"; - -Goal "(\\x \\ A. {x}) \\ bij(A, {{x}. x \\ A})"; -by (res_inst_tac [("d","%z. THE x. z={x}")] lam_bijective 1); -by (TRYALL (eresolve_tac [RepFunI, RepFunE])); -by (REPEAT (asm_full_simp_tac (simpset() addsimps [the_element]) 1)); -qed "lam_sing_bij"; - -val [major, minor] = Goalw [inj_def] - "[| f \\ inj(A, B); !!a. a \\ A ==> f`a \\ C |] ==> f \\ inj(A,C)"; -by (fast_tac (claset() addSEs [minor] - addSIs [major RS CollectD1 RS Pi_type, major RS CollectD2]) 1); -qed "inj_strengthen_type"; - -Goalw [Finite_def] "~Finite(nat)"; -by (fast_tac (claset() addSDs [eqpoll_imp_lepoll] - addIs [Ord_nat RSN (2, ltI) RS lt_not_lepoll RS notE]) 1); -qed "nat_not_Finite"; - -val le_imp_lepoll = le_imp_subset RS subset_imp_lepoll; - -(* ********************************************************************** *) -(* Another elimination rule for \\! *) -(* ********************************************************************** *) - -Goal "[| \\! x. P(x); P(x); P(y) |] ==> x=y"; -by (etac ex1E 1); -by (res_inst_tac [("b","xa")] (sym RSN (2, trans)) 1); -by (Fast_tac 1); -by (Fast_tac 1); -qed "ex1_two_eq"; - -(* ********************************************************************** *) -(* image of a surjection *) -(* ********************************************************************** *) - -Goalw [surj_def] "f \\ surj(A, B) ==> f``A = B"; -by (etac CollectE 1); -by (resolve_tac [subset_refl RSN (2, image_fun) RS ssubst] 1 - THEN (assume_tac 1)); -by (fast_tac (claset() addSEs [apply_type] addIs [equalityI]) 1); -qed "surj_image_eq"; - - -Goal "succ(x) lepoll y ==> y \\ 0"; -by (fast_tac (claset() addSDs [lepoll_0_is_0]) 1); -qed "succ_lepoll_imp_not_empty"; - -Goal "x eqpoll succ(n) ==> x \\ 0"; -by (fast_tac (claset() addSEs [eqpoll_sym RS eqpoll_0_is_0 RS succ_neq_0]) 1); -qed "eqpoll_succ_imp_not_empty"; - diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/AC_Equiv.thy --- a/src/ZF/AC/AC_Equiv.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/AC_Equiv.thy Wed Jan 16 17:52:06 2002 +0100 @@ -12,114 +12,231 @@ but slightly changed. *) +theory AC_Equiv = Main: (*obviously not Main_ZFC*) -AC_Equiv = Main + (*obviously not Main_ZFC*) - -consts +constdefs (* Well Ordering Theorems *) - WO1, WO2, WO3, WO5, WO6, WO7, WO8 :: o - WO4 :: i => o + WO1 :: o + "WO1 == \A. \R. well_ord(A,R)" + + WO2 :: o + "WO2 == \A. \a. Ord(a) & A\a" -(* Axioms of Choice *) - AC0, AC1, AC2, AC3, AC4, AC5, AC6, AC7, AC8, AC9, - AC11, AC12, AC14, AC15, AC17, AC19 :: o - AC10, AC13 :: i => o - AC16 :: [i, i] => o - AC18 :: prop ("AC18") + WO3 :: o + "WO3 == \A. \a. Ord(a) & (\b. b \ a & A\b)" -(* Auxiliary definitions used in definitions *) - pairwise_disjoint :: i => o - sets_of_size_between :: [i, i, i] => o + WO4 :: "i => o" + "WO4(m) == \A. \a f. Ord(a) & domain(f)=a & + (\bb m)" -defs - -(* Well Ordering Theorems *) + WO5 :: o + "WO5 == \m \ nat. 1\m & WO4(m)" - WO1_def "WO1 == \\A. \\R. well_ord(A,R)" - - WO2_def "WO2 == \\A. \\a. Ord(a) & A eqpoll a" + WO6 :: o + "WO6 == \A. \m \ nat. 1\m & (\a f. Ord(a) & domain(f)=a + & (\bb m))" - WO3_def "WO3 == \\A. \\a. Ord(a) & (\\b. b \\ a & A eqpoll b)" - - WO4_def "WO4(m) == \\A. \\a f. Ord(a) & domain(f)=a & - (\\bbA. Finite(A) <-> (\R. well_ord(A,R) --> well_ord(A,converse(R)))" - WO5_def "WO5 == \\m \\ nat. 1 le m & WO4(m)" + WO8 :: o + "WO8 == \A. (\f. f \ (\X \ A. X)) --> (\R. well_ord(A,R))" - WO6_def "WO6 == \\A. \\m \\ nat. 1 le m & (\\a f. Ord(a) & domain(f)=a - & (\\bbA. Finite(A) <-> (\\R. well_ord(A,R) --> - well_ord(A,converse(R)))" +(* Auxiliary concepts needed below *) + pairwise_disjoint :: "i => o" + "pairwise_disjoint(A) == \A1 \ A. \A2 \ A. A1 Int A2 \ 0 --> A1=A2" - WO8_def "WO8 == \\A. (\\f. f \\ (\\X \\ A. X)) --> (\\R. well_ord(A,R))" + sets_of_size_between :: "[i, i, i] => o" + "sets_of_size_between(A,m,n) == \B \ A. m \ B & B \ n" + (* Axioms of Choice *) + AC0 :: o + "AC0 == \A. \f. f \ (\X \ Pow(A)-{0}. X)" - AC0_def "AC0 == \\A. \\f. f \\ (\\X \\ Pow(A)-{0}. X)" + AC1 :: o + "AC1 == \A. 0\A --> (\f. f \ (\X \ A. X))" + + AC2 :: o + "AC2 == \A. 0\A & pairwise_disjoint(A) + --> (\C. \B \ A. \y. B Int C = {y})" + AC3 :: o + "AC3 == \A B. \f \ A->B. \g. g \ (\x \ {a \ A. f`a\0}. f`x)" - AC1_def "AC1 == \\A. 0\\A --> (\\f. f \\ (\\X \\ A. X))" + AC4 :: o + "AC4 == \R A B. (R \ A*B --> (\f. f \ (\x \ domain(R). R``{x})))" + + AC5 :: o + "AC5 == \A B. \f \ A->B. \g \ range(f)->A. \x \ domain(g). f`(g`x) = x" + + AC6 :: o + "AC6 == \A. 0\A --> (\B \ A. B)\0" - AC2_def "AC2 == \\A. 0\\A & pairwise_disjoint(A) - --> (\\C. \\B \\ A. \\y. B Int C = {y})" + AC7 :: o + "AC7 == \A. 0\A & (\B1 \ A. \B2 \ A. B1\B2) --> (\B \ A. B) \ 0" - AC3_def "AC3 == \\A B. \\f \\ A->B. \\g. g \\ (\\x \\ {a \\ A. f`a\\0}. f`x)" + AC8 :: o + "AC8 == \A. (\B \ A. \B1 B2. B= & B1\B2) + --> (\f. \B \ A. f`B \ bij(fst(B),snd(B)))" + + AC9 :: o + "AC9 == \A. (\B1 \ A. \B2 \ A. B1\B2) --> + (\f. \B1 \ A. \B2 \ A. f` \ bij(B1,B2))" - AC4_def "AC4 == \\R A B. (R \\ A*B --> (\\f. f \\ (\\x \\ domain(R). R``{x})))" + AC10 :: "i => o" + "AC10(n) == \A. (\B \ A. ~Finite(B)) --> + (\f. \B \ A. (pairwise_disjoint(f`B) & + sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B))" - AC5_def "AC5 == \\A B. \\f \\ A->B. \\g \\ range(f)->A. - \\x \\ domain(g). f`(g`x) = x" + AC11 :: o + "AC11 == \n \ nat. 1\n & AC10(n)" + + AC12 :: o + "AC12 == \A. (\B \ A. ~Finite(B)) --> + (\n \ nat. 1\n & (\f. \B \ A. (pairwise_disjoint(f`B) & + sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B)))" - AC6_def "AC6 == \\A. 0\\A --> (\\B \\ A. B)\\0" + AC13 :: "i => o" + "AC13(m) == \A. 0\A --> (\f. \B \ A. f`B\0 & f`B \ B & f`B \ m)" + + AC14 :: o + "AC14 == \m \ nat. 1\m & AC13(m)" + + AC15 :: o + "AC15 == \A. 0\A --> + (\m \ nat. 1\m & (\f. \B \ A. f`B\0 & f`B \ B & f`B \ m))" - AC7_def "AC7 == \\A. 0\\A & (\\B1 \\ A. \\B2 \\ A. B1 eqpoll B2) - --> (\\B \\ A. B)\\0" + AC16 :: "[i, i] => o" + "AC16(n, k) == + \A. ~Finite(A) --> + (\T. T \ {X \ Pow(A). X\succ(n)} & + (\X \ {X \ Pow(A). X\succ(k)}. \! Y. Y \ T & X \ Y))" - AC8_def "AC8 == \\A. (\\B \\ A. \\B1 B2. B= & B1 eqpoll B2) - --> (\\f. \\B \\ A. f`B \\ bij(fst(B),snd(B)))" + AC17 :: o + "AC17 == \A. \g \ (Pow(A)-{0} -> A) -> Pow(A)-{0}. + \f \ Pow(A)-{0} -> A. f`(g`f) \ g`f" - AC9_def "AC9 == \\A. (\\B1 \\ A. \\B2 \\ A. B1 eqpoll B2) --> - (\\f. \\B1 \\ A. \\B2 \\ A. f` \\ bij(B1,B2))" + AC18 :: "prop" ("AC18") + "AC18 == (!!X A B. A\0 & (\a \ A. B(a) \ 0) --> + ((\a \ A. \b \ B(a). X(a,b)) = + (\f \ \a \ A. B(a). \a \ A. X(a, f`a))))" + --"AC18 can be expressed only using meta-level quantification" + + AC19 :: o + "AC19 == \A. A\0 & 0\A --> ((\a \ A. \b \ a. b) = + (\f \ (\B \ A. B). \a \ A. f`a))" + + - AC10_def "AC10(n) == \\A. (\\B \\ A. ~Finite(B)) --> - (\\f. \\B \\ A. (pairwise_disjoint(f`B) & - sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B))" +(* ********************************************************************** *) +(* Theorems concerning ordinals *) +(* ********************************************************************** *) - AC11_def "AC11 == \\n \\ nat. 1 le n & AC10(n)" +(* lemma for ordertype_Int *) +lemma rvimage_id: "rvimage(A,id(A),r) = r Int A*A" +apply (unfold rvimage_def) +apply (rule equalityI, safe) +apply (drule_tac P = "%a. :r" in id_conv [THEN subst], + (assumption)) +apply (drule_tac P = "%a. :r" in id_conv [THEN subst], (assumption+)) +apply (fast intro: id_conv [THEN ssubst]) +done - AC12_def "AC12 == \\A. (\\B \\ A. ~Finite(B)) --> - (\\n \\ nat. 1 le n & (\\f. \\B \\ A. (pairwise_disjoint(f`B) & - sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B)))" +(* used only in Hartog.ML *) +lemma ordertype_Int: + "well_ord(A,r) ==> ordertype(A, r Int A*A) = ordertype(A,r)" +apply (rule_tac P = "%a. ordertype (A,a) =ordertype (A,r) " in rvimage_id [THEN subst]) +apply (erule id_bij [THEN bij_ordertype_vimage]) +done + +lemma the_element: "(THE z. {x}={z}) = x" +by (fast elim!: singleton_eq_iff [THEN iffD1, symmetric]) - AC13_def "AC13(m) == \\A. 0\\A --> (\\f. \\B \\ A. f`B\\0 & - f`B \\ B & f`B lepoll m)" +lemma lam_sing_bij: "(\x \ A. {x}) \ bij(A, {{x}. x \ A})" +apply (rule_tac d = "%z. THE x. z={x}" in lam_bijective) +apply (auto simp add: the_element) +done + +lemma inj_strengthen_type: + "[| f \ inj(A, B); !!a. a \ A ==> f`a \ C |] ==> f \ inj(A,C)" +by (unfold inj_def, blast intro: Pi_type) + +lemma nat_not_Finite: "~ Finite(nat)" +by (unfold Finite_def, blast dest: eqpoll_imp_lepoll ltI lt_not_lepoll) - AC14_def "AC14 == \\m \\ nat. 1 le m & AC13(m)" +lemmas le_imp_lepoll = le_imp_subset [THEN subset_imp_lepoll] + +(* ********************************************************************** *) +(* Another elimination rule for \! *) +(* ********************************************************************** *) + +lemma ex1_two_eq: "[| \! x. P(x); P(x); P(y) |] ==> x=y" +by blast - AC15_def "AC15 == \\A. 0\\A --> (\\m \\ nat. 1 le m & (\\f. \\B \\ A. - f`B\\0 & f`B \\ B & f`B lepoll m))" +(* ********************************************************************** *) +(* image of a surjection *) +(* ********************************************************************** *) - AC16_def "AC16(n, k) == \\A. ~Finite(A) --> - (\\T. T \\ {X \\ Pow(A). X eqpoll succ(n)} & - (\\X \\ {X \\ Pow(A). X eqpoll succ(k)}. \\! Y. Y \\ T & X \\ Y))" +lemma surj_image_eq: "f \ surj(A, B) ==> f``A = B" +apply (unfold surj_def) +apply (erule CollectE) +apply (rule image_fun [THEN ssubst], (assumption), rule subset_refl) +apply (blast dest: apply_type) +done + + +(* ********************************************************************** *) +(* Lemmas used in the proofs like WO? ==> AC? *) +(* ********************************************************************** *) - AC17_def "AC17 == \\A. \\g \\ (Pow(A)-{0} -> A) -> Pow(A)-{0}. - \\f \\ Pow(A)-{0} -> A. f`(g`f) \\ g`f" +lemma first_in_B: + "[| well_ord(Union(A),r); 0\A; B \ A |] ==> (THE b. first(b,B,r)) \ B" +by (blast dest!: well_ord_imp_ex1_first + [THEN theI, THEN first_def [THEN def_imp_iff, THEN iffD1]]) + +lemma ex_choice_fun: "[| well_ord(Union(A), R); 0\A |] ==> \f. f:(\X \ A. X)" +by (fast elim!: first_in_B intro!: lam_type) - AC18_def "AC18 == (!!X A B. A\\0 & (\\a \\ A. B(a) \\ 0) --> - ((\\a \\ A. \\b \\ B(a). X(a,b)) = - (\\f \\ \\a \\ A. B(a). \\a \\ A. X(a, f`a))))" +lemma ex_choice_fun_Pow: "well_ord(A, R) ==> \f. f:(\X \ Pow(A)-{0}. X)" +by (fast elim!: well_ord_subset [THEN ex_choice_fun]) + - AC19_def "AC19 == \\A. A\\0 & 0\\A --> ((\\a \\ A. \\b \\ a. b) = - (\\f \\ (\\B \\ A. B). \\a \\ A. f`a))" +(* ********************************************************************** *) +(* Lemmas needed to state when a finite relation is a function. *) +(* The criteria are cardinalities of the relation and its domain. *) +(* Used in WO6WO1.ML *) +(* ********************************************************************** *) -(* Auxiliary definitions used in the above definitions *) - - pairwise_disjoint_def "pairwise_disjoint(A) - == \\A1 \\ A. \\A2 \\ A. A1 Int A2 \\ 0 --> A1=A2" +(*Using AC we could trivially prove, for all u, domain(u) \ u*) +lemma lepoll_m_imp_domain_lepoll_m: + "[| m \ nat; u \ m |] ==> domain(u) \ m" +apply (unfold lepoll_def) +apply (erule exE) +apply (rule_tac x = "\x \ domain(u). LEAST i. \y. \ u & f` = i" + in exI) +apply (rule_tac d = "%y. fst (converse(f) ` y) " in lam_injective) +apply (fast intro: LeastI2 nat_into_Ord [THEN Ord_in_Ord] + inj_is_fun [THEN apply_type]) +apply (erule domainE) +apply (frule inj_is_fun [THEN apply_type], (assumption)) +apply (rule LeastI2) +apply (auto elim!: nat_into_Ord [THEN Ord_in_Ord]) +done - sets_of_size_between_def "sets_of_size_between(A,m,n) - == \\B \\ A. m lepoll B & B lepoll n" - +lemma rel_domain_ex1: + "[| succ(m) \ domain(r); r \ succ(m); m \ nat |] ==> function(r)" +apply (unfold function_def, safe) +apply (rule ccontr) +apply (fast elim!: lepoll_trans [THEN succ_lepoll_natE] + lepoll_m_imp_domain_lepoll_m [OF _ Diff_sing_lepoll] + elim: domain_Diff_eq [OF _ not_sym, THEN subst]) +done + +lemma rel_is_fun: + "[| succ(m) \ domain(r); r \ succ(m); m \ nat; + r \ A*B; A=domain(r) |] ==> r \ A->B" +by (simp add: Pi_iff rel_domain_ex1) + end diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/Cardinal_aux.ML --- a/src/ZF/AC/Cardinal_aux.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,209 +0,0 @@ -(* Title: ZF/AC/Cardinal_aux.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -Auxiliary lemmas concerning cardinalities -*) - -(* ********************************************************************** *) -(* Lemmas involving ordinals and cardinalities used in the proofs *) -(* concerning AC16 and DC *) -(* ********************************************************************** *) - -(* j=|A| *) -Goal "[| A lepoll i; Ord(i) |] ==> \\j. j le i & A eqpoll j"; -by (blast_tac (claset() addSIs [lepoll_cardinal_le, well_ord_Memrel, - well_ord_cardinal_eqpoll RS eqpoll_sym] - addDs [lepoll_well_ord]) 1); -qed "lepoll_imp_ex_le_eqpoll"; - -(* j=|A| *) -Goalw [lesspoll_def] - "[| A lesspoll i; Ord(i) |] ==> \\j. j InfCard(|i|)"; -by (rtac conjI 1); -by (rtac Card_cardinal 1); -by (resolve_tac [Card_nat RS (Card_def RS def_imp_iff RS iffD1 RS ssubst)] 1); -by (resolve_tac [nat_le_infinite_Ord RS le_imp_lepoll - RSN (2, well_ord_Memrel RS well_ord_lepoll_imp_Card_le)] 1 - THEN REPEAT (assume_tac 1)); -qed "Inf_Ord_imp_InfCard_cardinal"; - -Goal "[| A eqpoll i; B eqpoll i; ~Finite(i); Ord(i) |] \ -\ ==> A Un B eqpoll i"; -by (rtac eqpollI 1); -by (eresolve_tac [subset_imp_lepoll RSN (2, eqpoll_sym RS eqpoll_imp_lepoll - RS lepoll_trans)] 2); -by (Fast_tac 2); -by (resolve_tac [Un_lepoll_sum RS lepoll_trans] 1); -by (resolve_tac [lepoll_imp_sum_lepoll_prod RS lepoll_trans] 1); -by (eresolve_tac [eqpoll_sym RSN (2, eqpoll_trans) RS eqpoll_imp_lepoll] 1 - THEN (assume_tac 1)); -by (resolve_tac [nat_le_infinite_Ord RS le_imp_lepoll RS - (Ord_nat RS (nat_2I RS OrdmemD) RS subset_imp_lepoll RS lepoll_trans) - RS (eqpoll_sym RS eqpoll_imp_lepoll RSN (2, lepoll_trans))] 1 - THEN (REPEAT (assume_tac 1))); -by (eresolve_tac [prod_eqpoll_cong RS eqpoll_imp_lepoll RS lepoll_trans] 1 - THEN (assume_tac 1)); -by (resolve_tac [Inf_Ord_imp_InfCard_cardinal RSN (2, well_ord_Memrel RS - well_ord_InfCard_square_eq) RS eqpoll_imp_lepoll] 1 - THEN REPEAT (assume_tac 1)); -qed "Un_eqpoll_Inf_Ord"; - - -Goal "?f \\ bij({{y,z}. y \\ x}, x)"; -by (rtac RepFun_bijective 1); -by (simp_tac (simpset() addsimps [doubleton_eq_iff]) 1); -by (Blast_tac 1); -qed "paired_bij"; - -Goalw [eqpoll_def] "{{y,z}. y \\ x} eqpoll x"; -by (fast_tac (claset() addSIs [paired_bij]) 1); -qed "paired_eqpoll"; - -Goal "\\B. B eqpoll A & B Int C = 0"; -by (fast_tac (claset() addSIs [paired_eqpoll, equals0I] addEs [mem_asym]) 1); -qed "ex_eqpoll_disjoint"; - -Goal "[| A lepoll i; B lepoll i; ~Finite(i); Ord(i) |] \ -\ ==> A Un B lepoll i"; -by (res_inst_tac [("A1","i"), ("C1","i")] (ex_eqpoll_disjoint RS exE) 1); -by (etac conjE 1); -by (dresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RSN (2, lepoll_trans)] 1); -by (assume_tac 1); -by (resolve_tac [Un_lepoll_Un RS lepoll_trans] 1 THEN (REPEAT (assume_tac 1))); -by (eresolve_tac [eqpoll_refl RSN (2, Un_eqpoll_Inf_Ord) RS - eqpoll_imp_lepoll] 1 - THEN (REPEAT (assume_tac 1))); -qed "Un_lepoll_Inf_Ord"; - -Goal "[| P(i); i \\ j; Ord(j) |] ==> (LEAST i. P(i)) \\ j"; -by (eresolve_tac [Least_le RS leE] 1); -by (etac Ord_in_Ord 1 THEN (assume_tac 1)); -by (etac ltE 1); -by (fast_tac (claset() addDs [OrdmemD]) 1); -by (etac subst_elem 1 THEN (assume_tac 1)); -qed "Least_in_Ord"; - -Goal "[| well_ord(x,r); y \\ x; y lepoll succ(n); n \\ nat |] \ -\ ==> y-{THE b. first(b,y,r)} lepoll n"; -by (res_inst_tac [("Q","y=0")] (excluded_middle RS disjE) 1); -by (fast_tac (claset() addSIs [Diff_sing_lepoll, the_first_in]) 1); -by (res_inst_tac [("b","y-{THE b. first(b, y, r)}")] subst 1); -by (rtac empty_lepollI 2); -by (Fast_tac 1); -qed "Diff_first_lepoll"; - -Goal "(\\x \\ X. P(x)) \\ (\\x \\ X. P(x)-Q(x)) Un (\\x \\ X. Q(x))"; -by (Fast_tac 1); -qed "UN_subset_split"; - -Goalw [lepoll_def] "Ord(a) ==> (\\x \\ a. {P(x)}) lepoll a"; -by (res_inst_tac [("x","\\z \\ (\\x \\ a. {P(x)}). (LEAST i. P(i)=z)")] exI 1); -by (res_inst_tac [("d","%z. P(z)")] lam_injective 1); -by (fast_tac (claset() addSIs [Least_in_Ord]) 1); -by (fast_tac (claset() addIs [LeastI] addSEs [Ord_in_Ord]) 1); -qed "UN_sing_lepoll"; - -Goal "[| well_ord(T, R); ~Finite(a); Ord(a); n \\ nat |] ==> \ -\ \\f. (\\b \\ a. f`b lepoll n & f`b \\ T) --> (\\b \\ a. f`b) lepoll a"; -by (induct_tac "n" 1); -by (rtac allI 1); -by (rtac impI 1); -by (res_inst_tac [("b","\\b \\ a. f`b")] subst 1); -by (rtac empty_lepollI 2); -by (resolve_tac [equals0I RS sym] 1); -by (REPEAT (eresolve_tac [UN_E, allE] 1)); -by (fast_tac (claset() addDs [lepoll_0_is_0 RS subst]) 1); -by (rtac allI 1); -by (rtac impI 1); -by (eres_inst_tac [("x","\\x \\ a. f`x - {THE b. first(b,f`x,R)}")] allE 1); -by (etac impE 1); -by (Asm_full_simp_tac 1); -by (fast_tac (claset() addSIs [Diff_first_lepoll]) 1); -by (Asm_full_simp_tac 1); -by (resolve_tac [UN_subset_split RS subset_imp_lepoll RS lepoll_trans] 1); -by (rtac Un_lepoll_Inf_Ord 1 THEN (REPEAT_FIRST assume_tac)); -by (etac UN_sing_lepoll 1); -qed "UN_fun_lepoll_lemma"; - -Goal "[| \\b \\ a. f`b lepoll n & f`b \\ T; well_ord(T, R); \ -\ ~Finite(a); Ord(a); n \\ nat |] ==> (\\b \\ a. f`b) lepoll a"; -by (eresolve_tac [UN_fun_lepoll_lemma RS allE] 1 THEN (REPEAT (assume_tac 1))); -by (Fast_tac 1); -qed "UN_fun_lepoll"; - -Goal "[| \\b \\ a. F(b) lepoll n & F(b) \\ T; well_ord(T, R); \ -\ ~Finite(a); Ord(a); n \\ nat |] ==> (\\b \\ a. F(b)) lepoll a"; -by (rtac impE 1 THEN (assume_tac 3)); -by (res_inst_tac [("f","\\b \\ a. F(b)")] (UN_fun_lepoll) 2 - THEN (TRYALL assume_tac)); -by Auto_tac; -qed "UN_lepoll"; - -Goal "Ord(a) ==> (\\b \\ a. F(b)) = (\\b \\ a. F(b) - (\\c \\ b. F(c)))"; -by (rtac equalityI 1); -by (Fast_tac 2); -by (rtac subsetI 1); -by (etac UN_E 1); -by (rtac UN_I 1); -by (res_inst_tac [("P","%z. x \\ F(z)")] Least_in_Ord 1 THEN (REPEAT (assume_tac 1))); -by (rtac DiffI 1); -by (resolve_tac [Ord_in_Ord RSN (2, LeastI)] 1 THEN (REPEAT (assume_tac 1))); -by (rtac notI 1); -by (etac UN_E 1); -by (eres_inst_tac [("P","%z. x \\ F(z)"),("i","c")] less_LeastE 1); -by (eresolve_tac [Ord_Least RSN (2, ltI)] 1); -qed "UN_eq_UN_Diffs"; - -Goalw [lepoll_def, eqpoll_def] - "a lepoll X ==> \\Y. Y \\ X & a eqpoll Y"; -by (etac exE 1); -by (forward_tac [subset_refl RSN (2, restrict_bij)] 1); -by (res_inst_tac [("x","f``a")] exI 1); -by (fast_tac (claset() addSEs [inj_is_fun RS fun_is_rel RS image_subset]) 1); -qed "lepoll_imp_eqpoll_subset"; - -(* ********************************************************************** *) -(* Diff_lesspoll_eqpoll_Card *) -(* ********************************************************************** *) - -Goal "[| A\\a; ~Finite(a); Card(a); B lesspoll a; A-B lesspoll a |] ==> P"; -by (REPEAT (eresolve_tac [lesspoll_imp_ex_lt_eqpoll RS exE, - Card_is_Ord, conjE] 1)); -by (forw_inst_tac [("j","xa")] ([lt_Ord, lt_Ord] MRS Un_upper1_le) 1 - THEN (assume_tac 1)); -by (forw_inst_tac [("j","xa")] ([lt_Ord, lt_Ord] MRS Un_upper2_le) 1 - THEN (assume_tac 1)); -by (dtac Un_least_lt 1 THEN (assume_tac 1)); -by (dresolve_tac [le_imp_lepoll RSN - (2, eqpoll_imp_lepoll RS lepoll_trans)] 1 - THEN (assume_tac 1)); -by (dresolve_tac [le_imp_lepoll RSN - (2, eqpoll_imp_lepoll RS lepoll_trans)] 1 - THEN (assume_tac 1)); -by (res_inst_tac [("Q","Finite(x Un xa)")] (excluded_middle RS disjE) 1); -by (dresolve_tac [[lepoll_Finite, lepoll_Finite] MRS Finite_Un] 2 - THEN (REPEAT (assume_tac 2))); -by (dresolve_tac [subset_Un_Diff RS subset_imp_lepoll RS lepoll_Finite] 2); -by (fast_tac (claset() - addDs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_Finite]) 2); -by (dresolve_tac [ Un_lepoll_Inf_Ord] 1 THEN (REPEAT (assume_tac 1))); -by (fast_tac (claset() addSEs [ltE, Ord_in_Ord]) 1); -by (dresolve_tac [subset_Un_Diff RS subset_imp_lepoll RS lepoll_trans RS - (lt_Card_imp_lesspoll RSN (2, lesspoll_trans1))] 1 - THEN (TRYALL assume_tac)); -by (fast_tac (claset() addSDs [lesspoll_def RS def_imp_iff RS iffD1]) 1); -qed "Diff_lesspoll_eqpoll_Card_lemma"; - -Goal "[| A eqpoll a; ~Finite(a); Card(a); B lesspoll a |] \ -\ ==> A - B eqpoll a"; -by (rtac swap 1 THEN (Fast_tac 1)); -by (rtac Diff_lesspoll_eqpoll_Card_lemma 1 THEN (REPEAT (assume_tac 1))); -by (fast_tac (claset() addSIs [lesspoll_def RS def_imp_iff RS iffD2, - subset_imp_lepoll RS (eqpoll_imp_lepoll RSN (2, lepoll_trans))]) 1); -qed "Diff_lesspoll_eqpoll_Card"; - diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/Cardinal_aux.thy --- a/src/ZF/AC/Cardinal_aux.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/Cardinal_aux.thy Wed Jan 16 17:52:06 2002 +0100 @@ -1,3 +1,212 @@ -(*Dummy theory to document dependencies *) +(* Title: ZF/AC/Cardinal_aux.thy + ID: $Id$ + Author: Krzysztof Grabczewski + +Auxiliary lemmas concerning cardinalities +*) + +theory Cardinal_aux = AC_Equiv: + +lemma Diff_lepoll: "[| A \ succ(m); B \ A; B\0 |] ==> A-B \ m" +apply (rule not_emptyE, (assumption)) +apply (blast intro: lepoll_trans [OF subset_imp_lepoll Diff_sing_lepoll]) +done + + +(* ********************************************************************** *) +(* Lemmas involving ordinals and cardinalities used in the proofs *) +(* concerning AC16 and DC *) +(* ********************************************************************** *) + + +(* j=|A| *) +lemma lepoll_imp_ex_le_eqpoll: + "[| A \ i; Ord(i) |] ==> \j. j le i & A \ j" +by (blast intro!: lepoll_cardinal_le well_ord_Memrel + well_ord_cardinal_eqpoll [THEN eqpoll_sym] + dest: lepoll_well_ord); + +(* j=|A| *) +lemma lesspoll_imp_ex_lt_eqpoll: + "[| A \ i; Ord(i) |] ==> \j. j j" +by (unfold lesspoll_def, blast dest!: lepoll_imp_ex_le_eqpoll elim!: leE) + +lemma Inf_Ord_imp_InfCard_cardinal: "[| ~Finite(i); Ord(i) |] ==> InfCard(|i|)" +apply (unfold InfCard_def) +apply (rule conjI) +apply (rule Card_cardinal) +apply (rule Card_nat + [THEN Card_def [THEN def_imp_iff, THEN iffD1, THEN ssubst]]) + -- "rewriting would loop!" +apply (rule well_ord_Memrel [THEN well_ord_lepoll_imp_Card_le], assumption) +apply (rule nat_le_infinite_Ord [THEN le_imp_lepoll], assumption+) +done + +text{*An alternative and more general proof goes like this: A and B are both +well-ordered (because they are injected into an ordinal), either A lepoll B +or B lepoll A. Also both are equipollent to their cardinalities, so +(if A and B are infinite) then A Un B lepoll |A|+|B| = max(|A|,|B|) lepoll i. +In fact, the correctly strengthened version of this theorem appears below.*} +lemma Un_lepoll_Inf_Ord_weak: + "[|A \ i; B \ i; \ Finite(i); Ord(i)|] ==> A \ B \ i" +apply (rule Un_lepoll_sum [THEN lepoll_trans]) +apply (rule lepoll_imp_sum_lepoll_prod [THEN lepoll_trans]) +apply (erule eqpoll_trans [THEN eqpoll_imp_lepoll]) +apply (erule eqpoll_sym) +apply (rule subset_imp_lepoll [THEN lepoll_trans, THEN lepoll_trans]) +apply (rule nat_2I [THEN OrdmemD], rule Ord_nat) +apply (rule nat_le_infinite_Ord [THEN le_imp_lepoll], assumption+) +apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll]) +apply (erule prod_eqpoll_cong [THEN eqpoll_imp_lepoll, THEN lepoll_trans], + assumption) +apply (rule eqpoll_imp_lepoll) +apply (rule well_ord_Memrel [THEN well_ord_InfCard_square_eq], assumption) +apply (rule Inf_Ord_imp_InfCard_cardinal, assumption+) +done + +lemma Un_eqpoll_Inf_Ord: + "[| A \ i; B \ i; ~Finite(i); Ord(i) |] ==> A Un B \ i" +apply (rule eqpollI) +apply (blast intro: Un_lepoll_Inf_Ord_weak) +apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_trans]) +apply (rule Un_upper1 [THEN subset_imp_lepoll]) +done + +lemma paired_bij: "?f \ bij({{y,z}. y \ x}, x)" +apply (rule RepFun_bijective) +apply (simp add: doubleton_eq_iff, blast) +done + +lemma paired_eqpoll: "{{y,z}. y \ x} \ x" +by (unfold eqpoll_def, fast intro!: paired_bij) + +lemma ex_eqpoll_disjoint: "\B. B \ A & B Int C = 0" +by (fast intro!: paired_eqpoll equals0I elim: mem_asym) + +(*Finally we reach this result. Surely there's a simpler proof, as sketched + above?*) +lemma Un_lepoll_Inf_Ord: + "[| A \ i; B \ i; ~Finite(i); Ord(i) |] ==> A Un B \ i" +apply (rule_tac A1 = "i" and C1 = "i" in ex_eqpoll_disjoint [THEN exE]) +apply (erule conjE) +apply (drule lepoll_trans) +apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll]) +apply (rule Un_lepoll_Un [THEN lepoll_trans], (assumption+)) +apply (blast intro: eqpoll_refl Un_eqpoll_Inf_Ord eqpoll_imp_lepoll) +done + +lemma Least_in_Ord: "[| P(i); i \ j; Ord(j) |] ==> (LEAST i. P(i)) \ j" +apply (erule Least_le [THEN leE]) +apply (erule Ord_in_Ord, assumption) +apply (erule ltE) +apply (fast dest: OrdmemD) +apply (erule subst_elem, assumption) +done -Cardinal_aux = AC_Equiv +lemma Diff_first_lepoll: + "[| well_ord(x,r); y \ x; y \ succ(n); n \ nat |] + ==> y - {THE b. first(b,y,r)} \ n" +apply (case_tac "y=0", simp add: empty_lepollI) +apply (fast intro!: Diff_sing_lepoll the_first_in) +done + +lemma UN_subset_split: + "(\x \ X. P(x)) \ (\x \ X. P(x)-Q(x)) Un (\x \ X. Q(x))" +by blast + +lemma UN_sing_lepoll: "Ord(a) ==> (\x \ a. {P(x)}) \ a" +apply (unfold lepoll_def) +apply (rule_tac x = "\z \ (\x \ a. {P (x) }) . (LEAST i. P (i) =z) " in exI) +apply (rule_tac d = "%z. P (z) " in lam_injective) +apply (fast intro!: Least_in_Ord) +apply (fast intro: LeastI elim!: Ord_in_Ord) +done + +lemma UN_fun_lepoll_lemma [rule_format]: + "[| well_ord(T, R); ~Finite(a); Ord(a); n \ nat |] + ==> \f. (\b \ a. f`b \ n & f`b \ T) --> (\b \ a. f`b) \ a" +apply (induct_tac "n") +apply (rule allI) +apply (rule impI) +apply (rule_tac b = "\b \ a. f`b" in subst) +apply (rule_tac [2] empty_lepollI) +apply (rule equals0I [symmetric], clarify) +apply (fast dest: lepoll_0_is_0 [THEN subst]) +apply (rule allI) +apply (rule impI) +apply (erule_tac x = "\x \ a. f`x - {THE b. first (b,f`x,R) }" in allE) +apply (erule impE, simp) +apply (fast intro!: Diff_first_lepoll, simp) +apply (rule UN_subset_split [THEN subset_imp_lepoll, THEN lepoll_trans]) +apply (fast intro: Un_lepoll_Inf_Ord UN_sing_lepoll) +done + +lemma UN_fun_lepoll: + "[| \b \ a. f`b \ n & f`b \ T; well_ord(T, R); + ~Finite(a); Ord(a); n \ nat |] ==> (\b \ a. f`b) \ a" +by (blast intro: UN_fun_lepoll_lemma); + +lemma UN_lepoll: + "[| \b \ a. F(b) \ n & F(b) \ T; well_ord(T, R); + ~Finite(a); Ord(a); n \ nat |] + ==> (\b \ a. F(b)) \ a" +apply (rule rev_mp) +apply (rule_tac f="\b \ a. F (b)" in UN_fun_lepoll); +apply auto +done + +lemma UN_eq_UN_Diffs: + "Ord(a) ==> (\b \ a. F(b)) = (\b \ a. F(b) - (\c \ b. F(c)))" +apply (rule equalityI) + prefer 2 apply fast +apply (rule subsetI) +apply (erule UN_E) +apply (rule UN_I) + apply (rule_tac P = "%z. x \ F (z) " in Least_in_Ord, (assumption+)) +apply (rule DiffI, best intro: Ord_in_Ord LeastI, clarify) +apply (erule_tac P = "%z. x \ F (z) " and i = "c" in less_LeastE) +apply (blast intro: Ord_Least ltI) +done + +lemma lepoll_imp_eqpoll_subset: + "a \ X ==> \Y. Y \ X & a \ Y" +apply (unfold lepoll_def eqpoll_def, clarify) +apply (blast intro: restrict_bij + dest: inj_is_fun [THEN fun_is_rel, THEN image_subset]) +done + +(* ********************************************************************** *) +(* Diff_lesspoll_eqpoll_Card *) +(* ********************************************************************** *) + +lemma Diff_lesspoll_eqpoll_Card_lemma: + "[| A\a; ~Finite(a); Card(a); B \ a; A-B \ a |] ==> P" +apply (elim lesspoll_imp_ex_lt_eqpoll [THEN exE] Card_is_Ord conjE) +apply (frule_tac j=xa in Un_upper1_le [OF lt_Ord lt_Ord], assumption) +apply (frule_tac j=xa in Un_upper2_le [OF lt_Ord lt_Ord], assumption) +apply (drule Un_least_lt, assumption) +apply (drule eqpoll_imp_lepoll [THEN lepoll_trans], + rule le_imp_lepoll, assumption)+ +apply (case_tac "Finite(x Un xa)"); +txt{*finite case*} + apply (drule Finite_Un [OF lepoll_Finite lepoll_Finite], assumption+) + apply (drule subset_Un_Diff [THEN subset_imp_lepoll, THEN lepoll_Finite]) + apply (fast dest: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_Finite]) +txt{*infinite case*} +apply (drule Un_lepoll_Inf_Ord, (assumption+)) +apply (blast intro: le_Ord2) +apply (drule lesspoll_trans1 + [OF subset_Un_Diff [THEN subset_imp_lepoll, THEN lepoll_trans] + lt_Card_imp_lesspoll], assumption+) +apply (simp add: lesspoll_def) +done + +lemma Diff_lesspoll_eqpoll_Card: + "[| A \ a; ~Finite(a); Card(a); B \ a |] ==> A - B \ a" +apply (rule ccontr) +apply (rule Diff_lesspoll_eqpoll_Card_lemma, (assumption+)) +apply (blast intro: lesspoll_def [THEN def_imp_iff, THEN iffD2] + subset_imp_lepoll eqpoll_imp_lepoll lepoll_trans) +done + +end diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/DC.ML --- a/src/ZF/AC/DC.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,533 +0,0 @@ -(* Title: ZF/AC/DC.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -The proofs concerning the Axiom of Dependent Choice -*) - -(* ********************************************************************** *) -(* DC ==> DC(omega) *) -(* *) -(* The scheme of the proof: *) -(* *) -(* Assume DC. Let R and X satisfy the premise of DC(omega). *) -(* *) -(* Define XX and RR as follows: *) -(* *) -(* XX = (\\n \\ nat. {f \\ n->X. \\k \\ n. \\ R}) *) -(* f RR g iff domain(g)=succ(domain(f)) & *) -(* restrict(g, domain(f)) = f *) -(* *) -(* Then RR satisfies the hypotheses of DC. *) -(* So applying DC: *) -(* *) -(* \\f \\ nat->XX. \\n \\ nat. f`n RR f`succ(n) *) -(* *) -(* Thence *) -(* *) -(* ff = {. n \\ nat} *) -(* *) -(* is the desired function. *) -(* *) -(* ********************************************************************** *) - -Open_locale "DC0_imp"; - -val all_ex = thm "all_ex"; -val XX_def = thm "XX_def"; -val RR_def = thm "RR_def"; - -Goalw [RR_def] "RR \\ XX*XX"; -by (Fast_tac 1); -qed "lemma1_1"; - -Goalw [RR_def, XX_def] "RR \\ 0"; -by (rtac (all_ex RS ballE) 1); -by (eresolve_tac [empty_subsetI RS PowI RSN (2, notE)] 2); -by (eresolve_tac [nat_0I RS n_lesspoll_nat RSN (2, impE)] 1); -by (etac bexE 1); -by (res_inst_tac [("a","<0, {<0, x>}>")] not_emptyI 1); -by (rtac CollectI 1); -by (rtac SigmaI 1); -by (rtac (nat_0I RS UN_I) 1); -by (asm_simp_tac (simpset() addsimps [nat_0I RS UN_I]) 1); -by (rtac (nat_1I RS UN_I) 1); -by (asm_simp_tac (simpset() addsimps [singleton_0]) 2); -by (force_tac (claset() addSIs [singleton_fun RS Pi_type], - simpset() addsimps [singleton_0 RS sym]) 1); -qed "lemma1_2"; - -Goalw [RR_def, XX_def] "range(RR) \\ domain(RR)"; -by (rtac range_subset_domain 1); -by (Blast_tac 2); -by (Clarify_tac 1); -by (forward_tac [fun_is_rel RS image_subset RS PowI RS (all_ex RS bspec)] 1); -by (eresolve_tac [[nat_into_Ord RSN (2, image_Ord_lepoll), n_lesspoll_nat] - MRS lesspoll_trans1 RSN (2, impE)] 1 - THEN REPEAT (assume_tac 1)); -by (etac bexE 1); -by (res_inst_tac [("x","cons(, g)")] exI 1); -by (rtac CollectI 1); -by (force_tac (claset() addSEs [cons_fun_type2], - simpset() addsimps [cons_image_n, cons_val_n, - cons_image_k, cons_val_k]) 1); -by (asm_full_simp_tac (simpset() - addsimps [domain_of_fun, succ_def, restrict_cons_eq]) 1); -qed "lemma1_3"; - - -Goal "[| \\n \\ nat. \\ RR; f \\ nat -> XX; n \\ nat |] \ -\ ==> \\k \\ nat. f`succ(n) \\ k -> X & n \\ k \ -\ & \\ R"; -by (induct_tac "n" 1); -by (dresolve_tac [nat_1I RSN (2, apply_type)] 1); -by (dresolve_tac [nat_0I RSN (2, bspec)] 1); -by (asm_full_simp_tac (simpset() addsimps [XX_def]) 1); -by Safe_tac; -by (rtac bexI 1 THEN (assume_tac 2)); -by (best_tac (claset() addIs [ltD] - addSEs [nat_0_le RS leE] - addEs [sym RS trans RS succ_neq_0, domain_of_fun] - addss (simpset() addsimps [RR_def])) 1); -(** LEVEL 7, other subgoal **) -by (dresolve_tac [nat_succI RSN (2, bspec)] 1 THEN (assume_tac 1)); -by (subgoal_tac "f ` succ(succ(x)) \\ succ(k)->X" 1); -by (dresolve_tac [nat_succI RS nat_succI RSN (2, apply_type)] 1 - THEN (assume_tac 1)); -by (full_simp_tac (simpset() addsimps [XX_def,RR_def]) 1); -by Safe_tac; -by (forw_inst_tac [("a","succ(k)")] (domain_of_fun RS sym RS trans) 1 THEN - (assume_tac 1)); -by (forw_inst_tac [("a","xa")] (domain_of_fun RS sym RS trans) 1 THEN - (assume_tac 1)); -by (fast_tac (claset() addSEs [nat_into_Ord RS succ_in_succ] - addSDs [nat_into_Ord RS succ_in_succ RSN (2, bspec)]) 1); -by (dtac domain_of_fun 1); -by (full_simp_tac (simpset() addsimps [XX_def,RR_def]) 1); -by (deepen_tac (claset() addDs [domain_of_fun RS sym RS trans]) 0 1); -qed "lemma2"; - -Goal "[| \\n \\ nat. \\ RR; f \\ nat -> XX; m \\ nat |] \ -\ ==> {f`succ(x)`x. x \\ m} = {f`succ(m)`x. x \\ m}"; -by (subgoal_tac "\\x \\ m. f`succ(m)`x = f`succ(x)`x" 1); -by (Asm_full_simp_tac 1); -by (induct_tac "m" 1); -by (Fast_tac 1); -by (rtac ballI 1); -by (etac succE 1); -by (rtac restrict_eq_imp_val_eq 1); -by (dresolve_tac [nat_succI RSN (2, bspec)] 1 THEN (assume_tac 1)); -by (asm_full_simp_tac (simpset() addsimps [RR_def]) 1); -by (dtac lemma2 1 THEN REPEAT (assume_tac 1)); -by (fast_tac (claset() addSDs [domain_of_fun]) 1); -by (dres_inst_tac [("x","xa")] bspec 1 THEN (assume_tac 1)); -by (eresolve_tac [sym RS trans RS sym] 1); -by (resolve_tac [restrict_eq_imp_val_eq RS sym] 1); -by (dresolve_tac [nat_succI RSN (2, bspec)] 1 THEN (assume_tac 1)); -by (asm_full_simp_tac (simpset() addsimps [RR_def]) 1); -by (dtac lemma2 1 THEN REPEAT (assume_tac 1)); -by (blast_tac (claset() addSDs [domain_of_fun] - addIs [nat_into_Ord RSN (2, OrdmemD) RS subsetD]) 1); -qed "lemma3_1"; - -Goal "[| \\n \\ nat. \\ RR; f \\ nat -> XX; m \\ nat |] \ -\ ==> (\\x \\ nat. f`succ(x)`x) `` m = f`succ(m)``m"; -by (etac natE 1); -by (Asm_simp_tac 1); -by (stac image_lam 1); -by (fast_tac (claset() addSEs [[nat_succI, Ord_nat] MRS OrdmemD]) 1); -by (stac lemma3_1 1 THEN REPEAT (assume_tac 1)); -by (Fast_tac 1); -by (fast_tac (claset() addSDs [lemma2] - addSEs [nat_into_Ord RSN (2, OrdmemD) RSN - (2, image_fun RS sym)]) 1); -qed "lemma3"; - -Close_locale "DC0_imp"; - -Goalw [DC_def, DC0_def] "DC0 ==> DC(nat)"; -by (Clarify_tac 1); -by (REPEAT (etac allE 1)); -by (etac impE 1); - (*these three results comprise Lemma 1*) -by (blast_tac (claset() addSIs (map export [lemma1_1, lemma1_2, lemma1_3])) 1); -by (etac bexE 1); -by (res_inst_tac [("x","\\n \\ nat. f`succ(n)`n")] bexI 1); -by (fast_tac (claset() addSIs [lam_type] addSDs [export lemma2] - addSEs [fun_weaken_type, apply_type]) 2); -by (rtac oallI 1); -by (forward_tac [ltD RSN (3, export lemma2)] 1 - THEN assume_tac 2); -by (fast_tac (claset() addSEs [fun_weaken_type]) 1); -(** LEVEL 10: last subgoal **) -by (stac (ltD RSN (3, export lemma3)) 1); -by (Force_tac 4); -by (assume_tac 3); -by (assume_tac 1); -by (fast_tac (claset() addSEs [fun_weaken_type]) 1); -qed "DC0_imp_DC_nat"; - - -(* ************************************************************************ - DC(omega) ==> DC - - The scheme of the proof: - - Assume DC(omega). Let R and x satisfy the premise of DC. - - Define XX and RR as follows: - - XX = (\\n \\ nat. {f \\ succ(n)->domain(R). \\k \\ n. \\ R}) - - RR = {:Fin(XX)*XX. - (domain(z2)=succ(\\f \\ z1. domain(f)) & - (\\f \\ z1. restrict(z2, domain(f)) = f)) | - (~ (\\g \\ XX. domain(g)=succ(\\f \\ z1. domain(f)) & - (\\f \\ z1. restrict(g, domain(f)) = f)) & - z2={<0,x>})} - - Then XX and RR satisfy the hypotheses of DC(omega). - So applying DC: - - \\f \\ nat->XX. \\n \\ nat. f``n RR f`n - - Thence - - ff = {. n \\ nat} - - is the desired function. - -************************************************************************* *) - -Goal "n \\ nat ==> \\A. (A eqpoll n & A \\ X) --> A \\ Fin(X)"; -by (induct_tac "n" 1); -by (rtac allI 1); -by (fast_tac (claset() addSIs [Fin.emptyI] - addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]) 1); -by (rtac allI 1); -by (rtac impI 1); -by (etac conjE 1); -by (resolve_tac [eqpoll_succ_imp_not_empty RS not_emptyE] 1 - THEN (assume_tac 1)); -by (ftac Diff_sing_eqpoll 1 THEN (assume_tac 1)); -by (etac allE 1); -by (etac impE 1); -by (Fast_tac 1); -by (dtac subsetD 1 THEN (assume_tac 1)); -by (dresolve_tac [Fin.consI] 1 THEN (assume_tac 1)); -by (asm_full_simp_tac (simpset() addsimps [cons_Diff]) 1); -qed "Finite_Fin_lemma"; - -Goalw [Finite_def] "[| Finite(A); A \\ X |] ==> A \\ Fin(X)"; -by (etac bexE 1); -by (dtac Finite_Fin_lemma 1); -by (etac allE 1); -by (etac impE 1); -by (assume_tac 2); -by (Fast_tac 1); -qed "Finite_Fin"; - -Goal - "x \\ X ==> {<0,x>}: (\\n \\ nat. {f \\ succ(n)->X. \\k \\ n. \\ R})"; -by (rtac (nat_0I RS UN_I) 1); -by (fast_tac (claset() addSIs [singleton_fun RS Pi_type] - addss (simpset() addsimps [singleton_0 RS sym])) 1); -qed "singleton_in_funs"; - - -Open_locale "imp_DC0"; - -val XX_def = thm "XX_def"; -val RR_def = thm "RR_def"; -val allRR_def = thm "allRR_def"; - -Goal "[| range(R) \\ domain(R); x \\ domain(R) |] \ -\ ==> RR \\ Pow(XX)*XX & \ -\ (\\Y \\ Pow(XX). Y lesspoll nat --> (\\x \\ XX. :RR))"; -by (rtac conjI 1); -by (force_tac (claset() addSDs [FinD RS PowI], - simpset() addsimps [RR_def]) 1); -by (rtac (impI RS ballI) 1); -by (dresolve_tac [[lesspoll_nat_is_Finite, PowD] MRS Finite_Fin] 1 - THEN (assume_tac 1)); -by (excluded_middle_tac "\\g \\ XX. domain(g)=succ(\\f \\ Y. domain(f)) \ -\ & (\\f \\ Y. restrict(g, domain(f)) = f)" 1); -by (fast_tac (claset() addss (simpset() addsimps [RR_def])) 2); -by (safe_tac (claset() delrules [domainE])); -by (rewrite_goals_tac [XX_def,RR_def]); -by (swap_res_tac [bexI] 1 THEN etac singleton_in_funs 2); -by (asm_full_simp_tac (simpset() addsimps [nat_0I RSN (2, bexI), - cons_fun_type2]) 1); -qed "lemma4"; - -Goal "[| f \\ nat->X; n \\ nat |] ==> \ -\ (\\x \\ f``succ(n). P(x)) = P(f`n) Un (\\x \\ f``n. P(x))"; -by (asm_full_simp_tac (simpset() - addsimps [Ord_nat RSN (2, OrdmemD) RSN (2, image_fun), - [nat_succI, Ord_nat] MRS OrdmemD RSN (2, image_fun)]) 1); -qed "UN_image_succ_eq"; - -Goal "[| (\\x \\ f``n. P(x)) = y; P(f`n) = succ(y); \ -\ f \\ nat -> X; n \\ nat |] ==> (\\x \\ f``succ(n). P(x)) = succ(y)"; -by (asm_full_simp_tac (simpset() addsimps [UN_image_succ_eq]) 1); -by (Fast_tac 1); -qed "UN_image_succ_eq_succ"; - -Goal "[| f \\ succ(n) -> D; n \\ nat; \ -\ domain(f)=succ(x); x=y |] ==> f`y \\ D"; -by (fast_tac (claset() addEs [apply_type] - addSDs [[sym, domain_of_fun] MRS trans]) 1); -qed "apply_domain_type"; - -Goal "[| f \\ nat -> X; n \\ nat |] ==> f``succ(n) = cons(f`n, f``n)"; -by (asm_full_simp_tac (simpset() - addsimps [nat_succI, Ord_nat RSN (2, OrdmemD), image_fun]) 1); -qed "image_fun_succ"; - -Goalw [XX_def] "[| domain(f`n) = succ(u); f \\ nat -> XX; u=k; n \\ nat |] \ -\ ==> f`n \\ succ(k) -> domain(R)"; -by (dtac apply_type 1 THEN (assume_tac 1)); -by (fast_tac (claset() addEs [domain_eq_imp_fun_type]) 1); -qed "f_n_type"; - -Goalw [XX_def] - "[| f \\ nat -> XX; domain(f`n) = succ(k); n \\ nat |] \ -\ ==> \\i \\ k. \\ R"; -by (dtac apply_type 1 THEN (assume_tac 1)); -by (etac UN_E 1); -by (etac CollectE 1); -by (dresolve_tac [domain_of_fun RS sym RS trans] 1 THEN (assume_tac 1)); -by (Asm_full_simp_tac 1); -qed "f_n_pairs_in_R"; - -Goalw [restrict_def] - "[| restrict(f, domain(x))=x; f \\ n->X; domain(x) \\ n |] \ -\ ==> restrict(cons(, f), domain(x)) = x"; -by (eresolve_tac [sym RS trans RS sym] 1); -by (rtac fun_extension 1); -by (fast_tac (claset() addSIs [lam_type]) 1); -by (fast_tac (claset() addSIs [lam_type]) 1); -by (asm_full_simp_tac (simpset() addsimps [subsetD RS cons_val_k]) 1); -qed "restrict_cons_eq_restrict"; - -Goal "[| \\x \\ f``n. restrict(f`n, domain(x))=x; \ -\ f \\ nat -> XX; \ -\ n \\ nat; domain(f`n) = succ(n); \ -\ (\\x \\ f``n. domain(x)) \\ n |] \ -\ ==> \\x \\ f``succ(n). restrict(cons(, f`n), domain(x)) = x"; -by (rtac ballI 1); -by (asm_full_simp_tac (simpset() addsimps [image_fun_succ]) 1); -by (dtac f_n_type 1 THEN REPEAT (ares_tac [refl] 1)); -by (etac disjE 1); -by (asm_full_simp_tac (simpset() addsimps [domain_of_fun,restrict_cons_eq]) 1); -by (dtac bspec 1 THEN (assume_tac 1)); -by (fast_tac (claset() addSEs [restrict_cons_eq_restrict]) 1); -qed "all_in_image_restrict_eq"; - -Goalw [RR_def, allRR_def] - "[| \\b \\ RR; \ -\ f \\ nat -> XX; range(R) \\ domain(R); x \\ domain(R)|] \ -\ ==> allRR"; -by (rtac oallI 1); -by (dtac ltD 1); -by (etac nat_induct 1); -by (dresolve_tac [[nat_0I, Ord_nat] MRS ltI RSN (2, ospec)] 1); -by (fast_tac (FOL_cs addss - (simpset() addsimps [singleton_fun RS domain_of_fun, - singleton_0, singleton_in_funs])) 1); -(*induction step*) (** LEVEL 5 **) -by (full_simp_tac (*prevent simplification of ~\\to \\~*) - (FOL_ss addsimps [separation, split]) 1); -by (dresolve_tac [[nat_succI, Ord_nat] MRS ltI RSN (2, ospec)] 1 - THEN (assume_tac 1)); -by (REPEAT (eresolve_tac [conjE, disjE] 1)); -by (force_tac (FOL_cs addSEs [trans, subst_context] - addSIs [UN_image_succ_eq_succ], simpset()) 1); -by (etac conjE 1); -by (etac notE 1); -by (asm_lr_simp_tac (simpset() addsimps [XX_def, UN_image_succ_eq_succ]) 1); -(** LEVEL 12 **) -by (REPEAT (eresolve_tac [conjE, bexE] 1)); -by (dtac apply_domain_type 1 THEN REPEAT (assume_tac 1)); -by (etac domainE 1); -by (etac domainE 1); -by (forward_tac [export f_n_type] 1 THEN REPEAT (assume_tac 1)); -by (rtac bexI 1); -by (etac nat_succI 2); -by (res_inst_tac [("x","cons(, f`xa)")] bexI 1); -by (rtac conjI 1); -by (fast_tac (FOL_cs - addEs [subst_context RSN (2, trans) RS domain_cons_eq_succ, - subst_context, export all_in_image_restrict_eq, - trans, equalityD1]) 2); -by (eresolve_tac [rangeI RSN (2, subsetD) RSN (2, cons_fun_type2)] 2 - THEN REPEAT (assume_tac 2)); -by (rtac ballI 1); -by (etac succE 1); -(** LEVEL 25 **) -by (EVERY [dtac (domain_of_fun RSN (2, export f_n_pairs_in_R)) 2, - REPEAT (assume_tac 2), dtac bspec 2, assume_tac 2]); -by (asm_full_simp_tac (simpset() - addsimps [nat_into_Ord RS succ_in_succ, succI2, cons_val_k]) 2); -by (asm_full_simp_tac (simpset() addsimps [cons_val_n, cons_val_k]) 1); -qed "simplify_recursion"; - - -Goalw [allRR_def] - "[| allRR; f \\ nat -> XX; range(R) \\ domain(R); x \\ domain(R); n \\ nat |] \ -\ ==> f`n \\ succ(n) -> domain(R) & (\\i \\ n. :R)"; -by (dtac ospec 1); -by (eresolve_tac [Ord_nat RSN (2, ltI)] 1); -by (etac CollectE 1); -by (Asm_full_simp_tac 1); -by (rtac conjI 1); -by (fast_tac (FOL_cs addSEs [conjE, f_n_pairs_in_R, trans, subst_context]) 2); -by (rewtac XX_def); -by (fast_tac (claset() - addSEs [trans RS domain_eq_imp_fun_type, subst_context]) 1); -qed "lemma2"; - -Goal "[| allRR; f \\ nat -> XX; n \\ nat; range(R) \\ domain(R); x \\ domain(R) \ -\ |] ==> f`n`n = f`succ(n)`n"; -by (forward_tac [lemma2 RS conjunct1 RS domain_of_fun] 1 - THEN REPEAT (assume_tac 1)); -by (rewtac allRR_def); -by (dresolve_tac [[nat_succI, Ord_nat] MRS ltI RSN (2, ospec)] 1 - THEN (assume_tac 1)); -by (Asm_full_simp_tac 1); -by (REPEAT (etac conjE 1)); -by (etac ballE 1); -by (eresolve_tac [restrict_eq_imp_val_eq RS sym] 1); -by (fast_tac (claset() addSEs [ssubst]) 1); -by (asm_full_simp_tac (simpset() - addsimps [[nat_succI, Ord_nat] MRS OrdmemD RSN (2, image_fun)]) 1); -qed "lemma3"; - -Close_locale "imp_DC0"; - - -Goalw [DC_def, DC0_def] "DC(nat) ==> DC0"; -by (REPEAT (resolve_tac [allI, impI] 1)); -by (REPEAT (eresolve_tac [asm_rl, conjE, ex_in_domain RS exE, allE] 1)); -by (eresolve_tac [export lemma4 RSN (2, impE)] 1 - THEN REPEAT (assume_tac 1)); -by (etac bexE 1); -by (dresolve_tac [export simplify_recursion] 1 - THEN REPEAT (assume_tac 1)); -by (res_inst_tac [("x","\\n \\ nat. f`n`n")] bexI 1); -by (rtac lam_type 2); -by (eresolve_tac [[export lemma2 RS conjunct1, succI1] MRS apply_type] 2 - THEN REPEAT (assume_tac 2)); -by (rtac ballI 1); -by (forward_tac [export (nat_succI RSN (5,lemma2)) RS conjunct2] 1 - THEN REPEAT (assume_tac 1)); -by (dresolve_tac [export lemma3] 1 THEN REPEAT (assume_tac 1)); -by (asm_full_simp_tac (simpset() addsimps [nat_succI]) 1); -qed "DC_nat_imp_DC0"; - -(* ********************************************************************** *) -(* \\K. Card(K) --> DC(K) ==> WO3 *) -(* ********************************************************************** *) - -Goalw [lesspoll_def] - "[| ~ A lesspoll B; C lesspoll B |] ==> A - C \\ 0"; -by (fast_tac (claset() addSDs [Diff_eq_0_iff RS iffD1 RS subset_imp_lepoll] - addSIs [eqpollI] addEs [notE] addSEs [eqpollE, lepoll_trans]) 1); -qed "lesspoll_lemma"; - -val [f_type, Ord_a, not_eq] = goalw thy [inj_def] - "[| f \\ a->X; Ord(a); (!!b c. [| b a |] ==> f`b\\f`c) |] \ -\ ==> f \\ inj(a, X)"; -by (resolve_tac [f_type RS CollectI] 1); -by (REPEAT (resolve_tac [ballI,impI] 1)); -by (resolve_tac [Ord_a RS Ord_in_Ord RS Ord_linear_lt] 1 - THEN (assume_tac 1)); -by (eres_inst_tac [("j","x")] (Ord_a RS Ord_in_Ord) 1); -by (REPEAT (fast_tac (claset() addDs [not_eq, not_eq RS not_sym]) 1)); -qed "fun_Ord_inj"; - -Goal "[| f \\ X->Y; A \\ X; a \\ A |] ==> f`a \\ f``A"; -by (fast_tac (claset() addSEs [image_fun RS ssubst]) 1); -qed "value_in_image"; - -Goalw [DC_def, WO3_def] "\\K. Card(K) --> DC(K) ==> WO3"; -by (rtac allI 1); -by (excluded_middle_tac "A lesspoll Hartog(A)" 1); -by (fast_tac (claset() addSDs [lesspoll_imp_ex_lt_eqpoll] - addSIs [Ord_Hartog, leI RS le_imp_subset]) 2); -by (REPEAT (eresolve_tac [allE, impE] 1)); -by (rtac Card_Hartog 1); -by (eres_inst_tac [("x","A")] allE 1); -by (eres_inst_tac [("x","{:Pow(A)*A . z1 \ -\ lesspoll Hartog(A) & z2 \\ z1}")] allE 1); -by (Asm_full_simp_tac 1); -by (etac impE 1); -by (fast_tac (claset() addEs [lesspoll_lemma RS not_emptyE]) 1); -by (etac bexE 1); -by (resolve_tac [exI RS (lepoll_def RS (def_imp_iff RS iffD2)) - RS (HartogI RS notE)] 1); -by (resolve_tac [Ord_Hartog RSN (2, fun_Ord_inj)] 1 THEN (assume_tac 1)); -by (dresolve_tac [Ord_Hartog RSN (2, OrdmemD) RSN (2, - ltD RSN (3, value_in_image))] 1 - THEN REPEAT (assume_tac 1)); -by (force_tac (claset() addSDs [Ord_Hartog RSN (2, ltI) RSN (2, ospec)], - simpset()) 1); -qed "DC_WO3"; - -(* ********************************************************************** *) -(* WO1 ==> \\K. Card(K) --> DC(K) *) -(* ********************************************************************** *) - -Goal "[| Ord(a); b \\ a |] ==> (\\x \\ a. P(x))``b = (\\x \\ b. P(x))``b"; -by (rtac images_eq 1); -by (REPEAT (fast_tac (claset() addSEs [RepFunI, OrdmemD] - addSIs [lam_type]) 2)); -by (rtac ballI 1); -by (dresolve_tac [OrdmemD RS subsetD] 1 - THEN REPEAT (assume_tac 1)); -by (Asm_full_simp_tac 1); -qed "lam_images_eq"; - -Goalw [lesspoll_def] "[| Card(K); b \\ K |] ==> b lesspoll K"; -by (asm_full_simp_tac (simpset() addsimps [Card_iff_initial]) 1); -by (fast_tac (claset() addSIs [le_imp_lepoll, ltI, leI]) 1); -qed "in_Card_imp_lesspoll"; - -Goal "(\\b \\ a. P(b)) \\ a -> {P(b). b \\ a}"; -by (fast_tac (claset() addSIs [lam_type, RepFunI]) 1); -qed "lam_type_RepFun"; - -Goal "[| \\Y \\ Pow(X). Y lesspoll K --> (\\x \\ X. \\ R); \ -\ b \\ K; Z \\ Pow(X); Z lesspoll K |] \ -\ ==> {x \\ X. \\ R} \\ 0"; -by (Blast_tac 1); -qed "lemmaX"; - -Goal "[| Card(K); well_ord(X,Q); \ -\ \\Y \\ Pow(X). Y lesspoll K --> (\\x \\ X. \\ R); b \\ K |] \ -\ ==> ff(b, X, Q, R) \\ {x \\ X. <(\\c \\ b. ff(c, X, Q, R))``b, x> \\ R}"; -by (res_inst_tac [("P","b \\ K")] impE 1 THEN TRYALL assume_tac); -by (res_inst_tac [("i","b")] (Card_is_Ord RS Ord_in_Ord RS trans_induct) 1 - THEN REPEAT (assume_tac 1)); -by (rtac impI 1); -by (resolve_tac [ff_def RS def_transrec RS ssubst] 1); -by (etac the_first_in 1); -by (Fast_tac 1); -by (asm_full_simp_tac (simpset() - addsimps [[lam_type_RepFun, subset_refl] MRS image_fun]) 1); -by (etac lemmaX 1 THEN assume_tac 1); -by (blast_tac (claset() addIs [Card_is_Ord RSN (2, OrdmemD) RS subsetD]) 1); -by (blast_tac (claset() addIs [lesspoll_trans1, in_Card_imp_lesspoll, - RepFun_lepoll]) 1); -qed "lemma"; - -Goalw [DC_def, WO1_def] "WO1 ==> \\K. Card(K) --> DC(K)"; -by (REPEAT (resolve_tac [allI,impI] 1)); -by (REPEAT (eresolve_tac [allE,exE,conjE] 1)); -by (res_inst_tac [("x","\\b \\ K. ff(b, X, Ra, R)")] bexI 1); -by (rtac lam_type 2); -by (resolve_tac [lemma RS CollectD1] 2 THEN REPEAT (assume_tac 2)); -by (asm_full_simp_tac (simpset() - addsimps [[Card_is_Ord, ltD] MRS lam_images_eq]) 1); -by (fast_tac (claset() addSEs [ltE, lemma RS CollectD2]) 1); -qed "WO1_DC_Card"; diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/DC.thy --- a/src/ZF/AC/DC.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/DC.thy Wed Jan 16 17:52:06 2002 +0100 @@ -2,69 +2,578 @@ ID: $Id$ Author: Krzysztof Grabczewski -Theory file for the proofs concernind the Axiom of Dependent Choice +The proofs concerning the Axiom of Dependent Choice *) -DC = AC_Equiv + Hartog + Cardinal_aux + DC_lemmas + +theory DC = AC_Equiv + Hartog + Cardinal_aux: + +lemma RepFun_lepoll: "Ord(a) ==> {P(b). b \ a} \ a" +apply (unfold lepoll_def) +apply (rule_tac x = "\z \ RepFun (a,P) . LEAST i. z=P (i) " in exI) +apply (rule_tac d="%z. P (z)" in lam_injective) + apply (fast intro!: Least_in_Ord) +apply (rule sym) +apply (fast intro: LeastI Ord_in_Ord) +done -consts +text{*Trivial in the presence of AC, but here we need a wellordering of X*} +lemma image_Ord_lepoll: "[| f \ X->Y; Ord(X) |] ==> f``X \ X" +apply (unfold lepoll_def) +apply (rule_tac x = "\x \ f``X. LEAST y. f`y = x" in exI) +apply (rule_tac d = "%z. f`z" in lam_injective) +apply (fast intro!: Least_in_Ord apply_equality, clarify) +apply (rule LeastI) + apply (erule apply_equality, assumption+) +apply (blast intro: Ord_in_Ord) +done - DC :: i => o - DC0 :: o - ff :: [i, i, i, i] => i +lemma range_subset_domain: + "[| R \ X*X; !!g. g \ X ==> \u. \ R |] + ==> range(R) \ domain(R)" +by (blast ); + +lemma cons_fun_type: "g \ n->X ==> cons(, g) \ succ(n) -> cons(x, X)" +apply (unfold succ_def) +apply (fast intro!: fun_extend elim!: mem_irrefl) +done -rules +lemma cons_fun_type2: + "[| g \ n->X; x \ X |] ==> cons(, g) \ succ(n) -> X" +by (erule cons_absorb [THEN subst], erule cons_fun_type) + +lemma cons_image_n: "n \ nat ==> cons(, g)``n = g``n" +by (fast elim!: mem_irrefl) + +lemma cons_val_n: "g \ n->X ==> cons(, g)`n = x" +by (fast intro!: apply_equality elim!: cons_fun_type) + +lemma cons_image_k: "k \ n ==> cons(, g)``k = g``k" +by (fast elim: mem_asym) - DC_def "DC(a) == - \\X R. R \\ Pow(X)*X & - (\\Y \\ Pow(X). Y lesspoll a --> (\\x \\ X. \\ R)) - --> (\\f \\ a->X. \\b \\ R)" +lemma cons_val_k: "[| k \ n; g \ n->X |] ==> cons(, g)`k = g`k" +by (fast intro!: apply_equality consI2 elim!: cons_fun_type apply_Pair) + +lemma domain_cons_eq_succ: "domain(f)=x ==> domain(cons(, f)) = succ(x)" +by (simp add: domain_cons succ_def) + +lemma restrict_cons_eq: "g \ n->X ==> restrict(cons(, g), n) = g" +apply (unfold restrict_def) +apply (rule fun_extension) +apply (rule lam_type) +apply (erule cons_fun_type [THEN apply_type]) +apply (erule succI2, assumption) +apply (simp add: cons_val_k) +done + +lemma succ_in_succ: "[| Ord(k); i \ k |] ==> succ(i) \ succ(k)" +apply (rule Ord_linear [of "succ(i)" "succ(k)", THEN disjE]) +apply (fast elim: Ord_in_Ord mem_irrefl mem_asym)+ +done - DC0_def "DC0 == \\A B R. R \\ A*B & R\\0 & range(R) \\ domain(R) - --> (\\f \\ nat->domain(R). \\n \\ nat. :R)" +lemma restrict_eq_imp_val_eq: + "[| restrict(f, domain(g)) = g; x \ domain(g) |] ==> f`x = g`x" +apply (unfold restrict_def) +apply (erule subst, simp) +done + +lemma domain_eq_imp_fun_type: "[| domain(f)=A; f \ B->C |] ==> f \ A->C" +by (frule domain_of_fun, fast) + +lemma ex_in_domain: "[| R \ A * B; R \ 0 |] ==> \x. x \ domain(R)" +by (fast elim!: not_emptyE) + - ff_def "ff(b, X, Q, R) == - transrec(b, %c r. THE x. first(x, {x \\ X. \\ R}, Q))" - +constdefs + + DC :: "i => o" + "DC(a) == \X R. R \ Pow(X)*X & + (\Y \ Pow(X). Y \ a --> (\x \ X. \ R)) + --> (\f \ a->X. \b \ R)" + + DC0 :: o + "DC0 == \A B R. R \ A*B & R\0 & range(R) \ domain(R) + --> (\f \ nat->domain(R). \n \ nat. :R)" + + ff :: "[i, i, i, i] => i" + "ff(b, X, Q, R) == + transrec(b, %c r. THE x. first(x, {x \ X. \ R}, Q))" + locale DC0_imp = - fixes - XX :: i - RR :: i - X :: i - R :: i + fixes XX and RR and X and R + + assumes all_ex: "\Y \ Pow(X). Y \ nat --> (\x \ X. \ R)" + + defines XX_def: "XX == (\n \ nat. {f \ n->X. \k \ n. \ R})" + and RR_def: "RR == {:XX*XX. domain(z2)=succ(domain(z1)) + & restrict(z2, domain(z1)) = z1}" + + +(* ********************************************************************** *) +(* DC ==> DC(omega) *) +(* *) +(* The scheme of the proof: *) +(* *) +(* Assume DC. Let R and X satisfy the premise of DC(omega). *) +(* *) +(* Define XX and RR as follows: *) +(* *) +(* XX = (\n \ nat. {f \ n->X. \k \ n. \ R}) *) +(* f RR g iff domain(g)=succ(domain(f)) & *) +(* restrict(g, domain(f)) = f *) +(* *) +(* Then RR satisfies the hypotheses of DC. *) +(* So applying DC: *) +(* *) +(* \f \ nat->XX. \n \ nat. f`n RR f`succ(n) *) +(* *) +(* Thence *) +(* *) +(* ff = {. n \ nat} *) +(* *) +(* is the desired function. *) +(* *) +(* ********************************************************************** *) + +lemma (in DC0_imp) lemma1_1: "RR \ XX*XX" +by (unfold RR_def, fast) + +lemma (in DC0_imp) lemma1_2: "RR \ 0" +apply (unfold RR_def XX_def) +apply (rule all_ex [THEN ballE]) +apply (erule_tac [2] notE [OF _ empty_subsetI [THEN PowI]]) +apply (erule_tac impE [OF _ nat_0I [THEN n_lesspoll_nat]]) +apply (erule bexE) +apply (rule_tac a = "<0, {<0, x>}>" in not_emptyI) +apply (rule CollectI) +apply (rule SigmaI) +apply (rule nat_0I [THEN UN_I]) +apply (simp (no_asm_simp) add: nat_0I [THEN UN_I]) +apply (rule nat_1I [THEN UN_I]) +apply (force intro!: singleton_fun [THEN Pi_type] + simp add: singleton_0 [symmetric]) +apply (simp add: singleton_0) +done + +lemma (in DC0_imp) lemma1_3: "range(RR) \ domain(RR)" +apply (unfold RR_def XX_def) +apply (rule range_subset_domain, blast, clarify) +apply (frule fun_is_rel [THEN image_subset, THEN PowI, + THEN all_ex [THEN bspec]]) +apply (erule impE[OF _ lesspoll_trans1[OF image_Ord_lepoll + [OF _ nat_into_Ord] n_lesspoll_nat]], + assumption+) +apply (erule bexE) +apply (rule_tac x = "cons (, g) " in exI) +apply (rule CollectI) +apply (force elim!: cons_fun_type2 + simp add: cons_image_n cons_val_n cons_image_k cons_val_k) +apply (simp add: domain_of_fun succ_def restrict_cons_eq) +done - assumes - all_ex "\\Y \\ Pow(X). Y lesspoll nat --> (\\x \\ X. \\ R)" +lemma (in DC0_imp) lemma2: + "[| \n \ nat. \ RR; f \ nat -> XX; n \ nat |] + ==> \k \ nat. f`succ(n) \ k -> X & n \ k + & \ R" +apply (induct_tac "n") +apply (drule apply_type [OF _ nat_1I]) +apply (drule bspec [OF _ nat_0I]) +apply (simp add: XX_def, safe) +apply (rule rev_bexI, assumption) +apply (subgoal_tac "0 \ x", force) +apply (force simp add: RR_def + intro: ltD elim!: nat_0_le [THEN leE]) +(** LEVEL 7, other subgoal **) +apply (drule bspec [OF _ nat_succI], assumption) +apply (subgoal_tac "f ` succ (succ (x)) \ succ (k) ->X") +apply (drule apply_type [OF _ nat_succI [THEN nat_succI]], assumption) +apply (simp (no_asm_use) add: XX_def RR_def) +apply safe +apply (frule_tac a="succ(k)" in domain_of_fun [symmetric, THEN trans], + assumption) +apply (frule_tac a="xa" in domain_of_fun [symmetric, THEN trans], + assumption) +apply (fast elim!: nat_into_Ord [THEN succ_in_succ] + dest!: bspec [OF _ nat_into_Ord [THEN succ_in_succ]]) +apply (drule domain_of_fun) +apply (simp add: XX_def RR_def, clarify) +apply (blast dest: domain_of_fun [symmetric, THEN trans] ) +done + +lemma (in DC0_imp) lemma3_1: + "[| \n \ nat. \ RR; f \ nat -> XX; m \ nat |] + ==> {f`succ(x)`x. x \ m} = {f`succ(m)`x. x \ m}" +apply (subgoal_tac "\x \ m. f`succ (m) `x = f`succ (x) `x") +apply simp +apply (induct_tac "m", blast) +apply (rule ballI) +apply (erule succE) + apply (rule restrict_eq_imp_val_eq) + apply (drule bspec [OF _ nat_succI], assumption) + apply (simp add: RR_def) + apply (drule lemma2, assumption+) + apply (fast dest!: domain_of_fun) +apply (drule_tac x = "xa" in bspec, assumption) +apply (erule sym [THEN trans, symmetric]) +apply (rule restrict_eq_imp_val_eq [symmetric]) + apply (drule bspec [OF _ nat_succI], assumption) + apply (simp add: RR_def) +apply (drule lemma2, assumption+) +apply (blast dest!: domain_of_fun + intro: nat_into_Ord OrdmemD [THEN subsetD]) +done - defines - XX_def "XX == (\\n \\ nat. {f \\ n->X. \\k \\ n. \\ R})" - RR_def "RR == {:XX*XX. domain(z2)=succ(domain(z1)) - & restrict(z2, domain(z1)) = z1}" +lemma (in DC0_imp) lemma3: + "[| \n \ nat. \ RR; f \ nat -> XX; m \ nat |] + ==> (\x \ nat. f`succ(x)`x) `` m = f`succ(m)``m" +apply (erule natE, simp) +apply (subst image_lam) + apply (fast elim!: OrdmemD [OF nat_succI Ord_nat]) +apply (subst lemma3_1, assumption+) + apply fast +apply (fast dest!: lemma2 + elim!: image_fun [symmetric, OF _ OrdmemD [OF _ nat_into_Ord]]) +done + + +theorem DC0_imp_DC_nat: "DC0 ==> DC(nat)" +apply (unfold DC_def DC0_def, clarify) +apply (elim allE) +apply (erule impE) + (*these three results comprise Lemma 1*) +apply (blast intro!: DC0_imp.lemma1_1 DC0_imp.lemma1_2 DC0_imp.lemma1_3) +apply (erule bexE) +apply (rule_tac x = "\n \ nat. f`succ (n) `n" in rev_bexI) + apply (rule lam_type, blast dest!: DC0_imp.lemma2 intro: fun_weaken_type) +apply (rule oallI) +apply (frule DC0_imp.lemma2, assumption) + apply (blast intro: fun_weaken_type) + apply (erule ltD) +(** LEVEL 11: last subgoal **) +apply (subst DC0_imp.lemma3, assumption+) + apply (fast elim!: fun_weaken_type) + apply (erule ltD, force) +done + + +(* ************************************************************************ + DC(omega) ==> DC + + The scheme of the proof: + + Assume DC(omega). Let R and x satisfy the premise of DC. + + Define XX and RR as follows: + + XX = (\n \ nat. {f \ succ(n)->domain(R). \k \ n. \ R}) + + RR = {:Fin(XX)*XX. + (domain(z2)=succ(\f \ z1. domain(f)) & + (\f \ z1. restrict(z2, domain(f)) = f)) | + (~ (\g \ XX. domain(g)=succ(\f \ z1. domain(f)) & + (\f \ z1. restrict(g, domain(f)) = f)) & + z2={<0,x>})} + + Then XX and RR satisfy the hypotheses of DC(omega). + So applying DC: + + \f \ nat->XX. \n \ nat. f``n RR f`n + + Thence + + ff = {. n \ nat} + + is the desired function. + +************************************************************************* *) + +lemma singleton_in_funs: + "x \ X ==> {<0,x>} \ + (\n \ nat. {f \ succ(n)->X. \k \ n. \ R})" +apply (rule nat_0I [THEN UN_I]) +apply (force simp add: singleton_0 [symmetric] + intro!: singleton_fun [THEN Pi_type]) +done locale imp_DC0 = - fixes - XX :: i - RR :: i - x :: i - R :: i - f :: i - allRR :: o + fixes XX and RR and x and R and f and allRR + defines XX_def: "XX == (\n \ nat. + {f \ succ(n)->domain(R). \k \ n. \ R})" + and RR_def: + "RR == {:Fin(XX)*XX. + (domain(z2)=succ(\f \ z1. domain(f)) + & (\f \ z1. restrict(z2, domain(f)) = f)) + | (~ (\g \ XX. domain(g)=succ(\f \ z1. domain(f)) + & (\f \ z1. restrict(g, domain(f)) = f)) & z2={<0,x>})}" + and allRR_def: + "allRR == \b \ + {\Fin(XX)*XX. (domain(z2)=succ(\f \ z1. domain(f)) + & (\f \ z1. domain(f)) = b + & (\f \ z1. restrict(z2,domain(f)) = f))}" + +lemma (in imp_DC0) lemma4: + "[| range(R) \ domain(R); x \ domain(R) |] + ==> RR \ Pow(XX)*XX & + (\Y \ Pow(XX). Y \ nat --> (\x \ XX. :RR))" +apply (rule conjI) +apply (force dest!: FinD [THEN PowI] simp add: RR_def) +apply (rule impI [THEN ballI]) +apply (drule Finite_Fin [OF lesspoll_nat_is_Finite PowD], assumption) +apply (case_tac + "\g \ XX. domain (g) = + succ(\f \ Y. domain(f)) & (\f\Y. restrict(g, domain(f)) = f)") +apply (simp add: RR_def, blast) +apply (safe del: domainE) +apply (unfold XX_def RR_def) +apply (rule rev_bexI, erule singleton_in_funs) +apply (simp add: nat_0I [THEN rev_bexI] cons_fun_type2) +done + +lemma (in imp_DC0) UN_image_succ_eq: + "[| f \ nat->X; n \ nat |] + ==> (\x \ f``succ(n). P(x)) = P(f`n) Un (\x \ f``n. P(x))" +by (simp add: image_fun OrdmemD) + +lemma (in imp_DC0) UN_image_succ_eq_succ: + "[| (\x \ f``n. P(x)) = y; P(f`n) = succ(y); + f \ nat -> X; n \ nat |] ==> (\x \ f``succ(n). P(x)) = succ(y)" +by (simp add: UN_image_succ_eq, blast) + +lemma (in imp_DC0) apply_domain_type: + "[| h \ succ(n) -> D; n \ nat; domain(h)=succ(y) |] ==> h`y \ D" +by (fast elim: apply_type dest!: trans [OF sym domain_of_fun]) + +lemma (in imp_DC0) image_fun_succ: + "[| h \ nat -> X; n \ nat |] ==> h``succ(n) = cons(h`n, h``n)" +by (simp add: image_fun OrdmemD) + +lemma (in imp_DC0) f_n_type: + "[| domain(f`n) = succ(k); f \ nat -> XX; n \ nat |] + ==> f`n \ succ(k) -> domain(R)" +apply (unfold XX_def) +apply (drule apply_type, assumption) +apply (fast elim: domain_eq_imp_fun_type) +done + +lemma (in imp_DC0) f_n_pairs_in_R [rule_format]: + "[| h \ nat -> XX; domain(h`n) = succ(k); n \ nat |] + ==> \i \ k. \ R" +apply (unfold XX_def) +apply (drule apply_type, assumption) +apply (elim UN_E CollectE) +apply (drule domain_of_fun [symmetric, THEN trans], assumption) +apply simp +done + +lemma (in imp_DC0) restrict_cons_eq_restrict: + "[| restrict(h, domain(u))=u; h \ n->X; domain(u) \ n |] + ==> restrict(cons(, h), domain(u)) = u" +apply (unfold restrict_def) +apply (erule sym [THEN trans, symmetric]) +apply (rule fun_extension) +apply (fast intro!: lam_type) +apply (fast intro!: lam_type) +apply (simp add: subsetD [THEN cons_val_k]) +done + +lemma (in imp_DC0) all_in_image_restrict_eq: + "[| \x \ f``n. restrict(f`n, domain(x))=x; + f \ nat -> XX; + n \ nat; domain(f`n) = succ(n); + (\x \ f``n. domain(x)) \ n |] + ==> \x \ f``succ(n). restrict(cons(, f`n), domain(x)) = x" +apply (rule ballI) +apply (simp add: image_fun_succ) +apply (drule f_n_type, assumption+) +apply (erule disjE) + apply (simp add: domain_of_fun restrict_cons_eq) +apply (blast intro!: restrict_cons_eq_restrict) +done + +lemma (in imp_DC0) simplify_recursion: + "[| \b \ RR; + f \ nat -> XX; range(R) \ domain(R); x \ domain(R)|] + ==> allRR" +apply (unfold RR_def allRR_def) +apply (rule oallI, drule ltD) +apply (erule nat_induct) +apply (drule_tac x="0" in ospec, blast intro: Limit_has_0) +apply (force simp add: singleton_fun [THEN domain_of_fun] singleton_in_funs) +(*induction step*) (** LEVEL 5 **) +(*prevent simplification of ~\ to \~ *) +apply (simp only: separation split) +apply (drule_tac x="succ(xa)" in ospec, blast intro: ltI); +apply (elim conjE disjE) +apply (force elim!: trans subst_context + intro!: UN_image_succ_eq_succ) +apply (erule notE) +apply (simp add: XX_def UN_image_succ_eq_succ) +apply (elim conjE bexE) +apply (drule apply_domain_type, assumption+) +apply (erule domainE)+ +apply (frule f_n_type) +apply (simp add: XX_def, assumption+) +apply (rule rev_bexI, erule nat_succI) +apply (rule_tac x = "cons (, f`xa) " in bexI) +prefer 2 apply (blast intro: cons_fun_type2) +apply (rule conjI) +prefer 2 apply (fast del: ballI subsetI + elim: trans [OF _ subst_context, THEN domain_cons_eq_succ] + subst_context + all_in_image_restrict_eq [simplified XX_def] + trans equalityD1) +(*one remaining subgoal*) +apply (rule ballI) +apply (erule succE) +(** LEVEL 25 **) + apply (simp add: cons_val_n cons_val_k) +(*assumption+ will not perform the required backtracking!*) +apply (drule f_n_pairs_in_R [simplified XX_def, OF _ domain_of_fun], + assumption, assumption, assumption) +apply (simp add: nat_into_Ord [THEN succ_in_succ] succI2 cons_val_k) +done + - defines - XX_def "XX == (\\n \\ nat. - {f \\ succ(n)->domain(R). \\k \\ n. \\ R})" - RR_def - "RR == {:Fin(XX)*XX. - (domain(z2)=succ(\\f \\ z1. domain(f)) - & (\\f \\ z1. restrict(z2, domain(f)) = f)) - | (~ (\\g \\ XX. domain(g)=succ(\\f \\ z1. domain(f)) - & (\\f \\ z1. restrict(g, domain(f)) = f)) & z2={<0,x>})}" - allRR_def - "allRR == \\b \\ - {:Fin(XX)*XX. (domain(z2)=succ(\\f \\ z1. domain(f)) - & (\\f \\ z1. domain(f)) = b - & (\\f \\ z1. restrict(z2,domain(f)) = f))}" +lemma (in imp_DC0) lemma2: + "[| allRR; f \ nat->XX; range(R) \ domain(R); x \ domain(R); n \ nat |] + ==> f`n \ succ(n) -> domain(R) & (\i \ n. :R)" +apply (unfold allRR_def) +apply (drule ospec) +apply (erule ltI [OF _ Ord_nat]) +apply (erule CollectE, simp) +apply (rule conjI) +prefer 2 apply (fast elim!: f_n_pairs_in_R trans subst_context) +apply (unfold XX_def) +apply (fast elim!: trans [THEN domain_eq_imp_fun_type] subst_context) +done + +lemma (in imp_DC0) lemma3: + "[| allRR; f \ nat->XX; n\nat; range(R) \ domain(R); x \ domain(R) |] + ==> f`n`n = f`succ(n)`n" +apply (frule lemma2 [THEN conjunct1, THEN domain_of_fun], assumption+) +apply (unfold allRR_def) +apply (drule ospec) +apply (drule ltI [OF nat_succI Ord_nat], assumption) +apply simp +apply (elim conjE ballE) +apply (erule restrict_eq_imp_val_eq [symmetric], force) +apply (simp add: image_fun OrdmemD) +done + + +theorem DC_nat_imp_DC0: "DC(nat) ==> DC0" +apply (unfold DC_def DC0_def) +apply (intro allI impI) +apply (erule asm_rl conjE ex_in_domain [THEN exE] allE)+ +apply (erule impE [OF _ imp_DC0.lemma4], assumption+) +apply (erule bexE) +apply (drule imp_DC0.simplify_recursion, assumption+) +apply (rule_tac x = "\n \ nat. f`n`n" in bexI) +apply (rule_tac [2] lam_type) +apply (erule_tac [2] apply_type [OF imp_DC0.lemma2 [THEN conjunct1] succI1]) +apply (rule ballI) +apply (frule_tac n="succ(n)" in imp_DC0.lemma2, + (assumption|erule nat_succI)+) +apply (drule imp_DC0.lemma3, auto) +done + +(* ********************************************************************** *) +(* \K. Card(K) --> DC(K) ==> WO3 *) +(* ********************************************************************** *) + +lemma fun_Ord_inj: + "[| f \ a->X; Ord(a); + !!b c. [| b a |] ==> f`b\f`c |] + ==> f \ inj(a, X)" +apply (unfold inj_def, simp) +apply (intro ballI impI) +apply (rule_tac j=x in Ord_in_Ord [THEN Ord_linear_lt], assumption+) +apply (blast intro: Ord_in_Ord, auto) +apply (atomize, blast dest: not_sym) +done + +lemma value_in_image: "[| f \ X->Y; A \ X; a \ A |] ==> f`a \ f``A" +by (fast elim!: image_fun [THEN ssubst]); + +theorem DC_WO3: "(\K. Card(K) --> DC(K)) ==> WO3" +apply (unfold DC_def WO3_def) +apply (rule allI) +apply (case_tac "A \ Hartog (A)"); +apply (fast dest!: lesspoll_imp_ex_lt_eqpoll + intro!: Ord_Hartog leI [THEN le_imp_subset]) +apply (erule allE impE)+ +apply (rule Card_Hartog) +apply (erule_tac x = "A" in allE) +apply (erule_tac x = "{ \ Pow (A) *A . z1 \ Hartog (A) & z2 \ z1}" + in allE) +apply simp +apply (erule impE, fast elim: lesspoll_lemma [THEN not_emptyE]) +apply (erule bexE) +apply (rule Hartog_lepoll_selfE) +apply (rule lepoll_def [THEN def_imp_iff, THEN iffD2]) +apply (rule exI, rule fun_Ord_inj, assumption, rule Ord_Hartog) +apply (drule value_in_image) +apply (drule OrdmemD, rule Ord_Hartog, assumption+, erule ltD) +apply (drule ospec) +apply (blast intro: ltI Ord_Hartog, force) +done + +(* ********************************************************************** *) +(* WO1 ==> \K. Card(K) --> DC(K) *) +(* ********************************************************************** *) + +lemma images_eq: + "[| \x \ A. f`x=g`x; f \ Df->Cf; g \ Dg->Cg; A \ Df; A \ Dg |] + ==> f``A = g``A" +apply (simp (no_asm_simp) add: image_fun) +done + +lemma lam_images_eq: + "[| Ord(a); b \ a |] ==> (\x \ a. h(x))``b = (\x \ b. h(x))``b" +apply (rule images_eq) + apply (rule ballI) + apply (drule OrdmemD [THEN subsetD], assumption+) + apply simp + apply (fast elim!: RepFunI OrdmemD intro!: lam_type)+ +done + +lemma lam_type_RepFun: "(\b \ a. h(b)) \ a -> {h(b). b \ a}" +by (fast intro!: lam_type RepFunI) + +lemma lemmaX: + "[| \Y \ Pow(X). Y \ K --> (\x \ X. \ R); + b \ K; Z \ Pow(X); Z \ K |] + ==> {x \ X. \ R} \ 0" +by blast + + +lemma WO1_DC_lemma: + "[| Card(K); well_ord(X,Q); + \Y \ Pow(X). Y \ K --> (\x \ X. \ R); b \ K |] + ==> ff(b, X, Q, R) \ {x \ X. <(\c \ b. ff(c, X, Q, R))``b, x> \ R}" +apply (rule_tac P = "b \ K" in impE, (erule_tac [2] asm_rl)+) +apply (rule_tac i=b in Card_is_Ord [THEN Ord_in_Ord, THEN trans_induct], + assumption+) +apply (rule impI) +apply (rule ff_def [THEN def_transrec, THEN ssubst]) +apply (erule the_first_in, fast) +apply (simp add: image_fun [OF lam_type_RepFun subset_refl]) +apply (erule lemmaX, assumption) + apply (blast intro: Card_is_Ord OrdmemD [THEN subsetD]) +apply (blast intro: lesspoll_trans1 in_Card_imp_lesspoll RepFun_lepoll) +done + +theorem WO1_DC_Card: "WO1 ==> \K. Card(K) --> DC(K)" +apply (unfold DC_def WO1_def) +apply (rule allI impI)+ +apply (erule allE exE conjE)+ +apply (rule_tac x = "\b \ K. ff (b, X, Ra, R) " in bexI) + apply (simp add: lam_images_eq [OF Card_is_Ord ltD]) + apply (fast elim!: ltE WO1_DC_lemma [THEN CollectD2]) +apply (rule_tac lam_type) +apply (rule WO1_DC_lemma [THEN CollectD1], assumption+) +done + end diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/DC_lemmas.ML --- a/src/ZF/AC/DC_lemmas.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,105 +0,0 @@ -(* Title: ZF/AC/DC_lemmas.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -More general lemmas used in the proofs concerning DC - -*) - -val [prem] = goalw thy [lepoll_def] - "Ord(a) ==> {P(b). b \\ a} lepoll a"; -by (res_inst_tac [("x","\\z \\ RepFun(a,P). LEAST i. z=P(i)")] exI 1); -by (res_inst_tac [("d","%z. P(z)")] (sym RSN (2, lam_injective)) 1); -by (fast_tac (claset() addSIs [Least_in_Ord, prem]) 1); -by (REPEAT (eresolve_tac [RepFunE, LeastI, prem RS Ord_in_Ord] 1)); -qed "RepFun_lepoll"; - -Goalw [lesspoll_def] "n \\ nat ==> n lesspoll nat"; -by (rtac conjI 1); -by (eresolve_tac [Ord_nat RSN (2, OrdmemD) RS subset_imp_lepoll] 1); -by (rtac notI 1); -by (etac eqpollE 1); -by (rtac succ_lepoll_natE 1 THEN (assume_tac 2)); -by (eresolve_tac [nat_succI RS (Ord_nat RSN (2, OrdmemD) RS - subset_imp_lepoll) RS lepoll_trans] 1 - THEN (assume_tac 1)); -qed "n_lesspoll_nat"; - -Goalw [lepoll_def] - "[| f \\ X->Y; Ord(X) |] ==> f``X lepoll X"; -by (res_inst_tac [("x","\\x \\ f``X. LEAST y. f`y = x")] exI 1); -by (res_inst_tac [("d","%z. f`z")] lam_injective 1); -by (fast_tac (claset() addSIs [Least_in_Ord, apply_equality]) 1); -by (fast_tac (claset() addSEs [Ord_in_Ord] addSIs [LeastI, apply_equality]) 1); -qed "image_Ord_lepoll"; - -val [major, minor] = goal thy - "[| (!!g. g \\ X ==> \\u. :R); R \\ X*X \ -\ |] ==> range(R) \\ domain(R)"; -by (rtac subsetI 1); -by (etac rangeE 1); -by (dresolve_tac [minor RS subsetD RS SigmaD2 RS major] 1); -by (Fast_tac 1); -qed "range_subset_domain"; - -val prems = goal thy "!!k. k \\ n ==> k\\n"; -by (fast_tac (claset() addSEs [mem_irrefl]) 1); -qed "mem_not_eq"; - -Goalw [succ_def] "g \\ n->X ==> cons(, g) \\ succ(n) -> cons(x, X)"; -by (fast_tac (claset() addSIs [fun_extend] addSEs [mem_irrefl]) 1); -qed "cons_fun_type"; - -Goal "[| g \\ n->X; x \\ X |] ==> cons(, g) \\ succ(n) -> X"; -by (etac (cons_absorb RS subst) 1 THEN etac cons_fun_type 1); -qed "cons_fun_type2"; - -Goal "n \\ nat ==> cons(, g)``n = g``n"; -by (fast_tac (claset() addSEs [mem_irrefl]) 1); -qed "cons_image_n"; - -Goal "g \\ n->X ==> cons(, g)`n = x"; -by (fast_tac (claset() addSIs [apply_equality] addSEs [cons_fun_type]) 1); -qed "cons_val_n"; - -Goal "k \\ n ==> cons(, g)``k = g``k"; -by (fast_tac (claset() addEs [mem_asym]) 1); -qed "cons_image_k"; - -Goal "[| k \\ n; g \\ n->X |] ==> cons(, g)`k = g`k"; -by (fast_tac (claset() addSIs [apply_equality, consI2] addSEs [cons_fun_type, apply_Pair]) 1); -qed "cons_val_k"; - -Goal "domain(f)=x ==> domain(cons(, f)) = succ(x)"; -by (asm_full_simp_tac (simpset() addsimps [domain_cons, succ_def]) 1); -qed "domain_cons_eq_succ"; - -Goalw [restrict_def] "g \\ n->X ==> restrict(cons(, g), n)=g"; -by (rtac fun_extension 1); -by (rtac lam_type 1); -by (eresolve_tac [cons_fun_type RS apply_type] 1); -by (etac succI2 1); -by (assume_tac 1); -by (asm_full_simp_tac (simpset() addsimps [cons_val_k]) 1); -qed "restrict_cons_eq"; - -Goal "[| Ord(k); i \\ k |] ==> succ(i) \\ succ(k)"; -by (resolve_tac [Ord_linear RS disjE] 1 THEN (assume_tac 3)); -by (REPEAT (fast_tac (claset() addEs [Ord_in_Ord, mem_irrefl, mem_asym]) 1)); -qed "succ_in_succ"; - -Goalw [restrict_def] - "[| restrict(f, domain(g)) = g; x \\ domain(g) |] ==> f`x = g`x"; -by (etac subst 1); -by (Asm_full_simp_tac 1); -qed "restrict_eq_imp_val_eq"; - -Goal "[| domain(f)=A; f \\ B->C |] ==> f \\ A->C"; -by (ftac domain_of_fun 1); -by (Fast_tac 1); -qed "domain_eq_imp_fun_type"; - -Goal "[| R \\ A * B; R \\ 0 |] ==> \\x. x \\ domain(R)"; -by (fast_tac (claset() addSEs [not_emptyE]) 1); -qed "ex_in_domain"; - diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/DC_lemmas.thy --- a/src/ZF/AC/DC_lemmas.thy Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,3 +0,0 @@ -(*Dummy theory to document dependencies *) - -DC_lemmas = AC_Equiv + Cardinal_aux diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/HH.ML --- a/src/ZF/AC/HH.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,212 +0,0 @@ -(* Title: ZF/AC/HH.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -Some properties of the recursive definition of HH used in the proofs of - AC17 ==> AC1 - AC1 ==> WO2 - AC15 ==> WO6 -*) - -(* ********************************************************************** *) -(* Lemmas useful in each of the three proofs *) -(* ********************************************************************** *) - -Goal "HH(f,x,a) = \ -\ (let z = x - (\\b \\ a. HH(f,x,b)) \ -\ in if(f`z \\ Pow(z)-{0}, f`z, {x}))"; -by (resolve_tac [HH_def RS def_transrec RS trans] 1); -by (Simp_tac 1); -qed "HH_def_satisfies_eq"; - -Goal "HH(f,x,a) \\ Pow(x)-{0} | HH(f,x,a)={x}"; -by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1); -by (simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI]) 1); -by (Fast_tac 1); -qed "HH_values"; - -Goal "B \\ A ==> X-(\\a \\ A. P(a)) = X-(\\a \\ A-B. P(a))-(\\b \\ B. P(b))"; -by (Fast_tac 1); -qed "subset_imp_Diff_eq"; - -Goal "[| c \\ a-b; b c=b | b A ==> P(y) = {x}) ==> x - (\\y \\ A. P(y)) = x"; -by (asm_full_simp_tac (simpset() addsimps prems) 1); -by (fast_tac (claset() addSDs [prem] addSEs [mem_irrefl]) 1); -qed "Diff_UN_eq_self"; - -Goal "x - (\\b \\ a. HH(f,x,b)) = x - (\\b \\ a1. HH(f,x,b)) \ -\ ==> HH(f,x,a) = HH(f,x,a1)"; -by (resolve_tac [HH_def_satisfies_eq RS - (HH_def_satisfies_eq RS sym RSN (3, trans RS trans))] 1); -by (etac subst_context 1); -qed "HH_eq"; - -Goal "[| HH(f,x,b)={x}; b HH(f,x,a)={x}"; -by (res_inst_tac [("P","b Pow(x)-{0}; b HH(f,x,b) \\ Pow(x)-{0}"; -by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 1)); -by (dtac HH_is_x_gt_too 1 THEN (assume_tac 1)); -by (dtac subst 1 THEN (assume_tac 1)); -by (fast_tac (claset() addSEs [mem_irrefl]) 1); -qed "HH_subset_x_lt_too"; - -Goal "HH(f,x,a) \\ Pow(x)-{0} \ -\ ==> HH(f,x,a) \\ Pow(x - (\\b \\ a. HH(f,x,b)))-{0}"; -by (dresolve_tac [HH_def_satisfies_eq RS subst] 1); -by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1); -by (asm_full_simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI]) 1); -by (dresolve_tac [split_if RS iffD1] 1); -by (Simp_tac 1); -by (fast_tac (subset_cs addSEs [mem_irrefl]) 1); -qed "HH_subset_x_imp_subset_Diff_UN"; - -Goal "[| HH(f,x,v)=HH(f,x,w); HH(f,x,v): Pow(x)-{0}; v \\ w |] ==> P"; -by (forw_inst_tac [("P","%y. y \\ Pow(x)-{0}")] subst 1 THEN (assume_tac 1)); -by (dres_inst_tac [("a","w")] HH_subset_x_imp_subset_Diff_UN 1); -by (dtac subst_elem 1 THEN (assume_tac 1)); -by (fast_tac (claset() addSIs [singleton_iff RS iffD2, equals0I]) 1); -qed "HH_eq_arg_lt"; - -Goal "[| HH(f,x,v)=HH(f,x,w); HH(f,x,w): Pow(x)-{0}; \ -\ Ord(v); Ord(w) |] ==> v=w"; -by (res_inst_tac [("j","w")] Ord_linear_lt 1 THEN TRYALL assume_tac); -by (resolve_tac [sym RS (ltD RSN (3, HH_eq_arg_lt))] 2 - THEN REPEAT (assume_tac 2)); -by (dtac subst_elem 1 THEN (assume_tac 1)); -by (fast_tac (FOL_cs addDs [ltD] addSEs [HH_eq_arg_lt]) 1); -qed "HH_eq_imp_arg_eq"; - -Goalw [lepoll_def, inj_def] - "[| HH(f, x, i) \\ Pow(x)-{0}; Ord(i) |] ==> i lepoll Pow(x)-{0}"; -by (res_inst_tac [("x","\\j \\ i. HH(f, x, j)")] exI 1); -by (Asm_simp_tac 1); -by (fast_tac (FOL_cs addSEs [HH_eq_imp_arg_eq, Ord_in_Ord, HH_subset_x_lt_too] - addSIs [lam_type, ballI, ltI] addIs [bexI]) 1); -qed "HH_subset_x_imp_lepoll"; - -Goal "HH(f, x, Hartog(Pow(x)-{0})) = {x}"; -by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 2)); -by (fast_tac (FOL_cs addSDs [HH_subset_x_imp_lepoll] - addSIs [Ord_Hartog] addSEs [HartogE]) 1); -qed "HH_Hartog_is_x"; - -Goal "HH(f, x, LEAST i. HH(f, x, i) = {x}) = {x}"; -by (fast_tac (claset() addSIs [Ord_Hartog, HH_Hartog_is_x, LeastI]) 1); -qed "HH_Least_eq_x"; - -Goal "a \\ (LEAST i. HH(f,x,i)={x}) ==> HH(f,x,a) \\ Pow(x)-{0}"; -by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 1)); -by (rtac less_LeastE 1); -by (eresolve_tac [Ord_Least RSN (2, ltI)] 2); -by (assume_tac 1); -qed "less_Least_subset_x"; - -(* ********************************************************************** *) -(* Lemmas used in the proofs of AC1 ==> WO2 and AC17 ==> AC1 *) -(* ********************************************************************** *) - -Goalw [inj_def] - "(\\a \\ (LEAST i. HH(f,x,i)={x}). HH(f,x,a)) \ -\ \\ inj(LEAST i. HH(f,x,i)={x}, Pow(x)-{0})"; -by (Asm_full_simp_tac 1); -by (fast_tac (claset() addSIs [lam_type] addDs [less_Least_subset_x] - addSEs [HH_eq_imp_arg_eq, Ord_Least RS Ord_in_Ord]) 1); -qed "lam_Least_HH_inj_Pow"; - -Goal "\\a \\ (LEAST i. HH(f,x,i)={x}). \\z \\ x. HH(f,x,a) = {z} \ -\ ==> (\\a \\ (LEAST i. HH(f,x,i)={x}). HH(f,x,a)) \ -\ \\ inj(LEAST i. HH(f,x,i)={x}, {{y}. y \\ x})"; -by (resolve_tac [lam_Least_HH_inj_Pow RS inj_strengthen_type] 1); -by (Asm_full_simp_tac 1); -qed "lam_Least_HH_inj"; - -Goalw [surj_def] - "[| x - (\\a \\ A. F(a)) = 0; \ -\ \\a \\ A. \\z \\ x. F(a) = {z} |] \ -\ ==> (\\a \\ A. F(a)) \\ surj(A, {{y}. y \\ x})"; -by (asm_full_simp_tac (simpset() addsimps [lam_type, Diff_eq_0_iff]) 1); -by Safe_tac; -by (set_mp_tac 1); -by (deepen_tac (claset() addSIs [bexI] addSEs [equalityE]) 4 1); -qed "lam_surj_sing"; - -Goal "y \\ Pow(x)-{0} ==> x \\ 0"; -by Auto_tac; -qed "not_emptyI2"; - -Goal "f`(x - (\\j \\ i. HH(f,x,j))): Pow(x - (\\j \\ i. HH(f,x,j)))-{0} \ -\ ==> HH(f, x, i) \\ Pow(x) - {0}"; -by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1); -by (asm_full_simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI, - not_emptyI2 RS if_P]) 1); -by (Fast_tac 1); -qed "f_subset_imp_HH_subset"; - -val [prem] = goal thy "(!!z. z \\ Pow(x)-{0} ==> f`z \\ Pow(z)-{0}) ==> \ -\ x - (\\j \\ (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0"; -by (excluded_middle_tac "?P \\ {0}" 1); -by (Fast_tac 2); -by (dresolve_tac [Diff_subset RS PowI RS DiffI RS prem RS - f_subset_imp_HH_subset] 1); -by (fast_tac (claset() addSDs [HH_Least_eq_x RS sym RSN (2, subst_elem)] - addSEs [mem_irrefl]) 1); -qed "f_subsets_imp_UN_HH_eq_x"; - -Goal "HH(f,x,i)=f`(x - (\\j \\ i. HH(f,x,j))) | HH(f,x,i)={x}"; -by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1); -by (simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI]) 1); -qed "HH_values2"; - -Goal "HH(f,x,i): Pow(x)-{0} ==> HH(f,x,i)=f`(x - (\\j \\ i. HH(f,x,j)))"; -by (resolve_tac [HH_values2 RS disjE] 1 THEN (assume_tac 1)); -by (fast_tac (claset() addSEs [equalityE, mem_irrefl] - addSDs [singleton_subsetD]) 1); -qed "HH_subset_imp_eq"; - -Goal "[| f \\ (Pow(x)-{0}) -> {{z}. z \\ x}; \ -\ a \\ (LEAST i. HH(f,x,i)={x}) |] ==> \\z \\ x. HH(f,x,a) = {z}"; -by (dtac less_Least_subset_x 1); -by (ftac HH_subset_imp_eq 1); -by (dtac apply_type 1); -by (resolve_tac [Diff_subset RS PowI RS DiffI] 1); -by (fast_tac - (claset() addSDs [HH_subset_x_imp_subset_Diff_UN RS not_emptyI2]) 1); -by (fast_tac (claset() addss (simpset())) 1); -qed "f_sing_imp_HH_sing"; - -Goalw [bij_def] - "[| x - (\\j \\ (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0; \ -\ f \\ (Pow(x)-{0}) -> {{z}. z \\ x} |] \ -\ ==> (\\a \\ (LEAST i. HH(f,x,i)={x}). HH(f,x,a)) \ -\ \\ bij(LEAST i. HH(f,x,i)={x}, {{y}. y \\ x})"; -by (fast_tac (claset() addSIs [lam_Least_HH_inj, lam_surj_sing, - f_sing_imp_HH_sing]) 1); -qed "f_sing_lam_bij"; - -Goal "f \\ (\\X \\ Pow(x)-{0}. F(X)) \ -\ ==> (\\X \\ Pow(x)-{0}. {f`X}) \\ (\\X \\ Pow(x)-{0}. {{z}. z \\ F(X)})"; -by (fast_tac (FOL_cs addSIs [lam_type, RepFun_eqI, singleton_eq_iff RS iffD2] - addDs [apply_type]) 1); -qed "lam_singI"; - -val bij_Least_HH_x = - (lam_singI RSN (2, [f_sing_lam_bij, lam_sing_bij RS bij_converse_bij] - MRS comp_bij)) |> standard; diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/HH.thy --- a/src/ZF/AC/HH.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/HH.thy Wed Jan 16 17:52:06 2002 +0100 @@ -2,22 +2,255 @@ ID: $Id$ Author: Krzysztof Grabczewski -The theory file for the proofs of +Some properties of the recursive definition of HH used in the proofs of AC17 ==> AC1 AC1 ==> WO2 AC15 ==> WO6 *) -HH = AC_Equiv + Hartog + WO_AC + Let + +theory HH = AC_Equiv + Hartog: -consts +constdefs - HH :: [i, i, i] => i + HH :: "[i, i, i] => i" + "HH(f,x,a) == transrec(a, %b r. let z = x - (\c \ b. r`c) + in if f`z \ Pow(z)-{0} then f`z else {x})" + + +(* ********************************************************************** *) +(* Lemmas useful in each of the three proofs *) +(* ********************************************************************** *) + +lemma HH_def_satisfies_eq: + "HH(f,x,a) = (let z = x - (\b \ a. HH(f,x,b)) + in if f`z \ Pow(z)-{0} then f`z else {x})" +by (rule HH_def [THEN def_transrec, THEN trans], simp) + +lemma HH_values: "HH(f,x,a) \ Pow(x)-{0} | HH(f,x,a)={x}" +apply (rule HH_def_satisfies_eq [THEN ssubst]) +apply (simp add: Let_def Diff_subset [THEN PowI], fast) +done + +lemma subset_imp_Diff_eq: + "B \ A ==> X-(\a \ A. P(a)) = X-(\a \ A-B. P(a))-(\b \ B. P(b))" +by fast + +lemma Ord_DiffE: "[| c \ a-b; b c=b | b A ==> P(y) = {x}) ==> x - (\y \ A. P(y)) = x" +apply (simp, fast elim!: mem_irrefl) +done + +lemma HH_eq: "x - (\b \ a. HH(f,x,b)) = x - (\b \ a1. HH(f,x,b)) + ==> HH(f,x,a) = HH(f,x,a1)" +apply (subst HH_def_satisfies_eq) +apply (rule HH_def_satisfies_eq [THEN trans], simp) +done + +lemma HH_is_x_gt_too: "[| HH(f,x,b)={x}; b HH(f,x,a)={x}" +apply (rule_tac P = "b Pow(x)-{0}; b HH(f,x,b) \ Pow(x)-{0}" +apply (rule HH_values [THEN disjE], assumption) +apply (drule HH_is_x_gt_too, assumption) +apply (drule subst, assumption) +apply (fast elim!: mem_irrefl) +done + +lemma HH_subset_x_imp_subset_Diff_UN: + "HH(f,x,a) \ Pow(x)-{0} ==> HH(f,x,a) \ Pow(x - (\b \ a. HH(f,x,b)))-{0}" +apply (drule HH_def_satisfies_eq [THEN subst]) +apply (rule HH_def_satisfies_eq [THEN ssubst]) +apply (simp add: Let_def Diff_subset [THEN PowI]) +apply (drule split_if [THEN iffD1]) +apply (fast elim!: mem_irrefl) +done + +lemma HH_eq_arg_lt: + "[| HH(f,x,v)=HH(f,x,w); HH(f,x,v) \ Pow(x)-{0}; v \ w |] ==> P" +apply (frule_tac P = "%y. y \ Pow (x) -{0}" in subst, assumption) +apply (drule_tac a = "w" in HH_subset_x_imp_subset_Diff_UN) +apply (drule subst_elem, assumption) +apply (fast intro!: singleton_iff [THEN iffD2] equals0I) +done + +lemma HH_eq_imp_arg_eq: + "[| HH(f,x,v)=HH(f,x,w); HH(f,x,w) \ Pow(x)-{0}; Ord(v); Ord(w) |] ==> v=w" +apply (rule_tac j = "w" in Ord_linear_lt) +apply (simp_all (no_asm_simp)) + apply (drule subst_elem, assumption) + apply (blast dest: ltD HH_eq_arg_lt) +apply (blast dest: HH_eq_arg_lt [OF sym] ltD) +done + +lemma HH_subset_x_imp_lepoll: + "[| HH(f, x, i) \ Pow(x)-{0}; Ord(i) |] ==> i lepoll Pow(x)-{0}" +apply (unfold lepoll_def inj_def) +apply (rule_tac x = "\j \ i. HH (f, x, j) " in exI) +apply (simp (no_asm_simp)) +apply (fast del: DiffE + elim!: HH_eq_imp_arg_eq Ord_in_Ord HH_subset_x_lt_too + intro!: lam_type ballI ltI intro: bexI) +done + +lemma HH_Hartog_is_x: "HH(f, x, Hartog(Pow(x)-{0})) = {x}" +apply (rule HH_values [THEN disjE]) +prefer 2 apply assumption +apply (fast del: DiffE + intro!: Ord_Hartog + dest!: HH_subset_x_imp_lepoll + elim!: Hartog_lepoll_selfE) +done + +lemma HH_Least_eq_x: "HH(f, x, LEAST i. HH(f, x, i) = {x}) = {x}" +by (fast intro!: Ord_Hartog HH_Hartog_is_x LeastI) + +lemma less_Least_subset_x: + "a \ (LEAST i. HH(f,x,i)={x}) ==> HH(f,x,a) \ Pow(x)-{0}" +apply (rule HH_values [THEN disjE], assumption) +apply (rule less_LeastE) +apply (erule_tac [2] ltI [OF _ Ord_Least], assumption) +done -defs +(* ********************************************************************** *) +(* Lemmas used in the proofs of AC1 ==> WO2 and AC17 ==> AC1 *) +(* ********************************************************************** *) + +lemma lam_Least_HH_inj_Pow: + "(\a \ (LEAST i. HH(f,x,i)={x}). HH(f,x,a)) + \ inj(LEAST i. HH(f,x,i)={x}, Pow(x)-{0})" +apply (unfold inj_def, simp) +apply (fast intro!: lam_type dest: less_Least_subset_x + elim!: HH_eq_imp_arg_eq Ord_Least [THEN Ord_in_Ord]) +done + +lemma lam_Least_HH_inj: + "\a \ (LEAST i. HH(f,x,i)={x}). \z \ x. HH(f,x,a) = {z} + ==> (\a \ (LEAST i. HH(f,x,i)={x}). HH(f,x,a)) + \ inj(LEAST i. HH(f,x,i)={x}, {{y}. y \ x})" +by (rule lam_Least_HH_inj_Pow [THEN inj_strengthen_type], simp) + +lemma lam_surj_sing: + "[| x - (\a \ A. F(a)) = 0; \a \ A. \z \ x. F(a) = {z} |] + ==> (\a \ A. F(a)) \ surj(A, {{y}. y \ x})" +apply (simp add: surj_def lam_type Diff_eq_0_iff) +apply (blast elim: equalityE) +done + +lemma not_emptyI2: "y \ Pow(x)-{0} ==> x \ 0" +by auto + +lemma f_subset_imp_HH_subset: + "f`(x - (\j \ i. HH(f,x,j))) \ Pow(x - (\j \ i. HH(f,x,j)))-{0} + ==> HH(f, x, i) \ Pow(x) - {0}" +apply (rule HH_def_satisfies_eq [THEN ssubst]) +apply (simp add: Let_def Diff_subset [THEN PowI] not_emptyI2 [THEN if_P], fast) +done + +lemma f_subsets_imp_UN_HH_eq_x: + "\z \ Pow(x)-{0}. f`z \ Pow(z)-{0} + ==> x - (\j \ (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0" +apply (case_tac "?P \ {0}", fast) +apply (drule Diff_subset [THEN PowI, THEN DiffI]) +apply (drule bspec, assumption) +apply (drule f_subset_imp_HH_subset) +apply (blast dest!: subst_elem [OF _ HH_Least_eq_x [symmetric]] + elim!: mem_irrefl) +done + +lemma HH_values2: "HH(f,x,i) = f`(x - (\j \ i. HH(f,x,j))) | HH(f,x,i)={x}" +apply (rule HH_def_satisfies_eq [THEN ssubst]) +apply (simp add: Let_def Diff_subset [THEN PowI]) +done + +lemma HH_subset_imp_eq: + "HH(f,x,i): Pow(x)-{0} ==> HH(f,x,i)=f`(x - (\j \ i. HH(f,x,j)))" +apply (rule HH_values2 [THEN disjE], assumption) +apply (fast elim!: equalityE mem_irrefl dest!: singleton_subsetD) +done - HH_def "HH(f,x,a) == transrec(a, %b r. let z = x - (\\c \\ b. r`c) - in if(f`z \\ Pow(z)-{0}, f`z, {x}))" +lemma f_sing_imp_HH_sing: + "[| f \ (Pow(x)-{0}) -> {{z}. z \ x}; + a \ (LEAST i. HH(f,x,i)={x}) |] ==> \z \ x. HH(f,x,a) = {z}" +apply (drule less_Least_subset_x) +apply (frule HH_subset_imp_eq) +apply (drule apply_type) +apply (rule Diff_subset [THEN PowI, THEN DiffI]) +apply (fast dest!: HH_subset_x_imp_subset_Diff_UN [THEN not_emptyI2], force) +done + +lemma f_sing_lam_bij: + "[| x - (\j \ (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0; + f \ (Pow(x)-{0}) -> {{z}. z \ x} |] + ==> (\a \ (LEAST i. HH(f,x,i)={x}). HH(f,x,a)) + \ bij(LEAST i. HH(f,x,i)={x}, {{y}. y \ x})" +apply (unfold bij_def) +apply (fast intro!: lam_Least_HH_inj lam_surj_sing f_sing_imp_HH_sing) +done + +lemma lam_singI: + "f \ (\X \ Pow(x)-{0}. F(X)) + ==> (\X \ Pow(x)-{0}. {f`X}) \ (\X \ Pow(x)-{0}. {{z}. z \ F(X)})" +by (fast del: DiffI DiffE + intro!: lam_type singleton_eq_iff [THEN iffD2] dest: apply_type) + +(*FIXME: both uses have the form ...[THEN bij_converse_bij], so + simplification is needed!*) +lemmas bij_Least_HH_x = + comp_bij [OF f_sing_lam_bij [OF _ lam_singI] + lam_sing_bij [THEN bij_converse_bij], standard] + + +(* ********************************************************************** *) +(* The proof of AC1 ==> WO2 *) +(* ********************************************************************** *) + +(*Establishing the existence of a bijection, namely +converse + (converse(\x\x. {x}) O + Lambda + (LEAST i. HH(\X\Pow(x) - {0}. {f ` X}, x, i) = {x}, + HH(\X\Pow(x) - {0}. {f ` X}, x))) +Perhaps it could be simplified. *) + +lemma bijection: + "f \ (\X \ Pow(x) - {0}. X) + ==> \g. g \ bij(x, LEAST i. HH(\X \ Pow(x)-{0}. {f`X}, x, i) = {x})" +apply (rule exI) +apply (rule bij_Least_HH_x [THEN bij_converse_bij]) +apply (rule f_subsets_imp_UN_HH_eq_x) +apply (intro ballI apply_type) +apply (fast intro: lam_type apply_type del: DiffE) +apply assumption; +apply (fast intro: Pi_weaken_type) +done + +lemma AC1_WO2: "AC1 ==> WO2" +apply (unfold AC1_def WO2_def eqpoll_def) +apply (intro allI) +apply (drule_tac x = "Pow(A) - {0}" in spec) +apply (blast dest: bijection) +done end + diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/Hartog.ML --- a/src/ZF/AC/Hartog.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,82 +0,0 @@ -(* Title: ZF/AC/Hartog.ML - ID: $Id$ - Author: Krzysztof Grabczewski - - Some proofs on the Hartogs function. -*) - -Goal "\\a. Ord(a) --> a \\ X ==> P"; -by (res_inst_tac [("X1","{y \\ X. Ord(y)}")] (ON_class RS revcut_rl) 1); -by (Fast_tac 1); -qed "Ords_in_set"; - -Goalw [lepoll_def] "[| Ord(a); a lepoll X |] ==> \ -\ \\Y. Y \\ X & (\\R. well_ord(Y,R) & ordertype(Y,R)=a)"; -by (etac exE 1); -by (REPEAT (resolve_tac [exI, conjI] 1)); -by (eresolve_tac [inj_is_fun RS fun_is_rel RS image_subset] 1); -by (rtac exI 1); -by (rtac conjI 1); -by (eresolve_tac [well_ord_Memrel RSN (2, subset_refl RSN (2, - restrict_bij RS bij_converse_bij) RS bij_is_inj RS well_ord_rvimage)] 1 - THEN (assume_tac 1)); -by (resolve_tac [subset_refl RSN (2, restrict_bij RS bij_converse_bij) RS - (well_ord_Memrel RSN (2, bij_ordertype_vimage)) RS - (ordertype_Memrel RSN (2, trans))] 1 - THEN (REPEAT (assume_tac 1))); -qed "Ord_lepoll_imp_ex_well_ord"; - -Goal "[| Ord(a); a lepoll X |] ==> \ -\ \\Y. Y \\ X & (\\R. R \\ X*X & ordertype(Y,R)=a)"; -by (dtac Ord_lepoll_imp_ex_well_ord 1 THEN (assume_tac 1)); -by Safe_tac; -by (REPEAT (ares_tac [exI, conjI] 1)); -by (etac ordertype_Int 2); -by (Fast_tac 1); -qed "Ord_lepoll_imp_eq_ordertype"; - -Goal "\\a. Ord(a) --> a lepoll X ==> \ -\ \\a. Ord(a) --> \ -\ a:{a. Z \\ Pow(X)*Pow(X*X), \\Y R. Z= & ordertype(Y,R)=a}"; -by (REPEAT (resolve_tac [allI,impI] 1)); -by (REPEAT (eresolve_tac [allE, impE] 1)); -by (assume_tac 1); -by (dtac Ord_lepoll_imp_eq_ordertype 1 THEN (assume_tac 1)); -by (fast_tac (claset() addSIs [ReplaceI] addEs [sym]) 1); -qed "Ords_lepoll_set_lemma"; - -Goal "\\a. Ord(a) --> a lepoll X ==> P"; -by (eresolve_tac [Ords_lepoll_set_lemma RS Ords_in_set] 1); -qed "Ords_lepoll_set"; - -Goal "\\a. Ord(a) & ~a lepoll X"; -by (rtac swap 1); -by (Fast_tac 1); -by (rtac Ords_lepoll_set 1); -by (Fast_tac 1); -qed "ex_Ord_not_lepoll"; - -Goalw [Hartog_def] "~ Hartog(A) lepoll A"; -by (resolve_tac [ex_Ord_not_lepoll RS exE] 1); -by (rtac LeastI 1); -by (REPEAT (Fast_tac 1)); -qed "HartogI"; - -val HartogE = HartogI RS notE; - -Goalw [Hartog_def] "Ord(Hartog(A))"; -by (rtac Ord_Least 1); -qed "Ord_Hartog"; - -Goalw [Hartog_def] "[| i < Hartog(A); ~ i lepoll A |] ==> P"; -by (fast_tac (claset() addEs [less_LeastE]) 1); -qed "less_HartogE1"; - -Goal "[| i < Hartog(A); i eqpoll Hartog(A) |] ==> P"; -by (fast_tac (claset() addEs [less_HartogE1, eqpoll_sym RS eqpoll_imp_lepoll - RS lepoll_trans RS HartogE]) 1); -qed "less_HartogE"; - -Goal "Card(Hartog(A))"; -by (fast_tac (claset() addSIs [CardI, Ord_Hartog] addEs [less_HartogE]) 1); -qed "Card_Hartog"; diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/Hartog.thy --- a/src/ZF/AC/Hartog.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/Hartog.thy Wed Jan 16 17:52:06 2002 +0100 @@ -5,14 +5,79 @@ Hartog's function. *) -Hartog = Cardinal + +theory Hartog = AC_Equiv: + +constdefs + Hartog :: "i => i" + "Hartog(X) == LEAST i. ~ i \ X" + +lemma Ords_in_set: "\a. Ord(a) --> a \ X ==> P" +apply (rule_tac X1 = "{y \ X. Ord (y) }" in ON_class [THEN revcut_rl]) +apply fast +done -consts +lemma Ord_lepoll_imp_ex_well_ord: + "[| Ord(a); a \ X |] + ==> \Y. Y \ X & (\R. well_ord(Y,R) & ordertype(Y,R)=a)" +apply (unfold lepoll_def) +apply (erule exE) +apply (intro exI conjI) + apply (erule inj_is_fun [THEN fun_is_rel, THEN image_subset]) + apply (rule well_ord_rvimage [OF bij_is_inj well_ord_Memrel]) + apply (erule restrict_bij [THEN bij_converse_bij]) +apply (rule subset_refl, assumption); +apply (rule trans) +apply (rule bij_ordertype_vimage) +apply (erule restrict_bij [THEN bij_converse_bij]) +apply (rule subset_refl) +apply (erule well_ord_Memrel) +apply (erule ordertype_Memrel) +done + +lemma Ord_lepoll_imp_eq_ordertype: + "[| Ord(a); a \ X |] ==> \Y. Y \ X & (\R. R \ X*X & ordertype(Y,R)=a)" +apply (drule Ord_lepoll_imp_ex_well_ord, (assumption)) +apply clarify +apply (intro exI conjI) +apply (erule_tac [3] ordertype_Int, auto) +done - Hartog :: i => i +lemma Ords_lepoll_set_lemma: + "(\a. Ord(a) --> a \ X) ==> + \a. Ord(a) --> + a \ {b. Z \ Pow(X)*Pow(X*X), \Y R. Z= & ordertype(Y,R)=b}" +apply (intro allI impI) +apply (elim allE impE, assumption) +apply (blast dest!: Ord_lepoll_imp_eq_ordertype intro: sym) +done + +lemma Ords_lepoll_set: "\a. Ord(a) --> a \ X ==> P" +by (erule Ords_lepoll_set_lemma [THEN Ords_in_set]) + +lemma ex_Ord_not_lepoll: "\a. Ord(a) & ~a \ X" +apply (rule ccontr) +apply (best intro: Ords_lepoll_set) +done -defs +lemma not_Hartog_lepoll_self: "~ Hartog(A) \ A" +apply (unfold Hartog_def) +apply (rule ex_Ord_not_lepoll [THEN exE]) +apply (rule LeastI, auto) +done + +lemmas Hartog_lepoll_selfE = not_Hartog_lepoll_self [THEN notE, standard] - Hartog_def "Hartog(X) == LEAST i. ~i lepoll X" +lemma Ord_Hartog: "Ord(Hartog(A))" +by (unfold Hartog_def, rule Ord_Least) + +lemma less_HartogE1: "[| i < Hartog(A); ~ i \ A |] ==> P" +by (unfold Hartog_def, fast elim: less_LeastE) + +lemma less_HartogE: "[| i < Hartog(A); i \ Hartog(A) |] ==> P" +by (blast intro: less_HartogE1 eqpoll_sym eqpoll_imp_lepoll + lepoll_trans [THEN Hartog_lepoll_selfE]); + +lemma Card_Hartog: "Card(Hartog(A))" +by (fast intro!: CardI Ord_Hartog elim: less_HartogE) end diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/ROOT.ML --- a/src/ZF/AC/ROOT.ML Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/ROOT.ML Wed Jan 16 17:52:06 2002 +0100 @@ -6,15 +6,12 @@ Executes the proofs of the AC-equivalences, due to Krzysztof Grabczewski *) -time_use_thy "AC_Equiv"; - time_use_thy "WO6_WO1"; time_use_thy "WO1_WO7"; -time_use "AC7_AC9.ML"; +time_use_thy "AC7_AC9"; time_use_thy "WO1_AC"; -time_use_thy "AC1_WO2"; time_use_thy "AC15_WO6"; diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/WO1_AC.ML --- a/src/ZF/AC/WO1_AC.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,111 +0,0 @@ -(* Title: ZF/AC/WO1_AC.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -The proofs of WO1 ==> AC1 and WO1 ==> AC10(n) for n >= 1 - -The latter proof is referred to as clear by the Rubins. -However it seems to be quite complicated. -The formal proof presented below is a mechanisation of the proof -by Lawrence C. Paulson which is the following: - -Assume WO1. Let s be a set of infinite sets. - -Suppose x \\ s. Then x is equipollent to |x| (by WO1), an infinite cardinal; -call it K. Since K = K |+| K = |K+K| (by InfCard_cdouble_eq) there is an -isomorphism h \\ bij(K+K, x). (Here + means disjoint sum.) - -So there is a partition of x into 2-element sets, namely - - {{h(Inl(i)), h(Inr(i))} . i \\ K} - -So for all x \\ s the desired partition exists. By AC1 (which follows from WO1) -there exists a function f that chooses a partition for each x \\ s. Therefore we -have AC10(2). - -*) - -(* ********************************************************************** *) -(* WO1 ==> AC1 *) -(* ********************************************************************** *) - -Goalw [AC1_def, WO1_def] "WO1 ==> AC1"; -by (fast_tac (claset() addSEs [ex_choice_fun]) 1); -qed "WO1_AC1"; - -(* ********************************************************************** *) -(* WO1 ==> AC10(n) (n >= 1) *) -(* ********************************************************************** *) - -Goalw [WO1_def] "[| WO1; \\B \\ A. \\C \\ D(B). P(C,B) |] \ -\ ==> \\f. \\B \\ A. P(f`B,B)"; -by (eres_inst_tac [("x","Union({{C \\ D(B). P(C,B)}. B \\ A})")] allE 1); -by (etac exE 1); -by (dtac ex_choice_fun 1); -by (Fast_tac 1); -by (etac exE 1); -by (res_inst_tac [("x","\\x \\ A. f`{C \\ D(x). P(C,x)}")] exI 1); -by (Asm_full_simp_tac 1); -by (blast_tac (claset() addSDs [RepFunI RSN (2, apply_type)]) 1); -val lemma1 = result(); - -Goalw [WO1_def] "[| ~Finite(B); WO1 |] ==> |B| + |B| eqpoll B"; -by (rtac eqpoll_trans 1); -by (fast_tac (claset() addSEs [well_ord_cardinal_eqpoll]) 2); -by (resolve_tac [eqpoll_sym RS eqpoll_trans] 1); -by (fast_tac (claset() addSEs [well_ord_cardinal_eqpoll]) 1); -by (fold_tac [cadd_def]); -by (resolve_tac [Card_cardinal RSN (2, Inf_Card_is_InfCard) RS - InfCard_cdouble_eq RS ssubst] 1); -by (rtac eqpoll_refl 2); -by (rtac notI 1); -by (etac notE 1); -by (resolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_Finite] 1 - THEN (assume_tac 2)); -by (fast_tac (claset() addSEs [well_ord_cardinal_eqpoll]) 1); -val lemma2_1 = result(); - -Goal "f \\ bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i \\ D} \\ Pow(Pow(B))"; -by (fast_tac (claset() addSEs [bij_is_fun RS apply_type]) 1); -val lemma2_2 = result(); - -Goal "[| f \\ inj(A,B); f`a = f`b; a \\ A; b \\ A |] ==> a=b"; -by (rtac inj_equality 1); -by (TRYALL (fast_tac (claset() addSEs [inj_is_fun RS apply_Pair] addEs [subst]))); -val lemma = result(); - -Goalw AC_aux_defs - "f \\ bij(D+D, B) ==> pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i \\ D})"; -by (blast_tac (claset() addDs [bij_is_inj RS lemma]) 1); -val lemma2_3 = result(); - -val [major, minor] = goalw thy AC_aux_defs - "[| f \\ bij(D+D, B); 1 le n |] ==> \ -\ sets_of_size_between({{f`Inl(i), f`Inr(i)}. i \\ D}, 2, succ(n))"; -by (rewtac succ_def); -by (fast_tac (claset() - addSIs [cons_lepoll_cong, minor, lepoll_refl] - addIs [singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans, - le_imp_subset RS subset_imp_lepoll] - addDs [major RS bij_is_inj RS lemma] - addSEs [mem_irrefl]) 1); -val lemma2_4 = result(); - -Goalw [bij_def, surj_def] - "f \\ bij(D+D, B) ==> Union({{f`Inl(i), f`Inr(i)}. i \\ D})=B"; -by (fast_tac (claset() addSEs [inj_is_fun RS apply_type]) 1); -val lemma2_5 = result(); - -Goal "[| WO1; ~Finite(B); 1 le n |] \ -\ ==> \\C \\ Pow(Pow(B)). pairwise_disjoint(C) & \ -\ sets_of_size_between(C, 2, succ(n)) & \ -\ Union(C)=B"; -by (eresolve_tac [lemma2_1 RS (eqpoll_def RS def_imp_iff RS iffD1 RS exE)] 1 - THEN (assume_tac 1)); -by (fast_tac (FOL_cs addSIs [bexI] - addSEs [lemma2_2, lemma2_3, lemma2_4, lemma2_5]) 1); -val lemma2 = result(); - -Goalw AC_defs "[| WO1; 1 le n |] ==> AC10(n)"; -by (fast_tac (claset() addSIs [lemma1] addSEs [lemma2]) 1); -qed "WO1_AC10"; diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/WO1_AC.thy --- a/src/ZF/AC/WO1_AC.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/WO1_AC.thy Wed Jan 16 17:52:06 2002 +0100 @@ -1,3 +1,105 @@ -(*Dummy theory to document dependencies *) +(* Title: ZF/AC/WO1_AC.thy + ID: $Id$ + Author: Krzysztof Grabczewski + +The proofs of WO1 ==> AC1 and WO1 ==> AC10(n) for n >= 1 + +The latter proof is referred to as clear by the Rubins. +However it seems to be quite complicated. +The formal proof presented below is a mechanisation of the proof +by Lawrence C. Paulson which is the following: + +Assume WO1. Let s be a set of infinite sets. + +Suppose x \ s. Then x is equipollent to |x| (by WO1), an infinite cardinal +call it K. Since K = K |+| K = |K+K| (by InfCard_cdouble_eq) there is an +isomorphism h \ bij(K+K, x). (Here + means disjoint sum.) + +So there is a partition of x into 2-element sets, namely + + {{h(Inl(i)), h(Inr(i))} . i \ K} + +So for all x \ s the desired partition exists. By AC1 (which follows from WO1) +there exists a function f that chooses a partition for each x \ s. Therefore we +have AC10(2). + +*) + +theory WO1_AC = AC_Equiv: + +(* ********************************************************************** *) +(* WO1 ==> AC1 *) +(* ********************************************************************** *) + +theorem WO1_AC1: "WO1 ==> AC1" +by (unfold AC1_def WO1_def, fast elim!: ex_choice_fun) + +(* ********************************************************************** *) +(* WO1 ==> AC10(n) (n >= 1) *) +(* ********************************************************************** *) + +lemma lemma1: "[| WO1; \B \ A. \C \ D(B). P(C,B) |] ==> \f. \B \ A. P(f`B,B)" +apply (unfold WO1_def) +apply (erule_tac x = "Union ({{C \ D (B) . P (C,B) }. B \ A}) " in allE) +apply (erule exE, drule ex_choice_fun, fast) +apply (erule exE) +apply (rule_tac x = "\x \ A. f`{C \ D (x) . P (C,x) }" in exI) +apply (simp, blast dest!: apply_type [OF _ RepFunI]) +done -WO1_AC = AC_Equiv + WO_AC +lemma lemma2_1: "[| ~Finite(B); WO1 |] ==> |B| + |B| \ B" +apply (unfold WO1_def) +apply (rule eqpoll_trans) +prefer 2 apply (fast elim!: well_ord_cardinal_eqpoll) +apply (rule eqpoll_sym [THEN eqpoll_trans]) +apply (fast elim!: well_ord_cardinal_eqpoll) +apply (drule spec [of _ B]) +apply (clarify dest!: eqpoll_imp_Finite_iff [OF well_ord_cardinal_eqpoll]) +apply (simp add: cadd_def [symmetric] + eqpoll_refl InfCard_cdouble_eq Card_cardinal Inf_Card_is_InfCard) +done + +lemma lemma2_2: + "f \ bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i \ D} \ Pow(Pow(B))" +by (fast elim!: bij_is_fun [THEN apply_type]) + + +lemma lemma2_3: + "f \ bij(D+D, B) ==> pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i \ D})" +apply (unfold pairwise_disjoint_def) +apply (blast dest: bij_is_inj [THEN inj_apply_equality]) +done + +lemma lemma2_4: + "[| f \ bij(D+D, B); 1\n |] + ==> sets_of_size_between({{f`Inl(i), f`Inr(i)}. i \ D}, 2, succ(n))" +apply (simp (no_asm_simp) add: sets_of_size_between_def succ_def) +apply (blast intro!: cons_lepoll_cong + intro: singleton_eqpoll_1 [THEN eqpoll_imp_lepoll] + le_imp_subset [THEN subset_imp_lepoll] lepoll_trans + dest: bij_is_inj [THEN inj_apply_equality] elim!: mem_irrefl) +done + +lemma lemma2_5: + "f \ bij(D+D, B) ==> Union({{f`Inl(i), f`Inr(i)}. i \ D})=B" +apply (unfold bij_def surj_def) +apply (fast elim!: inj_is_fun [THEN apply_type]) +done + +lemma lemma2: + "[| WO1; ~Finite(B); 1\n |] + ==> \C \ Pow(Pow(B)). pairwise_disjoint(C) & + sets_of_size_between(C, 2, succ(n)) & + Union(C)=B" +apply (drule lemma2_1 [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], + assumption) +apply (blast intro!: lemma2_2 lemma2_3 lemma2_4 lemma2_5) +done + +theorem WO1_AC10: "[| WO1; 1\n |] ==> AC10(n)" +apply (unfold AC10_def) +apply (fast intro!: lemma1 elim!: lemma2) +done + +end + diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/WO1_WO7.ML --- a/src/ZF/AC/WO1_WO7.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,115 +0,0 @@ -(* Title: ZF/AC/WO1_WO7.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -WO7 <-> LEMMA <-> WO1 (Rubin & Rubin p. 5) -LEMMA is the sentence denoted by (**) - -Also, WO1 <-> WO8 -*) - -(* ********************************************************************** *) -(* It is easy to see, that WO7 is equivallent to (**) *) -(* ********************************************************************** *) - -Goalw [WO7_def, LEMMA_def] - "WO7 <-> LEMMA"; -by (fast_tac (claset() addSEs [Finite_well_ord_converse]) 1); -qed "WO7_iff_LEMMA"; - -(* ********************************************************************** *) -(* It is also easy to show that LEMMA implies WO1. *) -(* ********************************************************************** *) - -Goalw [WO1_def, LEMMA_def] "LEMMA ==> WO1"; -by (rtac allI 1); -by (etac allE 1); -by (excluded_middle_tac "Finite(A)" 1); -by (Fast_tac 1); -by (rewrite_goals_tac [Finite_def, eqpoll_def]); -by (fast_tac (claset() addSIs [[bij_is_inj, nat_implies_well_ord] MRS - well_ord_rvimage]) 1); -qed "LEMMA_imp_WO1"; - -(* ********************************************************************** *) -(* The Rubins' proof of the other implication is contained within the *) -(* following sentence \\ *) -(* "... each infinite ordinal is well ordered by < but not by >." *) -(* This statement can be proved by the following two theorems. *) -(* But moreover we need to show similar property for any well ordered *) -(* infinite set. It is not very difficult thanks to Isabelle order types *) -(* We show that if a set is well ordered by some relation and by its *) -(* converse, then apropriate order type is well ordered by the converse *) -(* of it's membership relation, which in connection with the previous *) -(* gives the conclusion. *) -(* ********************************************************************** *) - -Goalw [wf_on_def, wf_def] - "[| Ord(a); ~Finite(a) |] ==> ~wf[a](converse(Memrel(a)))"; -by (dresolve_tac [nat_le_infinite_Ord RS le_imp_subset] 1 - THEN (assume_tac 1)); -by (rtac notI 1); -by (eres_inst_tac [("x","nat")] allE 1); -by (Blast_tac 1); -qed "converse_Memrel_not_wf_on"; - -Goalw [well_ord_def] - "[| Ord(a); ~Finite(a) |] ==> ~well_ord(a,converse(Memrel(a)))"; -by (fast_tac (claset() addSDs [converse_Memrel_not_wf_on]) 1); -qed "converse_Memrel_not_well_ord"; - -Goal "[| well_ord(A,r); well_ord(A,converse(r)) |] \ -\ ==> well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r))))"; -by (rtac ([ordertype_ord_iso RS ord_iso_sym RS ord_iso_rvimage_eq, - Memrel_type RS (subset_Int_iff RS iffD1)] - MRS trans RS subst) 1 - THEN (assume_tac 1)); -by (rtac (rvimage_converse RS subst) 1); -by (etac (ordertype_ord_iso RS ord_iso_sym RS ord_iso_is_bij RS - bij_is_inj RS well_ord_rvimage) 1 - THEN (assume_tac 1)); -qed "well_ord_converse_Memrel"; - -Goalw [WO1_def, LEMMA_def] "WO1 ==> LEMMA"; -by (REPEAT (resolve_tac [allI,impI] 1)); -by (REPEAT (eresolve_tac [allE,exE] 1)); -by (REPEAT (ares_tac [exI,conjI,notI] 1)); -by (ftac well_ord_converse_Memrel 1 THEN (assume_tac 1)); -by (forward_tac [Ord_ordertype RS converse_Memrel_not_well_ord] 1); -by (contr_tac 2); -by (fast_tac (empty_cs addSEs [ordertype_ord_iso RS ord_iso_is_bij RS - bij_is_inj RS (exI RS (lepoll_def RS def_imp_iff RS iffD2)) - RS lepoll_Finite] - addSIs [notI] addEs [notE]) 1); -qed "WO1_imp_LEMMA"; - - -Goal "WO1 <-> WO7"; -by (simp_tac (simpset() addsimps [WO7_iff_LEMMA]) 1); -by (blast_tac (claset() addIs [LEMMA_imp_WO1, WO1_imp_LEMMA]) 1); -qed "WO1_iff_WO7"; - - - -(* ********************************************************************** *) - -(* The proof of WO8 <-> WO1 (Rubin & Rubin p. 6) *) - -(* ********************************************************************** *) - -Goalw WO_defs "WO1 ==> WO8"; -by (Fast_tac 1); -qed "WO1_WO8"; - - -(* The proof of "WO8 ==> WO1" - faithful image of Rubin & Rubin's proof *) -Goalw WO_defs "WO8 ==> WO1"; -by (rtac allI 1); -by (eres_inst_tac [("x","{{x}. x \\ A}")] allE 1); -by (etac impE 1); -by (fast_tac (claset() addSEs [lam_sing_bij RS bij_is_inj RS - well_ord_rvimage]) 2); -by (res_inst_tac [("x","\\a \\ {{x}. x \\ A}. THE x. a={x}")] exI 1); -by (force_tac (claset() addSIs [lam_type], - simpset() addsimps [singleton_eq_iff, the_equality]) 1); -qed "WO8_WO1"; diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/WO1_WO7.thy --- a/src/ZF/AC/WO1_WO7.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/WO1_WO7.thy Wed Jan 16 17:52:06 2002 +0100 @@ -5,13 +5,110 @@ WO7 <-> LEMMA <-> WO1 (Rubin & Rubin p. 5) LEMMA is the sentence denoted by (**) + +Also, WO1 <-> WO8 *) -WO1_WO7 = AC_Equiv + +theory WO1_WO7 = AC_Equiv: constdefs LEMMA :: o "LEMMA == - \\X. ~Finite(X) --> (\\R. well_ord(X,R) & ~well_ord(X,converse(R)))" + \X. ~Finite(X) --> (\R. well_ord(X,R) & ~well_ord(X,converse(R)))" + +(* ********************************************************************** *) +(* It is easy to see that WO7 is equivalent to (**) *) +(* ********************************************************************** *) + +lemma WO7_iff_LEMMA: "WO7 <-> LEMMA" +apply (unfold WO7_def LEMMA_def) +apply (blast intro: Finite_well_ord_converse) +done + +(* ********************************************************************** *) +(* It is also easy to show that LEMMA implies WO1. *) +(* ********************************************************************** *) + +lemma LEMMA_imp_WO1: "LEMMA ==> WO1" +apply (unfold WO1_def LEMMA_def Finite_def eqpoll_def) +apply (blast intro!: well_ord_rvimage [OF bij_is_inj nat_implies_well_ord]) +done + +(* ********************************************************************** *) +(* The Rubins' proof of the other implication is contained within the *) +(* following sentence \ *) +(* "... each infinite ordinal is well ordered by < but not by >." *) +(* This statement can be proved by the following two theorems. *) +(* But moreover we need to show similar property for any well ordered *) +(* infinite set. It is not very difficult thanks to Isabelle order types *) +(* We show that if a set is well ordered by some relation and by its *) +(* converse, then apropriate order type is well ordered by the converse *) +(* of it's membership relation, which in connection with the previous *) +(* gives the conclusion. *) +(* ********************************************************************** *) + +lemma converse_Memrel_not_wf_on: + "[| Ord(a); ~Finite(a) |] ==> ~wf[a](converse(Memrel(a)))" +apply (unfold wf_on_def wf_def) +apply (drule nat_le_infinite_Ord [THEN le_imp_subset], (assumption)) +apply (rule notI) +apply (erule_tac x = "nat" in allE, blast) +done + +lemma converse_Memrel_not_well_ord: + "[| Ord(a); ~Finite(a) |] ==> ~well_ord(a,converse(Memrel(a)))" +apply (unfold well_ord_def) +apply (blast dest: converse_Memrel_not_wf_on) +done + +lemma well_ord_rvimage_ordertype: + "well_ord(A,r) \ + rvimage (ordertype(A,r), converse(ordermap(A,r)),r) = + Memrel(ordertype(A,r))" +by (blast intro: ordertype_ord_iso [THEN ord_iso_sym] ord_iso_rvimage_eq + Memrel_type [THEN subset_Int_iff [THEN iffD1]] trans) + +lemma well_ord_converse_Memrel: + "[| well_ord(A,r); well_ord(A,converse(r)) |] + ==> well_ord(ordertype(A,r), converse(Memrel(ordertype(A,r))))" +apply (subst well_ord_rvimage_ordertype [symmetric], assumption) +apply (rule rvimage_converse [THEN subst]) +apply (blast intro: ordertype_ord_iso ord_iso_sym ord_iso_is_bij + bij_is_inj well_ord_rvimage) +done + +lemma WO1_imp_LEMMA: "WO1 ==> LEMMA" +apply (unfold WO1_def LEMMA_def, clarify) +apply (blast dest: well_ord_converse_Memrel + Ord_ordertype [THEN converse_Memrel_not_well_ord] + intro: ordertype_ord_iso ord_iso_is_bij bij_is_inj lepoll_Finite + lepoll_def [THEN def_imp_iff, THEN iffD2] ) +done + +lemma WO1_iff_WO7: "WO1 <-> WO7" +apply (simp add: WO7_iff_LEMMA) +apply (blast intro: LEMMA_imp_WO1 WO1_imp_LEMMA) +done + + + +(* ********************************************************************** *) +(* The proof of WO8 <-> WO1 (Rubin & Rubin p. 6) *) +(* ********************************************************************** *) + +lemma WO1_WO8: "WO1 ==> WO8" +by (unfold WO1_def WO8_def, fast) + + +(* The implication "WO8 ==> WO1": a faithful image of Rubin & Rubin's proof*) +lemma WO8_WO1: "WO8 ==> WO1" +apply (unfold WO1_def WO8_def) +apply (rule allI) +apply (erule_tac x = "{{x}. x \ A}" in allE) +apply (erule impE) + apply (rule_tac x = "\a \ {{x}. x \ A}. THE x. a={x}" in exI) + apply (force intro!: lam_type simp add: singleton_eq_iff the_equality) +apply (blast intro: lam_sing_bij bij_is_inj well_ord_rvimage) +done end diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/WO2_AC16.ML --- a/src/ZF/AC/WO2_AC16.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,591 +0,0 @@ -(* Title: ZF/AC/WO2_AC16.ML - ID: $Id$ - Author: Krzysztof Grabczewski - - The proof of WO2 ==> AC16(k #+ m, k) - - The main part of the proof is the inductive reasoning concerning - properties of constructed family T_gamma. - The proof deals with three cases for ordinals: 0, succ and limit ordinal. - The first instance is trivial, the third not difficult, but the second - is very complicated requiring many lemmas. - We also need to prove that at any stage gamma the set - (s - Union(...) - k_gamma) (Rubin & Rubin page 15) - contains m distinct elements (in fact is equipollent to s) -*) - -(* ********************************************************************** *) -(* case of limit ordinal *) -(* ********************************************************************** *) - - -Goal "[| \\yzY \\ F(y). f(z)<=Y) \ -\ --> (\\! Y. Y \\ F(y) & f(z)<=Y); \ -\ \\i j. i le j --> F(i) \\ F(j); j le i; i F(i); f(z)<=V; W \\ F(j); f(z)<=W |] \ -\ ==> V = W"; -by (REPEAT (eresolve_tac [asm_rl, allE, impE] 1)); -by (dtac subsetD 1 THEN (assume_tac 1)); -by (REPEAT (dtac ospec 1 THEN (assume_tac 1))); -by (eresolve_tac [disjI2 RSN (2, impE)] 1); -by (fast_tac (FOL_cs addSIs [bexI]) 1); -by (etac ex1_two_eq 1 THEN (REPEAT (ares_tac [conjI] 1))); -val lemma3_1 = result(); - - -Goal "[| \\yzY \\ F(y). f(z)<=Y) \ -\ --> (\\! Y. Y \\ F(y) & f(z)<=Y); \ -\ \\i j. i le j --> F(i) \\ F(j); i F(i); f(z)<=V; W \\ F(j); f(z)<=W |] \ -\ ==> V = W"; -by (res_inst_tac [("j","j")] ([lt_Ord, lt_Ord] MRS Ord_linear_le) 1 - THEN (REPEAT (assume_tac 1))); -by (eresolve_tac [lemma3_1 RS sym] 1 THEN (REPEAT (assume_tac 1))); -by (etac lemma3_1 1 THEN (REPEAT (assume_tac 1))); -val lemma3 = result(); - - -Goal "[| \\y X & \ -\ (\\xY \\ F(y). fa(x) \\ Y) --> \ -\ (\\! Y. Y \\ F(y) & fa(x) \\ Y)); x < a |] \ -\ ==> \\yzY \\ F(y). fa(z) \\ Y) --> \ -\ (\\! Y. Y \\ F(y) & fa(z) \\ Y)"; -by (REPEAT (resolve_tac [oallI, impI] 1)); -by (dtac ospec 1 THEN (assume_tac 1)); -by (fast_tac (FOL_cs addSEs [oallE]) 1); -val lemma4 = result(); - - -Goal "[| \\y X & \ -\ (\\xY \\ F(y). fa(x) \\ Y) --> \ -\ (\\! Y. Y \\ F(y) & fa(x) \\ Y)); \ -\ x < a; Limit(x); \\i j. i le j --> F(i) \\ F(j) |] \ -\ ==> (\\x X & \ -\ (\\xax \\ \\x x) \ -\ --> (\\! Y. Y \\ (\\x Y))"; -by (rtac conjI 1); -by (rtac subsetI 1); -by (etac OUN_E 1); -by (dtac ospec 1 THEN (assume_tac 1)); -by (Fast_tac 1); -by (dtac lemma4 1 THEN (assume_tac 1)); -by (rtac oallI 1); -by (rtac impI 1); -by (etac disjE 1); -by (forward_tac [Limit_has_succ RSN (2, ospec)] 1 THEN (REPEAT (assume_tac 1))); -by (dres_inst_tac [("A","a"),("x","xa")] ospec 1 THEN (assume_tac 1)); -by (eresolve_tac [lt_Ord RS le_refl RSN (2, disjI1 RSN (2, impE))] 1 - THEN (assume_tac 1)); -by (REPEAT (eresolve_tac [ex1E, conjE] 1)); -by (rtac ex1I 1); -by (rtac conjI 1 THEN (assume_tac 2)); -by (eresolve_tac [Limit_has_succ RS OUN_I] 1 THEN (TRYALL assume_tac)); -by (REPEAT (eresolve_tac [conjE, OUN_E] 1)); -by (etac lemma3 1 THEN (TRYALL assume_tac)); -by (etac Limit_has_succ 1 THEN (assume_tac 1)); -by (etac bexE 1); -by (rtac ex1I 1); -by (etac conjI 1 THEN (assume_tac 1)); -by (REPEAT (eresolve_tac [conjE, OUN_E] 1)); -by (etac lemma3 1 THEN (TRYALL assume_tac)); -val lemma5 = result(); - -(* ********************************************************************** *) -(* case of successor ordinal *) -(* ********************************************************************** *) - -(* - First quite complicated proof of the fact used in the recursive construction - of the family T_gamma (WO2 ==> AC16(k #+ m, k)) - the fact that at any stage - gamma the set (s - Union(...) - k_gamma) is equipollent to s - (Rubin & Rubin page 15). -*) - -(* ********************************************************************** *) -(* dbl_Diff_eqpoll_Card *) -(* ********************************************************************** *) - - -Goal "[| A eqpoll a; Card(a); ~Finite(a); B lesspoll a; \ -\ C lesspoll a |] ==> A - B - C eqpoll a"; -by (rtac Diff_lesspoll_eqpoll_Card 1 THEN (REPEAT (assume_tac 1))); -by (rtac Diff_lesspoll_eqpoll_Card 1 THEN (REPEAT (assume_tac 1))); -qed "dbl_Diff_eqpoll_Card"; - -(* ********************************************************************** *) -(* Case of finite ordinals *) -(* ********************************************************************** *) - - -Goalw [lesspoll_def] - "[| Finite(X); ~Finite(a); Ord(a) |] ==> X lesspoll a"; -by (rtac conjI 1); -by (dresolve_tac [nat_le_infinite_Ord RS le_imp_lepoll] 1 - THEN (assume_tac 1)); -by (rewtac Finite_def); -by (fast_tac (claset() addSEs [eqpoll_sym RS eqpoll_trans]) 2); -by (rtac lepoll_trans 1 THEN (assume_tac 2)); -by (fast_tac (claset() addSEs [Ord_nat RSN (2, ltI) RS leI RS le_imp_subset RS - subset_imp_lepoll RSN (2, eqpoll_imp_lepoll RS lepoll_trans)]) 1); -qed "Finite_lesspoll_infinite_Ord"; - -Goal "[| \\x \\ X. x lepoll n & x \\ T; well_ord(T, R); X lepoll b; \ -\ b nat |] \ -\ ==> Union(X) lesspoll a"; -by (excluded_middle_tac "Finite(X)" 1); -by (resolve_tac [Card_is_Ord RSN (3, Finite_lesspoll_infinite_Ord)] 2 - THEN (REPEAT (assume_tac 3))); -by (fast_tac (claset() addSEs [lepoll_nat_imp_Finite] - addSIs [Finite_Union]) 2); -by (dresolve_tac [lt_Ord RSN (2, lepoll_imp_ex_le_eqpoll)] 1 THEN (assume_tac 1)); -by (REPEAT (eresolve_tac [exE, conjE] 1)); -by (forward_tac [eqpoll_imp_lepoll RS lepoll_infinite] 1 THEN (assume_tac 1)); -by (eresolve_tac [eqpoll_sym RS (eqpoll_def RS def_imp_iff RS iffD1) RS - exE] 1); -by (forward_tac [bij_is_surj RS surj_image_eq] 1); -by (dresolve_tac [[bij_is_fun, subset_refl] MRS image_fun] 1); -by (dresolve_tac [sym RS trans] 1 THEN (assume_tac 1)); -by (blast_tac (claset() addIs [lesspoll_trans1, UN_lepoll, lt_Ord, - lt_trans1 RSN (2, lt_Card_imp_lesspoll)]) 1); -qed "Union_lesspoll"; - -(* ********************************************************************** *) -(* recfunAC16_lepoll_index *) -(* ********************************************************************** *) - - -Goal "A Un {a} = cons(a, A)"; -by (Fast_tac 1); -qed "Un_sing_eq_cons"; - - -Goal "A lepoll B ==> A Un {a} lepoll succ(B)"; -by (asm_simp_tac (simpset() addsimps [Un_sing_eq_cons, succ_def]) 1); -by (eresolve_tac [mem_not_refl RSN (2, cons_lepoll_cong)] 1); -qed "Un_lepoll_succ"; - - -Goal "Ord(a) ==> F(a) - (\\b F(a) Un X - (\\b X"; -by (fast_tac (claset() addSIs [OUN_I, le_refl]) 1); -qed "Diff_UN_succ_subset"; - - -Goal "Ord(x) ==> \ -\ recfunAC16(f, g, x, a) - (\\i F(x); Ord(x) |] \ -\ ==> z \\ F(LEAST i. z \\ F(i)) - (\\j<(LEAST i. z \\ F(i)). F(j))"; -by (fast_tac (claset() addEs [less_LeastE] addSEs [OUN_E, LeastI]) 1); -qed "in_Least_Diff"; - - -Goal "[| (LEAST i. w \\ F(i)) = (LEAST i. z \\ F(i)); \ -\ w \\ (\\i (\\i \\b (F(b) - (\\c (F(b) - (\\c A; b \\ A |] ==> a=b"; -by (fast_tac (claset() addSDs [lepoll_1_is_sing]) 1); -qed "two_in_lepoll_1"; - - -Goal "[| \\ij (\\xz \\ (\\x F(i)")] exI 1); -by (rewtac inj_def); -by (rtac CollectI 1); -by (rtac lam_type 1); -by (etac OUN_E 1); -by (etac Least_in_Ord 1); -by (etac ltD 1); -by (etac lt_Ord2 1); -by (rtac ballI 1); -by (rtac ballI 1); -by (Asm_simp_tac 1); -by (rtac impI 1); -by (dtac Least_eq_imp_ex 1 THEN (REPEAT (assume_tac 1))); -by (fast_tac (claset() addSEs [two_in_lepoll_1]) 1); -qed "UN_lepoll_index"; - - -Goal "Ord(y) ==> recfunAC16(f, fa, y, a) lepoll y"; -by (etac trans_induct 1); -by (etac Ord_cases 1); -by (asm_simp_tac (simpset() addsimps [recfunAC16_0, lepoll_refl]) 1); -by (asm_simp_tac (simpset() addsimps [recfunAC16_succ]) 1); -by (fast_tac (claset() - addSDs [succI1 RSN (2, bspec)] - addSEs [subset_succI RS subset_imp_lepoll RSN (2, lepoll_trans), - Un_lepoll_succ]) 1); -by (asm_simp_tac (simpset() addsimps [recfunAC16_Limit]) 1); -by (fast_tac (claset() addSEs [lt_Ord RS recfunAC16_Diff_lepoll_1] - addSIs [UN_lepoll_index]) 1); -qed "recfunAC16_lepoll_index"; - - -Goal "[| recfunAC16(f,g,y,a) \\ {X \\ Pow(A). X eqpoll n}; \ -\ A eqpoll a; y nat |] \ -\ ==> Union(recfunAC16(f,g,y,a)) lesspoll a"; -by (eresolve_tac [eqpoll_def RS def_imp_iff RS iffD1 RS exE] 1); -by (rtac Union_lesspoll 1 THEN (TRYALL assume_tac)); -by (eresolve_tac [lt_Ord RS recfunAC16_lepoll_index] 3); -by (eresolve_tac [[bij_is_inj, Card_is_Ord RS well_ord_Memrel] MRS - well_ord_rvimage] 2 - THEN (assume_tac 2)); -by (fast_tac (claset() addSEs [eqpoll_imp_lepoll]) 1); -qed "Union_recfunAC16_lesspoll"; - - -Goal "[| recfunAC16(f, fa, y, a) \\ {X \\ Pow(A) . X eqpoll succ(k #+ m)}; \ -\ Card(a); ~ Finite(a); A eqpoll a; \ -\ k \\ nat; y bij(a, {Y \\ Pow(A). Y eqpoll succ(k)}) |] \ -\ ==> A - Union(recfunAC16(f, fa, y, a)) - fa`y eqpoll a"; -by (rtac dbl_Diff_eqpoll_Card 1 THEN (TRYALL assume_tac)); -by (rtac Union_recfunAC16_lesspoll 1 THEN (REPEAT (assume_tac 1))); -by (Simp_tac 1); -by (resolve_tac [nat_succI RSN - (2, bexI RS (Finite_def RS def_imp_iff RS iffD2)) RS - (Card_is_Ord RSN (3, Finite_lesspoll_infinite_Ord))] 1 - THEN (TRYALL assume_tac)); -by (eresolve_tac [ltD RSN (2, bij_is_fun RS apply_type) RS CollectE] 1 - THEN (TRYALL assume_tac)); -qed "dbl_Diff_eqpoll"; - -(* back to the proof *) - -val disj_Un_eqpoll_nat_sum = - [disj_Un_eqpoll_sum, sum_eqpoll_cong, nat_sum_eqpoll_sum] MRS - (eqpoll_trans RS eqpoll_trans) |> standard; - - -Goal "[| x \\ Pow(A - B - fa`i); x eqpoll m; \ -\ fa \\ bij(a, {x \\ Pow(A) . x eqpoll k}); i nat; m \\ nat |] \ -\ ==> fa ` i Un x \\ {x \\ Pow(A) . x eqpoll k #+ m}"; -by (rtac CollectI 1); -by (fast_tac (claset() - addSEs [ltD RSN (2, bij_is_fun RS apply_type RS CollectE)]) 1); -by (rtac disj_Un_eqpoll_nat_sum 1 - THEN (TRYALL assume_tac)); -by (fast_tac (claset() addSIs [equals0I]) 1); -by (eresolve_tac [ltD RSN (2, bij_is_fun RS apply_type RS CollectE)] 1 - THEN (REPEAT (assume_tac 1))); -qed "Un_in_Collect"; - -(* ********************************************************************** *) -(* Lemmas simplifying assumptions *) -(* ********************************************************************** *) - - -Goal "[| \\yx Q(x,y)); succ(j) F(j)<=X & (\\x Q(x,j))"; -by (dtac ospec 1); -by (resolve_tac [lt_Ord RS (succI1 RS ltI RS lt_Ord RS le_refl)] 1 - THEN (REPEAT (assume_tac 1))); -val lemma6 = result(); - - -Goal "[| \\x Q(x,j); succ(j) P(j,j) --> (\\x Q(x,j))"; -by (fast_tac (claset() addSEs [leE]) 1); -val lemma7 = result(); - -(* ********************************************************************** *) -(* Lemmas needded to prove ex_next_set which means that for any successor *) -(* ordinal there is a set satisfying certain properties *) -(* ********************************************************************** *) - - -Goal "[| A eqpoll a; ~ Finite(a); Ord(a); m \\ nat |] \ -\ ==> \\X \\ Pow(A). X eqpoll m"; -by (eresolve_tac [Ord_nat RSN (2, ltI) RS - (nat_le_infinite_Ord RSN (2, lt_trans2)) RS - leI RS le_imp_lepoll RS - ((eqpoll_sym RS eqpoll_imp_lepoll) RSN (2, lepoll_trans)) RS - lepoll_imp_eqpoll_subset RS exE] 1 - THEN REPEAT (assume_tac 1)); -by (fast_tac (claset() addSEs [eqpoll_sym]) 1); -qed "ex_subset_eqpoll"; - - -Goal "[| A \\ B Un C; A Int C = 0 |] ==> A \\ B"; -by (Blast_tac 1); -qed "subset_Un_disjoint"; - - -Goal "[| X \\ Pow(A - Union(B) -C); T \\ B; F \\ T |] ==> F Int X = 0"; -by (Blast_tac 1); -qed "Int_empty"; - -(* ********************************************************************** *) -(* equipollent subset (and finite) is the whole set *) -(* ********************************************************************** *) - - -Goal "[| A \\ B; a \\ A; A - {a} = B - {a} |] ==> A = B"; -by (fast_tac (claset() addSEs [equalityE]) 1); -qed "Diffs_eq_imp_eq"; - - -Goal "m \\ nat ==> \\A B. A \\ B & m lepoll A & B lepoll m --> A=B"; -by (induct_tac "m" 1); -by (fast_tac (claset() addSDs [lepoll_0_is_0]) 1); -by (REPEAT (resolve_tac [allI, impI] 1)); -by (REPEAT (etac conjE 1)); -by (resolve_tac [succ_lepoll_imp_not_empty RS not_emptyE] 1 - THEN (assume_tac 1)); -by (forward_tac [subsetD RS Diff_sing_lepoll] 1 - THEN REPEAT (assume_tac 1)); -by (ftac lepoll_Diff_sing 1); -by (REPEAT (eresolve_tac [allE, impE] 1)); -by (rtac conjI 1); -by (Fast_tac 2); -by (Fast_tac 1); -by (etac Diffs_eq_imp_eq 1 - THEN REPEAT (assume_tac 1)); -qed "subset_imp_eq_lemma"; - - -Goal "[| A \\ B; m lepoll A; B lepoll m; m \\ nat |] ==> A=B"; -by (fast_tac (FOL_cs addSDs [subset_imp_eq_lemma]) 1); -qed "subset_imp_eq"; - - -Goal "[| f \\ bij(a, {Y \\ X. Y eqpoll succ(k)}); k \\ nat; f`b \\ f`y; b b=y"; -by (dtac subset_imp_eq 1); -by (etac nat_succI 3); -by (fast_tac (claset() addSEs [bij_is_fun RS (ltD RSN (2, apply_type)) RS - CollectE, eqpoll_sym RS eqpoll_imp_lepoll]) 1); -by (fast_tac (claset() addSEs [bij_is_fun RS (ltD RSN (2, apply_type)) RS - CollectE, eqpoll_imp_lepoll]) 1); -by (rewrite_goals_tac [bij_def, inj_def]); -by (fast_tac (claset() addSDs [ltD]) 1); -qed "bij_imp_arg_eq"; - - -Goal "[| recfunAC16(f, fa, y, a) \\ {X \\ Pow(A) . X eqpoll succ(k #+ m)}; \ -\ Card(a); ~ Finite(a); A eqpoll a; \ -\ k \\ nat; m \\ nat; y bij(a, {Y \\ Pow(A). Y eqpoll succ(k)}); \ -\ ~ (\\Y \\ recfunAC16(f, fa, y, a). fa`y \\ Y) |] \ -\ ==> \\X \\ {Y \\ Pow(A). Y eqpoll succ(k #+ m)}. fa`y \\ X & \ -\ (\\b X --> \ -\ (\\T \\ recfunAC16(f, fa, y, a). ~ fa`b \\ T))"; -by (eresolve_tac [dbl_Diff_eqpoll RS ex_subset_eqpoll RS bexE] 1 - THEN REPEAT (assume_tac 1)); -by (etac Card_is_Ord 1); -by (ftac Un_in_Collect 2 THEN REPEAT (assume_tac 2)); -by (etac CollectE 4); -by (rtac bexI 4); -by (rtac CollectI 5); -by (assume_tac 5); -by (eresolve_tac [add_succ RS subst] 5); -by (assume_tac 1); -by (etac nat_succI 1); -by (assume_tac 1); -by (rtac conjI 1); -by (Fast_tac 1); -by (REPEAT (resolve_tac [ballI, impI, oallI, notI] 1)); -by (dresolve_tac [Int_empty RSN (2, subset_Un_disjoint)] 1 - THEN REPEAT (assume_tac 1)); -by (dtac bij_imp_arg_eq 1 THEN REPEAT (assume_tac 1)); -by (hyp_subst_tac 1); -by (eresolve_tac [bexI RSN (2, notE)] 1 THEN TRYALL assume_tac); -qed "ex_next_set"; - -(* ********************************************************************** *) -(* Lemma ex_next_Ord states that for any successor *) -(* ordinal there is a number of the set satisfying certain properties *) -(* ********************************************************************** *) - - -Goal "[| recfunAC16(f, fa, y, a) \\ {X \\ Pow(A) . X eqpoll succ(k #+ m)}; \ -\ Card(a); ~ Finite(a); A eqpoll a; \ -\ k \\ nat; m \\ nat; y bij(a, {Y \\ Pow(A). Y eqpoll succ(k)}); \ -\ f \\ bij(a, {Y \\ Pow(A). Y eqpoll succ(k #+ m)}); \ -\ ~ (\\Y \\ recfunAC16(f, fa, y, a). fa`y \\ Y) |] \ -\ ==> \\c f`c & \ -\ (\\b f`c --> \ -\ (\\T \\ recfunAC16(f, fa, y, a). ~ fa`b \\ T))"; -by (dtac ex_next_set 1 THEN REPEAT (assume_tac 1)); -by (etac bexE 1); -by (resolve_tac [bij_converse_bij RS bij_is_fun RS apply_type RS ltI RSN - (2, oexI)] 1); -by (resolve_tac [right_inverse_bij RS ssubst] 1 - THEN REPEAT (ares_tac [Card_is_Ord] 1)); -qed "ex_next_Ord"; - - -Goal "[| \\! Y. Y \\ Z & P(Y); ~P(W) |] ==> \\! Y. Y \\ (Z Un {W}) & P(Y)"; -by (Fast_tac 1); -qed "ex1_in_Un_sing"; - -(* ********************************************************************** *) -(* Lemma simplifying assumptions *) -(* ********************************************************************** *) - - -Goal "[| \\xxa \\ F(j). P(x, xa)) \ -\ --> (\\! Y. Y \\ F(j) & P(x, Y)); F(j) \\ X; \ -\ L \\ X; P(j, L) & (\\x (\\xa \\ F(j). ~P(x, xa))) |] \ -\ ==> F(j) Un {L} \\ X & \ -\ (\\xxa \\ (F(j) Un {L}). P(x, xa)) --> \ -\ (\\! Y. Y \\ (F(j) Un {L}) & P(x, Y)))"; -by (rtac conjI 1); -by (fast_tac (claset() addSIs [singleton_subsetI]) 1); -by (rtac oallI 1); -by (etac oallE 1 THEN (contr_tac 2)); -by (blast_tac (claset() addSEs [leE]) 1); -val lemma8 = result(); - -(* ********************************************************************** *) -(* The main part of the proof: inductive proof of the property of T_gamma *) -(* lemma main_induct *) -(* ********************************************************************** *) - - -Goal "[| b < a; f \\ bij(a, {Y \\ Pow(A) . Y eqpoll succ(k #+ m)}); \ -\ fa \\ bij(a, {Y \\ Pow(A) . Y eqpoll succ(k)}); \ -\ ~Finite(a); Card(a); A eqpoll a; k \\ nat; m \\ nat |] \ -\ ==> recfunAC16(f, fa, b, a) \\ {X \\ Pow(A) . X eqpoll succ(k #+ m)} & \ -\ (\\xY \\ recfunAC16(f, fa, b, a). fa ` x \\ Y) --> \ -\ (\\! Y. Y \\ recfunAC16(f, fa, b, a) & fa ` x \\ Y))"; -by (etac lt_induct 1); -by (ftac lt_Ord 1); -by (etac Ord_cases 1); -(* case 0 *) -by (asm_simp_tac (simpset() addsimps [recfunAC16_0]) 1); -(* case Limit *) -by (asm_simp_tac (simpset() addsimps [recfunAC16_Limit]) 2); -by (rtac lemma5 2 THEN (REPEAT (assume_tac 2))); -by (fast_tac (FOL_cs addSEs [recfunAC16_mono]) 2); -(* case succ *) -by (hyp_subst_tac 1); -by (eresolve_tac [lemma6 RS conjE] 1 THEN (assume_tac 1)); -by (asm_simp_tac (simpset() delsplits [split_if] - addsimps [recfunAC16_succ]) 1); -by (resolve_tac [conjI RS (split_if RS iffD2)] 1); -by (Asm_simp_tac 1); -by (etac lemma7 1 THEN assume_tac 1); -by (rtac impI 1); -by (resolve_tac [ex_next_Ord RS oexE] 1 - THEN REPEAT (ares_tac [le_refl RS lt_trans] 1)); -by (etac lemma8 1 THEN (assume_tac 1)); -by (resolve_tac [bij_is_fun RS apply_type] 1 THEN (assume_tac 1)); -by (eresolve_tac [Least_le RS lt_trans2 RS ltD] 1 - THEN REPEAT (ares_tac [lt_Ord, succ_leI] 1)); -by (rtac (lt_Ord RSN (2, LeastI)) 1 THEN REPEAT (assume_tac 1)); -qed "main_induct"; - -(* ********************************************************************** *) -(* Lemma to simplify the inductive proof *) -(* - the desired property is a consequence of the inductive assumption *) -(* ********************************************************************** *) - -val [prem1, prem2, prem3, prem4] = goal thy - "[| (!!b. b F(b) \\ S & (\\xY \\ F(b). f`x \\ Y)) \ -\ --> (\\! Y. Y \\ F(b) & f`x \\ Y))); \ -\ f \\ a->f``(a); Limit(a); (!!i j. i le j ==> F(i) \\ F(j)) |] \ -\ ==> (\\j S & \ -\ (\\x \\ f``a. \\! Y. Y \\ (\\j Y)"; -by (rtac conjI 1); -by (rtac subsetI 1); -by (etac OUN_E 1); -by (dtac prem1 1); -by (Fast_tac 1); -(** LEVEL 5 **) -by (rtac ballI 1); -by (etac imageE 1); -by (dresolve_tac [prem3 RS Limit_is_Ord RSN (2, ltI) RS - (prem3 RS Limit_has_succ)] 1); -by (ftac prem1 1); -by (etac conjE 1); -(** LEVEL 10 **) -by (dresolve_tac [leI RS succ_leE RSN (2, ospec)] 1 THEN (assume_tac 1)); -by (etac impE 1); -by (fast_tac (claset() addSEs [leI RS succ_leE RS lt_Ord RS le_refl]) 1); -by (dresolve_tac [prem2 RSN (2, apply_equality)] 1); -by (REPEAT (eresolve_tac [conjE, ex1E] 1)); -(** LEVEL 15 **) -by (rtac ex1I 1); -by (fast_tac (claset() addSIs [OUN_I]) 1); -by (REPEAT (eresolve_tac [conjE, OUN_E] 1)); -by (eresolve_tac [lt_Ord RSN (2, lt_Ord RS Ord_linear_le)] 1 - THEN assume_tac 1); -by (dresolve_tac [prem4 RS subsetD] 2 THEN (assume_tac 2)); -(** LEVEL 20 **) -by (fast_tac FOL_cs 2); -by (ftac prem1 1); -by (ftac succ_leE 1); -by (dresolve_tac [prem4 RS subsetD] 1 THEN (assume_tac 1)); -by (etac conjE 1); -(** LEVEL 25 **) -by (dresolve_tac [lt_trans RSN (2, ospec)] 1 THEN (TRYALL assume_tac)); -by (dresolve_tac [disjI1 RSN (2, mp)] 1 THEN (assume_tac 1)); -by (etac ex1_two_eq 1); -by (REPEAT (Fast_tac 1)); -qed "lemma_simp_induct"; - -(* ********************************************************************** *) -(* The target theorem *) -(* ********************************************************************** *) - - -Goalw [AC16_def] "[| WO2; 0 nat; m \\ nat |] ==> AC16(k #+ m,k)"; -by (rtac allI 1); -by (rtac impI 1); -by (ftac WO2_infinite_subsets_eqpoll_X 1 - THEN (REPEAT (assume_tac 1))); -by (forw_inst_tac [("n","k #+ m")] (WO2_infinite_subsets_eqpoll_X) 1 - THEN (REPEAT (ares_tac [add_type] 1))); -by (ftac WO2_imp_ex_Card 1); -by (REPEAT (eresolve_tac [exE,conjE] 1)); -by (dresolve_tac [eqpoll_trans RS eqpoll_sym RS (eqpoll_def RS - def_imp_iff RS iffD1)] 1 THEN (assume_tac 1)); -by (dresolve_tac [eqpoll_trans RS eqpoll_sym RS (eqpoll_def RS - def_imp_iff RS iffD1)] 1 THEN (assume_tac 1)); -by (REPEAT (etac exE 1)); -by (res_inst_tac [("x","\\jx \\ z. ?Z(x))")] - (bij_is_surj RS surj_image_eq RS subst) 1 - THEN (assume_tac 1)); -by (rtac lemma_simp_induct 1); -by (eresolve_tac [bij_is_fun RS surj_image RS surj_is_fun] 2); -by (eresolve_tac [eqpoll_imp_lepoll RS lepoll_infinite RS - infinite_Card_is_InfCard RS InfCard_is_Limit] 2 - THEN REPEAT (assume_tac 2)); -by (etac recfunAC16_mono 2); -by (rtac main_induct 1 - THEN REPEAT (ares_tac [eqpoll_imp_lepoll RS lepoll_infinite] 1)); -qed "WO2_AC16"; diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/WO2_AC16.thy --- a/src/ZF/AC/WO2_AC16.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/WO2_AC16.thy Wed Jan 16 17:52:06 2002 +0100 @@ -1,3 +1,588 @@ -(*Dummy theory to document dependencies *) +(* Title: ZF/AC/WO2_AC16.thy + ID: $Id$ + Author: Krzysztof Grabczewski + + The proof of WO2 ==> AC16(k #+ m, k) + + The main part of the proof is the inductive reasoning concerning + properties of constructed family T_gamma. + The proof deals with three cases for ordinals: 0, succ and limit ordinal. + The first instance is trivial, the third not difficult, but the second + is very complicated requiring many lemmas. + We also need to prove that at any stage gamma the set + (s - Union(...) - k_gamma) (Rubin & Rubin page 15) + contains m distinct elements (in fact is equipollent to s) +*) + +theory WO2_AC16 = AC_Equiv + AC16_lemmas + Cardinal_aux: + +(**** A recursive definition used in the proof of WO2 ==> AC16 ****) + +constdefs + recfunAC16 :: "[i,i,i,i] => i" + "recfunAC16(f,h,i,a) == + transrec2(i, 0, + %g r. if (\y \ r. h`g \ y) then r + else r Un {f`(LEAST i. h`g \ f`i & + (\b f`i --> (\t \ r. ~ h`b \ t))))})" + +(* ********************************************************************** *) +(* Basic properties of recfunAC16 *) +(* ********************************************************************** *) + +lemma recfunAC16_0: "recfunAC16(f,h,0,a) = 0" +apply (simp add: recfunAC16_def); +done + +lemma recfunAC16_succ: + "recfunAC16(f,h,succ(i),a) = + (if (\y \ recfunAC16(f,h,i,a). h ` i \ y) then recfunAC16(f,h,i,a) + else recfunAC16(f,h,i,a) Un + {f ` (LEAST j. h ` i \ f ` j & + (\b f`j + --> (\t \ recfunAC16(f,h,i,a). ~ h`b \ t))))})" +apply (simp add: recfunAC16_def); +done + +lemma recfunAC16_Limit: "Limit(i) + ==> recfunAC16(f,h,i,a) = (\j B(g,r); Ord(i) |] + ==> j transrec2(j, 0, B) \ transrec2(i, 0, B)" +apply (erule trans_induct); +apply (rule Ord_cases , assumption+); +apply fast +apply (simp (no_asm_simp)) +apply (blast elim!: leE); +apply (simp add: transrec2_Limit); +apply (blast intro: OUN_I ltI Ord_in_Ord [THEN le_refl] + elim!: Limit_has_succ [THEN ltE]) +done + +lemma transrec2_mono: + "[| !!g r. r \ B(g,r); j\i |] + ==> transrec2(j, 0, B) \ transrec2(i, 0, B)" +apply (erule leE); +apply (rule transrec2_mono_lemma) +apply (auto intro: lt_Ord2 ); +done + +(* ********************************************************************** *) +(* Monotonicity of recfunAC16 *) +(* ********************************************************************** *) + +lemma recfunAC16_mono: + "i\j ==> recfunAC16(f, g, i, a) \ recfunAC16(f, g, j, a)" +apply (unfold recfunAC16_def) +apply (rule transrec2_mono) +apply (auto ); +done + +(* ********************************************************************** *) +(* case of limit ordinal *) +(* ********************************************************************** *) + + +lemma lemma3_1: + "[| \yzY \ F(y). f(z)<=Y) --> (\! Y. Y \ F(y) & f(z)<=Y); + \i j. i\j --> F(i) \ F(j); j\i; i F(i); f(z)<=V; W \ F(j); f(z)<=W |] + ==> V = W" +apply (erule asm_rl allE impE)+ +apply (drule subsetD, assumption) +apply blast +done + + +lemma lemma3: + "[| \yzY \ F(y). f(z)<=Y) --> (\! Y. Y \ F(y) & f(z)<=Y); + \i j. i\j --> F(i) \ F(j); i F(i); f(z)<=V; W \ F(j); f(z)<=W |] + ==> V = W" +apply (rule_tac j=j in Ord_linear_le [OF lt_Ord lt_Ord], assumption+) +apply (erule lemma3_1 [symmetric], assumption+) +apply (erule lemma3_1, assumption+) +done + +lemma lemma4: + "[| \y X & + (\xY \ F(y). h(x) \ Y) --> + (\! Y. Y \ F(y) & h(x) \ Y)); + x < a |] + ==> \yzY \ F(y). h(z) \ Y) --> + (\! Y. Y \ F(y) & h(z) \ Y)" +apply (intro oallI impI) +apply (drule ospec, assumption) +apply clarify +apply (blast elim!: oallE ) +done + +lemma lemma5: + "[| \y X & + (\xY \ F(y). h(x) \ Y) --> + (\! Y. Y \ F(y) & h(x) \ Y)); + x < a; Limit(x); \i j. i\j --> F(i) \ F(j) |] + ==> (\x X & + (\xax \ \x x) + --> (\! Y. Y \ (\x Y))" +apply (rule conjI) +apply (rule subsetI) +apply (erule OUN_E) +apply (drule ospec, assumption) +apply fast +apply (drule lemma4, assumption) +apply (rule oallI) +apply (rule impI) +apply (erule disjE) +apply (frule ospec, erule Limit_has_succ, assumption) +apply (drule_tac A = "a" and x = "xa" in ospec, assumption) +apply (erule impE, rule le_refl [THEN disjI1], erule lt_Ord) +apply (blast intro: lemma3 Limit_has_succ) +apply (blast intro: lemma3) +done + +(* ********************************************************************** *) +(* case of successor ordinal *) +(* ********************************************************************** *) + +(* + First quite complicated proof of the fact used in the recursive construction + of the family T_gamma (WO2 ==> AC16(k #+ m, k)) - the fact that at any stage + gamma the set (s - Union(...) - k_gamma) is equipollent to s + (Rubin & Rubin page 15). +*) + +(* ********************************************************************** *) +(* dbl_Diff_eqpoll_Card *) +(* ********************************************************************** *) + + +lemma dbl_Diff_eqpoll_Card: + "[| A\a; Card(a); ~Finite(a); B\a; C\a |] ==> A - B - C\a" +by (blast intro: Diff_lesspoll_eqpoll_Card) + +(* ********************************************************************** *) +(* Case of finite ordinals *) +(* ********************************************************************** *) + + +lemma Finite_lesspoll_infinite_Ord: + "[| Finite(X); ~Finite(a); Ord(a) |] ==> X\a" +apply (unfold lesspoll_def) +apply (rule conjI) +apply (drule nat_le_infinite_Ord [THEN le_imp_lepoll], assumption) +apply (unfold Finite_def) +apply (blast intro: leI [THEN le_imp_subset, THEN subset_imp_lepoll] + ltI eqpoll_imp_lepoll lepoll_trans) +apply (blast intro: eqpoll_sym [THEN eqpoll_trans]) +done + +lemma Union_lesspoll: + "[| \x \ X. x lepoll n & x \ T; well_ord(T, R); X lepoll b; + b nat |] + ==> Union(X)\a" +apply (case_tac "Finite (X)") +apply (blast intro: Card_is_Ord Finite_lesspoll_infinite_Ord + lepoll_nat_imp_Finite Finite_Union) +apply (drule lepoll_imp_ex_le_eqpoll) +apply (erule lt_Ord) +apply (elim exE conjE) +apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption) +apply (erule eqpoll_sym [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1], + THEN exE]) +apply (frule bij_is_surj [THEN surj_image_eq]) +apply (drule image_fun [OF bij_is_fun subset_refl]) +apply (drule sym [THEN trans], assumption) +apply (blast intro: lt_Ord UN_lepoll lt_Card_imp_lesspoll + lt_trans1 lesspoll_trans1) +done + +(* ********************************************************************** *) +(* recfunAC16_lepoll_index *) +(* ********************************************************************** *) + +lemma Un_sing_eq_cons: "A Un {a} = cons(a, A)" +by fast + +lemma Un_lepoll_succ: "A lepoll B ==> A Un {a} lepoll succ(B)" +apply (simp add: Un_sing_eq_cons succ_def) +apply (blast elim!: mem_irrefl intro: cons_lepoll_cong) +done + +lemma Diff_UN_succ_empty: "Ord(a) ==> F(a) - (\b F(a) Un X - (\b X" +by blast + +lemma recfunAC16_Diff_lepoll_1: + "Ord(x) + ==> recfunAC16(f, g, x, a) - (\i F(x); Ord(x) |] + ==> z \ F(LEAST i. z \ F(i)) - (\j<(LEAST i. z \ F(i)). F(j))" +by (fast elim: less_LeastE elim!: LeastI); + +lemma Least_eq_imp_ex: + "[| (LEAST i. w \ F(i)) = (LEAST i. z \ F(i)); + w \ (\i (\i \b (F(b) - (\c (F(b) - (\c A; b \ A |] ==> a=b" +by (fast dest!: lepoll_1_is_sing) + + +lemma UN_lepoll_index: + "[| \ij (\x recfunAC16(f, h, y, a) lepoll y" +apply (erule trans_induct3) +(*0 case*) +apply (simp (no_asm_simp) add: recfunAC16_0 lepoll_refl) +(*succ case*) +apply (simp (no_asm_simp) add: recfunAC16_succ) +apply (blast dest!: succI1 [THEN rev_bspec] + intro: subset_succI [THEN subset_imp_lepoll] Un_lepoll_succ + lepoll_trans) +apply (simp (no_asm_simp) add: recfunAC16_Limit) +apply (blast intro: lt_Ord [THEN recfunAC16_Diff_lepoll_1] UN_lepoll_index) +done + -WO2_AC16 = AC_Equiv + recfunAC16 + AC16_lemmas + Cardinal_aux +lemma Union_recfunAC16_lesspoll: + "[| recfunAC16(f,g,y,a) \ {X \ Pow(A). X\n}; + A\a; y nat |] + ==> Union(recfunAC16(f,g,y,a))\a" +apply (erule eqpoll_def [THEN def_imp_iff, THEN iffD1, THEN exE]) +apply (rule_tac T="A" in Union_lesspoll, simp_all) +apply (blast intro!: eqpoll_imp_lepoll) +apply (blast intro: bij_is_inj Card_is_Ord [THEN well_ord_Memrel] + well_ord_rvimage) +apply (erule lt_Ord [THEN recfunAC16_lepoll_index]) +done + +lemma dbl_Diff_eqpoll: + "[| recfunAC16(f, h, y, a) \ {X \ Pow(A) . X\succ(k #+ m)}; + Card(a); ~ Finite(a); A\a; + k \ nat; y bij(a, {Y \ Pow(A). Y\succ(k)}) |] + ==> A - Union(recfunAC16(f, h, y, a)) - h`y\a" +apply (rule dbl_Diff_eqpoll_Card, simp_all) +apply (simp add: Union_recfunAC16_lesspoll) +apply (rule Finite_lesspoll_infinite_Ord) +apply (rule Finite_def [THEN def_imp_iff, THEN iffD2]) +apply (blast dest: ltD bij_is_fun [THEN apply_type], assumption); +apply (blast intro: Card_is_Ord) +done; + +(* back to the proof *) + +lemmas disj_Un_eqpoll_nat_sum = + eqpoll_trans [THEN eqpoll_trans, + OF disj_Un_eqpoll_sum sum_eqpoll_cong nat_sum_eqpoll_sum, + standard]; + + +lemma Un_in_Collect: "[| x \ Pow(A - B - h`i); x\m; + h \ bij(a, {x \ Pow(A) . x\k}); i nat; m \ nat |] + ==> h ` i Un x \ {x \ Pow(A) . x\k #+ m}" +by (blast intro: disj_Un_eqpoll_nat_sum + dest: ltD bij_is_fun [THEN apply_type]) + + +(* ********************************************************************** *) +(* Lemmas simplifying assumptions *) +(* ********************************************************************** *) + +lemma lemma6: + "[| \yx Q(x,y)); succ(j) F(j)<=X & (\x Q(x,j))" +by (blast intro!: lt_Ord succI1 [THEN ltI, THEN lt_Ord, THEN le_refl]); + + +lemma lemma7: + "[| \x Q(x,j); succ(j) P(j,j) --> (\xj | P(x,j) --> Q(x,j))" +by (fast elim!: leE) + +(* ********************************************************************** *) +(* Lemmas needed to prove ex_next_set, which means that for any successor *) +(* ordinal there is a set satisfying certain properties *) +(* ********************************************************************** *) + +lemma ex_subset_eqpoll: + "[| A\a; ~ Finite(a); Ord(a); m \ nat |] ==> \X \ Pow(A). X\m" +apply (rule lepoll_imp_eqpoll_subset [of m A, THEN exE]) + apply (rule lepoll_trans, rule leI [THEN le_imp_lepoll]) + apply (blast intro: lt_trans2 [OF ltI nat_le_infinite_Ord] Ord_nat) +apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll]) +apply (fast elim!: eqpoll_sym) +done + +lemma subset_Un_disjoint: "[| A \ B Un C; A Int C = 0 |] ==> A \ B" +by blast + + +lemma Int_empty: + "[| X \ Pow(A - Union(B) -C); T \ B; F \ T |] ==> F Int X = 0" +by blast + + +(* ********************************************************************** *) +(* equipollent subset (and finite) is the whole set *) +(* ********************************************************************** *) + + +lemma subset_imp_eq_lemma: + "m \ nat ==> \A B. A \ B & m lepoll A & B lepoll m --> A=B" +apply (induct_tac "m") +apply (fast dest!: lepoll_0_is_0) +apply (intro allI impI) +apply (elim conjE) +apply (rule succ_lepoll_imp_not_empty [THEN not_emptyE], assumption) +apply (frule subsetD [THEN Diff_sing_lepoll], assumption+) +apply (frule lepoll_Diff_sing) +apply (erule allE impE)+ +apply (rule conjI) +prefer 2 apply fast +apply fast +apply (blast elim: equalityE) +done + + +lemma subset_imp_eq: "[| A \ B; m lepoll A; B lepoll m; m \ nat |] ==> A=B" +by (blast dest!: subset_imp_eq_lemma) + + +lemma bij_imp_arg_eq: + "[| f \ bij(a, {Y \ X. Y\succ(k)}); k \ nat; f`b \ f`y; b b=y" +apply (drule subset_imp_eq) +apply (erule_tac [3] nat_succI) +apply (unfold bij_def inj_def) +apply (blast intro: eqpoll_sym eqpoll_imp_lepoll + dest: ltD apply_type)+ +done + + +lemma ex_next_set: + "[| recfunAC16(f, h, y, a) \ {X \ Pow(A) . X\succ(k #+ m)}; + Card(a); ~ Finite(a); A\a; + k \ nat; m \ nat; y bij(a, {Y \ Pow(A). Y\succ(k)}); + ~ (\Y \ recfunAC16(f, h, y, a). h`y \ Y) |] + ==> \X \ {Y \ Pow(A). Y\succ(k #+ m)}. h`y \ X & + (\b X --> + (\T \ recfunAC16(f, h, y, a). ~ h`b \ T))" +apply (erule_tac m1=m in dbl_Diff_eqpoll [THEN ex_subset_eqpoll, THEN bexE], + assumption+) +apply (erule Card_is_Ord, assumption) +apply (frule Un_in_Collect, (erule asm_rl nat_succI)+) +apply (erule CollectE) +apply (rule rev_bexI, simp) +apply (rule conjI, blast) +apply (intro ballI impI oallI notI) +apply (drule subset_Un_disjoint, rule Int_empty, assumption+) +apply (blast dest: bij_imp_arg_eq) +done + +(* ********************************************************************** *) +(* Lemma ex_next_Ord states that for any successor *) +(* ordinal there is a number of the set satisfying certain properties *) +(* ********************************************************************** *) + +lemma ex_next_Ord: + "[| recfunAC16(f, h, y, a) \ {X \ Pow(A) . X\succ(k #+ m)}; + Card(a); ~ Finite(a); A\a; + k \ nat; m \ nat; y bij(a, {Y \ Pow(A). Y\succ(k)}); + f \ bij(a, {Y \ Pow(A). Y\succ(k #+ m)}); + ~ (\Y \ recfunAC16(f, h, y, a). h`y \ Y) |] + ==> \c f`c & + (\b f`c --> + (\T \ recfunAC16(f, h, y, a). ~ h`b \ T))" +apply (drule ex_next_set, assumption+) +apply (erule bexE) +apply (rule oexI) +apply (subst right_inverse_bij, blast, assumption+) +apply (blast intro: bij_converse_bij bij_is_fun [THEN apply_type] ltI + Card_is_Ord) +done + + +(* ********************************************************************** *) +(* Lemma simplifying assumptions *) +(* ********************************************************************** *) + +lemma lemma8: + "[| \xxa \ F(j). P(x, xa)) + --> (\! Y. Y \ F(j) & P(x, Y)); F(j) \ X; + L \ X; P(j, L) & (\x (\xa \ F(j). ~P(x, xa))) |] + ==> F(j) Un {L} \ X & + (\xj | (\xa \ (F(j) Un {L}). P(x, xa)) --> + (\! Y. Y \ (F(j) Un {L}) & P(x, Y)))" +apply (rule conjI) +apply (fast intro!: singleton_subsetI) +apply (rule oallI) +apply (blast elim!: leE oallE) +done + +(* ********************************************************************** *) +(* The main part of the proof: inductive proof of the property of T_gamma *) +(* lemma main_induct *) +(* ********************************************************************** *) + +lemma main_induct: + "[| b < a; f \ bij(a, {Y \ Pow(A) . Y\succ(k #+ m)}); + h \ bij(a, {Y \ Pow(A) . Y\succ(k)}); + ~Finite(a); Card(a); A\a; k \ nat; m \ nat |] + ==> recfunAC16(f, h, b, a) \ {X \ Pow(A) . X\succ(k #+ m)} & + (\xY \ recfunAC16(f, h, b, a). h ` x \ Y) --> + (\! Y. Y \ recfunAC16(f, h, b, a) & h ` x \ Y))" +apply (erule lt_induct) +apply (frule lt_Ord) +apply (erule Ord_cases) +(* case 0 *) +apply (simp add: recfunAC16_0) +(* case Limit *) +prefer 2 apply (simp add: recfunAC16_Limit) + apply (rule lemma5, assumption+) + apply (blast dest!: recfunAC16_mono) +(* case succ *) +apply clarify +apply (erule lemma6 [THEN conjE], assumption) +apply (simp (no_asm_simp) split del: split_if add: recfunAC16_succ) +apply (rule conjI [THEN split_if [THEN iffD2]]) + apply (simp, erule lemma7, assumption) +apply (rule impI) +apply (rule ex_next_Ord [THEN oexE], + assumption+, rule le_refl [THEN lt_trans], assumption+) +apply (erule lemma8, assumption) + apply (rule bij_is_fun [THEN apply_type], assumption) + apply (erule Least_le [THEN lt_trans2, THEN ltD]) + apply (erule lt_Ord) + apply (erule succ_leI) +apply (erule LeastI) +apply (erule lt_Ord) +done + +(* ********************************************************************** *) +(* Lemma to simplify the inductive proof *) +(* - the desired property is a consequence of the inductive assumption *) +(* ********************************************************************** *) + +lemma lemma_simp_induct: + "[| \b. b F(b) \ S & (\xY \ F(b). f`x \ Y)) + --> (\! Y. Y \ F(b) & f`x \ Y)); + f \ a->f``(a); Limit(a); + \i j. i\j --> F(i) \ F(j) |] + ==> (\j S & + (\x \ f``a. \! Y. Y \ (\j Y)" +apply (rule conjI) +apply (rule subsetI) +apply (erule OUN_E, blast) +apply (rule ballI) +apply (erule imageE) +apply (drule ltI, erule Limit_is_Ord) +apply (drule Limit_has_succ, assumption) +apply (frule_tac x1="succ(xa)" in spec [THEN mp], assumption); +apply (erule conjE) +apply (drule ospec) +(** LEVEL 10 **) +apply (erule leI [THEN succ_leE]) +apply (erule impE) +apply (fast elim!: leI [THEN succ_leE, THEN lt_Ord, THEN le_refl]) +apply (drule apply_equality, assumption) +apply (elim conjE ex1E) +(** LEVEL 15 **) +apply (rule ex1I, blast) +apply (elim conjE OUN_E) +apply (erule_tac i="succ(xa)" and j=aa + in Ord_linear_le [OF lt_Ord lt_Ord], assumption) +prefer 2 +apply (drule spec [THEN spec, THEN mp, THEN subsetD], assumption+, blast) +(** LEVEL 20 **) +apply (drule_tac x1="aa" in spec [THEN mp], assumption) +apply (frule succ_leE) +apply (drule spec [THEN spec, THEN mp, THEN subsetD], assumption+, blast) +done + +(* ********************************************************************** *) +(* The target theorem *) +(* ********************************************************************** *) + +theorem WO2_AC16: "[| WO2; 0 nat; m \ nat |] ==> AC16(k #+ m,k)" +apply (unfold AC16_def) +apply (rule allI) +apply (rule impI) +apply (frule WO2_infinite_subsets_eqpoll_X, assumption+) +apply (frule_tac n="k #+ m" in WO2_infinite_subsets_eqpoll_X, simp) +apply simp +apply (frule WO2_imp_ex_Card) +apply (elim exE conjE) +apply (drule eqpoll_trans [THEN eqpoll_sym, + THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], + assumption) +apply (drule eqpoll_trans [THEN eqpoll_sym, + THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], + assumption+) +apply (elim exE) +apply (rename_tac h) +apply (rule_tac x = "\j WO1. - - Every proof (except WO6 ==> WO1 and WO1 ==> WO2) are described as "clear" - by Rubin & Rubin (page 2). - They refer reader to a book by Gödel to see the proof WO1 ==> WO2. - Fortunately order types made this proof also very easy. -*) - -(* ********************************************************************** *) - -Goalw WO_defs "WO2 ==> WO3"; -by (Fast_tac 1); -qed "WO2_WO3"; - -(* ********************************************************************** *) - -Goalw (eqpoll_def::WO_defs) "WO3 ==> WO1"; -by (fast_tac (claset() addSEs [bij_is_inj RS well_ord_rvimage, - well_ord_Memrel RS well_ord_subset]) 1); -qed "WO3_WO1"; - -(* ********************************************************************** *) - -Goalw (eqpoll_def::WO_defs) "WO1 ==> WO2"; -by (fast_tac (claset() addSIs [Ord_ordertype, ordermap_bij]) 1); -qed "WO1_WO2"; - -(* ********************************************************************** *) - -Goal "f \\ A->B ==> (\\x \\ A. {f`x}): A -> {{b}. b \\ B}"; -by (fast_tac (claset() addSIs [lam_type, apply_type]) 1); -qed "lam_sets"; - -Goalw [surj_def] "f \\ surj(A,B) ==> (\\a \\ A. {f`a}) = B"; -by (fast_tac (claset() addSEs [apply_type]) 1); -qed "surj_imp_eq_"; - -Goal "[| f \\ surj(A,B); Ord(A) |] ==> (\\a WO4(1)"; -by (rtac allI 1); -by (eres_inst_tac [("x","A")] allE 1); -by (etac exE 1); -by (REPEAT (resolve_tac [exI, conjI] 1)); -by (etac Ord_ordertype 1); -by (rtac conjI 1); -by (eresolve_tac [ordermap_bij RS bij_converse_bij RS bij_is_fun RS - lam_sets RS domain_of_fun] 1); -by (asm_simp_tac (simpset() addsimps [singleton_eqpoll_1 RS eqpoll_imp_lepoll, - Ord_ordertype RSN (2, ordermap_bij RS bij_converse_bij RS - bij_is_surj RS surj_imp_eq)]) 1); -qed "WO1_WO4"; - -(* ********************************************************************** *) - -Goalw WO_defs "[| m le n; WO4(m) |] ==> WO4(n)"; -by (blast_tac (claset() addSDs [spec] - addIs [le_imp_lepoll RSN (2, lepoll_trans)]) 1); -qed "WO4_mono"; - -(* ********************************************************************** *) - -Goalw WO_defs "[| m \\ nat; 1 le m; WO4(m) |] ==> WO5"; -by (Blast_tac 1); -qed "WO4_WO5"; - -(* ********************************************************************** *) - -Goalw WO_defs "WO5 ==> WO6"; -by (Blast_tac 1); -qed "WO5_WO6"; - - -(* ********************************************************************** - The proof of "WO6 ==> WO1". Simplified by L C Paulson. - - From the book "Equivalents of the Axiom of Choice" by Rubin & Rubin, - pages 2-5 -************************************************************************* *) - -goal OrderType.thy - "!!i j k. [| k < i++j; Ord(i); Ord(j) |] ==> \ -\ k < i | (~ k f`b"; -by (Blast_tac 1); -qed "domain_uu_subset"; - -Goal "\\b \ -\ \\bgd f`b * f`g"; -by (Blast_tac 1); -qed "uu_subset1"; - -Goalw [uu_def] "uu(f,b,g,d) \\ f`d"; -by (Blast_tac 1); -qed "uu_subset2"; - -Goal "[| \\b uu(f,b,g,d) lepoll m"; -by (fast_tac (claset() - addSEs [uu_subset2 RS subset_imp_lepoll RS lepoll_trans]) 1); -qed "uu_lepoll_m"; - -(* ********************************************************************** *) -(* Two cases for lemma ii *) -(* ********************************************************************** *) -Goalw [lesspoll_def] - "\\bgd (\\b 0 --> \ -\ (\\gd 0 & u(f,b,g,d) lesspoll m)) \ -\ | (\\b 0 & (\\gd 0 --> \ -\ u(f,b,g,d) eqpoll m))"; -by (Asm_simp_tac 1); -by (blast_tac (claset() delrules [equalityI]) 1); -qed "cases"; - -(* ********************************************************************** *) -(* Lemmas used in both cases *) -(* ********************************************************************** *) -Goal "Ord(a) ==> (\\bb lemmas *) -(* ********************************************************************** *) - -Goalw [vv1_def] "vv1(f,m,b) \\ f`b"; -by (rtac (LetI RS LetI) 1); -by (simp_tac (simpset() addsimps [domain_uu_subset]) 1); -qed "vv1_subset"; - -(* ********************************************************************** *) -(* Case 1 \\ Union of images is the whole "y" *) -(* ********************************************************************** *) -Goalw [gg1_def] - "!! a f y. [| Ord(a); m \\ nat |] ==> \ -\ (\\bbx. Ord(x) & P(a, x)) |] \ -\ ==> P(Least_a, LEAST b. P(Least_a, b))"; -by (etac ssubst 1); -by (res_inst_tac [("Q","%z. P(z, LEAST b. P(z, b))")] LeastI2 1); -by (REPEAT (fast_tac (claset() addSEs [LeastI]) 1)); -qed "nested_LeastI"; - -bind_thm ("nested_Least_instance", - inst "P" - "%g d. domain(uu(f,b,g,d)) \\ 0 & domain(uu(f,b,g,d)) lepoll m" - nested_LeastI); - -Goalw [gg1_def] - "!!a. [| Ord(a); m \\ nat; \ -\ \\b0 --> \ -\ (\\gd 0 & \ -\ domain(uu(f,b,g,d)) lepoll m); \ -\ \\b gg1(f,a,m)`b lepoll m"; -by (Asm_simp_tac 1); -by (safe_tac (claset() addSEs [lt_oadd_odiff_cases])); -(*Case b show vv1(f,m,b) lepoll m *) -by (asm_simp_tac (simpset() addsimps [vv1_def, Let_def]) 1); -by (fast_tac (claset() addIs [nested_Least_instance RS conjunct2] - addSEs [lt_Ord] - addSIs [empty_lepollI]) 1); -(*Case a le b \\ show ww1(f,m,b--a) lepoll m *) -by (asm_simp_tac (simpset() addsimps [ww1_def]) 1); -by (excluded_middle_tac "f`(b--a) = 0" 1); -by (asm_simp_tac (simpset() addsimps [empty_lepollI]) 2); -by (rtac Diff_lepoll 1); -by (Blast_tac 1); -by (rtac vv1_subset 1); -by (dtac (ospec RS mp) 1); -by (REPEAT (eresolve_tac [asm_rl, oexE, conjE] 1)); -by (asm_simp_tac (simpset() - addsimps [vv1_def, Let_def, lt_Ord, - nested_Least_instance RS conjunct1]) 1); -qed "gg1_lepoll_m"; - -(* ********************************************************************** *) -(* Case 2 \\ lemmas *) -(* ********************************************************************** *) - -(* ********************************************************************** *) -(* Case 2 \\ vv2_subset *) -(* ********************************************************************** *) - -Goalw [uu_def] "[| b0; f`g\\0; \ -\ y*y \\ y; (\\b \\d 0"; -by (fast_tac (claset() addSIs [not_emptyI] - addSDs [SigmaI RSN (2, subsetD)] - addSEs [not_emptyE]) 1); -qed "ex_d_uu_not_empty"; - -Goal "[| b0; f`g\\0; \ -\ y*y \\ y; (\\b uu(f,b,g,LEAST d. (uu(f,b,g,d) \\ 0)) \\ 0"; -by (dtac ex_d_uu_not_empty 1 THEN REPEAT (assume_tac 1)); -by (fast_tac (claset() addSEs [LeastI, lt_Ord]) 1); -qed "uu_not_empty"; - -Goal "[| r \\ A*B; r\\0 |] ==> domain(r)\\0"; -by (Blast_tac 1); -qed "not_empty_rel_imp_domain"; - -Goal "[| b0; f`g\\0; \ -\ y*y \\ y; (\\b (LEAST d. uu(f,b,g,d) \\ 0) < a"; -by (eresolve_tac [ex_d_uu_not_empty RS oexE] 1 - THEN REPEAT (assume_tac 1)); -by (resolve_tac [Least_le RS lt_trans1] 1 - THEN (REPEAT (ares_tac [lt_Ord] 1))); -qed "Least_uu_not_empty_lt_a"; - -Goal "[| B \\ A; a\\B |] ==> B \\ A-{a}"; -by (Blast_tac 1); -qed "subset_Diff_sing"; - -(*Could this be proved more directly?*) -Goal "[| A lepoll m; m lepoll B; B \\ A; m \\ nat |] ==> A=B"; -by (etac natE 1); -by (fast_tac (claset() addSDs [lepoll_0_is_0] addSIs [equalityI]) 1); -by (safe_tac (claset() addSIs [equalityI])); -by (rtac ccontr 1); -by (etac (subset_Diff_sing RS subset_imp_lepoll - RSN (2, Diff_sing_lepoll RSN (3, lepoll_trans RS lepoll_trans)) RS - succ_lepoll_natE) 1 - THEN REPEAT (assume_tac 1)); -qed "supset_lepoll_imp_eq"; - -Goal "[| \\gd0 --> \ -\ domain(uu(f, b, g, d)) eqpoll succ(m); \ -\ \\b y; \ -\ (\\b0; f`g\\0; m \\ nat; s \\ f`b \ -\ |] ==> uu(f, b, g, LEAST d. uu(f,b,g,d)\\0) \\ f`b -> f`g"; -by (dres_inst_tac [("x2","g")] (ospec RS ospec RS mp) 1); -by (rtac ([uu_subset1, uu_not_empty] MRS not_empty_rel_imp_domain) 3); -by (rtac Least_uu_not_empty_lt_a 2 THEN TRYALL assume_tac); -by (resolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS - (Least_uu_not_empty_lt_a RSN (2, uu_lepoll_m) RSN (2, - uu_subset1 RSN (4, rel_is_fun)))] 1 - THEN TRYALL assume_tac); -by (rtac (eqpoll_sym RS eqpoll_imp_lepoll RSN (2, supset_lepoll_imp_eq)) 1); -by (REPEAT_FIRST assume_tac); -by (REPEAT (fast_tac (claset() addSIs [domain_uu_subset]) 1)); -qed "uu_Least_is_fun"; - -Goalw [vv2_def] - "!!a. [| \\gd0 --> \ -\ domain(uu(f, b, g, d)) eqpoll succ(m); \ -\ \\b y; \ -\ (\\b nat; s \\ f`b \ -\ |] ==> vv2(f,b,g,s) \\ f`g"; -by (split_tac [split_if] 1); -by Safe_tac; -by (etac (uu_Least_is_fun RS apply_type) 1); -by (REPEAT_SOME (fast_tac (claset() addSIs [not_emptyI, singleton_subsetI]))); -qed "vv2_subset"; - -(* ********************************************************************** *) -(* Case 2 \\ Union of images is the whole "y" *) -(* ********************************************************************** *) -Goalw [gg2_def] - "!!a. [| \\gd 0 --> \ -\ domain(uu(f,b,g,d)) eqpoll succ(m); \ -\ \\b y; \ -\ (\\b nat; s \\ f`b; b (\\g nat; m\\0 |] ==> vv2(f,b,g,s) lepoll m"; -by (asm_simp_tac (simpset() addsimps [empty_lepollI]) 1); -by (fast_tac (claset() - addSDs [le_imp_subset RS subset_imp_lepoll RS lepoll_0_is_0] - addSIs [singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans, - not_lt_imp_le RS le_imp_subset RS subset_imp_lepoll, - nat_into_Ord, nat_1I]) 1); -qed "vv2_lepoll"; - -Goalw [ww2_def] - "[| \\b nat; vv2(f,b,g,d) \\ f`g |] \ -\ ==> ww2(f,b,g,d) lepoll m"; -by (excluded_middle_tac "f`g = 0" 1); -by (asm_simp_tac (simpset() addsimps [empty_lepollI]) 2); -by (dtac ospec 1 THEN (assume_tac 1)); -by (rtac Diff_lepoll 1 THEN (TRYALL assume_tac)); -by (asm_simp_tac (simpset() addsimps [vv2_def, not_emptyI]) 1); -qed "ww2_lepoll"; - -Goalw [gg2_def] - "!!a. [| \\gd 0 --> \ -\ domain(uu(f,b,g,d)) eqpoll succ(m); \ -\ \\b y; \ -\ (\\b f`b; m \\ nat; m\\ 0; g gg2(f,a,b,s) ` g lepoll m"; -by (Asm_simp_tac 1); -by (safe_tac (claset() addSEs [lt_oadd_odiff_cases, lt_Ord2])); -by (asm_simp_tac (simpset() addsimps [vv2_lepoll]) 1); -by (asm_simp_tac (simpset() addsimps [ww2_lepoll, vv2_subset]) 1); -qed "gg2_lepoll_m"; - -(* ********************************************************************** *) -(* lemma ii *) -(* ********************************************************************** *) -Goalw [NN_def] "[| succ(m) \\ NN(y); y*y \\ y; m \\ nat; m\\0 |] ==> m \\ NN(y)"; -by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1)); -by (resolve_tac [quant_domain_uu_lepoll_m RS cases RS disjE] 1 - THEN (assume_tac 1)); -(* case 1 *) -by (asm_full_simp_tac (simpset() addsimps [lesspoll_succ_iff]) 1); -by (res_inst_tac [("x","a++a")] exI 1); -by (fast_tac (claset() addSIs [Ord_oadd, domain_gg1, UN_gg1_eq, - gg1_lepoll_m]) 1); -(* case 2 *) -by (REPEAT (eresolve_tac [oexE, conjE] 1)); -by (res_inst_tac [("A","f`?B")] not_emptyE 1 THEN (assume_tac 1)); -by (rtac CollectI 1); -by (etac succ_natD 1); -by (res_inst_tac [("x","a++a")] exI 1); -by (res_inst_tac [("x","gg2(f,a,b,x)")] exI 1); -(*Calling fast_tac might get rid of the res_inst_tac calls, but it - is just too slow.*) -by (asm_simp_tac (simpset() addsimps - [Ord_oadd, domain_gg2, UN_gg2_eq, gg2_lepoll_m]) 1); -qed "lemma_ii"; - - -(* ********************************************************************** *) -(* lemma iv - p. 4 \\ *) -(* For every set x there is a set y such that x Un (y * y) \\ y *) -(* ********************************************************************** *) - -(* the quantifier \\looks inelegant but makes the proofs shorter *) -(* (used only in the following two lemmas) *) - -Goal "\\n \\ nat. rec(n, x, %k r. r Un r*r) \\ \ -\ rec(succ(n), x, %k r. r Un r*r)"; -by (fast_tac (claset() addIs [rec_succ RS ssubst]) 1); -qed "z_n_subset_z_succ_n"; - -Goal "[| \\n \\ nat. f(n)<=f(succ(n)); n le m; n \\ nat; m \\ nat |] \ -\ ==> f(n)<=f(m)"; -by (eres_inst_tac [("P","n le m")] rev_mp 1); -by (res_inst_tac [("P","%z. n le z --> f(n) \\ f(z)")] nat_induct 1); -by (REPEAT (fast_tac le_cs 1)); -qed "le_subsets"; - -Goal "[| n le m; m \\ nat |] ==> \ -\ rec(n, x, %k r. r Un r*r) \\ rec(m, x, %k r. r Un r*r)"; -by (resolve_tac [z_n_subset_z_succ_n RS le_subsets] 1 - THEN (TRYALL assume_tac)); -by (eresolve_tac [Ord_nat RSN (2, ltI) RSN (2, lt_trans1) RS ltD] 1 - THEN (assume_tac 1)); -qed "le_imp_rec_subset"; - -Goal "\\y. x Un y*y \\ y"; -by (res_inst_tac [("x","\\n \\ nat. rec(n, x, %k r. r Un r*r)")] exI 1); -by Safe_tac; -by (rtac (nat_0I RS UN_I) 1); -by (Asm_simp_tac 1); -by (res_inst_tac [("a","succ(n Un na)")] UN_I 1); -by (eresolve_tac [Un_nat_type RS nat_succI] 1 THEN (assume_tac 1)); -by (fast_tac (ZF_cs addIs [le_imp_rec_subset RS subsetD] - addSIs [Un_upper1_le, Un_upper2_le, Un_nat_type] - addSEs [nat_into_Ord] addss (simpset())) 1); -qed "lemma_iv"; - -(* ********************************************************************** *) -(* Rubin & Rubin wrote \\ *) -(* "It follows from (ii) and mathematical induction that if y*y \\ y then *) -(* y can be well-ordered" *) - -(* In fact we have to prove \\ *) -(* * WO6 ==> NN(y) \\ 0 *) -(* * reverse induction which lets us infer that 1 \\ NN(y) *) -(* * 1 \\ NN(y) ==> y can be well-ordered *) -(* ********************************************************************** *) - -(* ********************************************************************** *) -(* WO6 ==> NN(y) \\ 0 *) -(* ********************************************************************** *) - -Goalw [WO6_def, NN_def] "WO6 ==> NN(y) \\ 0"; -by (fast_tac ZF_cs 1); (*SLOW if current claset is used*) -qed "WO6_imp_NN_not_empty"; - -(* ********************************************************************** *) -(* 1 \\ NN(y) ==> y can be well-ordered *) -(* ********************************************************************** *) - -Goal "[| (\\b y; \\b \\cb y; \\b f` (LEAST i. f`i = {x}) = {x}"; -by (dtac lemma1 1 THEN REPEAT (assume_tac 1)); -by (fast_tac (claset() addSEs [lt_Ord] addIs [LeastI]) 1); -val lemma2 = result(); - -Goalw [NN_def] "1 \\ NN(y) ==> \\a f. Ord(a) & f \\ inj(y, a)"; -by (etac CollectE 1); -by (REPEAT (eresolve_tac [exE, conjE] 1)); -by (res_inst_tac [("x","a")] exI 1); -by (res_inst_tac [("x","\\x \\ y. LEAST i. f`i = {x}")] exI 1); -by (rtac conjI 1 THEN (assume_tac 1)); -by (res_inst_tac [("d","%i. THE x. x \\ f`i")] lam_injective 1); -by (dtac lemma1 1 THEN REPEAT (assume_tac 1)); -by (fast_tac (claset() addSEs [Least_le RS lt_trans1 RS ltD, lt_Ord]) 1); -by (resolve_tac [lemma2 RS ssubst] 1 THEN REPEAT (assume_tac 1)); -by (Blast_tac 1); -qed "NN_imp_ex_inj"; - -Goal "[| y*y \\ y; 1 \\ NN(y) |] ==> \\r. well_ord(y, r)"; -by (dtac NN_imp_ex_inj 1); -by (fast_tac (claset() addSEs [well_ord_Memrel RSN (2, well_ord_rvimage)]) 1); -qed "y_well_ord"; - -(* ********************************************************************** *) -(* reverse induction which lets us infer that 1 \\ NN(y) *) -(* ********************************************************************** *) - -val [prem1, prem2] = goal thy - "[| n \\ nat; !!m. [| m \\ nat; m\\0; P(succ(m)) |] ==> P(m) |] \ -\ ==> n\\0 --> P(n) --> P(1)"; -by (rtac (prem1 RS nat_induct) 1); -by (Blast_tac 1); -by (excluded_middle_tac "x=0" 1); -by (Blast_tac 2); -by (fast_tac (claset() addSIs [prem2]) 1); -qed "rev_induct_lemma"; - -val prems = -Goal "[| P(n); n \\ nat; n\\0; \ -\ !!m. [| m \\ nat; m\\0; P(succ(m)) |] ==> P(m) |] \ -\ ==> P(1)"; -by (resolve_tac [rev_induct_lemma RS impE] 1); -by (etac impE 4 THEN (assume_tac 5)); -by (REPEAT (ares_tac prems 1)); -qed "rev_induct"; - -Goalw [NN_def] "n \\ NN(y) ==> n \\ nat"; -by (etac CollectD1 1); -qed "NN_into_nat"; - -Goal "[| n \\ NN(y); y*y \\ y; n\\0 |] ==> 1 \\ NN(y)"; -by (rtac rev_induct 1 THEN REPEAT (ares_tac [NN_into_nat] 1)); -by (rtac lemma_ii 1 THEN REPEAT (assume_tac 1)); -val lemma3 = result(); - -(* ********************************************************************** *) -(* Main theorem "WO6 ==> WO1" *) -(* ********************************************************************** *) - -(* another helpful lemma *) -Goalw [NN_def] "0 \\ NN(y) ==> y=0"; -by (fast_tac (claset() addSIs [equalityI] - addSDs [lepoll_0_is_0] addEs [subst]) 1); -qed "NN_y_0"; - -Goalw [WO1_def] "WO6 ==> WO1"; -by (rtac allI 1); -by (excluded_middle_tac "A=0" 1); -by (fast_tac (claset() addSIs [well_ord_Memrel, nat_0I RS nat_into_Ord]) 2); -by (res_inst_tac [("x1","A")] (lemma_iv RS revcut_rl) 1); -by (etac exE 1); -by (dtac WO6_imp_NN_not_empty 1); -by (eresolve_tac [Un_subset_iff RS iffD1 RS conjE] 1); -by (eres_inst_tac [("A","NN(y)")] not_emptyE 1); -by (ftac y_well_ord 1); -by (fast_tac (claset() addEs [well_ord_subset]) 2); -by (fast_tac (claset() addSIs [lemma3] addSDs [NN_y_0] addSEs [not_emptyE]) 1); -qed "WO6_imp_WO1"; - diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/WO6_WO1.thy --- a/src/ZF/AC/WO6_WO1.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/AC/WO6_WO1.thy Wed Jan 16 17:52:06 2002 +0100 @@ -2,49 +2,532 @@ ID: $Id$ Author: Krzysztof Grabczewski -The proof of "WO6 ==> WO1". +Proofs needed to state that formulations WO1,...,WO6 are all equivalent. +The only hard one is WO6 ==> WO1. -From the book "Equivalents of the Axiom of Choice" by Rubin & Rubin, -pages 2-5 +Every proof (except WO6 ==> WO1 and WO1 ==> WO2) are described as "clear" +by Rubin & Rubin (page 2). +They refer reader to a book by Gödel to see the proof WO1 ==> WO2. +Fortunately order types made this proof also very easy. *) -WO6_WO1 = "rel_is_fun" + AC_Equiv + Let + +theory WO6_WO1 = Cardinal_aux: -consts +constdefs (* Auxiliary definitions used in proof *) - NN :: i => i - uu :: [i, i, i, i] => i - vv1, ww1, gg1 :: [i, i, i] => i - vv2, ww2, gg2 :: [i, i, i, i] => i + NN :: "i => i" + "NN(y) == {m \ nat. \a. \f. Ord(a) & domain(f)=a & + (\bb m)}" + + uu :: "[i, i, i, i] => i" + "uu(f, beta, gamma, delta) == (f`beta * f`gamma) Int f`delta" + + (** Definitions for case 1 **) + vv1 :: "[i, i, i] => i" + "vv1(f,m,b) == + let g = LEAST g. (\d. Ord(d) & (domain(uu(f,b,g,d)) \ 0 & + domain(uu(f,b,g,d)) \ m)); + d = LEAST d. domain(uu(f,b,g,d)) \ 0 & + domain(uu(f,b,g,d)) \ m + in if f`b \ 0 then domain(uu(f,b,g,d)) else 0" + + ww1 :: "[i, i, i] => i" + "ww1(f,m,b) == f`b - vv1(f,m,b)" + + gg1 :: "[i, i, i] => i" + "gg1(f,a,m) == \b \ a++a. if b i" + "vv2(f,b,g,s) == + if f`g \ 0 then {uu(f, b, g, LEAST d. uu(f,b,g,d) \ 0)`s} else 0" + + ww2 :: "[i, i, i, i] => i" + "ww2(f,b,g,s) == f`g - vv2(f,b,g,s)" + + gg2 :: "[i, i, i, i] => i" + "gg2(f,a,b,s) == + \g \ a++a. if g WO3" +by (unfold WO2_def WO3_def, fast) + +(* ********************************************************************** *) + +lemma WO3_WO1: "WO3 ==> WO1" +apply (unfold eqpoll_def WO1_def WO3_def) +apply (intro allI) +apply (drule_tac x=A in spec) +apply (blast intro: bij_is_inj well_ord_rvimage + well_ord_Memrel [THEN well_ord_subset]) +done + +(* ********************************************************************** *) + +lemma WO1_WO2: "WO1 ==> WO2" +apply (unfold eqpoll_def WO1_def WO2_def) +apply (blast intro!: Ord_ordertype ordermap_bij) +done + +(* ********************************************************************** *) + +lemma lam_sets: "f \ A->B ==> (\x \ A. {f`x}): A -> {{b}. b \ B}" +by (fast intro!: lam_type apply_type) + +lemma surj_imp_eq_: "f \ surj(A,B) ==> (\a \ A. {f`a}) = B" +apply (unfold surj_def) +apply (fast elim!: apply_type) +done + +lemma surj_imp_eq: "[| f \ surj(A,B); Ord(A) |] ==> (\a WO4(1)" +apply (unfold WO1_def WO4_def) +apply (rule allI) +apply (erule_tac x = "A" in allE) +apply (erule exE) +apply (intro exI conjI) +apply (erule Ord_ordertype) +apply (erule ordermap_bij [THEN bij_converse_bij, THEN bij_is_fun, THEN lam_sets, THEN domain_of_fun]) +apply (simp_all add: singleton_eqpoll_1 eqpoll_imp_lepoll Ord_ordertype + ordermap_bij [THEN bij_converse_bij, THEN bij_is_surj, THEN surj_imp_eq]) +done + +(* ********************************************************************** *) + +lemma WO4_mono: "[| m\n; WO4(m) |] ==> WO4(n)" +apply (unfold WO4_def) +apply (blast dest!: spec intro: lepoll_trans [OF _ le_imp_lepoll]) +done + +(* ********************************************************************** *) + +lemma WO4_WO5: "[| m \ nat; 1\m; WO4(m) |] ==> WO5" +by (unfold WO4_def WO5_def, blast) + +(* ********************************************************************** *) + +lemma WO5_WO6: "WO5 ==> WO6" +by (unfold WO4_def WO5_def WO6_def, blast) + + +(* ********************************************************************** + The proof of "WO6 ==> WO1". Simplified by L C Paulson. + + From the book "Equivalents of the Axiom of Choice" by Rubin & Rubin, + pages 2-5 +************************************************************************* *) + +lemma lt_oadd_odiff_disj: + "[| k < i++j; Ord(i); Ord(j) |] + ==> k < i | (~ k f`b" +by (unfold uu_def, blast) + +lemma quant_domain_uu_lepoll_m: + "\b m ==> \bgd m" +by (blast intro: domain_uu_subset [THEN subset_imp_lepoll] lepoll_trans) + +lemma uu_subset1: "uu(f,b,g,d) \ f`b * f`g" +by (unfold uu_def, blast) + +lemma uu_subset2: "uu(f,b,g,d) \ f`d" +by (unfold uu_def, blast) + +lemma uu_lepoll_m: "[| \b m; d uu(f,b,g,d) \ m" +by (blast intro: uu_subset2 [THEN subset_imp_lepoll] lepoll_trans) + +(* ********************************************************************** *) +(* Two cases for lemma ii *) +(* ********************************************************************** *) +lemma cases: + "\bgd m + ==> (\b 0 --> + (\gd 0 & u(f,b,g,d) \ m)) + | (\b 0 & (\gd 0 --> + u(f,b,g,d) \ m))" +apply (unfold lesspoll_def) +apply (blast del: equalityI) +done + +(* ********************************************************************** *) +(* Lemmas used in both cases *) +(* ********************************************************************** *) +lemma UN_oadd: "Ord(a) ==> (\bb f`b" +by (simp add: vv1_def Let_def domain_uu_subset) + +(* ********************************************************************** *) +(* Case 1: Union of images is the whole "y" *) +(* ********************************************************************** *) +lemma UN_gg1_eq: + "[| Ord(a); m \ nat |] ==> (\bb nat. \\a. \\f. Ord(a) & domain(f)=a & - (\\bbx. Ord(x) & P(a, x)) |] + ==> P(Least_a, LEAST b. P(Least_a, b))" +apply (erule ssubst) +apply (rule_tac Q = "%z. P (z, LEAST b. P (z, b))" in LeastI2) +apply (fast elim!: LeastI)+ +done + +lemmas nested_Least_instance = + nested_LeastI [of "%g d. domain(uu(f,b,g,d)) \ 0 & + domain(uu(f,b,g,d)) \ m", + standard] (*generalizes the Variables!*) - uu_def "uu(f, beta, gamma, delta) == (f`beta * f`gamma) Int f`delta" +lemma gg1_lepoll_m: + "[| Ord(a); m \ nat; + \b0 --> + (\gd 0 & + domain(uu(f,b,g,d)) \ m); + \b succ(m); b gg1(f,a,m)`b \ m" +apply (unfold gg1_def, simp) +apply (safe dest!: lt_oadd_odiff_disj) +(*Case b show vv1(f,m,b) \ m *) +apply (simp add: vv1_def Let_def) +apply (fast intro: nested_Least_instance [THEN conjunct2] + elim!: lt_Ord intro!: empty_lepollI) +(*Case a\b \ show ww1(f,m,b--a) \ m *) +apply (simp add: ww1_def) +apply (case_tac "f` (b--a) = 0", simp add: empty_lepollI) +apply (rule Diff_lepoll, blast) +apply (rule vv1_subset) +apply (drule ospec [THEN mp], assumption+) +apply (elim oexE conjE) +apply (simp add: vv1_def Let_def lt_Ord nested_Least_instance [THEN conjunct1]) +done + +(* ********************************************************************** *) +(* Case 2: lemmas *) +(* ********************************************************************** *) - (*Definitions for case 1*) +(* ********************************************************************** *) +(* Case 2: vv2_subset *) +(* ********************************************************************** *) + +lemma ex_d_uu_not_empty: + "[| b0; f`g\0; + y*y \ y; (\b \d 0" +by (unfold uu_def, blast) + +lemma uu_not_empty: + "[| b0; f`g\0; y*y \ y; (\b uu(f,b,g,LEAST d. (uu(f,b,g,d) \ 0)) \ 0" +apply (drule ex_d_uu_not_empty, assumption+) +apply (fast elim!: LeastI lt_Ord) +done + +lemma not_empty_rel_imp_domain: "[| r \ A*B; r\0 |] ==> domain(r)\0" +by blast + +lemma Least_uu_not_empty_lt_a: + "[| b0; f`g\0; y*y \ y; (\b (LEAST d. uu(f,b,g,d) \ 0) < a" +apply (erule ex_d_uu_not_empty [THEN oexE], assumption+) +apply (blast intro: Least_le [THEN lt_trans1] lt_Ord) +done + +lemma subset_Diff_sing: "[| B \ A; a\B |] ==> B \ A-{a}" +by blast - vv1_def "vv1(f,m,b) == - let g = LEAST g. (\\d. Ord(d) & (domain(uu(f,b,g,d)) \\ 0 & - domain(uu(f,b,g,d)) lepoll m)); - d = LEAST d. domain(uu(f,b,g,d)) \\ 0 & - domain(uu(f,b,g,d)) lepoll m - in if f`b \\ 0 then domain(uu(f,b,g,d)) else 0" +(*Could this be proved more directly?*) +lemma supset_lepoll_imp_eq: + "[| A \ m; m \ B; B \ A; m \ nat |] ==> A=B" +apply (erule natE) +apply (fast dest!: lepoll_0_is_0 intro!: equalityI) +apply (safe intro!: equalityI) +apply (rule ccontr) +apply (rule succ_lepoll_natE) + apply (erule lepoll_trans) + apply (rule lepoll_trans) + apply (erule subset_Diff_sing [THEN subset_imp_lepoll], assumption) + apply (rule Diff_sing_lepoll, assumption+) +done + +lemma uu_Least_is_fun: + "[| \gd0 --> + domain(uu(f, b, g, d)) \ succ(m); + \b succ(m); y*y \ y; + (\b0; f`g\0; m \ nat; s \ f`b |] + ==> uu(f, b, g, LEAST d. uu(f,b,g,d)\0) \ f`b -> f`g" +apply (drule_tac x2=g in ospec [THEN ospec, THEN mp]) + apply (rule_tac [3] not_empty_rel_imp_domain [OF uu_subset1 uu_not_empty]) + apply (rule_tac [2] Least_uu_not_empty_lt_a, assumption+) +apply (rule rel_is_fun) + apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll]) + apply (erule uu_lepoll_m) + apply (rule Least_uu_not_empty_lt_a, assumption+) +apply (rule uu_subset1) +apply (rule supset_lepoll_imp_eq [OF _ eqpoll_sym [THEN eqpoll_imp_lepoll]]) +apply (fast intro!: domain_uu_subset)+ +done + +lemma vv2_subset: + "[| \gd0 --> + domain(uu(f, b, g, d)) \ succ(m); + \b succ(m); y*y \ y; + (\b nat; s \ f`b |] + ==> vv2(f,b,g,s) \ f`g" +apply (simp add: vv2_def) +apply (blast intro: uu_Least_is_fun [THEN apply_type]) +done + +(* ********************************************************************** *) +(* Case 2: Union of images is the whole "y" *) +(* ********************************************************************** *) +lemma UN_gg2_eq: + "[| \gd 0 --> + domain(uu(f,b,g,d)) \ succ(m); + \b succ(m); y*y \ y; + (\b nat; s \ f`b; b (\g nat; m\0 |] ==> vv2(f,b,g,s) \ m" +apply (unfold vv2_def) +apply (simp add: empty_lepollI) +apply (fast dest!: le_imp_subset [THEN subset_imp_lepoll, THEN lepoll_0_is_0] + intro!: singleton_eqpoll_1 [THEN eqpoll_imp_lepoll, THEN lepoll_trans] + not_lt_imp_le [THEN le_imp_subset, THEN subset_imp_lepoll] + nat_into_Ord nat_1I) +done + +lemma ww2_lepoll: + "[| \b succ(m); g nat; vv2(f,b,g,d) \ f`g |] + ==> ww2(f,b,g,d) \ m" +apply (unfold ww2_def) +apply (case_tac "f`g = 0") +apply (simp add: empty_lepollI) +apply (drule ospec, assumption) +apply (rule Diff_lepoll, assumption+) +apply (simp add: vv2_def not_emptyI) +done + +lemma gg2_lepoll_m: + "[| \gd 0 --> + domain(uu(f,b,g,d)) \ succ(m); + \b succ(m); y*y \ y; + (\b f`b; m \ nat; m\ 0; g gg2(f,a,b,s) ` g \ m" +apply (unfold gg2_def, simp) +apply (safe elim!: lt_Ord2 dest!: lt_oadd_odiff_disj) + apply (simp add: vv2_lepoll) +apply (simp add: ww2_lepoll vv2_subset) +done + +(* ********************************************************************** *) +(* lemma ii *) +(* ********************************************************************** *) +lemma lemma_ii: "[| succ(m) \ NN(y); y*y \ y; m \ nat; m\0 |] ==> m \ NN(y)" +apply (unfold NN_def) +apply (elim CollectE exE conjE) +apply (rule quant_domain_uu_lepoll_m [THEN cases, THEN disjE], assumption) +(* case 1 *) +apply (simp add: lesspoll_succ_iff) +apply (rule_tac x = "a++a" in exI) +apply (fast intro!: Ord_oadd domain_gg1 UN_gg1_eq gg1_lepoll_m) +(* case 2 *) +apply (elim oexE conjE) +apply (rule_tac A = "f`?B" in not_emptyE, assumption) +apply (rule CollectI) +apply (erule succ_natD) +apply (rule_tac x = "a++a" in exI) +apply (rule_tac x = "gg2 (f,a,b,x) " in exI) +(*Calling fast_tac might get rid of the res_inst_tac calls, but it + is just too slow.*) +apply (simp add: Ord_oadd domain_gg2 UN_gg2_eq gg2_lepoll_m) +done + + +(* ********************************************************************** *) +(* lemma iv - p. 4: *) +(* For every set x there is a set y such that x Un (y * y) \ y *) +(* ********************************************************************** *) + + +(* The leading \-quantifier looks odd but makes the proofs shorter + (used only in the following two lemmas) *) - gg1_def "gg1(f,a,m) == \\b \\ a++a. if bn \ nat. rec(n, x, %k r. r Un r*r) \ rec(succ(n), x, %k r. r Un r*r)" +by (fast intro: rec_succ [THEN ssubst]) + +lemma le_subsets: + "[| \n \ nat. f(n)<=f(succ(n)); n\m; n \ nat; m \ nat |] + ==> f(n)<=f(m)" +apply (erule_tac P = "n\m" in rev_mp) +apply (rule_tac P = "%z. n\z --> f (n) \ f (z) " in nat_induct) +apply (auto simp add: le_iff) +done + +lemma le_imp_rec_subset: + "[| n\m; m \ nat |] + ==> rec(n, x, %k r. r Un r*r) \ rec(m, x, %k r. r Un r*r)" +apply (rule z_n_subset_z_succ_n [THEN le_subsets]) +apply (blast intro: lt_nat_in_nat)+ +done + +lemma lemma_iv: "\y. x Un y*y \ y" +apply (rule_tac x = "\n \ nat. rec (n, x, %k r. r Un r*r) " in exI) +apply safe +apply (rule nat_0I [THEN UN_I], simp) +apply (rule_tac a = "succ (n Un na) " in UN_I) +apply (erule Un_nat_type [THEN nat_succI], assumption) +apply (auto intro: le_imp_rec_subset [THEN subsetD] + intro!: Un_upper1_le Un_upper2_le Un_nat_type + elim!: nat_into_Ord) +done - vv2_def "vv2(f,b,g,s) == - if f`g \\ 0 then {uu(f, b, g, LEAST d. uu(f,b,g,d) \\ 0)`s} else 0" +(* ********************************************************************** *) +(* Rubin & Rubin wrote, *) +(* "It follows from (ii) and mathematical induction that if y*y \ y then *) +(* y can be well-ordered" *) + +(* In fact we have to prove *) +(* * WO6 ==> NN(y) \ 0 *) +(* * reverse induction which lets us infer that 1 \ NN(y) *) +(* * 1 \ NN(y) ==> y can be well-ordered *) +(* ********************************************************************** *) + +(* ********************************************************************** *) +(* WO6 ==> NN(y) \ 0 *) +(* ********************************************************************** *) + +lemma WO6_imp_NN_not_empty: "WO6 ==> NN(y) \ 0" +apply (unfold WO6_def NN_def, clarify) +apply blast +done + +(* ********************************************************************** *) +(* 1 \ NN(y) ==> y can be well-ordered *) +(* ********************************************************************** *) + +lemma lemma1: + "[| (\b y; \b 1; Ord(a) |] ==> \cb y; \b 1; Ord(a) |] + ==> f` (LEAST i. f`i = {x}) = {x}" +apply (drule lemma1, assumption+) +apply (fast elim!: lt_Ord intro: LeastI) +done - ww2_def "ww2(f,b,g,s) == f`g - vv2(f,b,g,s)" +lemma NN_imp_ex_inj: "1 \ NN(y) ==> \a f. Ord(a) & f \ inj(y, a)" +apply (unfold NN_def) +apply (elim CollectE exE conjE) +apply (rule_tac x = "a" in exI) +apply (rule_tac x = "\x \ y. LEAST i. f`i = {x}" in exI) +apply (rule conjI, assumption) +apply (rule_tac d = "%i. THE x. x \ f`i" in lam_injective) +apply (drule lemma1, assumption+) +apply (fast elim!: Least_le [THEN lt_trans1, THEN ltD] lt_Ord) +apply (rule lemma2 [THEN ssubst], assumption+, blast) +done + +lemma y_well_ord: "[| y*y \ y; 1 \ NN(y) |] ==> \r. well_ord(y, r)" +apply (drule NN_imp_ex_inj) +apply (fast elim!: well_ord_rvimage [OF _ well_ord_Memrel]) +done + +(* ********************************************************************** *) +(* reverse induction which lets us infer that 1 \ NN(y) *) +(* ********************************************************************** *) + +lemma rev_induct_lemma [rule_format]: + "[| n \ nat; !!m. [| m \ nat; m\0; P(succ(m)) |] ==> P(m) |] + ==> n\0 --> P(n) --> P(1)" +by (erule nat_induct, blast+) + +lemma rev_induct: + "[| n \ nat; P(n); n\0; + !!m. [| m \ nat; m\0; P(succ(m)) |] ==> P(m) |] + ==> P(1)" +by (rule rev_induct_lemma, blast+) - gg2_def "gg2(f,a,b,s) == - \\g \\ a++a. if g NN(y) ==> n \ nat" +by (simp add: NN_def) + +lemma lemma3: "[| n \ NN(y); y*y \ y; n\0 |] ==> 1 \ NN(y)" +apply (rule rev_induct [OF NN_into_nat], assumption+) +apply (rule lemma_ii, assumption+) +done + +(* ********************************************************************** *) +(* Main theorem "WO6 ==> WO1" *) +(* ********************************************************************** *) + +(* another helpful lemma *) +lemma NN_y_0: "0 \ NN(y) ==> y=0" +apply (unfold NN_def) +apply (fast intro!: equalityI dest!: lepoll_0_is_0 elim: subst) +done + +lemma WO6_imp_WO1: "WO6 ==> WO1" +apply (unfold WO1_def) +apply (rule allI) +apply (case_tac "A=0") +apply (fast intro!: well_ord_Memrel nat_0I [THEN nat_into_Ord]) +apply (rule_tac x1 = "A" in lemma_iv [THEN revcut_rl]) +apply (erule exE) +apply (drule WO6_imp_NN_not_empty) +apply (erule Un_subset_iff [THEN iffD1, THEN conjE]) +apply (erule_tac A = "NN (y) " in not_emptyE) +apply (frule y_well_ord) + apply (fast intro!: lemma3 dest!: NN_y_0 elim!: not_emptyE) +apply (fast elim: well_ord_subset) +done + end diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/WO_AC.ML --- a/src/ZF/AC/WO_AC.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,20 +0,0 @@ -(* Title: ZF/AC/WO_AC.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -Lemmas used in the proofs like WO? ==> AC? -*) - -Goal "[| well_ord(Union(A),r); 0\\A; B \\ A |] \ -\ ==> (THE b. first(b,B,r)) \\ B"; -by (fast_tac (claset() addSEs [well_ord_imp_ex1_first RS theI RS - (first_def RS def_imp_iff RS iffD1 RS conjunct1)]) 1); -qed "first_in_B"; - -Goal "[| well_ord(Union(A), R); 0\\A |] ==> \\f. f:(\\X \\ A. X)"; -by (fast_tac (claset() addSEs [first_in_B] addSIs [lam_type]) 1); -qed "ex_choice_fun"; - -Goal "well_ord(A, R) ==> \\f. f:(\\X \\ Pow(A)-{0}. X)"; -by (fast_tac (claset() addSEs [well_ord_subset RS ex_choice_fun]) 1); -qed "ex_choice_fun_Pow"; diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/WO_AC.thy --- a/src/ZF/AC/WO_AC.thy Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,3 +0,0 @@ -(*Dummy theory to document dependencies *) - -WO_AC = Order diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/recfunAC16.ML --- a/src/ZF/AC/recfunAC16.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,65 +0,0 @@ -(* Title: ZF/AC/recfunAC16.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -Properties of the recursive definition used in the proof of WO2 ==> AC16 -*) - -(* ********************************************************************** *) -(* Basic properties of recfunAC16 *) -(* ********************************************************************** *) - -Goalw [recfunAC16_def] "recfunAC16(f,fa,0,a) = 0"; -by (rtac transrec2_0 1); -qed "recfunAC16_0"; - -Goalw [recfunAC16_def] - "recfunAC16(f,fa,succ(i),a) = \ -\ (if (\\y \\ recfunAC16(f,fa,i,a). fa ` i \\ y) then recfunAC16(f,fa,i,a) \ -\ else recfunAC16(f,fa,i,a) Un \ -\ {f ` (LEAST j. fa ` i \\ f ` j & \ -\ (\\b f`j \ -\ --> (\\t \\ recfunAC16(f,fa,i,a). ~ fa`b \\ t))))})"; -by (rtac transrec2_succ 1); -qed "recfunAC16_succ"; - -Goalw [recfunAC16_def] "Limit(i) \ -\ ==> recfunAC16(f,fa,i,a) = (\\j B(g,r); Ord(i) |] \ -\ ==> j transrec2(j, 0, B) \\ transrec2(i, 0, B)"; -by (resolve_tac [prem2 RS trans_induct] 1); -by (rtac Ord_cases 1 THEN (REPEAT (assume_tac 1))); -by (Fast_tac 1); -by (Asm_simp_tac 1); -by (fast_tac (FOL_cs addSIs [succI1, prem1] - addSEs [ballE, leE, prem1 RSN (2, subset_trans)]) 1); -by (fast_tac (claset() addIs [OUN_I, ltI] - addSEs [Limit_has_succ RS ltE, succI1 RSN (2, Ord_in_Ord) RS le_refl, - transrec2_Limit RS ssubst]) 1); -qed "transrec2_mono_lemma"; - -val [prem1, prem2] = goal thy "[| !!g r. r \\ B(g,r); j le i |] \ -\ ==> transrec2(j, 0, B) \\ transrec2(i, 0, B)"; -by (resolve_tac [prem2 RS leE] 1); -by (resolve_tac [transrec2_mono_lemma RS impE] 1); -by (TRYALL (fast_tac (claset() addSIs [prem1, prem2, lt_Ord2]))); -qed "transrec2_mono"; - -(* ********************************************************************** *) -(* Monotonicity of recfunAC16 *) -(* ********************************************************************** *) - -Goalw [recfunAC16_def] - "!!i. i le j ==> recfunAC16(f, g, i, a) \\ recfunAC16(f, g, j, a)"; -by (rtac transrec2_mono 1); -by (REPEAT (fast_tac (claset() addIs [split_if RS iffD2]) 1)); -qed "recfunAC16_mono"; - diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/recfunAC16.thy --- a/src/ZF/AC/recfunAC16.thy Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,19 +0,0 @@ -(* Title: ZF/AC/recfunAC16.thy - ID: $Id$ - Author: Krzysztof Grabczewski - -A recursive definition used in the proof of WO2 ==> AC16 -*) - -recfunAC16 = Cardinal + Epsilon + - -constdefs - recfunAC16 :: [i, i, i, i] => i - - "recfunAC16(f,fa,i,a) == - transrec2(i, 0, - %g r. if(\\y \\ r. fa`g \\ y, r, - r Un {f`(LEAST i. fa`g \\ f`i & - (\\b f`i --> (\\t \\ r. ~ fa`b \\ t))))}))" - -end diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/rel_is_fun.ML --- a/src/ZF/AC/rel_is_fun.ML Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,42 +0,0 @@ -(* Title: ZF/AC/rel_is_fun.ML - ID: $Id$ - Author: Krzysztof Grabczewski - -Lemmas needed to state when a finite relation is a function. - -The criteria are cardinalities of the relation and its domain. -Used in WO6WO1.ML -*) - -(*Using AC we could trivially prove, for all u, domain(u) lepoll u*) -goalw Cardinal.thy [lepoll_def] - "!!m. [| m \\ nat; u lepoll m |] ==> domain(u) lepoll m"; -by (etac exE 1); -by (res_inst_tac [("x", - "\\x \\ domain(u). LEAST i. \\y. \\ u & f` = i")] exI 1); -by (res_inst_tac [("d","%y. fst(converse(f)`y)")] lam_injective 1); -by (fast_tac (claset() addIs [LeastI2, nat_into_Ord RS Ord_in_Ord, - inj_is_fun RS apply_type]) 1); -by (etac domainE 1); -by (forward_tac [inj_is_fun RS apply_type] 1 THEN (atac 1)); -by (rtac LeastI2 1); -by (auto_tac (claset() addSEs [nat_into_Ord RS Ord_in_Ord], - simpset())); -qed "lepoll_m_imp_domain_lepoll_m"; - -goalw Cardinal.thy [function_def] - "!!r. [| succ(m) lepoll domain(r); r lepoll succ(m); m \\ nat |] ==> \ -\ function(r)"; -by Safe_tac; -by (resolve_tac [excluded_middle RS disjE] 1 THEN (atac 2)); -by (fast_tac (claset() addSEs [lepoll_trans RS succ_lepoll_natE, - Diff_sing_lepoll RSN (2, lepoll_m_imp_domain_lepoll_m)] - addEs [not_sym RSN (2, domain_Diff_eq) RS subst]) 1); -qed "rel_domain_ex1"; - -goal Cardinal.thy - "!!r. [| succ(m) lepoll domain(r); r lepoll succ(m); m \\ nat; \ -\ r \\ A*B; A=domain(r) |] ==> r \\ A->B"; -by (hyp_subst_tac 1); -by (asm_simp_tac (simpset() addsimps [Pi_iff, rel_domain_ex1]) 1); -qed "rel_is_fun"; diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/AC/rel_is_fun.thy --- a/src/ZF/AC/rel_is_fun.thy Wed Jan 16 15:04:37 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1 +0,0 @@ -rel_is_fun = Cardinal diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/CardinalArith.thy --- a/src/ZF/CardinalArith.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/CardinalArith.thy Wed Jan 16 17:52:06 2002 +0100 @@ -102,4 +102,60 @@ "(!!x. x:A ==> Card(K(x))) ==> Card(UN x nat ==> n \ nat" +apply (unfold lesspoll_def) +apply (rule conjI) +apply (erule OrdmemD [THEN subset_imp_lepoll], rule Ord_nat) +apply (rule notI) +apply (erule eqpollE) +apply (rule succ_lepoll_natE) +apply (blast intro: nat_succI [THEN OrdmemD, THEN subset_imp_lepoll] + lepoll_trans, assumption); +done + +lemma in_Card_imp_lesspoll: "[| Card(K); b \ K |] ==> b \ K" +apply (unfold lesspoll_def) +apply (simp add: Card_iff_initial) +apply (fast intro!: le_imp_lepoll ltI leI) +done + +lemma succ_lepoll_imp_not_empty: "succ(x) \ y ==> y \ 0" +by (fast dest!: lepoll_0_is_0) + +lemma eqpoll_succ_imp_not_empty: "x \ succ(n) ==> x \ 0" +by (fast elim!: eqpoll_sym [THEN eqpoll_0_is_0, THEN succ_neq_0]) + +lemma Finite_Fin_lemma [rule_format]: + "n \ nat ==> \A. (A\n & A \ X) --> A \ Fin(X)" +apply (induct_tac "n") +apply (rule allI) +apply (fast intro!: Fin.emptyI dest!: eqpoll_imp_lepoll [THEN lepoll_0_is_0]) +apply (rule allI) +apply (rule impI) +apply (erule conjE) +apply (rule eqpoll_succ_imp_not_empty [THEN not_emptyE], (assumption)) +apply (frule Diff_sing_eqpoll, (assumption)) +apply (erule allE) +apply (erule impE, fast) +apply (drule subsetD, (assumption)) +apply (drule Fin.consI, (assumption)) +apply (simp add: cons_Diff) +done + +lemma Finite_Fin: "[| Finite(A); A \ X |] ==> A \ Fin(X)" +by (unfold Finite_def, blast intro: Finite_Fin_lemma) + +lemma lesspoll_lemma: + "[| ~ A \ B; C \ B |] ==> A - C \ 0" +apply (unfold lesspoll_def) +apply (fast dest!: Diff_eq_0_iff [THEN iffD1, THEN subset_imp_lepoll] + intro!: eqpollI elim: notE + elim!: eqpollE lepoll_trans) +done + +lemma eqpoll_imp_Finite_iff: "A \ B ==> Finite(A) <-> Finite(B)" +apply (unfold Finite_def) +apply (blast intro: eqpoll_trans eqpoll_sym) +done + end diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/IsaMakefile --- a/src/ZF/IsaMakefile Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/IsaMakefile Wed Jan 16 17:52:06 2002 +0100 @@ -58,16 +58,11 @@ ZF-AC: ZF $(LOG)/ZF-AC.gz $(LOG)/ZF-AC.gz: $(OUT)/ZF \ - AC/AC15_WO6.ML AC/AC15_WO6.thy AC/AC16_WO4.ML AC/AC16_WO4.thy \ - AC/AC16_lemmas.ML AC/AC16_lemmas.thy AC/AC17_AC1.ML AC/AC17_AC1.thy \ - AC/AC18_AC19.ML AC/AC18_AC19.thy AC/AC1_WO2.ML \ - AC/AC1_WO2.thy AC/AC7_AC9.ML AC/AC_Equiv.ML \ - AC/AC_Equiv.thy AC/Cardinal_aux.ML AC/Cardinal_aux.thy AC/DC.ML \ - AC/DC.thy AC/DC_lemmas.ML AC/DC_lemmas.thy AC/HH.ML AC/HH.thy \ - AC/Hartog.ML AC/Hartog.thy AC/ROOT.ML AC/WO1_AC.ML AC/WO1_AC.thy \ - AC/WO1_WO7.ML AC/WO1_WO7.thy AC/WO2_AC16.ML \ - AC/WO2_AC16.thy AC/WO6_WO1.ML AC/WO6_WO1.thy AC/WO_AC.ML AC/WO_AC.thy \ - AC/recfunAC16.ML AC/recfunAC16.thy AC/rel_is_fun.ML AC/rel_is_fun.thy + AC/ROOT.ML AC/AC15_WO6.thy AC/AC16_WO4.thy \ + AC/AC16_lemmas.thy AC/AC17_AC1.thy AC/AC18_AC19.thy AC/AC7_AC9.thy \ + AC/AC_Equiv.thy AC/Cardinal_aux.thy \ + AC/DC.thy AC/HH.thy AC/Hartog.thy AC/WO1_AC.thy AC/WO1_WO7.thy \ + AC/WO2_AC16.thy AC/WO6_WO1.thy @$(ISATOOL) usedir $(OUT)/ZF AC diff -r 1748c16c2df3 -r 249600a63ba9 src/ZF/Main.thy --- a/src/ZF/Main.thy Wed Jan 16 15:04:37 2002 +0100 +++ b/src/ZF/Main.thy Wed Jan 16 17:52:06 2002 +0100 @@ -57,4 +57,15 @@ (Limit(K) --> f(K) = (UN j K \ K \ K" +apply (rule well_ord_InfCard_square_eq); + apply (erule InfCard_is_Card [THEN Card_is_Ord, THEN well_ord_Memrel]); +apply (simp add: InfCard_is_Card [THEN Card_cardinal_eq]); +done + +lemma Inf_Card_is_InfCard: "[| ~Finite(i); Card(i) |] ==> InfCard(i)" +by (simp add: InfCard_def Card_is_Ord [THEN nat_le_infinite_Ord]) + end