(* Title: ZF/AC/WO2_AC16.thy 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 imports AC_Equiv AC16_lemmas Cardinal_aux begin
(**** A recursive definition used in the proof of WO2 ==> AC16 ****)
definition
recfunAC16 :: "[i,i,i,i] => i"where "recfunAC16(f,h,i,a) ==
transrec2(i, 0,
%g r. if (\<exists>y \<in> r. h`g \<subseteq> y) then r
else r \<union> {f`(\<mu> i. h`g \<subseteq> f`i &
(\<forall>b<a. (h`b \<subseteq> f`i \<longrightarrow> (\<forall>t \<in> r. ~ h`b \<subseteq> t))))})"
(* ********************************************************************** *) (* case of limit ordinal *) (* ********************************************************************** *)
lemma lemma3_1: "[| \yzY \ F(y). f(z)<=Y) \ (\! Y. Y \ F(y) & f(z)<=Y); \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j); j\<le>i; i<x; z<a;
V \<in> F(i); f(z)<=V; W \<in> F(j); f(z)<=W |]
==> V = W" apply (erule asm_rl allE impE)+ apply (drule subsetD, assumption, blast) done
lemma lemma3: "[| \yzY \ F(y). f(z)<=Y) \ (\! Y. Y \ F(y) & f(z)<=Y); \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j); i<x; j<x; z<a;
V \<in> F(i); f(z)<=V; W \<in> 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 &
(\<forall>x<a. x < y | (\<exists>Y \<in> F(y). h(x) \<subseteq> Y) \<longrightarrow>
(\<exists>! Y. Y \<in> F(y) & h(x) \<subseteq> Y));
x < a |]
==> \<forall>y<x. \<forall>z<a. z < y | (\<exists>Y \<in> F(y). h(z) \<subseteq> Y) \<longrightarrow>
(\<exists>! Y. Y \<in> F(y) & h(z) \<subseteq> Y)" apply (intro oallI impI) apply (drule ospec, assumption, clarify) apply (blast elim!: oallE ) done
lemma lemma5: "[| \y X &
(\<forall>x<a. x < y | (\<exists>Y \<in> F(y). h(x) \<subseteq> Y) \<longrightarrow>
(\<exists>! Y. Y \<in> F(y) & h(x) \<subseteq> Y));
x < a; Limit(x); \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j) |]
==> (\<Union>x<x. F(x)) \<subseteq> X &
(\<forall>xa<a. xa < x | (\<exists>x \<in> \<Union>x<x. F(x). h(xa) \<subseteq> x) \<longrightarrow> (\<exists>! Y. Y \<in> (\<Union>x<x. F(x)) & h(xa) \<subseteq> Y))" apply (rule conjI) apply (rule subsetI) apply (erule OUN_E) apply (drule ospec, assumption, 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).
*)
lemma in_Least_Diff: "[| z \ F(x); Ord(x) |]
==> z \<in> F(\<mu> i. z \<in> F(i)) - (\<Union>j<(\<mu> i. z \<in> F(i)). F(j))" by (fast elim: less_LeastE elim!: LeastI)
lemma Least_eq_imp_ex: "[| (\ i. w \ F(i)) = (\ i. z \ F(i));
w \<in> (\<Union>i<a. F(i)); z \<in> (\<Union>i<a. F(i)) |]
==> \<exists>b<a. w \<in> (F(b) - (\<Union>c<b. F(c))) & z \<in> (F(b) - (\<Union>c<b. F(c)))" apply (elim OUN_E) apply (drule in_Least_Diff, erule lt_Ord) apply (frule in_Least_Diff, erule lt_Ord) apply (rule oexI, force) apply (blast intro: lt_Ord Least_le [THEN lt_trans1]) done
lemma two_in_lepoll_1: "[| A \ 1; a \ A; b \ A |] ==> a=b" by (fast dest!: lepoll_1_is_sing)