diff --git a/thys/Independence_CH/Absolute_Versions.thy b/thys/Independence_CH/Absolute_Versions.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Absolute_Versions.thy @@ -0,0 +1,148 @@ +section\From $M$ to $\calV$\ + +theory Absolute_Versions + imports + CH + ZF.Cardinal_AC +begin + +subsection\Locales of a class \<^term>\M\ hold in \<^term>\\\\ + +interpretation V: M_trivial \ + using Union_ax_absolute upair_ax_absolute + by unfold_locales auto + +lemmas bad_simps = V.nonempty V.Forall_in_M_iff V.Inl_in_M_iff V.Inr_in_M_iff + V.succ_in_M_iff V.singleton_in_M_iff V.Equal_in_M_iff V.Member_in_M_iff V.Nand_in_M_iff + V.Cons_in_M_iff V.pair_in_M_iff V.upair_in_M_iff + +lemmas bad_M_trivial_simps[simp del] = V.Forall_in_M_iff V.Equal_in_M_iff + V.nonempty + +lemmas bad_M_trivial_rules[rule del] = V.pair_in_MI V.singleton_in_MI V.pair_in_MD V.nat_into_M + V.depth_closed V.length_closed V.nat_case_closed V.separation_closed + V.Un_closed V.strong_replacement_closed V.nonempty + +interpretation V:M_basic \ + using power_ax_absolute separation_absolute replacement_absolute + by unfold_locales auto + +interpretation V:M_eclose \ + by unfold_locales (auto intro:separation_absolute replacement_absolute + simp:iterates_replacement_def wfrec_replacement_def) + +lemmas bad_M_basic_rules[simp del, rule del] = + V.cartprod_closed V.finite_funspace_closed V.converse_closed + V.list_case'_closed V.pred_closed + +interpretation V:M_cardinal_arith \ + by unfold_locales (auto intro:separation_absolute replacement_absolute + simp add:iterates_replacement_def wfrec_replacement_def lam_replacement_def) + +lemmas bad_M_cardinals_rules[simp del, rule del] = + V.iterates_closed V.M_nat V.trancl_closed V.rvimage_closed + +interpretation V:M_cardinal_arith_jump \ + by unfold_locales (auto intro:separation_absolute replacement_absolute + simp:wfrec_replacement_def) + +lemma choice_ax_Universe: "choice_ax(\)" +proof - + { + fix x + obtain f where "f \ surj(|x|,x)" + using cardinal_eqpoll unfolding eqpoll_def bij_def by fast + moreover + have "Ord(|x|)" by simp + ultimately + have "\a. Ord(a) \ (\f. f \ surj(a,x))" + by fast + } + then + show ?thesis unfolding choice_ax_def rall_def rex_def + by simp +qed + +interpretation V:M_master \ + using choice_ax_Universe + by unfold_locales (auto intro:separation_absolute replacement_absolute + simp:lam_replacement_def transrec_replacement_def wfrec_replacement_def + is_wfrec_def M_is_recfun_def) + +named_theorems V_simps + +\ \To work systematically, ASCII versions of "\_absolute" theorems as + those below are preferable.\ +lemma eqpoll_rel_absolute[V_simps]: "x \\<^bsup>\\<^esup> y \ x \ y" + unfolding eqpoll_def using V.def_eqpoll_rel by auto + +lemma cardinal_rel_absolute[V_simps]: "|x|\<^bsup>\\<^esup> = |x|" + unfolding cardinal_def cardinal_rel_def by (simp add:V_simps) + +lemma Card_rel_absolute[V_simps]:"Card\<^bsup>\\<^esup>(a) \ Card(a)" + unfolding Card_rel_def Card_def by (simp add:V_simps) + +lemma csucc_rel_absolute[V_simps]:"(a\<^sup>+)\<^bsup>\\<^esup> = a\<^sup>+" + unfolding csucc_rel_def csucc_def by (simp add:V_simps) + +lemma function_space_rel_absolute[V_simps]:"x \\<^bsup>\\<^esup> y = x \ y" + using V.function_space_rel_char by (simp add:V_simps) + +lemma cexp_rel_absolute[V_simps]:"x\<^bsup>\y,\\<^esup> = x\<^bsup>\y\<^esup>" + unfolding cexp_rel_def cexp_def by (simp add:V_simps) + +lemma HAleph_rel_absolute[V_simps]:"HAleph_rel(\,a,b) = HAleph(a,b)" + unfolding HAleph_rel_def HAleph_def by (auto simp add:V_simps) + +lemma Aleph_rel_absolute[V_simps]: "Ord(x) \ \\<^bsub>x\<^esub>\<^bsup>\\<^esup> = \\<^bsub>x\<^esub>" +proof - + assume "Ord(x)" + have "\\<^bsub>x\<^esub>\<^bsup>\\<^esup> = transrec(x, \a b. HAleph_rel(\,a,b))" + unfolding Aleph_rel_def by simp + also + have "\ = transrec(x, HAleph)" + by (simp add:V_simps) + also from \Ord(x)\ + have "\ = \\<^bsub>x\<^esub>" + using Aleph'_eq_Aleph unfolding Aleph'_def by simp + finally + show ?thesis . +qed + +txt\Example of absolute lemmas obtained from the relative versions. + Note the \<^emph>\only\ declarations\ +lemma Ord_cardinal_idem': "Ord(A) \ ||A|| = |A|" + using V.Ord_cardinal_rel_idem by (simp only:V_simps) + +lemma Aleph_succ': "Ord(\) \ \\<^bsub>succ(\)\<^esub> = \\<^bsub>\\<^esub>\<^sup>+" + using V.Aleph_rel_succ by (simp only:V_simps) + +txt\These two results are new, first obtained in relative form + (not ported).\ +lemma csucc_cardinal: + assumes "Ord(\)" shows "|\|\<^sup>+ = \\<^sup>+" + using assms V.csucc_rel_cardinal_rel by (simp add:V_simps) + +lemma csucc_le_mono: + assumes "\ \ \" shows "\\<^sup>+ \ \\<^sup>+" + using assms V.csucc_rel_le_mono by (simp add:V_simps) + +txt\Example of transferring results from a transitive model to \<^term>\\\\ +lemma (in M_Perm) eqpoll_rel_transfer_absolute: + assumes "M(A)" "M(B)" "A \\<^bsup>M\<^esup> B" + shows "A \ B" +proof - + interpret M_N_Perm M \ + by (unfold_locales, simp only:V_simps) + from assms + show ?thesis using eqpoll_rel_transfer + by (simp only:V_simps) +qed + +txt\The “relationalized” $\CH$ with respect to \<^term>\\\ corresponds + to the real $\CH$.\ +lemma is_ContHyp_iff_CH: "is_ContHyp(\) \ ContHyp" + using V.is_ContHyp_iff + by (auto simp add:ContHyp_rel_def ContHyp_def V_simps) + +end \ No newline at end of file diff --git a/thys/Independence_CH/CH.thy b/thys/Independence_CH/CH.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/CH.thy @@ -0,0 +1,529 @@ +section\Forcing extension satisfying the Continuum Hypothesis\ + +theory CH + imports + Kappa_Closed_Notions + Cohen_Posets_Relative +begin + +context M_ctm3_AC +begin + +declare Fn_rel_closed[simp del, rule del, simplified setclass_iff, simp, intro] +declare Fnle_rel_closed[simp del, rule del, simplified setclass_iff, simp, intro] + +abbreviation + Coll :: "i" where + "Coll \ Fn\<^bsup>M\<^esup>(\\<^bsub>1\<^esub>\<^bsup>M\<^esup>, \\<^bsub>1\<^esub>\<^bsup>M\<^esup>, \ \\<^bsup>M\<^esup> 2)" + +abbreviation + Colleq :: "i" where + "Colleq \ Fnle\<^bsup>M\<^esup>(\\<^bsub>1\<^esub>\<^bsup>M\<^esup>, \\<^bsub>1\<^esub>\<^bsup>M\<^esup>, \ \\<^bsup>M\<^esup> 2)" + +lemma Coll_in_M[intro,simp]: "Coll \ M" by simp + +lemma Colleq_refl : "refl(Coll,Colleq)" + unfolding Fnle_rel_def Fnlerel_def refl_def + using RrelI by simp + +subsection\Collapse forcing is sufficiently closed\ + +\ \Kunen IV.7.14, only for \<^term>\\\<^bsub>1\<^esub>\\ +lemma succ_omega_closed_Coll: "succ(\)-closed\<^bsup>M\<^esup>(Coll,Colleq)" +proof - + \ \Case for finite sequences\ + have "n\\ \ \f \ n \<^sub><\\<^bsup>M\<^esup> (Coll,converse(Colleq)). + \q\M. q \ Coll \ (\\\M. \ \ n \ \q, f ` \\ \ Colleq)" for n + proof (induct rule:nat_induct) + case 0 + then + show ?case + using zero_lt_Aleph_rel1 zero_in_Fn_rel + by (auto simp del:setclass_iff) (rule bexI[OF _ zero_in_M], auto) + next + case (succ x) + then + have "\f\succ(x) \<^sub><\\<^bsup>M\<^esup> (Coll,converse(Colleq)). \\ \ succ(x). \f`x, f ` \\ \ Colleq" + proof(intro ballI) + fix f \ + assume "f\succ(x) \<^sub><\\<^bsup>M\<^esup> (Coll,converse(Colleq))" "\\succ(x)" + moreover from \x\\\ this + have "f\succ(x) \<^sub><\ (Coll,converse(Colleq))" + using mono_seqspace_rel_char nat_into_M + by simp + moreover from calculation succ + consider "\\x" | "\=x" + by auto + then + show "\f`x, f ` \\ \ Colleq" + proof(cases) + case 1 + then + have "\\, x\ \ Memrel(succ(x))" "x\succ(x)" "\\succ(x)" + by auto + with \f\succ(x) \<^sub><\ (Coll,converse(Colleq))\ + show ?thesis + using mono_mapD(2)[OF _ \\\succ(x)\ _ \\\, x\ \ Memrel(succ(x))\] + unfolding mono_seqspace_def + by auto + next + case 2 + with \f\succ(x) \<^sub><\ (Coll,converse(Colleq))\ + show ?thesis + using Colleq_refl mono_seqspace_is_fun[THEN apply_type] + unfolding refl_def + by simp + qed + qed + moreover + note \x\\\ + moreover from this + have "f`x \ Coll" if "f: succ(x) \<^sub><\\<^bsup>M\<^esup> (Coll,converse(Colleq))" for f + using that mono_seqspace_rel_char[simplified, of "succ(x)" Coll "converse(Colleq)"] + nat_into_M[simplified] mono_seqspace_is_fun[of "converse(Colleq)"] + by (intro apply_type[of _ "succ(x)"]) (auto simp del:setclass_iff) + ultimately + show ?case + using transM[of _ Coll] + by (auto dest:transM simp del:setclass_iff, rule_tac x="f`x" in bexI) + (auto simp del:setclass_iff, simp) + qed + moreover + \ \Interesting case: Countably infinite sequences.\ + have "\f\M. f \ \ \<^sub><\\<^bsup>M\<^esup> (Coll,converse(Colleq)) \ + (\q\M. q \ Coll \ (\\\M. \ \ \ \ \q, f ` \\ \ Colleq))" + proof(intro ballI impI) + fix f + let ?G="f``\" + assume "f\M" "f \ \ \<^sub><\\<^bsup>M\<^esup> (Coll,converse(Colleq))" + moreover from this + have "f\\ \<^sub><\ (Coll,converse(Colleq))" "f\\ \ Coll" + using mono_seqspace_rel_char mono_mapD(2) nat_in_M + by auto + moreover from this + have "countable\<^bsup>M\<^esup>(f`x)" if "x\\" for x + using that Fn_rel_is_function countable_iff_lesspoll_rel_Aleph_rel_one + by auto + moreover from calculation + have "?G \ M" "f\\\Coll" + using nat_in_M image_closed Pi_iff + by simp_all + moreover from calculation + have 1:"\d\?G. d \ h \ d \ g" if "h \ ?G" "g \ ?G" for h g + proof - + from calculation + have "?G={f`x . x\\}" + using image_function[of f \] Pi_iff domain_of_fun + by auto + from \?G=_\ that + obtain m n where eq:"h=f`m" "g=f`n" "n\\" "m\\" + by auto + then + have "m\n\\" "m\m\n" "n\m\n" + using Un_upper1_le Un_upper2_le nat_into_Ord by simp_all + with calculation eq \?G=_\ + have "f`(m\n)\?G" "f`(m\n) \ h" "f`(m\n) \ g" + using Fnle_relD mono_map_lt_le_is_mono converse_refl[OF Colleq_refl] + by auto + then + show ?thesis + by auto + qed + moreover from calculation + have "?G \ (\\<^bsub>1\<^esub>\<^bsup>M\<^esup> \\<^bsup>##M\<^esup> (nat \\<^bsup>M\<^esup> 2))" + using subset_trans[OF image_subset[OF \f\_\,of \] Fn_rel_subset_PFun_rel] + by simp + moreover + have "\ ?G \ (\\<^bsub>1\<^esub>\<^bsup>M\<^esup> \\<^bsup>##M\<^esup> (nat \\<^bsup>M\<^esup> 2))" + using pfun_Un_filter_closed'[OF \?G\_\ 1] \?G\M\ + by simp + moreover from calculation + have "\?G \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" + using countable_fun_imp_countable_image[of f] + mem_function_space_rel_abs[simplified,OF nat_in_M Coll_in_M \f\M\] + countableI[OF lepoll_rel_refl] + countable_iff_lesspoll_rel_Aleph_rel_one[of "\?G"] + by auto + moreover from calculation + have "\?G\Coll" + unfolding Fn_rel_def + by simp + moreover from calculation + have "\?G \ f ` \ " if "\\\" for \ + using that image_function[OF fun_is_function] domain_of_fun + by auto + ultimately + show "\q\M. q \ Coll \ (\\\M. \ \ \ \ \q, f ` \\ \ Colleq)" + using Fn_rel_is_function Fnle_relI + by auto + qed + ultimately + show ?thesis + unfolding kappa_closed_rel_def by (auto elim!:leE dest:ltD) +qed + +end \ \\<^locale>\M_ctm3_AC\\ + +locale collapse_generic4 = G_generic4_AC "Fn\<^bsup>M\<^esup>(\\<^bsub>1\<^esub>\<^bsup>##M\<^esup>, \\<^bsub>1\<^esub>\<^bsup>M\<^esup>, \ \\<^bsup>M\<^esup> 2)" "Fnle\<^bsup>M\<^esup>(\\<^bsub>1\<^esub>\<^bsup>##M\<^esup>, \\<^bsub>1\<^esub>\<^bsup>M\<^esup>, \ \\<^bsup>M\<^esup> 2)" 0 + +sublocale collapse_generic4 \ forcing_notion "Coll" "Colleq" 0 + using zero_lt_Aleph_rel1 by unfold_locales + +context collapse_generic4 +begin + +notation Leq (infixl "\" 50) +notation Incompatible (infixl "\" 50) +notation GenExt_at_P ("_[_]" [71,1]) + +abbreviation + f_G :: "i" (\f\<^bsub>G\<^esub>\) where + "f\<^bsub>G\<^esub> \ \G" + +lemma f_G_in_MG[simp]: + shows "f\<^bsub>G\<^esub> \ M[G]" + using G_in_MG by simp + +abbreviation + dom_dense :: "i\i" where + "dom_dense(x) \ { p\Coll . x \ domain(p) }" + +lemma Coll_into_countable_rel: "p \ Coll \ countable\<^bsup>M\<^esup>(p)" +proof - + assume "p\Coll" + then + have "p \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "p\M" + using Fn_rel_is_function by simp_all + moreover from this + have "p \\<^bsup>M\<^esup> \" + using lesspoll_rel_Aleph_rel_succ[of 0] Aleph_rel_zero + by simp + ultimately + show ?thesis + using countableI eqpoll_rel_imp_lepoll_rel eqpoll_rel_sym cardinal_rel_eqpoll_rel + by simp +qed + +(* TODO: Should be more general, cf. @{thm add_generic.dense_dom_dense} *) +lemma dense_dom_dense: "x \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ dense(dom_dense(x))" +proof + fix p + assume "x \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "p \ Coll" + show "\d\dom_dense(x). d \ p" + proof (cases "x \ domain(p)") + case True + with \x \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ \p \ Coll\ + show ?thesis using refl_leq by auto + next + case False + note \p \ Coll\ + moreover from this and False and \x \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ + have "cons(\x,\n\\. 0\, p) \ Coll" "x\M" + using function_space_rel_char + function_space_rel_closed lam_replacement_constant + lam_replacement_iff_lam_closed InfCard_rel_Aleph_rel + by (auto intro!: cons_in_Fn_rel dest:transM intro:function_space_nonempty) + ultimately + show ?thesis + using Fn_relD by blast + qed +qed + +lemma dom_dense_closed[intro,simp]: "x\M \ dom_dense(x) \ M" + using separation_in_domain[of x] + by simp + +lemma domain_f_G: assumes "x \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" + shows "x \ domain(f\<^bsub>G\<^esub>)" +proof - + from assms + have "dense(dom_dense(x))" "x\M" + using dense_dom_dense transitivity[OF _ + Aleph_rel_closed[of 1,THEN setclass_iff[THEN iffD1]]] + by simp_all + with assms + obtain p where "p\dom_dense(x)" "p\G" + using generic[THEN M_generic_denseD, of "dom_dense(x)"] + by auto + then + show "x \ domain(f\<^bsub>G\<^esub>)" by blast +qed + +lemma rex_mono : assumes "\ d \ A . P(d)" "A\B" + shows "\ d \ B. P(d)" + using assms by auto + +lemma Un_filter_is_function: + assumes "filter(G)" + shows "function(\G)" +proof - + have "Coll \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \\<^bsup>##M\<^esup> (\ \\<^bsup>M\<^esup> 2)" + using Fn_rel_subset_PFun_rel + by simp + moreover + have "\ d \ Coll. d \ f \ d \ g" if "f\G" "g\G" for f g + using filter_imp_compat[OF assms \f\G\ \g\G\] filterD[OF assms] + unfolding compat_def compat_in_def + by auto + ultimately + have "\d \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \\<^bsup>##M\<^esup> (\ \\<^bsup>M\<^esup> 2). d \ f \ d \ g" if "f\G" "g\G" for f g + using rex_mono[of Coll] that by simp + moreover + have "G\Coll" + using assms + unfolding filter_def + by simp + moreover from this + have "G \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \\<^bsup>##M\<^esup> (\ \\<^bsup>M\<^esup> 2)" + using assms unfolding Fn_rel_def + by auto + ultimately + show ?thesis + using pfun_Un_filter_closed[of G] + by simp +qed + +lemma f_G_funtype: + shows "f\<^bsub>G\<^esub> : \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ \ \\<^bsup>M[G]\<^esup> 2" +proof - + have "x \ B \ B \ G \ x \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ (\ \\<^bsup>M[G]\<^esup> 2)" for B x + proof - + assume "x\B" "B\G" + moreover from this + have "x \ M[G]" + by (auto dest!:generic_dests ext.transM) + (intro generic_simps(2)[of Coll], simp) + moreover from calculation + have "x \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ (\ \ 2)" + using Fn_rel_subset_Pow[of "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "\ \\<^bsup>M\<^esup> 2", + OF _ _ function_space_rel_closed] function_space_rel_char + by (auto dest!:generic_dests) + moreover from this + obtain z w where "x=\z,w\" "z\\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "w:\ \ 2" by auto + moreover from calculation + have "w \ M[G]" by (auto dest:ext.transM) + ultimately + show ?thesis using ext.function_space_rel_char by auto + qed + moreover + have "function(f\<^bsub>G\<^esub>)" + using Un_filter_is_function generic + unfolding M_generic_def by fast + ultimately + show ?thesis + using generic domain_f_G unfolding Pi_def by auto +qed + +abbreviation + surj_dense :: "i\i" where + "surj_dense(x) \ { p\Coll . x \ range(p) }" + +(* TODO: write general versions of this for \<^term>\Fn\<^bsup>M\<^esup>(\,I,J)\ *) +lemma dense_surj_dense: + assumes "x \ \ \\<^bsup>M\<^esup> 2" + shows "dense(surj_dense(x))" +proof + fix p + assume "p \ Coll" + then + have "countable\<^bsup>M\<^esup>(p)" using Coll_into_countable_rel by simp + show "\d\surj_dense(x). d \ p" + proof - + from \p \ Coll\ + have "domain(p) \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "p\M" + using transM[of _ Coll] domain_of_fun + by (auto del:Fn_relD dest!:Fn_relD del:domainE) + moreover from \countable\<^bsup>M\<^esup>(p)\ + have "domain(p) \ {fst(x) . x \ p }" by (auto intro!: rev_bexI) + moreover from calculation + have "{ fst(x) . x \ p } \ M" + using lam_replacement_fst[THEN lam_replacement_imp_strong_replacement] + by (auto simp flip:setclass_iff intro!:RepFun_closed dest:transM) + moreover from calculation and \countable\<^bsup>M\<^esup>(p)\ + have "countable\<^bsup>M\<^esup>({fst(x) . x \ p })" + using cardinal_rel_RepFun_le lam_replacement_fst + countable_rel_iff_cardinal_rel_le_nat[THEN iffD1, THEN [2] le_trans, of _ p] + by (rule_tac countable_rel_iff_cardinal_rel_le_nat[THEN iffD2]) simp_all + moreover from calculation + have "countable\<^bsup>M\<^esup>(domain(p))" + using uncountable_rel_not_subset_countable_rel[of "{fst(x) . x \ p }" "domain(p)"] + by auto + ultimately + obtain \ where "\ \ domain(p)" "\\\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" + using lt_cardinal_rel_imp_not_subset[of "domain(p)" "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>"] + Ord_Aleph_rel countable_iff_le_rel_Aleph_rel_one[THEN iffD1, + THEN lesspoll_cardinal_lt_rel, of "domain(p)"] + cardinal_rel_idem by auto + moreover note assms + moreover from calculation and \p \ Coll\ + have "cons(\\,x\, p) \ Coll" "x\M" "cons(\\,x\, p) \ p" + using InfCard_rel_Aleph_rel + by (auto del:Fnle_relI intro!: cons_in_Fn_rel Fnle_relI dest:transM) + ultimately + show ?thesis by blast + qed +qed + +lemma surj_dense_closed[intro,simp]: + "x \ \ \\<^bsup>M\<^esup> 2 \ surj_dense(x) \ M" + using separation_in_range transM[of x] by simp + +lemma reals_sub_image_f_G: + assumes "x\\ \\<^bsup>M\<^esup> 2" + shows "\\\\\<^bsub>1\<^esub>\<^bsup>M\<^esup>. f\<^bsub>G\<^esub> ` \ = x" +proof - + from assms + have "dense(surj_dense(x))" using dense_surj_dense by simp + with \x \ \ \\<^bsup>M\<^esup> 2\ + obtain p where "p\surj_dense(x)" "p\G" + using generic[THEN M_generic_denseD, of "surj_dense(x)"] + by blast + then + show ?thesis + using succ_omega_closed_Coll f_G_funtype function_apply_equality[of _ x f_G] + succ_omega_closed_imp_no_new_reals[symmetric] + by (auto, rule_tac bexI) (auto simp:Pi_def) +qed + +lemma f_G_surj_Aleph_rel1_reals: "f\<^bsub>G\<^esub> \ surj\<^bsup>M[G]\<^esup>(\\<^bsub>1\<^esub>\<^bsup>M\<^esup>, \ \\<^bsup>M[G]\<^esup> 2)" + using Aleph_rel_sub_closed +proof (intro ext.mem_surj_abs[THEN iffD2]) + show "f\<^bsub>G\<^esub> \ surj(\\<^bsub>1\<^esub>\<^bsup>M\<^esup>, \ \\<^bsup>M[G]\<^esup> 2)" + unfolding surj_def + proof (intro ballI CollectI impI) + show "f\<^bsub>G\<^esub> \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ \ \\<^bsup>M[G]\<^esup> 2" + using f_G_funtype G_in_MG ext.nat_into_M f_G_in_MG by simp + fix x + assume "x \ \ \\<^bsup>M[G]\<^esup> 2" + then + show "\\\\\<^bsub>1\<^esub>\<^bsup>M\<^esup>. f\<^bsub>G\<^esub> ` \ = x" + using reals_sub_image_f_G succ_omega_closed_Coll + f_G_funtype succ_omega_closed_imp_no_new_reals by simp + qed +qed simp_all + +lemma continuum_rel_le_Aleph1_extension: + includes G_generic1_lemmas + shows "2\<^bsup>\\\<^bsub>0\<^esub>\<^bsup>M[G]\<^esup>,M[G]\<^esup> \ \\<^bsub>1\<^esub>\<^bsup>M[G]\<^esup>" +proof - + have "\\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ M[G]" "Ord(\\<^bsub>1\<^esub>\<^bsup>M\<^esup>)" + using Card_rel_is_Ord by auto + moreover from this + have "\ \\<^bsup>M[G]\<^esup> 2 \\<^bsup>M[G]\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" + using ext.surj_rel_implies_inj_rel[OF f_G_surj_Aleph_rel1_reals] + f_G_in_MG unfolding lepoll_rel_def by auto + with \Ord(\\<^bsub>1\<^esub>\<^bsup>M\<^esup>)\ + have "|\ \\<^bsup>M[G]\<^esup> 2|\<^bsup>M[G]\<^esup> \ |\\<^bsub>1\<^esub>\<^bsup>M\<^esup>|\<^bsup>M[G]\<^esup>" + using M_subset_MG[OF one_in_G, OF generic] Aleph_rel_closed[of 1] + by (rule_tac ext.lepoll_rel_imp_cardinal_rel_le) simp_all + ultimately + have "2\<^bsup>\\\<^bsub>0\<^esub>\<^bsup>M[G]\<^esup>,M[G]\<^esup> \ |\\<^bsub>1\<^esub>\<^bsup>M[G]\<^esup>|\<^bsup>M[G]\<^esup>" + using ext.lepoll_rel_imp_cardinal_rel_le[of "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "\ \\<^bsup>M[G]\<^esup> 2"] + ext.Aleph_rel_zero succ_omega_closed_Coll + succ_omega_closed_imp_Aleph_1_preserved + unfolding cexp_rel_def by simp + then + show "2\<^bsup>\\\<^bsub>0\<^esub>\<^bsup>M[G]\<^esup>,M[G]\<^esup> \ \\<^bsub>1\<^esub>\<^bsup>M[G]\<^esup>" by simp +qed + +theorem CH: "\\<^bsub>1\<^esub>\<^bsup>M[G]\<^esup> = 2\<^bsup>\\\<^bsub>0\<^esub>\<^bsup>M[G]\<^esup>,M[G]\<^esup>" + using continuum_rel_le_Aleph1_extension ext.Aleph_rel_succ[of 0] + ext.Aleph_rel_zero ext.csucc_rel_le[of "2\<^bsup>\\\<^bsub>0\<^esub>\<^bsup>M[G]\<^esup>,M[G]\<^esup>" \] + ext.Card_rel_cexp_rel ext.cantor_cexp_rel[of \] ext.Card_rel_nat + le_anti_sym + by auto + +end \ \\<^locale>\collapse_generic4\\ + +subsection\Models of fragments of $\ZFC + \CH$\ + +theorem ctm_of_CH: + assumes + "M \ \" "Transset(M)" "M \ ZC \ {\Replacement(p)\ . p \ overhead}" + "\ \ formula" "M \ { \Replacement(ground_repl_fm(\))\ . \ \ \}" + shows + "\N. + M \ N \ N \ \ \ Transset(N) \ N \ ZC \ {\CH\} \ { \Replacement(\)\ . \ \ \} \ + (\\. Ord(\) \ (\ \ M \ \ \ N))" +proof - + from \M \ ZC \ {\Replacement(p)\ . p \ overhead}\ + interpret M_ZFC4 M + using M_satT_overhead_imp_M_ZF4 by simp + from \Transset(M)\ + interpret M_ZFC4_trans M + using M_satT_imp_M_ZF4 + by unfold_locales + from \M \ \\ + obtain enum where "enum \ bij(\,M)" + using eqpoll_sym unfolding eqpoll_def by blast + then + interpret M_ctm3_AC M enum by unfold_locales + interpret forcing_data1 "Coll" "Colleq" 0 M enum + using zero_in_Fn_rel[of "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "\ \\<^bsup>M\<^esup> 2"] + zero_top_Fn_rel[of _ "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "\ \\<^bsup>M\<^esup> 2"] + preorder_on_Fnle_rel[of "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "\ \\<^bsup>M\<^esup> 2"] + zero_lt_Aleph_rel1 + by unfold_locales simp_all + obtain G where "M_generic(G)" + using generic_filter_existence[OF one_in_P] + by auto + moreover from this + interpret collapse_generic4 M enum G by unfold_locales + have "\\<^bsub>1\<^esub>\<^bsup>M[G]\<^esup> = 2\<^bsup>\\\<^bsub>0\<^esub>\<^bsup>M[G]\<^esup>,M[G]\<^esup>" + using CH . + then + have "M[G], [] \ \CH\" + using ext.is_ContHyp_iff + by (simp add:ContHyp_rel_def) + then + have "M[G] \ ZC \ {\CH\}" + using ext.M_satT_ZC by auto + moreover + have "Transset(M[G])" using Transset_MG . + moreover + have "M \ M[G]" using M_subset_MG[OF one_in_G] generic by simp + moreover + note \M \ { \Replacement(ground_repl_fm(\))\ . \ \ \}\ \\ \ formula\ + ultimately + show ?thesis + using Ord_MG_iff MG_eqpoll_nat satT_ground_repl_fm_imp_satT_ZF_replacement_fm[of \] + by (rule_tac x="M[G]" in exI,blast) +qed + +corollary ctm_ZFC_imp_ctm_CH: + assumes + "M \ \" "Transset(M)" "M \ ZFC" + shows + "\N. + M \ N \ N \ \ \ Transset(N) \ N \ ZFC \ {\CH\} \ + (\\. Ord(\) \ (\ \ M \ \ \ N))" +proof - + from assms + have "\N. + M \ N \ + N \ \ \ + Transset(N) \ + N \ ZC \ N \ {\CH\} \ N \ {\Replacement(x)\ . x \ formula} \ (\\. Ord(\) \ \ \ M \ \ \ N)" + using ctm_of_CH[of M formula] satT_ZFC_imp_satT_ZC[of M] + satT_mono[OF _ ground_repl_fm_sub_ZFC, of M] + satT_mono[OF _ ZF_replacement_overhead_sub_ZFC, of M] + satT_mono[OF _ ZF_replacement_fms_sub_ZFC, of M] + by (simp add: satT_Un_iff) + then + obtain N where "N \ ZC" "N \ {\CH\}" "N \ {\Replacement(x)\ . x \ formula}" + "M \ N" "N \ \" "Transset(N)" "(\\. Ord(\) \ \ \ M \ \ \ N)" + by auto + moreover from this + have "N \ ZFC" + using satT_ZC_ZF_replacement_imp_satT_ZFC + by auto + moreover from this and \N \ {\CH\}\ + have "N \ ZFC \ {\CH\}" + using satT_ZC_ZF_replacement_imp_satT_ZFC + by auto + ultimately + show ?thesis + by auto +qed + +end \ No newline at end of file diff --git a/thys/Independence_CH/Cardinal_Preservation.thy b/thys/Independence_CH/Cardinal_Preservation.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Cardinal_Preservation.thy @@ -0,0 +1,523 @@ +section\Preservation of cardinals in generic extensions\ + +theory Cardinal_Preservation + imports + Forcing_Main +begin + +context forcing_notion +begin + +definition + antichain :: "i\o" where + "antichain(A) \ A\P \ (\p\A. \q\A. p \ q \ p \ q)" + +definition + ccc :: "o" where + "ccc \ \A. antichain(A) \ |A| \ \" + +end \ \\<^locale>\forcing_notion\\ + +context forcing_data1 + +begin +abbreviation + antichain_r' :: "i \ o" where + "antichain_r'(A) \ antichain_rel(##M,P,leq,A)" + +lemma antichain_abs' [absolut]: + "\ A\M \ \ antichain_r'(A) \ antichain(A)" + unfolding antichain_rel_def antichain_def compat_def + using P_in_M leq_in_M transitivity[of _ A] + by (auto simp add:absolut) + +lemma (in forcing_notion) Incompatible_imp_not_eq: "\ p \ q; p\P; q\P \\ p \ q" + using refl_leq by blast + +lemma inconsistent_imp_incompatible: + assumes "p \ \ env" "q \ Neg(\) env" "p\P" "q\P" + "arity(\) \ length(env)" "\ \ formula" "env \ list(M)" + shows "p \ q" +proof + assume "compat(p,q)" + then + obtain d where "d \ p" "d \ q" "d \ P" by blast + moreover + note assms + moreover from calculation + have "d \ \ env" "d \ Neg(\) env" + using strengthening_lemma by simp_all + ultimately + show "False" + using Forces_Neg[of d env \] refl_leq P_in_M + by (auto dest:transM; drule_tac bspec; auto dest:transM) +qed + +notation check (\_\<^sup>v\ [101] 100) + +end \ \\<^locale>\forcing_data1\\ + +locale G_generic2 = G_generic1 + forcing_data2 +locale G_generic2_AC = G_generic1_AC + G_generic2 + M_ctm2_AC + +locale G_generic3 = G_generic2 + forcing_data3 +locale G_generic3_AC = G_generic2_AC + G_generic3 + +locale G_generic4 = G_generic3 + forcing_data4 +locale G_generic4_AC = G_generic3_AC + G_generic4 + +sublocale G_generic4_AC \ ext:M_ZFC3_trans "M[G]" + using ground_replacements4 replacement_assm_MG + by unfold_locales simp_all + +lemma (in forcing_data1) forces_neq_apply_imp_incompatible: + assumes + "p \ \0`1 is 2\ [f,a,b\<^sup>v]" + "q \ \0`1 is 2\ [f,a,b'\<^sup>v]" + "b \ b'" + \ \More general version: taking general names + \<^term>\b\<^sup>v\ and \<^term>\b'\<^sup>v\, satisfying + \<^term>\p \ \\\0 = 1\\ [b\<^sup>v, b'\<^sup>v]\ and + \<^term>\q \ \\\0 = 1\\ [b\<^sup>v, b'\<^sup>v]\.\ + and + types:"f\M" "a\M" "b\M" "b'\M" "p\P" "q\P" + shows + "p \ q" +proof - + { + fix G + assume "M_generic(G)" + then + interpret G_generic1 _ _ _ _ _ G by unfold_locales + include G_generic1_lemmas + (* NOTE: might be useful to have a locale containg two \M_ZF1_trans\ + instances, one for \<^term>\M\ and one for \<^term>\M[G]\ *) + assume "q\G" + with assms \M_generic(G)\ + have "M[G], map(val(P,G),[f,a,b'\<^sup>v]) \ \0`1 is 2\" + using truth_lemma[of "\0`1 is 2\" G "[f,a,b'\<^sup>v]"] + by (auto simp add:ord_simp_union arity_fun_apply_fm + fun_apply_type) + with \b \ b'\ types + have "M[G], map(val(P,G),[f,a,b\<^sup>v]) \ \\\0`1 is 2\\" + using GenExtI by auto + } + with types + have "q \ \\\0`1 is 2\\ [f,a,b\<^sup>v]" + using definition_of_forcing[where \="\\\0`1 is 2\\" ] check_in_M + by (auto simp add:ord_simp_union arity_fun_apply_fm) + with \p \ \0`1 is 2\ [f,a,b\<^sup>v]\ and types + show "p \ q" + using check_in_M inconsistent_imp_incompatible + by (simp add:ord_simp_union arity_fun_apply_fm fun_apply_type) +qed + +context M_ctm3_AC +begin + +\ \Simplifying simp rules (because of the occurrence of "\#\#")\ +lemmas sharp_simps = Card_rel_Union Card_rel_cardinal_rel Collect_abs + Cons_abs Cons_in_M_iff Diff_closed Equal_abs Equal_in_M_iff Finite_abs + Forall_abs Forall_in_M_iff Inl_abs Inl_in_M_iff Inr_abs Inr_in_M_iff + Int_closed Inter_abs Inter_closed M_nat Member_abs Member_in_M_iff + Memrel_closed Nand_abs Nand_in_M_iff Nil_abs Nil_in_M Ord_cardinal_rel + Pow_rel_closed Un_closed Union_abs Union_closed and_abs and_closed + apply_abs apply_closed bij_rel_closed bijection_abs bool_of_o_abs + bool_of_o_closed cadd_rel_0 cadd_rel_closed cardinal_rel_0_iff_0 + cardinal_rel_closed cardinal_rel_idem cartprod_abs cartprod_closed + cmult_rel_0 cmult_rel_1 cmult_rel_closed comp_closed composition_abs + cons_abs cons_closed converse_abs converse_closed csquare_lam_closed + csquare_rel_closed depth_closed domain_abs domain_closed eclose_abs + eclose_closed empty_abs field_abs field_closed finite_funspace_closed + finite_ordinal_abs formula_N_abs formula_N_closed formula_abs + formula_case_abs formula_case_closed formula_closed + formula_functor_abs fst_closed function_abs function_space_rel_closed + hd_abs image_abs image_closed inj_rel_closed injection_abs inter_abs + irreflexive_abs is_depth_apply_abs is_eclose_n_abs is_funspace_abs + iterates_closed length_abs length_closed lepoll_rel_refl + limit_ordinal_abs linear_rel_abs list_N_abs list_N_closed list_abs + list_case'_closed list_case_abs list_closed list_functor_abs + mem_bij_abs mem_eclose_abs mem_inj_abs mem_list_abs membership_abs + minimum_closed nat_case_abs nat_case_closed nonempty not_abs + not_closed nth_abs number1_abs number2_abs number3_abs omega_abs + or_abs or_closed order_isomorphism_abs ordermap_closed + ordertype_closed ordinal_abs pair_abs pair_in_M_iff powerset_abs + pred_closed pred_set_abs quasilist_abs quasinat_abs radd_closed + rall_abs range_abs range_closed relation_abs restrict_closed + restriction_abs rex_abs rmult_closed rtrancl_abs rtrancl_closed + rvimage_closed separation_closed setdiff_abs singleton_abs + singleton_in_M_iff snd_closed strong_replacement_closed subset_abs + succ_in_M_iff successor_abs successor_ordinal_abs sum_abs sum_closed + surj_rel_closed surjection_abs tl_abs trancl_abs trancl_closed + transitive_rel_abs transitive_set_abs typed_function_abs union_abs + upair_abs upair_in_M_iff vimage_abs vimage_closed well_ord_abs + mem_formula_abs nth_closed Aleph_rel_closed csucc_rel_closed + Card_rel_Aleph_rel + +declare sharp_simps[simp del, simplified setclass_iff, simp] + +lemmas sharp_intros = nat_into_M Aleph_rel_closed Card_rel_Aleph_rel + +declare sharp_intros[rule del, simplified setclass_iff, intro] + +end \ \\<^locale>\M_ctm3_AC\\ + +context G_generic4_AC begin + +context + includes G_generic1_lemmas +begin + +lemmas mg_sharp_simps = ext.Card_rel_Union ext.Card_rel_cardinal_rel + ext.Collect_abs ext.Cons_abs ext.Cons_in_M_iff ext.Diff_closed + ext.Equal_abs ext.Equal_in_M_iff ext.Finite_abs ext.Forall_abs + ext.Forall_in_M_iff ext.Inl_abs ext.Inl_in_M_iff ext.Inr_abs + ext.Inr_in_M_iff ext.Int_closed ext.Inter_abs ext.Inter_closed + ext.M_nat ext.Member_abs ext.Member_in_M_iff ext.Memrel_closed + ext.Nand_abs ext.Nand_in_M_iff ext.Nil_abs ext.Nil_in_M + ext.Ord_cardinal_rel ext.Pow_rel_closed ext.Un_closed + ext.Union_abs ext.Union_closed ext.and_abs ext.and_closed + ext.apply_abs ext.apply_closed ext.bij_rel_closed + ext.bijection_abs ext.bool_of_o_abs ext.bool_of_o_closed + ext.cadd_rel_0 ext.cadd_rel_closed ext.cardinal_rel_0_iff_0 + ext.cardinal_rel_closed ext.cardinal_rel_idem ext.cartprod_abs + ext.cartprod_closed ext.cmult_rel_0 ext.cmult_rel_1 + ext.cmult_rel_closed ext.comp_closed ext.composition_abs + ext.cons_abs ext.cons_closed ext.converse_abs ext.converse_closed + ext.csquare_lam_closed ext.csquare_rel_closed ext.depth_closed + ext.domain_abs ext.domain_closed ext.eclose_abs ext.eclose_closed + ext.empty_abs ext.field_abs ext.field_closed + ext.finite_funspace_closed ext.finite_ordinal_abs ext.formula_N_abs + ext.formula_N_closed ext.formula_abs ext.formula_case_abs + ext.formula_case_closed ext.formula_closed ext.formula_functor_abs + ext.fst_closed ext.function_abs ext.function_space_rel_closed + ext.hd_abs ext.image_abs ext.image_closed ext.inj_rel_closed + ext.injection_abs ext.inter_abs ext.irreflexive_abs + ext.is_depth_apply_abs ext.is_eclose_n_abs ext.is_funspace_abs + ext.iterates_closed ext.length_abs ext.length_closed + ext.lepoll_rel_refl ext.limit_ordinal_abs ext.linear_rel_abs + ext.list_N_abs ext.list_N_closed ext.list_abs + ext.list_case'_closed ext.list_case_abs ext.list_closed + ext.list_functor_abs ext.mem_bij_abs ext.mem_eclose_abs + ext.mem_inj_abs ext.mem_list_abs ext.membership_abs + ext.nat_case_abs ext.nat_case_closed + ext.nonempty ext.not_abs ext.not_closed ext.nth_abs + ext.number1_abs ext.number2_abs ext.number3_abs ext.omega_abs + ext.or_abs ext.or_closed ext.order_isomorphism_abs + ext.ordermap_closed ext.ordertype_closed ext.ordinal_abs + ext.pair_abs ext.pair_in_M_iff ext.powerset_abs ext.pred_closed + ext.pred_set_abs ext.quasilist_abs ext.quasinat_abs + ext.radd_closed ext.rall_abs ext.range_abs ext.range_closed + ext.relation_abs ext.restrict_closed ext.restriction_abs + ext.rex_abs ext.rmult_closed ext.rtrancl_abs ext.rtrancl_closed + ext.rvimage_closed ext.separation_closed ext.setdiff_abs + ext.singleton_abs ext.singleton_in_M_iff ext.snd_closed + ext.strong_replacement_closed ext.subset_abs ext.succ_in_M_iff + ext.successor_abs ext.successor_ordinal_abs ext.sum_abs + ext.sum_closed ext.surj_rel_closed ext.surjection_abs ext.tl_abs + ext.trancl_abs ext.trancl_closed ext.transitive_rel_abs + ext.transitive_set_abs ext.typed_function_abs ext.union_abs + ext.upair_abs ext.upair_in_M_iff ext.vimage_abs ext.vimage_closed + ext.well_ord_abs ext.mem_formula_abs ext.nth_closed ext.Aleph_rel_closed + ext.csucc_rel_closed ext.Card_rel_Aleph_rel + +\ \The following was motivated by the fact that + @{thm ext.apply_closed} did not simplify appropriately.\ +declare mg_sharp_simps[simp del, simplified setclass_iff, simp] + +lemmas mg_sharp_intros = ext.nat_into_M ext.Aleph_rel_closed + ext.Card_rel_Aleph_rel + +declare mg_sharp_intros[rule del, simplified setclass_iff, intro] + +\ \Kunen IV.2.31\ +lemma forces_below_filter: + assumes "M[G], map(val(P,G),env) \ \" "p \ G" + "arity(\) \ length(env)" "\ \ formula" "env \ list(M)" + shows "\q\G. q \ p \ q \ \ env" +proof - + note assms + moreover from this + obtain r where "r \ \ env" "r\G" + using generic truth_lemma[of \ _ env] + by blast + moreover from this and \p\G\ + obtain q where "q \ p" "q \ r" "q \ G" by auto + ultimately + show ?thesis + using strengthening_lemma[of r \ _ env] by blast +qed + +subsection\Preservation by ccc forcing notions\ + +\ \This definition has the arguments in the expected order by most of the lemmas: +first the parameters, the only argument in the penultimate place and the result in +the last place.\ +definition check_fm' where + "check_fm'(ofm,arg,res) \ check_fm(arg,ofm,res)" + +lemma ccc_fun_closed_lemma_aux: + assumes "f_dot\M" "p\M" "a\M" "b\M" + shows "{q \ P . q \ p \ (M, [q, P, leq, \, f_dot, a\<^sup>v, b\<^sup>v] \ forces(\0`1 is 2\ ))} \ M" +proof - + have "\0`1 is 2\ \ formula" "arity(\0`1 is 2\ ) = 3" + using arity_fun_apply_fm union_abs1 + by simp_all + then + show ?thesis + using separation_forces[where env="[f_dot, a\<^sup>v, b\<^sup>v]" and \="\0`1 is 2\",simplified] + assms G_subset_M[THEN subsetD] generic one_in_M P_in_M + separation_in lam_replacement_constant lam_replacement_identity + lam_replacement_Pair[THEN[5] lam_replacement_hcomp2] leq_in_M check_in_M + separation_conj separation_forces + by simp_all +qed + +lemma ccc_fun_closed_lemma_aux2: + assumes "B\M" "f_dot\M" "p\M" "a\M" + shows "(##M)(\b\B. {q \ P . q \ p \ (M, [q, P, leq, \, f_dot, a\<^sup>v, b\<^sup>v] \ forces(\0`1 is 2\ ))})" +proof - + have "separation(##M, \z. M, [snd(z), P, leq, \, f_dot, \, fst(fst(z))\<^sup>v] \ forces(\0`1 is 2\ ))" + if "\\M" for \ + proof - + let ?f_fm="snd_fm(1,0)" + let ?g_fm="hcomp_fm(check_fm'(6),hcomp_fm(fst_fm,fst_fm),2,0)" + note types = assms leq_in_M P_in_M one_in_M + have "arity(forces(\0`1 is 2\ )) \ 7" + using arity_fun_apply_fm union_abs1 arity_forces[of "\0`1 is 2\ "] + by simp + moreover + have "?f_fm \ formula" "arity(?f_fm) \ 7" "?g_fm \ formula" "arity(?g_fm) \ 8" + using ord_simp_union + unfolding hcomp_fm_def check_fm'_def + by (simp_all add:arity) + ultimately + show ?thesis + using separation_sat_after_function types that sats_fst_fm + snd_abs types sats_snd_fm sats_check_fm check_abs check_in_M fst_abs + unfolding hcomp_fm_def check_fm'_def + by simp + qed + then + show ?thesis + using lam_replacement_imp_lam_closed lam_replacement_Collect + separation_conj separation_in separation_forces separation_ball separation_iff' + lam_replacement_Pair[THEN [5] lam_replacement_hcomp2] lam_replacement_identity + lam_replacement_constant lam_replacement_snd lam_replacement_fst lam_replacement_hcomp + ccc_fun_closed_lemma_aux arity_fun_apply_fm union_abs1 + transitivity[of _ B] leq_in_M assms + by simp +qed + +lemma ccc_fun_closed_lemma: + assumes "A\M" "B\M" "f_dot\M" "p\M" + shows "(\a\A. {b\B. \q\P. q \ p \ (q \ \0`1 is 2\ [f_dot, a\<^sup>v, b\<^sup>v])}) \ M" +proof - + have "separation(##M, \z. M, [snd(z), P, leq, \, f_dot, fst(fst(fst(z)))\<^sup>v, snd(fst(z))\<^sup>v] \ forces(\0`1 is 2\ ))" + proof - + note types = assms leq_in_M P_in_M one_in_M + let ?f_fm="snd_fm(1,0)" + let ?g="\z . fst(fst(fst(z)))\<^sup>v" + let ?g_fm="hcomp_fm(check_fm'(6),hcomp_fm(fst_fm,hcomp_fm(fst_fm,fst_fm)),2,0)" + let ?h_fm="hcomp_fm(check_fm'(7),hcomp_fm(snd_fm,fst_fm),3,0)" + have "arity(forces(\0`1 is 2\ )) \ 7" + using arity_fun_apply_fm union_abs1 arity_forces[of "\0`1 is 2\ "] by simp + moreover + have "?f_fm \ formula" "arity(?f_fm) \ 6" "?g_fm \ formula" "arity(?g_fm) \ 7" + "?h_fm \ formula" "arity(?h_fm) \ 8" + using ord_simp_union + unfolding hcomp_fm_def check_fm'_def + by (simp_all add:arity) + ultimately + show ?thesis + using separation_sat_after_function3 assms types sats_check_fm check_abs check_in_M + fst_abs snd_abs + unfolding hcomp_fm_def check_fm'_def + by simp + qed + moreover + have "separation(##M, \z. M, [snd(z), P, leq, \, f_dot, \, fst(z)\<^sup>v] \ forces(\0`1 is 2\ ))" + if "\\M" for \ + proof - + let ?f_fm="snd_fm(1,0)" + let ?g_fm="hcomp_fm(check_fm'(6),fst_fm,2,0)" + note types = assms leq_in_M P_in_M one_in_M + have "arity(forces(\0`1 is 2\ )) \ 7" + using arity_forces[of "\0`1 is 2\ "] arity_fun_apply_fm union_abs1 + by simp + moreover + have "?f_fm \ formula" "arity(?f_fm) \ 7" "?g_fm \ formula" "arity(?g_fm) \ 8" + using ord_simp_union + unfolding hcomp_fm_def check_fm'_def + by (simp_all add:arity) + ultimately + show ?thesis + using separation_sat_after_function assms types that fst_abs + snd_abs types sats_check_fm check_abs check_in_M + unfolding hcomp_fm_def check_fm'_def + by simp + qed + ultimately + show ?thesis + using lam_replacement_imp_lam_closed lam_replacement_Collect + lam_replacement_constant lam_replacement_identity lam_replacement_snd lam_replacement_fst + lam_replacement_hcomp lam_replacement_Pair[THEN [5] lam_replacement_hcomp2] + separation_conj separation_in separation_ball separation_bex separation_iff' + transitivity[of _ A] leq_in_M assms + by simp +qed + +\ \Kunen IV.3.5\ +lemma ccc_fun_approximation_lemma: + notes le_trans[trans] + assumes "ccc\<^bsup>M\<^esup>(P,leq)" "A\M" "B\M" "f\M[G]" "f : A \ B" + shows + "\F\M. F : A \ Pow\<^bsup>M\<^esup>(B) \ (\a\A. f`a \ F`a \ |F`a|\<^bsup>M\<^esup> \ \)" +proof - + from \f\M[G]\ + obtain f_dot where "f = val(P,G,f_dot)" "f_dot\M" using GenExtD by force + with assms + obtain p where "p \ \0:1\2\ [f_dot, A\<^sup>v, B\<^sup>v]" "p\G" "p\M" + using transitivity[OF M_genericD P_in_M] + generic truth_lemma[of "\0:1\2\" G "[f_dot, A\<^sup>v, B\<^sup>v]"] + by (auto simp add:ord_simp_union arity_typed_function_fm + \ \NOTE: type-checking is not performed here by the Simplifier\ + typed_function_type) + define F where "F\\a\A. {b\B. \q\P. q \ p \ (q \ \0`1 is 2\ [f_dot, a\<^sup>v, b\<^sup>v])}" + from assms \f_dot\_\ \p\M\ + have "F \ M" + unfolding F_def using ccc_fun_closed_lemma by simp + moreover from calculation + have "f`a \ F`a" if "a \ A" for a + proof - + note \f: A \ B\ \a \ A\ + moreover from this + have "f ` a \ B" by simp + moreover + note \f\M[G]\ \A\M\ + moreover from calculation + have "M[G], [f, a, f`a] \ \0`1 is 2\" + by (auto dest:transM) + moreover + note \B\M\ \f = val(P,G,f_dot)\ + moreover from calculation + have "a\M" "val(P,G, f_dot)`a\M" + by (auto dest:transM) + moreover + note \f_dot\M\ \p\G\ + ultimately + obtain q where "q \ p" "q \ \0`1 is 2\ [f_dot, a\<^sup>v, (f`a)\<^sup>v]" "q\G" + using forces_below_filter[of "\0`1 is 2\" "[f_dot, a\<^sup>v, (f`a)\<^sup>v]" p] + by (auto simp add: ord_simp_union arity_fun_apply_fm + fun_apply_type) + with \f`a \ B\ + have "f`a \ {b\B . \q\P. q \ p \ q \ \0`1 is 2\ [f_dot, a\<^sup>v, b\<^sup>v]}" + by blast + with \a\A\ + show ?thesis unfolding F_def by simp + qed + moreover + have "|F`a|\<^bsup>M\<^esup> \ \ \ F`a\M" if "a \ A" for a + proof - + let ?Q="\b. {q\P. q \ p \ (q \ \0`1 is 2\ [f_dot, a\<^sup>v, b\<^sup>v])}" + from \F \ M\ \a\A\ \A\M\ + have "F`a \ M" "a\M" + using transitivity[OF _ \A\M\] by simp_all + moreover + have 2:"\x. x\F`a \ x\M" + using transitivity[OF _ \F`a\M\] by simp + moreover + have 3:"\x. x\F`a \ (##M)(?Q(x))" + using ccc_fun_closed_lemma_aux[OF \f_dot\M\ \p\M\ \a\M\ 2] transitivity[of _ "F`a"] + by simp + moreover + have 4:"lam_replacement(##M,\b. {q \ P . q \ p \ (M, [q, P, leq, \, f_dot, a\<^sup>v, b\<^sup>v] \ forces(\0`1 is 2\ ))})" + using ccc_fun_closed_lemma_aux2[OF _ \f_dot\M\ \p\M\ \a\M\] + lam_replacement_iff_lam_closed[THEN iffD2] + ccc_fun_closed_lemma_aux[OF \f_dot\M\ \p\M\ \a\M\] + by simp + ultimately + interpret M_Pi_assumptions_choice "##M" "F`a" ?Q + using Pi_replacement1[OF _ 3] lam_replacement_Sigfun[OF 4] + lam_replacement_imp_strong_replacement + ccc_fun_closed_lemma_aux[OF \f_dot\M\ \p\M\ \a\M\] + lam_replacement_hcomp2[OF lam_replacement_constant 4 _ _ + lam_replacement_minimum,unfolded lam_replacement_def] + by unfold_locales simp_all + from \F`a \ M\ + interpret M_Pi_assumptions2 "##M" "F`a" ?Q "\_ . P" + using P_in_M lam_replacement_imp_strong_replacement[OF + lam_replacement_Sigfun[OF lam_replacement_constant]] + Pi_replacement1 transM[of _ "F`a"] + by unfold_locales simp_all + from \p \ \0:1\2\ [f_dot, A\<^sup>v, B\<^sup>v]\ \a\A\ + have "\y. y \ ?Q(b)" if "b \ F`a" for b + using that unfolding F_def by auto + then + obtain q where "q \ Pi\<^bsup>M\<^esup>(F`a,?Q)" "q\M" using AC_Pi_rel by auto + moreover + note \F`a \ M\ + moreover from calculation + have "q : F`a \\<^bsup>M\<^esup> P" + using Pi_rel_weaken_type def_function_space_rel by auto + moreover from calculation + have "q : F`a \ range(q)" "q : F`a \ P" "q : F`a \\<^bsup>M\<^esup> range(q)" + using mem_function_space_rel_abs range_of_fun by simp_all + moreover + have "q`b \ q`c" if "b \ F`a" "c \ F`a" "b \ c" + \ \For the next step, if the premise \<^term>\b \ c\ is first, + the proof breaks down badly\ + for b c + proof - + from \b \ F`a\ \c \ F`a\ \q \ Pi\<^bsup>M\<^esup>(F`a,?Q)\ \q\M\ + have "q`b \ \0`1 is 2\ [f_dot, a\<^sup>v, b\<^sup>v]" + "q`c \ \0`1 is 2\ [f_dot, a\<^sup>v, c\<^sup>v]" + using mem_Pi_rel_abs[of q] apply_type[of _ _ ?Q] + by simp_all + with \b \ c\ \q : F`a \ P\ \a\A\ \b\_\ \c\_\ + \A\M\ \f_dot\M\ \F`a\M\ + show ?thesis + using forces_neq_apply_imp_incompatible + transitivity[of _ A] transitivity[of _ "F`a"] + by auto + qed + moreover from calculation + have "antichain(range(q))" + using Pi_range_eq[of _ _ "\_ . P"] + unfolding antichain_def by auto + moreover from this and \q\M\ + have "antichain_r'(range(q))" + by (simp add:absolut) + moreover from calculation + have "q`b \ q`c" if "b \ c" "b \ F`a" "c \ F`a" for b c + using that Incompatible_imp_not_eq apply_type + mem_function_space_rel_abs by simp + ultimately + have "q \ inj\<^bsup>M\<^esup>(F`a,range(q))" + using def_inj_rel by auto + with \F`a \ M\ \q\M\ + have "|F`a|\<^bsup>M\<^esup> \ |range(q)|\<^bsup>M\<^esup>" + using def_lepoll_rel + by (rule_tac lepoll_rel_imp_cardinal_rel_le) auto + also from \antichain_r'(range(q))\ \ccc\<^bsup>M\<^esup>(P,leq)\ \q\M\ + have "|range(q)|\<^bsup>M\<^esup> \ \" + using def_ccc_rel by simp + finally + show ?thesis using \F`a\M\ by auto + qed + moreover from this + have "F`a\M" if "a\A" for a + using that by simp + moreover from this \B\M\ + have "F : A \ Pow\<^bsup>M\<^esup>(B)" + using Pow_rel_char + unfolding F_def by (rule_tac lam_type) auto + ultimately + show ?thesis by auto +qed + +end \ \G\_generic1\_lemmas bundle\ + +end \ \\<^locale>\G_generic4_AC\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Choice_Axiom.thy b/thys/Independence_CH/Choice_Axiom.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Choice_Axiom.thy @@ -0,0 +1,355 @@ +section\The Axiom of Choice in $M[G]$\ + +theory Choice_Axiom + imports + Powerset_Axiom + Extensionality_Axiom + Foundation_Axiom + Replacement_Axiom + Infinity_Axiom +begin + +definition + induced_surj :: "i\i\i\i" where + "induced_surj(f,a,e) \ f-``(range(f)-a)\{e} \ restrict(f,f-``a)" + +lemma domain_induced_surj: "domain(induced_surj(f,a,e)) = domain(f)" + unfolding induced_surj_def using domain_restrict domain_of_prod by auto + +lemma range_restrict_vimage: + assumes "function(f)" + shows "range(restrict(f,f-``a)) \ a" +proof + from assms + have "function(restrict(f,f-``a))" + using function_restrictI by simp + fix y + assume "y \ range(restrict(f,f-``a))" + then + obtain x where "\x,y\ \ restrict(f,f-``a)" "x \ f-``a" "x\domain(f)" + using domain_restrict domainI[of _ _ "restrict(f,f-``a)"] by auto + moreover + note \function(restrict(f,f-``a))\ + ultimately + have "y = restrict(f,f-``a)`x" + using function_apply_equality by blast + also from \x \ f-``a\ + have "restrict(f,f-``a)`x = f`x" + by simp + finally + have "y = f`x" . + moreover from assms \x\domain(f)\ + have "\x,f`x\ \ f" + using function_apply_Pair by auto + moreover + note assms \x \ f-``a\ + ultimately + show "y\a" + using function_image_vimage[of f a] by auto +qed + +lemma induced_surj_type: + assumes "function(f)" (* "relation(f)" (* a function can contain non-pairs *) *) + shows + "induced_surj(f,a,e): domain(f) \ {e} \ a" + and + "x \ f-``a \ induced_surj(f,a,e)`x = f`x" +proof - + let ?f1="f-``(range(f)-a) \ {e}" and ?f2="restrict(f, f-``a)" + have "domain(?f2) = domain(f) \ f-``a" + using domain_restrict by simp + moreover from assms + have "domain(?f1) = f-``(range(f))-f-``a" + using domain_of_prod function_vimage_Diff by simp + ultimately + have "domain(?f1) \ domain(?f2) = 0" + by auto + moreover + have "function(?f1)" "relation(?f1)" "range(?f1) \ {e}" + unfolding function_def relation_def range_def by auto + moreover from this and assms + have "?f1: domain(?f1) \ range(?f1)" + using function_imp_Pi by simp + moreover from assms + have "?f2: domain(?f2) \ range(?f2)" + using function_imp_Pi[of "restrict(f, f -`` a)"] function_restrictI by simp + moreover from assms + have "range(?f2) \ a" + using range_restrict_vimage by simp + ultimately + have "induced_surj(f,a,e): domain(?f1) \ domain(?f2) \ {e} \ a" + unfolding induced_surj_def using fun_is_function fun_disjoint_Un fun_weaken_type by simp + moreover + have "domain(?f1) \ domain(?f2) = domain(f)" + using domain_restrict domain_of_prod by auto + ultimately + show "induced_surj(f,a,e): domain(f) \ {e} \ a" + by simp + assume "x \ f-``a" + then + have "?f2`x = f`x" + using restrict by simp + moreover from \x \ f-``a\ \domain(?f1) = _\ + have "x \ domain(?f1)" + by simp + ultimately + show "induced_surj(f,a,e)`x = f`x" + unfolding induced_surj_def using fun_disjoint_apply2[of x ?f1 ?f2] by simp +qed + +lemma induced_surj_is_surj : + assumes + "e\a" "function(f)" "domain(f) = \" "\y. y \ a \ \x\\. f ` x = y" + shows "induced_surj(f,a,e) \ surj(\,a)" + unfolding surj_def +proof (intro CollectI ballI) + from assms + show "induced_surj(f,a,e): \ \ a" + using induced_surj_type[of f a e] cons_eq cons_absorb by simp + fix y + assume "y \ a" + with assms + have "\x\\. f ` x = y" + by simp + then + obtain x where "x\\" "f ` x = y" by auto + with \y\a\ assms + have "x\f-``a" + using vimage_iff function_apply_Pair[of f x] by auto + with \f ` x = y\ assms + have "induced_surj(f, a, e) ` x = y" + using induced_surj_type by simp + with \x\\\ show + "\x\\. induced_surj(f, a, e) ` x = y" by auto +qed + +context G_generic1 +begin + +lemma upair_name_abs : + assumes "x\M" "y\M" "z\M" "o\M" + shows "is_upair_name(##M,x,y,o,z) \ z = upair_name(x,y,o)" + unfolding is_upair_name_def upair_name_def + using assms zero_in_M pair_in_M_iff Upair_eq_cons + by simp + +lemma upair_name_closed : + "\ x\M; y\M ; o\M\ \ upair_name(x,y,o)\M" + unfolding upair_name_def + using upair_in_M_iff pair_in_M_iff Upair_eq_cons + by simp + +lemma opair_name_abs : + assumes "x\M" "y\M" "z\M" "o\M" + shows "is_opair_name(##M,x,y,o,z) \ z = opair_name(x,y,o)" + unfolding is_opair_name_def opair_name_def + using assms upair_name_abs upair_name_closed + by simp + +lemma opair_name_closed : + "\ x\M; y\M ; o\M \ \ opair_name(x,y,o)\M" + unfolding opair_name_def + using upair_name_closed by simp + +lemma val_upair_name : "val(P,G,upair_name(\,\,\)) = {val(P,G,\),val(P,G,\)}" + unfolding upair_name_def + using val_Upair Upair_eq_cons generic one_in_G one_in_P + by simp + +lemma val_opair_name : "val(P,G,opair_name(\,\,\)) = \val(P,G,\),val(P,G,\)\" + unfolding opair_name_def Pair_def + using val_upair_name by simp + +lemma val_RepFun_one: "val(P,G,{\f(x),\\ . x\a}) = {val(P,G,f(x)) . x\a}" +proof - + let ?A = "{f(x) . x \ a}" + let ?Q = "\\x,p\ . p = \" + have "\ \ P\G" using generic one_in_G one_in_P by simp + have "{\f(x),\\ . x \ a} = {t \ ?A \ P . ?Q(t)}" + using one_in_P by force + then + have "val(P,G,{\f(x),\\ . x \ a}) = val(P,G,{t \ ?A \ P . ?Q(t)})" + by simp + also + have "... = {z . t \ ?A , (\p\P\G . ?Q(\t,p\)) \ z= val(P,G,t)}" + using val_of_name_alt by simp + also + have "... = {val(P,G,t) . t \ ?A }" + using \\\P\G\ by force + also + have "... = {val(P,G,f(x)) . x \ a}" + by auto + finally + show ?thesis + by simp +qed + +\ \NOTE: The following bundled additions to the simpset might be of + use later on, perhaps add them globally to some appropriate + locale.\ +lemmas generic_simps = generic[THEN one_in_G, THEN valcheck, OF one_in_P] + generic[THEN one_in_G, THEN M_subset_MG, THEN subsetD] + check_in_M GenExtI P_in_M +lemmas generic_dests = M_genericD[OF generic] M_generic_compatD[OF generic] + +bundle G_generic1_lemmas = generic_simps[simp] generic_dests[dest] + +end\ \\<^locale>\G_generic1\\ + +subsection\$M[G]$ is a transitive model of ZF\ + +sublocale G_generic1 \ ext:M_Z_trans "M[G]" + using Transset_MG generic pairing_in_MG Union_MG + extensionality_in_MG power_in_MG foundation_in_MG + replacement_assm_MG separation_in_MG infinity_in_MG + replacement_ax1 by unfold_locales + +context G_generic1 +begin + +lemma opname_check_abs : + assumes "s\M" "x\M" "y\M" + shows "is_opname_check(##M,\,s,x,y) \ y = opair_name(check(x),s`x,\)" + unfolding is_opname_check_def + using assms check_abs check_in_M opair_name_abs apply_abs apply_closed one_in_M + by simp + +lemma repl_opname_check : + assumes "A\M" "f\M" + shows "{opair_name(check(x),f`x,\). x\A}\M" +proof - + have "arity(is_opname_check_fm(3,2,0,1))= 4" + using arity_is_opname_check_fm + by (simp add:ord_simp_union arity) + moreover + have "opair_name(check(x), f ` x,\)\M" if "x\A" for x + using assms opair_name_closed apply_closed transitivity check_in_M one_in_M that + by simp + ultimately + show ?thesis + using assms opname_check_abs[of f] is_opname_check_iff_sats + one_in_M zero_in_M transitivity + Replace_relativized_in_M[of "is_opname_check_fm(3,2,0,1)" + "[f,\]" _ "is_opname_check(##M,\,f)"] replacement_ax1(14) + by simp +qed + +theorem choice_in_MG: + assumes "choice_ax(##M)" + shows "choice_ax(##M[G])" +proof - + { + fix a + assume "a\M[G]" + then + obtain \ where "\\M" "val(P,G,\) = a" + using GenExt_def by auto + with \\\M\ + have "domain(\)\M" + using domain_closed by simp + then + obtain s \ where "s\surj(\,domain(\))" "Ord(\)" "s\M" "\\M" + using assms choice_ax_abs + by auto + then + have "\\M[G]" + using M_subset_MG generic one_in_G subsetD + by blast + let ?A="domain(\)\P" + let ?g = "{opair_name(check(\),s`\,\). \\\}" + have "?g \ M" + using \s\M\ \\\M\ repl_opname_check + by simp + let ?f_dot="{\opair_name(check(\),s`\,\),\\. \\\}" + have "?f_dot = ?g \ {\}" by blast + define f where + "f \ val(P,G,?f_dot)" + from \?g\M\ \?f_dot = ?g\{\}\ + have "?f_dot\M" + using cartprod_closed singleton_closed one_in_M + by simp + then + have "f \ M[G]" + unfolding f_def + by (blast intro:GenExtI) + have "f = {val(P,G,opair_name(check(\),s`\,\)) . \\\}" + unfolding f_def + using val_RepFun_one + by simp + also + have "... = {\\,val(P,G,s`\)\ . \\\}" + using val_opair_name valcheck generic one_in_G one_in_P + by simp + finally + have "f = {\\,val(P,G,s`\)\ . \\\}" . + then + have 1: "domain(f) = \" "function(f)" + unfolding function_def by auto + have 2: "y \ a \ \x\\. f ` x = y" for y + proof - + fix y + assume + "y \ a" + with \val(P,G,\) = a\ + obtain \ where "\\domain(\)" "val(P,G,\) = y" + using elem_of_val[of y _ \] + by blast + with \s\surj(\,domain(\))\ + obtain \ where "\\\" "s`\ = \" + unfolding surj_def + by auto + with \val(P,G,\) = y\ + have "val(P,G,s`\) = y" + by simp + with \f = {\\,val(P,G,s`\)\ . \\\}\ \\\\\ + have "\\,y\\f" + by auto + with \function(f)\ + have "f`\ = y" + using function_apply_equality by simp + with \\\\\ show + "\\\\. f ` \ = y" + by auto + qed + then + have "\\\(M[G]). \f'\(M[G]). Ord(\) \ f' \ surj(\,a)" + proof (cases "a=0") + case True + then + show ?thesis + unfolding surj_def + using zero_in_MG + by auto + next + case False + with \a\M[G]\ + obtain e where "e\a" "e\M[G]" + using transitivity_MG + by blast + with 1 and 2 + have "induced_surj(f,a,e) \ surj(\,a)" + using induced_surj_is_surj by simp + moreover from \f\M[G]\ \a\M[G]\ \e\M[G]\ + have "induced_surj(f,a,e) \ M[G]" + unfolding induced_surj_def + by (simp flip: setclass_iff) + moreover + note \\\M[G]\ \Ord(\)\ + ultimately + show ?thesis + by auto + qed + } + then + show ?thesis + using ext.choice_ax_abs + by simp +qed + +end \ \\<^locale>\G_generic1\\ + +sublocale G_generic1_AC \ ext:M_ZC_basic "M[G]" + using choice_ax choice_in_MG + by unfold_locales + +end \ No newline at end of file diff --git a/thys/Independence_CH/Cohen_Posets_Relative.thy b/thys/Independence_CH/Cohen_Posets_Relative.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Cohen_Posets_Relative.thy @@ -0,0 +1,658 @@ +section\Cohen forcing notions\ + +theory Cohen_Posets_Relative + imports + Forcing_Notions + Transitive_Models.Delta_System_Relative + Transitive_Models.Partial_Functions_Relative +begin + +locale cohen_data = + fixes \ I J::i + assumes zero_lt_kappa: "0<\" +begin + +lemmas zero_lesspoll_kappa = zero_lesspoll[OF zero_lt_kappa] + +end \ \\<^locale>\cohen_data\\ + +locale add_reals = cohen_data nat _ 2 + +subsection\Combinatorial results on Cohen posets\ + +sublocale cohen_data \ forcing_notion "Fn(\,I,J)" "Fnle(\,I,J)" 0 +proof + show "0 \ Fn(\, I, J)" + using zero_lt_kappa zero_in_Fn by simp + then + show "\p\Fn(\, I, J). \p, 0\ \ Fnle(\, I, J)" + unfolding preorder_on_def refl_def trans_on_def + by auto +next + show "preorder_on(Fn(\, I, J), Fnle(\, I, J))" + unfolding preorder_on_def refl_def trans_on_def + by blast +qed + +context cohen_data +begin + +lemma compat_imp_Un_is_function: + assumes "G \ Fn(\, I, J)" "\p q. p \ G \ q \ G \ compat(p,q)" + shows "function(\G)" + unfolding function_def +proof (intro allI ballI impI) + fix x y y' + assume "\x, y\ \ \G" "\x, y'\ \ \G" + then + obtain p q where "\x, y\ \ p" "\x, y'\ \ q" "p \ G" "q \ G" + by auto + moreover from this and assms + obtain r where "r \ Fn(\, I, J)" "r \ p" "r \ q" + using compatD[of p q] by force + moreover from this + have "function(r)" + using Fn_is_function by simp + ultimately + show "y = y'" + unfolding function_def by fastforce +qed + +(* MOVE THIS to an appropriate place *) +lemma filter_subset_notion: "filter(G) \ G \ Fn(\, I, J)" + unfolding filter_def by simp + +lemma Un_filter_is_function: "filter(G) \ function(\G)" + using compat_imp_Un_is_function filter_imp_compat[of G] + filter_subset_notion + by simp + +end \ \\<^locale>\cohen_data\\ + +locale M_cohen = M_delta + + assumes + countable_lepoll_assms2: + "M(A') \ M(A) \ M(b) \ M(f) \ separation(M, \y. \x\A'. y = \x, \ i. x \ if_range_F_else_F(\a. {p \ A . domain(p) = a}, b, f, i)\)" + and + countable_lepoll_assms3: + "M(A) \ M(f) \ M(b) \ M(D) \ M(r') \ M(A')\ + separation(M, \y. \x\A'. y = \x, \ i. x \ if_range_F_else_F(drSR_Y(r', D, A), b, f, i)\)" + +context M_cardinal_library +begin + +lemma lesspoll_nat_imp_lesspoll_rel: + assumes "A \ \" "M(A)" + shows "A \\<^bsup>M\<^esup> \" +proof - + note assms + moreover from this + obtain f n where "f\bij\<^bsup>M\<^esup>(A,n)" "n\\" "A \\<^bsup>M\<^esup> n" + using lesspoll_nat_is_Finite Finite_rel_def[of M A] + by auto + moreover from calculation + have "A \\<^bsup>M\<^esup> \" + using lesspoll_nat_is_Finite Infinite_imp_nats_lepoll_rel[of \ n] + nat_not_Finite eq_lepoll_rel_trans[of A n \] + by auto + moreover from calculation + have "\ g \ bij\<^bsup>M\<^esup>(A,\)" for g + using mem_bij_rel unfolding lesspoll_def by auto + ultimately + show ?thesis + unfolding lesspoll_rel_def + by auto +qed + +lemma Finite_imp_lesspoll_rel_nat: + assumes "Finite(A)" "M(A)" + shows "A \\<^bsup>M\<^esup> \" + using Finite_imp_lesspoll_nat assms lesspoll_nat_imp_lesspoll_rel + by auto + +lemma InfCard_rel_lesspoll_rel_Un: + includes Ord_dests + assumes "InfCard_rel(M,\)" "A \\<^bsup>M\<^esup> \" "B \\<^bsup>M\<^esup> \" + and types: "M(\)" "M(A)" "M(B)" + shows "A \ B \\<^bsup>M\<^esup> \" +proof - + from assms + have "|A|\<^bsup>M\<^esup> < \" "|B|\<^bsup>M\<^esup> < \" + using lesspoll_rel_cardinal_rel_lt InfCard_rel_is_Card_rel + by auto + show ?thesis + proof (cases "Finite(A) \ Finite(B)") + case True + with assms + show ?thesis + using lesspoll_rel_trans2[OF _ le_imp_lepoll_rel, of _ nat \] + Finite_imp_lesspoll_rel_nat[OF Finite_Un] + unfolding InfCard_rel_def + by simp + next + case False + with types + have "InfCard_rel(M,max(|A|\<^bsup>M\<^esup>,|B|\<^bsup>M\<^esup>))" + using Infinite_InfCard_rel_cardinal_rel InfCard_rel_is_Card_rel + le_trans[of nat] not_le_iff_lt[THEN iffD1, THEN leI, of "|A|\<^bsup>M\<^esup>" "|B|\<^bsup>M\<^esup>"] + unfolding max_def InfCard_rel_def + by auto + with \M(A)\ \M(B)\ + have "|A \ B|\<^bsup>M\<^esup> \ max(|A|\<^bsup>M\<^esup>,|B|\<^bsup>M\<^esup>)" + using cardinal_rel_Un_le[of "max(|A|\<^bsup>M\<^esup>,|B|\<^bsup>M\<^esup>)" A B] + not_le_iff_lt[THEN iffD1, THEN leI, of "|A|\<^bsup>M\<^esup>" "|B|\<^bsup>M\<^esup>" ] + unfolding max_def + by simp + moreover from \|A|\<^bsup>M\<^esup> < \\ \|B|\<^bsup>M\<^esup> < \\ + have "max(|A|\<^bsup>M\<^esup>,|B|\<^bsup>M\<^esup>) < \" + unfolding max_def + by simp + ultimately + have "|A \ B|\<^bsup>M\<^esup> < \" + using lt_trans1 + by blast + moreover + note \InfCard_rel(M,\)\ + moreover from calculation types + have "|A \ B|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> \" + using lt_Card_rel_imp_lesspoll_rel InfCard_rel_is_Card_rel + by simp + ultimately + show ?thesis + using cardinal_rel_eqpoll_rel eq_lesspoll_rel_trans[of "A \ B" "|A \ B|\<^bsup>M\<^esup>" \] + eqpoll_rel_sym types + by simp + qed +qed + +end \ \\<^locale>\M_cardinal_library\\ + +lemma (in M_cohen) Fn_rel_unionI: + assumes "p \ Fn\<^bsup>M\<^esup>(\,I,J)" "q\Fn\<^bsup>M\<^esup>(\,I,J)" "InfCard\<^bsup>M\<^esup>(\)" + "M(\)" "M(I)" "M(J)" "domain(p) \ domain(q) = 0" + shows "p\q \ Fn\<^bsup>M\<^esup>(\,I,J)" +proof - + note assms + moreover from calculation + have "p \\<^bsup>M\<^esup> \" "q \\<^bsup>M\<^esup> \" "M(p)" "M(q)" + using Fn_rel_is_function by simp_all + moreover from calculation + have "p\q \\<^bsup>M\<^esup> \" + using eqpoll_rel_sym cardinal_rel_eqpoll_rel InfCard_rel_lesspoll_rel_Un + by simp_all + ultimately + show ?thesis + unfolding Fn_rel_def + using pfun_unionI cardinal_rel_eqpoll_rel eq_lesspoll_rel_trans[OF _ \p\q \\<^bsup>M\<^esup> _\] + by auto +qed + +lemma (in M_cohen) restrict_eq_imp_compat_rel: + assumes "p \ Fn\<^bsup>M\<^esup>(\, I, J)" "q \ Fn\<^bsup>M\<^esup>(\, I, J)" "InfCard\<^bsup>M\<^esup>(\)" "M(J)" "M(\)" + "restrict(p, domain(p) \ domain(q)) = restrict(q, domain(p) \ domain(q))" + shows "p \ q \ Fn\<^bsup>M\<^esup>(\, I, J)" +proof - + note assms + moreover from calculation + have "p \\<^bsup>M\<^esup> \" "q \\<^bsup>M\<^esup> \" "M(p)" "M(q)" + using Fn_rel_is_function by simp_all + moreover from calculation + have "p\q \\<^bsup>M\<^esup> \" + using InfCard_rel_lesspoll_rel_Un cardinal_rel_eqpoll_rel[THEN eqpoll_rel_sym] + by auto + ultimately + show ?thesis + unfolding Fn_rel_def + using pfun_restrict_eq_imp_compat cardinal_rel_eqpoll_rel + eq_lesspoll_rel_trans[OF _ \p\q \\<^bsup>M\<^esup> _\] + by auto +qed + +lemma (in M_cohen) InfCard_rel_imp_n_lesspoll_rel : + assumes "InfCard\<^bsup>M\<^esup>(\)" "M(\)" "n\\" + shows "n \\<^bsup>M\<^esup> \" +proof - + note assms + moreover from this + have "n \\<^bsup>M\<^esup> \" + using n_lesspoll_rel_nat by simp + ultimately + show ?thesis + using lesspoll_rel_trans2 Infinite_iff_lepoll_rel_nat InfCard_rel_imp_Infinite nat_into_M + by simp +qed + +lemma (in M_cohen) cons_in_Fn_rel: + assumes "x \ domain(p)" "p \ Fn\<^bsup>M\<^esup>(\,I,J)" "x \ I" "j \ J" "InfCard\<^bsup>M\<^esup>(\)" + "M(\)" "M(I)" "M(J)" + shows "cons(\x,j\, p) \ Fn\<^bsup>M\<^esup>(\,I,J)" + using assms cons_eq Fn_rel_unionI[OF Fn_rel_singletonI[of x I j J] \p\_\] + InfCard_rel_imp_n_lesspoll_rel + by auto + +lemma (in M_library) Fnle_rel_Aleph_rel1_closed[intro,simp]: + "M(Fnle\<^bsup>M\<^esup>(\\<^bsub>1\<^esub>\<^bsup>M\<^esup>, \\<^bsub>1\<^esub>\<^bsup>M\<^esup>, \ \\<^bsup>M\<^esup> 2))" + by simp + +locale M_add_reals = M_cohen + add_reals +begin + +lemmas zero_lesspoll_rel_kappa = zero_lesspoll_rel[OF zero_lt_kappa] + +end \ \\<^locale>\M_add_reals\\ + +(* MOVE THIS to some appropriate place. Notice that in Forcing_Notions +we don't import anything relative. *) +relativize relational "compat_in" "is_compat_in" external +synthesize "compat_in" from_definition "is_compat_in" assuming "nonempty" +context + notes Un_assoc[simp] Un_trasposition_aux2[simp] +begin +arity_theorem for "compat_in_fm" +end + +lemma (in M_trivial) compat_in_abs[absolut]: + assumes + "M(A)" "M(r)" "M(p)" "M(q)" + shows + "is_compat_in(M,A,r,p,q) \ compat_in(A,r,p,q)" + using assms unfolding is_compat_in_def compat_in_def by simp + +definition + antichain :: "i\i\i\o" where + "antichain(P,leq,A) \ A\P \ (\p\A. \q\A. p\q \ \compat_in(P,leq,p,q))" + +relativize relational "antichain" "antichain_rel" +definition + ccc :: "i \ i \ o" where + "ccc(P,leq) \ \A. antichain(P,leq,A) \ |A| \ nat" + +abbreviation + antichain_rel_abbr :: "[i\o,i,i,i] \ o" (\antichain\<^bsup>_\<^esup>'(_,_,_')\) where + "antichain\<^bsup>M\<^esup>(P,leq,A) \ antichain_rel(M,P,leq,A)" + +abbreviation + antichain_r_set :: "[i,i,i,i] \ o" (\antichain\<^bsup>_\<^esup>'(_,_,_')\) where + "antichain\<^bsup>M\<^esup>(P,leq,A) \ antichain_rel(##M,P,leq,A)" + +context M_trivial +begin + +lemma antichain_abs [absolut]: + "\ M(A); M(P); M(leq) \ \ antichain\<^bsup>M\<^esup>(P,leq,A) \ antichain(P,leq,A)" + unfolding antichain_rel_def antichain_def by (simp add:absolut) + +end \ \\<^locale>\M_trivial\\ + +relativize relational "ccc" "ccc_rel" + +abbreviation + ccc_rel_abbr :: "[i\o,i,i]\o" (\ccc\<^bsup>_\<^esup>'(_,_')\) where + "ccc_rel_abbr(M) \ ccc_rel(M)" + +abbreviation + ccc_r_set :: "[i,i,i]\o" (\ccc\<^bsup>_\<^esup>'(_,_')\) where + "ccc_r_set(M) \ ccc_rel(##M)" + +context M_cardinals +begin + +lemma def_ccc_rel: + shows + "ccc\<^bsup>M\<^esup>(P,leq) \ (\A[M]. antichain\<^bsup>M\<^esup>(P,leq,A) \ |A|\<^bsup>M\<^esup> \ \)" + using is_cardinal_iff + unfolding ccc_rel_def by (simp add:absolut) + +end \ \\<^locale>\M_cardinals\\ + +context M_FiniteFun +begin + +lemma Fnle_nat_closed[intro,simp]: + assumes "M(I)" "M(J)" + shows "M(Fnle(\,I,J))" + unfolding Fnle_def Fnlerel_def Rrel_def + using supset_separation FiniteFun_closed Fn_nat_eq_FiniteFun assms by simp + +lemma Fn_nat_closed: + assumes "M(A)" "M(B)" shows "M(Fn(\,A,B))" + using assms Fn_nat_eq_FiniteFun + by simp + +end \ \\<^locale>\M_FiniteFun\\ + +context M_add_reals +begin + +lemma lam_replacement_drSR_Y: "M(A) \ M(D) \ M(r') \ lam_replacement(M, drSR_Y(r',D,A))" + using lam_replacement_drSR_Y + by simp + +lemma (in M_trans) mem_F_bound3: + fixes F A + defines "F \ dC_F" + shows "x\F(A,c) \ c \ (range(f) \ {domain(x). x\A})" + using apply_0 unfolding F_def + by (cases "M(c)", auto simp:F_def drSR_Y_def dC_F_def) + +lemma ccc_rel_Fn_nat: + assumes "M(I)" + shows "ccc\<^bsup>M\<^esup>(Fn(nat,I,2), Fnle(nat,I,2))" +proof - + have repFun_dom_closed:"M({domain(p) . p \ A})" if "M(A)" for A + using RepFun_closed domain_replacement_simp transM[OF _ \M(A)\] that + by auto + from assms + have "M(Fn(nat,I,2))" using Fn_nat_eq_FiniteFun by simp + { + fix A + assume "\ |A|\<^bsup>M\<^esup> \ nat" "M(A)" "A \ Fn(nat, I, 2)" + moreover from this + have "countable_rel(M,{p\A. domain(p) = d})" if "M(d)" for d + proof (cases "d\\<^bsup>M\<^esup>nat \ d \ I") + case True + with \A \ Fn(nat, I, 2)\ \M(A)\ + have "{p \ A . domain(p) = d} \ d \\<^bsup>M\<^esup> 2" + using domain_of_fun function_space_rel_char[of _ 2] + by (auto,subgoal_tac "M(domain(x))",simp_all add:transM[of _ A],force) + moreover from True \M(d)\ + have "Finite(d \\<^bsup>M\<^esup> 2)" + using Finite_Pi[THEN [2] subset_Finite, of _ d "\_. 2"] + lesspoll_rel_nat_is_Finite_rel function_space_rel_char[of _ 2] + by auto + moreover from \M(d)\ + have "M(d \\<^bsup>M\<^esup> 2)" + by simp + moreover from \M(A)\ + have "M({p \ A . domain(p) = d})" + using separation_closed domain_eq_separation[OF \M(d)\] by simp + ultimately + show ?thesis using subset_Finite[of _ "d\\<^bsup>M\<^esup>2" ] + by (auto intro!:Finite_imp_countable_rel) + next + case False + with \A \ Fn(nat, I, 2)\ \M(A)\ + have "domain(p) \ d" if "p\A" for p + proof - + note False that \M(A)\ + moreover from this + obtain d' where "d' \ I" "p\d' \ 2" "d' \ \" + using FnD[OF subsetD[OF \A\_\ \p\A\]] + by auto + moreover from this + have "p \ d'" "domain(p) = d'" + using function_eqpoll Pi_iff + by auto + ultimately + show ?thesis + using lesspoll_nat_imp_lesspoll_rel transM[of p] + by auto + qed + then + show ?thesis + using empty_lepoll_relI by auto + qed + have 2:"M(x) \ x \ dC_F(X, i) \ M(i)" for x X i + unfolding dC_F_def + by auto + moreover + have "uncountable_rel(M,{domain(p) . p \ A})" + proof + interpret M_replacement_lepoll M dC_F + using lam_replacement_dC_F domain_eq_separation lam_replacement_inj_rel + unfolding dC_F_def + proof(unfold_locales,simp_all) + fix X b f + assume "M(X)" "M(b)" "M(f)" + with 2[of X] + show "lam_replacement(M, \x. \ i. x \ if_range_F_else_F(\d. {p \ X . domain(p) = d}, b, f, i))" + using lam_replacement_dC_F domain_eq_separation + mem_F_bound3 countable_lepoll_assms2 repFun_dom_closed + by (rule_tac lam_Least_assumption_general[where U="\_. {domain(x). x\X}"],auto) + qed (auto) + have "\a\A. x = domain(a) \ M(dC_F(A,x))" for x + using \M(A)\ transM[OF _ \M(A)\] by (auto) + moreover + have "w \ A \ domain(w) = x \ M(x)" for w x + using transM[OF _ \M(A)\] by auto + ultimately + interpret M_cardinal_UN_lepoll _ "dC_F(A)" "{domain(p). p\A}" + using lam_replacement_dC_F lam_replacement_inj_rel \M(A)\ + lepoll_assumptions domain_eq_separation + countable_lepoll_assms2 repFun_dom_closed + lepoll_assumptions1[OF \M(A)\ repFun_dom_closed[OF \M(A)\]] + apply(unfold_locales) + by(simp_all del:if_range_F_else_F_def, + rule_tac lam_Least_assumption_general[where U="\_. {domain(x). x\A}"]) + (auto simp del:if_range_F_else_F_def simp add:dC_F_def) + from \A \ Fn(nat, I, 2)\ + have x:"(\d\{domain(p) . p \ A}. {p\A. domain(p) = d}) = A" + by auto + moreover + assume "countable_rel(M,{domain(p) . p \ A})" + moreover + note \\d. M(d) \ countable_rel(M,{p\A. domain(p) = d})\ + moreover from \M(A)\ + have "p\A \ M(domain(p))" for p + by (auto dest: transM) + ultimately + have "countable_rel(M,A)" + using countable_rel_imp_countable_rel_UN + unfolding dC_F_def + by auto + with \\ |A|\<^bsup>M\<^esup> \ nat\ \M(A)\ + show False + using countable_rel_iff_cardinal_rel_le_nat by simp + qed + moreover from \A \ Fn(nat, I, 2)\ \M(A)\ + have "p \ A \ Finite(domain(p))" for p + using lesspoll_rel_nat_is_Finite_rel[of "domain(p)"] + lesspoll_nat_imp_lesspoll_rel[of "domain(p)"] + domain_of_fun[of p _ "\_. 2"] by (auto dest:transM) + moreover + note repFun_dom_closed[OF \M(A)\] + ultimately + obtain D where "delta_system(D)" "D \ {domain(p) . p \ A}" "D \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "M(D)" + using delta_system_uncountable_rel[of "{domain(p) . p \ A}"] by auto + then + have delta:"\d1\D. \d2\D. d1 \ d2 \ d1 \ d2 = \D" + using delta_system_root_eq_Inter + by simp + moreover from \D \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ \M(D)\ + have "uncountable_rel(M,D)" + using uncountable_rel_iff_subset_eqpoll_rel_Aleph_rel1 by auto + moreover from this and \D \ {domain(p) . p \ A}\ + obtain p1 where "p1 \ A" "domain(p1) \ D" + using uncountable_rel_not_empty[of D] by blast + moreover from this and \p1 \ A \ Finite(domain(p1))\ + have "Finite(domain(p1))" + using Finite_domain by simp + moreover + define r where "r \ \D" + moreover from \M(D)\ + have "M(r)" "M(r\2)" + unfolding r_def by simp_all + ultimately + have "Finite(r)" using subset_Finite[of "r" "domain(p1)"] + by auto + have "countable_rel(M,{restrict(p,r) . p\A})" + proof - + note \M(Fn(nat, I, 2))\ \M(r)\ + moreover from this + have "f\Fn(nat, I, 2) \ M(restrict(f,r))" for f + by (blast dest: transM) + ultimately + have "f\Fn(nat, I, 2) \ restrict(f,r) \ Pow_rel(M,r \ 2)" for f + using restrict_subset_Sigma[of f _ "\_. 2" r] Pow_rel_char + by (auto del:FnD dest!:FnD simp: Pi_def) (auto dest:transM) + with \A \ Fn(nat, I, 2)\ + have "{restrict(f,r) . f \ A } \ Pow_rel(M,r \ 2)" + by fast + moreover from \M(A)\ \M(r)\ + have "M({restrict(f,r) . f \ A })" + using RepFun_closed restrict_strong_replacement transM[OF _ \M(A)\] by auto + moreover + note \Finite(r)\ \M(r)\ + ultimately + show ?thesis + using Finite_Sigma[THEN Finite_Pow_rel, of r "\_. 2"] + by (intro Finite_imp_countable_rel) (auto intro:subset_Finite) + qed + moreover from \M(A)\ \M(D)\ + have "M({p\A. domain(p) \ D})" + using domain_mem_separation by simp + have "uncountable_rel(M,{p\A. domain(p) \ D})" (is "uncountable_rel(M,?X)") + proof + from \D \ {domain(p) . p \ A}\ + have "(\p\?X. domain(p)) \ surj(?X, D)" + using lam_type unfolding surj_def by auto + moreover from \M(A)\ \M(?X)\ + have "M(\p\?X. domain(p))" + using lam_closed[OF domain_replacement \M(?X)\] transM[OF _ \M(?X)\] by simp + moreover + note \M(?X)\ \M(D)\ + moreover from calculation + have surjection:"(\p\?X. domain(p)) \ surj\<^bsup>M\<^esup>(?X, D)" + using surj_rel_char by simp + moreover + assume "countable_rel(M,?X)" + moreover + note \uncountable_rel(M,D)\ + ultimately + show False + using surj_rel_countable_rel[OF _ surjection] by auto + qed + moreover + have "D = (\f\Pow_rel(M,r\2) . {y . p\A, restrict(p,r) = f \ y=domain(p) \ domain(p) \ D})" + proof - + { + fix z + assume "z \ D" + with \M(D)\ + have \M(z)\ by (auto dest:transM) + from \z\D\ \D \ _\ \M(A)\ + obtain p where "domain(p) = z" "p \ A" "M(p)" + by (auto dest:transM) + moreover from \A \ Fn(nat, I, 2)\ \M(z)\ and this + have "p \ z \\<^bsup>M\<^esup> 2" + using domain_of_fun function_space_rel_char by (auto del:FnD dest!:FnD) + moreover from this \M(z)\ + have "p \ z \ 2" + using domain_of_fun function_space_rel_char by (auto) + moreover from this \M(r)\ + have "restrict(p,r) \ r \ 2" + using function_restrictI[of p r] fun_is_function[of p z "\_. 2"] + restrict_subset_Sigma[of p z "\_. 2" r] function_space_rel_char + by (auto simp:Pi_def) + moreover from \M(p)\ \M(r)\ + have "M(restrict(p,r))" by simp + moreover + note \M(r)\ + ultimately + have "\p\A. restrict(p,r) \ Pow_rel(M,r\2) \ domain(p) = z" + using Pow_rel_char by auto + } + then + show ?thesis + by (intro equalityI) (force)+ + qed + from \M(D)\\M(r)\ + have "M({y . p\A, restrict(p,r) = f \ y=domain(p) \ domain(p) \ D})" (is "M(?Y(A,f))") + if "M(f)" "M(A)" for f A + using drSR_Y_closed[unfolded drSR_Y_def] that + by simp + then + obtain f where "uncountable_rel(M,?Y(A,f))" "M(f)" + proof - + have 1:"M(i)" + if "M(B)" "M(x)" + "x \ {y . x \ B, restrict(x, r) = i \ y = domain(x) \ domain(x) \ D}" + for B x i + using that \M(r)\ + by (auto dest:transM) + note \M(r)\ + moreover from this + have "M(Pow\<^bsup>M\<^esup>(r \ 2))" by simp + moreover + note \M(A)\ \\f A. M(f) \ M(A) \ M(?Y(A,f))\ \M(D)\ + moreover from calculation + interpret M_replacement_lepoll M "drSR_Y(r,D)" + using countable_lepoll_assms3 lam_replacement_inj_rel lam_replacement_drSR_Y + drSR_Y_closed lam_Least_assumption_drSR_Y + by (unfold_locales,simp_all add:drSR_Y_def,rule_tac 1,simp_all) + from calculation + have "x \ Pow\<^bsup>M\<^esup>(r \ 2) \ M(drSR_Y(r, D, A, x))" for x + unfolding drSR_Y_def by (auto dest:transM) + ultimately + interpret M_cardinal_UN_lepoll _ "?Y(A)" "Pow_rel(M,r\2)" + using countable_lepoll_assms3 lam_replacement_drSR_Y + lepoll_assumptions[where S="Pow_rel(M,r\2)", unfolded lepoll_assumptions_defs] + lam_Least_assumption_drSR_Y[unfolded drSR_Y_def] + unfolding drSR_Y_def + apply unfold_locales + apply (simp_all add:lam_replacement_inj_rel del: if_range_F_else_F_def,rule_tac 1,simp_all) + by ((fastforce dest:transM[OF _ \M(A)\])+) + { + from \Finite(r)\ \M(r)\ + have "countable_rel(M,Pow_rel(M,r\2))" + using Finite_Sigma[THEN Finite_Pow_rel] Finite_imp_countable_rel + by simp + moreover + assume "M(f) \ countable_rel(M,?Y(A,f))" for f + moreover + note \D = (\f\Pow_rel(M,r\2) .?Y(A,f))\ \M(r)\ + moreover + note \uncountable_rel(M,D)\ + ultimately + have "False" + using countable_rel_imp_countable_rel_UN by (auto dest: transM) + } + with that + show ?thesis + by auto + qed + moreover from this \M(A)\ and \M(f) \ M(A) \ M(?Y(A,f))\ + have "M(?Y(A,f))" + by blast + ultimately + obtain j where "j \ inj_rel(M,nat, ?Y(A,f))" "M(j)" + using uncountable_rel_iff_nat_lt_cardinal_rel[THEN iffD1, THEN leI, + THEN cardinal_rel_le_imp_lepoll_rel, THEN lepoll_relD] + by auto + with \M(?Y(A,f))\ + have "j`0 \ j`1" "j`0 \ ?Y(A,f)" "j`1 \ ?Y(A,f)" + using inj_is_fun[THEN apply_type, of j nat "?Y(A,f)"] + inj_rel_char + unfolding inj_def by auto + then + obtain p q where "domain(p) \ domain(q)" "p \ A" "q \ A" + "domain(p) \ D" "domain(q) \ D" + "restrict(p,r) = restrict(q,r)" by auto + moreover from this and delta + have "domain(p) \ domain(q) = r" + unfolding r_def by simp + moreover + note \A \ Fn(nat, I, 2)\ Fn_nat_abs[OF \M(I)\ nat_into_M[of 2],simplified] + moreover from calculation + have "p \ Fn\<^bsup>M\<^esup>(nat, I, 2)" "q \ Fn\<^bsup>M\<^esup>(nat, I, 2)" + by auto + moreover from calculation + have "p \ q \ Fn(nat, I, 2)" + using restrict_eq_imp_compat_rel InfCard_rel_nat + by simp + ultimately + have "\p\A. \q\A. p \ q \ compat_in(Fn(nat, I, 2), Fnle(nat, I, 2), p, q)" + unfolding compat_in_def + by (rule_tac bexI[of _ p], rule_tac bexI[of _ q]) blast + } + moreover from assms + have "M(Fnle(\,I,2))" + by simp + moreover note \M(Fn(\,I,2))\ + ultimately + show ?thesis using def_ccc_rel by (auto simp:absolut antichain_def) fastforce +qed + +end \ \\<^locale>\M_add_reals\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Definitions_Main.thy b/thys/Independence_CH/Definitions_Main.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Definitions_Main.thy @@ -0,0 +1,610 @@ +section\Main definitions of the development\label{sec:main-definitions}\ + +theory Definitions_Main + imports + Absolute_Versions +begin + +text\This theory gathers the main definitions of the Forcing session. + +It might be considered as the bare minimum reading requisite to +trust that our development indeed formalizes the theory of +forcing. This should be mathematically clear since this is the +only known method for obtaining proper extensions of ctms while +preserving the ordinals. + +The main theorem of this session and all of its relevant definitions +appear in Section~\ref{sec:def-main-forcing}. The reader trusting +all the libraries in which our development is based, might jump +directly there. But in case one wants to dive deeper, the following +sections treat some basic concepts in the ZF logic +(Section~\ref{sec:def-main-ZF}) and in the +ZF-Constructible library (Section~\ref{sec:def-main-relative}) +on which our definitions are built. +\ + +declare [[show_question_marks=false]] + +subsection\ZF\label{sec:def-main-ZF}\ + +text\For the basic logic ZF we restrict ourselves to just a few +concepts.\ + +thm bij_def[unfolded inj_def surj_def] +text\@{thm [display] bij_def[unfolded inj_def surj_def]}\ +(* + bij(A, B) \ {f \ A \ B . \w\A. \x\A. f ` w = f ` x \ w = x} + \ {f \ A \ B . \y\B. \x\A. f ` x = y} +*) + +thm eqpoll_def +text\@{thm [display] eqpoll_def}\ +(* + A \ B \ \f. f \ bij(A, B) +*) + +thm Transset_def +text\@{thm [display] Transset_def}\ +(* + Transset(i) \ \x\i. x \ i +*) + +thm Ord_def +text\@{thm [display] Ord_def}\ +(* + Ord(i) \ Transset(i) \ (\x\i. Transset(x)) +*) + +thm lt_def le_iff +text\@{thm [display] lt_def le_iff}\ +(* + i < j \ i \ j \ Ord(j) + i \ j \ i < j \ i = j \ Ord(j) +*) + +text\With the concepts of empty set and successor in place,\ +lemma empty_def': "\x. x \ 0" by simp +lemma succ_def': "succ(i) = i \ {i}" by blast + +text\we can define the set of natural numbers \<^term>\\\. In the +sources, it is defined as a fixpoint, but here we just write +its characterization as the first limit ordinal.\ +thm Limit_nat[unfolded Limit_def] nat_le_Limit[unfolded Limit_def] +text\@{thm [display] Limit_nat[unfolded Limit_def] + nat_le_Limit[unfolded Limit_def]}\ +(* + Ord(\) \ 0 < \ \ (\y. y < \ \ succ(y) < \) + Ord(i) \ 0 < i \ (\y. y < i \ succ(y) < i) \ \ \ i +*) + +text\Then, addition and predecessor are inductively characterized +as follows:\ +thm add_0_right add_succ_right pred_0 pred_succ_eq +text\@{thm [display] add_succ_right add_0_right pred_0 pred_succ_eq}\ +(* + m \ \ \ m +\<^sub>\ 0 = m + m +\<^sub>\ succ(n) = succ(m +\<^sub>\ n) + + pred(0) = 0 + pred(succ(y)) = y +*) + +text\Lists on a set \<^term>\A\ can be characterized by being +recursively generated from the empty list \<^term>\[]\ and the +operation \<^term>\Cons\ that adds a new element to the left end; +the induction theorem for them shows that the characterization is +“complete”.\ + +thm Nil Cons list.induct +text\@{thm [display] Nil Cons list.induct }\ +(* + [] \ list(A) + a \ A \ l \ list(A) \ Cons(a, l) \ list(A) + x \ list(A) \ P([]) \ (\a l. a \ A \ l \ list(A) \ P(l) \ P(Cons(a, l))) \ P(x) +*) + +text\Length, concatenation, and \<^term>\n\th element of lists are +recursively characterized as follows.\ +thm length.simps app.simps nth_0 nth_Cons +text\@{thm [display] length.simps app.simps nth_0 nth_Cons}\ +(* + length([]) = 0 + length(Cons(a, l)) = succ(length(l)) + + [] @ ys = ys + Cons(a, l) @ ys = Cons(a, l @ ys) + + nth(0, Cons(a, l)) = a + n \ \ \ nth(succ(n), Cons(a, l)) = nth(n, l) +*) +txt\We have the usual Haskell-like notation for iterated applications +of \<^term>\Cons\:\ +lemma Cons_app: "[a,b,c] = Cons(a,Cons(b,Cons(c,[])))" .. + +txt\Relative quantifiers restrict the range of the bound variable to a +class \<^term>\M\ of type \<^typ>\i\o\; that is, a truth-valued function with +set arguments.\ +lemma "\x[M]. P(x) \ \x. M(x) \ P(x)" + "\x[M]. P(x) \ \x. M(x) \ P(x)" + unfolding rall_def rex_def . + +txt\Finally, a set can be viewed (“cast”) as a class using the +following function of type \<^typ>\i\(i\o)\.\ +thm setclass_iff +text\@{thm [display] setclass_iff}\ +(* + (##A)(x) \ x \ A +*) + +subsection\Relative concepts\label{sec:def-main-relative}\ +txt\A list of relative concepts (mostly from the ZF-Constructible + library) follows next.\ + +thm big_union_def +text\@{thm [display] big_union_def}\ +(* + big_union(M, A, z) \ \x[M]. x \ z \ (\y[M]. y \ A \ x \ y) +*) + +thm upair_def +text\@{thm [display] upair_def}\ +(* + upair(M, a, b, z) \ a \ z \ b \ z \ (\x[M]. x \ z \ x = a \ x = b) +*) + +thm pair_def +text\@{thm [display] pair_def}\ +(* + pair(M, a, b, z) \ \x[M]. upair(M, a, a, x) \ + (\y[M]. upair(M, a, b, y) \ upair(M, x, y, z)) +*) + +thm successor_def[unfolded is_cons_def union_def] +text\@{thm [display] successor_def[unfolded is_cons_def union_def]}\ +(* + successor(M, a, z) \ \x[M]. upair(M, a, a, x) \ (\xa[M]. xa \ z \ xa \ x \ xa \ a) +*) + +thm empty_def +text\@{thm [display] empty_def}\ +(* + empty(M, z) \ \x[M]. x \ z +*) + +thm transitive_set_def[unfolded subset_def] +text\@{thm [display] transitive_set_def[unfolded subset_def]}\ +(* + transitive_set(M, a) \ \x[M]. x \ a \ (\xa[M]. xa \ x \ xa \ a) +*) + + +thm ordinal_def +text\@{thm [display] ordinal_def}\ +(* + ordinal(M, a) \ transitive_set(M, a) \ (\x[M]. x \ a \ + transitive_set(M, x)) +*) + +thm image_def +text\@{thm [display] image_def}\ +(* + image(M, r, A, z) \ \y[M]. y \ z \ + (\w[M]. w \ r \ (\x[M]. x \ A \ pair(M, x, y, w))) +*) + +thm fun_apply_def +text\@{thm [display] fun_apply_def}\ +(* + fun_apply(M, f, x, y) \ \xs[M]. \fxs[M]. upair(M, x, x, xs) \ + image(M, f, xs, fxs) \ big_union(M, fxs, y) +*) + +thm is_function_def +text\@{thm [display] is_function_def}\ +(* + is_function(M, r) \ \x[M]. \y[M]. \y'[M]. \p[M]. \p'[M]. + pair(M, x, y, p) \ pair(M, x, y', p') \ p \ r \ p' \ r \ y = y' +*) + +thm is_relation_def +text\@{thm [display] is_relation_def}\ +(* + is_relation(M, r) \ \z[M]. z \ r \ (\x[M]. \y[M]. pair(M, x, y, z)) +*) + +thm is_domain_def +text\@{thm [display] is_domain_def}\ +(* + is_domain(M, r, z) \ \x[M]. x \ z \ + (\w[M]. w \ r \ (\y[M]. pair(M, x, y, w))) +*) + +thm typed_function_def +text\@{thm [display] typed_function_def}\ +(* + typed_function(M, A, B, r) \ is_function(M, r) \ is_relation(M, r) \ + is_domain(M, r, A) \ + (\u[M]. u \ r \ (\x[M]. \y[M]. pair(M, x, y, u) \ y \ B)) +*) + +thm is_function_space_def[unfolded is_funspace_def] + function_space_rel_def surjection_def +text\@{thm [display] is_function_space_def[unfolded is_funspace_def] + function_space_rel_def surjection_def}\ +(* + is_function_space(M, A, B, fs) \ + M(fs) \ (\f[M]. f \ fs \ typed_function(M, A, B, f)) + + A \\<^bsup>M\<^esup> B \ THE d. is_function_space(M, A, B, d) + + surjection(M, A, B, f) \ + typed_function(M, A, B, f) \ + (\y[M]. y \ B \ (\x[M]. x \ A \ is_apply(M, f, x, y))) +*) + + +txt\Relative version of the $\ZFC$ axioms\ +thm extensionality_def +text\@{thm [display] extensionality_def}\ +(* + extensionality(M) \ \x[M]. \y[M]. (\z[M]. z \ x \ z \ y) \ x = y +*) + +thm foundation_ax_def +text\@{thm [display] foundation_ax_def}\ +(* + foundation_ax(M) \ \x[M]. (\y[M]. y \ x) \ (\y[M]. y \ x \ \ (\z[M]. z \ x \ z \ y)) +*) + +thm upair_ax_def +text\@{thm [display] upair_ax_def}\ +(* + upair_ax(M) \ \x[M]. \y[M]. \z[M]. upair(M, x, y, z) +*) + +thm Union_ax_def +text\@{thm [display] Union_ax_def}\ +(* + Union_ax(M) \ \x[M]. \z[M]. \xa[M]. xa \ z \ (\y[M]. y \ x \ xa \ y) +*) + +thm power_ax_def[unfolded powerset_def subset_def] +text\@{thm [display] power_ax_def[unfolded powerset_def subset_def]}\ +(* + power_ax(M) \ \x[M]. \z[M]. \xa[M]. xa \ z \ (\xb[M]. xb \ xa \ xb \ x) +*) + +thm infinity_ax_def +text\@{thm [display] infinity_ax_def}\ +(* + infinity_ax(M) \ \I[M]. (\z[M]. empty(M, z) \ z \ I) \ (\y[M]. y \ I \ + (\sy[M]. successor(M, y, sy) \ sy \ I)) +*) + +thm choice_ax_def +text\@{thm [display] choice_ax_def}\ +(* + choice_ax(M) \ \x[M]. \a[M]. \f[M]. ordinal(M, a) \ surjection(M, a, x, f) +*) + +thm separation_def +text\@{thm [display] separation_def}\ +(* + separation(M, P) \ \z[M]. \y[M]. \x[M]. x \ y \ x \ z \ P(x) +*) + +thm univalent_def +text\@{thm [display] univalent_def}\ +(* + univalent(M, A, P) \ \x[M]. x \ A \ + (\y[M]. \z[M]. P(x, y) \ P(x, z) \ y = z) +*) + +thm strong_replacement_def +text\@{thm [display] strong_replacement_def}\ +(* + strong_replacement(M, P) \ \A[M]. univalent(M, A, P) \ + (\Y[M]. \b[M]. b \ Y \ (\x[M]. x \ A \ P(x, b))) +*) + +text\Internalized formulas\ + +txt\“Codes” for formulas (as sets) are constructed from natural +numbers using \<^term>\Member\, \<^term>\Equal\, \<^term>\Nand\, +and \<^term>\Forall\.\ + +thm Member Equal Nand Forall formula.induct +text\@{thm [display] Member Equal Nand Forall formula.induct}\ +(* + x \ \ \ y \ \ \ \x \ y\ \ formula + x \ \ \ y \ \ \ \x = y\ \ formula + p \ formula \ q \ formula \ \\(p \ q)\ \ formula + p \ formula \ (\p) \ formula + + x \ formula \ + (\x y. x \ \ \ y \ \ \ P(\x \ y\)) \ + (\x y. x \ \ \ y \ \ \ P(\x = y\)) \ + (\p q. p \ formula \ P(p) \ q \ formula \ P(q) \ P(\\(p \ q)\)) \ + (\p. p \ formula \ P(p) \ P((\p))) \ P(x) +*) + +txt\Definitions for the other connectives and the internal existential +quantifier are also provided. For instance, negation:\ +thm Neg_def +text\@{thm [display] Neg_def}\ + +thm arity.simps +text\@{thm [display] arity.simps}\ +(* + arity(\x \ y\) = succ(x) \ succ(y) + arity(\x = y\) = succ(x) \ succ(y) + arity(\\(p \ q)\) = arity(p) \ arity(q) + arity((\p)) = pred(arity(p)) +*) + +txt\We have the satisfaction relation between $\in$-models and + first order formulas (given a “environment” list representing + the assignment of free variables),\ +thm mem_iff_sats equal_iff_sats sats_Nand_iff sats_Forall_iff +text\@{thm [display] mem_iff_sats equal_iff_sats sats_Nand_iff sats_Forall_iff}\ +(* + nth(i, env) = x \ nth(j, env) = y \ env \ list(A) \ x \ y \ A, env \ \i \ j\ + nth(i, env) = x \ nth(j, env) = y \ env \ list(A) \ x = y \ A, env \ \i = j\ + env \ list(A) \ (A, env \ \\(p \ q)\) \ \ ((A, env \ p) \ (A, env \ q)) + env \ list(A) \ (A, env \ (\\p\)) \ (\x\A. A, Cons(x, env) \ p)*) + +txt\as well as the satisfaction of an arbitrary set of sentences.\ +thm satT_def +text\@{thm [display] satT_def}\ +(* + A \ \ \ \\\\. A, [] \ \ +*) + +txt\The internalized (viz. as elements of the set \<^term>\formula\) + version of the axioms follow next.\ + +thm ZF_union_iff_sats ZF_power_iff_sats ZF_pairing_iff_sats + ZF_foundation_iff_sats ZF_extensionality_iff_sats + ZF_infinity_iff_sats sats_ZF_separation_fm_iff + sats_ZF_replacement_fm_iff ZF_choice_iff_sats +text\@{thm [display] ZF_union_iff_sats ZF_power_iff_sats + ZF_pairing_iff_sats + ZF_foundation_iff_sats ZF_extensionality_iff_sats + ZF_infinity_iff_sats sats_ZF_separation_fm_iff + sats_ZF_replacement_fm_iff ZF_choice_iff_sats}\ +(* + Union_ax(##A) \ A, [] \ \Union Ax\ + power_ax(##A) \ A, [] \ \Powerset Ax\ + upair_ax(##A) \ A, [] \ \Pairing\ + foundation_ax(##A) \ A, [] \ \Foundation\ + extensionality(##A) \ A, [] \ \Extensionality\ + infinity_ax(##A) \ A, [] \ \Infinity\ + + \ \ formula \ + (M, [] \ \Separation(\)\) \ + (\env\list(M). + arity(\) \ 1 +\<^sub>\ length(env) \ separation(##M, \x. M, [x] @ env \ \)) + + \ \ formula \ + (M, [] \ \Replacement(\)\) \ + (\env\list(M). + arity(\) \ 2 +\<^sub>\ length(env) \ + strong_replacement(##M, \x y. M, [x, y] @ env \ \)) + + choice_ax(##A) \ A, [] \ \AC\ +*) + +thm ZF_fin_def ZF_schemes_def Zermelo_fms_def ZC_def ZF_def ZFC_def +text\@{thm [display] ZF_fin_def ZF_schemes_def Zermelo_fms_def ZC_def ZF_def + ZFC_def}\ +(* + ZF_fin \ {\Extensionality\, \Foundation\, \Pairing\, \Union Ax\, \Infinity\, \Powerset Ax\} + ZF_schemes \ {\Separation(p)\ . p \ formula} \ {\Replacement(p)\ . p \ formula} + \Z\ \ ZF_fin \ {\Separation(p)\ . p \ formula} + ZC \ \Z\ \ {\AC\} + ZF \ ZF_schemes \ ZF_fin + ZFC \ ZF \ {\AC\} +*) + +subsection\Relativization of infinitary arithmetic\ + +txt\In order to state the defining property of the relative + equipotence relation, we work under the assumptions of the + locale \<^term>\M_cardinals\. They comprise a finite set + of instances of Separation and Replacement to prove + closure properties of the transitive class \<^term>\M\.\ + +lemma (in M_cardinals) eqpoll_def': + assumes "M(A)" "M(B)" shows "A \\<^bsup>M\<^esup> B \ (\f[M]. f \ bij(A,B))" + using assms unfolding eqpoll_rel_def by auto + +txt\Below, $\mu$ denotes the minimum operator on the ordinals.\ +lemma cardinalities_defs: + fixes M::"i\o" + shows + "|A|\<^bsup>M\<^esup> \ \ i. M(i) \ i \\<^bsup>M\<^esup> A" + "Card\<^bsup>M\<^esup>(\) \ \ = |\|\<^bsup>M\<^esup>" + "\\<^bsup>\\,M\<^esup> \ |\ \\<^bsup>M\<^esup> \|\<^bsup>M\<^esup>" + "(\\<^sup>+)\<^bsup>M\<^esup> \ \ x. M(x) \ Card\<^bsup>M\<^esup>(x) \ \ < x" + unfolding cardinal_rel_def cexp_rel_def + csucc_rel_def Card_rel_def . + +context M_aleph +begin + +txt\As in the previous Lemma @{thm [source] eqpoll_def'}, we are now under + the assumptions of the locale \<^term>\M_aleph\. The axiom instances + included are sufficient to state and prove the defining + properties of the relativized \<^term>\Aleph\ function + (in particular, the required ability to perform transfinite recursions).\ + +thm Aleph_rel_zero Aleph_rel_succ Aleph_rel_limit +text\@{thm [display] Aleph_rel_zero Aleph_rel_succ Aleph_rel_limit}\ +(* + \\<^bsub>0\<^esub>\<^bsup>M\<^esup> = \ + Ord(\) \ M(\) \ \\<^bsub>succ(\)\<^esub>\<^bsup>M\<^esup> = (\\<^bsub>\\<^esub>\<^bsup>M\<^esup>\<^sup>+)\<^bsup>M\<^esup> + Limit(\) \ M(\) \ \\<^bsub>\\<^esub>\<^bsup>M\<^esup> = (\j\\. \\<^bsub>j\<^esub>\<^bsup>M\<^esup>) +*) + +end \ \\<^locale>\M_aleph\\ + +lemma ContHyp_rel_def': + fixes N::"i\o" + shows + "CH\<^bsup>N\<^esup> \ \\<^bsub>1\<^esub>\<^bsup>N\<^esup> = 2\<^bsup>\\\<^bsub>0\<^esub>\<^bsup>N\<^esup>,N\<^esup>" + unfolding ContHyp_rel_def . + +txt\Under appropriate hypothesis (this time, from the locale \<^term>\M_master\), + \<^term>\CH\<^bsup>M\<^esup>\ is equivalent to its fully relational version \<^term>\is_ContHyp\. + As a sanity check, we see that if the transitive class is indeed \<^term>\\\, + we recover the original $\CH$.\ + +thm M_master.is_ContHyp_iff is_ContHyp_iff_CH[unfolded ContHyp_def] +text\@{thm [display] M_master.is_ContHyp_iff + is_ContHyp_iff_CH[unfolded ContHyp_def]}\ +(* + M_master(M) \ is_ContHyp(M) \ CH\<^bsup>M\<^esup> + is_ContHyp(\) \ \\<^bsub>1\<^esub> = 2\<^bsup>\\\<^bsub>0\<^esub>\<^esup> +*) + +txt\In turn, the fully relational version evaluated on a nonempty + transitive \<^term>\A\ is equivalent to the satisfaction of the + first-order formula \<^term>\\CH\\.\ +thm is_ContHyp_iff_sats +text\@{thm [display] is_ContHyp_iff_sats}\ +(* + env \ list(A) \ 0 \ A \ is_ContHyp(##A) \ A, env \ \CH\ +*) + + +subsection\Forcing \label{sec:def-main-forcing}\ + +txt\Our first milestone was to obtain a proper extension using forcing. +It's original proof didn't required the previous developments involving +the relativization of material on cardinal arithmetic. Now it is +derived from a stronger result, namely @{thm [source] extensions_of_ctms} +below.\ + +thm extensions_of_ctms_ZF +text\@{thm [display] extensions_of_ctms_ZF}\ +(* + M \ \ \ + Transset(M) \ + M \ ZF \ + \N. M \ N \ N \ \ \ Transset(N) \ N \ ZF \ M \ N \ + (\\. Ord(\) \ \ \ M \ \ \ N) \ ((M, [] \ \AC\) \ N \ ZFC) +*) + +txt\We can finally state our main results, namely, the existence of models +for $\ZFC + \CH$ and $\ZFC + \neg\CH$ under the assumption of a ctm of $\ZFC$.\ + +thm ctm_ZFC_imp_ctm_not_CH +text\@{thm [display] ctm_ZFC_imp_ctm_not_CH}\ +(* + M \ \ \ + Transset(M) \ + M \ ZFC \ + \N. M \ N \ + N \ \ \ Transset(N) \ N \ ZFC \ {\\\CH\\} \ + (\\. Ord(\) \ \ \ M \ \ \ N) +*) + +thm ctm_ZFC_imp_ctm_CH +text\@{thm [display] ctm_ZFC_imp_ctm_CH}\ +(* + M \ \ \ + Transset(M) \ + M \ ZFC \ + \N. M \ N \ + N \ \ \ + Transset(N) \ N \ ZFC \ {\CH\} \ (\\. Ord(\) \ \ \ M \ \ \ N) +*) + +txt\These results can be strengthened by enumerating four finite sets of +replacement instances which are sufficient to develop forcing and for +the construction of the aforementioned models: \<^term>\instances1_fms\ +through \<^term>\instances4_fms\, which are then collected into +\<^term>\overhead\. For example, we have:\ + +thm instances1_fms_def +text\@{thm [display] instances1_fms_def}\ +(* +instances1_fms \ +{ wfrec_Hfrc_at_fm, list_repl1_intf_fm, list_repl2_intf_fm, + formula_repl2_intf_fm, eclose_repl2_intf_fm, powapply_repl_fm, + phrank_repl_fm, wfrec_rank_fm, trans_repl_HVFrom_fm, wfrec_Hcheck_fm, + repl_PHcheck_fm, check_replacement_fm, G_dot_in_M_fm, repl_opname_check_fm, + tl_repl_intf_fm, formula_repl1_intf_fm, eclose_repl1_intf_fm } +*) + +thm overhead_def +text\@{thm [display] overhead_def}\ +(* +overhead \ instances1_fms \ instances2_fms \ instances3_fms \ instances4_fms +*) + +thm extensions_of_ctms +text\@{thm [display] extensions_of_ctms}\ +(* +M \ \ \ +Transset(M) \ +M \ \Z\ \ {\Replacement(p)\ . p \ instances1_fms \ instances2_fms} \ +\ \ formula \ +M \ {\Replacement(ground_repl_fm(\))\ . \ \ \} \ +\N. M \ N \ + N \ \ \ + Transset(N) \ + M \ N \ + (\\. Ord(\) \ \ \ M \ \ \ N) \ + ((M, [] \ \AC\) \ N, [] \ \AC\) \ + N \ \Z\ \ {\Replacement(\)\ . \ \ \} +*) + +thm ctm_of_not_CH +text\@{thm [display] ctm_of_not_CH}\ +(* +M \ \ \ +Transset(M) \ +M \ ZC \ {\Replacement(p)\ . p \ overhead} \ +\ \ formula \ +M \ {\Replacement(ground_repl_fm(\))\ . \ \ \} \ +\N. M \ N \ + N \ \ \ + Transset(N) \ + N \ ZC \ {\\\CH\\} \ {\Replacement(\)\ . \ \ \} \ + (\\. Ord(\) \ \ \ M \ \ \ N)*) + +thm ctm_of_CH +text\@{thm [display] ctm_of_CH}\ +(* +M \ \ \ +Transset(M) \ +M \ ZC \ {\Replacement(p)\ . p \ overhead} \ +\ \ formula \ +M \ {\Replacement(ground_repl_fm(\))\ . \ \ \} \ +\N. M \ N \ + N \ \ \ + Transset(N) \ + N \ ZC \ {\CH\} \ {\Replacement(\)\ . \ \ \} \ + (\\. Ord(\) \ \ \ M \ \ \ N) +*) + +txt\In the above three statements, the function \<^term>\ground_repl_fm\ +takes an element \<^term>\\\of \<^term>\formula\ and returns the +replacement instance in the ground model that produces the +\<^term>\\\-replacement instance in the generic extension. The next +result is stated in the context \<^locale>\G_generic1\, which assumes +the existence of a generic filter.\ + +context G_generic1 +begin + +thm sats_ground_repl_fm_imp_sats_ZF_replacement_fm +text\@{thm [display] sats_ground_repl_fm_imp_sats_ZF_replacement_fm}\ +(* +\ \ formula \ M, [] \ \Replacement(ground_repl_fm(\))\ \ M[G], [] \ \Replacement(\)\ +*) + +end \ \\<^locale>\G_generic1\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Demonstrations.thy b/thys/Independence_CH/Demonstrations.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Demonstrations.thy @@ -0,0 +1,166 @@ +section\Some demonstrations\ + +theory Demonstrations + imports + Definitions_Main +begin + +txt\The following theory is only intended to explore some details of the +formalization and to show the appearance of relevant internalized formulas. +It is \<^bold>\not\ intended as the entry point of the session. For that purpose, +consult \<^theory>\Independence_CH.Definitions_Main\\ + +locale Demo = M_trivial + M_AC + + fixes t\<^sub>1 t\<^sub>2 + assumes + ts_in_nat[simp]: "t\<^sub>1\\" "t\<^sub>2\\" + and + power_infty: "power_ax(M)" "M(\)" +begin + +txt\The next fake lemma is intended to explore the instances of the axiom +schemes that are needed to build our forcing models. They are categorized as +plain replacements (using \<^term>\strong_replacement\), “lambda-replacements” with +using a higher order function, replacements to perform +transfinite and general well-founded recursion (using \<^term>\transrec_replacement\ and +\<^term>\wfrec_replacement\ respectively) and for the construction of fixpoints +(using \<^term>\iterates_replacement\). Lastly, separations instances.\ + +lemma + assumes + sorried_replacements: + "\P. strong_replacement(M,P)" + "\F. lam_replacement(M,F)" + "\Q S. iterates_replacement(M,Q,S)" + "\Q S. wfrec_replacement(M,Q,S)" + "\Q S. transrec_replacement(M,Q,S)" + and + sorried_separations: + "\Q. separation(M,Q)"shows "M_master(M)" + apply unfold_locales + apply + (simp_all add: + sorried_replacements(1-2) + sorried_separations + power_infty) + oops + +\ \NOTE: Only for pretty-printing purposes, overrides previous + fundamental notations\ +no_notation mem (infixl \\\ 50) +no_notation conj (infixr \\\ 35) +no_notation disj (infixr \\\ 30) +no_notation iff (infixr \\\ 25) +no_notation imp (infixr \\\ 25) +no_notation not (\\ _\ [40] 40) +no_notation All (\'(\_')\) +no_notation Ex (\'(\_')\) + +no_notation Member (\\_ \/ _\\) +no_notation Equal (\\_ =/ _\\) +no_notation Nand (\\\'(_ \/ _')\\) +no_notation And (\\_ \/ _\\) +no_notation Or (\\_ \/ _\\) +no_notation Iff (\\_ \/ _\\) +no_notation Implies (\\_ \/ _\\) +no_notation Neg (\\\_\\) +no_notation Forall (\'(\\(/_)\')\) +no_notation Exists (\'(\\(/_)\')\) + +notation Member (infixl \\\ 50) +notation Equal (infixl \\\ 50) +notation Nand (\\'(_ \/ _')\) +notation And (infixr \\\ 35) +notation Or (infixr \\\ 30) +notation Iff (infixr \\\ 25) +notation Implies (infixr \\\ 25) +notation Neg (\\ _\ [40] 40) +notation Forall (\'(\_')\) +notation Exists (\'(\_')\) + +lemma "forces(t\<^sub>1\t\<^sub>2) = (0 \ 1 \ forces_mem_fm(1, 2, 0, t\<^sub>1+\<^sub>\4, t\<^sub>2+\<^sub>\4))" + unfolding forces_def by simp + +(* +\ \Prefix abbreviated notation\ +notation Member (\M\) +notation Equal (\Eq\) +notation Nand (\Na\) +notation And (\A\) +notation Or (\O\) +notation Iff (\If\) +notation Implies (\Im\) +notation Neg (\Ne\) +notation Forall (\Fo\) +notation Exists (\Ex\) +*) + +(* forces_mem_fm(1, 2, 0, t\<^sub>1+\<^sub>\4, t\<^sub>1+\<^sub>\4) + = forces_mem_fm(1, 2, 0, succ(succ(succ(succ(t\<^sub>1)))), succ(succ(succ(succ(t\<^sub>2))))) + = \ *) + +definition forces_0_mem_1 where "forces_0_mem_1\forces_mem_fm(1,2,0,t\<^sub>1+\<^sub>\4,t\<^sub>2+\<^sub>\4)" +thm forces_0_mem_1_def[ + unfolded frc_at_fm_def ftype_fm_def + name1_fm_def name2_fm_def snd_snd_fm_def hcomp_fm_def + ecloseN_fm_def eclose_n1_fm_def eclose_n2_fm_def + is_eclose_fm_def mem_eclose_fm_def eclose_n_fm_def + is_If_fm_def least_fm_def Replace_fm_def Collect_fm_def + fm_definitions,simplified] + (* NOTE: in view of the above, @{thm fm_definitions} might be incomplete *) + +named_theorems incr_bv_new_simps + +schematic_goal incr_bv_Neg(* [incr_bv_new_simps] *): + "mem(n,\) \ mem(\,formula) \ incr_bv(Neg(\))`n = ?x" + unfolding Neg_def by simp + +schematic_goal incr_bv_Exists [incr_bv_new_simps]: + "mem(n,\) \ mem(\,formula) \ incr_bv(Exists(\))`n = ?x" + unfolding Exists_def by (simp add: incr_bv_Neg) +(* +schematic_goal incr_bv_And [incr_bv_new_simps]: + "mem(n,\) \ mem(\,formula) \mem(\,formula)\ incr_bv(And(\,\))`n = ?x" + unfolding And_def by (simp add: incr_bv_Neg) + +schematic_goal incr_bv_Or [incr_bv_new_simps]: + "mem(n,\) \ mem(\,formula) \mem(\,formula)\ incr_bv(Or(\,\))`n = ?x" + unfolding Or_def by (simp add: incr_bv_Neg) + +schematic_goal incr_bv_Implies [incr_bv_new_simps]: + "mem(n,\) \ mem(\,formula) \mem(\,formula)\ incr_bv(Implies(\,\))`n = ?x" + unfolding Implies_def by (simp add: incr_bv_Neg) +*) + +\ \The two renamings involved in the definition of \<^term>\forces\ depend on +the recursive function \<^term>\incr_bv\. Here we have an apparently +exponential bottleneck, since all the propositional connectives (even \<^term>\Neg\) +duplicate the appearances of \<^term>\incr_bv\. + +Not even the double negation of an atomic formula can be managed by the system.\ +(* + +schematic_goal "forces(\\0\1) = ?x" + unfolding forces_def Neg_def + by (simp add:ren_forces_nand_def ren_forces_forall_def + frc_at_fm_def ftype_fm_def + name1_fm_def name2_fm_def snd_snd_fm_def hcomp_fm_def + ecloseN_fm_def eclose_n1_fm_def eclose_n2_fm_def + is_eclose_fm_def mem_eclose_fm_def eclose_n_fm_def + is_If_fm_def least_fm_def Collect_fm_def + fm_definitions incr_bv_Neg incr_bv_Exists) +(* exception Size raised (line 183 of "./basis/LibrarySupport.sml") *) + +*) + +(* +declare is_ContHyp_fm_def[fm_definitions del] + +thm is_ContHyp_fm_def[unfolded is_eclose_fm_def mem_eclose_fm_def eclose_n_fm_def + is_If_fm_def least_fm_def Replace_fm_def Collect_fm_def + fm_definitions, simplified] *) + +end \ \\<^locale>\Demo\\ + + +end \ No newline at end of file diff --git a/thys/Independence_CH/Edrel.thy b/thys/Independence_CH/Edrel.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Edrel.thy @@ -0,0 +1,257 @@ +theory Edrel + imports + Transitive_Models.ZF_Miscellanea + Transitive_Models.Recursion_Thms + +begin + +subsection\The well-founded relation \<^term>\ed\\ + +lemma eclose_sing : "x \ eclose(a) \ x \ eclose({a})" + using subsetD[OF mem_eclose_subset] + by simp + +lemma ecloseE : + assumes "x \ eclose(A)" + shows "x \ A \ (\ B \ A . x \ eclose(B))" + using assms +proof (induct rule:eclose_induct_down) + case (1 y) + then + show ?case + using arg_into_eclose by auto +next + case (2 y z) + from \y \ A \ (\B\A. y \ eclose(B))\ + consider (inA) "y \ A" | (exB) "(\B\A. y \ eclose(B))" + by auto + then show ?case + proof (cases) + case inA + then + show ?thesis using 2 arg_into_eclose by auto + next + case exB + then obtain B where "y \ eclose(B)" "B\A" + by auto + then + show ?thesis using 2 ecloseD[of y B z] by auto + qed +qed + +lemma eclose_singE : "x \ eclose({a}) \ x = a \ x \ eclose(a)" + by(blast dest: ecloseE) + +lemma in_eclose_sing : + assumes "x \ eclose({a})" "a \ eclose(z)" + shows "x \ eclose({z})" +proof - + from \x\eclose({a})\ + consider "x=a" | "x\eclose(a)" + using eclose_singE by auto + then + show ?thesis + using eclose_sing mem_eclose_trans assms + by (cases, auto) +qed + +lemma in_dom_in_eclose : + assumes "x \ domain(z)" + shows "x \ eclose(z)" +proof - + from assms + obtain y where "\x,y\ \ z" + unfolding domain_def by auto + then + show ?thesis + unfolding Pair_def + using ecloseD[of "{x,x}"] ecloseD[of "{{x,x},{x,y}}"] arg_into_eclose + by auto +qed + +text\term\ed\ is the well-founded relation on which \<^term>\val\ is defined.\ +definition + ed :: "[i,i] \ o" where + "ed(x,y) \ x \ domain(y)" + +definition + edrel :: "i \ i" where + "edrel(A) \ Rrel(ed,A)" + +lemma edI[intro!]: "t\domain(x) \ ed(t,x)" + unfolding ed_def . + +lemma edD[dest!]: "ed(t,x) \ t\domain(x)" + unfolding ed_def . + +lemma rank_ed: + assumes "ed(y,x)" + shows "succ(rank(y)) \ rank(x)" +proof + from assms + obtain p where "\y,p\\x" by auto + moreover + obtain s where "y\s" "s\\y,p\" unfolding Pair_def by auto + ultimately + have "rank(y) < rank(s)" "rank(s) < rank(\y,p\)" "rank(\y,p\) < rank(x)" + using rank_lt by blast+ + then + show "rank(y) < rank(x)" + using lt_trans by blast +qed + +lemma edrel_dest [dest]: "x \ edrel(A) \ \ a\ A. \ b \ A. x =\a,b\" + by(auto simp add:ed_def edrel_def Rrel_def) + +lemma edrelD : "x \ edrel(A) \ \ a\ A. \ b \ A. x =\a,b\ \ a \ domain(b)" + by(auto simp add:ed_def edrel_def Rrel_def) + +lemma edrelI [intro!]: "x\A \ y\A \ x \ domain(y) \ \x,y\\edrel(A)" + by (simp add:ed_def edrel_def Rrel_def) + +lemma edrel_trans: "Transset(A) \ y\A \ x \ domain(y) \ \x,y\\edrel(A)" + by (rule edrelI, auto simp add:Transset_def domain_def Pair_def) + +lemma domain_trans: "Transset(A) \ y\A \ x \ domain(y) \ x\A" + by (auto simp add: Transset_def domain_def Pair_def) + +lemma relation_edrel : "relation(edrel(A))" + by(auto simp add: relation_def) + +lemma field_edrel : "field(edrel(A))\A" + by blast + +lemma edrel_sub_memrel: "edrel(A) \ trancl(Memrel(eclose(A)))" +proof + let + ?r="trancl(Memrel(eclose(A)))" + fix z + assume "z\edrel(A)" + then + obtain x y where "x\A" "y\A" "z=\x,y\" "x\domain(y)" + using edrelD + by blast + moreover from this + obtain u v where "x\u" "u\v" "v\y" + unfolding domain_def Pair_def by auto + moreover from calculation + have "x\eclose(A)" "y\eclose(A)" "y\eclose(A)" + using arg_into_eclose Transset_eclose[unfolded Transset_def] + by simp_all + moreover from calculation + have "v\eclose(A)" + by auto + moreover from calculation + have "u\eclose(A)" + using Transset_eclose[unfolded Transset_def] + by auto + moreover from calculation + have"\x,u\\?r" "\u,v\\?r" "\v,y\\?r" + by (auto simp add: r_into_trancl) + moreover from this + have "\x,y\\?r" + using trancl_trans[OF _ trancl_trans[of _ v _ y]] + by simp + ultimately + show "z\?r" + by simp +qed + +lemma wf_edrel : "wf(edrel(A))" + using wf_subset [of "trancl(Memrel(eclose(A)))"] edrel_sub_memrel + wf_trancl wf_Memrel + by auto + +lemma ed_induction: + assumes "\x. \\y. ed(y,x) \ Q(y) \ \ Q(x)" + shows "Q(a)" +proof(induct rule: wf_induct2[OF wf_edrel[of "eclose({a})"] ,of a "eclose({a})"]) + case 1 + then show ?case using arg_into_eclose by simp +next + case 2 + then show ?case using field_edrel . +next + case (3 x) + then + show ?case + using assms[of x] edrelI domain_trans[OF Transset_eclose 3(1)] by blast +qed + +lemma dom_under_edrel_eclose: "edrel(eclose({x})) -`` {x} = domain(x)" +proof(intro equalityI) + show "edrel(eclose({x})) -`` {x} \ domain(x)" + unfolding edrel_def Rrel_def ed_def + by auto +next + show "domain(x) \ edrel(eclose({x})) -`` {x}" + unfolding edrel_def Rrel_def + using in_dom_in_eclose eclose_sing arg_into_eclose + by blast +qed + +lemma ed_eclose : "\y,z\ \ edrel(A) \ y \ eclose(z)" + by(drule edrelD,auto simp add:domain_def in_dom_in_eclose) + +lemma tr_edrel_eclose : "\y,z\ \ edrel(eclose({x}))^+ \ y \ eclose(z)" + by(rule trancl_induct,(simp add: ed_eclose mem_eclose_trans)+) + +lemma restrict_edrel_eq : + assumes "z \ domain(x)" + shows "edrel(eclose({x})) \ eclose({z})\eclose({z}) = edrel(eclose({z}))" +proof(intro equalityI subsetI) + let ?ec="\ y . edrel(eclose({y}))" + let ?ez="eclose({z})" + let ?rr="?ec(x) \ ?ez \ ?ez" + fix y + assume "y \ ?rr" + then + obtain a b where "\a,b\ \ ?rr" "a \ ?ez" "b \ ?ez" "\a,b\ \ ?ec(x)" "y=\a,b\" + by blast + moreover from this + have "a \ domain(b)" + using edrelD by blast + ultimately + show "y \ edrel(eclose({z}))" + by blast +next + let ?ec="\ y . edrel(eclose({y}))" + let ?ez="eclose({z})" + let ?rr="?ec(x) \ ?ez \ ?ez" + fix y + assume "y \ edrel(?ez)" + then + obtain a b where "a \ ?ez" "b \ ?ez" "y=\a,b\" "a \ domain(b)" + using edrelD by blast + moreover + from this assms + have "z \ eclose(x)" + using in_dom_in_eclose by simp + moreover + from assms calculation + have "a \ eclose({x})" "b \ eclose({x})" + using in_eclose_sing by simp_all + moreover from calculation + have "\a,b\ \ edrel(eclose({x}))" + by blast + ultimately + show "y \ ?rr" + by simp +qed + +lemma tr_edrel_subset : + assumes "z \ domain(x)" + shows "tr_down(edrel(eclose({x})),z) \ eclose({z})" +proof(intro subsetI) + let ?r="\ x . edrel(eclose({x}))" + fix y + assume "y \ tr_down(?r(x),z)" + then + have "\y,z\ \ ?r(x)^+" + using tr_downD by simp + with assms + show "y \ eclose({z})" + using tr_edrel_eclose eclose_sing by simp +qed + +end \ No newline at end of file diff --git a/thys/Independence_CH/Extensionality_Axiom.thy b/thys/Independence_CH/Extensionality_Axiom.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Extensionality_Axiom.thy @@ -0,0 +1,29 @@ +section\The Axiom of Extensionality in $M[G]$\ + +theory Extensionality_Axiom + imports + Names +begin + +context forcing_data1 +begin + +lemma extensionality_in_MG : "extensionality(##(M[G]))" + unfolding extensionality_def +proof(clarsimp) + fix x y + assume "x\M[G]" "y\M[G]" "(\w\M[G] . w \ x \ w \ y)" + moreover from this + have "z\x \ z\M[G] \ z\y" for z + using transitivity_MG by auto + moreover from calculation + have "z\M[G] \ z\x \ z\y" for z + using transitivity_MG by auto + ultimately + show "x=y" + by auto +qed + +end \ \\<^locale>\forcing_data1\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Fm_Definitions.thy b/thys/Independence_CH/Fm_Definitions.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Fm_Definitions.thy @@ -0,0 +1,1017 @@ +section\Concepts involved in instances of Replacement\ +theory Fm_Definitions + imports + Transitive_Models.Renaming_Auto + Transitive_Models.Aleph_Relative + FrecR_Arities +begin + +txt\In this theory we put every concept that should be synthesized in a formula +to have an instance of replacement. + +The automatic synthesis of a concept /foo/ requires that every concept used to +define /foo/ is already synthesized. We try to use our meta-programs to synthesize +concepts: given the absolute concept /foo/ we relativize in relational form +obtaining /is\_foo/ and the we synthesize the formula /is\_foo\_fm/. +The meta-program that synthesizes formulas also produce satisfactions lemmas. + +Having one file to collect every formula needed for replacements breaks +the reading flow: we need to introduce the concept in this theory in order +to use the meta-programs; moreover there are some concepts for which we prove +here the satisfaction lemmas manually, while for others we prove them +on its theory. +\ + +declare arity_subset_fm [simp del] arity_ordinal_fm[simp del, arity] arity_transset_fm[simp del] + FOL_arities[simp del] + +txt\Formulas for particular replacement instances\ + +text\Now we introduce some definitions used in the definition of check; which +is defined by well-founded recursion using replacement in the recursive call.\ + +\ \The well-founded relation for defining check.\ +definition + rcheck :: "i \ i" where + "rcheck(x) \ Memrel(eclose({x}))^+" + +relativize "rcheck" "is_rcheck" +synthesize "is_rcheck" from_definition +arity_theorem for "is_rcheck_fm" + +\ \The function used for the replacement.\ +definition + PHcheck :: "[i\o,i,i,i,i] \ o" where + "PHcheck(M,o,f,y,p) \ M(p) \ (\fy[M]. fun_apply(M,f,y,fy) \ pair(M,fy,o,p))" + +synthesize "PHcheck" from_definition assuming "nonempty" +arity_theorem for "PHcheck_fm" + +\ \The recursive call for check. We could use the meta-program relationalize for +this; but it makes some proofs more involved.\ +definition + is_Hcheck :: "[i\o,i,i,i,i] \ o" where + "is_Hcheck(M,o,z,f,hc) \ is_Replace(M,z,PHcheck(M,o,f),hc)" + +synthesize "is_Hcheck" from_definition assuming "nonempty" + +lemma arity_is_Hcheck_fm: + assumes "m\nat" "n\nat" "p\nat" "o\nat" + shows "arity(is_Hcheck_fm(m,n,p,o)) = succ(o) \ succ(n) \ succ(p) \ succ(m) " + unfolding is_Hcheck_fm_def + using assms arity_Replace_fm[rule_format,OF PHcheck_fm_type _ _ _ arity_PHcheck_fm] + pred_Un_distrib Un_assoc Un_nat_type + by simp + +\ \The relational version of check is hand-made because our automatic tool +does not handle \<^term>\wfrec\.\ +definition + is_check :: "[i\o,i,i,i] \ o" where + "is_check(M,o,x,z) \ \rch[M]. is_rcheck(M,x,rch) \ + is_wfrec(M,is_Hcheck(M,o),rch,x,z)" + +\ \Finally, we internalize the formula.\ +definition + check_fm :: "[i,i,i] \ i" where + "check_fm(x,o,z) \ Exists(And(is_rcheck_fm(1+\<^sub>\x,0), + is_wfrec_fm(is_Hcheck_fm(6+\<^sub>\o,2,1,0),0,1+\<^sub>\x,1+\<^sub>\z)))" + +lemma check_fm_type[TC]: "x\nat \ o\nat \ z\nat \ check_fm(x,o,z) \ formula" + by (simp add:check_fm_def) + +lemma sats_check_fm : + assumes + "o\nat" "x\nat" "z\nat" "env\list(M)" "0\M" + shows + "(M , env \ check_fm(x,o,z)) \ is_check(##M,nth(o,env),nth(x,env),nth(z,env))" +proof - + have sats_is_Hcheck_fm: + "\a0 a1 a2 a3 a4 a6. \ a0\M; a1\M; a2\M; a3\M; a4\M;a6 \M\ \ + is_Hcheck(##M,a6,a2, a1, a0) \ + (M , [a0,a1,a2,a3,a4,r,a6]@env \ is_Hcheck_fm(6,2,1,0))" if "r\M" for r + using that assms + by simp + then + have "(M , [r]@env \ is_wfrec_fm(is_Hcheck_fm(6+\<^sub>\o,2,1,0),0,1+\<^sub>\x,1+\<^sub>\z)) + \ is_wfrec(##M,is_Hcheck(##M,nth(o,env)),r,nth(x,env),nth(z,env))" + if "r\M" for r + using that assms is_wfrec_iff_sats'[symmetric] + by simp + then + show ?thesis + unfolding is_check_def check_fm_def + using assms is_rcheck_iff_sats[symmetric] + by simp +qed + +lemma iff_sats_check_fm[iff_sats] : + assumes + "nth(o, env) = oa" "nth(x, env) = xa" "nth(z, env) = za" "o \ nat" "x \ nat" "z \ nat" "env \ list(A)" "0 \ A" + shows "is_check(##A, oa,xa, za) \ A, env \ check_fm(x,o, z)" + using assms sats_check_fm[symmetric] + by auto + +lemma arity_check_fm[arity]: + assumes "m\nat" "n\nat" "o\nat" + shows "arity(check_fm(m,n,o)) = succ(o) \ succ(n) \ succ(m) " + unfolding check_fm_def + using assms arity_is_wfrec_fm[rule_format,OF _ _ _ _ _ arity_is_Hcheck_fm] + pred_Un_distrib Un_assoc arity_tran_closure_fm + by (auto simp add:arity) + +notation check_fm (\\_\<^sup>v_ is _\\) + +subsection\Names for forcing the Axiom of Choice.\ +definition + upair_name :: "i \ i \ i \ i" where + "upair_name(\,\,on) \ Upair(\\,on\,\\,on\)" + +relativize "upair_name" "is_upair_name" +synthesize "upair_name" from_definition "is_upair_name" +arity_theorem for "upair_name_fm" + +definition + opair_name :: "i \ i \ i \ i" where + "opair_name(\,\,on) \ upair_name(upair_name(\,\,on),upair_name(\,\,on),on)" + +relativize "opair_name" "is_opair_name" +synthesize "opair_name" from_definition "is_opair_name" +arity_theorem for "opair_name_fm" + +definition + is_opname_check :: "[i\o,i,i,i,i] \ o" where + "is_opname_check(M,on,s,x,y) \ \chx[M]. \sx[M]. is_check(M,on,x,chx) \ + fun_apply(M,s,x,sx) \ is_opair_name(M,chx,sx,on,y)" + +synthesize "is_opname_check" from_definition assuming "nonempty" +arity_theorem for "is_opname_check_fm" + +\ \The pair of elements belongs to some set. The intended set is the preorder.\ +definition + is_leq :: "[i\o,i,i,i] \ o" where + "is_leq(A,l,q,p) \ \qp[A]. (pair(A,q,p,qp) \ qp\l)" + +synthesize "is_leq" from_definition assuming "nonempty" +arity_theorem for "is_leq_fm" + +abbreviation + fm_leq :: "[i,i,i] \ i" (\\_\\<^bsup>_\<^esup>_\\) where + "fm_leq(A,l,B) \ is_leq_fm(l,A,B)" + +subsection\Formulas used to prove some generic instances.\ + +definition \_repl :: "i\i" where + "\_repl(l) \ rsum({\0, 1\, \1, 0\}, id(l), 2, 3, l)" + +lemma f_type : "{\0, 1\, \1, 0\} \ 2 \ 3" + using Pi_iff unfolding function_def by auto + +\ \thm\Internalize.sum_type\ clashes with thm\Renaming.sum_type\.\ +hide_fact Internalize.sum_type + +lemma ren_type : + assumes "l\nat" + shows "\_repl(l) : 2+\<^sub>\l \ 3+\<^sub>\l" + using sum_type[of 2 3 l l "{\0, 1\, \1, 0\}" "id(l)"] f_type assms id_type + unfolding \_repl_def by auto + +definition Lambda_in_M_fm where [simp]:"Lambda_in_M_fm(\,len) \ + \(\\\pair_fm(1, 0, 2) \ + ren(\) ` (2 +\<^sub>\ len) ` (3 +\<^sub>\ len) ` \_repl(len) \\) \ \0 \ len +\<^sub>\ 2\\" + +lemma Lambda_in_M_fm_type[TC]: "\\formula \ len\nat \ Lambda_in_M_fm(\,len) \formula" + using ren_tc[of \ "2+\<^sub>\len" "3+\<^sub>\len" "\_repl(len)"] ren_type + unfolding Lambda_in_M_fm_def + by simp + +definition \_pair_repl :: "i\i" where + "\_pair_repl(l) \ rsum({\0, 0\, \1, 1\, \2, 3\}, id(l), 3, 4, l)" + +definition LambdaPair_in_M_fm where "LambdaPair_in_M_fm(\,len) \ + \(\\\pair_fm(1, 0, 2) \ + ren((\\(\\\\fst(2) is 0\ \ \\snd(2) is 1\ \ ren(\) ` (3 +\<^sub>\ len) ` (4 +\<^sub>\ len) ` \_pair_repl(len) \\\)\)) ` (2 +\<^sub>\ len) ` + (3 +\<^sub>\ len) ` + \_repl(len) \\) \ + \0 \ len +\<^sub>\ 2\\ " + +lemma f_type' : "{\0,0 \, \1, 1\, \2, 3\} \ 3 \ 4" + using Pi_iff unfolding function_def by auto + +lemma ren_type' : + assumes "l\nat" + shows "\_pair_repl(l) : 3+\<^sub>\l \ 4+\<^sub>\l" + using sum_type[of 3 4 l l "{\0, 0\, \1, 1\, \2, 3\}" "id(l)"] f_type' assms id_type + unfolding \_pair_repl_def by auto + +lemma LambdaPair_in_M_fm_type[TC]: "\\formula \ len\nat \ LambdaPair_in_M_fm(\,len) \formula" + using ren_tc[OF _ _ _ ren_type',of \ "len"] Lambda_in_M_fm_type + unfolding LambdaPair_in_M_fm_def + by simp + +subsection\The relation \<^term>\frecrel\\ + +definition + frecrelP :: "[i\o,i] \ o" where + "frecrelP(M,xy) \ (\x[M]. \y[M]. pair(M,x,y,xy) \ is_frecR(M,x,y))" + +synthesize "frecrelP" from_definition +arity_theorem for "frecrelP_fm" + +definition + is_frecrel :: "[i\o,i,i] \ o" where + "is_frecrel(M,A,r) \ \A2[M]. cartprod(M,A,A,A2) \ is_Collect(M,A2, frecrelP(M) ,r)" + +synthesize "frecrel" from_definition "is_frecrel" +arity_theorem for "frecrel_fm" + +definition + names_below :: "i \ i \ i" where + "names_below(P,x) \ 2\ecloseN(x)\ecloseN(x)\P" + +lemma names_belowsD: + assumes "x \ names_below(P,z)" + obtains f n1 n2 p where + "x = \f,n1,n2,p\" "f\2" "n1\ecloseN(z)" "n2\ecloseN(z)" "p\P" + using assms unfolding names_below_def by auto + +synthesize "number2" from_definition + +lemma number2_iff : + "(A)(c) \ number2(A,c) \ (\b[A]. \a[A]. successor(A, b, c) \ successor(A, a, b) \ empty(A, a))" + unfolding number2_def number1_def by auto +arity_theorem for "number2_fm" + +reldb_add "ecloseN" "is_ecloseN" +relativize "names_below" "is_names_below" +synthesize "is_names_below" from_definition +arity_theorem for "is_names_below_fm" + +definition + is_tuple :: "[i\o,i,i,i,i,i] \ o" where + "is_tuple(M,z,t1,t2,p,t) \ \t1t2p[M]. \t2p[M]. pair(M,t2,p,t2p) \ pair(M,t1,t2p,t1t2p) \ + pair(M,z,t1t2p,t)" + +synthesize "is_tuple" from_definition +arity_theorem for "is_tuple_fm" + +subsection\Definition of Forces\ + +subsubsection\Definition of \<^term>\forces\ for equality and membership\ +text\$p\forces \tau = \theta$ if every $q\leqslant p$ both $q\forces \sigma \in \tau$ +and $q\forces \sigma \in \theta$ for every $\sigma \in \dom(\tau)\cup \dom(\theta)$.\ +definition + eq_case :: "[i,i,i,i,i,i] \ o" where + "eq_case(\,\,p,P,leq,f) \ \\. \ \ domain(\) \ domain(\) \ + (\q. q\P \ \q,p\\leq \ (f`\1,\,\,q\=1 \ f`\1,\,\,q\ =1))" + +relativize "eq_case" "is_eq_case" +synthesize "eq_case" from_definition "is_eq_case" + +text\$p\forces \tau \in \theta$ if for every $v\leqslant p$ + there exists $q$, $r$, and $\sigma$ such that + $v\leqslant q$, $q\leqslant r$, $\langle \sigma,r\rangle \in \tau$, and + $q\forces \pi = \sigma$.\ +definition + mem_case :: "[i,i,i,i,i,i] \ o" where + "mem_case(\,\,p,P,leq,f) \ \v\P. \v,p\\leq \ + (\q. \\. \r. r\P \ q\P \ \q,v\\leq \ \\,r\ \ \ \ \q,r\\leq \ f`\0,\,\,q\ = 1)" + +relativize "mem_case" "is_mem_case" +synthesize "mem_case" from_definition "is_mem_case" +arity_theorem intermediate for "eq_case_fm" +lemma arity_eq_case_fm[arity]: + assumes + "n1\nat" "n2\nat" "p\nat" "P\nat" "leq\nat" "f\nat" + shows + "arity(eq_case_fm(n1,n2,p,P,leq,f)) = + succ(n1) \ succ(n2) \ succ(p) \ succ(P) \ succ(leq) \ succ(f)" + using assms arity_eq_case_fm' + by auto + +arity_theorem intermediate for "mem_case_fm" +lemma arity_mem_case_fm[arity] : + assumes + "n1\nat" "n2\nat" "p\nat" "P\nat" "leq\nat" "f\nat" + shows + "arity(mem_case_fm(n1,n2,p,P,leq,f)) = + succ(n1) \ succ(n2) \ succ(p) \ succ(P) \ succ(leq) \ succ(f)" + using assms arity_mem_case_fm' + by auto + +definition + Hfrc :: "[i,i,i,i] \ o" where + "Hfrc(P,leq,fnnc,f) \ \ft. \\. \\. \p. p\P \ fnnc = \ft,\,\,p\ \ + ( ft = 0 \ eq_case(\,\,p,P,leq,f) + \ ft = 1 \ mem_case(\,\,p,P,leq,f))" + +relativize "Hfrc" "is_Hfrc" +synthesize "Hfrc" from_definition "is_Hfrc" + +definition + is_Hfrc_at :: "[i\o,i,i,i,i,i] \ o" where + "is_Hfrc_at(M,P,leq,fnnc,f,b) \ + (empty(M,b) \ \ is_Hfrc(M,P,leq,fnnc,f)) + \ (number1(M,b) \ is_Hfrc(M,P,leq,fnnc,f))" + +synthesize "Hfrc_at" from_definition "is_Hfrc_at" +arity_theorem intermediate for "Hfrc_fm" + +lemma arity_Hfrc_fm[arity] : + assumes + "P\nat" "leq\nat" "fnnc\nat" "f\nat" + shows + "arity(Hfrc_fm(P,leq,fnnc,f)) = succ(P) \ succ(leq) \ succ(fnnc) \ succ(f)" + using assms arity_Hfrc_fm' + by auto + +arity_theorem for "Hfrc_at_fm" + +subsubsection\The well-founded relation \<^term>\forcerel\\ +definition + forcerel :: "i \ i \ i" where + "forcerel(P,x) \ frecrel(names_below(P,x))^+" + +definition + is_forcerel :: "[i\o,i,i,i] \ o" where + "is_forcerel(M,P,x,z) \ \r[M]. \nb[M]. tran_closure(M,r,z) \ + (is_names_below(M,P,x,nb) \ is_frecrel(M,nb,r))" +synthesize "is_forcerel" from_definition +arity_theorem for "is_forcerel_fm" + +subsection\\<^term>\frc_at\, forcing for atomic formulas\ +definition + frc_at :: "[i,i,i] \ i" where + "frc_at(P,leq,fnnc) \ wfrec(frecrel(names_below(P,fnnc)),fnnc, + \x f. bool_of_o(Hfrc(P,leq,x,f)))" + +\ \The relational form is defined manually because it uses \<^term>\wfrec\.\ +definition + is_frc_at :: "[i\o,i,i,i,i] \ o" where + "is_frc_at(M,P,leq,x,z) \ \r[M]. is_forcerel(M,P,x,r) \ + is_wfrec(M,is_Hfrc_at(M,P,leq),r,x,z)" + +definition + frc_at_fm :: "[i,i,i,i] \ i" where + "frc_at_fm(p,l,x,z) \ Exists(And(is_forcerel_fm(succ(p),succ(x),0), + is_wfrec_fm(Hfrc_at_fm(6+\<^sub>\p,6+\<^sub>\l,2,1,0),0,succ(x),succ(z))))" + +lemma frc_at_fm_type [TC] : + "\p\nat;l\nat;x\nat;z\nat\ \ frc_at_fm(p,l,x,z)\formula" + unfolding frc_at_fm_def by simp + +lemma arity_frc_at_fm[arity] : + assumes "p\nat" "l\nat" "x\nat" "z\nat" + shows "arity(frc_at_fm(p,l,x,z)) = succ(p) \ succ(l) \ succ(x) \ succ(z)" +proof - + let ?\ = "Hfrc_at_fm(6 +\<^sub>\ p, 6 +\<^sub>\ l, 2, 1, 0)" + note assms + moreover from this + have "arity(?\) = (7+\<^sub>\p) \ (7+\<^sub>\l)" "?\ \ formula" + using arity_Hfrc_at_fm ord_simp_union + by auto + moreover from calculation + have "arity(is_wfrec_fm(?\, 0, succ(x), succ(z))) = 2+\<^sub>\p \ (2+\<^sub>\l) \ (2+\<^sub>\x) \ (2+\<^sub>\z)" + using arity_is_wfrec_fm[OF \?\\_\ _ _ _ _ \arity(?\) = _\] pred_Un_distrib pred_succ_eq + union_abs1 + by auto + moreover from assms + have "arity(is_forcerel_fm(succ(p),succ(x),0)) = 2+\<^sub>\p \ (2+\<^sub>\x)" + using arity_is_forcerel_fm ord_simp_union + by auto + ultimately + show ?thesis + unfolding frc_at_fm_def + using arity_is_forcerel_fm pred_Un_distrib + by (auto simp:FOL_arities) +qed + +lemma sats_frc_at_fm : + assumes + "p\nat" "l\nat" "i\nat" "j\nat" "env\list(A)" "i < length(env)" "j < length(env)" + shows + "(A , env \ frc_at_fm(p,l,i,j)) \ + is_frc_at(##A,nth(p,env),nth(l,env),nth(i,env),nth(j,env))" +proof - + { + fix r pp ll + assume "r\A" + have "is_Hfrc_at(##A,nth(p,env),nth(l,env),a2, a1, a0) \ + (A, [a0,a1,a2,a3,a4,r]@env \ Hfrc_at_fm(6+\<^sub>\p,6+\<^sub>\l,2,1,0))" + if "a0\A" "a1\A" "a2\A" "a3\A" "a4\A" for a0 a1 a2 a3 a4 + using that assms \r\A\ + Hfrc_at_iff_sats[of "6+\<^sub>\p" "6+\<^sub>\l" 2 1 0 "[a0,a1,a2,a3,a4,r]@env" A] by simp + with \r\A\ + have "(A,[r]@env \ is_wfrec_fm(Hfrc_at_fm(6+\<^sub>\p, 6+\<^sub>\l,2,1,0),0, i+\<^sub>\1, j+\<^sub>\1)) \ + is_wfrec(##A, is_Hfrc_at(##A, nth(p,env), nth(l,env)), r,nth(i, env), nth(j, env))" + using assms sats_is_wfrec_fm + by simp + } + moreover + have "(A, Cons(r, env) \ is_forcerel_fm(succ(p), succ(i), 0)) \ + is_forcerel(##A,nth(p,env),nth(i,env),r)" if "r\A" for r + using assms sats_is_forcerel_fm that + by simp + ultimately + show ?thesis + unfolding is_frc_at_def frc_at_fm_def + using assms + by simp +qed + +lemma frc_at_fm_iff_sats: + assumes "nth(i,env) = w" "nth(j,env) = x" "nth(k,env) = y" "nth(l,env) = z" + "i \ nat" "j \ nat" "k \ nat" "l\nat" "env \ list(A)" "k (A , env \ frc_at_fm(i,j,k,l))" + using assms sats_frc_at_fm + by simp + +declare frc_at_fm_iff_sats [iff_sats] + +definition + forces_eq' :: "[i,i,i,i,i] \ o" where + "forces_eq'(P,l,p,t1,t2) \ frc_at(P,l,\0,t1,t2,p\) = 1" + +definition + forces_mem' :: "[i,i,i,i,i] \ o" where + "forces_mem'(P,l,p,t1,t2) \ frc_at(P,l,\1,t1,t2,p\) = 1" + +definition + forces_neq' :: "[i,i,i,i,i] \ o" where + "forces_neq'(P,l,p,t1,t2) \ \ (\q\P. \q,p\\l \ forces_eq'(P,l,q,t1,t2))" + +definition + forces_nmem' :: "[i,i,i,i,i] \ o" where + "forces_nmem'(P,l,p,t1,t2) \ \ (\q\P. \q,p\\l \ forces_mem'(P,l,q,t1,t2))" + +\ \The following definitions are explicitly defined to avoid the expansion +of concepts.\ +definition + is_forces_eq' :: "[i\o,i,i,i,i,i] \ o" where + "is_forces_eq'(M,P,l,p,t1,t2) \ \o[M]. \z[M]. \t[M]. number1(M,o) \ empty(M,z) \ + is_tuple(M,z,t1,t2,p,t) \ is_frc_at(M,P,l,t,o)" + +definition + is_forces_mem' :: "[i\o,i,i,i,i,i] \ o" where + "is_forces_mem'(M,P,l,p,t1,t2) \ \o[M]. \t[M]. number1(M,o) \ + is_tuple(M,o,t1,t2,p,t) \ is_frc_at(M,P,l,t,o)" + +definition + is_forces_neq' :: "[i\o,i,i,i,i,i] \ o" where + "is_forces_neq'(M,P,l,p,t1,t2) \ + \ (\q[M]. q\P \ (\qp[M]. pair(M,q,p,qp) \ qp\l \ is_forces_eq'(M,P,l,q,t1,t2)))" + +definition + is_forces_nmem' :: "[i\o,i,i,i,i,i] \ o" where + "is_forces_nmem'(M,P,l,p,t1,t2) \ + \ (\q[M]. \qp[M]. q\P \ pair(M,q,p,qp) \ qp\l \ is_forces_mem'(M,P,l,q,t1,t2))" + +synthesize "forces_eq" from_definition "is_forces_eq'" +synthesize "forces_mem" from_definition "is_forces_mem'" +synthesize "forces_neq" from_definition "is_forces_neq'" assuming "nonempty" +synthesize "forces_nmem" from_definition "is_forces_nmem'" assuming "nonempty" + +context + notes Un_assoc[simp] Un_trasposition_aux2[simp] +begin +arity_theorem for "forces_eq_fm" +arity_theorem for "forces_mem_fm" +arity_theorem for "forces_neq_fm" +arity_theorem for "forces_nmem_fm" +end + +subsection\Forcing for general formulas\ + +definition + ren_forces_nand :: "i\i" where + "ren_forces_nand(\) \ Exists(And(Equal(0,1),iterates(\p. incr_bv(p)`1 , 2, \)))" + +lemma ren_forces_nand_type[TC] : + "\\formula \ ren_forces_nand(\) \formula" + unfolding ren_forces_nand_def + by simp + +lemma arity_ren_forces_nand : + assumes "\\formula" + shows "arity(ren_forces_nand(\)) \ succ(arity(\))" +proof - + consider (lt) "1)" | (ge) "\ 1 < arity(\)" + by auto + then + show ?thesis + proof cases + case lt + with \\\_\ + have "2 < succ(arity(\))" "2)+\<^sub>\2" + using succ_ltI by auto + with \\\_\ + have "arity(iterates(\p. incr_bv(p)`1,2,\)) = 2+\<^sub>\arity(\)" + using arity_incr_bv_lemma lt + by auto + with \\\_\ + show ?thesis + unfolding ren_forces_nand_def + using lt pred_Un_distrib union_abs1 Un_assoc[symmetric] Un_le_compat + by (simp add:FOL_arities) + next + case ge + with \\\_\ + have "arity(\) \ 1" "pred(arity(\)) \ 1" + using not_lt_iff_le le_trans[OF le_pred] + by simp_all + with \\\_\ + have "arity(iterates(\p. incr_bv(p)`1,2,\)) = (arity(\))" + using arity_incr_bv_lemma ge + by simp + with \arity(\) \ 1\ \\\_\ \pred(_) \ 1\ + show ?thesis + unfolding ren_forces_nand_def + using pred_Un_distrib union_abs1 Un_assoc[symmetric] union_abs2 + by (simp add:FOL_arities) + qed +qed + +lemma sats_ren_forces_nand: + "[q,P,leq,o,p] @ env \ list(M) \ \\formula \ + (M, [q,p,P,leq,o] @ env \ ren_forces_nand(\)) \ (M, [q,P,leq,o] @ env \ \)" + unfolding ren_forces_nand_def + using sats_incr_bv_iff [of _ _ M _ "[q]"] + by simp + + +definition + ren_forces_forall :: "i\i" where + "ren_forces_forall(\) \ + Exists(Exists(Exists(Exists(Exists( + And(Equal(0,6),And(Equal(1,7),And(Equal(2,8),And(Equal(3,9), + And(Equal(4,5),iterates(\p. incr_bv(p)`5 , 5, \)))))))))))" + +lemma arity_ren_forces_all : + assumes "\\formula" + shows "arity(ren_forces_forall(\)) = 5 \ arity(\)" +proof - + consider (lt) "5)" | (ge) "\ 5 < arity(\)" + by auto + then + show ?thesis + proof cases + case lt + with \\\_\ + have "5 < succ(arity(\))" "5)+\<^sub>\2" "5)+\<^sub>\3" "5)+\<^sub>\4" + using succ_ltI by auto + with \\\_\ + have "arity(iterates(\p. incr_bv(p)`5,5,\)) = 5+\<^sub>\arity(\)" + using arity_incr_bv_lemma lt + by simp + with \\\_\ + show ?thesis + unfolding ren_forces_forall_def + using pred_Un_distrib union_abs1 Un_assoc[symmetric] union_abs2 + by (simp add:FOL_arities) + next + case ge + with \\\_\ + have "arity(\) \ 5" "pred^5(arity(\)) \ 5" + using not_lt_iff_le le_trans[OF le_pred] + by simp_all + with \\\_\ + have "arity(iterates(\p. incr_bv(p)`5,5,\)) = arity(\)" + using arity_incr_bv_lemma ge + by simp + with \arity(\) \ 5\ \\\_\ \pred^5(_) \ 5\ + show ?thesis + unfolding ren_forces_forall_def + using pred_Un_distrib union_abs1 Un_assoc[symmetric] union_abs2 + by (simp add:FOL_arities) + qed +qed + +lemma ren_forces_forall_type[TC] : + "\\formula \ ren_forces_forall(\) \formula" + unfolding ren_forces_forall_def by simp + +lemma sats_ren_forces_forall : + "[x,P,leq,o,p] @ env \ list(M) \ \\formula \ + (M, [x,p,P,leq,o] @ env \ ren_forces_forall(\)) \ (M, [p,P,leq,o,x] @ env \ \)" + unfolding ren_forces_forall_def + using sats_incr_bv_iff [of _ _ M _ "[p,P,leq,o,x]"] + by simp + +subsubsection\The primitive recursion\ + +consts forces' :: "i\i" +primrec + "forces'(Member(x,y)) = forces_mem_fm(1,2,0,x+\<^sub>\4,y+\<^sub>\4)" + "forces'(Equal(x,y)) = forces_eq_fm(1,2,0,x+\<^sub>\4,y+\<^sub>\4)" + "forces'(Nand(p,q)) = + Neg(Exists(And(Member(0,2),And(is_leq_fm(3,0,1),And(ren_forces_nand(forces'(p)), + ren_forces_nand(forces'(q)))))))" + "forces'(Forall(p)) = Forall(ren_forces_forall(forces'(p)))" + + +definition + forces :: "i\i" where + "forces(\) \ And(Member(0,1),forces'(\))" + +lemma forces'_type [TC]: "\\formula \ forces'(\) \ formula" + by (induct \ set:formula; simp) + +lemma forces_type[TC] : "\\formula \ forces(\) \ formula" + unfolding forces_def by simp + +subsection\The arity of \<^term>\forces\\ + +lemma arity_forces_at: + assumes "x \ nat" "y \ nat" + shows "arity(forces(Member(x, y))) = (succ(x) \ succ(y)) +\<^sub>\ 4" + "arity(forces(Equal(x, y))) = (succ(x) \ succ(y)) +\<^sub>\ 4" + unfolding forces_def + using assms arity_forces_mem_fm arity_forces_eq_fm succ_Un_distrib ord_simp_union + by (auto simp:FOL_arities,(rule_tac le_anti_sym,simp_all,(rule_tac not_le_anti_sym,simp_all))+) + +lemma arity_forces': + assumes "\\formula" + shows "arity(forces'(\)) \ arity(\) +\<^sub>\ 4" + using assms +proof (induct set:formula) + case (Member x y) + then + show ?case + using arity_forces_mem_fm succ_Un_distrib ord_simp_union leI not_le_iff_lt + by simp +next + case (Equal x y) + then + show ?case + using arity_forces_eq_fm succ_Un_distrib ord_simp_union leI not_le_iff_lt + by simp +next + case (Nand \ \) + let ?\' = "ren_forces_nand(forces'(\))" + let ?\' = "ren_forces_nand(forces'(\))" + have "arity(is_leq_fm(3, 0, 1)) = 4" + using arity_is_leq_fm succ_Un_distrib ord_simp_union + by simp + have "3 \ (4+\<^sub>\arity(\)) \ (4+\<^sub>\arity(\))" (is "_ \ ?rhs") + using ord_simp_union by simp + from \\\_\ Nand + have "pred(arity(?\')) \ ?rhs" "pred(arity(?\')) \ ?rhs" + proof - + from \\\_\ \\\_\ + have A:"pred(arity(?\')) \ arity(forces'(\))" + "pred(arity(?\')) \ arity(forces'(\))" + using pred_mono[OF _ arity_ren_forces_nand] pred_succ_eq + by simp_all + from Nand + have "3 \ arity(forces'(\)) \ arity(\) +\<^sub>\ 4" + "3 \ arity(forces'(\)) \ arity(\) +\<^sub>\ 4" + using Un_le by simp_all + with Nand + show "pred(arity(?\')) \ ?rhs" + "pred(arity(?\')) \ ?rhs" + using le_trans[OF A(1)] le_trans[OF A(2)] le_Un_iff + by simp_all + qed + with Nand \_=4\ + show ?case + using pred_Un_distrib Un_assoc[symmetric] succ_Un_distrib union_abs1 Un_leI3[OF \3 \ ?rhs\] + by (simp add:FOL_arities) +next + case (Forall \) + let ?\' = "ren_forces_forall(forces'(\))" + show ?case + proof (cases "arity(\) = 0") + case True + with Forall + show ?thesis + proof - + from Forall True + have "arity(forces'(\)) \ 5" + using le_trans[of _ 4 5] by auto + with \\\_\ + have "arity(?\') \ 5" + using arity_ren_forces_all[OF forces'_type[OF \\\_\]] union_abs2 + by auto + with Forall True + show ?thesis + using pred_mono[OF _ \arity(?\') \ 5\] + by simp + qed + next + case False + with Forall + show ?thesis + proof - + from Forall False + have "arity(?\') = 5 \ arity(forces'(\))" + "arity(forces'(\)) \ 5 +\<^sub>\ arity(\)" + "4 \ 3+\<^sub>\arity(\)" + using Ord_0_lt arity_ren_forces_all + le_trans[OF _ add_le_mono[of 4 5, OF _ le_refl]] + by auto + with \\\_\ + have "5 \ arity(forces'(\)) \ 5+\<^sub>\arity(\)" + using ord_simp_union by auto + with \\\_\ \arity(?\') = 5 \ _\ + show ?thesis + using pred_Un_distrib succ_pred_eq[OF _ \arity(\)\0\] + pred_mono[OF _ Forall(2)] Un_le[OF \4\3+\<^sub>\arity(\)\] + by simp + qed + qed +qed + +lemma arity_forces : + assumes "\\formula" + shows "arity(forces(\)) \ 4+\<^sub>\arity(\)" + unfolding forces_def + using assms arity_forces' le_trans ord_simp_union FOL_arities by auto + +lemma arity_forces_le : + assumes "\\formula" "n\nat" "arity(\) \ n" + shows "arity(forces(\)) \ 4+\<^sub>\n" + using assms le_trans[OF _ add_le_mono[OF le_refl[of 5] \arity(\)\_\]] arity_forces + by auto + +definition rename_split_fm where + "rename_split_fm(\) \ (\\(\\(\\(\\(\\(\\\\snd(9) is 0\ \ \\fst(9) is 4\ \ \\1=11\ \ + \\2=12\ \ \\3=13\ \ \\5=7\ \ + (\p. incr_bv(p)`6)^8(forces(\)) \\\\\\\)\)\)\)\)\)" + +lemma rename_split_fm_type[TC]: "\\formula \ rename_split_fm(\)\formula" + unfolding rename_split_fm_def by simp + +schematic_goal arity_rename_split_fm: "\\formula \ arity(rename_split_fm(\)) = ?m" + using arity_forces[of \] forces_type unfolding rename_split_fm_def + by (simp add:arity Un_assoc[symmetric] union_abs1) + +lemma arity_rename_split_fm_le: + assumes "\\formula" + shows "arity(rename_split_fm(\)) \ 8 \ (6 +\<^sub>\ arity(\))" +proof - + from assms + have arity_forces_6: "\ 1 < arity(\) \ 6 \ n \ arity(forces(\)) \ n" for n + using le_trans lt_trans[of _ 5 n] not_lt_iff_le[of 1 "arity(\)"] + by (auto intro!:le_trans[OF arity_forces]) + have pred1_arity_forces: "\ 1 < arity(\) \ pred^n(arity(forces(\))) \ 8" if "n\nat" for n + using that pred_le[of 7] le_succ[THEN [2] le_trans] arity_forces_6 + by (induct rule:nat_induct) auto + have arity_forces_le_succ6: "pred^n(arity(forces(\))) \ succ(succ(succ(succ(succ(succ(arity(\)))))))" + if "n\nat" for n + using that assms arity_forces[of \, THEN le_trans, + OF _ le_succ, THEN le_trans, OF _ _ le_succ] le_trans[OF pred_le[OF _ le_succ]] + by (induct rule:nat_induct) auto + note trivial_arities = arity_forces_6 + arity_forces_le_succ6[of 1, simplified] arity_forces_le_succ6[of 2, simplified] + arity_forces_le_succ6[of 3, simplified] arity_forces_le_succ6[of 4, simplified] + arity_forces_le_succ6[of 5, simplified] arity_forces_le_succ6[of 6, simplified] + pred1_arity_forces[of 1, simplified] pred1_arity_forces[of 2, simplified] + pred1_arity_forces[of 3, simplified] pred1_arity_forces[of 4, simplified] + pred1_arity_forces[of 5, simplified] pred1_arity_forces[of 6, simplified] + show ?thesis + using assms arity_forces[of \] arity_forces[of \, THEN le_trans, OF _ le_succ] + arity_forces[of \, THEN le_trans, OF _ le_succ, THEN le_trans, OF _ _ le_succ] + unfolding rename_split_fm_def + by (simp add:arity Un_assoc[symmetric] union_abs1 arity_forces[of \] forces_type) + ((subst arity_incr_bv_lemma; auto simp: arity ord_simp_union forces_type trivial_arities)+) +qed + +definition body_ground_repl_fm where + "body_ground_repl_fm(\) \ (\\(\\\is_Vset_fm(2, 0) \ \\1 \ 0\ \ rename_split_fm(\) \\\)\)" + +lemma body_ground_repl_fm_type[TC]: "\\formula \ body_ground_repl_fm(\)\formula" + unfolding body_ground_repl_fm_def by simp + +lemma arity_body_ground_repl_fm_le: + notes le_trans[trans] + assumes "\\formula" + shows "arity(body_ground_repl_fm(\)) \ 6 \ (arity(\) +\<^sub>\ 4)" +proof - + from \\\formula\ + have ineq: "n \ pred(pred(arity(rename_split_fm(\)))) + \ m \ pred(pred(8 \ (arity(\) +\<^sub>\6 )))" if "n \ m" "n\nat" "m\nat" for n m + using that arity_rename_split_fm_le[of \, THEN [2] pred_mono, THEN [2] pred_mono, + THEN [2] Un_mono[THEN subset_imp_le, OF _ le_imp_subset]] le_imp_subset + by auto + moreover + have eq1: "pred(pred(pred(4 \ 2 \ pred(pred(pred( + pred(pred(pred(pred(pred(9 \ 1 \ 3 \ 2))))))))))) = 1" + by (auto simp:pred_Un_distrib) + ultimately + have "pred(pred(pred(4 \ 2 \ pred(pred(pred( + pred(pred(pred(pred(pred(9 \ 1 \ 3 \ 2))))))))))) \ + pred(pred(arity(rename_split_fm(\)))) \ + 1 \ pred(pred(8 \ (arity(\) +\<^sub>\6 )))" + by auto + also from \\\formula\ + have "1 \ pred(pred(8 \ (arity(\) +\<^sub>\6 ))) \ 6 \ (4+\<^sub>\arity(\))" + by (auto simp:pred_Un_distrib Un_assoc[symmetric] ord_simp_union) + finally + show ?thesis + using \\\formula\ unfolding body_ground_repl_fm_def + by (simp add:arity pred_Un_distrib, subst arity_transrec_fm[of "is_HVfrom_fm(8,2,1,0)" 3 1]) + (simp add:arity pred_Un_distrib,simp_all, + auto simp add:eq1 arity_is_HVfrom_fm[of 8 2 1 0]) +qed + +definition ground_repl_fm where + "ground_repl_fm(\) \ least_fm(body_ground_repl_fm(\), 1)" + +lemma ground_repl_fm_type[TC]: + "\\formula \ ground_repl_fm(\) \ formula" + unfolding ground_repl_fm_def by simp + +lemma arity_ground_repl_fm: + assumes "\\formula" + shows "arity(ground_repl_fm(\)) \ 5 \ (3 +\<^sub>\ arity(\))" +proof - + from assms + have "pred(arity(body_ground_repl_fm(\))) \ 5 \ (3 +\<^sub>\ arity(\))" + using arity_body_ground_repl_fm_le pred_mono succ_Un_distrib + by (rule_tac pred_le) auto + with assms + have "2 \ pred(arity(body_ground_repl_fm(\))) \ 5 \ (3 +\<^sub>\ arity(\))" + using Un_le le_Un_iff by auto + then + show ?thesis + using assms arity_forces arity_body_ground_repl_fm_le + unfolding least_fm_def ground_repl_fm_def + apply (auto simp add:arity Un_assoc[symmetric]) + apply (simp add: pred_Un Un_assoc, simp add: Un_assoc[symmetric] union_abs1 pred_Un) + by(simp only: Un_commute, subst Un_commute, simp add:ord_simp_union,force) +qed + +simple_rename "ren_F" src "[x_P, x_leq, x_o, x_f, y_c, x_bc, p, x, b]" + tgt "[x_bc, y_c,b,x, x_P, x_leq, x_o, x_f, p]" + +simple_rename "ren_G" src "[x,x_P, x_leq, x_one, x_f,x_p,y,x_B]" + tgt "[x,y,x_P, x_leq, x_one, x_f,x_p,x_B]" + +simple_rename "ren_F_aux" src "[q,x_P, x_leq, x_one, f_dot, x_a, x_bc,x_p,x_b]" + tgt "[x_bc, q, x_b, x_P, x_leq, x_one, f_dot,x_a,x_p]" + +simple_rename "ren_G_aux" src "[ x_b, x_P, x_leq, x_one, f_dot,x_a,x_p,y]" + tgt "[ x_b, y, x_P, x_leq, x_one, f_dot,x_a,x_p]" + +definition ccc_fun_closed_lemma_aux2_fm where [simp]: + "ccc_fun_closed_lemma_aux2_fm \ ren(Collect_fm(1, (\\\\2\<^sup>v5 is 0\ \ ren(\\0\\<^bsup>2\<^esup>7\ + \ forces(\0`1 is 2\ ) \ ) ` 9 ` 9 ` ren_F_aux_fn\\), 7)) ` 8 ` 8 ` ren_G_aux_fn" + +lemma ccc_fun_closed_lemma_aux2_fm_type [TC] : + "ccc_fun_closed_lemma_aux2_fm \ formula" +proof - + let ?\="\\0\\<^bsup>2\<^esup>7\ \ forces(\0`1 is 2\ ) \ " + let ?G="(\\\\2\<^sup>v5 is 0\ \ ren(?\) ` 9 ` 9 ` ren_F_aux_fn\\)" + have "ren(?\)`9`9`ren_F_aux_fn \ formula" + using ren_tc ren_F_aux_thm check_fm_type is_leq_fm_type ren_F_aux_fn_def pred_le + by simp_all + then + show ?thesis + using ren_tc ren_G_aux_thm ren_G_aux_fn_def + by simp +qed + +definition ccc_fun_closed_lemma_fm where [simp]: + "ccc_fun_closed_lemma_fm \ ren(Collect_fm(7, (\\\\2\<^sup>v5 is 0\ \ (\\\\2\<^sup>v6 is 0\ \ + ren((\\\\0 \ 1\ \ \\0\\<^bsup>2\<^esup>7\ \ forces(\0`1 is 2\ ) \\\)) ` 9 ` 9 ` ren_F_fn\\)\\), 6)) + ` 8 ` 8 ` ren_G_fn" + +lemma ccc_fun_closed_lemma_fm_type [TC] : + "ccc_fun_closed_lemma_fm \ formula" +proof - + let ?\="(\\\\0 \ 1\ \ \ \0 \\<^bsup>2\<^esup> 7\ \ forces(\0`1 is 2\ ) \\\)" + let ?G="(\\\\2\<^sup>v5 is 0\ \ (\\\\2\<^sup>v6 is 0\ \ ren(?\) ` 9 ` 9 ` ren_F_fn\\)\\)" + have "ren(?\)`9`9`ren_F_fn \ formula" + using ren_tc ren_F_thm check_fm_type is_leq_fm_type ren_F_fn_def pred_le + by simp_all + then + show ?thesis + using ren_tc ren_G_thm ren_G_fn_def + by simp +qed + +definition is_order_body + where "is_order_body(M,X,x,z) \ \A[M]. cartprod(M,X,X,A) \ subset(M,x,A) \ M(z) \ M(x) \ + is_well_ord(M,X, x) \ is_ordertype(M,X, x,z)" + +synthesize "is_order_body" from_definition assuming "nonempty" + +definition omap_wfrec_body where + "omap_wfrec_body(A,r) \ (\\\image_fm(2, 0, 1) \ pred_set_fm(9+\<^sub>\A, 3,9+\<^sub>\r, 0) \\)" + +lemma type_omap_wfrec_body_fm :"A\nat \ r\nat \ omap_wfrec_body(A,r)\formula" + unfolding omap_wfrec_body_def by simp + +lemma arity_omap_wfrec_aux : "A\nat \ r\nat \ arity(omap_wfrec_body(A,r)) = (9+\<^sub>\A) \ (9+\<^sub>\r)" + unfolding omap_wfrec_body_def + using arity_image_fm arity_pred_set_fm pred_Un_distrib union_abs2[of 3] union_abs1 + by (simp add:FOL_arities, auto simp add:Un_assoc[symmetric] union_abs1) + +lemma arity_omap_wfrec: "A\nat \ r\nat \ + arity(is_wfrec_fm(omap_wfrec_body(A,r),r+\<^sub>\3, 1, 0)) = (4+\<^sub>\A) \ (4+\<^sub>\r)" + using Arities.arity_is_wfrec_fm[OF _ _ _ _ _ arity_omap_wfrec_aux,of A r "3+\<^sub>\r" 1 0] + pred_Un_distrib union_abs1 union_abs2 type_omap_wfrec_body_fm + by auto + +lemma arity_isordermap: "A\nat \ r\nat \d\nat\ + arity(is_ordermap_fm(A,r,d)) = succ(d) \ (succ(A) \ succ(r))" + unfolding is_ordermap_fm_def + using arity_lambda_fm[where i="(4+\<^sub>\A) \ (4+\<^sub>\r)",OF _ _ _ _ arity_omap_wfrec, + unfolded omap_wfrec_body_def] pred_Un_distrib union_abs1 + by auto + + +lemma arity_is_ordertype: "A\nat \ r\nat \d\nat\ + arity(is_ordertype_fm(A,r,d)) = succ(d) \ (succ(A) \ succ(r))" + unfolding is_ordertype_fm_def + using arity_isordermap arity_image_fm pred_Un_distrib FOL_arities + by auto + +arity_theorem for "is_order_body_fm" + +lemma arity_is_order_body: "arity(is_order_body_fm(2,0,1)) = 3" + using arity_is_order_body_fm arity_is_ordertype ord_simp_union + by (simp add:FOL_arities) + +definition H_order_pred where + "H_order_pred(A,r) \ \x f . f `` Order.pred(A, x, r)" + +relationalize "H_order_pred" "is_H_order_pred" + +synthesize "is_H_order_pred" from_definition assuming "nonempty" + +definition order_pred_wfrec_body where + "order_pred_wfrec_body(M,A,r,z,x) \ \y[M]. + pair(M, x, y, z) \ + (\f[M]. + (\z[M]. + z \ f \ + (\xa[M]. + \y[M]. + \xaa[M]. + \sx[M]. + \r_sx[M]. + \f_r_sx[M]. + pair(M, xa, y, z) \ + pair(M, xa, x, xaa) \ + upair(M, xa, xa, sx) \ + pre_image(M, r, sx, r_sx) \ + restriction(M, f, r_sx, f_r_sx) \ + xaa \ r \ (\a[M]. image(M, f_r_sx, a, y) \ pred_set(M, A, xa, r, a)))) \ + (\a[M]. image(M, f, a, y) \ pred_set(M, A, x, r, a)))" + + +synthesize "order_pred_wfrec_body" from_definition +arity_theorem for "order_pred_wfrec_body_fm" + +definition replacement_is_order_body_fm where "replacement_is_order_body_fm \ is_order_body_fm(2,0,1)" +definition wfrec_replacement_order_pred_fm where "wfrec_replacement_order_pred_fm \ order_pred_wfrec_body_fm(3,2,1,0)" +definition replacement_is_jump_cardinal_body_fm where "replacement_is_jump_cardinal_body_fm \ is_jump_cardinal_body'_fm(0,1)" +definition replacement_is_aleph_fm where "replacement_is_aleph_fm \ \\0 is ordinal\ \ \\(0) is 1\\" + +definition + funspace_succ_rep_intf where + "funspace_succ_rep_intf \ \p z n. \f b. p = & z = {cons(, f)}" + +relativize functional "funspace_succ_rep_intf" "funspace_succ_rep_intf_rel" + +\ \The definition obtained next uses \<^term>\is_cons\ instead of \<^term>\upair\ + as in Paulson's \<^file>\~~/src/ZF/Constructible/Relative.thy\.\ +relationalize "funspace_succ_rep_intf_rel" "is_funspace_succ_rep_intf" + +synthesize "is_funspace_succ_rep_intf" from_definition + +arity_theorem for "is_funspace_succ_rep_intf_fm" + +definition + PHrank :: "[i\o,i,i,i] \ o" where + "PHrank(M,f,y,z) \ (\fy[M]. fun_apply(M,f,y,fy) \ successor(M,fy,z))" + +synthesize "PHrank" from_definition assuming "nonempty" + +definition wfrec_Hfrc_at_fm where "wfrec_Hfrc_at_fm \ (\\\pair_fm(1, 0, 2) \ is_wfrec_fm(Hfrc_at_fm(8, 9, 2, 1, 0), 5, 1, 0) \\)" +definition list_repl1_intf_fm where "list_repl1_intf_fm \ (\\\pair_fm(1, 0, 2) \ is_wfrec_fm(iterates_MH_fm(list_functor_fm(13, 1, 0), 10, 2, 1, 0), 3, 1, 0) \\)" +definition list_repl2_intf_fm where "list_repl2_intf_fm \ \\0 \ 4\ \ is_iterates_fm(list_functor_fm(13, 1, 0), 3, 0, 1) \" +definition formula_repl2_intf_fm where "formula_repl2_intf_fm \ \\0 \ 3\ \ is_iterates_fm(formula_functor_fm(1, 0), 2, 0, 1) \" +definition eclose_repl2_intf_fm where "eclose_repl2_intf_fm \ \\0 \ 3\ \ is_iterates_fm(\\1 is 0\, 2, 0, 1) \" +definition powapply_repl_fm where "powapply_repl_fm \ is_Powapply_fm(2,0,1)" +definition phrank_repl_fm where "phrank_repl_fm \ PHrank_fm(2,0,1)" +definition wfrec_rank_fm where "wfrec_rank_fm \ (\\\pair_fm(1, 0, 2) \ is_wfrec_fm(is_Hrank_fm(2, 1, 0), 3, 1, 0) \\)" +definition trans_repl_HVFrom_fm where "trans_repl_HVFrom_fm \ (\\\pair_fm(1, 0, 2) \ is_wfrec_fm(is_HVfrom_fm(8, 2, 1, 0), 4, 1, 0) \\)" +definition wfrec_Hcheck_fm where "wfrec_Hcheck_fm \ (\\\pair_fm(1, 0, 2) \ is_wfrec_fm(is_Hcheck_fm(8, 2, 1, 0), 4, 1, 0) \\) " +definition repl_PHcheck_fm where "repl_PHcheck_fm \ PHcheck_fm(2,3,0,1)" +definition check_replacement_fm where "check_replacement_fm \ \check_fm(0,2,1) \ \0 \ 3\\" +definition G_dot_in_M_fm where "G_dot_in_M_fm \ \(\\\\1\<^sup>v3 is 0\ \ pair_fm(0, 1, 2) \\) \ \0 \ 3\\" +definition repl_opname_check_fm where "repl_opname_check_fm \ \is_opname_check_fm(3,2,0,1) \ \0 \ 4\\" +definition tl_repl_intf_fm where "tl_repl_intf_fm \ (\\\pair_fm(1, 0, 2) \ is_wfrec_fm(iterates_MH_fm(tl_fm(1,0), 9, 2, 1, 0), 3, 1, 0) \\)" +definition formula_repl1_intf_fm where "formula_repl1_intf_fm \ (\\\pair_fm(1, 0, 2) \ is_wfrec_fm(iterates_MH_fm(formula_functor_fm(1,0), 9, 2, 1, 0), 3, 1, 0) \\)" +definition eclose_repl1_intf_fm where "eclose_repl1_intf_fm \ (\\\pair_fm(1, 0, 2) \ is_wfrec_fm(iterates_MH_fm(big_union_fm(1,0), 9, 2, 1, 0), 3, 1, 0) \\)" + +definition replacement_assm where + "replacement_assm(M,env,\) \ \ \ formula \ env \ list(M) \ + arity(\) \ 2 +\<^sub>\ length(env) \ + strong_replacement(##M,\x y. (M , [x,y]@env \ \))" + +definition ground_replacement_assm where + "ground_replacement_assm(M,env,\) \ replacement_assm(M,env,ground_repl_fm(\))" + +end diff --git a/thys/Independence_CH/Forces_Definition.thy b/thys/Independence_CH/Forces_Definition.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Forces_Definition.thy @@ -0,0 +1,858 @@ +section\The definition of \<^term>\forces\\ + +theory Forces_Definition + imports + Forcing_Data +begin + +text\This is the core of our development.\ + +subsection\The relation \<^term>\frecrel\\ + +lemma names_belowsD: + assumes "x \ names_below(P,z)" + obtains f n1 n2 p where + "x = \f,n1,n2,p\" "f\2" "n1\ecloseN(z)" "n2\ecloseN(z)" "p\P" + using assms unfolding names_below_def by auto + +context forcing_data1 +begin + +(* Absoluteness of components *) +lemma ftype_abs: + "\x\M; y\M \ \ is_ftype(##M,x,y) \ y = ftype(x)" + unfolding ftype_def is_ftype_def by (simp add:absolut) + +lemma name1_abs: + "\x\M; y\M \ \ is_name1(##M,x,y) \ y = name1(x)" + unfolding name1_def is_name1_def + by (rule is_hcomp_abs[OF fst_abs],simp_all add: fst_snd_closed[simplified] absolut) + +lemma snd_snd_abs: + "\x\M; y\M \ \ is_snd_snd(##M,x,y) \ y = snd(snd(x))" + unfolding is_snd_snd_def + by (rule is_hcomp_abs[OF snd_abs], + simp_all add: conjunct2[OF fst_snd_closed,simplified] absolut) + +lemma name2_abs: + "\x\M; y\M \ \ is_name2(##M,x,y) \ y = name2(x)" + unfolding name2_def is_name2_def + by (rule is_hcomp_abs[OF fst_abs snd_snd_abs],simp_all add:absolut conjunct2[OF fst_snd_closed,simplified]) + +lemma cond_of_abs: + "\x\M; y\M \ \ is_cond_of(##M,x,y) \ y = cond_of(x)" + unfolding cond_of_def is_cond_of_def + by (rule is_hcomp_abs[OF snd_abs snd_snd_abs];simp_all add:fst_snd_closed[simplified]) + +lemma tuple_abs: + "\z\M;t1\M;t2\M;p\M;t\M\ \ + is_tuple(##M,z,t1,t2,p,t) \ t = \z,t1,t2,p\" + unfolding is_tuple_def using pair_in_M_iff by simp + +lemmas components_abs = ftype_abs name1_abs name2_abs cond_of_abs + tuple_abs + +lemma comp_in_M: + "p \ q \ p\M" + "p \ q \ q\M" + using leq_in_M transitivity[of _ leq] pair_in_M_iff by auto + +(* Absoluteness of Hfrc *) + +lemma eq_case_abs [simp]: + assumes "t1\M" "t2\M" "p\M" "f\M" + shows "is_eq_case(##M,t1,t2,p,P,leq,f) \ eq_case(t1,t2,p,P,leq,f)" +proof - + have "q \ p \ q\M" for q + using comp_in_M by simp + moreover + have "\s,y\\t \ s\domain(t)" if "t\M" for s y t + using that unfolding domain_def by auto + ultimately + have + "(\s\M. s \ domain(t1) \ s \ domain(t2) \ (\q\M. q\P \ q \ p \ + (f ` \1, s, t1, q\ =1 \ f ` \1, s, t2, q\=1))) \ + (\s. s \ domain(t1) \ s \ domain(t2) \ (\q. q\P \ q \ p \ + (f ` \1, s, t1, q\ =1 \ f ` \1, s, t2, q\=1)))" + using assms domain_trans[OF trans_M,of t1] domain_trans[OF trans_M,of t2] + by auto + then + show ?thesis + unfolding eq_case_def is_eq_case_def + using assms pair_in_M_iff nat_into_M domain_closed apply_closed leq_in_M zero_in_M Un_closed + by (simp add:components_abs) +qed + +lemma mem_case_abs [simp]: + assumes "t1\M" "t2\M" "p\M" "f\M" + shows "is_mem_case(##M,t1,t2,p,P,leq,f) \ mem_case(t1,t2,p,P,leq,f)" +proof + { + fix v + assume "v\P" "v \ p" "is_mem_case(##M,t1,t2,p,P,leq,f)" + moreover + from this + have "v\M" "\v,p\ \ M" "(##M)(v)" + using transitivity[OF _ P_in_M,of v] transitivity[OF _ leq_in_M] + by simp_all + moreover + from calculation assms + obtain q r s where + "r \ P \ q \ P \ \q, v\ \ M \ \s, r\ \ M \ \q, r\ \ M \ 0 \ M \ + \0, t1, s, q\ \ M \ q \ v \ \s, r\ \ t2 \ q \ r \ f ` \0, t1, s, q\ = 1" + unfolding is_mem_case_def by (auto simp add:components_abs) + then + have "\q s r. r \ P \ q \ P \ q \ v \ \s, r\ \ t2 \ q \ r \ f ` \0, t1, s, q\ = 1" + by auto + } + then + show "mem_case(t1, t2, p, P, leq, f)" if "is_mem_case(##M, t1, t2, p, P, leq, f)" + unfolding mem_case_def using that assms by auto +next + { fix v + assume "v \ M" "v \ P" "\v, p\ \ M" "v \ p" "mem_case(t1, t2, p, P, leq, f)" + moreover + from this + obtain q s r where "r \ P \ q \ P \ q \ v \ \s, r\ \ t2 \ q \ r \ f ` \0, t1, s, q\ = 1" + unfolding mem_case_def by auto + moreover + from this \t2\M\ + have "r\M" "q\M" "s\M" "r \ P \ q \ P \ q \ v \ \s, r\ \ t2 \ q \ r \ f ` \0, t1, s, q\ = 1" + using transitivity domainI[of s r] P_in_M domain_closed + by auto + moreover + note \t1\M\ + ultimately + have "\q\M . \s\M. \r\M. + r \ P \ q \ P \ \q, v\ \ M \ \s, r\ \ M \ \q, r\ \ M \ 0 \ M \ + \0, t1, s, q\ \ M \ q \ v \ \s, r\ \ t2 \ q \ r \ f ` \0, t1, s, q\ = 1" + using pair_in_M_iff zero_in_M by auto + } + then + show "is_mem_case(##M, t1, t2, p, P, leq, f)" if "mem_case(t1, t2, p, P, leq, f)" + unfolding is_mem_case_def + using assms that zero_in_M pair_in_M_iff apply_closed nat_into_M + by (auto simp add:components_abs) +qed + +lemma Hfrc_abs: + "\fnnc\M; f\M\ \ + is_Hfrc(##M,P,leq,fnnc,f) \ Hfrc(P,leq,fnnc,f)" + unfolding is_Hfrc_def Hfrc_def using pair_in_M_iff zero_in_M + by (auto simp add:components_abs) + +lemma Hfrc_at_abs: + "\fnnc\M; f\M ; z\M\ \ + is_Hfrc_at(##M,P,leq,fnnc,f,z) \ z = bool_of_o(Hfrc(P,leq,fnnc,f)) " + unfolding is_Hfrc_at_def using Hfrc_abs + by auto + +lemma components_closed : + "x\M \ (##M)(ftype(x))" + "x\M \ (##M)(name1(x))" + "x\M \ (##M)(name2(x))" + "x\M \ (##M)(cond_of(x))" + unfolding ftype_def name1_def name2_def cond_of_def using fst_snd_closed by simp_all + +lemma ecloseN_closed: + "(##M)(A) \ (##M)(ecloseN(A))" + "(##M)(A) \ (##M)(eclose_n(name1,A))" + "(##M)(A) \ (##M)(eclose_n(name2,A))" + unfolding ecloseN_def eclose_n_def + using components_closed eclose_closed singleton_closed Un_closed by auto + +lemma eclose_n_abs : + assumes "x\M" "ec\M" + shows "is_eclose_n(##M,is_name1,ec,x) \ ec = eclose_n(name1,x)" + "is_eclose_n(##M,is_name2,ec,x) \ ec = eclose_n(name2,x)" + unfolding is_eclose_n_def eclose_n_def + using assms name1_abs name2_abs eclose_abs singleton_closed components_closed + by auto + + +lemma ecloseN_abs : + "\x\M;ec\M\ \ is_ecloseN(##M,x,ec) \ ec = ecloseN(x)" + unfolding is_ecloseN_def ecloseN_def + using eclose_n_abs Un_closed union_abs ecloseN_closed + by auto + +lemma frecR_abs : + "x\M \ y\M \ frecR(x,y) \ is_frecR(##M,x,y)" + unfolding frecR_def is_frecR_def + using zero_in_M domain_closed Un_closed components_closed nat_into_M + by (auto simp add: components_abs) + +lemma frecrelP_abs : + "z\M \ frecrelP(##M,z) \ (\x y. z = \x,y\ \ frecR(x,y))" + using pair_in_M_iff frecR_abs unfolding frecrelP_def by auto + +lemma frecrel_abs: + assumes "A\M" "r\M" + shows "is_frecrel(##M,A,r) \ r = frecrel(A)" +proof - + from \A\M\ + have "z\M" if "z\A\A" for z + using cartprod_closed transitivity that by simp + then + have "Collect(A\A,frecrelP(##M)) = Collect(A\A,\z. (\x y. z = \x,y\ \ frecR(x,y)))" + using Collect_cong[of "A\A" "A\A" "frecrelP(##M)"] assms frecrelP_abs by simp + with assms + show ?thesis + unfolding is_frecrel_def def_frecrel using cartprod_closed + by simp +qed + +lemma frecrel_closed: + assumes "x\M" + shows "frecrel(x)\M" +proof - + have "Collect(x\x,\z. (\x y. z = \x,y\ \ frecR(x,y)))\M" + using Collect_in_M[of "frecrelP_fm(0)" "[]"] arity_frecrelP_fm sats_frecrelP_fm + frecrelP_abs \x\M\ cartprod_closed + by simp + then + show ?thesis + unfolding frecrel_def Rrel_def frecrelP_def by simp +qed + +lemma field_frecrel : "field(frecrel(names_below(P,x))) \ names_below(P,x)" + unfolding frecrel_def + using field_Rrel by simp + +lemma forcerelD : "uv \ forcerel(P,x) \ uv\ names_below(P,x) \ names_below(P,x)" + unfolding forcerel_def + using trancl_type field_frecrel by blast + +lemma wf_forcerel : + "wf(forcerel(P,x))" + unfolding forcerel_def using wf_trancl wf_frecrel . + +lemma restrict_trancl_forcerel: + assumes "frecR(w,y)" + shows "restrict(f,frecrel(names_below(P,x))-``{y})`w + = restrict(f,forcerel(P,x)-``{y})`w" + unfolding forcerel_def frecrel_def using assms restrict_trancl_Rrel[of frecR] + by simp + +lemma names_belowI : + assumes "frecR(\ft,n1,n2,p\,\a,b,c,d\)" "p\P" + shows "\ft,n1,n2,p\ \ names_below(P,\a,b,c,d\)" (is "?x \ names_below(_,?y)") +proof - + from assms + have "ft \ 2" "a \ 2" + unfolding frecR_def by (auto simp add:components_simp) + from assms + consider (e) "n1 \ domain(b) \ domain(c) \ (n2 = b \ n2 =c)" + | (m) "n1 = b \ n2 \ domain(c)" + unfolding frecR_def by (auto simp add:components_simp) + then show ?thesis + proof cases + case e + then + have "n1 \ eclose(b) \ n1 \ eclose(c)" + using Un_iff in_dom_in_eclose by auto + with e + have "n1 \ ecloseN(?y)" "n2 \ ecloseN(?y)" + using ecloseNI components_in_eclose by auto + with \ft\2\ \p\P\ + show ?thesis + unfolding names_below_def by auto + next + case m + then + have "n1 \ ecloseN(?y)" "n2 \ ecloseN(?y)" + using mem_eclose_trans ecloseNI + in_dom_in_eclose components_in_eclose by auto + with \ft\2\ \p\P\ + show ?thesis + unfolding names_below_def + by auto + qed +qed + +lemma names_below_tr : + assumes "x\ names_below(P,y)" "y\ names_below(P,z)" + shows "x\ names_below(P,z)" +proof - + let ?A="\y . names_below(P,y)" + note assms + moreover from this + obtain fx x1 x2 px where "x = \fx,x1,x2,px\" "fx\2" "x1\ecloseN(y)" "x2\ecloseN(y)" "px\P" + unfolding names_below_def by auto + moreover from calculation + obtain fy y1 y2 py where "y = \fy,y1,y2,py\" "fy\2" "y1\ecloseN(z)" "y2\ecloseN(z)" "py\P" + unfolding names_below_def by auto + moreover from calculation + have "x1\ecloseN(z)" "x2\ecloseN(z)" + using ecloseN_mono names_simp by auto + ultimately + have "x\?A(z)" + unfolding names_below_def by simp + then + show ?thesis using subsetI by simp +qed + +lemma arg_into_names_below2 : + assumes "\x,y\ \ frecrel(names_below(P,z))" + shows "x \ names_below(P,y)" +proof - + from assms + have "x\names_below(P,z)" "y\names_below(P,z)" "frecR(x,y)" + unfolding frecrel_def Rrel_def + by auto + obtain f n1 n2 p where "x = \f,n1,n2,p\" "f\2" "n1\ecloseN(z)" "n2\ecloseN(z)" "p\P" + using \x\names_below(P,z)\ + unfolding names_below_def by auto + moreover + obtain fy m1 m2 q where "q\P" "y = \fy,m1,m2,q\" + using \y\names_below(P,z)\ + unfolding names_below_def by auto + moreover + note \frecR(x,y)\ + ultimately + show ?thesis + using names_belowI by simp +qed + +lemma arg_into_names_below : + assumes "\x,y\ \ frecrel(names_below(P,z))" + shows "x \ names_below(P,x)" +proof - + from assms + have "x\names_below(P,z)" + unfolding frecrel_def Rrel_def + by auto + from \x\names_below(P,z)\ + obtain f n1 n2 p where + "x = \f,n1,n2,p\" "f\2" "n1\ecloseN(z)" "n2\ecloseN(z)" "p\P" + unfolding names_below_def by auto + then + have "n1\ecloseN(x)" "n2\ecloseN(x)" + using components_in_eclose by simp_all + with \f\2\ \p\P\ \x = \f,n1,n2,p\\ + show ?thesis + unfolding names_below_def by simp +qed + +lemma forcerel_arg_into_names_below : + assumes "\x,y\ \ forcerel(P,z)" + shows "x \ names_below(P,x)" + using assms + unfolding forcerel_def + by(rule trancl_induct;auto simp add: arg_into_names_below) + +lemma names_below_mono : + assumes "\x,y\ \ frecrel(names_below(P,z))" + shows "names_below(P,x) \ names_below(P,y)" +proof - + from assms + have "x\names_below(P,y)" + using arg_into_names_below2 by simp + then + show ?thesis + using names_below_tr subsetI by simp +qed + +lemma frecrel_mono : + assumes "\x,y\ \ frecrel(names_below(P,z))" + shows "frecrel(names_below(P,x)) \ frecrel(names_below(P,y))" + unfolding frecrel_def + using Rrel_mono names_below_mono assms by simp + +lemma forcerel_mono2 : + assumes "\x,y\ \ frecrel(names_below(P,z))" + shows "forcerel(P,x) \ forcerel(P,y)" + unfolding forcerel_def + using trancl_mono frecrel_mono assms by simp + +lemma forcerel_mono_aux : + assumes "\x,y\ \ frecrel(names_below(P, w))^+" + shows "forcerel(P,x) \ forcerel(P,y)" + using assms + by (rule trancl_induct,simp_all add: subset_trans forcerel_mono2) + +lemma forcerel_mono : + assumes "\x,y\ \ forcerel(P,z)" + shows "forcerel(P,x) \ forcerel(P,y)" + using forcerel_mono_aux assms unfolding forcerel_def by simp + +lemma forcerel_eq_aux: "x \ names_below(P, w) \ \x,y\ \ forcerel(P,z) \ + (y \ names_below(P, w) \ \x,y\ \ forcerel(P,w))" + unfolding forcerel_def +proof(rule_tac a=x and b=y and P="\ y . y \ names_below(P, w) \ \x,y\ \ frecrel(names_below(P,w))^+" in trancl_induct,simp) + let ?A="\ a . names_below(P, a)" + let ?R="\ a . frecrel(?A(a))" + let ?fR="\ a .forcerel(a)" + show "u\?A(w) \ \x,u\\?R(w)^+" if "x\?A(w)" "\x,y\\?R(z)^+" "\x,u\\?R(z)" for u + using that frecrelD frecrelI r_into_trancl + unfolding names_below_def by simp + { + fix u v + assume "x \ ?A(w)" + "\x, y\ \ ?R(z)^+" + "\x, u\ \ ?R(z)^+" + "\u, v\ \ ?R(z)" + "u \ ?A(w) \ \x, u\ \ ?R(w)^+" + then + have "v \ ?A(w) \ \x, v\ \ ?R(w)^+" + proof - + assume "v \?A(w)" + from \\u,v\\_\ + have "u\?A(v)" + using arg_into_names_below2 by simp + with \v \?A(w)\ + have "u\?A(w)" + using names_below_tr by simp + with \v\_\ \\u,v\\_\ + have "\u,v\\ ?R(w)" + using frecrelD frecrelI r_into_trancl unfolding names_below_def by simp + with \u \ ?A(w) \ \x, u\ \ ?R(w)^+\ \u\?A(w)\ + have "\x, u\ \ ?R(w)^+" + by simp + with \\u,v\\ ?R(w)\ + show "\x,v\\ ?R(w)^+" using trancl_trans r_into_trancl + by simp + qed + } + then + show "v \ ?A(w) \ \x, v\ \ ?R(w)^+" + if "x \ ?A(w)" + "\x, y\ \ ?R(z)^+" + "\x, u\ \ ?R(z)^+" + "\u, v\ \ ?R(z)" + "u \ ?A(w) \ \x, u\ \ ?R(w)^+" for u v + using that + by simp +qed + +lemma forcerel_eq : + assumes "\z,x\ \ forcerel(P,x)" + shows "forcerel(P,z) = forcerel(P,x) \ names_below(P,z)\names_below(P,z)" + using assms forcerel_eq_aux forcerelD forcerel_mono[of z x x] subsetI + by auto + +lemma forcerel_below_aux : + assumes "\z,x\ \ forcerel(P,x)" "\u,z\ \ forcerel(P,x)" + shows "u \ names_below(P,z)" + using assms(2) + unfolding forcerel_def +proof(rule trancl_induct) + show "u \ names_below(P,y)" if " \u, y\ \ frecrel(names_below(P, x))" for y + using that vimage_singleton_iff arg_into_names_below2 by simp +next + show "u \ names_below(P,z)" + if "\u, y\ \ frecrel(names_below(P, x))^+" + "\y, z\ \ frecrel(names_below(P, x))" + "u \ names_below(P, y)" + for y z + using that arg_into_names_below2[of y z x] names_below_tr by simp +qed + +lemma forcerel_below : + assumes "\z,x\ \ forcerel(P,x)" + shows "forcerel(P,x) -`` {z} \ names_below(P,z)" + using vimage_singleton_iff assms forcerel_below_aux by auto + +lemma relation_forcerel : + shows "relation(forcerel(P,z))" "trans(forcerel(P,z))" + unfolding forcerel_def using relation_trancl trans_trancl by simp_all + +lemma Hfrc_restrict_trancl: "bool_of_o(Hfrc(P, leq, y, restrict(f,frecrel(names_below(P,x))-``{y}))) + = bool_of_o(Hfrc(P, leq, y, restrict(f,(frecrel(names_below(P,x))^+)-``{y})))" + unfolding Hfrc_def bool_of_o_def eq_case_def mem_case_def + using restrict_trancl_forcerel frecRI1 frecRI2 frecRI3 + unfolding forcerel_def + by simp + +(* Recursive definition of forces for atomic formulas using a transitive relation *) +lemma frc_at_trancl: "frc_at(P,leq,z) = wfrec(forcerel(P,z),z,\x f. bool_of_o(Hfrc(P,leq,x,f)))" + unfolding frc_at_def forcerel_def using wf_eq_trancl Hfrc_restrict_trancl by simp + +lemma forcerelI1 : + assumes "n1 \ domain(b) \ n1 \ domain(c)" "p\P" "d\P" + shows "\\1, n1, b, p\, \0,b,c,d\\\ forcerel(P,\0,b,c,d\)" +proof - + let ?x="\1, n1, b, p\" + let ?y="\0,b,c,d\" + from assms + have "frecR(?x,?y)" + using frecRI1 by simp + then + have "?x\names_below(P,?y)" "?y \ names_below(P,?y)" + using names_belowI assms components_in_eclose + unfolding names_below_def by auto + with \frecR(?x,?y)\ + show ?thesis + unfolding forcerel_def frecrel_def + using subsetD[OF r_subset_trancl[OF relation_Rrel]] RrelI + by auto +qed + +lemma forcerelI2 : + assumes "n1 \ domain(b) \ n1 \ domain(c)" "p\P" "d\P" + shows "\\1, n1, c, p\, \0,b,c,d\\\ forcerel(P,\0,b,c,d\)" +proof - + let ?x="\1, n1, c, p\" + let ?y="\0,b,c,d\" + note assms + moreover from this + have "frecR(?x,?y)" + using frecRI2 by simp + moreover from calculation + have "?x\names_below(P,?y)" "?y \ names_below(P,?y)" + using names_belowI components_in_eclose + unfolding names_below_def by auto + ultimately + show ?thesis + unfolding forcerel_def frecrel_def + using subsetD[OF r_subset_trancl[OF relation_Rrel]] RrelI + by auto +qed + +lemma forcerelI3 : + assumes "\n2, r\ \ c" "p\P" "d\P" "r \ P" + shows "\\0, b, n2, p\,\1, b, c, d\\ \ forcerel(P,\1,b,c,d\)" +proof - + let ?x="\0, b, n2, p\" + let ?y="\1, b, c, d\" + note assms + moreover from this + have "frecR(?x,?y)" + using frecRI3 by simp + moreover from calculation + have "?x\names_below(P,?y)" "?y \ names_below(P,?y)" + using names_belowI components_in_eclose + unfolding names_below_def by auto + ultimately + show ?thesis + unfolding forcerel_def frecrel_def + using subsetD[OF r_subset_trancl[OF relation_Rrel]] RrelI + by auto +qed + +lemmas forcerelI = forcerelI1[THEN vimage_singleton_iff[THEN iffD2]] + forcerelI2[THEN vimage_singleton_iff[THEN iffD2]] + forcerelI3[THEN vimage_singleton_iff[THEN iffD2]] + +lemma aux_def_frc_at: + assumes "z \ forcerel(P,x) -`` {x}" + shows "wfrec(forcerel(P,x), z, H) = wfrec(forcerel(P,z), z, H)" +proof - + let ?A="names_below(P,z)" + from assms + have "\z,x\ \ forcerel(P,x)" + using vimage_singleton_iff by simp + moreover from this + have "z \ ?A" + using forcerel_arg_into_names_below by simp + moreover from calculation + have "forcerel(P,z) = forcerel(P,x) \ (?A\?A)" + "forcerel(P,x) -`` {z} \ ?A" + using forcerel_eq forcerel_below + by auto + moreover from calculation + have "wfrec(forcerel(P,x), z, H) = wfrec[?A](forcerel(P,x), z, H)" + using wfrec_trans_restr[OF relation_forcerel(1) wf_forcerel relation_forcerel(2), of x z ?A] + by simp + ultimately + show ?thesis + using wfrec_restr_eq by simp +qed + +subsection\Recursive expression of \<^term>\frc_at\\ + +lemma def_frc_at : + assumes "p\P" + shows + "frc_at(P,leq,\ft,n1,n2,p\) = + bool_of_o( p \P \ + ( ft = 0 \ (\s. s\domain(n1) \ domain(n2) \ + (\q. q\P \ q \ p \ (frc_at(P,leq,\1,s,n1,q\) =1 \ frc_at(P,leq,\1,s,n2,q\) =1))) + \ ft = 1 \ ( \v\P. v \ p \ + (\q. \s. \r. r\P \ q\P \ q \ v \ \s,r\ \ n2 \ q \ r \ frc_at(P,leq,\0,n1,s,q\) = 1))))" +proof - + let ?r="\y. forcerel(P,y)" and ?Hf="\x f. bool_of_o(Hfrc(P,leq,x,f))" + let ?t="\y. ?r(y) -`` {y}" + let ?arg="\ft,n1,n2,p\" + from wf_forcerel + have wfr: "\w . wf(?r(w))" .. + with wfrec [of "?r(?arg)" ?arg ?Hf] + have "frc_at(P,leq,?arg) = ?Hf( ?arg, \x\?r(?arg) -`` {?arg}. wfrec(?r(?arg), x, ?Hf))" + using frc_at_trancl by simp + also + have " ... = ?Hf( ?arg, \x\?r(?arg) -`` {?arg}. frc_at(P,leq,x))" + using aux_def_frc_at frc_at_trancl by simp + finally + show ?thesis + unfolding Hfrc_def mem_case_def eq_case_def + using forcerelI assms + by auto +qed + + +subsection\Absoluteness of \<^term>\frc_at\\ + +lemma forcerel_in_M : + assumes "x\M" + shows "forcerel(P,x)\M" + unfolding forcerel_def def_frecrel names_below_def +proof - + let ?Q = "2 \ ecloseN(x) \ ecloseN(x) \ P" + have "?Q \ ?Q \ M" + using \x\M\ P_in_M nat_into_M ecloseN_closed cartprod_closed by simp + moreover + have "separation(##M,\z. frecrelP(##M,z))" + using separation_in_ctm[of "frecrelP_fm(0)",OF _ _ _ sats_frecrelP_fm] + arity_frecrelP_fm frecrelP_fm_type + by auto + moreover from this + have "separation(##M,\z. \x y. z = \x, y\ \ frecR(x, y))" + using separation_cong[OF frecrelP_abs] + by force + ultimately + show "{z \ ?Q \ ?Q . \x y. z = \x, y\ \ frecR(x, y)}^+ \ M" + using separation_closed frecrelP_abs trancl_closed + by simp +qed + +lemma relation2_Hfrc_at_abs: + "relation2(##M,is_Hfrc_at(##M,P,leq),\x f. bool_of_o(Hfrc(P,leq,x,f)))" + unfolding relation2_def using Hfrc_at_abs + by simp + +lemma Hfrc_at_closed : + "\x\M. \g\M. function(g) \ bool_of_o(Hfrc(P,leq,x,g))\M" + unfolding bool_of_o_def using zero_in_M nat_into_M[of 1] by simp + +lemma wfrec_Hfrc_at : + assumes "X\M" + shows "wfrec_replacement(##M,is_Hfrc_at(##M,P,leq),forcerel(P,X))" +proof - + have 0:"is_Hfrc_at(##M,P,leq,a,b,c) \ + sats(M,Hfrc_at_fm(8,9,2,1,0),[c,b,a,d,e,y,x,z,P,leq,forcerel(P,X)])" + if "a\M" "b\M" "c\M" "d\M" "e\M" "y\M" "x\M" "z\M" + for a b c d e y x z + using that P_in_M leq_in_M \X\M\ forcerel_in_M + Hfrc_at_iff_sats[of concl:M P leq a b c 8 9 2 1 0 + "[c,b,a,d,e,y,x,z,P,leq,forcerel(P,X)]"] by simp + have 1:"sats(M,is_wfrec_fm(Hfrc_at_fm(8,9,2,1,0),5,1,0),[y,x,z,P,leq,forcerel(P,X)]) \ + is_wfrec(##M, is_Hfrc_at(##M,P,leq),forcerel(P,X), x, y)" + if "x\M" "y\M" "z\M" for x y z + using that \X\M\ forcerel_in_M P_in_M leq_in_M sats_is_wfrec_fm[OF 0] + by simp + let + ?f="Exists(And(pair_fm(1,0,2),is_wfrec_fm(Hfrc_at_fm(8,9,2,1,0),5,1,0)))" + have satsf:"sats(M, ?f, [x,z,P,leq,forcerel(P,X)]) \ + (\y\M. pair(##M,x,y,z) & is_wfrec(##M, is_Hfrc_at(##M,P,leq),forcerel(P,X), x, y))" + if "x\M" "z\M" for x z + using that 1 \X\M\ forcerel_in_M P_in_M leq_in_M by (simp del:pair_abs) + have artyf:"arity(?f) = 5" + using arity_wfrec_replacement_fm[where p="Hfrc_at_fm(8,9,2,1,0)" and i=10] + arity_Hfrc_at_fm ord_simp_union + by simp + moreover + have "?f\formula" by simp + ultimately + have "strong_replacement(##M,\x z. sats(M,?f,[x,z,P,leq,forcerel(P,X)]))" + using replacement_ax1(1) 1 artyf \X\M\ forcerel_in_M P_in_M leq_in_M + unfolding replacement_assm_def by simp + then + have "strong_replacement(##M,\x z. + \y\M. pair(##M,x,y,z) & is_wfrec(##M, is_Hfrc_at(##M,P,leq),forcerel(P,X), x, y))" + using repl_sats[of M ?f "[P,leq,forcerel(P,X)]"] satsf by (simp del:pair_abs) + then + show ?thesis unfolding wfrec_replacement_def by simp +qed + +lemma names_below_abs : + "\Q\M;x\M;nb\M\ \ is_names_below(##M,Q,x,nb) \ nb = names_below(Q,x)" + unfolding is_names_below_def names_below_def + using succ_in_M_iff zero_in_M cartprod_closed ecloseN_abs ecloseN_closed + by auto + +lemma names_below_closed: + "\Q\M;x\M\ \ names_below(Q,x) \ M" + unfolding names_below_def + using zero_in_M cartprod_closed ecloseN_closed succ_in_M_iff + by simp + +lemma "names_below_productE" : + assumes "Q \ M" "x \ M" + "\A1 A2 A3 A4. A1 \ M \ A2 \ M \ A3 \ M \ A4 \ M \ R(A1 \ A2 \ A3 \ A4)" + shows "R(names_below(Q,x))" + unfolding names_below_def using assms nat_into_M ecloseN_closed[of x] by auto + +lemma forcerel_abs : + "\x\M;z\M\ \ is_forcerel(##M,P,x,z) \ z = forcerel(P,x)" + unfolding is_forcerel_def forcerel_def + using frecrel_abs names_below_abs trancl_abs P_in_M ecloseN_closed names_below_closed + names_below_productE[of concl:"\p. is_frecrel(##M,p,_) \ _ = frecrel(p)"] frecrel_closed + by simp + +lemma frc_at_abs: + assumes "fnnc\M" "z\M" + shows "is_frc_at(##M,P,leq,fnnc,z) \ z = frc_at(P,leq,fnnc)" +proof - + from assms + have "(\r\M. is_forcerel(##M,P,fnnc, r) \ is_wfrec(##M, is_Hfrc_at(##M, P, leq), r, fnnc, z)) + \ is_wfrec(##M, is_Hfrc_at(##M, P, leq), forcerel(P,fnnc), fnnc, z)" + using forcerel_abs forcerel_in_M by simp + then + show ?thesis + unfolding frc_at_trancl is_frc_at_def + using assms wfrec_Hfrc_at[of fnnc] wf_forcerel relation_forcerel forcerel_in_M + Hfrc_at_closed relation2_Hfrc_at_abs + trans_wfrec_abs[of "forcerel(P,fnnc)" fnnc z "is_Hfrc_at(##M,P,leq)" "\x f. bool_of_o(Hfrc(P,leq,x,f))"] + by (simp flip:setclass_iff) +qed + +lemma forces_eq'_abs : + "\p\M ; t1\M ; t2\M\ \ is_forces_eq'(##M,P,leq,p,t1,t2) \ forces_eq'(P,leq,p,t1,t2)" + unfolding is_forces_eq'_def forces_eq'_def + using frc_at_abs nat_into_M pair_in_M_iff by (auto simp add:components_abs) + +lemma forces_mem'_abs : + "\p\M ; t1\M ; t2\M\ \ is_forces_mem'(##M,P,leq,p,t1,t2) \ forces_mem'(P,leq,p,t1,t2)" + unfolding is_forces_mem'_def forces_mem'_def + using frc_at_abs nat_into_M pair_in_M_iff by (auto simp add:components_abs) + +lemma forces_neq'_abs : + assumes "p\M" "t1\M" "t2\M" + shows "is_forces_neq'(##M,P,leq,p,t1,t2) \ forces_neq'(P,leq,p,t1,t2)" +proof - + have "q\M" if "q\P" for q + using that transitivity P_in_M by simp + with assms + show ?thesis + unfolding is_forces_neq'_def forces_neq'_def + using forces_eq'_abs pair_in_M_iff + by (auto simp add:components_abs,blast) +qed + + +lemma forces_nmem'_abs : + assumes "p\M" "t1\M" "t2\M" + shows "is_forces_nmem'(##M,P,leq,p,t1,t2) \ forces_nmem'(P,leq,p,t1,t2)" +proof - + have "q\M" if "q\P" for q + using that transitivity P_in_M by simp + with assms + show ?thesis + unfolding is_forces_nmem'_def forces_nmem'_def + using forces_mem'_abs pair_in_M_iff + by (auto simp add:components_abs,blast) +qed + +subsection\Forcing for general formulas\ + +lemma leq_abs: + "\ l\M ; q\M ; p\M \ \ is_leq(##M,l,q,p) \ \q,p\\l" + unfolding is_leq_def using pair_in_M_iff by simp + +(* TODO: MOVE THIS to an appropriate place: subsubsection\The primitive recursion\ *) + +subsection\Forcing for atomic formulas in context\ + +definition + forces_eq :: "[i,i,i] \ o" (\_ forces\<^sub>a '(_ = _')\ [36,1,1] 60) where + "forces_eq \ forces_eq'(P,leq)" + +definition + forces_mem :: "[i,i,i] \ o" (\_ forces\<^sub>a '(_ \ _')\ [36,1,1] 60) where + "forces_mem \ forces_mem'(P,leq)" + +(* frc_at(P,leq,\0,t1,t2,p\) = 1*) +abbreviation is_forces_eq + where "is_forces_eq \ is_forces_eq'(##M,P,leq)" + +(* frc_at(P,leq,\1,t1,t2,p\) = 1*) +abbreviation + is_forces_mem :: "[i,i,i] \ o" where + "is_forces_mem \ is_forces_mem'(##M,P,leq)" + +lemma def_forces_eq: "p\P \ p forces\<^sub>a (t1 = t2) \ + (\s\domain(t1) \ domain(t2). \q. q\P \ q \ p \ + (q forces\<^sub>a (s \ t1) \ q forces\<^sub>a (s \ t2)))" + unfolding forces_eq_def forces_mem_def forces_eq'_def forces_mem'_def + using def_frc_at[of p 0 t1 t2 ] + unfolding bool_of_o_def + by auto + +lemma def_forces_mem: "p\P \ p forces\<^sub>a (t1 \ t2) \ + (\v\P. v \ p \ + (\q. \s. \r. r\P \ q\P \ q \ v \ \s,r\ \ t2 \ q \ r \ q forces\<^sub>a (t1 = s)))" + unfolding forces_eq'_def forces_mem'_def forces_eq_def forces_mem_def + using def_frc_at[of p 1 t1 t2] + unfolding bool_of_o_def + by auto + +lemma forces_eq_abs : + "\p\M ; t1\M ; t2\M\ \ is_forces_eq(p,t1,t2) \ p forces\<^sub>a (t1 = t2)" + unfolding forces_eq_def + using forces_eq'_abs by simp + +lemma forces_mem_abs : + "\p\M ; t1\M ; t2\M\ \ is_forces_mem(p,t1,t2) \ p forces\<^sub>a (t1 \ t2)" + unfolding forces_mem_def + using forces_mem'_abs + by simp + +definition + forces_neq :: "[i,i,i] \ o" (\_ forces\<^sub>a '(_ \ _')\ [36,1,1] 60) where + "p forces\<^sub>a (t1 \ t2) \ \ (\q\P. q\p \ q forces\<^sub>a (t1 = t2))" + +definition + forces_nmem :: "[i,i,i] \ o" (\_ forces\<^sub>a '(_ \ _')\ [36,1,1] 60) where + "p forces\<^sub>a (t1 \ t2) \ \ (\q\P. q\p \ q forces\<^sub>a (t1 \ t2))" + +lemma forces_neq : + "p forces\<^sub>a (t1 \ t2) \ forces_neq'(P,leq,p,t1,t2)" + unfolding forces_neq_def forces_neq'_def forces_eq_def by simp + +lemma forces_nmem : + "p forces\<^sub>a (t1 \ t2) \ forces_nmem'(P,leq,p,t1,t2)" + unfolding forces_nmem_def forces_nmem'_def forces_mem_def by simp + +abbreviation Forces :: "[i, i, i] \ o" ("_ \ _ _" [36,36,36] 60) where + "p \ \ env \ M, ([p,P,leq,\] @ env) \ forces(\)" + +lemma sats_forces_Member : + assumes "x\nat" "y\nat" "env\list(M)" + "nth(x,env)=xx" "nth(y,env)=yy" "q\M" + shows "q \ \x \ y\ env \ q \ P \ is_forces_mem(q, xx, yy)" + unfolding forces_def + using assms P_in_M leq_in_M one_in_M + by simp + +lemma sats_forces_Equal : + assumes "a\nat" "b\nat" "env\list(M)" "nth(a,env)=x" "nth(b,env)=y" "q\M" + shows "q \ \a = b\ env \ q \ P \ is_forces_eq(q, x, y)" + unfolding forces_def + using assms P_in_M leq_in_M one_in_M + by simp + +lemma sats_forces_Nand : + assumes "\\formula" "\\formula" "env\list(M)" "p\M" + shows "p \ \\(\ \ \)\ env \ + p\P \ \(\q\M. q\P \ is_leq(##M,leq,q,p) \ + (M,[q,P,leq,\]@env \ forces(\)) \ (M,[q,P,leq,\]@env \ forces(\)))" + unfolding forces_def + using sats_is_leq_fm_auto assms sats_ren_forces_nand P_in_M leq_in_M one_in_M zero_in_M + by simp + +lemma sats_forces_Neg : + assumes "\\formula" "env\list(M)" "p\M" + shows "p \ \\\\ env \ + (p\P \ \(\q\M. q\P \ is_leq(##M,leq,q,p) \ (M, [q, P, leq, \] @ env \ forces(\))))" + unfolding Neg_def using assms sats_forces_Nand + by simp + +lemma sats_forces_Forall : + assumes "\\formula" "env\list(M)" "p\M" + shows "p \ (\\\\) env \ p \ P \ (\x\M. M,[p,P,leq,\,x] @ env \ forces(\))" + unfolding forces_def using assms sats_ren_forces_forall P_in_M leq_in_M one_in_M + by simp + +end \ \\<^locale>\forcing_data1\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Forcing_Data.thy b/thys/Independence_CH/Forcing_Data.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Forcing_Data.thy @@ -0,0 +1,143 @@ +section\Transitive set models of ZF\ +text\This theory defines locales for countable transitive models of $\ZF$, +and on top of that, one that includes a forcing notion. Weakened versions +of both locales are included, that only assume finitely many replacement +instances.\ + +theory Forcing_Data + imports + Forcing_Notions + Cohen_Posets_Relative + Interface +begin + +locale M_ctm1 = M_ZF1_trans + + fixes enum + assumes M_countable: "enum\bij(nat,M)" + +locale M_ctm1_AC = M_ctm1 + M_ZFC1_trans + +subsection\A forcing locale and generic filters\ + +txt\Ideally, countability should be separated from the assumption of this locale. +The fact is that our present proofs of the "definition of forces" (and many +consequences) and of the lemma for “forcing a value” of function +unnecessarily depend on the countability of the ground model. \ + +locale forcing_data1 = forcing_notion + M_ctm1 + + assumes P_in_M: "P \ M" + and leq_in_M: "leq \ M" + +context forcing_data1 +begin + +(* P \ M *) +lemma P_sub_M : "P\M" + using transitivity P_in_M by auto + +definition + M_generic :: "i\o" where + "M_generic(G) \ filter(G) \ (\D\M. D\P \ dense(D)\D\G\0)" + +lemma M_genericD [dest]: "M_generic(G) \ x\G \ x\P" + unfolding M_generic_def by (blast dest:filterD) + +lemma M_generic_leqD [dest]: "M_generic(G) \ p\G \ q\P \ p\q \ q\G" + unfolding M_generic_def by (blast dest:filter_leqD) + +lemma M_generic_compatD [dest]: "M_generic(G) \ p\G \ r\G \ \q\G. q\p \ q\r" + unfolding M_generic_def by (blast dest:low_bound_filter) + +lemma M_generic_denseD [dest]: "M_generic(G) \ dense(D) \ D\P \ D\M \ \q\G. q\D" + unfolding M_generic_def by blast + +lemma G_nonempty: "M_generic(G) \ G\0" + using P_in_M P_dense subset_refl[of P] + unfolding M_generic_def + by auto + +lemma one_in_G : + assumes "M_generic(G)" + shows "\ \ G" +proof - + from assms + have "G\P" + unfolding M_generic_def filter_def by simp + from \M_generic(G)\ + have "increasing(G)" + unfolding M_generic_def filter_def by simp + with \G\P\ \M_generic(G)\ + show ?thesis + using G_nonempty one_in_P one_max + unfolding increasing_def by blast +qed + +lemma G_subset_M: "M_generic(G) \ G \ M" + using transitivity[OF _ P_in_M] by auto + +declare iff_trans [trans] + +lemma generic_filter_existence: + "p\P \ \G. p\G \ M_generic(G)" +proof - + assume "p\P" + let ?D="\n\nat. (if (enum`n\P \ dense(enum`n)) then enum`n else P)" + have "\n\nat. ?D`n \ Pow(P)" + by auto + then + have "?D:nat\Pow(P)" + using lam_type by auto + have "\n\nat. dense(?D`n)" + proof(intro ballI) + fix n + assume "n\nat" + then + have "dense(?D`n) \ dense(if enum`n \ P \ dense(enum`n) then enum`n else P)" + by simp + also + have "... \ (\(enum`n \ P \ dense(enum`n)) \ dense(P)) " + using split_if by simp + finally + show "dense(?D`n)" + using P_dense \n\nat\ by auto + qed + with \?D\_\ + interpret cg: countable_generic P leq \ ?D + by (unfold_locales, auto) + from \p\P\ + obtain G where 1: "p\G \ filter(G) \ (\n\nat.(?D`n)\G\0)" + using cg.countable_rasiowa_sikorski[where M="\_. M"] P_sub_M + M_countable[THEN bij_is_fun] M_countable[THEN bij_is_surj, THEN surj_range] + unfolding cg.D_generic_def by blast + then + have "(\D\M. D\P \ dense(D)\D\G\0)" + proof (intro ballI impI) + fix D + assume "D\M" and 2: "D \ P \ dense(D) " + moreover + have "\y\M. \x\nat. enum`x= y" + using M_countable and bij_is_surj unfolding surj_def by (simp) + moreover from calculation + obtain n where Eq10: "n\nat \ enum`n = D" + by auto + moreover from calculation if_P + have "?D`n = D" + by simp + moreover + note 1 + ultimately + show "D\G\0" + by auto + qed + with 1 + show ?thesis + unfolding M_generic_def by auto +qed + +lemma one_in_M: "\ \ M" + using one_in_P P_in_M transitivity + by simp + +end \ \\<^locale>\forcing_data1\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Forcing_Main.thy b/thys/Independence_CH/Forcing_Main.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Forcing_Main.thy @@ -0,0 +1,245 @@ +section\The main theorem\ + +theory Forcing_Main + imports + Ordinals_In_MG + Choice_Axiom + ZF_Trans_Interpretations + +begin + +subsection\The generic extension is countable\ + +lemma (in forcing_data1) surj_nat_MG : "\f. f \ surj(\,M[G])" +proof - + let ?f="\n\\. val(P,G,enum`n)" + have "x \ \ \ val(P,G, enum ` x)\ M[G]" for x + using GenExt_iff[THEN iffD2, of _ G] bij_is_fun[OF M_countable] by force + then + have "?f: \ \ M[G]" + using lam_type[of \ "\n. val(P,G,enum`n)" "\_.M[G]"] by simp + moreover + have "\n\\. ?f`n = x" if "x\M[G]" for x + using that GenExt_iff[of _ G] bij_is_surj[OF M_countable] + unfolding surj_def by auto + ultimately + show ?thesis + unfolding surj_def by blast +qed + +lemma (in G_generic1) MG_eqpoll_nat: "M[G] \ \" +proof - + obtain f where "f \ surj(\,M[G])" + using surj_nat_MG by blast + then + have "M[G] \ \" + using well_ord_surj_imp_lepoll well_ord_Memrel[of \] by simp + moreover + have "\ \ M[G]" + using ext.nat_into_M subset_imp_lepoll by (auto del:lepollI) + ultimately + show ?thesis + using eqpollI by simp +qed + +subsection\Extensions of ctms of fragments of $\ZFC$\ + +lemma M_satT_imp_M_ZF2: "(M \ ZF) \ M_ZF2(M)" +proof - + assume "M \ ZF" + then + have fin: "upair_ax(##M)" "Union_ax(##M)" "power_ax(##M)" + "extensionality(##M)" "foundation_ax(##M)" "infinity_ax(##M)" + unfolding ZF_def ZF_fin_def ZFC_fm_defs satT_def + using ZFC_fm_sats[of M] by simp_all + { + fix \ env + assume "\ \ formula" "env\list(M)" + moreover from \M \ ZF\ + have "\p\formula. (M, [] \ (ZF_separation_fm(p)))" + "\p\formula. (M, [] \ (ZF_replacement_fm(p)))" + unfolding ZF_def ZF_schemes_def by auto + moreover from calculation + have "arity(\) \ succ(length(env)) \ separation(##M, \x. (M, Cons(x, env) \ \))" + "arity(\) \ succ(succ(length(env))) \ strong_replacement(##M,\x y. sats(M,\,Cons(x,Cons(y, env))))" + using sats_ZF_separation_fm_iff sats_ZF_replacement_fm_iff by simp_all + } + with fin + show "M_ZF2(M)" + by unfold_locales (simp_all add:replacement_assm_def ground_replacement_assm_def) +qed + +lemma M_satT_imp_M_ZFC2: + shows "(M \ ZFC) \ M_ZFC2(M)" +proof - + have "(M \ ZF) \ choice_ax(##M) \ M_ZFC2(M)" + using M_satT_imp_M_ZF2[of M] unfolding M_ZF2_def M_ZFC1_def M_ZFC2_def + M_ZC_basic_def M_ZF1_def M_AC_def by auto + then + show ?thesis + unfolding ZFC_def by auto +qed + +lemma M_satT_instances12_imp_M_ZF2: + assumes "(M \ \Z\ \ {\Replacement(p)\ . p \ instances1_fms \ instances2_fms})" + shows "M_ZF2(M)" +proof - + from assms + have fin: "upair_ax(##M)" "Union_ax(##M)" "power_ax(##M)" + "extensionality(##M)" "foundation_ax(##M)" "infinity_ax(##M)" + unfolding ZF_fin_def Zermelo_fms_def ZFC_fm_defs satT_def + using ZFC_fm_sats[of M] by simp_all + moreover + { + fix \ env + from \M \ \Z\ \ {\Replacement(p)\ . p \ instances1_fms \ instances2_fms}\ + have "\p\formula. (M, [] \ (ZF_separation_fm(p)))" + unfolding Zermelo_fms_def ZF_def instances1_fms_def + instances2_fms_def by auto + moreover + assume "\ \ formula" "env\list(M)" + ultimately + have "arity(\) \ succ(length(env)) \ separation(##M, \x. (M, Cons(x, env) \ \))" + using sats_ZF_separation_fm_iff by simp_all + } + moreover + { + fix \ env + assume "\ \ instances1_fms \ instances2_fms" "env\list(M)" + moreover from this and \M \ \Z\ \ {\Replacement(p)\ . p \ instances1_fms \ instances2_fms}\ + have "M, [] \ \Replacement(\)\" by auto + ultimately + have "arity(\) \ succ(succ(length(env))) \ strong_replacement(##M,\x y. sats(M,\,Cons(x,Cons(y, env))))" + using sats_ZF_replacement_fm_iff[of \] instances1_fms_type instances2_fms_type by auto + } + ultimately + show ?thesis + unfolding instances1_fms_def instances2_fms_def + by unfold_locales (simp_all add:replacement_assm_def ground_replacement_assm_def) +qed + +context G_generic1 +begin + +lemma sats_ground_repl_fm_imp_sats_ZF_replacement_fm: + assumes + "\\formula" "M, [] \ \Replacement(ground_repl_fm(\))\" + shows + "M[G], [] \ \Replacement(\)\" + using assms sats_ZF_replacement_fm_iff + by (auto simp:replacement_assm_def ground_replacement_assm_def + intro:strong_replacement_in_MG[simplified]) + +lemma satT_ground_repl_fm_imp_satT_ZF_replacement_fm: + assumes + "\ \ formula" "M \ { \Replacement(ground_repl_fm(\))\ . \ \ \}" + shows + "M[G] \ { \Replacement(\)\ . \ \ \}" + using assms sats_ground_repl_fm_imp_sats_ZF_replacement_fm + by auto + +end \ \\<^locale>\G_generic1\\ + +theorem extensions_of_ctms: + assumes + "M \ \" "Transset(M)" "M \ \Z\ \ {\Replacement(p)\ . p \ instances1_fms \ instances2_fms}" + "\ \ formula" "M \ { \Replacement(ground_repl_fm(\))\ . \ \ \}" + shows + "\N. + M \ N \ N \ \ \ Transset(N) \ M\N \ + (\\. Ord(\) \ (\ \ M \ \ \ N)) \ + ((M, []\ \AC\) \ N, [] \ \AC\) \ N \ \Z\ \ { \Replacement(\)\ . \ \ \}" +proof - + from \M \ \Z\ \ {\Replacement(p)\ . p \ instances1_fms \ instances2_fms}\ + interpret M_ZF2 M + using M_satT_instances12_imp_M_ZF2 + by simp + from \Transset(M)\ + interpret M_ZF1_trans M + using M_satT_imp_M_ZF2 + by unfold_locales + from \M \ \\ + obtain enum where "enum \ bij(\,M)" + using eqpoll_sym unfolding eqpoll_def by blast + then + interpret M_ctm2 M enum by unfold_locales simp_all + interpret forcing_data1 "2\<^bsup><\\<^esup>" seqle 0 M enum + using nat_into_M seqspace_closed seqle_in_M + by unfold_locales simp + obtain G where "M_generic(G)" "M \ M\<^bsup>s\<^esup>[G]" + using cohen_extension_is_proper + by blast + txt\Recall that \<^term>\M\<^bsup>s\<^esup>[G]\ denotes the generic extension \<^term>\M\<^bsup>2\<^bsup><\\<^esup>\<^esup>[G]\ + of \<^term>\M\ using the poset of sequences \<^term>\2\<^bsup><\\<^esup>\.\ + then + interpret G_generic1 "2\<^bsup><\\<^esup>" seqle 0 _ enum G by unfold_locales + interpret MG: M_Z_basic "M\<^bsup>s\<^esup>[G]" + using generic pairing_in_MG + Union_MG extensionality_in_MG power_in_MG + foundation_in_MG replacement_assm_MG + separation_in_MG infinity_in_MG replacement_ax1 + by unfold_locales simp + have "M, []\ \AC\ \ M\<^bsup>s\<^esup>[G], [] \ \AC\" + proof - + assume "M, [] \ \AC\" + then + have "choice_ax(##M)" + unfolding ZF_choice_fm_def using ZF_choice_auto by simp + then + have "choice_ax(##M\<^bsup>s\<^esup>[G])" using choice_in_MG by simp + then + show "M\<^bsup>s\<^esup>[G], [] \ \AC\" + using ZF_choice_auto sats_ZFC_iff_sats_ZF_AC + unfolding ZF_choice_fm_def by simp + qed + moreover + note \M \ M\<^bsup>s\<^esup>[G]\ \M \ { \Replacement(ground_repl_fm(\))\ . \ \ \}\ \\ \ formula\ + moreover + have "Transset(M\<^bsup>s\<^esup>[G])" using Transset_MG . + moreover + have "M \ M\<^bsup>s\<^esup>[G]" using M_subset_MG[OF one_in_G] generic by simp + ultimately + show ?thesis + using Ord_MG_iff MG_eqpoll_nat ext.M_satT_Zermelo_fms + satT_ground_repl_fm_imp_satT_ZF_replacement_fm[of \] + by (rule_tac x="M\<^bsup>s\<^esup>[G]" in exI, auto) +qed + +lemma ZF_replacement_instances12_sub_ZF: "{\Replacement(p)\ . p \ instances1_fms \ instances2_fms} \ ZF" + using instances1_fms_type instances2_fms_type unfolding ZF_def ZF_schemes_def by auto + +theorem extensions_of_ctms_ZF: + assumes + "M \ \" "Transset(M)" "M \ ZF" + shows + "\N. + M \ N \ N \ \ \ Transset(N) \ N \ ZF \ M\N \ + (\\. Ord(\) \ (\ \ M \ \ \ N)) \ + ((M, []\ \AC\) \ N \ ZFC)" +proof - + from assms + have "\N. + M \ N \ N \ \ \ Transset(N) \ M\N \ + (\\. Ord(\) \ (\ \ M \ \ \ N)) \ + ((M, []\ \AC\) \ N, [] \ \AC\) \ N \ \Z\ \ { \Replacement(\)\ . \ \ formula}" + using extensions_of_ctms[of M formula] satT_ZF_imp_satT_Z[of M] + satT_mono[OF _ ground_repl_fm_sub_ZF, of M] + satT_mono[OF _ ZF_replacement_instances12_sub_ZF, of M] + by (auto simp: satT_Un_iff) + then + obtain N where "N \ \Z\ \ { \Replacement(\)\ . \ \ formula}" "M \ N" "N \ \" "Transset(N)" + "M \ N" "(\\. Ord(\) \ \ \ M \ \ \ N)" + "(M, []\ \AC\) \ N, [] \ \AC\" + by blast + moreover from \N \ \Z\ \ { \Replacement(\)\ . \ \ formula}\ + have "N \ ZF" + using satT_Z_ZF_replacement_imp_satT_ZF by auto + moreover from this and \(M, []\ \AC\) \ N, [] \ \AC\\ + have "(M, []\ \AC\) \ N \ ZFC" + using sats_ZFC_iff_sats_ZF_AC by simp + ultimately + show ?thesis + by auto +qed + +end \ No newline at end of file diff --git a/thys/Independence_CH/Forcing_Notions.thy b/thys/Independence_CH/Forcing_Notions.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Forcing_Notions.thy @@ -0,0 +1,424 @@ +section\Forcing notions\ +text\This theory defines a locale for forcing notions, that is, + preorders with a distinguished maximum element.\ + +theory Forcing_Notions + imports + "ZF-Constructible.Relative" + "Delta_System_Lemma.ZF_Library" +begin + +subsection\Basic concepts\ +text\We say that two elements $p,q$ are + \<^emph>\compatible\ if they have a lower bound in $P$\ +definition compat_in :: "i\i\i\i\o" where + "compat_in(A,r,p,q) \ \d\A . \d,p\\r \ \d,q\\r" + +lemma compat_inI : + "\ d\A ; \d,p\\r ; \d,g\\r \ \ compat_in(A,r,p,g)" + by (auto simp add: compat_in_def) + +lemma refl_compat: + "\ refl(A,r) ; \p,q\ \ r | p=q | \q,p\ \ r ; p\A ; q\A\ \ compat_in(A,r,p,q)" + by (auto simp add: refl_def compat_inI) + +lemma chain_compat: + "refl(A,r) \ linear(A,r) \ (\p\A.\q\A. compat_in(A,r,p,q))" + by (simp add: refl_compat linear_def) + +lemma subset_fun_image: "f:N\P \ f``N\P" + by (auto simp add: image_fun apply_funtype) + +lemma refl_monot_domain: "refl(B,r) \ A\B \ refl(A,r)" + unfolding refl_def by blast + +locale forcing_notion = + fixes P leq one + assumes one_in_P: "one \ P" + and leq_preord: "preorder_on(P,leq)" + and one_max: "\p\P. \p,one\\leq" +begin + +notation one (\\\) + +abbreviation Leq :: "[i, i] \ o" (infixl "\" 50) + where "x \ y \ \x,y\\leq" + +lemma refl_leq: + "r\P \ r\r" + using leq_preord unfolding preorder_on_def refl_def by simp + +text\A set $D$ is \<^emph>\dense\ if every element $p\in P$ has a lower +bound in $D$.\ +definition + dense :: "i\o" where + "dense(D) \ \p\P. \d\D . d\p" + +text\There is also a weaker definition which asks for +a lower bound in $D$ only for the elements below some fixed +element $q$.\ +definition + dense_below :: "i\i\o" where + "dense_below(D,q) \ \p\P. p\q \ (\d\D. d\P \ d\p)" + +lemma P_dense: "dense(P)" + by (insert leq_preord, auto simp add: preorder_on_def refl_def dense_def) + +definition + increasing :: "i\o" where + "increasing(F) \ \x\F. \ p \ P . x\p \ p\F" + +definition + compat :: "i\i\o" where + "compat(p,q) \ compat_in(P,leq,p,q)" + +lemma leq_transD: "a\b \ b\c \ a \ P\ b \ P\ c \ P\ a\c" + using leq_preord trans_onD unfolding preorder_on_def by blast + +lemma leq_transD': "A\P \ a\b \ b\c \ a \ A \ b \ P\ c \ P\ a\c" + using leq_preord trans_onD subsetD unfolding preorder_on_def by blast + +lemma compatD[dest!]: "compat(p,q) \ \d\P. d\p \ d\q" + unfolding compat_def compat_in_def . + +abbreviation Incompatible :: "[i, i] \ o" (infixl "\" 50) + where "p \ q \ \ compat(p,q)" + +lemma compatI[intro!]: "d\P \ d\p \ d\q \ compat(p,q)" + unfolding compat_def compat_in_def by blast + +lemma denseD [dest]: "dense(D) \ p\P \ \d\D. d\ p" + unfolding dense_def by blast + +lemma denseI [intro!]: "\ \p. p\P \ \d\D. d\ p \ \ dense(D)" + unfolding dense_def by blast + +lemma dense_belowD [dest]: + assumes "dense_below(D,p)" "q\P" "q\p" + shows "\d\D. d\P \ d\q" + using assms unfolding dense_below_def by simp + +lemma dense_belowI [intro!]: + assumes "\q. q\P \ q\p \ \d\D. d\P \ d\q" + shows "dense_below(D,p)" + using assms unfolding dense_below_def by simp + +lemma dense_below_cong: "p\P \ D = D' \ dense_below(D,p) \ dense_below(D',p)" + by blast + +lemma dense_below_cong': "p\P \ \\x. x\P \ Q(x) \ Q'(x)\ \ + dense_below({q\P. Q(q)},p) \ dense_below({q\P. Q'(q)},p)" + by blast + +lemma dense_below_mono: "p\P \ D \ D' \ dense_below(D,p) \ dense_below(D',p)" + by blast + +lemma dense_below_under: + assumes "dense_below(D,p)" "p\P" "q\P" "q\p" + shows "dense_below(D,q)" + using assms leq_transD by blast + +lemma ideal_dense_below: + assumes "\q. q\P \ q\p \ q\D" + shows "dense_below(D,p)" + using assms refl_leq by blast + +lemma dense_below_dense_below: + assumes "dense_below({q\P. dense_below(D,q)},p)" "p\P" + shows "dense_below(D,p)" + using assms leq_transD refl_leq by blast + +text\A filter is an increasing set $G$ with all its elements +being compatible in $G$.\ +definition + filter :: "i\o" where + "filter(G) \ G\P \ increasing(G) \ (\p\G. \q\G. compat_in(G,leq,p,q))" + +lemma filterD : "filter(G) \ x \ G \ x \ P" + by (auto simp add : subsetD filter_def) + +lemma filter_leqD : "filter(G) \ x \ G \ y \ P \ x\y \ y \ G" + by (simp add: filter_def increasing_def) + +lemma filter_imp_compat: "filter(G) \ p\G \ q\G \ compat(p,q)" + unfolding filter_def compat_in_def compat_def by blast + +lemma low_bound_filter: \ \says the compatibility is attained inside G\ + assumes "filter(G)" and "p\G" and "q\G" + shows "\r\G. r\p \ r\q" + using assms + unfolding compat_in_def filter_def by blast + +text\We finally introduce the upward closure of a set +and prove that the closure of $A$ is a filter if its elements are +compatible in $A$.\ +definition + upclosure :: "i\i" where + "upclosure(A) \ {p\P.\a\A. a\p}" + +lemma upclosureI [intro] : "p\P \ a\A \ a\p \ p\upclosure(A)" + by (simp add:upclosure_def, auto) + +lemma upclosureE [elim] : + "p\upclosure(A) \ (\x a. x\P \ a\A \ a\x \ R) \ R" + by (auto simp add:upclosure_def) + +lemma upclosureD [dest] : + "p\upclosure(A) \ \a\A.(a\p) \ p\P" + by (simp add:upclosure_def) + +lemma upclosure_increasing : + assumes "A\P" + shows "increasing(upclosure(A))" + unfolding increasing_def upclosure_def + using leq_transD'[OF \A\P\] by auto + +lemma upclosure_in_P: "A \ P \ upclosure(A) \ P" + using subsetI upclosure_def by simp + +lemma A_sub_upclosure: "A \ P \ A\upclosure(A)" + using subsetI leq_preord + unfolding upclosure_def preorder_on_def refl_def by auto + +lemma elem_upclosure: "A\P \ x\A \ x\upclosure(A)" + by (blast dest:A_sub_upclosure) + +lemma closure_compat_filter: + assumes "A\P" "(\p\A.\q\A. compat_in(A,leq,p,q))" + shows "filter(upclosure(A))" + unfolding filter_def +proof(auto) + show "increasing(upclosure(A))" + using assms upclosure_increasing by simp +next + let ?UA="upclosure(A)" + show "compat_in(upclosure(A), leq, p, q)" if "p\?UA" "q\?UA" for p q + proof - + from that + obtain a b where 1:"a\A" "b\A" "a\p" "b\q" "p\P" "q\P" + using upclosureD[OF \p\?UA\] upclosureD[OF \q\?UA\] by auto + with assms(2) + obtain d where "d\A" "d\a" "d\b" + unfolding compat_in_def by auto + with 1 + have "d\p" "d\q" "d\?UA" + using A_sub_upclosure[THEN subsetD] \A\P\ + leq_transD'[of A d a] leq_transD'[of A d b] by auto + then + show ?thesis unfolding compat_in_def by auto + qed +qed + +lemma aux_RS1: "f \ N \ P \ n\N \ f`n \ upclosure(f ``N)" + using elem_upclosure[OF subset_fun_image] image_fun + by (simp, blast) + +lemma decr_succ_decr: + assumes "f \ nat \ P" "preorder_on(P,leq)" + "\n\nat. \f ` succ(n), f ` n\ \ leq" + "m\nat" + shows "n\nat \ n\m \ \f ` m, f ` n\ \ leq" + using \m\_\ +proof(induct m) + case 0 + then show ?case using assms refl_leq by simp +next + case (succ x) + then + have 1:"f`succ(x) \ f`x" "f`n\P" "f`x\P" "f`succ(x)\P" + using assms by simp_all + consider (lt) "n nat \ P" + "\n\nat. \f ` succ(n), f ` n\ \ leq" + "trans[P](leq)" + shows "linear(f `` nat, leq)" +proof - + have "preorder_on(P,leq)" + unfolding preorder_on_def using assms by simp + { + fix n m + assume "n\nat" "m\nat" + then + have "f`m \ f`n \ f`n \ f`m" + proof(cases "m\n") + case True + with \n\_\ \m\_\ + show ?thesis + using decr_succ_decr[of f n m] assms leI \preorder_on(P,leq)\ by simp + next + case False + with \n\_\ \m\_\ + show ?thesis + using decr_succ_decr[of f m n] assms leI not_le_iff_lt \preorder_on(P,leq)\ by simp + qed + } + then + show ?thesis + unfolding linear_def using ball_image_simp assms by auto +qed + +end \ \\<^locale>\forcing_notion\\ + +subsection\Towards Rasiowa-Sikorski Lemma (RSL)\ +locale countable_generic = forcing_notion + + fixes \ + assumes countable_subs_of_P: "\ \ nat\Pow(P)" + and seq_of_denses: "\n \ nat. dense(\`n)" + +begin + +definition + D_generic :: "i\o" where + "D_generic(G) \ filter(G) \ (\n\nat.(\`n)\G\0)" + +text\The next lemma identifies a sufficient condition for obtaining +RSL.\ +lemma RS_sequence_imp_rasiowa_sikorski: + assumes + "p\P" "f : nat\P" "f ` 0 = p" + "\n. n\nat \ f ` succ(n)\ f ` n \ f ` succ(n) \ \ ` n" + shows + "\G. p\G \ D_generic(G)" +proof - + note assms + moreover from this + have "f``nat \ P" + by (simp add:subset_fun_image) + moreover from calculation + have "refl(f``nat, leq) \ trans[P](leq)" + using leq_preord unfolding preorder_on_def by (blast intro:refl_monot_domain) + moreover from calculation + have "\n\nat. f ` succ(n)\ f ` n" by (simp) + moreover from calculation + have "linear(f``nat, leq)" + using leq_preord and decr_seq_linear unfolding preorder_on_def by (blast) + moreover from calculation + have "(\p\f``nat.\q\f``nat. compat_in(f``nat,leq,p,q))" + using chain_compat by (auto) + ultimately + have "filter(upclosure(f``nat))" (is "filter(?G)") + using closure_compat_filter by simp + moreover + have "\n\nat. \ ` n \ ?G \ 0" + proof + fix n + assume "n\nat" + with assms + have "f`succ(n) \ ?G \ f`succ(n) \ \ ` n" + using aux_RS1 by simp + then + show "\ ` n \ ?G \ 0" by blast + qed + moreover from assms + have "p \ ?G" + using aux_RS1 by auto + ultimately + show ?thesis unfolding D_generic_def by auto +qed + +end \ \\<^locale>\countable_generic\\ + +text\Now, the following recursive definition will fulfill the +requirements of lemma \<^term>\RS_sequence_imp_rasiowa_sikorski\ \ + +consts RS_seq :: "[i,i,i,i,i,i] \ i" +primrec + "RS_seq(0,P,leq,p,enum,\) = p" + "RS_seq(succ(n),P,leq,p,enum,\) = + enum`(\ m. \enum`m, RS_seq(n,P,leq,p,enum,\)\ \ leq \ enum`m \ \ ` n)" + +context countable_generic +begin + +lemma countable_RS_sequence_aux: + fixes p enum + defines "f(n) \ RS_seq(n,P,leq,p,enum,\)" + and "Q(q,k,m) \ enum`m\ q \ enum`m \ \ ` k" + assumes "n\nat" "p\P" "P \ range(enum)" "enum:nat\M" + "\x k. x\P \ k\nat \ \q\P. q\ x \ q \ \ ` k" + shows + "f(succ(n)) \ P \ f(succ(n))\ f(n) \ f(succ(n)) \ \ ` n" + using \n\nat\ +proof (induct) + case 0 + from assms + obtain q where "q\P" "q\ p" "q \ \ ` 0" by blast + moreover from this and \P \ range(enum)\ + obtain m where "m\nat" "enum`m = q" + using Pi_rangeD[OF \enum:nat\M\] by blast + moreover + have "\`0 \ P" + using apply_funtype[OF countable_subs_of_P] by simp + moreover note \p\P\ + ultimately + show ?case + using LeastI[of "Q(p,0)" m] unfolding Q_def f_def by auto +next + case (succ n) + with assms + obtain q where "q\P" "q\ f(succ(n))" "q \ \ ` succ(n)" by blast + moreover from this and \P \ range(enum)\ + obtain m where "m\nat" "enum`m\ f(succ(n))" "enum`m \ \ ` succ(n)" + using Pi_rangeD[OF \enum:nat\M\] by blast + moreover note succ + moreover from calculation + have "\`succ(n) \ P" + using apply_funtype[OF countable_subs_of_P] by auto + ultimately + show ?case + using LeastI[of "Q(f(succ(n)),succ(n))" m] unfolding Q_def f_def by auto +qed + +lemma countable_RS_sequence: + fixes p enum + defines "f \ \n\nat. RS_seq(n,P,leq,p,enum,\)" + and "Q(q,k,m) \ enum`m\ q \ enum`m \ \ ` k" + assumes "n\nat" "p\P" "P \ range(enum)" "enum:nat\M" + shows + "f`0 = p" "f`succ(n)\ f`n \ f`succ(n) \ \ ` n" "f`succ(n) \ P" +proof - + from assms + show "f`0 = p" by simp + { + fix x k + assume "x\P" "k\nat" + then + have "\q\P. q\ x \ q \ \ ` k" + using seq_of_denses apply_funtype[OF countable_subs_of_P] + unfolding dense_def by blast + } + with assms + show "f`succ(n)\ f`n \ f`succ(n) \ \ ` n" "f`succ(n)\P" + unfolding f_def using countable_RS_sequence_aux by simp_all +qed + +lemma RS_seq_type: + assumes "n \ nat" "p\P" "P \ range(enum)" "enum:nat\M" + shows "RS_seq(n,P,leq,p,enum,\) \ P" + using assms countable_RS_sequence(1,3) + by (induct;simp) + +lemma RS_seq_funtype: + assumes "p\P" "P \ range(enum)" "enum:nat\M" + shows "(\n\nat. RS_seq(n,P,leq,p,enum,\)): nat \ P" + using assms lam_type RS_seq_type by auto + +lemmas countable_rasiowa_sikorski = + RS_sequence_imp_rasiowa_sikorski[OF _ RS_seq_funtype countable_RS_sequence(1,2)] + +end \ \\<^locale>\countable_generic\\ + +end diff --git a/thys/Independence_CH/Forcing_Theorems.thy b/thys/Independence_CH/Forcing_Theorems.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Forcing_Theorems.thy @@ -0,0 +1,1516 @@ +section\The Forcing Theorems\ + +theory Forcing_Theorems + imports + Cohen_Posets_Relative + Forces_Definition + Names + +begin + +context forcing_data1 +begin + +subsection\The forcing relation in context\ + +lemma separation_forces : + assumes + fty: "\\formula" and + far: "arity(\)\length(env)" and + envty: "env\list(M)" + shows + "separation(##M,\p. (p \ \ env))" + using separation_ax arity_forces far fty P_in_M leq_in_M one_in_M envty arity_forces_le + transitivity[of _ P] + by simp + +lemma Collect_forces : + assumes + "\\formula" and + "arity(\)\length(env)" and + "env\list(M)" + shows + "{p\P . p \ \ env} \ M" + using assms separation_forces separation_closed P_in_M + by simp + +lemma forces_mem_iff_dense_below: "p\P \ p forces\<^sub>a (t1 \ t2) \ dense_below( + {q\P. \s. \r. r\P \ \s,r\ \ t2 \ q\r \ q forces\<^sub>a (t1 = s)} + ,p)" + using def_forces_mem[of p t1 t2] by blast + +subsection\Kunen 2013, Lemma IV.2.37(a)\ + +lemma strengthening_eq: + assumes "p\P" "r\P" "r\p" "p forces\<^sub>a (t1 = t2)" + shows "r forces\<^sub>a (t1 = t2)" + using assms def_forces_eq[of _ t1 t2] leq_transD by blast + (* Long proof *) + (* +proof - + { + fix s q + assume "q\ r" "q\P" + with assms + have "q\p" + using leq_preord unfolding preorder_on_def trans_on_def by blast + moreover + note \q\P\ assms + moreover + assume "s\domain(t1) \ domain(t2)" + ultimately + have "q forces\<^sub>a ( s \ t1) \ q forces\<^sub>a ( s \ t2)" + using def_forces_eq[of p t1 t2] by simp + } + with \r\P\ + show ?thesis using def_forces_eq[of r t1 t2] by blast +qed +*) + +subsection\Kunen 2013, Lemma IV.2.37(a)\ +lemma strengthening_mem: + assumes "p\P" "r\P" "r\p" "p forces\<^sub>a (t1 \ t2)" + shows "r forces\<^sub>a (t1 \ t2)" + using assms forces_mem_iff_dense_below dense_below_under by auto + +subsection\Kunen 2013, Lemma IV.2.37(b)\ +lemma density_mem: + assumes "p\P" + shows "p forces\<^sub>a (t1 \ t2) \ dense_below({q\P. q forces\<^sub>a (t1 \ t2)},p)" +proof + assume "p forces\<^sub>a (t1 \ t2)" + with assms + show "dense_below({q\P. q forces\<^sub>a (t1 \ t2)},p)" + using forces_mem_iff_dense_below strengthening_mem[of p] ideal_dense_below by auto +next + assume "dense_below({q \ P . q forces\<^sub>a ( t1 \ t2)}, p)" + with assms + have "dense_below({q\P. + dense_below({q'\P. \s r. r \ P \ \s,r\\t2 \ q'\r \ q' forces\<^sub>a (t1 = s)},q) + },p)" + using forces_mem_iff_dense_below by simp + with assms + show "p forces\<^sub>a (t1 \ t2)" + using dense_below_dense_below forces_mem_iff_dense_below[of p t1 t2] by blast +qed + +lemma aux_density_eq: + assumes + "dense_below( + {q'\P. \q. q\P \ q\q' \ q forces\<^sub>a (s \ t1) \ q forces\<^sub>a (s \ t2)} + ,p)" + "q forces\<^sub>a (s \ t1)" "q\P" "p\P" "q\p" + shows + "dense_below({r\P. r forces\<^sub>a (s \ t2)},q)" +proof + fix r + assume "r\P" "r\q" + moreover from this and \p\P\ \q\p\ \q\P\ + have "r\p" + using leq_transD by simp + moreover + note \q forces\<^sub>a (s \ t1)\ \dense_below(_,p)\ \q\P\ + ultimately + obtain q1 where "q1\r" "q1\P" "q1 forces\<^sub>a (s \ t2)" + using strengthening_mem[of q _ s t1] refl_leq leq_transD[of _ r q] by blast + then + show "\d\{r \ P . r forces\<^sub>a ( s \ t2)}. d \ P \ d\ r" + by blast +qed + +(* Kunen 2013, Lemma IV.2.37(b) *) +lemma density_eq: + assumes "p\P" + shows "p forces\<^sub>a (t1 = t2) \ dense_below({q\P. q forces\<^sub>a (t1 = t2)},p)" +proof + assume "p forces\<^sub>a (t1 = t2)" + with \p\P\ + show "dense_below({q\P. q forces\<^sub>a (t1 = t2)},p)" + using strengthening_eq ideal_dense_below by auto +next + assume "dense_below({q\P. q forces\<^sub>a (t1 = t2)},p)" + { + fix s q + let ?D1="{q'\P. \s\domain(t1) \ domain(t2). \q. q \ P \ q\q' \ + q forces\<^sub>a (s \ t1)\q forces\<^sub>a (s \ t2)}" + let ?D2="{q'\P. \q. q\P \ q\q' \ q forces\<^sub>a (s \ t1) \ q forces\<^sub>a (s \ t2)}" + assume "s\domain(t1) \ domain(t2)" + then + have "?D1\?D2" by blast + with \dense_below(_,p)\ + have "dense_below({q'\P. \s\domain(t1) \ domain(t2). \q. q \ P \ q\q' \ + q forces\<^sub>a (s \ t1)\q forces\<^sub>a (s \ t2)},p)" + using dense_below_cong'[OF \p\P\ def_forces_eq[of _ t1 t2]] by simp + with \p\P\ \?D1\?D2\ + have "dense_below({q'\P. \q. q\P \ q\q' \ + q forces\<^sub>a (s \ t1) \ q forces\<^sub>a (s \ t2)},p)" + using dense_below_mono by simp + moreover from this (* Automatic tools can't handle this symmetry + in order to apply aux_density_eq below *) + have "dense_below({q'\P. \q. q\P \ q\q' \ + q forces\<^sub>a (s \ t2) \ q forces\<^sub>a (s \ t1)},p)" + by blast + moreover + assume "q \ P" "q\p" + moreover + note \p\P\ + ultimately (*We can omit the next step but it is slower *) + have "q forces\<^sub>a (s \ t1) \ dense_below({r\P. r forces\<^sub>a (s \ t2)},q)" + "q forces\<^sub>a (s \ t2) \ dense_below({r\P. r forces\<^sub>a (s \ t1)},q)" + using aux_density_eq by simp_all + then + have "q forces\<^sub>a ( s \ t1) \ q forces\<^sub>a ( s \ t2)" + using density_mem[OF \q\P\] by blast + } + with \p\P\ + show "p forces\<^sub>a (t1 = t2)" using def_forces_eq by blast +qed + +subsection\Kunen 2013, Lemma IV.2.38\ +lemma not_forces_neq: + assumes "p\P" + shows "p forces\<^sub>a (t1 = t2) \ \ (\q\P. q\p \ q forces\<^sub>a (t1 \ t2))" + using assms density_eq unfolding forces_neq_def by blast + +lemma not_forces_nmem: + assumes "p\P" + shows "p forces\<^sub>a (t1 \ t2) \ \ (\q\P. q\p \ q forces\<^sub>a (t1 \ t2))" + using assms density_mem unfolding forces_nmem_def by blast + +subsection\The relation of forcing and atomic formulas\ +lemma Forces_Equal: + assumes + "p\P" "t1\M" "t2\M" "env\list(M)" "nth(n,env) = t1" "nth(m,env) = t2" "n\nat" "m\nat" + shows + "(p \ Equal(n,m) env) \ p forces\<^sub>a (t1 = t2)" + using assms sats_forces_Equal forces_eq_abs transitivity P_in_M + by simp + +lemma Forces_Member: + assumes + "p\P" "t1\M" "t2\M" "env\list(M)" "nth(n,env) = t1" "nth(m,env) = t2" "n\nat" "m\nat" + shows + "(p \ Member(n,m) env) \ p forces\<^sub>a (t1 \ t2)" + using assms sats_forces_Member forces_mem_abs transitivity P_in_M + by simp + +lemma Forces_Neg: + assumes + "p\P" "env \ list(M)" "\\formula" + shows + "(p \ Neg(\) env) \ \(\q\M. q\P \ q\p \ (q \ \ env))" + using assms sats_forces_Neg transitivity P_in_M pair_in_M_iff leq_in_M leq_abs + by simp + +subsection\The relation of forcing and connectives\ + +lemma Forces_Nand: + assumes + "p\P" "env \ list(M)" "\\formula" "\\formula" + shows + "(p \ Nand(\,\) env) \ \(\q\M. q\P \ q\p \ (q \ \ env) \ (q \ \ env))" + using assms sats_forces_Nand transitivity + P_in_M pair_in_M_iff leq_in_M leq_abs by simp + +lemma Forces_And_aux: + assumes + "p\P" "env \ list(M)" "\\formula" "\\formula" + shows + "p \ And(\,\) env \ + (\q\M. q\P \ q\p \ (\r\M. r\P \ r\q \ (r \ \ env) \ (r \ \ env)))" + unfolding And_def using assms Forces_Neg Forces_Nand by (auto simp only:) + +lemma Forces_And_iff_dense_below: + assumes + "p\P" "env \ list(M)" "\\formula" "\\formula" + shows + "(p \ And(\,\) env) \ dense_below({r\P. (r \ \ env) \ (r \ \ env) },p)" + unfolding dense_below_def using Forces_And_aux assms + by (auto dest:transitivity[OF _ P_in_M]; rename_tac q; drule_tac x=q in bspec)+ + +lemma Forces_Forall: + assumes + "p\P" "env \ list(M)" "\\formula" + shows + "(p \ Forall(\) env) \ (\x\M. (p \ \ ([x] @ env)))" + using sats_forces_Forall assms transitivity[OF _ P_in_M] + by simp + +(* "x\val(P,G,\) \ \\. \p\G. \\,p\\\ \ val(P,G,\) = x" *) +bundle some_rules = elem_of_val_pair [dest] + +context + includes some_rules +begin + +lemma elem_of_valI: "\\. \p\P. p\G \ \\,p\\\ \ val(P,G,\) = x \ x\val(P,G,\)" + by (subst def_val, auto) + +lemma GenExt_iff: "x\M[G] \ (\\\M. x = val(P,G,\))" + unfolding GenExt_def by simp + +subsection\Kunen 2013, Lemma IV.2.29\ +lemma generic_inter_dense_below: + assumes "D\M" "M_generic(G)" "dense_below(D,p)" "p\G" + shows "D \ G \ 0" +proof - + let ?D="{q\P. p\q \ q\D}" + have "dense(?D)" + proof + fix r + assume "r\P" + show "\d\{q \ P . p \ q \ q \ D}. d \ r" + proof (cases "p \ r") + case True + with \r\P\ + (* Automatic tools can't handle this case for some reason... *) + show ?thesis using refl_leq[of r] by (intro bexI) (blast+) + next + case False + then + obtain s where "s\P" "s\p" "s\r" by blast + with assms \r\P\ + show ?thesis + using dense_belowD[OF assms(3), of s] leq_transD[of _ s r] + by blast + qed + qed + have "?D\P" by auto + (* D\M *) + let ?d_fm="\\\compat_in_fm(1, 2, 3, 0) \ \ \0 \ 4\\" + have 1:"p\M" + using \M_generic(G)\ M_genericD transitivity[OF _ P_in_M] + \p\G\ by simp + moreover + have "?d_fm\formula" by simp + moreover + have "arity(?d_fm) = 5" + by (auto simp add: arity) + moreover + have "(M, [q,P,leq,p,D] \ ?d_fm) \ (\ is_compat_in(##M,P,leq,p,q) \ q\D)" + if "q\M" for q + using that sats_compat_in_fm P_in_M leq_in_M 1 \D\M\ zero_in_M + by simp + moreover + have "(\ is_compat_in(##M,P,leq,p,q) \ q\D) \ p\q \ q\D" if "q\M" for q + unfolding compat_def + using that compat_in_abs P_in_M leq_in_M 1 + by simp + ultimately + have "?D\M" + using Collect_in_M[of ?d_fm "[P,leq,p,D]"] P_in_M leq_in_M \D\M\ + by simp + note asm = \M_generic(G)\ \dense(?D)\ \?D\P\ \?D\M\ + obtain x where "x\G" "x\?D" using M_generic_denseD[OF asm] + by force (* by (erule bexE) does it, but the other automatic tools don't *) + moreover from this and \M_generic(G)\ + have "x\D" + using M_generic_compatD[OF _ \p\G\, of x] refl_leq compatI[of _ p x] + by force + ultimately + show ?thesis by auto +qed + +subsection\Auxiliary results for Lemma IV.2.40(a)\ +lemma IV240a_mem_Collect: + assumes + "\\M" "\\M" + shows + "{q\P. \\. \r. r\P \ \\,r\ \ \ \ q\r \ q forces\<^sub>a (\ = \)}\M" +proof - + let ?rel_pred= "\M x a1 a2 a3 a4. \\[M]. \r[M]. \\r[M]. + r\a1 \ pair(M,\,r,\r) \ \r\a4 \ is_leq(M,a2,x,r) \ is_forces_eq'(M,a1,a2,x,a3,\)" + let ?\="Exists(Exists(Exists(And(Member(1,4),And(pair_fm(2,1,0), + And(Member(0,7),And(is_leq_fm(5,3,1),forces_eq_fm(4,5,3,6,2))))))))" + have "\\M \ r\M" if "\\, r\ \ \" for \ r + using that \\\M\ pair_in_M_iff transitivity[of "\\,r\" \] by simp + then + have "?rel_pred(##M,q,P,leq,\,\) \ (\\. \r. r\P \ \\,r\ \ \ \ q\r \ q forces\<^sub>a (\ = \))" + if "q\M" for q + unfolding forces_eq_def + using assms that P_in_M leq_in_M leq_abs forces_eq'_abs pair_in_M_iff + by auto + moreover + have "(M, [q,P,leq,\,\] \ ?\) \ ?rel_pred(##M,q,P,leq,\,\)" if "q\M" for q + using assms that sats_forces_eq_fm sats_is_leq_fm P_in_M leq_in_M zero_in_M + by simp + moreover + have "?\\formula" by simp + moreover + have "arity(?\)=5" + using arity_forces_eq_fm + by (simp add:ord_simp_union arity) + ultimately + show ?thesis + unfolding forces_eq_def using P_in_M leq_in_M assms Collect_in_M[of ?\ "[P,leq,\,\]"] + by simp +qed + +(* Lemma IV.2.40(a), membership *) +lemma IV240a_mem: + assumes + "M_generic(G)" "p\G" "\\M" "\\M" "p forces\<^sub>a (\ \ \)" + "\q \. q\P \ q\G \ \\domain(\) \ q forces\<^sub>a (\ = \) \ + val(P,G,\) = val(P,G,\)" (* inductive hypothesis *) + shows + "val(P,G,\)\val(P,G,\)" +proof (intro elem_of_valI) + let ?D="{q\P. \\. \r. r\P \ \\,r\ \ \ \ q\r \ q forces\<^sub>a (\ = \)}" + from \M_generic(G)\ \p\G\ + have "p\P" by blast + moreover + note \\\M\ \\\M\ + ultimately + have "?D \ M" using IV240a_mem_Collect by simp + moreover from assms \p\P\ + have "dense_below(?D,p)" + using forces_mem_iff_dense_below by simp + moreover + note \M_generic(G)\ \p\G\ + ultimately + obtain q where "q\G" "q\?D" using generic_inter_dense_below by blast + then + obtain \ r where "r\P" "\\,r\ \ \" "q\r" "q forces\<^sub>a (\ = \)" by blast + moreover from this and \q\G\ assms + have "r \ G" "val(P,G,\) = val(P,G,\)" by blast+ + ultimately + show "\ \. \p\P. p \ G \ \\, p\ \ \ \ val(P,G, \) = val(P,G, \)" by auto +qed + +(* Example IV.2.36 (next two lemmas) *) +lemma refl_forces_eq:"p\P \ p forces\<^sub>a (x = x)" + using def_forces_eq by simp + +lemma forces_memI: "\\,r\\\ \ p\P \ r\P \ p\r \ p forces\<^sub>a (\ \ \)" + using refl_forces_eq[of _ \] leq_transD refl_leq + by (blast intro:forces_mem_iff_dense_below[THEN iffD2]) + +(* Lemma IV.2.40(a), equality, first inclusion *) +lemma IV240a_eq_1st_incl: + assumes + "M_generic(G)" "p\G" "p forces\<^sub>a (\ = \)" + and + IH:"\q \. q\P \ q\G \ \\domain(\) \ domain(\) \ + (q forces\<^sub>a (\ \ \) \ val(P,G,\) \ val(P,G,\)) \ + (q forces\<^sub>a (\ \ \) \ val(P,G,\) \ val(P,G,\))" + (* Strong enough for this case: *) + (* IH:"\q \. q\P \ \\domain(\) \ q forces\<^sub>a (\ \ \) \ + val(P,G,\) \ val(P,G,\)" *) + shows + "val(P,G,\) \ val(P,G,\)" +proof + fix x + assume "x\val(P,G,\)" + then + obtain \ r where "\\,r\\\" "r\G" "val(P,G,\)=x" by blast + moreover from this and \p\G\ \M_generic(G)\ + obtain q where "q\G" "q\p" "q\r" by force + moreover from this and \p\G\ \M_generic(G)\ + have "q\P" "p\P" by blast+ + moreover from calculation and \M_generic(G)\ + have "q forces\<^sub>a (\ \ \)" + using forces_memI by blast + moreover + note \p forces\<^sub>a (\ = \)\ + ultimately + have "q forces\<^sub>a (\ \ \)" + using def_forces_eq by blast + with \q\P\ \q\G\ IH[of q \] \\\,r\\\\ \val(P,G,\) = x\ + show "x\val(P,G,\)" by (blast) +qed + +(* Lemma IV.2.40(a), equality, second inclusion--- COPY-PASTE *) +lemma IV240a_eq_2nd_incl: + assumes + "M_generic(G)" "p\G" "p forces\<^sub>a (\ = \)" + and + IH:"\q \. q\P \ q\G \ \\domain(\) \ domain(\) \ + (q forces\<^sub>a (\ \ \) \ val(P,G,\) \ val(P,G,\)) \ + (q forces\<^sub>a (\ \ \) \ val(P,G,\) \ val(P,G,\))" + shows + "val(P,G,\) \ val(P,G,\)" +proof + fix x + assume "x\val(P,G,\)" + then + obtain \ r where "\\,r\\\" "r\G" "val(P,G,\)=x" by blast + moreover from this and \p\G\ \M_generic(G)\ + obtain q where "q\G" "q\p" "q\r" by force + moreover from this and \p\G\ \M_generic(G)\ + have "q\P" "p\P" by blast+ + moreover from calculation and \M_generic(G)\ + have "q forces\<^sub>a (\ \ \)" + using forces_memI by blast + moreover + note \p forces\<^sub>a (\ = \)\ + ultimately + have "q forces\<^sub>a (\ \ \)" + using def_forces_eq by blast + with \q\P\ \q\G\ IH[of q \] \\\,r\\\\ \val(P,G,\) = x\ + show "x\val(P,G,\)" by (blast) +qed + +(* Lemma IV.2.40(a), equality, second inclusion--- COPY-PASTE *) +lemma IV240a_eq: + assumes + "M_generic(G)" "p\G" "p forces\<^sub>a (\ = \)" + and + IH:"\q \. q\P \ q\G \ \\domain(\) \ domain(\) \ + (q forces\<^sub>a (\ \ \) \ val(P,G,\) \ val(P,G,\)) \ + (q forces\<^sub>a (\ \ \) \ val(P,G,\) \ val(P,G,\))" + shows + "val(P,G,\) = val(P,G,\)" + using IV240a_eq_1st_incl[OF assms] IV240a_eq_2nd_incl[OF assms] IH by blast + +subsection\Induction on names\ + +lemma core_induction: + assumes + "\\ \ p. p \ P \ \\q \. \q\P ; \\domain(\)\ \ Q(0,\,\,q)\ \ Q(1,\,\,p)" + "\\ \ p. p \ P \ \\q \. \q\P ; \\domain(\) \ domain(\)\ \ Q(1,\,\,q) \ Q(1,\,\,q)\ \ Q(0,\,\,p)" + "ft \ 2" "p \ P" + shows + "Q(ft,\,\,p)" +proof - + { + fix ft p \ \ + have "Transset(eclose({\,\}))" (is "Transset(?e)") + using Transset_eclose by simp + have "\ \ ?e" "\ \ ?e" + using arg_into_eclose by simp_all + moreover + assume "ft \ 2" "p \ P" + ultimately + have "\ft,\,\,p\\ 2\?e\?e\P" (is "?a\2\?e\?e\P") by simp + then + have "Q(ftype(?a), name1(?a), name2(?a), cond_of(?a))" + using core_induction_aux[of ?e P Q ?a,OF \Transset(?e)\ assms(1,2) \?a\_\] + by (clarify) (blast) + then have "Q(ft,\,\,p)" by (simp add:components_simp) + } + then show ?thesis using assms by simp +qed + +lemma forces_induction_with_conds: + assumes + "\\ \ p. p \ P \ \\q \. \q\P ; \\domain(\)\ \ Q(q,\,\)\ \ R(p,\,\)" + "\\ \ p. p \ P \ \\q \. \q\P ; \\domain(\) \ domain(\)\ \ R(q,\,\) \ R(q,\,\)\ \ Q(p,\,\)" + "p \ P" + shows + "Q(p,\,\) \ R(p,\,\)" +proof - + let ?Q="\ft \ \ p. (ft = 0 \ Q(p,\,\)) \ (ft = 1 \ R(p,\,\))" + from assms(1) + have "\\ \ p. p \ P \ \\q \. \q\P ; \\domain(\)\ \ ?Q(0,\,\,q)\ \ ?Q(1,\,\,p)" + by simp + moreover from assms(2) + have "\\ \ p. p \ P \ \\q \. \q\P ; \\domain(\) \ domain(\)\ \ ?Q(1,\,\,q) \ ?Q(1,\,\,q)\ \ ?Q(0,\,\,p)" + by simp + moreover + note \p\P\ + ultimately + have "?Q(ft,\,\,p)" if "ft\2" for ft + by (rule core_induction[OF _ _ that, of ?Q]) + then + show ?thesis by auto +qed + +lemma forces_induction: + assumes + "\\ \. \\\. \\domain(\) \ Q(\,\)\ \ R(\,\)" + "\\ \. \\\. \\domain(\) \ domain(\) \ R(\,\) \ R(\,\)\ \ Q(\,\)" + shows + "Q(\,\) \ R(\,\)" +proof (intro forces_induction_with_conds[OF _ _ one_in_P ]) + fix \ \ p + assume "q \ P \ \ \ domain(\) \ Q(\, \)" for q \ + with assms(1) + show "R(\,\)" + using one_in_P by simp +next + fix \ \ p + assume "q \ P \ \ \ domain(\) \ domain(\) \ R(\,\) \ R(\,\)" for q \ + with assms(2) + show "Q(\,\)" + using one_in_P by simp +qed + +subsection\Lemma IV.2.40(a), in full\ +lemma IV240a: + assumes + "M_generic(G)" + shows + "(\\M \ \\M \ (\p\G. p forces\<^sub>a (\ = \) \ val(P,G,\) = val(P,G,\))) \ + (\\M \ \\M \ (\p\G. p forces\<^sub>a (\ \ \) \ val(P,G,\) \ val(P,G,\)))" + (is "?Q(\,\) \ ?R(\,\)") +proof (intro forces_induction[of ?Q ?R] impI) + fix \ \ + assume "\\M" "\\M" "\\domain(\) \ ?Q(\,\)" for \ + moreover from this + have "\\domain(\) \ q forces\<^sub>a (\ = \) \ val(P,G, \) = val(P,G, \)" + if "q\P" "q\G" for q \ + using that domain_closed[of \] transitivity by auto + moreover + note assms + ultimately + show "\p\G. p forces\<^sub>a (\ \ \) \ val(P,G,\) \ val(P,G,\)" + using IV240a_mem domain_closed transitivity by (simp) +next + fix \ \ + assume "\\M" "\\M" "\ \ domain(\) \ domain(\) \ ?R(\,\) \ ?R(\,\)" for \ + moreover from this + have IH':"\ \ domain(\) \ domain(\) \ q\G \ + (q forces\<^sub>a (\ \ \) \ val(P,G, \) \ val(P,G, \)) \ + (q forces\<^sub>a (\ \ \) \ val(P,G, \) \ val(P,G, \))" for q \ + by (auto intro: transitivity[OF _ domain_closed[simplified]]) + ultimately + show "\p\G. p forces\<^sub>a (\ = \) \ val(P,G,\) = val(P,G,\)" + using IV240a_eq[OF assms(1) _ _ IH'] by (simp) +qed + +subsection\Lemma IV.2.40(b)\ + (* Lemma IV.2.40(b), membership *) +lemma IV240b_mem: + assumes + "M_generic(G)" "val(P,G,\)\val(P,G,\)" "\\M" "\\M" + and + IH:"\\. \\domain(\) \ val(P,G,\) = val(P,G,\) \ + \p\G. p forces\<^sub>a (\ = \)" (* inductive hypothesis *) + shows + "\p\G. p forces\<^sub>a (\ \ \)" +proof - + from \val(P,G,\)\val(P,G,\)\ + obtain \ r where "r\G" "\\,r\\\" "val(P,G,\) = val(P,G,\)" by auto + moreover from this and IH + obtain p' where "p'\G" "p' forces\<^sub>a (\ = \)" by blast + moreover + note \M_generic(G)\ + ultimately + obtain p where "p\r" "p\G" "p forces\<^sub>a (\ = \)" + using M_generic_compatD strengthening_eq[of p'] by blast + moreover + note \M_generic(G)\ + moreover from calculation + have "q forces\<^sub>a (\ = \)" if "q\P" "q\p" for q + using that strengthening_eq by blast + moreover + note \\\,r\\\\ \r\G\ + ultimately + have "r\P \ \\,r\ \ \ \ q\r \ q forces\<^sub>a (\ = \)" if "q\P" "q\p" for q + using that leq_transD[of _ p r] by blast + then + have "dense_below({q\P. \s r. r\P \ \s,r\ \ \ \ q\r \ q forces\<^sub>a (\ = s)},p)" + using refl_leq by blast + moreover + note \M_generic(G)\ \p\G\ + moreover from calculation + have "p forces\<^sub>a (\ \ \)" + using forces_mem_iff_dense_below by blast + ultimately + show ?thesis by blast +qed + +end \ \includes some\_rules\ + +lemma Collect_forces_eq_in_M: + assumes "\ \ M" "\ \ M" + shows "{p\P. p forces\<^sub>a (\ = \)} \ M" + using assms Collect_in_M[of "forces_eq_fm(1,2,0,3,4)" "[P,leq,\,\]"] + arity_forces_eq_fm P_in_M leq_in_M sats_forces_eq_fm forces_eq_abs forces_eq_fm_type + by (simp add: union_abs1 Un_commute) + +lemma IV240b_eq_Collects: + assumes "\ \ M" "\ \ M" + shows "{p\P. \\\domain(\) \ domain(\). p forces\<^sub>a (\ \ \) \ p forces\<^sub>a (\ \ \)}\M" and + "{p\P. \\\domain(\) \ domain(\). p forces\<^sub>a (\ \ \) \ p forces\<^sub>a (\ \ \)}\M" +proof - + let ?rel_pred="\M x a1 a2 a3 a4. + \\[M]. \u[M]. \da3[M]. \da4[M]. is_domain(M,a3,da3) \ is_domain(M,a4,da4) \ + union(M,da3,da4,u) \ \\u \ is_forces_mem'(M,a1,a2,x,\,a3) \ + is_forces_nmem'(M,a1,a2,x,\,a4)" + let ?\="Exists(Exists(Exists(Exists(And(domain_fm(7,1),And(domain_fm(8,0), + And(union_fm(1,0,2),And(Member(3,2),And(forces_mem_fm(5,6,4,3,7), + forces_nmem_fm(5,6,4,3,8))))))))))" + have 1:"\\M" if "\\,y\\\" "\\M" for \ \ y + using that pair_in_M_iff transitivity[of "\\,y\" \] by simp + have abs1:"?rel_pred(##M,p,P,leq,\,\) \ + (\\\domain(\) \ domain(\). forces_mem'(P,leq,p,\,\) \ forces_nmem'(P,leq,p,\,\))" + if "p\M" for p + unfolding forces_mem_def forces_nmem_def + using assms that forces_mem'_abs forces_nmem'_abs P_in_M leq_in_M + domain_closed Un_closed + by (auto simp add:1[of _ _ \] 1[of _ _ \]) + have abs2:"?rel_pred(##M,p,P,leq,\,\) \ (\\\domain(\) \ domain(\). + forces_nmem'(P,leq,p,\,\) \ forces_mem'(P,leq,p,\,\))" if "p\M" for p + unfolding forces_mem_def forces_nmem_def + using assms that forces_mem'_abs forces_nmem'_abs P_in_M leq_in_M + domain_closed Un_closed + by (auto simp add:1[of _ _ \] 1[of _ _ \]) + have fsats1:"(M,[p,P,leq,\,\] \ ?\) \ ?rel_pred(##M,p,P,leq,\,\)" if "p\M" for p + using that assms sats_forces_mem_fm sats_forces_nmem_fm P_in_M leq_in_M zero_in_M + domain_closed Un_closed by simp + have fsats2:"(M,[p,P,leq,\,\] \ ?\) \ ?rel_pred(##M,p,P,leq,\,\)" if "p\M" for p + using that assms sats_forces_mem_fm sats_forces_nmem_fm P_in_M leq_in_M zero_in_M + domain_closed Un_closed by simp + have fty:"?\\formula" by simp + have farit:"arity(?\)=5" + by (simp add:ord_simp_union arity) + show + "{p \ P . \\\domain(\) \ domain(\). p forces\<^sub>a (\ \ \) \ p forces\<^sub>a (\ \ \)} \ M" + and "{p \ P . \\\domain(\) \ domain(\). p forces\<^sub>a (\ \ \) \ p forces\<^sub>a (\ \ \)} \ M" + unfolding forces_mem_def + using abs1 fty fsats1 farit P_in_M leq_in_M assms forces_nmem + Collect_in_M[of ?\ "[P,leq,\,\]"] + using abs2 fty fsats2 farit P_in_M leq_in_M assms forces_nmem domain_closed Un_closed + Collect_in_M[of ?\ "[P,leq,\,\]"] + by simp_all +qed + +(* Lemma IV.2.40(b), equality *) +lemma IV240b_eq: + assumes + "M_generic(G)" "val(P,G,\) = val(P,G,\)" "\\M" "\\M" + and + IH:"\\. \\domain(\)\domain(\) \ + (val(P,G,\)\val(P,G,\) \ (\q\G. q forces\<^sub>a (\ \ \))) \ + (val(P,G,\)\val(P,G,\) \ (\q\G. q forces\<^sub>a (\ \ \)))" + (* inductive hypothesis *) + shows + "\p\G. p forces\<^sub>a (\ = \)" +proof - + let ?D1="{p\P. p forces\<^sub>a (\ = \)}" + let ?D2="{p\P. \\\domain(\) \ domain(\). p forces\<^sub>a (\ \ \) \ p forces\<^sub>a (\ \ \)}" + let ?D3="{p\P. \\\domain(\) \ domain(\). p forces\<^sub>a (\ \ \) \ p forces\<^sub>a (\ \ \)}" + let ?D="?D1 \ ?D2 \ ?D3" + note assms + moreover from this + have "domain(\) \ domain(\)\M" (is "?B\M") using domain_closed Un_closed by auto + moreover from calculation + have "?D2\M" and "?D3\M" using IV240b_eq_Collects by simp_all + ultimately + have "?D\M" using Collect_forces_eq_in_M Un_closed by auto + moreover + have "dense(?D)" + proof + fix p + assume "p\P" + have "\d\P. (d forces\<^sub>a (\ = \) \ + (\\\domain(\) \ domain(\). d forces\<^sub>a (\ \ \) \ d forces\<^sub>a (\ \ \)) \ + (\\\domain(\) \ domain(\). d forces\<^sub>a (\ \ \) \ d forces\<^sub>a (\ \ \))) \ + d \ p" + proof (cases "p forces\<^sub>a (\ = \)") + case True + with \p\P\ + show ?thesis using refl_leq by blast + next + case False + moreover note \p\P\ + moreover from calculation + obtain \ q where "\\domain(\)\domain(\)" "q\P" "q\p" + "(q forces\<^sub>a (\ \ \) \ \ q forces\<^sub>a (\ \ \)) \ + (\ q forces\<^sub>a (\ \ \) \ q forces\<^sub>a (\ \ \))" + using def_forces_eq by blast + moreover from this + obtain r where "r\q" "r\P" + "(r forces\<^sub>a (\ \ \) \ r forces\<^sub>a (\ \ \)) \ + (r forces\<^sub>a (\ \ \) \ r forces\<^sub>a (\ \ \))" + using not_forces_nmem strengthening_mem by blast + ultimately + show ?thesis using leq_transD by blast + qed + then + show "\d\?D1 \ ?D2 \ ?D3. d \ p" by blast + qed + moreover + have "?D \ P" + by auto + moreover + note \M_generic(G)\ + ultimately + obtain p where "p\G" "p\?D" + unfolding M_generic_def by blast + then + consider + (1) "p forces\<^sub>a (\ = \)" | + (2) "\\\domain(\) \ domain(\). p forces\<^sub>a (\ \ \) \ p forces\<^sub>a (\ \ \)" | + (3) "\\\domain(\) \ domain(\). p forces\<^sub>a (\ \ \) \ p forces\<^sub>a (\ \ \)" + by blast + then + show ?thesis + proof (cases) + case 1 + with \p\G\ + show ?thesis by blast + next + case 2 + then + obtain \ where "\\domain(\) \ domain(\)" "p forces\<^sub>a (\ \ \)" "p forces\<^sub>a (\ \ \)" + by blast + moreover from this and \p\G\ and assms + have "val(P,G,\)\val(P,G,\)" + using IV240a[of G \ \] transitivity[OF _ domain_closed[simplified]] by blast + moreover note IH \val(P,G,\) = _\ + ultimately + obtain q where "q\G" "q forces\<^sub>a (\ \ \)" by auto + moreover from this and \p\G\ \M_generic(G)\ + obtain r where "r\P" "r\p" "r\q" + by blast + moreover + note \M_generic(G)\ + ultimately + have "r forces\<^sub>a (\ \ \)" + using strengthening_mem by blast + with \r\p\ \p forces\<^sub>a (\ \ \)\ \r\P\ + have "False" + unfolding forces_nmem_def by blast + then + show ?thesis by simp + next (* copy-paste from case 2 mutatis mutandis*) + case 3 + then + obtain \ where "\\domain(\) \ domain(\)" "p forces\<^sub>a (\ \ \)" "p forces\<^sub>a (\ \ \)" + by blast + moreover from this and \p\G\ and assms + have "val(P,G,\)\val(P,G,\)" + using IV240a[of G \ \] transitivity[OF _ domain_closed[simplified]] by blast + moreover note IH \val(P,G,\) = _\ + ultimately + obtain q where "q\G" "q forces\<^sub>a (\ \ \)" by auto + moreover from this and \p\G\ \M_generic(G)\ + obtain r where "r\P" "r\p" "r\q" + by blast + moreover + note \M_generic(G)\ + ultimately + have "r forces\<^sub>a (\ \ \)" + using strengthening_mem by blast + with \r\p\ \p forces\<^sub>a (\ \ \)\ \r\P\ + have "False" + unfolding forces_nmem_def by blast + then + show ?thesis by simp + qed +qed + +(* Lemma IV.2.40(b), full *) +lemma IV240b: + assumes + "M_generic(G)" + shows + "(\\M\\\M\val(P,G,\) = val(P,G,\) \ (\p\G. p forces\<^sub>a (\ = \))) \ + (\\M\\\M\val(P,G,\) \ val(P,G,\) \ (\p\G. p forces\<^sub>a (\ \ \)))" + (is "?Q(\,\) \ ?R(\,\)") +proof (intro forces_induction) + fix \ \ p + assume "\\domain(\) \ ?Q(\, \)" for \ + with assms + show "?R(\, \)" + using IV240b_mem domain_closed transitivity by (simp) +next + fix \ \ p + assume "\ \ domain(\) \ domain(\) \ ?R(\,\) \ ?R(\,\)" for \ + moreover from this + have IH':"\\M \ \\M \ \ \ domain(\) \ domain(\) \ + (val(P,G, \) \ val(P,G, \) \ (\q\G. q forces\<^sub>a (\ \ \))) \ + (val(P,G, \) \ val(P,G, \) \ (\q\G. q forces\<^sub>a (\ \ \)))" for \ + using domain_trans[OF trans_M] + by (blast) + ultimately + show "?Q(\,\)" + using IV240b_eq[OF assms(1)] by (auto) +qed + +lemma map_val_in_MG: + assumes + "env\list(M)" + shows + "map(val(P,G),env)\list(M[G])" + unfolding GenExt_def using assms map_type2 by simp + +lemma truth_lemma_mem: + assumes + "env\list(M)" "M_generic(G)" + "n\nat" "m\nat" "np\G. p \ Member(n,m) env) \ M[G], map(val(P,G),env) \ Member(n,m)" + using assms IV240a[OF assms(2), of "nth(n,env)" "nth(m,env)"] + IV240b[OF assms(2), of "nth(n,env)" "nth(m,env)"] + P_in_M leq_in_M one_in_M + Forces_Member[of _ "nth(n,env)" "nth(m,env)" env n m] map_val_in_MG + by (auto) + +lemma truth_lemma_eq: + assumes + "env\list(M)" "M_generic(G)" + "n\nat" "m\nat" "np\G. p \ Equal(n,m) env) \ M[G], map(val(P,G),env) \ Equal(n,m)" + using assms IV240a(1)[OF assms(2), of "nth(n,env)" "nth(m,env)"] + IV240b(1)[OF assms(2), of "nth(n,env)" "nth(m,env)"] + P_in_M leq_in_M one_in_M + Forces_Equal[of _ "nth(n,env)" "nth(m,env)" env n m] map_val_in_MG + by (auto) + +lemma arities_at_aux: + assumes + "n \ nat" "m \ nat" "env \ list(M)" "succ(n) \ succ(m) \ length(env)" + shows + "n < length(env)" "m < length(env)" + using assms succ_leE[OF Un_leD1, of n "succ(m)" "length(env)"] + succ_leE[OF Un_leD2, of "succ(n)" m "length(env)"] by auto + +subsection\The Strenghtening Lemma\ + +lemma strengthening_lemma: + assumes + "p\P" "\\formula" "r\P" "r\p" + "env\list(M)" "arity(\)\length(env)" + shows + "p \ \ env \ r \ \ env" + using assms(2-) +proof (induct arbitrary:env) + case (Member n m) + then + have "nlist(M)" + moreover + note assms Member + ultimately + show ?case + using Forces_Member[of _ "nth(n,env)" "nth(m,env)" env n m] + strengthening_mem[of p r "nth(n,env)" "nth(m,env)"] by simp +next + case (Equal n m) + then + have "nlist(M)" + moreover + note assms Equal + ultimately + show ?case + using Forces_Equal[of _ "nth(n,env)" "nth(m,env)" env n m] + strengthening_eq[of p r "nth(n,env)" "nth(m,env)"] by simp +next + case (Nand \ \) + with assms + show ?case + using Forces_Nand transitivity[OF _ P_in_M] pair_in_M_iff + transitivity[OF _ leq_in_M] leq_transD by auto +next + case (Forall \) + with assms + have "p \ \ ([x] @ env)" if "x\M" for x + using that Forces_Forall by simp + with Forall + have "r \ \ ([x] @ env)" if "x\M" for x + using that pred_le2 by (simp) + with assms Forall + show ?case + using Forces_Forall by simp +qed + +subsection\The Density Lemma\ + +lemma arity_Nand_le: + assumes "\ \ formula" "\ \ formula" "arity(Nand(\, \)) \ length(env)" "env\list(A)" + shows "arity(\) \ length(env)" "arity(\) \ length(env)" + using assms + by (rule_tac Un_leD1, rule_tac [5] Un_leD2, auto) + +lemma dense_below_imp_forces: + assumes + "p\P" "\\formula" + "env\list(M)" "arity(\)\length(env)" + shows + "dense_below({q\P. (q \ \ env)},p) \ (p \ \ env)" + using assms(2-) +proof (induct arbitrary:env) + case (Member n m) + then + have "nlist(M)" + moreover + note assms Member + ultimately + show ?case + using Forces_Member[of _ "nth(n,env)" "nth(m,env)" env n m] + density_mem[of p "nth(n,env)" "nth(m,env)"] by simp +next + case (Equal n m) + then + have "nlist(M)" + moreover + note assms Equal + ultimately + show ?case + using Forces_Equal[of _ "nth(n,env)" "nth(m,env)" env n m] + density_eq[of p "nth(n,env)" "nth(m,env)"] by simp +next + case (Nand \ \) + { + fix q + assume "q\M" "q\P" "q\ p" "q \ \ env" + moreover + note Nand + moreover from calculation + obtain d where "d\P" "d \ Nand(\, \) env" "d\ q" + using dense_belowI by auto + moreover from calculation + have "\(d\ \ env)" if "d \ \ env" + using that Forces_Nand refl_leq transitivity[OF _ P_in_M, of d] by auto + moreover + note arity_Nand_le[of \ \] + moreover from calculation + have "d \ \ env" + using strengthening_lemma[of q \ d env] Un_leD1 by auto + ultimately + have "\ (q \ \ env)" + using strengthening_lemma[of q \ d env] by auto + } + with \p\P\ + show ?case + using Forces_Nand[symmetric, OF _ Nand(6,1,3)] by blast +next + case (Forall \) + have "dense_below({q\P. q \ \ ([a]@env)},p)" if "a\M" for a + proof + fix r + assume "r\P" "r\p" + with \dense_below(_,p)\ + obtain q where "q\P" "q\r" "q \ Forall(\) env" + by blast + moreover + note Forall \a\M\ + moreover from calculation + have "q \ \ ([a]@env)" + using Forces_Forall by simp + ultimately + show "\d \ {q\P. q \ \ ([a]@env)}. d \ P \ d\r" + by auto + qed + moreover + note Forall(2)[of "Cons(_,env)"] Forall(1,3-5) + ultimately + have "p \ \ ([a]@env)" if "a\M" for a + using that pred_le2 by simp + with assms Forall + show ?case using Forces_Forall by simp +qed + +lemma density_lemma: + assumes + "p\P" "\\formula" "env\list(M)" "arity(\)\length(env)" + shows + "p \ \ env \ dense_below({q\P. (q \ \ env)},p)" +proof + assume "dense_below({q\P. (q \ \ env)},p)" + with assms + show "(p \ \ env)" + using dense_below_imp_forces by simp +next + assume "p \ \ env" + with assms + show "dense_below({q\P. q \ \ env},p)" + using strengthening_lemma refl_leq by auto +qed + +subsection\The Truth Lemma\ + +lemma Forces_And: + assumes + "p\P" "env \ list(M)" "\\formula" "\\formula" + "arity(\) \ length(env)" "arity(\) \ length(env)" + shows + "p \ And(\,\) env \ (p \ \ env) \ (p \ \ env)" +proof + assume "p \ And(\, \) env" + with assms + have "dense_below({r \ P . (r \ \ env) \ (r \ \ env)}, p)" + using Forces_And_iff_dense_below by simp + then + have "dense_below({r \ P . (r \ \ env)}, p)" "dense_below({r \ P . (r \ \ env)}, p)" + by blast+ + with assms + show "(p \ \ env) \ (p \ \ env)" + using density_lemma[symmetric] by simp +next + assume "(p \ \ env) \ (p \ \ env)" + have "dense_below({r \ P . (r \ \ env) \ (r \ \ env)}, p)" + proof (intro dense_belowI bexI conjI, assumption) + fix q + assume "q\P" "q\ p" + with assms \(p \ \ env) \ (p \ \ env)\ + show "q\{r \ P . (r \ \ env) \ (r \ \ env)}" "q\ q" + using strengthening_lemma refl_leq by auto + qed + with assms + show "p \ And(\,\) env" + using Forces_And_iff_dense_below by simp +qed + +lemma Forces_Nand_alt: + assumes + "p\P" "env \ list(M)" "\\formula" "\\formula" + "arity(\) \ length(env)" "arity(\) \ length(env)" + shows + "(p \ Nand(\,\) env) \ (p \ Neg(And(\,\)) env)" + using assms Forces_Nand Forces_And Forces_Neg by auto + +lemma truth_lemma_Neg: + assumes + "\\formula" "M_generic(G)" "env\list(M)" "arity(\)\length(env)" and + IH: "(\p\G. p \ \ env) \ M[G], map(val(P,G),env) \ \" + shows + "(\p\G. p \ Neg(\) env) \ M[G], map(val(P,G),env) \ Neg(\)" +proof (intro iffI, elim bexE, rule ccontr) + (* Direct implication by contradiction *) + fix p + assume "p\G" "p \ Neg(\) env" "\(M[G],map(val(P,G),env) \ Neg(\))" + moreover + note assms + moreover from calculation + have "M[G], map(val(P,G),env) \ \" + using map_val_in_MG by simp + with IH + obtain r where "r \ \ env" "r\G" by blast + moreover from this and \M_generic(G)\ \p\G\ + obtain q where "q\p" "q\r" "q\G" + by blast + moreover from calculation + have "q \ \ env" + using strengthening_lemma[where \=\] by blast + ultimately + show "False" + using Forces_Neg[where \=\] transitivity[OF _ P_in_M] by blast +next + assume "M[G], map(val(P,G),env) \ Neg(\)" + with assms + have "\ (M[G], map(val(P,G),env) \ \)" + using map_val_in_MG by simp + let ?D="{p\P. (p \ \ env) \ (p \ Neg(\) env)}" + have "separation(##M,\p. (p \ \ env))" + using separation_ax[of "forces(\)"] arity_forces assms P_in_M leq_in_M one_in_M arity_forces_le + by simp + moreover + have "separation(##M,\p. (p \ \\\\ env))" + using separation_ax[of "forces( \\\\ )"] arity_forces assms P_in_M leq_in_M one_in_M arity_forces_le + by simp + ultimately + have "separation(##M,\p. (p \ \ env) \ (p \ Neg(\) env))" + using separation_disj by simp + then + have "?D \ M" + using separation_closed P_in_M by simp + moreover + have "?D \ P" by auto + moreover + have "dense(?D)" + proof + fix q + assume "q\P" + show "\d\{p \ P . (p \ \ env) \ (p \ Neg(\) env)}. d\ q" + proof (cases "q \ Neg(\) env") + case True + with \q\P\ + show ?thesis using refl_leq by blast + next + case False + with \q\P\ and assms + show ?thesis using Forces_Neg by auto + qed + qed + moreover + note \M_generic(G)\ + ultimately + obtain p where "p\G" "(p \ \ env) \ (p \ Neg(\) env)" + by blast + then + consider (1) "p \ \ env" | (2) "p \ Neg(\) env" by blast + then + show "\p\G. (p \ Neg(\) env)" + proof (cases) + case 1 + with \\ (M[G],map(val(P,G),env) \ \)\ \p\G\ IH + show ?thesis + by blast + next + case 2 + with \p\G\ + show ?thesis by blast + qed +qed + +lemma truth_lemma_And: + assumes + "env\list(M)" "\\formula" "\\formula" + "arity(\)\length(env)" "arity(\) \ length(env)" "M_generic(G)" + and + IH: "(\p\G. p \ \ env) \ M[G], map(val(P,G),env) \ \" + "(\p\G. p \ \ env) \ M[G], map(val(P,G),env) \ \" + shows + "(\p\G. (p \ And(\,\) env)) \ M[G] , map(val(P,G),env) \ And(\,\)" + using assms map_val_in_MG Forces_And[OF M_genericD assms(1-5)] +proof (intro iffI, elim bexE) + fix p + assume "p\G" "p \ And(\,\) env" + with assms + show "M[G], map(val(P,G),env) \ And(\,\)" + using Forces_And[OF M_genericD, of _ _ _ \ \] map_val_in_MG by auto +next + assume "M[G], map(val(P,G),env) \ And(\,\)" + moreover + note assms + moreover from calculation + obtain q r where "q \ \ env" "r \ \ env" "q\G" "r\G" + using map_val_in_MG Forces_And[OF M_genericD assms(1-5)] by auto + moreover from calculation + obtain p where "p\q" "p\r" "p\G" + by blast + moreover from calculation + have "(p \ \ env) \ (p \ \ env)" (* can't solve as separate goals *) + using strengthening_lemma by (blast) + ultimately + show "\p\G. (p \ And(\,\) env)" + using Forces_And[OF M_genericD assms(1-5)] by auto +qed + +definition + ren_truth_lemma :: "i\i" where + "ren_truth_lemma(\) \ + Exists(Exists(Exists(Exists(Exists( + And(Equal(0,5),And(Equal(1,8),And(Equal(2,9),And(Equal(3,10),And(Equal(4,6), + iterates(\p. incr_bv(p)`5 , 6, \)))))))))))" + +lemma ren_truth_lemma_type[TC] : + "\\formula \ ren_truth_lemma(\) \formula" + unfolding ren_truth_lemma_def + by simp + +lemma arity_ren_truth : + assumes "\\formula" + shows "arity(ren_truth_lemma(\)) \ 6 \ succ(arity(\))" +proof - + consider (lt) "5 )" | (ge) "\ 5 < arity(\)" + by auto + then + show ?thesis + proof cases + case lt + consider (a) "5)+\<^sub>\5" | (b) "arity(\)+\<^sub>\5 \ 5" + using not_lt_iff_le \\\_\ by force + then + show ?thesis + proof cases + case a + with \\\_\ lt + have "5 < succ(arity(\))" "5)+\<^sub>\2" "5)+\<^sub>\3" "5)+\<^sub>\4" + using succ_ltI by auto + with \\\_\ + have c:"arity(iterates(\p. incr_bv(p)`5,5,\)) = 5+\<^sub>\arity(\)" (is "arity(?\') = _") + using arity_incr_bv_lemma lt a + by simp + with \\\_\ + have "arity(incr_bv(?\')`5) = 6+\<^sub>\arity(\)" + using arity_incr_bv_lemma[of ?\' 5] a by auto + with \\\_\ + show ?thesis + unfolding ren_truth_lemma_def + using pred_Un_distrib union_abs1 Un_assoc[symmetric] a c union_abs2 + by (simp add:arity) + next + case b + with \\\_\ lt + have "5 < succ(arity(\))" "5)+\<^sub>\2" "5)+\<^sub>\3" "5)+\<^sub>\4" "5)+\<^sub>\5" + using succ_ltI by auto + with \\\_\ + have "arity(iterates(\p. incr_bv(p)`5,6,\)) = 6+\<^sub>\arity(\)" (is "arity(?\') = _") + using arity_incr_bv_lemma lt + by simp + with \\\_\ + show ?thesis + unfolding ren_truth_lemma_def + using pred_Un_distrib union_abs1 Un_assoc[symmetric] union_abs2 + by (simp add:arity) + qed + next + case ge + with \\\_\ + have "arity(\) \ 5" "pred^5(arity(\)) \ 5" + using not_lt_iff_le le_trans[OF le_pred] + by auto + with \\\_\ + have "arity(iterates(\p. incr_bv(p)`5,6,\)) = arity(\)" "arity(\)\6" "pred^5(arity(\)) \ 6" + using arity_incr_bv_lemma ge le_trans[OF \arity(\)\5\] le_trans[OF \pred^5(arity(\))\5\] + by auto + with \arity(\) \ 5\ \\\_\ \pred^5(_) \ 5\ + show ?thesis + unfolding ren_truth_lemma_def + using pred_Un_distrib union_abs1 Un_assoc[symmetric] union_abs2 + by (simp add:arity) + qed +qed + +lemma sats_ren_truth_lemma: + "[q,b,d,a1,a2,a3] @ env \ list(M) \ \ \ formula \ + (M, [q,b,d,a1,a2,a3] @ env \ ren_truth_lemma(\) ) \ + (M, [q,a1,a2,a3,b] @ env \ \)" + unfolding ren_truth_lemma_def + by (insert sats_incr_bv_iff [of _ _ M _ "[q,a1,a2,a3,b]"], simp) + +lemma truth_lemma' : + assumes + "\\formula" "env\list(M)" "arity(\) \ succ(length(env))" + shows + "separation(##M,\d. \b\M. \q\P. q\d \ \(q \ \ ([b]@env)))" +proof - + let ?rel_pred="\M x a1 a2 a3. \b\M. \q\M. q\a1 \ is_leq(##M,a2,q,x) \ + \(M, [q,a1,a2,a3,b] @ env \ forces(\))" + let ?\="Exists(Forall(Implies(And(Member(0,3),is_leq_fm(4,0,2)), + Neg(ren_truth_lemma(forces(\))))))" + have "q\M" if "q\P" for q using that transitivity[OF _ P_in_M] by simp + then + have 1:"\q\M. q\P \ R(q) \ Q(q) \ (\q\P. R(q) \ Q(q))" for R Q + by auto + then + have "\b \ M; \q\M. q \ P \ q \ d \ \(q \ \ ([b]@env))\ \ + \c\M. \q\P. q \ d \ \(q \ \ ([c]@env))" for b d + by (rule bexI,simp_all) + then + have "?rel_pred(M,d,P,leq,\) \ (\b\M. \q\P. q\d \ \(q \ \ ([b]@env)))" if "d\M" for d + using that leq_abs leq_in_M P_in_M one_in_M assms + by auto + moreover + have "?\\formula" using assms by simp + moreover + have "(M, [d,P,leq,\]@env \ ?\) \ ?rel_pred(M,d,P,leq,\)" if "d\M" for d + using assms that P_in_M leq_in_M one_in_M sats_is_leq_fm sats_ren_truth_lemma zero_in_M + by simp + moreover + have "arity(?\) \ 4+\<^sub>\length(env)" + proof - + have eq:"arity(is_leq_fm(4, 0, 2)) = 5" + using arity_is_leq_fm succ_Un_distrib ord_simp_union + by simp + with \\\_\ + have "arity(?\) = 3 \ (pred^2(arity(ren_truth_lemma(forces(\)))))" + using union_abs1 pred_Un_distrib by (simp add:arity) + moreover + have "... \ 3 \ (pred(pred(6 \ succ(arity(forces(\))))))" (is "_ \ ?r") + using \\\_\ Un_le_compat[OF le_refl[of 3]] + le_imp_subset arity_ren_truth[of "forces(\)"] + pred_mono + by auto + finally + have "arity(?\) \ ?r" by simp + have i:"?r \ 4 \ pred(arity(forces(\)))" + using pred_Un_distrib pred_succ_eq \\\_\ Un_assoc[symmetric] union_abs1 by simp + have h:"4 \ pred(arity(forces(\))) \ 4 \ (4+\<^sub>\length(env))" + using \env\_\ add_commute \\\_\ + Un_le_compat[of 4 4,OF _ pred_mono[OF _ arity_forces_le[OF _ _ \arity(\)\_\]] ] + \env\_\ by auto + with \\\_\ \env\_\ + show ?thesis + using le_trans[OF \arity(?\) \ ?r\ le_trans[OF i h]] ord_simp_union by simp + qed + ultimately + show ?thesis using assms P_in_M leq_in_M one_in_M + separation_ax[of "?\" "[P,leq,\]@env"] + separation_cong[of "##M" "\y. (M, [y,P,leq,\]@env \?\)"] + by simp +qed + + +lemma truth_lemma: + assumes + "\\formula" "M_generic(G)" + "env\list(M)" "arity(\)\length(env)" + shows + "(\p\G. p \ \ env) \ M[G], map(val(P,G),env) \ \" + using assms +proof (induct arbitrary:env) + case (Member x y) + then + show ?case + using assms truth_lemma_mem[OF \env\list(M)\ assms(2) \x\nat\ \y\nat\] + arities_at_aux by simp +next + case (Equal x y) + then + show ?case + using assms truth_lemma_eq[OF \env\list(M)\ assms(2) \x\nat\ \y\nat\] + arities_at_aux by simp +next + case (Nand \ \) + moreover + note \M_generic(G)\ + ultimately + show ?case + using truth_lemma_And truth_lemma_Neg[of "\\ \ \\"] Forces_Nand_alt + M_genericD map_val_in_MG arity_Nand_le[of \ \] FOL_arities by auto +next + case (Forall \) + with \M_generic(G)\ + show ?case + proof (intro iffI) + assume "\p\G. (p \ Forall(\) env)" + with \M_generic(G)\ + obtain p where "p\G" "p\M" "p\P" "p \ Forall(\) env" + using transitivity[OF _ P_in_M] by auto + with \env\list(M)\ \\\formula\ + have "p \ \ ([x]@env)" if "x\M" for x + using that Forces_Forall by simp + with \p\G\ \\\formula\ \env\_\ \arity(Forall(\)) \ length(env)\ + Forall(2)[of "Cons(_,env)"] \M_generic(G)\ + show "M[G], map(val(P,G),env) \ Forall(\)" + using pred_le2 map_val_in_MG + by (auto iff:GenExt_iff) + next + assume "M[G], map(val(P,G),env) \ Forall(\)" + let ?D1="{d\P. (d \ Forall(\) env)}" + let ?D2="{d\P. \b\M. \q\P. q\d \ \(q \ \ ([b]@env))}" + define D where "D \ ?D1 \ ?D2" + have ar\:"arity(\)\succ(length(env))" + using assms \arity(Forall(\)) \ length(env)\ \\\formula\ \env\list(M)\ pred_le2 + by simp + then + have "arity(Forall(\)) \ length(env)" + using pred_le \\\formula\ \env\list(M)\ by simp + then + have "?D1\M" using Collect_forces ar\ \\\formula\ \env\list(M)\ by simp + moreover from \env\list(M)\ \\\formula\ + have "?D2\M" + using truth_lemma'[of \] separation_closed ar\ P_in_M + by simp + ultimately + have "D\M" unfolding D_def using Un_closed by simp + moreover + have "D \ P" unfolding D_def by auto + moreover + have "dense(D)" + proof + fix p + assume "p\P" + show "\d\D. d\ p" + proof (cases "p \ Forall(\) env") + case True + with \p\P\ + show ?thesis unfolding D_def using refl_leq by blast + next + case False + with Forall \p\P\ + obtain b where "b\M" "\(p \ \ ([b]@env))" + using Forces_Forall by blast + moreover from this \p\P\ Forall + have "\dense_below({q\P. q \ \ ([b]@env)},p)" + using density_lemma pred_le2 by auto + moreover from this + obtain d where "d\p" "\q\P. q\d \ \(q \ \ ([b] @ env))" + "d\P" by blast + ultimately + show ?thesis unfolding D_def by auto + qed + qed + moreover + note \M_generic(G)\ + ultimately + obtain d where "d \ D" "d \ G" by blast + then + consider (1) "d\?D1" | (2) "d\?D2" unfolding D_def by blast + then + show "\p\G. (p \ Forall(\) env)" + proof (cases) + case 1 + with \d\G\ + show ?thesis by blast + next + case 2 + then + obtain b where "b\M" "\q\P. q\d \\(q \ \ ([b] @ env))" + by blast + moreover from this(1) and \M[G], _ \ Forall(\)\ and + Forall(2)[of "Cons(b,env)"] Forall(1,3-5) \M_generic(G)\ + obtain p where "p\G" "p\P" "p \ \ ([b] @ env)" + using pred_le2 using map_val_in_MG by (auto iff:GenExt_iff) + moreover + note \d\G\ \M_generic(G)\ + ultimately + obtain q where "q\G" "q\P" "q\d" "q\p" by blast + moreover from this and \p \ \ ([b] @ env)\ + Forall \b\M\ \p\P\ + have "q \ \ ([b] @ env)" + using pred_le2 strengthening_lemma by simp + moreover + note \\q\P. q\d \\(q \ \ ([b] @ env))\ + ultimately + show ?thesis by simp + qed + qed +qed + +subsection\The ``Definition of forcing''\ +lemma definition_of_forcing: + assumes + "p\P" "\\formula" "env\list(M)" "arity(\)\length(env)" + shows + "(p \ \ env) \ + (\G. M_generic(G) \ p\G \ M[G], map(val(P,G),env) \ \)" +proof (intro iffI allI impI, elim conjE) + fix G + assume "(p \ \ env)" "M_generic(G)" "p \ G" + with assms + show "M[G], map(val(P,G),env) \ \" + using truth_lemma[of \] by blast +next + assume 1: "\G.(M_generic(G)\ p\G)\ M[G] , map(val(P,G),env) \ \" + { + fix r + assume 2: "r\P" "r\p" + then + obtain G where "r\G" "M_generic(G)" + text\Here we're using countability (via the existence of + generic filters) of \<^term>\M\ as a shortcut.\ + using generic_filter_existence by auto + moreover from calculation 2 \p\P\ + have "p\G" + unfolding M_generic_def using filter_leqD by simp + moreover note 1 + ultimately + have "M[G], map(val(P,G),env) \ \" + by simp + with assms \M_generic(G)\ + obtain s where "s\G" "(s \ \ env)" + using truth_lemma[of \] by blast + moreover from this and \M_generic(G)\ \r\G\ + obtain q where "q\G" "q\s" "q\r" + by blast + moreover from calculation \s\G\ \M_generic(G)\ + have "s\P" "q\P" + unfolding M_generic_def filter_def by auto + moreover + note assms + ultimately + have "\q\P. q\r \ (q \ \ env)" + using strengthening_lemma by blast + } + then + have "dense_below({q\P. (q \ \ env)},p)" + unfolding dense_below_def by blast + with assms + show "(p \ \ env)" + using density_lemma by blast +qed + +lemmas definability = forces_type + +end \ \\<^locale>\forcing_data1\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Foundation_Axiom.thy b/thys/Independence_CH/Foundation_Axiom.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Foundation_Axiom.thy @@ -0,0 +1,40 @@ +section\The Axiom of Foundation in $M[G]$\ + +theory Foundation_Axiom + imports + Names +begin + +context forcing_data1 +begin + +(* Slick proof essentially by Paulson (adapted from L) *) +lemma foundation_in_MG : "foundation_ax(##(M[G]))" + unfolding foundation_ax_def + by (rule rallI, cut_tac A=x in foundation, auto intro: transitivity_MG) + +(* Same theorem as above, declarative proof, + without using transitivity *) +lemma "foundation_ax(##(M[G]))" +proof - + { + fix x + assume "x\M[G]" "\y\M[G] . y\x" + then + have "\y\M[G] . y\x\M[G]" + by simp + then + obtain y where "y\x\M[G]" "\z\y. z \ x\M[G]" + using foundation[of "x\M[G]"] by blast + then + have "\y\M[G] . y \ x \ (\z\M[G] . z \ x \ z \ y)" + by auto + } + then + show ?thesis + unfolding foundation_ax_def by auto +qed + +end \ \\<^locale>\forcing_data1\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/FrecR.thy b/thys/Independence_CH/FrecR.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/FrecR.thy @@ -0,0 +1,640 @@ +section\Well-founded relation on names\ +theory FrecR + imports + Transitive_Models.Discipline_Function + Edrel +begin + +text\\<^term>\frecR\ is the well-founded relation on names that allows +us to define forcing for atomic formulas.\ + +definition + ftype :: "i\i" where + "ftype \ fst" + +definition + name1 :: "i\i" where + "name1(x) \ fst(snd(x))" + +definition + name2 :: "i\i" where + "name2(x) \ fst(snd(snd(x)))" + +definition + cond_of :: "i\i" where + "cond_of(x) \ snd(snd(snd((x))))" + +lemma components_simp: + "ftype(\f,n1,n2,c\) = f" + "name1(\f,n1,n2,c\) = n1" + "name2(\f,n1,n2,c\) = n2" + "cond_of(\f,n1,n2,c\) = c" + unfolding ftype_def name1_def name2_def cond_of_def + by simp_all + +definition eclose_n :: "[i\i,i] \ i" where + "eclose_n(name,x) = eclose({name(x)})" + +definition + ecloseN :: "i \ i" where + "ecloseN(x) = eclose_n(name1,x) \ eclose_n(name2,x)" + +lemma components_in_eclose : + "n1 \ ecloseN(\f,n1,n2,c\)" + "n2 \ ecloseN(\f,n1,n2,c\)" + unfolding ecloseN_def eclose_n_def + using components_simp arg_into_eclose by auto + +lemmas names_simp = components_simp(2) components_simp(3) + +lemma ecloseNI1 : + assumes "x \ eclose(n1) \ x\eclose(n2)" + shows "x \ ecloseN(\f,n1,n2,c\)" + unfolding ecloseN_def eclose_n_def + using assms eclose_sing names_simp + by auto + +lemmas ecloseNI = ecloseNI1 + +lemma ecloseN_mono : + assumes "u \ ecloseN(x)" "name1(x) \ ecloseN(y)" "name2(x) \ ecloseN(y)" + shows "u \ ecloseN(y)" +proof - + from \u\_\ + consider (a) "u\eclose({name1(x)})" | (b) "u \ eclose({name2(x)})" + unfolding ecloseN_def eclose_n_def by auto + then + show ?thesis + proof cases + case a + with \name1(x) \ _\ + show ?thesis + unfolding ecloseN_def eclose_n_def + using eclose_singE[OF a] mem_eclose_trans[of u "name1(x)" ] by auto + next + case b + with \name2(x) \ _\ + show ?thesis + unfolding ecloseN_def eclose_n_def + using eclose_singE[OF b] mem_eclose_trans[of u "name2(x)"] by auto + qed +qed + +definition + is_ftype :: "(i\o)\i\i\o" where + "is_ftype \ is_fst" + +definition + ftype_fm :: "[i,i] \ i" where + "ftype_fm \ fst_fm" + +lemma is_ftype_iff_sats [iff_sats]: + assumes + "nth(a,env) = x" "nth(b,env) = y" "a\nat" "b\nat" "env \ list(A)" + shows + "is_ftype(##A,x,y) \ sats(A,ftype_fm(a,b), env)" + unfolding ftype_fm_def is_ftype_def + using assms sats_fst_fm + by simp + +definition + is_name1 :: "(i\o)\i\i\o" where + "is_name1(M,x,t2) \ is_hcomp(M,is_fst(M),is_snd(M),x,t2)" + +definition + name1_fm :: "[i,i] \ i" where + "name1_fm(x,t) \ hcomp_fm(fst_fm,snd_fm,x,t)" + +lemma sats_name1_fm [simp]: + "\ x \ nat; y \ nat;env \ list(A) \ \ + (A, env \ name1_fm(x,y)) \ is_name1(##A, nth(x,env), nth(y,env))" + unfolding name1_fm_def is_name1_def + using sats_fst_fm sats_snd_fm sats_hcomp_fm[of A "is_fst(##A)" _ fst_fm "is_snd(##A)"] + by simp + +lemma is_name1_iff_sats [iff_sats]: + assumes + "nth(a,env) = x" "nth(b,env) = y" "a\nat" "b\nat" "env \ list(A)" + shows + "is_name1(##A,x,y) \ A , env \ name1_fm(a,b)" + using assms sats_name1_fm + by simp + +definition + is_snd_snd :: "(i\o)\i\i\o" where + "is_snd_snd(M,x,t) \ is_hcomp(M,is_snd(M),is_snd(M),x,t)" + +definition + snd_snd_fm :: "[i,i]\i" where + "snd_snd_fm(x,t) \ hcomp_fm(snd_fm,snd_fm,x,t)" + +lemma sats_snd2_fm [simp]: + "\ x \ nat; y \ nat;env \ list(A) \ \ + (A, env \ snd_snd_fm(x,y)) \ is_snd_snd(##A, nth(x,env), nth(y,env))" + unfolding snd_snd_fm_def is_snd_snd_def + using sats_snd_fm sats_hcomp_fm[of A "is_snd(##A)" _ snd_fm "is_snd(##A)"] + by simp + +definition + is_name2 :: "(i\o)\i\i\o" where + "is_name2(M,x,t3) \ is_hcomp(M,is_fst(M),is_snd_snd(M),x,t3)" + +definition + name2_fm :: "[i,i] \ i" where + "name2_fm(x,t3) \ hcomp_fm(fst_fm,snd_snd_fm,x,t3)" + +lemma sats_name2_fm : + "\ x \ nat; y \ nat;env \ list(A) \ + \ (A , env \ name2_fm(x,y)) \ is_name2(##A, nth(x,env), nth(y,env))" + unfolding name2_fm_def is_name2_def + using sats_fst_fm sats_snd2_fm sats_hcomp_fm[of A "is_fst(##A)" _ fst_fm "is_snd_snd(##A)"] + by simp + +lemma is_name2_iff_sats [iff_sats]: + assumes + "nth(a,env) = x" "nth(b,env) = y" "a\nat" "b\nat" "env \ list(A)" + shows + "is_name2(##A,x,y) \ A, env \ name2_fm(a,b)" + using assms sats_name2_fm + by simp + +definition + is_cond_of :: "(i\o)\i\i\o" where + "is_cond_of(M,x,t4) \ is_hcomp(M,is_snd(M),is_snd_snd(M),x,t4)" + +definition + cond_of_fm :: "[i,i] \ i" where + "cond_of_fm(x,t4) \ hcomp_fm(snd_fm,snd_snd_fm,x,t4)" + +lemma sats_cond_of_fm : + "\ x \ nat; y \ nat;env \ list(A) \ \ + (A, env \ cond_of_fm(x,y)) \ is_cond_of(##A, nth(x,env), nth(y,env))" + unfolding cond_of_fm_def is_cond_of_def + using sats_snd_fm sats_snd2_fm sats_hcomp_fm[of A "is_snd(##A)" _ snd_fm "is_snd_snd(##A)"] + by simp + +lemma is_cond_of_iff_sats [iff_sats]: + assumes + "nth(a,env) = x" "nth(b,env) = y" "a\nat" "b\nat" "env \ list(A)" + shows + "is_cond_of(##A,x,y) \ A, env \ cond_of_fm(a,b)" + using assms sats_cond_of_fm + by simp + +lemma components_type[TC] : + assumes "a\nat" "b\nat" + shows + "ftype_fm(a,b)\formula" + "name1_fm(a,b)\formula" + "name2_fm(a,b)\formula" + "cond_of_fm(a,b)\formula" + using assms + unfolding ftype_fm_def fst_fm_def snd_fm_def snd_snd_fm_def name1_fm_def name2_fm_def + cond_of_fm_def hcomp_fm_def + by simp_all + +lemmas components_iff_sats = is_ftype_iff_sats is_name1_iff_sats is_name2_iff_sats + is_cond_of_iff_sats + +lemmas components_defs = ftype_fm_def snd_snd_fm_def hcomp_fm_def + name1_fm_def name2_fm_def cond_of_fm_def + +definition + is_eclose_n :: "[i\o,[i\o,i,i]\o,i,i] \ o" where + "is_eclose_n(N,is_name,en,t) \ + \n1[N].\s1[N]. is_name(N,t,n1) \ is_singleton(N,n1,s1) \ is_eclose(N,s1,en)" + +definition + eclose_n1_fm :: "[i,i] \ i" where + "eclose_n1_fm(m,t) \ Exists(Exists(And(And(name1_fm(t+\<^sub>\2,0),singleton_fm(0,1)), + is_eclose_fm(1,m+\<^sub>\2))))" + +definition + eclose_n2_fm :: "[i,i] \ i" where + "eclose_n2_fm(m,t) \ Exists(Exists(And(And(name2_fm(t+\<^sub>\2,0),singleton_fm(0,1)), + is_eclose_fm(1,m+\<^sub>\2))))" + +definition + is_ecloseN :: "[i\o,i,i] \ o" where + "is_ecloseN(N,t,en) \ \en1[N].\en2[N]. + is_eclose_n(N,is_name1,en1,t) \ is_eclose_n(N,is_name2,en2,t)\ + union(N,en1,en2,en)" + +definition + ecloseN_fm :: "[i,i] \ i" where + "ecloseN_fm(en,t) \ Exists(Exists(And(eclose_n1_fm(1,t+\<^sub>\2), + And(eclose_n2_fm(0,t+\<^sub>\2),union_fm(1,0,en+\<^sub>\2)))))" + +lemma ecloseN_fm_type [TC] : + "\ en \ nat ; t \ nat \ \ ecloseN_fm(en,t) \ formula" + unfolding ecloseN_fm_def eclose_n1_fm_def eclose_n2_fm_def by simp + +lemma sats_ecloseN_fm [simp]: + "\ en \ nat; t \ nat ; env \ list(A) \ + \ (A, env \ ecloseN_fm(en,t)) \ is_ecloseN(##A,nth(t,env),nth(en,env))" + unfolding ecloseN_fm_def is_ecloseN_def eclose_n1_fm_def eclose_n2_fm_def is_eclose_n_def + using nth_0 nth_ConsI sats_name1_fm sats_name2_fm singleton_iff_sats[symmetric] + by auto + +lemma is_ecloseN_iff_sats [iff_sats]: + "\ nth(en, env) = ena; nth(t, env) = ta; en \ nat; t \ nat ; env \ list(A) \ + \ is_ecloseN(##A,ta,ena) \ A, env \ ecloseN_fm(en,t)" + by simp + +(* Relation of forces *) +definition + frecR :: "i \ i \ o" where + "frecR(x,y) \ + (ftype(x) = 1 \ ftype(y) = 0 + \ (name1(x) \ domain(name1(y)) \ domain(name2(y)) \ (name2(x) = name1(y) \ name2(x) = name2(y)))) + \ (ftype(x) = 0 \ ftype(y) = 1 \ name1(x) = name1(y) \ name2(x) \ domain(name2(y)))" + +lemma frecR_ftypeD : + assumes "frecR(x,y)" + shows "(ftype(x) = 0 \ ftype(y) = 1) \ (ftype(x) = 1 \ ftype(y) = 0)" + using assms unfolding frecR_def by auto + +lemma frecRI1: "s \ domain(n1) \ s \ domain(n2) \ frecR(\1, s, n1, q\, \0, n1, n2, q'\)" + unfolding frecR_def by (simp add:components_simp) + +lemma frecRI1': "s \ domain(n1) \ domain(n2) \ frecR(\1, s, n1, q\, \0, n1, n2, q'\)" + unfolding frecR_def by (simp add:components_simp) + +lemma frecRI2: "s \ domain(n1) \ s \ domain(n2) \ frecR(\1, s, n2, q\, \0, n1, n2, q'\)" + unfolding frecR_def by (simp add:components_simp) + +lemma frecRI2': "s \ domain(n1) \ domain(n2) \ frecR(\1, s, n2, q\, \0, n1, n2, q'\)" + unfolding frecR_def by (simp add:components_simp) + +lemma frecRI3: "\s, r\ \ n2 \ frecR(\0, n1, s, q\, \1, n1, n2, q'\)" + unfolding frecR_def by (auto simp add:components_simp) + +lemma frecRI3': "s \ domain(n2) \ frecR(\0, n1, s, q\, \1, n1, n2, q'\)" + unfolding frecR_def by (auto simp add:components_simp) + +lemma frecR_D1 : + "frecR(x,y) \ ftype(y) = 0 \ ftype(x) = 1 \ + (name1(x) \ domain(name1(y)) \ domain(name2(y)) \ (name2(x) = name1(y) \ name2(x) = name2(y)))" + unfolding frecR_def + by auto + +lemma frecR_D2 : + "frecR(x,y) \ ftype(y) = 1 \ ftype(x) = 0 \ + ftype(x) = 0 \ ftype(y) = 1 \ name1(x) = name1(y) \ name2(x) \ domain(name2(y))" + unfolding frecR_def + by auto + +lemma frecR_DI : + assumes "frecR(\a,b,c,d\,\ftype(y),name1(y),name2(y),cond_of(y)\)" + shows "frecR(\a,b,c,d\,y)" + using assms + unfolding frecR_def + by (force simp add:components_simp) + +reldb_add "ftype" "is_ftype" +reldb_add "name1" "is_name1" +reldb_add "name2" "is_name2" + +relativize "frecR" "is_frecR" + +schematic_goal sats_frecR_fm_auto: + assumes + "i\nat" "j\nat" "env\list(A)" + shows + "is_frecR(##A,nth(i,env),nth(j,env)) \ A, env \ ?fr_fm(i,j)" + unfolding is_frecR_def + by (insert assms ; (rule sep_rules' cartprod_iff_sats components_iff_sats + | simp del:sats_cartprod_fm)+) + +synthesize "frecR" from_schematic sats_frecR_fm_auto + +text\Third item of Kunen's observations (p. 257) about the trcl relation.\ +lemma eq_ftypep_not_frecrR: + assumes "ftype(x) = ftype(y)" + shows "\ frecR(x,y)" + using assms frecR_ftypeD by force + +definition + rank_names :: "i \ i" where + "rank_names(x) \ max(rank(name1(x)),rank(name2(x)))" + +lemma rank_names_types [TC]: + shows "Ord(rank_names(x))" + unfolding rank_names_def max_def using Ord_rank Ord_Un by auto + +definition + mtype_form :: "i \ i" where + "mtype_form(x) \ if rank(name1(x)) < rank(name2(x)) then 0 else 2" + +definition + type_form :: "i \ i" where + "type_form(x) \ if ftype(x) = 0 then 1 else mtype_form(x)" + +lemma type_form_tc [TC]: + shows "type_form(x) \ 3" + unfolding type_form_def mtype_form_def by auto + +lemma frecR_le_rnk_names : + assumes "frecR(x,y)" + shows "rank_names(x)\rank_names(y)" +proof - + obtain a b c d where + H: "a = name1(x)" "b = name2(x)" + "c = name1(y)" "d = name2(y)" + "(a \ domain(c)\domain(d) \ (b=c \ b = d)) \ (a = c \ b \ domain(d))" + using assms + unfolding frecR_def + by force + then + consider + (m) "a \ domain(c) \ (b = c \ b = d) " + | (n) "a \ domain(d) \ (b = c \ b = d)" + | (o) "b \ domain(d) \ a = c" + by auto + then + show ?thesis + proof(cases) + case m + then + have "rank(a) < rank(c)" + using eclose_rank_lt in_dom_in_eclose + by simp + with \rank(a) < rank(c)\ H m + show ?thesis + unfolding rank_names_def + using Ord_rank max_cong max_cong2 leI + by auto + next + case n + then + have "rank(a) < rank(d)" + using eclose_rank_lt in_dom_in_eclose + by simp + with \rank(a) < rank(d)\ H n + show ?thesis + unfolding rank_names_def + using Ord_rank max_cong2 max_cong max_commutes[of "rank(c)" "rank(d)"] leI + by auto + next + case o + then + have "rank(b) < rank(d)" (is "?b < ?d") "rank(a) = rank(c)" (is "?a = _") + using eclose_rank_lt in_dom_in_eclose + by simp_all + with H + show ?thesis + unfolding rank_names_def + using Ord_rank max_commutes max_cong2[OF leI[OF \?b < ?d\], of ?a] + by simp + qed +qed + +definition + \ :: "i \ i" where + "\(x) = 3 ** rank_names(x) ++ type_form(x)" + +lemma \_type [TC]: + shows "Ord(\(x))" + unfolding \_def by simp + +lemma \_mono : + assumes "frecR(x,y)" + shows "\(x) < \(y)" +proof - + have F: "type_form(x) < 3" "type_form(y) < 3" + using ltI + by simp_all + from assms + have A: "rank_names(x) \ rank_names(y)" (is "?x \ ?y") + using frecR_le_rnk_names + by simp + then + have "Ord(?y)" + unfolding rank_names_def + using Ord_rank max_def + by simp + note leE[OF \?x\?y\] + then + show ?thesis + proof(cases) + case 1 + then + show ?thesis + unfolding \_def + using oadd_lt_mono2 \?x < ?y\ F + by auto + next + case 2 + consider (a) "ftype(x) = 0 \ ftype(y) = 1" | (b) "ftype(x) = 1 \ ftype(y) = 0" + using frecR_ftypeD[OF \frecR(x,y)\] + by auto + then show ?thesis + proof(cases) + case b + moreover from this + have "type_form(y) = 1" + using type_form_def by simp + moreover from calculation + have "name2(x) = name1(y) \ name2(x) = name2(y) " (is "?\ = ?\' \ ?\ = ?\'") + "name1(x) \ domain(name1(y)) \ domain(name2(y))" (is "?\ \ domain(?\') \ domain(?\')") + using assms unfolding type_form_def frecR_def by auto + moreover from calculation + have E: "rank(?\) = rank(?\') \ rank(?\) = rank(?\')" by auto + from calculation + consider (c) "rank(?\) < rank(?\')" | (d) "rank(?\) < rank(?\')" + using eclose_rank_lt in_dom_in_eclose by force + then + have "rank(?\) < rank(?\)" + proof (cases) + case c + with \rank_names(x) = rank_names(y) \ + show ?thesis + unfolding rank_names_def mtype_form_def type_form_def + using max_D2[OF E c] E assms Ord_rank + by simp + next + case d + with \rank_names(x) = rank_names(y) \ + show ?thesis + unfolding rank_names_def mtype_form_def type_form_def + using max_D2[OF _ d] max_commutes E assms Ord_rank disj_commute + by simp + qed + with b + have "type_form(x) = 0" unfolding type_form_def mtype_form_def by simp + with \rank_names(x) = rank_names(y) \ \type_form(y) = 1\ \type_form(x) = 0\ + show ?thesis + unfolding \_def by auto + next + case a + then + have "name1(x) = name1(y)" (is "?\ = ?\'") + "name2(x) \ domain(name2(y))" (is "?\ \ domain(?\')") + "type_form(x) = 1" + using assms + unfolding type_form_def frecR_def + by auto + then + have "rank(?\) = rank(?\')" "rank(?\) < rank(?\')" + using eclose_rank_lt in_dom_in_eclose + by simp_all + with \rank_names(x) = rank_names(y) \ + have "rank(?\') \ rank(?\')" + using Ord_rank max_D1 + unfolding rank_names_def + by simp + with a + have "type_form(y) = 2" + unfolding type_form_def mtype_form_def + using not_lt_iff_le assms + by simp + with \rank_names(x) = rank_names(y) \ \type_form(y) = 2\ \type_form(x) = 1\ + show ?thesis + unfolding \_def by auto + qed + qed +qed + +definition + frecrel :: "i \ i" where + "frecrel(A) \ Rrel(frecR,A)" + +lemma frecrelI : + assumes "x \ A" "y\A" "frecR(x,y)" + shows "\x,y\\frecrel(A)" + using assms unfolding frecrel_def Rrel_def by auto + +lemma frecrelD : + assumes "\x,y\ \ frecrel(A1\A2\A3\A4)" + shows + "ftype(x) \ A1" "ftype(x) \ A1" + "name1(x) \ A2" "name1(y) \ A2" + "name2(x) \ A3" "name2(x) \ A3" + "cond_of(x) \ A4" "cond_of(y) \ A4" + "frecR(x,y)" + using assms + unfolding frecrel_def Rrel_def ftype_def by (auto simp add:components_simp) + +lemma wf_frecrel : + shows "wf(frecrel(A))" +proof - + have "frecrel(A) \ measure(A,\)" + unfolding frecrel_def Rrel_def measure_def + using \_mono + by force + then + show ?thesis + using wf_subset wf_measure by auto +qed + +lemma core_induction_aux: + fixes A1 A2 :: "i" + assumes + "Transset(A1)" + "\\ \ p. p \ A2 \ \\q \. \ q\A2 ; \\domain(\)\ \ Q(0,\,\,q)\ \ Q(1,\,\,p)" + "\\ \ p. p \ A2 \ \\q \. \ q\A2 ; \\domain(\) \ domain(\)\ \ Q(1,\,\,q) \ Q(1,\,\,q)\ \ Q(0,\,\,p)" + shows "a\2\A1\A1\A2 \ Q(ftype(a),name1(a),name2(a),cond_of(a))" +proof (induct a rule:wf_induct[OF wf_frecrel[of "2\A1\A1\A2"]]) + case (1 x) + let ?\ = "name1(x)" + let ?\ = "name2(x)" + let ?D = "2\A1\A1\A2" + assume "x \ ?D" + then + have "cond_of(x)\A2" + by (auto simp add:components_simp) + from \x\?D\ + consider (eq) "ftype(x)=0" | (mem) "ftype(x)=1" + by (auto simp add:components_simp) + then + show ?case + proof cases + case eq + then + have "Q(1, \, ?\, q) \ Q(1, \, ?\, q)" if "\ \ domain(?\) \ domain(?\)" and "q\A2" for q \ + proof - + from 1 + have "?\\A1" "?\\A1" "?\\eclose(A1)" "?\\eclose(A1)" + using arg_into_eclose + by (auto simp add:components_simp) + moreover from \Transset(A1)\ that(1) + have "\\eclose(?\) \ eclose(?\)" + using in_dom_in_eclose + by auto + then + have "\\A1" + using mem_eclose_subset[OF \?\\A1\] mem_eclose_subset[OF \?\\A1\] + Transset_eclose_eq_arg[OF \Transset(A1)\] + by auto + with \q\A2\ \?\ \ A1\ \cond_of(x)\A2\ \?\\A1\ + have "frecR(\1, \, ?\, q\, x)" (is "frecR(?T,_)") + "frecR(\1, \, ?\, q\, x)" (is "frecR(?U,_)") + using frecRI1'[OF that(1)] frecR_DI \ftype(x) = 0\ + frecRI2'[OF that(1)] + by (auto simp add:components_simp) + with \x\?D\ \\\A1\ \q\A2\ + have "\?T,x\\ frecrel(?D)" "\?U,x\\ frecrel(?D)" + using frecrelI[of ?T ?D x] frecrelI[of ?U ?D x] + by (auto simp add:components_simp) + with \q\A2\ \\\A1\ \?\\A1\ \?\\A1\ + have "Q(1, \, ?\, q)" + using 1 + by (force simp add:components_simp) + moreover from \q\A2\ \\\A1\ \?\\A1\ \?\\A1\ \\?U,x\\ frecrel(?D)\ + have "Q(1, \, ?\, q)" + using 1 by (force simp add:components_simp) + ultimately + show ?thesis + by simp + qed + with assms(3) \ftype(x) = 0\ \cond_of(x)\A2\ + show ?thesis + by auto + next + case mem + have "Q(0, ?\, \, q)" if "\ \ domain(?\)" and "q\A2" for q \ + proof - + from 1 assms + have "?\\A1" "?\\A1" "cond_of(x)\A2" "?\\eclose(A1)" "?\\eclose(A1)" + using arg_into_eclose + by (auto simp add:components_simp) + with \Transset(A1)\ that(1) + have "\\ eclose(?\)" + using in_dom_in_eclose + by auto + then + have "\\A1" + using mem_eclose_subset[OF \?\\A1\] Transset_eclose_eq_arg[OF \Transset(A1)\] + by auto + with \q\A2\ \?\ \ A1\ \cond_of(x)\A2\ \?\\A1\ \ftype(x) = 1\ + have "frecR(\0, ?\, \, q\, x)" (is "frecR(?T,_)") + using frecRI3'[OF that(1)] frecR_DI + by (auto simp add:components_simp) + with \x\?D\ \\\A1\ \q\A2\ \?\\A1\ + have "\?T,x\\ frecrel(?D)" "?T\?D" + using frecrelI[of ?T ?D x] + by (auto simp add:components_simp) + with \q\A2\ \\\A1\ \?\\A1\ \?\\A1\ 1 + show ?thesis + by (force simp add:components_simp) + qed + with assms(2) \ftype(x) = 1\ \cond_of(x)\A2\ + show ?thesis + by auto + qed +qed + +lemma def_frecrel : "frecrel(A) = {z\A\A. \x y. z = \x, y\ \ frecR(x,y)}" + unfolding frecrel_def Rrel_def .. + +lemma frecrel_fst_snd: + "frecrel(A) = {z \ A\A . + ftype(fst(z)) = 1 \ + ftype(snd(z)) = 0 \ name1(fst(z)) \ domain(name1(snd(z))) \ domain(name2(snd(z))) \ + (name2(fst(z)) = name1(snd(z)) \ name2(fst(z)) = name2(snd(z))) + \ (ftype(fst(z)) = 0 \ + ftype(snd(z)) = 1 \ name1(fst(z)) = name1(snd(z)) \ name2(fst(z)) \ domain(name2(snd(z))))}" + unfolding def_frecrel frecR_def + by (intro equalityI subsetI CollectI; elim CollectE; auto) + +end \ No newline at end of file diff --git a/thys/Independence_CH/FrecR_Arities.thy b/thys/Independence_CH/FrecR_Arities.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/FrecR_Arities.thy @@ -0,0 +1,79 @@ +theory FrecR_Arities + imports + FrecR +begin + +context + notes FOL_arities[simp] +begin + +arity_theorem intermediate for "fst_fm" +lemma arity_fst_fm [arity] : + "\x\nat ; t\nat\ \ arity(fst_fm(x,t)) = succ(x) \ succ(t)" + using arity_fst_fm' + by auto + +arity_theorem intermediate for "snd_fm" +lemma arity_snd_fm [arity] : + "\x\nat ; t\nat\ \ arity(snd_fm(x,t)) = succ(x) \ succ(t)" + using arity_snd_fm' + by auto + +lemma arity_snd_snd_fm [arity] : + "\x\nat ; t\nat\ \ arity(snd_snd_fm(x,t)) = succ(x) \ succ(t)" + unfolding snd_snd_fm_def hcomp_fm_def + using arity_snd_fm arity_empty_fm union_abs2 pred_Un_distrib + by auto + +lemma arity_ftype_fm [arity] : + "\x\nat ; t\nat\ \ arity(ftype_fm(x,t)) = succ(x) \ succ(t)" + unfolding ftype_fm_def + using arity_fst_fm + by auto + +lemma arity_name1_fm [arity] : + "\x\nat ; t\nat\ \ arity(name1_fm(x,t)) = succ(x) \ succ(t)" + unfolding name1_fm_def hcomp_fm_def + using arity_fst_fm arity_snd_fm union_abs2 pred_Un_distrib + by auto + +lemma arity_name2_fm [arity] : + "\x\nat ; t\nat\ \ arity(name2_fm(x,t)) = succ(x) \ succ(t)" + unfolding name2_fm_def hcomp_fm_def + using arity_fst_fm arity_snd_snd_fm union_abs2 pred_Un_distrib + by auto + +lemma arity_cond_of_fm [arity] : + "\x\nat ; t\nat\ \ arity(cond_of_fm(x,t)) = succ(x) \ succ(t)" + unfolding cond_of_fm_def hcomp_fm_def + using arity_snd_fm arity_snd_snd_fm union_abs2 pred_Un_distrib + by auto + +lemma arity_eclose_n1_fm [arity] : + "\x\nat ; t\nat\ \ arity(eclose_n1_fm(x,t)) = succ(x) \ succ(t)" + unfolding eclose_n1_fm_def + using arity_is_eclose_fm arity_singleton_fm arity_name1_fm union_abs2 pred_Un_distrib + by auto + +lemma arity_eclose_n2_fm [arity] : + "\x\nat ; t\nat\ \ arity(eclose_n2_fm(x,t)) = succ(x) \ succ(t)" + unfolding eclose_n2_fm_def + using arity_is_eclose_fm arity_singleton_fm arity_name2_fm union_abs2 pred_Un_distrib + by auto + +lemma arity_ecloseN_fm [arity] : + "\x\nat ; t\nat\ \ arity(ecloseN_fm(x,t)) = succ(x) \ succ(t)" + unfolding ecloseN_fm_def + using arity_eclose_n1_fm arity_eclose_n2_fm arity_union_fm union_abs2 pred_Un_distrib + by auto + +lemma arity_frecR_fm [arity]: + "\a\nat;b\nat\ \ arity(frecR_fm(a,b)) = succ(a) \ succ(b)" + unfolding frecR_fm_def + using arity_ftype_fm arity_name1_fm arity_name2_fm arity_domain_fm + arity_empty_fm arity_union_fm pred_Un_distrib arity_succ_fm + by auto + +end \ \@{thm [source] FOL_arities}\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Infinity_Axiom.thy b/thys/Independence_CH/Infinity_Axiom.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Infinity_Axiom.thy @@ -0,0 +1,37 @@ +section\The Axiom of Infinity in $M[G]$\ +theory Infinity_Axiom + imports Separation_Axiom Union_Axiom Pairing_Axiom +begin + +context G_generic1 begin + +interpretation mg_triv: M_trivial"##M[G]" + using transitivity_MG zero_in_MG generic Union_MG pairing_in_MG + by unfold_locales auto + +lemma infinity_in_MG : "infinity_ax(##M[G])" +proof - + from infinity_ax + obtain I where "I\M" "0 \ I" "\y\M. y \ I \ succ(y) \ I" + unfolding infinity_ax_def by auto + then + have "check(I) \ M" + using check_in_M by simp + then + have "I\ M[G]" + using valcheck generic one_in_G one_in_P GenExtI[of "check(I)" G] by simp + moreover from this \I\M[G]\ \\y\M. y \ I \ succ(y) \ I\ + have "succ(y) \ I \ M[G]" if "y \ I" for y + using that transitivity_MG transitivity[OF _ \I\M\] by blast + moreover + note \0\I\ + ultimately + show ?thesis + using transitivity_MG[of _ I] + unfolding infinity_ax_def + by auto +qed + +end \ \\<^locale>\G_generic1\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Interface.thy b/thys/Independence_CH/Interface.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Interface.thy @@ -0,0 +1,1535 @@ +section\Interface between set models and Constructibility\ + +text\This theory provides an interface between Paulson's +relativization results and set models of ZFC. In particular, +it is used to prove that the locale \<^term>\forcing_data\ is +a sublocale of all relevant locales in \<^session>\ZF-Constructible\ +(\<^term>\M_trivial\, \<^term>\M_basic\, \<^term>\M_eclose\, etc). + +In order to interpret the locales in \<^session>\ZF-Constructible\ we +introduce new locales, each stronger than the previous one, assuming +only the instances of Replacement needed to interpret the subsequent +locales of that session. From the start we assume Separation for +every internalized formula (with one parameter, but this is not a +problem since we can use pairing).\ + +theory Interface + imports + Fm_Definitions + Transitive_Models.Cardinal_AC_Relative + Transitive_Models.M_Basic_No_Repl +begin + +locale M_Z_basic = + fixes M + assumes + upair_ax: "upair_ax(##M)" and + Union_ax: "Union_ax(##M)" and + power_ax: "power_ax(##M)" and + extensionality:"extensionality(##M)" and + foundation_ax: "foundation_ax(##M)" and + infinity_ax: "infinity_ax(##M)" and + separation_ax: "\ \ formula \ env \ list(M) \ + arity(\) \ 1 +\<^sub>\ length(env) \ + separation(##M,\x. (M, [x] @ env \ \))" + +locale M_transset = + fixes M + assumes + trans_M: "Transset(M)" + +locale M_Z_trans = M_Z_basic + M_transset + +locale M_ZF1 = M_Z_basic + + assumes + replacement_ax1: + "replacement_assm(M,env,wfrec_Hfrc_at_fm)" + "replacement_assm(M,env,list_repl1_intf_fm)" + "replacement_assm(M,env,list_repl2_intf_fm)" + "replacement_assm(M,env,formula_repl2_intf_fm)" + "replacement_assm(M,env,eclose_repl2_intf_fm)" + "replacement_assm(M,env,powapply_repl_fm)" + "replacement_assm(M,env,phrank_repl_fm)" + "replacement_assm(M,env,wfrec_rank_fm)" + "replacement_assm(M,env,trans_repl_HVFrom_fm)" + "replacement_assm(M,env,wfrec_Hcheck_fm)" + "replacement_assm(M,env,repl_PHcheck_fm)" + "replacement_assm(M,env,check_replacement_fm)" + "replacement_assm(M,env,G_dot_in_M_fm)" + "replacement_assm(M,env,repl_opname_check_fm)" + "replacement_assm(M,env,tl_repl_intf_fm)" + "replacement_assm(M,env,formula_repl1_intf_fm)" + "replacement_assm(M,env,eclose_repl1_intf_fm)" + +definition instances1_fms where "instances1_fms \ + { wfrec_Hfrc_at_fm, + list_repl1_intf_fm, + list_repl2_intf_fm, + formula_repl2_intf_fm, + eclose_repl2_intf_fm, + powapply_repl_fm, + phrank_repl_fm, + wfrec_rank_fm, + trans_repl_HVFrom_fm, + wfrec_Hcheck_fm, + repl_PHcheck_fm, + check_replacement_fm, + G_dot_in_M_fm, + repl_opname_check_fm, + tl_repl_intf_fm, + formula_repl1_intf_fm, + eclose_repl1_intf_fm }" + +txt\This set has 17 internalized formulas.\ + +lemmas replacement_instances1_defs = tl_repl_intf_fm_def formula_repl1_intf_fm_def + eclose_repl1_intf_fm_def wfrec_Hfrc_at_fm_def + list_repl1_intf_fm_def list_repl2_intf_fm_def formula_repl2_intf_fm_def + eclose_repl2_intf_fm_def powapply_repl_fm_def phrank_repl_fm_def wfrec_rank_fm_def + trans_repl_HVFrom_fm_def wfrec_Hcheck_fm_def repl_PHcheck_fm_def check_replacement_fm_def + G_dot_in_M_fm_def repl_opname_check_fm_def + +lemma instances1_fms_type[TC]: "instances1_fms \ formula" + unfolding replacement_instances1_defs instances1_fms_def by simp + +declare (in M_ZF1) replacement_instances1_defs[simp] + +locale M_ZF1_trans = M_ZF1 + M_Z_trans + +context M_Z_trans +begin + +lemmas transitivity = Transset_intf[OF trans_M] + +subsection\Interface with \<^term>\M_trivial\\ + +lemma zero_in_M: "0 \ M" +proof - + obtain z where "empty(##M,z)" "z\M" + using empty_intf[OF infinity_ax] + by auto + moreover from this + have "z=0" + using transitivity empty_def + by auto + ultimately + show ?thesis + by simp +qed + +lemma separation_in_ctm : + assumes + "\ \ formula" "env\list(M)" + "arity(\) \ 1 +\<^sub>\ length(env)" and + satsQ: "\x. x\M \ (M, [x]@env \ \) \ Q(x)" + shows + "separation(##M,Q)" + using assms separation_ax satsQ transitivity + separation_cong[of "##M" "\y. (M, [y]@env \ \)" "Q"] + by simp + +end \ \\<^locale>\M_Z_trans\\ + +locale M_ZC_basic = M_Z_basic + M_AC "##M" + +locale M_ZFC1 = M_ZF1 + M_ZC_basic + +locale M_ZFC1_trans = M_ZF1_trans + M_ZFC1 + +sublocale M_Z_trans \ M_trans "##M" + using transitivity zero_in_M exI[of "\x. x\M"] + by unfold_locales simp_all + +sublocale M_Z_trans \ M_trivial "##M" + using upair_ax Union_ax by unfold_locales + +subsection\Interface with \<^term>\M_basic\\ + +definition Intersection where + "Intersection(N,B,x) \ (\y[N]. y\B \ x\y)" + +synthesize "Intersection" from_definition "Intersection" assuming "nonempty" +arity_theorem for "Intersection_fm" + +definition CartProd where + "CartProd(N,B,C,z) \ (\x[N]. x\B \ (\y[N]. y\C \ pair(N,x,y,z)))" + +synthesize "CartProd" from_definition "CartProd" assuming "nonempty" +arity_theorem for "CartProd_fm" + +definition Image where + "Image(N,B,r,y) \ (\p[N]. p\r \ (\x[N]. x\B \ pair(N,x,y,p)))" + +synthesize "Image" from_definition "Image" assuming "nonempty" +arity_theorem for "Image_fm" + +definition Converse where + "Converse(N,R,z) \ \p[N]. p\R \ (\x[N].\y[N]. pair(N,x,y,p) \ pair(N,y,x,z))" + +synthesize "Converse" from_definition "Converse" assuming "nonempty" +arity_theorem for "Converse_fm" + +definition Restrict where + "Restrict(N,A,z) \ \x[N]. x\A \ (\y[N]. pair(N,x,y,z))" + +synthesize "Restrict" from_definition "Restrict" assuming "nonempty" +arity_theorem for "Restrict_fm" + +definition Comp where + "Comp(N,R,S,xz) \ \x[N]. \y[N]. \z[N]. \xy[N]. \yz[N]. + pair(N,x,z,xz) \ pair(N,x,y,xy) \ pair(N,y,z,yz) \ xy\S \ yz\R" + +synthesize "Comp" from_definition "Comp" assuming "nonempty" +arity_theorem for "Comp_fm" + +definition Pred where + "Pred(N,R,X,y) \ \p[N]. p\R \ pair(N,y,X,p)" + +synthesize "Pred" from_definition "Pred" assuming "nonempty" +arity_theorem for "Pred_fm" + +definition is_Memrel where + "is_Memrel(N,z) \ \x[N]. \y[N]. pair(N,x,y,z) \ x \ y" + +synthesize "is_Memrel" from_definition "is_Memrel" assuming "nonempty" +arity_theorem for "is_Memrel_fm" + +definition RecFun where + "RecFun(N,r,f,g,a,b,x) \ \xa[N]. \xb[N]. + pair(N,x,a,xa) \ xa \ r \ pair(N,x,b,xb) \ xb \ r \ + (\fx[N]. \gx[N]. fun_apply(N,f,x,fx) \ fun_apply(N,g,x,gx) \ + fx \ gx)" + +synthesize "RecFun" from_definition "RecFun" assuming "nonempty" +arity_theorem for "RecFun_fm" + +arity_theorem for "rtran_closure_mem_fm" + +synthesize "wellfounded_trancl" from_definition assuming "nonempty" +arity_theorem for "wellfounded_trancl_fm" + +context M_Z_trans +begin + +lemma inter_sep_intf : + assumes "A\M" + shows "separation(##M,\x . \y\M . y\A \ x\y)" + using assms separation_in_ctm[of "Intersection_fm(1,0)" "[A]" "Intersection(##M,A)"] + Intersection_iff_sats[of 1 "[_,A]" A 0 _ M] arity_Intersection_fm Intersection_fm_type + ord_simp_union zero_in_M + unfolding Intersection_def + by simp + +lemma diff_sep_intf : + assumes "B\M" + shows "separation(##M,\x . x\B)" + using assms separation_in_ctm[of "Neg(Member(0,1))" "[B]" "\x . x\B"] ord_simp_union + by simp + +lemma cartprod_sep_intf : + assumes "A\M" and "B\M" + shows "separation(##M,\z. \x\M. x\A \ (\y\M. y\B \ pair(##M,x,y,z)))" + using assms separation_in_ctm[of "CartProd_fm(1,2,0)" "[A,B]" "CartProd(##M,A,B)"] + CartProd_iff_sats[of 1 "[_,A,B]" A 2 B 0 _ M] arity_CartProd_fm CartProd_fm_type + ord_simp_union zero_in_M + unfolding CartProd_def + by simp + +lemma image_sep_intf : + assumes "A\M" and "B\M" + shows "separation(##M, \y. \p\M. p\B \ (\x\M. x\A \ pair(##M,x,y,p)))" + using assms separation_in_ctm[of "Image_fm(1,2,0)" "[A,B]" "Image(##M,A,B)"] + Image_iff_sats[of 1 "[_,A,B]" _ 2 _ 0 _ M] arity_Image_fm Image_fm_type + ord_simp_union zero_in_M + unfolding Image_def + by simp + +lemma converse_sep_intf : + assumes "R\M" + shows "separation(##M,\z. \p\M. p\R \ (\x\M.\y\M. pair(##M,x,y,p) \ pair(##M,y,x,z)))" + using assms separation_in_ctm[of "Converse_fm(1,0)" "[R]" "Converse(##M,R)"] + Converse_iff_sats[of 1 "[_,R]" _ 0 _ M] arity_Converse_fm Converse_fm_type + ord_simp_union zero_in_M + unfolding Converse_def + by simp + +lemma restrict_sep_intf : + assumes "A\M" + shows "separation(##M,\z. \x\M. x\A \ (\y\M. pair(##M,x,y,z)))" + using assms separation_in_ctm[of "Restrict_fm(1,0)" "[A]" "Restrict(##M,A)"] + Restrict_iff_sats[of 1 "[_,A]" _ 0 _ M] arity_Restrict_fm Restrict_fm_type + ord_simp_union zero_in_M + unfolding Restrict_def + by simp + +lemma comp_sep_intf : + assumes "R\M" and "S\M" + shows "separation(##M,\xz. \x\M. \y\M. \z\M. \xy\M. \yz\M. + pair(##M,x,z,xz) \ pair(##M,x,y,xy) \ pair(##M,y,z,yz) \ xy\S \ yz\R)" + using assms separation_in_ctm[of "Comp_fm(1,2,0)" "[R,S]" "Comp(##M,R,S)"] + Comp_iff_sats[of 1 "[_,R,S]" _ 2 _ 0 _ M] arity_Comp_fm Comp_fm_type + ord_simp_union zero_in_M + unfolding Comp_def + by simp + +lemma pred_sep_intf: + assumes "R\M" and "X\M" + shows "separation(##M, \y. \p\M. p\R \ pair(##M,y,X,p))" + using assms separation_in_ctm[of "Pred_fm(1,2,0)" "[R,X]" "Pred(##M,R,X)"] + Pred_iff_sats[of 1 "[_,R,X]" _ 2 _ 0 _ M] arity_Pred_fm Pred_fm_type + ord_simp_union zero_in_M + unfolding Pred_def + by simp + +lemma memrel_sep_intf: + "separation(##M, \z. \x\M. \y\M. pair(##M,x,y,z) \ x \ y)" + using separation_in_ctm[of "is_Memrel_fm(0)" "[]" "is_Memrel(##M)"] + is_Memrel_iff_sats[of 0 "[_]" _ M] arity_is_Memrel_fm is_Memrel_fm_type + ord_simp_union zero_in_M + unfolding is_Memrel_def + by simp + +lemma is_recfun_sep_intf : + assumes "r\M" "f\M" "g\M" "a\M" "b\M" + shows "separation(##M,\x. \xa\M. \xb\M. + pair(##M,x,a,xa) \ xa \ r \ pair(##M,x,b,xb) \ xb \ r \ + (\fx\M. \gx\M. fun_apply(##M,f,x,fx) \ fun_apply(##M,g,x,gx) \ + fx \ gx))" + using assms separation_in_ctm[of "RecFun_fm(1,2,3,4,5,0)" "[r,f,g,a,b]" "RecFun(##M,r,f,g,a,b)"] + RecFun_iff_sats[of 1 "[_,r,f,g,a,b]" _ 2 _ 3 _ 4 _ 5 _ 0 _ M] arity_RecFun_fm RecFun_fm_type + ord_simp_union zero_in_M + unfolding RecFun_def + by simp + +lemmas M_basic_sep_instances = + inter_sep_intf diff_sep_intf cartprod_sep_intf + image_sep_intf converse_sep_intf restrict_sep_intf + pred_sep_intf memrel_sep_intf comp_sep_intf is_recfun_sep_intf +end \ \\<^locale>\M_Z_trans\\ + +sublocale M_Z_trans \ M_basic_no_repl "##M" + using power_ax M_basic_sep_instances + by unfold_locales simp_all + +lemma Replace_eq_Collect: + assumes "\x y y'. x\A \ P(x,y) \ P(x,y') \ y=y'" "{y . x \ A, P(x, y)} \ B" + shows "{y . x \ A, P(x, y)} = {y\B . \x\A. P(x,y)}" + using assms by blast + +context M_Z_trans +begin + +lemma Pow_inter_M_closed: assumes "A \ M" shows "Pow(A) \ M \ M" +proof - + have "{a \ Pow(A) . a \ M} = Pow(A) \ M" by auto + then + show ?thesis + using power_ax powerset_abs assms unfolding power_ax_def + by auto +qed + +lemma Pow'_inter_M_closed: assumes "A \ M" shows "{a \ Pow(A) . a \ M} \ M" + using power_ax powerset_abs assms unfolding power_ax_def by auto + +end \ \\<^locale>\M_Z_trans\\ + +context M_basic_no_repl +begin + +lemma Replace_funspace_succ_rep_intf_sub: + assumes + "M(A)" "M(n)" + shows + "{z . p \ A, funspace_succ_rep_intf_rel(M,p,z,n)} + \ Pow\<^bsup>M\<^esup>(Pow\<^bsup>M\<^esup>(\domain(A) \ ({n} \ range(A)) \ (\({n} \ range(A)))))" + unfolding funspace_succ_rep_intf_rel_def using assms mem_Pow_rel_abs + by clarsimp (auto simp: cartprod_def) + +lemma funspace_succ_rep_intf_uniq: + assumes + "funspace_succ_rep_intf_rel(M,p,z,n)" "funspace_succ_rep_intf_rel(M,p,z',n)" + shows + "z = z'" + using assms unfolding funspace_succ_rep_intf_rel_def by auto + +lemma Replace_funspace_succ_rep_intf_eq: + assumes + "M(A)" "M(n)" + shows + "{z . p \ A, funspace_succ_rep_intf_rel(M,p,z,n)} = + {z \ Pow\<^bsup>M\<^esup>(Pow\<^bsup>M\<^esup>(\domain(A) \ ({n} \ range(A)) \ (\({n} \ range(A))))) . + \p\A. funspace_succ_rep_intf_rel(M,p,z,n)}" + using assms Replace_eq_Collect[OF funspace_succ_rep_intf_uniq, of A, + OF _ _ Replace_funspace_succ_rep_intf_sub[of A n], of "\x y z. x" "\x y z. n"] + by (intro equalityI) + (auto dest:transM simp:funspace_succ_rep_intf_rel_def) + +end \ \\<^locale>\M_basic_no_repl\\ + +definition fsri where + "fsri(N,A,B) \ \z. \p\A. \f[N]. \b[N]. p = \f, b\ \ z = {cons(\B, b\, f)}" + +relationalize "fsri" "is_fsri" +synthesize "is_fsri" from_definition assuming "nonempty" +arity_theorem for "is_fsri_fm" + + +context M_Z_trans +begin + +lemma separation_fsri: + "(##M)(A) \ (##M)(B) \ separation(##M, is_fsri(##M,A,B))" + using separation_in_ctm[where env="[A,B]" and \="is_fsri_fm(1,2,0)"] + zero_in_M is_fsri_iff_sats[symmetric] arity_is_fsri_fm is_fsri_fm_type + by (simp_all add: ord_simp_union) + +lemma separation_funspace_succ_rep_intf_rel: + "(##M)(A) \ (##M)(B) \ separation(##M, \z. \p\A. funspace_succ_rep_intf_rel(##M,p,z,B))" + using separation_fsri zero_in_M + by (rule_tac separation_cong[THEN iffD1, of _ "is_fsri(##M,A,B)"]) + (auto simp flip:setclass_iff dest:transM + simp:is_fsri_def funspace_succ_rep_intf_rel_def, force) + +lemma Replace_funspace_succ_rep_intf_in_M: + assumes + "A \ M" "n \ M" + shows + "{z . p \ A, funspace_succ_rep_intf_rel(##M,p,z,n)} \ M" +proof - + have "(##M)({z \ Pow\<^bsup>M\<^esup>(Pow\<^bsup>M\<^esup>(\domain(A) \ ({n} \ range(A)) \ (\({n} \ range(A))))) . + \p\A. funspace_succ_rep_intf_rel(##M,p,z,n)})" + using assms separation_funspace_succ_rep_intf_rel + by (intro separation_closed) (auto simp flip:setclass_iff) + with assms + show ?thesis + using Replace_funspace_succ_rep_intf_eq by auto +qed + +lemma funspace_succ_rep_intf: + assumes "n\M" + shows + "strong_replacement(##M, + \p z. \f\M. \b\M. \nb\M. \cnbf\M. + pair(##M,f,b,p) \ pair(##M,n,b,nb) \ is_cons(##M,nb,f,cnbf) \ + upair(##M,cnbf,cnbf,z))" + using assms + unfolding strong_replacement_def univalent_def + apply (simp add:pair_in_M_iff[simplified]) + apply clarsimp + apply (rule_tac x="{z . p \ A, funspace_succ_rep_intf_rel(##M,p,z,n)}" in bexI) + apply (auto simp:funspace_succ_rep_intf_rel_def + Replace_funspace_succ_rep_intf_in_M[unfolded funspace_succ_rep_intf_rel_def, simplified]) + apply (rule_tac x="\f, ba\" in bexI) + apply (auto dest:transM simp:pair_in_M_iff[simplified] cons_closed[simplified]) + done + +end \ \\<^locale>\M_Z_trans\\ + +sublocale M_Z_trans \ M_basic "##M" + using power_ax M_basic_sep_instances funspace_succ_rep_intf + by unfold_locales auto + +subsection\Interface with \<^term>\M_trancl\\ + +lemma (in M_ZF1_trans) rtrancl_separation_intf: + assumes "r\M" "A\M" + shows "separation (##M, rtran_closure_mem(##M,A,r))" + using assms separation_in_ctm[of "rtran_closure_mem_fm(1,2,0)" "[A,r]" "rtran_closure_mem(##M,A,r)"] + arity_rtran_closure_mem_fm ord_simp_union zero_in_M + by simp + +context M_ZF1_trans +begin + +lemma wftrancl_separation_intf: + assumes "r\M" and "Z\M" + shows "separation (##M, wellfounded_trancl(##M,Z,r))" + using assms separation_in_ctm[of "wellfounded_trancl_fm(1,2,0)" "[Z,r]" "wellfounded_trancl(##M,Z,r)"] + arity_wellfounded_trancl_fm ord_simp_union zero_in_M + by simp + +text\To prove \<^term>\nat \ M\ we get an infinite set \<^term>\I\ from \<^term>\infinity_ax\ +closed under \<^term>\0\ and \<^term>\succ\; that shows \<^term>\nat\I\. Then we +can separate \<^term>\I\ with the predicate \<^term>\\x. x\nat\.\ +lemma finite_sep_intf: "separation(##M, \x. x\nat)" +proof - + have "(\v\M. separation(##M,\x. (M, [x,v] \ finite_ordinal_fm(0))))" + using separation_ax arity_finite_ordinal_fm + by simp + then + have "(\v\M. separation(##M,finite_ordinal(##M)))" + unfolding separation_def + by simp + then + have "separation(##M,finite_ordinal(##M))" + using separation_in_ctm zero_in_M + by auto + then + show ?thesis + unfolding separation_def + by simp +qed + +lemma nat_subset_I: "\I\M. nat \ I" +proof - + have "nat \ I" + if "I\M" and "0\I" and "\x. x\I \ succ(x)\I" for I + using that + by (rule_tac subsetI,induct_tac x,simp_all) + moreover + obtain I where + "I\M" "0\I" "\x. x\I \ succ(x)\I" + using infinity_ax transitivity + unfolding infinity_ax_def + by auto + ultimately + show ?thesis + by auto +qed + +lemma nat_in_M: "nat \ M" +proof - + have "{x\B . x\A}=A" if "A\B" for A B + using that by auto + moreover + obtain I where + "I\M" "nat\I" + using nat_subset_I by auto + moreover from this + have "{x\I . x\nat} \ M" + using finite_sep_intf separation_closed[of "\x . x\nat"] + by simp + ultimately + show ?thesis + by simp +qed + +end \ \\<^locale>\M_ZF1_trans\\ + +sublocale M_ZF1_trans \ M_trancl "##M" + using rtrancl_separation_intf wftrancl_separation_intf nat_in_M + wellfounded_trancl_def + by unfold_locales auto + +subsection\Interface with \<^term>\M_eclose\\ + +lemma repl_sats: + assumes + sat:"\x z. x\M \ z\M \ (M, Cons(x,Cons(z,env)) \ \) \ P(x,z)" + shows + "strong_replacement(##M,\x z. (M, Cons(x,Cons(z,env)) \ \)) \ + strong_replacement(##M,P)" + by (rule strong_replacement_cong,simp add:sat) + +arity_theorem for "list_functor_fm" + +lemma (in M_ZF1_trans) list_repl1_intf: + assumes "A\M" + shows "iterates_replacement(##M, is_list_functor(##M,A), 0)" +proof - + let ?f="Exists(And(pair_fm(1,0,2), + is_wfrec_fm(iterates_MH_fm(list_functor_fm(13,1,0),10,2,1,0),3,1,0)))" + have "arity(?f) = 5" + using arity_iterates_MH_fm[where isF="list_functor_fm(13,1,0)" and i=14] + arity_wfrec_replacement_fm[where i=11] arity_list_functor_fm ord_simp_union + by simp + { + fix n + assume "n\nat" + then + have "Memrel(succ(n))\M" + using nat_into_M Memrel_closed + by simp + moreover + note assms zero_in_M + moreover from calculation + have "is_list_functor(##M, A, a, b) + \ (M, [b,a,c,d,a0,a1,a2,a3,a4,y,x,z,Memrel(succ(n)),A,0] \ list_functor_fm(13,1,0))" + if "a\M" "b\M" "c\M" "d\M" "a0\M" "a1\M" "a2\M" "a3\M" "a4\M" "y\M" "x\M" "z\M" + for a b c d a0 a1 a2 a3 a4 y x z + using that + by simp + moreover from calculation + have "(M, [a0,a1,a2,a3,a4,y,x,z,Memrel(succ(n)),A,0] \ + iterates_MH_fm(list_functor_fm(13,1,0),10,2,1,0)) \ + iterates_MH(##M,is_list_functor(##M,A),0,a2, a1, a0)" + if "a0\M" "a1\M" "a2\M" "a3\M" "a4\M" "y\M" "x\M" "z\M" + for a0 a1 a2 a3 a4 y x z + using that sats_iterates_MH_fm[of M "is_list_functor(##M,A)" _] + by simp + moreover from calculation + have "(M, [y,x,z,Memrel(succ(n)),A,0] \ + is_wfrec_fm(iterates_MH_fm(list_functor_fm(13,1,0),10,2,1,0),3,1,0)) \ + is_wfrec(##M, iterates_MH(##M,is_list_functor(##M,A),0) , Memrel(succ(n)), x, y)" + if "y\M" "x\M" "z\M" for y x z + using that sats_is_wfrec_fm + by simp + moreover from calculation + have "(M, [x,z,Memrel(succ(n)),A,0] \ ?f) \ + + (\y\M. pair(##M,x,y,z) \ + is_wfrec(##M, iterates_MH(##M,is_list_functor(##M,A),0) , Memrel(succ(n)), x, y))" + if "x\M" "z\M" for x z + using that + by (simp del:pair_abs) + moreover + note \arity(?f) = 5\ + moreover from calculation + have "strong_replacement(##M,\x z. (M, [x,z,Memrel(succ(n)),A,0] \ ?f))" + using replacement_ax1(2)[unfolded replacement_assm_def] + by simp + moreover from calculation + have "strong_replacement(##M,\x z. + \y\M. pair(##M,x,y,z) \ is_wfrec(##M, iterates_MH(##M,is_list_functor(##M,A),0) , + Memrel(succ(n)), x, y))" + using repl_sats[of M ?f "[Memrel(succ(n)),A,0]"] + by (simp del:pair_abs) + } + then + show ?thesis + unfolding iterates_replacement_def wfrec_replacement_def + by simp +qed + +text\This lemma obtains \<^term>\iterates_replacement\ for predicates +without parameters.\ +lemma (in M_ZF1_trans) iterates_repl_intf : + assumes + "v\M" and + isfm:"is_F_fm \ formula" and + arty:"arity(is_F_fm)=2" and + satsf: "\a b env'. \ a\M ; b\M ; env'\list(M) \ + \ is_F(a,b) \ (M, [b,a]@env' \ is_F_fm)" + and is_F_fm_replacement: + "\env. (\\\\\1,0\ is 2\ \ is_wfrec_fm(iterates_MH_fm(is_F_fm,9,2,1,0),3,1,0) \\) \ formula \ env \ list(M) \ + arity((\\\\\1,0\ is 2\ \ is_wfrec_fm(iterates_MH_fm(is_F_fm,9,2,1,0),3,1,0) \\)) \ 2 +\<^sub>\ length(env) \ + strong_replacement(##M,\x y. + M, [x,y] @ env \ (\\\\\1,0\ is 2\ \ is_wfrec_fm(iterates_MH_fm(is_F_fm,9,2,1,0),3,1,0) \\))" + shows + "iterates_replacement(##M,is_F,v)" +proof - + let ?f="(\\\\\1,0\ is 2\ \ is_wfrec_fm(iterates_MH_fm(is_F_fm,9,2,1,0),3,1,0) \\)" + have "arity(?f) = 4" "?f\formula" + using arity_iterates_MH_fm[where isF=is_F_fm and i=2] + arity_wfrec_replacement_fm[where i=10] isfm arty ord_simp_union + by simp_all + { + fix n + assume "n\nat" + then + have "Memrel(succ(n))\M" + using nat_into_M Memrel_closed + by simp + moreover + { + fix a0 a1 a2 a3 a4 y x z + assume "[a0,a1,a2,a3,a4,y,x,z]\list(M)" + moreover + note \v\M\ \Memrel(succ(n))\M\ + moreover from calculation + have "(M, [b,a,c,d,a0,a1,a2,a3,a4,y,x,z,Memrel(succ(n)),v] \ is_F_fm) \ + is_F(a,b)" + if "a\M" "b\M" "c\M" "d\M" for a b c d + using that satsf[of a b "[c,d,a0,a1,a2,a3,a4,y,x,z,Memrel(succ(n)),v]"] + by simp + moreover from calculation + have "(M, [a0,a1,a2,a3,a4,y,x,z,Memrel(succ(n)),v] \ iterates_MH_fm(is_F_fm,9,2,1,0)) \ + iterates_MH(##M,is_F,v,a2, a1, a0)" + using sats_iterates_MH_fm[of M "is_F" "is_F_fm"] + by simp + } + moreover from calculation + have "(M, [y,x,z,Memrel(succ(n)),v] \ is_wfrec_fm(iterates_MH_fm(is_F_fm,9,2,1,0),3,1,0)) \ + is_wfrec(##M, iterates_MH(##M,is_F,v),Memrel(succ(n)), x, y)" + if "y\M" "x\M" "z\M" for y x z + using that sats_is_wfrec_fm \v\M\ by simp + moreover from calculation + have "(M, [x,z,Memrel(succ(n)),v] \ ?f) \ + + (\y\M. pair(##M,x,y,z) \ + is_wfrec(##M, iterates_MH(##M,is_F,v) , Memrel(succ(n)), x, y))" + if "x\M" "z\M" for x z + using that \v\M\ + by (simp del:pair_abs) + moreover + note \arity(?f) = 4\ \?f\formula\ + moreover from calculation \v\_\ + have "strong_replacement(##M,\x z. (M, [x,z,Memrel(succ(n)),v] \ ?f))" + using is_F_fm_replacement + by simp + ultimately + have "strong_replacement(##M,\x z. + \y\M. pair(##M,x,y,z) \ is_wfrec(##M, iterates_MH(##M,is_F,v) , + Memrel(succ(n)), x, y))" + using repl_sats[of M ?f "[Memrel(succ(n)),v]"] + by (simp del:pair_abs) + } + then + show ?thesis + unfolding iterates_replacement_def wfrec_replacement_def + by simp +qed + +arity_theorem for "formula_functor_fm" +lemma (in M_ZF1_trans) formula_repl1_intf : + "iterates_replacement(##M, is_formula_functor(##M), 0)" + using arity_formula_functor_fm zero_in_M ord_simp_union + iterates_repl_intf[where is_F_fm="formula_functor_fm(1,0)"] + replacement_ax1(16)[unfolded replacement_assm_def] + by simp + +arity_theorem for "Inl_fm" +arity_theorem for "Inr_fm" +arity_theorem for "Nil_fm" +arity_theorem for "Cons_fm" +arity_theorem for "quasilist_fm" +arity_theorem for "tl_fm" + +lemma (in M_ZF1_trans) tl_repl_intf: + assumes "l \ M" + shows "iterates_replacement(##M,\l' t. is_tl(##M,l',t),l)" + using assms arity_tl_fm ord_simp_union + iterates_repl_intf[where is_F_fm="tl_fm(1,0)"] + replacement_ax1(15)[unfolded replacement_assm_def] + by simp + +arity_theorem for "big_union_fm" + +lemma (in M_ZF1_trans) eclose_repl1_intf: + assumes "A\M" + shows "iterates_replacement(##M, big_union(##M), A)" + using assms arity_big_union_fm + iterates_repl_intf[where is_F_fm="big_union_fm(1,0)"] + replacement_ax1(17)[unfolded replacement_assm_def] + ord_simp_union + by simp + +lemma (in M_ZF1_trans) list_repl2_intf: + assumes "A\M" + shows "strong_replacement(##M,\n y. n\nat \ + is_iterates(##M, is_list_functor(##M,A), 0, n, y))" +proof - + let ?f = "And(Member(0,4),is_iterates_fm(list_functor_fm(13,1,0),3,0,1))" + note zero_in_M nat_in_M \A\M\ + moreover from this + have "is_list_functor(##M,A,a,b) \ + (M, [b,a,c,d,e,f,g,h,i,j,k,n,y,A,0,nat] \ list_functor_fm(13,1,0))" + if "a\M" "b\M" "c\M" "d\M" "e\M" "f\M""g\M""h\M""i\M""j\M" "k\M" "n\M" "y\M" + for a b c d e f g h i j k n y + using that + by simp + moreover from calculation + have "(M, [n,y,A,0,nat] \ is_iterates_fm(list_functor_fm(13,1,0),3,0,1)) \ + is_iterates(##M, is_list_functor(##M,A), 0, n , y)" + if "n\M" "y\M" for n y + using that sats_is_iterates_fm[of M "is_list_functor(##M,A)"] + by simp + moreover from calculation + have "(M, [n,y,A,0,nat] \ ?f) \ + n\nat \ is_iterates(##M, is_list_functor(##M,A), 0, n, y)" + if "n\M" "y\M" for n y + using that + by simp + moreover + have "arity(?f) = 5" + using arity_is_iterates_fm[where p="list_functor_fm(13,1,0)" and i=14] + arity_list_functor_fm arity_And ord_simp_union + by simp + ultimately + show ?thesis + using replacement_ax1(3)[unfolded replacement_assm_def] repl_sats[of M ?f "[A,0,nat]"] + by simp +qed + +lemma (in M_ZF1_trans) formula_repl2_intf: + "strong_replacement(##M,\n y. n\nat \ is_iterates(##M, is_formula_functor(##M), 0, n, y))" +proof - + let ?f = "And(Member(0,3),is_iterates_fm(formula_functor_fm(1,0),2,0,1))" + note zero_in_M nat_in_M + moreover from this + have "is_formula_functor(##M,a,b) \ + (M, [b,a,c,d,e,f,g,h,i,j,k,n,y,0,nat] \ formula_functor_fm(1,0))" + if "a\M" "b\M" "c\M" "d\M" "e\M" "f\M""g\M""h\M""i\M""j\M" "k\M" "n\M" "y\M" + for a b c d e f g h i j k n y + using that + by simp + moreover from calculation + have "(M, [n,y,0,nat] \ is_iterates_fm(formula_functor_fm(1,0),2,0,1)) \ + is_iterates(##M, is_formula_functor(##M), 0, n , y)" + if "n\M" "y\M" for n y + using that sats_is_iterates_fm[of M "is_formula_functor(##M)"] + by simp + moreover from calculation + have "(M, [n,y,0,nat] \ ?f) \ + n\nat \ is_iterates(##M, is_formula_functor(##M), 0, n, y)" + if "n\M" "y\M" for n y + using that + by simp + moreover + have "arity(?f) = 4" + using arity_is_iterates_fm[where p="formula_functor_fm(1,0)" and i=2] + arity_formula_functor_fm arity_And ord_simp_union + by simp + ultimately + show ?thesis + using replacement_ax1(4)[unfolded replacement_assm_def] repl_sats[of M ?f "[0,nat]"] + by simp +qed + + +lemma (in M_ZF1_trans) eclose_repl2_intf: + assumes "A\M" + shows "strong_replacement(##M,\n y. n\nat \ is_iterates(##M, big_union(##M), A, n, y))" +proof - + let ?f = "And(Member(0,3),is_iterates_fm(big_union_fm(1,0),2,0,1))" + note nat_in_M \A\M\ + moreover from this + have "big_union(##M,a,b) \ + (M, [b,a,c,d,e,f,g,h,i,j,k,n,y,A,nat] \ big_union_fm(1,0))" + if "a\M" "b\M" "c\M" "d\M" "e\M" "f\M""g\M""h\M""i\M""j\M" "k\M" "n\M" "y\M" + for a b c d e f g h i j k n y + using that by simp + moreover from calculation + have "(M, [n,y,A,nat] \ is_iterates_fm(big_union_fm(1,0),2,0,1)) \ + is_iterates(##M, big_union(##M), A, n , y)" + if "n\M" "y\M" for n y + using that sats_is_iterates_fm[of M "big_union(##M)"] + by simp + moreover from calculation + have "(M, [n,y,A,nat] \ ?f) \ + n\nat \ is_iterates(##M, big_union(##M), A, n, y)" + if "n\M" "y\M" for n y + using that + by simp + moreover + have "arity(?f) = 4" + using arity_is_iterates_fm[where p="big_union_fm(1,0)" and i=2] + arity_big_union_fm arity_And ord_simp_union + by simp + ultimately + show ?thesis + using repl_sats[of M ?f "[A,nat]"] replacement_ax1(5)[unfolded replacement_assm_def] + by simp +qed + +sublocale M_ZF1_trans \ M_datatypes "##M" + using list_repl1_intf list_repl2_intf formula_repl1_intf + formula_repl2_intf tl_repl_intf + by unfold_locales auto + +sublocale M_ZF1_trans \ M_eclose "##M" + using eclose_repl1_intf eclose_repl2_intf + by unfold_locales auto + +text\Interface with \<^locale>\M_eclose\.\ + +lemma (in M_ZF1_trans) Powapply_repl : + assumes "f\M" + shows "strong_replacement(##M,\x y. y=Powapply_rel(##M,f,x))" +proof - + note assms + moreover + have "arity(is_Powapply_fm(2,0,1)) = 3" + unfolding is_Powapply_fm_def + by (simp add:arity ord_simp_union) + moreover from calculation + have iff:"z=Powapply_rel(##M,f,p) \ (M, [p,z,f] \ is_Powapply_fm(2,0,1) )" + if "p\M" "z\M" for p z + using that zero_in_M sats_is_Powapply_fm[of 2 0 1 "[p,z,f]" M] is_Powapply_iff + replacement_ax1[unfolded replacement_assm_def] + by simp + ultimately + show ?thesis + using replacement_ax1(6)[unfolded replacement_assm_def] + by (rule_tac strong_replacement_cong[THEN iffD2,OF iff],simp_all) +qed + +lemma (in M_ZF1_trans) phrank_repl : + assumes + "f\M" + shows + "strong_replacement(##M, \x y. y = succ(f`x))" +proof - + note assms + moreover from this + have iff:"y = succ(f ` x) \ M, [x, y, f] \ PHrank_fm(2, 0, 1)" if "x\M" "y\M" for x y + using PHrank_iff_sats[of 2 "[x,y,f]" f 0 _ 1 _ M] zero_in_M that + apply_closed + unfolding PHrank_def + by simp + moreover + have "arity(PHrank_fm(2,0,1)) = 3" + unfolding PHrank_fm_def + by (simp add:arity ord_simp_union) + ultimately + show ?thesis + using replacement_ax1(7)[unfolded replacement_assm_def] + unfolding PHrank_def + by(rule_tac strong_replacement_cong[THEN iffD2,OF iff],simp_all) +qed + +declare is_Hrank_fm_def[fm_definitions add] +declare PHrank_fm_def[fm_definitions add] + +lemma (in M_ZF1_trans) wfrec_rank : + assumes "X\M" + shows "wfrec_replacement(##M,is_Hrank(##M),rrank(X))" +proof - + let ?f="Exists(And(pair_fm(1,0,2),is_wfrec_fm(is_Hrank_fm(2,1,0),3,1,0)))" + note assms zero_in_M + moreover from this + have + "is_Hrank(##M,a2, a1, a0) \ + (M, [a0,a1,a2,a3,a4,y,x,z,rrank(X)] \ is_Hrank_fm(2,1,0))" + if "a4\M" "a3\M" "a2\M" "a1\M" "a0\M" "y\M" "x\M" "z\M" for a4 a3 a2 a1 a0 y x z + using that rrank_in_M is_Hrank_iff_sats + by simp + moreover from calculation + have "(M, [y,x,z,rrank(X)] \ is_wfrec_fm(is_Hrank_fm(2,1,0),3,1,0)) \ + is_wfrec(##M, is_Hrank(##M) ,rrank(X), x, y)" + if "y\M" "x\M" "z\M" for y x z + using that rrank_in_M sats_is_wfrec_fm + by simp + moreover from calculation + have "(M, [x,z,rrank(X)] \ ?f) \ + (\y\M. pair(##M,x,y,z) \ is_wfrec(##M, is_Hrank(##M) , rrank(X), x, y))" + if "x\M" "z\M" for x z + using that rrank_in_M + by (simp del:pair_abs) + moreover + have "arity(?f) = 3" + using arity_wfrec_replacement_fm[where p="is_Hrank_fm(2,1,0)" and i=3,simplified] + arity_is_Hrank_fm[of 2 1 0,simplified] ord_simp_union + by simp + moreover from calculation + have "strong_replacement(##M,\x z. (M, [x,z,rrank(X)] \ ?f))" + using replacement_ax1(8)[unfolded replacement_assm_def] rrank_in_M + by simp + ultimately + show ?thesis + using repl_sats[of M ?f "[rrank(X)]"] + unfolding wfrec_replacement_def + by (simp del:pair_abs) +qed + +schematic_goal sats_is_Vset_fm_auto: + assumes + "i\nat" "v\nat" "env\list(A)" "0\A" + "i < length(env)" "v < length(env)" + shows + "is_Vset(##A,nth(i, env),nth(v, env)) \ (A, env \ ?ivs_fm(i,v))" + unfolding is_Vset_def is_Vfrom_def + by (insert assms; (rule sep_rules is_HVfrom_iff_sats is_transrec_iff_sats | simp)+) + +synthesize "is_Vset" from_schematic "sats_is_Vset_fm_auto" +arity_theorem for "is_Vset_fm" + +lemma (in M_ZF1_trans) trans_repl_HVFrom : + assumes "A\M" "i\M" + shows "transrec_replacement(##M,is_HVfrom(##M,A),i)" +proof - + let ?f="Exists(And(pair_fm(1,0,2),is_wfrec_fm(is_HVfrom_fm(8,2,1,0),4,1,0)))" + note facts = assms zero_in_M + moreover + have "\sa\M. \esa\M. \mesa\M. + upair(##M,a,a,sa) \ is_eclose(##M,sa,esa) \ membership(##M,esa,mesa)" + if "a\M" for a + using that upair_ax eclose_closed Memrel_closed + unfolding upair_ax_def + by (simp del:upair_abs) + moreover + { + fix mesa + assume "mesa\M" + moreover + note facts + moreover from calculation + have "is_HVfrom(##M,A,a2, a1, a0) \ + (M, [a0,a1,a2,a3,a4,y,x,z,A,mesa] \ is_HVfrom_fm(8,2,1,0))" + if "a4\M" "a3\M" "a2\M" "a1\M" "a0\M" "y\M" "x\M" "z\M" for a4 a3 a2 a1 a0 y x z + using that sats_is_HVfrom_fm + by simp + moreover from calculation + have "(M, [y,x,z,A,mesa] \ is_wfrec_fm(is_HVfrom_fm(8,2,1,0),4,1,0)) \ + is_wfrec(##M, is_HVfrom(##M,A),mesa, x, y)" + if "y\M" "x\M" "z\M" for y x z + using that sats_is_wfrec_fm + by simp + moreover from calculation + have "(M, [x,z,A,mesa] \ ?f) \ + (\y\M. pair(##M,x,y,z) \ is_wfrec(##M, is_HVfrom(##M,A) , mesa, x, y))" + if "x\M" "z\M" for x z + using that + by (simp del:pair_abs) + moreover + have "arity(?f) = 4" + using arity_wfrec_replacement_fm[where p="is_HVfrom_fm(8,2,1,0)" and i=9] + arity_is_HVfrom_fm ord_simp_union + by simp + moreover from calculation + have "strong_replacement(##M,\x z. (M, [x,z,A,mesa] \ ?f))" + using replacement_ax1(9)[unfolded replacement_assm_def] + by simp + ultimately + have "wfrec_replacement(##M,is_HVfrom(##M,A),mesa)" + using repl_sats[of M ?f "[A,mesa]"] + unfolding wfrec_replacement_def + by (simp del:pair_abs) + } + ultimately + show ?thesis + unfolding transrec_replacement_def + by simp +qed + +sublocale M_ZF1_trans \ M_Vfrom "##M" + using power_ax Powapply_repl phrank_repl trans_repl_HVFrom wfrec_rank + by unfold_locales auto + + +subsection\Interface for proving Collects and Replace in M.\ +context M_ZF1_trans +begin + +lemma Collect_in_M : + assumes + "\ \ formula" "env\list(M)" + "arity(\) \ 1 +\<^sub>\ length(env)" "A\M" and + satsQ: "\x. x\M \ (M, [x]@env \ \) \ Q(x)" + shows + "{y\A . Q(y)}\M" +proof - + have "separation(##M,\x. (M, [x] @ env \ \))" + using assms separation_ax by simp + then + show ?thesis + using \A\M\ satsQ transitivity separation_closed + separation_cong[of "##M" "\y. (M, [y]@env \ \)" "Q"] + by simp +qed + +\ \This version has a weaker assumption.\ +lemma separation_in_M : + assumes + "\ \ formula" "env\list(M)" + "arity(\) \ 1 +\<^sub>\ length(env)" "A\M" and + satsQ: "\x. x\A \ (M, [x]@env \ \) \ Q(x)" + shows + "{y\A . Q(y)} \ M" +proof - + let ?\' = "And(\,Member(0,length(env)+\<^sub>\1))" + note assms + moreover + have "arity(?\') \ 1 +\<^sub>\ length(env@[A])" + using assms Un_le le_trans[of "arity(\)" "1+\<^sub>\length(env)" "2+\<^sub>\length(env)"] + by (force simp:FOL_arities) + moreover from calculation + have "?\'\formula" "nth(length(env), env @ [A]) = A" + using nth_append + by auto + moreover from calculation + have "\ x . x \ M \ (M, [x]@env@[A] \ ?\') \ Q(x) \ x\A" + using arity_sats_iff[of _ "[A]" _ "[_]@env"] + by auto + ultimately + show ?thesis + using Collect_in_M[of ?\' "env@[A]" _ "\x . Q(x) \ x\A", OF _ _ _ \A\M\] + by auto +qed + +end \ \\<^locale>\M_ZF1_trans\\ + +context M_Z_trans +begin + +lemma strong_replacement_in_ctm: + assumes + f_fm: "\ \ formula" and + f_ar: "arity(\)\ 2 +\<^sub>\ length(env)" and + fsats: "\x y. x\M \ y\M \ (M,[x,y]@env \ \) \ y = f(x)" and + fclosed: "\x. x\M \ f(x) \ M" and + phi_replacement:"replacement_assm(M,env,\)" and + "env\list(M)" + shows "strong_replacement(##M, \x y . y = f(x))" + using assms + strong_replacement_cong[of "##M" "\x y. M,[x,y]@env\\" "\x y. y = f(x)"] + unfolding replacement_assm_def + by auto + +lemma strong_replacement_rel_in_ctm : + assumes + f_fm: "\ \ formula" and + f_ar: "arity(\)\ 2 +\<^sub>\ length(env)" and + fsats: "\x y. x\M \ y\M \ (M,[x,y]@env \ \) \ f(x,y)" and + phi_replacement:"replacement_assm(M,env,\)" and + "env\list(M)" + shows "strong_replacement(##M, f)" + using assms + strong_replacement_cong[of "##M" "\x y. M,[x,y]@env\\" "f"] + unfolding replacement_assm_def + by auto + +lemma Replace_in_M : + assumes + f_fm: "\ \ formula" and + f_ar: "arity(\)\ 2 +\<^sub>\ length(env)" and + fsats: "\x y. x\A \ y\M \ (M,[x,y]@env \ \) \ y = f(x)" and + fclosed: "\x. x\A \ f(x) \ M" and + "A\M" "env\list(M)" and + phi'_replacement:"replacement_assm(M,env@[A], \\ \ \0 \ length(env) +\<^sub>\ 2\\ )" + shows "{f(x) . x\A}\M" +proof - + let ?\' = "And(\,Member(0,length(env)+\<^sub>\2))" + note assms + moreover from this + have "arity(?\') \ 2 +\<^sub>\ length(env@[A])" + using Un_le le_trans[of "arity(\)" "2+\<^sub>\(length(env))" "3+\<^sub>\length(env)"] + by (force simp:FOL_arities) + moreover from calculation + have "?\'\formula" "nth(length(env), env @ [A]) = A" + using nth_append by auto + moreover from calculation + have "\ x y. x \ M \ y\M \ (M,[x,y]@env@[A]\?\') \ y=f(x) \x\A" + using arity_sats_iff[of _ "[A]" _ "[_,_]@env"] + by auto + moreover from calculation + have "strong_replacement(##M, \x y. M,[x,y]@env@[A] \ ?\')" + using phi'_replacement assms(1-6) unfolding replacement_assm_def by simp + ultimately + have 4:"strong_replacement(##M, \x y. y = f(x) \ x\A)" + using + strong_replacement_cong[of "##M" "\x y. M,[x,y]@env@[A]\?\'" "\x y. y = f(x) \ x\A"] + by simp + then + have "{y . x\A , y = f(x)} \ M" + using \A\M\ strong_replacement_closed[OF 4,of A] fclosed by simp + moreover + have "{f(x). x\A} = { y . x\A , y = f(x)}" + by auto + ultimately + show ?thesis by simp +qed + +lemma Replace_relativized_in_M : + assumes + f_fm: "\ \ formula" and + f_ar: "arity(\)\ 2 +\<^sub>\ length(env)" and + fsats: "\x y. x\A \ y\M \ (M,[x,y]@env \ \) \ is_f(x,y)" and + fabs: "\x y. x\A \ y\M \ is_f(x,y) \ y = f(x)" and + fclosed: "\x. x\A \ f(x) \ M" and + "A\M" "env\list(M)" and + phi'_replacement:"replacement_assm(M,env@[A], \\ \ \0 \ length(env) +\<^sub>\ 2\\ )" + shows "{f(x) . x\A}\M" + using assms Replace_in_M[of \] by auto + +lemma ren_action : + assumes + "env\list(M)" "x\M" "y\M" "z\M" + shows "\ i . i < 2+\<^sub>\length(env) \ + nth(i,[x,z]@env) = nth(\_repl(length(env))`i,[z,x,y]@env)" +proof - + let ?f="{\0, 1\, \1, 0\}" + have 1:"(\j. j < length(env) \ nth(j, env) = nth(id(length(env)) ` j, env))" + using assms ltD by simp + have 2:"nth(j, [x,z]) = nth(?f ` j, [z,x,y])" if "j<2" for j + proof - + consider "j=0" | "j=1" using ltD[OF \j<2\] by auto + then show ?thesis + proof(cases) + case 1 + then show ?thesis using apply_equality f_type by simp + next + case 2 + then show ?thesis using apply_equality f_type by simp + qed + qed + show ?thesis + using sum_action[OF _ _ _ _ f_type id_type _ _ _ _ _ _ _ 2 1,simplified] assms + unfolding \_repl_def by simp +qed + +lemma Lambda_in_M : + assumes + f_fm: "\ \ formula" and + f_ar: "arity(\)\ 2 +\<^sub>\ length(env)" and + fsats: "\x y. x\A \ y\M \ (M,[x,y]@env \ \) \ is_f(x,y)" and + fabs: "\x y. x\A \ y\M \ is_f(x,y) \ y = f(x)" and + fclosed: "\x. x\A \ f(x) \ M" and + "A\M" "env\list(M)" and + phi'_replacement2: "replacement_assm(M,env@[A],Lambda_in_M_fm(\,length(env)))" + shows "(\x\A . f(x)) \M" + unfolding lam_def +proof - + let ?ren="\_repl(length(env))" + let ?j="2+\<^sub>\length(env)" + let ?k="3+\<^sub>\length(env)" + let ?\="ren(\)`?j`?k`?ren" + let ?\'="Exists(And(pair_fm(1,0,2),?\))" + let ?p="\x y. \z\M. pair(##M,x,z,y) \ is_f(x,z)" + have "?\'\formula" "?\\formula" + using \env\_\ length_type f_fm ren_type ren_tc[of \ "2+\<^sub>\length(env)" "3+\<^sub>\length(env)" ?ren] + by simp_all + moreover from this + have "arity(?\)\3+\<^sub>\(length(env))" "arity(?\)\nat" + using assms arity_ren[OF f_fm _ _ ren_type,of "length(env)"] by simp_all + then + have "arity(?\') \ 2+\<^sub>\(length(env))" + using Un_le pred_Un_distrib assms pred_le + by (simp add:arity) + moreover from this calculation + have "x\A \ y\M \ (M,[x,y]@env \ ?\') \ ?p(x,y)" for x y + using \env\_\ length_type[OF \env\_\] assms transitivity[OF _ \A\M\] + sats_iff_sats_ren[OF f_fm _ _ _ _ ren_type f_ar ren_action[rule_format,of _ x y],of _ M ] + by auto + moreover + have "x\A \ y\M \ ?p(x,y) \ y = " for x y + using assms transitivity[OF _ \A\_\] fclosed + by simp + moreover + have "\ x . x\A \ \ M" + using transitivity[OF _ \A\M\] pair_in_M_iff fclosed by simp + ultimately + show "{\x,f(x)\ . x\A } \ M" + using Replace_in_M[of ?\' env A] phi'_replacement2 \A\M\ \env\_\ + by simp +qed + +lemma ren_action' : + assumes + "env\list(M)" "x\M" "y\M" "z\M" "u\M" + shows "\ i . i < 3+\<^sub>\length(env) \ + nth(i,[x,z,u]@env) = nth(\_pair_repl(length(env))`i,[x,z,y,u]@env)" +proof - + let ?f="{\0, 0\, \1, 1\, \2,3\}" + have 1:"(\j. j < length(env) \ nth(j, env) = nth(id(length(env)) ` j, env))" + using assms ltD by simp + have 2:"nth(j, [x,z,u]) = nth(?f ` j, [x,z,y,u])" if "j<3" for j + proof - + consider "j=0" | "j=1" | "j=2" using ltD[OF \j<3\] by auto + then show ?thesis + proof(cases) + case 1 + then show ?thesis using apply_equality f_type' by simp + next + case 2 + then show ?thesis using apply_equality f_type' by simp + next + case 3 + then show ?thesis using apply_equality f_type' by simp + qed + qed + show ?thesis + using sum_action[OF _ _ _ _ f_type' id_type _ _ _ _ _ _ _ 2 1,simplified] assms + unfolding \_pair_repl_def by simp +qed + +lemma LambdaPair_in_M : + assumes + f_fm: "\ \ formula" and + f_ar: "arity(\)\ 3 +\<^sub>\ length(env)" and + fsats: "\x z r. x\M \ z\M \ r\M \ (M,[x,z,r]@env \ \) \ is_f(x,z,r)" and + fabs: "\x z r. x\M \ z\M \ r\M \ is_f(x,z,r) \ r = f(x,z)" and + fclosed: "\x z. x\M \ z\M \ f(x,z) \ M" and + "A\M" "env\list(M)" and + phi'_replacement3: "replacement_assm(M,env@[A],LambdaPair_in_M_fm(\,length(env)))" + shows "(\x\A . f(fst(x),snd(x))) \M" +proof - + let ?ren="\_pair_repl(length(env))" + let ?j="3+\<^sub>\length(env)" + let ?k="4+\<^sub>\length(env)" + let ?\="ren(\)`?j`?k`?ren" + let ?\'="Exists(Exists(And(fst_fm(2,0),(And(snd_fm(2,1),?\)))))" + let ?p="\x y. is_f(fst(x),snd(x),y)" + have "?\'\formula" "?\\formula" + using \env\_\ length_type f_fm ren_type' ren_tc[of \ ?j ?k ?ren] + by simp_all + moreover from this + have "arity(?\)\4+\<^sub>\(length(env))" "arity(?\)\nat" + using assms arity_ren[OF f_fm _ _ ren_type',of "length(env)"] by simp_all + moreover from calculation + have 1:"arity(?\') \ 2+\<^sub>\(length(env))" "?\'\formula" + using Un_le pred_Un_distrib assms pred_le + by (simp_all add:arity) + moreover from this calculation + have 2:"x\A \ y\M \ (M,[x,y]@env \ ?\') \ ?p(x,y)" for x y + using + sats_iff_sats_ren[OF f_fm _ _ _ _ ren_type' f_ar + ren_action'[rule_format,of _ "fst(x)" x "snd(x)" y],simplified] + \env\_\ length_type[OF \env\_\] transitivity[OF _ \A\M\] + fst_snd_closed pair_in_M_iff fsats[of "fst(x)" "snd(x)" y,symmetric] + fst_abs snd_abs + by auto + moreover from assms + have 3:"x\A \ y\M \ ?p(x,y) \ y = f(fst(x),snd(x))" for x y + using fclosed fst_snd_closed pair_in_M_iff fabs transitivity + by auto + moreover + have 4:"\ x . x\A \ \ M" "\ x . x\A \ f(fst(x),snd(x)) \ M" + using transitivity[OF _ \A\M\] pair_in_M_iff fclosed fst_snd_closed + by simp_all + ultimately + show ?thesis + using Lambda_in_M[unfolded Lambda_in_M_fm_def, of ?\', OF _ _ _ _ _ _ _ + phi'_replacement3[unfolded LambdaPair_in_M_fm_def]] + \env\_\ \A\_\ by simp + +qed + +lemma (in M_ZF1_trans) lam_replacement2_in_ctm : + assumes + f_fm: "\ \ formula" and + f_ar: "arity(\)\ 3 +\<^sub>\ length(env)" and + fsats: "\x z r. x\M \ z\M \ r\M \ (M,[x,z,r]@env \ \) \ is_f(x,z,r)" and + fabs: "\x z r. x\M \ z\M \ r\M \ is_f(x,z,r) \ r = f(x,z)" and + fclosed: "\x z. x\M \ z\M \ f(x,z) \ M" and + "env\list(M)" and + phi'_replacement3: "\A. A\M \ replacement_assm(M,env@[A],LambdaPair_in_M_fm(\,length(env)))" + shows "lam_replacement(##M , \x . f(fst(x),snd(x)))" + using + LambdaPair_in_M fabs + f_ar ord_simp_union transitivity assms fst_snd_closed + by (rule_tac lam_replacement_iff_lam_closed[THEN iffD2],simp_all) + +simple_rename "ren_U" src "[z1,x_P, x_leq, x_o, x_t, z2_c]" + tgt "[z2_c,z1,z,x_P, x_leq, x_o, x_t]" + +simple_rename "ren_V" src "[fz,x_P, x_leq, x_o,x_f, x_t, gz]" + tgt "[gz,fz,z,x_P, x_leq, x_o,x_f, x_t]" + +simple_rename "ren_V3" src "[fz,x_P, x_leq, x_o,x_f, gz, hz]" + tgt "[hz,gz,fz,z,x_P, x_leq, x_o,x_f]" + +lemma separation_sat_after_function_1: + assumes "[a,b,c,d]\list(M)" and "\\formula" and "arity(\) \ 6" + and + f_fm: "f_fm \ formula" and + f_ar: "arity(f_fm) \ 6" and + fsats: "\ fx x. fx\M \ x\M \ (M,[fx,x]@[a, b, c, d] \ f_fm) \ fx=f(x)" and + fclosed: "\x . x\M \ f(x) \ M" and + g_fm: "g_fm \ formula" and + g_ar: "arity(g_fm) \ 7" and + gsats: "\ gx fx x. gx\M \ fx\M \ x\M \ (M,[gx,fx,x]@[a, b, c, d] \ g_fm) \ gx=g(x)" and + gclosed: "\x . x\M \ g(x) \ M" + shows "separation(##M, \r. M, [f(r), a, b, c, d, g(r)] \ \)" +proof - + note types = assms(1-4) + let ?\="ren(\)`6`7`ren_U_fn" + let ?\'="Exists(And(f_fm,Exists(And(g_fm,?\))))" + let ?\="\z.[f(z), a, b, c, d, g(z)]" + let ?env="[a, b, c, d]" + let ?\="\z.[g(z),f(z),z]@?env" + note types + moreover from this + have "arity(\) \ 7" "?\\formula" + using ord_simp_union ren_tc ren_U_thm(2)[folded ren_U_fn_def] le_trans[of "arity(\)" 6] + by simp_all + moreover from calculation + have "arity(?\) \ 7" "?\'\formula" + using arity_ren ren_U_thm(2)[folded ren_U_fn_def] f_fm g_fm + by simp_all + moreover from calculation f_ar g_ar f_fm g_fm + have "arity(?\') \ 5" + using ord_simp_union pred_le arity_type + by (simp add:arity) + moreover from calculation fclosed gclosed + have 0:"(M, [f(z), a, b, c, d, g(z)] \ \) \ (M,?\(z)\ ?\)" if "(##M)(z)" for z + using sats_iff_sats_ren[of \ 6 7 _ _ "?\(z)"] + ren_U_thm(1)[where A=M,folded ren_U_fn_def] ren_U_thm(2)[folded ren_U_fn_def] that + by simp + moreover from calculation + have 1:"(M,?\(z)\ ?\) \ M,[z]@?env\?\'" if "(##M)(z)" for z + using that fsats[OF fclosed[of z],of z] gsats[of "g(z)" "f(z)" z] fclosed gclosed f_fm g_fm + proof(rule_tac iffI,simp,rule_tac rev_bexI[where x="f(z)"],simp,(auto)[1]) + assume "M, [z] @ [a, b, c, d] \ (\\\f_fm \ (\\\g_fm \ ren(\) ` 6 ` 7 ` ren_U_fn\\)\\)" + then + have "\xa\M. (M, [xa, z, a, b, c, d] \ f_fm) \ + (\x\M. (M, [x, xa, z, a, b, c, d] \ g_fm) \ + (M, [x, xa, z, a, b, c, d] \ ren(\) ` 6 ` 7 ` ren_U_fn))" + using that calculation by auto + then + obtain xa x where "x\M" "xa\M" "M, [xa, z, a, b, c, d] \ f_fm" + "(M, [x, xa, z, a, b, c, d] \ g_fm)" + "(M, [x, xa, z, a, b, c, d] \ ren(\) ` 6 ` 7 ` ren_U_fn)" + using that calculation by auto + moreover from this + have "xa=f(z)" "x=g(z)" using fsats[of xa] gsats[of x xa] that by simp_all + ultimately + show "M, [g(z), f(z), z] @ [a, b, c, d] \ ren(\) ` 6 ` 7 ` ren_U_fn" + by auto + qed + moreover from calculation + have "separation(##M, \z. (M,[z]@?env \ ?\'))" + using separation_ax + by simp_all + ultimately + show ?thesis + by(rule_tac separation_cong[THEN iffD2,OF iff_trans[OF 0 1]],clarify,force) +qed + +lemma separation_sat_after_function3: + assumes "[a, b, c, d]\list(M)" and "\\formula" and "arity(\) \ 7" + and + f_fm: "f_fm \ formula" and + f_ar: "arity(f_fm) \ 6" and + fsats: "\ fx x. fx\M \ x\M \ (M,[fx,x]@[a, b, c, d] \ f_fm) \ fx=f(x)" and + fclosed: "\x . x\M \ f(x) \ M" and + g_fm: "g_fm \ formula" and + g_ar: "arity(g_fm) \ 7" and + gsats: "\ gx fx x. gx\M \ fx\M \ x\M \ (M,[gx,fx,x]@[a, b, c, d] \ g_fm) \ gx=g(x)" and + gclosed: "\x . x\M \ g(x) \ M" and + h_fm: "h_fm \ formula" and + h_ar: "arity(h_fm) \ 8" and + hsats: "\ hx gx fx x. hx\M \ gx\M \ fx\M \ x\M \ (M,[hx,gx,fx,x]@[a, b, c, d] \ h_fm) \ hx=h(x)" and + hclosed: "\x . x\M \ h(x) \ M" + shows "separation(##M, \r. M, [f(r), a, b, c, d, g(r), h(r)] \ \)" +proof - + note types = assms(1-3) + let ?\="\" + let ?\="ren(?\)`7`8`ren_V3_fn" + let ?\'="Exists(And(f_fm,Exists(And(g_fm,Exists(And(h_fm,?\))))))" + let ?\="\z.[f(z), a, b, c, d,g(z), h(z)]" + let ?env="[a, b, c, d]" + let ?\="\z.[h(z),g(z),f(z),z]@?env" + note types + moreover from this + have "?\\formula" by simp + moreover from calculation + have "arity(?\) \ 9" "?\\formula" + using ord_simp_union ren_tc ren_V3_thm(2)[folded ren_V3_fn_def] le_trans[of "arity(\)" 7] + by simp_all + moreover from calculation + have "arity(?\) \ 8" "?\'\formula" + using arity_ren ren_V3_thm(2)[folded ren_V3_fn_def] f_fm g_fm h_fm + by (simp_all) + moreover from this f_ar g_ar f_fm g_fm h_fm h_ar \?\'\_\ + have "arity(?\') \ 5" + using ord_simp_union arity_type nat_into_Ord + by (simp add:arity,(rule_tac pred_le,simp,rule_tac Un_le,simp)+,simp_all add: \?\\_\) + moreover from calculation fclosed gclosed hclosed + have 0:"(M, ?\(z) \ ?\) \ (M,?\(z)\ ?\)" if "(##M)(z)" for z + using sats_iff_sats_ren[of ?\ 7 8 "?\(z)" M "?\(z)"] + ren_V3_thm(1)[where A=M,folded ren_V3_fn_def,simplified] ren_V3_thm(2)[folded ren_V3_fn_def] that + by simp + moreover from calculation + have 1:"(M,?\(z)\ ?\) \ M,[z]@?env\?\'" if "(##M)(z)" for z + using that fsats[OF fclosed[of z],of z] gsats[of "g(z)" "f(z)" z] + hsats[of "h(z)" "g(z)" "f(z)" z] + fclosed gclosed hclosed f_fm g_fm h_fm + apply(rule_tac iffI,simp,rule_tac rev_bexI[where x="f(z)"],simp) + apply(rule_tac conjI,simp,rule_tac rev_bexI[where x="g(z)"],simp) + apply(rule_tac conjI,simp,rule_tac rev_bexI[where x="h(z)"],simp,rule_tac conjI,simp,simp) + proof - + assume "M, [z] @ [a, b, c, d] \ (\\\f_fm \ (\\\g_fm \ (\\\h_fm \ ren(\) ` 7 ` 8 ` ren_V3_fn\\)\\)\\)" + with calculation that + have "\x\M. (M, [x, z, a, b, c, d] \ f_fm) \ + (\xa\M. (M, [xa, x, z, a, b, c, d] \ g_fm) \ (\xb\M. (M, [xb, xa, x, z, a, b, c, d] \ h_fm) \ (M, [xb, xa, x, z, a, b, c, d] \ ren(\) ` 7 ` 8 ` ren_V3_fn)))" + by auto + with calculation + obtain x where "x\M" "(M, [x, z, a, b, c, d] \ f_fm)" + "(\xa\M. (M, [xa, x, z, a, b, c, d] \ g_fm) \ (\xb\M. (M, [xb, xa, x, z, a, b, c, d] \ h_fm) \ (M, [xb, xa, x, z, a, b, c, d] \ ren(\) ` 7 ` 8 ` ren_V3_fn)))" + by force + moreover from this + have "x=f(z)" using fsats[of x] that by simp + moreover from calculation + obtain xa where "xa\M" "(M, [xa, x, z, a, b, c, d] \ g_fm)" + "(\xb\M. (M, [xb, xa, x, z, a, b, c, d] \ h_fm) \ (M, [xb, xa, x, z, a, b, c, d] \ ren(\) ` 7 ` 8 ` ren_V3_fn))" + by auto + moreover from calculation + have "xa=g(z)" using gsats[of xa x] that by simp + moreover from calculation + obtain xb where "xb\M" "(M, [xb, xa, x, z, a, b, c, d] \ h_fm)" + "(M, [xb, xa, x, z, a, b, c, d] \ ren(\) ` 7 ` 8 ` ren_V3_fn)" + by auto + moreover from calculation + have "xb=h(z)" using hsats[of xb xa x] that by simp + ultimately + show "M, [h(z), g(z), f(z), z] @ [a, b, c, d] \ ren(\) ` 7 ` 8 ` ren_V3_fn" + by auto + qed + moreover from calculation \?\'\_\ + have "separation(##M, \z. (M,[z]@?env \ ?\'))" + using separation_ax + by simp + ultimately + show ?thesis + by(rule_tac separation_cong[THEN iffD2,OF iff_trans[OF 0 1]],clarify,force) +qed + +lemma separation_sat_after_function: + assumes "[a, b, c, d, \]\list(M)" and "\\formula" and "arity(\) \ 7" + and + f_fm: "f_fm \ formula" and + f_ar: "arity(f_fm) \ 7" and + fsats: "\ fx x. fx\M \ x\M \ (M,[fx,x]@[a, b, c, d, \] \ f_fm) \ fx=f(x)" and + fclosed: "\x . x\M \ f(x) \ M" and + g_fm: "g_fm \ formula" and + g_ar: "arity(g_fm) \ 8" and + gsats: "\ gx fx x. gx\M \ fx\M \ x\M \ (M,[gx,fx,x]@[a, b, c, d, \] \ g_fm) \ gx=g(x)" and + gclosed: "\x . x\M \ g(x) \ M" + shows "separation(##M, \r. M, [f(r), a, b, c, d, \, g(r)] \ \)" +proof - + note types = assms(1-3) + let ?\="\" + let ?\="ren(?\)`7`8`ren_V_fn" + let ?\'="Exists(And(f_fm,Exists(And(g_fm,?\))))" + let ?\="\z.[f(z), a, b, c, d, \, g(z)]" + let ?env="[a, b, c, d, \]" + let ?\="\z.[g(z),f(z),z]@?env" + note types + moreover from this + have "?\\formula" by simp + moreover from calculation + have "arity(?\) \ 8" "?\\formula" + using ord_simp_union ren_tc ren_V_thm(2)[folded ren_V_fn_def] le_trans[of "arity(\)" 7] + by simp_all + moreover from calculation + have "arity(?\) \ 8" "?\'\formula" + using arity_ren ren_V_thm(2)[folded ren_V_fn_def] f_fm g_fm + by (simp_all) + moreover from calculation f_ar g_ar f_fm g_fm + have "arity(?\') \ 6" + using ord_simp_union pred_le arity_type + by (simp add:arity) + moreover from calculation fclosed gclosed + have 0:"(M, ?\(z) \ ?\) \ (M,?\(z)\ ?\)" if "(##M)(z)" for z + using sats_iff_sats_ren[of ?\ 7 8 "?\(z)" _ "?\(z)"] + ren_V_thm(1)[where A=M,folded ren_V_fn_def] ren_V_thm(2)[folded ren_V_fn_def] that + by simp + moreover from calculation + have 1:"(M,?\(z)\ ?\) \ M,[z]@?env\?\'" if "(##M)(z)" for z + using that fsats[OF fclosed[of z],of z] gsats[of "g(z)" "f(z)" z] + fclosed gclosed f_fm g_fm + apply(rule_tac iffI,simp,rule_tac rev_bexI[where x="f(z)"],simp) + apply(auto)[1] + proof - + assume "M, [z] @ [a, b, c, d, \] \ (\\\f_fm \ (\\\g_fm \ ren(\) ` 7 ` 8 ` ren_V_fn\\)\\)" + then have "\xa\M. (M, [xa, z, a, b, c, d, \] \ f_fm) \ + (\x\M. (M, [x, xa, z, a, b, c, d, \] \ g_fm) \ (M, [x, xa, z, a, b, c, d, \] \ ren(\) ` 7 ` 8 ` ren_V_fn))" + using that calculation by auto + then + obtain xa where "xa\M" "M, [xa, z, a, b, c, d, \] \ f_fm" + "(\x\M. (M, [x, xa, z, a, b, c, d, \] \ g_fm) \ (M, [x, xa, z, a, b, c, d, \] \ ren(\) ` 7 ` 8 ` ren_V_fn))" + by auto + moreover from this + have "xa=f(z)" using fsats[of xa] that by simp + moreover from calculation + obtain x where "x\M" "M, [x, xa, z, a, b, c, d, \] \ g_fm" "M, [x, xa, z, a, b, c, d, \] \ ren(\) ` 7 ` 8 ` ren_V_fn" + by auto + moreover from calculation + have "x=g(z)" using gsats[of x xa] that by simp + ultimately + show "M, [g(z), f(z), z] @ [a, b, c, d, \] \ ren(\) ` 7 ` 8 ` ren_V_fn" + by auto + qed + moreover from calculation + have "separation(##M, \z. (M,[z]@?env \ ?\'))" + using separation_ax + by simp_all + ultimately + show ?thesis + by(rule_tac separation_cong[THEN iffD2,OF iff_trans[OF 0 1]],clarify,force) +qed + +end \ \\<^locale>\M_Z_trans\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Internal_ZFC_Axioms.thy b/thys/Independence_CH/Internal_ZFC_Axioms.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Internal_ZFC_Axioms.thy @@ -0,0 +1,520 @@ +section\The ZFC axioms, internalized\ +theory Internal_ZFC_Axioms + imports + Forcing_Data + +begin + +schematic_goal ZF_union_auto: + "Union_ax(##A) \ (A, [] \ ?zfunion)" + unfolding Union_ax_def + by ((rule sep_rules | simp)+) + +synthesize "ZF_union" from_schematic ZF_union_auto +notation ZF_union_fm (\\Union Ax\\) + +schematic_goal ZF_power_auto: + "power_ax(##A) \ (A, [] \ ?zfpow)" + unfolding power_ax_def powerset_def subset_def + by ((rule sep_rules | simp)+) + +synthesize "ZF_power" from_schematic ZF_power_auto +notation ZF_power_fm (\\Powerset Ax\\) + +schematic_goal ZF_pairing_auto: + "upair_ax(##A) \ (A, [] \ ?zfpair)" + unfolding upair_ax_def + by ((rule sep_rules | simp)+) + +synthesize "ZF_pairing" from_schematic ZF_pairing_auto +notation ZF_pairing_fm (\\Pairing\\) + +schematic_goal ZF_foundation_auto: + "foundation_ax(##A) \ (A, [] \ ?zffound)" + unfolding foundation_ax_def + by ((rule sep_rules | simp)+) + +synthesize "ZF_foundation" from_schematic ZF_foundation_auto +notation ZF_foundation_fm (\\Foundation\\) + +schematic_goal ZF_extensionality_auto: + "extensionality(##A) \ (A, [] \ ?zfext)" + unfolding extensionality_def + by ((rule sep_rules | simp)+) + +synthesize "ZF_extensionality" from_schematic ZF_extensionality_auto +notation ZF_extensionality_fm (\\Extensionality\\) + +schematic_goal ZF_infinity_auto: + "infinity_ax(##A) \ (A, [] \ (?\(i,j,h)))" + unfolding infinity_ax_def + by ((rule sep_rules | simp)+) + +synthesize "ZF_infinity" from_schematic ZF_infinity_auto +notation ZF_infinity_fm (\\Infinity\\) + +schematic_goal ZF_choice_auto: + "choice_ax(##A) \ (A, [] \ (?\(i,j,h)))" + unfolding choice_ax_def + by ((rule sep_rules | simp)+) + +synthesize "ZF_choice" from_schematic ZF_choice_auto +notation ZF_choice_fm (\\AC\\) + +lemmas ZFC_fm_defs = ZF_extensionality_fm_def ZF_foundation_fm_def ZF_pairing_fm_def + ZF_union_fm_def ZF_infinity_fm_def ZF_power_fm_def ZF_choice_fm_def + +lemmas ZFC_fm_sats = ZF_extensionality_auto ZF_foundation_auto ZF_pairing_auto + ZF_union_auto ZF_infinity_auto ZF_power_auto ZF_choice_auto + +definition + ZF_fin :: "i" where + "ZF_fin \ {\Extensionality\, \Foundation\, \Pairing\, + \Union Ax\, \Infinity\, \Powerset Ax\}" + +subsection\The Axiom of Separation, internalized\ +lemma iterates_Forall_type [TC]: + "\ n \ nat; p \ formula \ \ Forall^n(p) \ formula" + by (induct set:nat, auto) + +lemma last_init_eq : + assumes "l \ list(A)" "length(l) = succ(n)" + shows "\ a\A. \l'\list(A). l = l'@[a]" +proof- + from \l\_\ \length(_) = _\ + have "rev(l) \ list(A)" "length(rev(l)) = succ(n)" + by simp_all + then + obtain a l' where "a\A" "l'\list(A)" "rev(l) = Cons(a,l')" + by (cases;simp) + then + have "l = rev(l') @ [a]" "rev(l') \ list(A)" + using rev_rev_ident[OF \l\_\] by auto + with \a\_\ + show ?thesis by blast +qed + +lemma take_drop_eq : + assumes "l\list(M)" + shows "\ n . n < succ(length(l)) \ l = take(n,l) @ drop(n,l)" + using \l\list(M)\ +proof induct + case Nil + then show ?case by auto +next + case (Cons a l) + then show ?case + proof - + { + fix i + assume "il\list(M)\ + consider (lt) "i = 0" | (eq) "\k\nat. i = succ(k) \ k < succ(length(l))" + using \l\list(M)\ le_natI nat_imp_quasinat + by (cases rule:nat_cases[of i];auto) + then + have "take(i,Cons(a,l)) @ drop(i,Cons(a,l)) = Cons(a,l)" + using Cons + by (cases;auto) + } + then show ?thesis using Cons by auto + qed +qed + +lemma list_split : +assumes "n \ succ(length(rest))" "rest \ list(M)" +shows "\re\list(M). \st\list(M). rest = re @ st \ length(re) = pred(n)" +proof - + from assms + have "pred(n) \ length(rest)" + using pred_mono[OF _ \n\_\] pred_succ_eq by auto + with \rest\_\ + have "pred(n)\nat" "rest = take(pred(n),rest) @ drop(pred(n),rest)" (is "_ = ?re @ ?st") + using take_drop_eq[OF \rest\_\] le_natI by auto + then + have "length(?re) = pred(n)" "?re\list(M)" "?st\list(M)" + using length_take[rule_format,OF _ \pred(n)\_\] \pred(n) \ _\ \rest\_\ + unfolding min_def + by auto + then + show ?thesis + using rev_bexI[of _ _ "\ re. \st\list(M). rest = re @ st \ length(re) = pred(n)"] + \length(?re) = _\ \rest = _\ + by auto +qed + +lemma sats_nForall: + assumes + "\ \ formula" + shows + "n\nat \ ms \ list(M) \ + (M, ms \ (Forall^n(\))) \ + (\rest \ list(M). length(rest) = n \ M, rest @ ms \ \)" +proof (induct n arbitrary:ms set:nat) + case 0 + with assms + show ?case by simp +next + case (succ n) + have "(\rest\list(M). length(rest) = succ(n) \ P(rest,n)) \ + (\t\M. \res\list(M). length(res) = n \ P(res @ [t],n))" + if "n\nat" for n P + using that last_init_eq by force + from this[of _ "\rest _. (M, rest @ ms \ \)"] \n\nat\ + have "(\rest\list(M). length(rest) = succ(n) \ M, rest @ ms \ \) \ + (\t\M. \res\list(M). length(res) = n \ M, (res @ [t]) @ ms \ \)" + by simp + with assms succ(1,3) succ(2)[of "Cons(_,ms)"] + show ?case + using arity_sats_iff[of \ _ M "Cons(_, ms @ _)"] app_assoc + by (simp) +qed + +definition + sep_body_fm :: "i \ i" where + "sep_body_fm(p) \ (\\(\\(\\\\0 \ 1\ \ \\0 \ 2\ \ incr_bv1^2 (p) \\\)\)\)" + +lemma sep_body_fm_type [TC]: "p \ formula \ sep_body_fm(p) \ formula" + by (simp add: sep_body_fm_def) + +lemma sats_sep_body_fm: + assumes + "\ \ formula" "ms\list(M)" "rest\list(M)" + shows + "(M, rest @ ms \ sep_body_fm(\)) \ + separation(##M,\x. M, [x] @ rest @ ms \ \)" + using assms formula_add_params1[of _ 2 _ _ "[_,_]" ] + unfolding sep_body_fm_def separation_def by simp + +definition + ZF_separation_fm :: "i \ i" (\\Separation'(_')\\) where + "ZF_separation_fm(p) \ Forall^(pred(arity(p)))(sep_body_fm(p))" + +lemma ZF_separation_fm_type [TC]: "p \ formula \ ZF_separation_fm(p) \ formula" + by (simp add: ZF_separation_fm_def) + +lemma sats_ZF_separation_fm_iff: + assumes + "\\formula" + shows + "(M, [] \ \Separation(\)\) + \ + (\env\list(M). arity(\) \ 1 +\<^sub>\ length(env) \ + separation(##M,\x. M, [x] @ env \ \))" +proof (intro iffI ballI impI) + let ?n="pred(arity(\))" + fix env + assume "M, [] \ ZF_separation_fm(\)" + assume "arity(\) \ 1 +\<^sub>\ length(env)" "env\list(M)" + moreover from this + have "arity(\) \ succ(length(env))" by simp + then + obtain some rest where "some\list(M)" "rest\list(M)" + "env = some @ rest" "length(some) = pred(arity(\))" + using list_split[OF \arity(\) \ succ(_)\ \env\_\] by force + moreover from \\\_\ + have "arity(\) \ succ(pred(arity(\)))" + using succpred_leI by simp + moreover + note assms + moreover + assume "M, [] \ ZF_separation_fm(\)" + moreover from calculation + have "M, some \ sep_body_fm(\)" + using sats_nForall[of "sep_body_fm(\)" ?n] + unfolding ZF_separation_fm_def by simp + ultimately + show "separation(##M, \x. M, [x] @ env \ \)" + unfolding ZF_separation_fm_def + using sats_sep_body_fm[of \ "[]" M some] + arity_sats_iff[of \ rest M "[_] @ some"] + separation_cong[of "##M" "\x. M, Cons(x, some @ rest) \ \" _ ] + by simp +next \ \almost equal to the previous implication\ + let ?n="pred(arity(\))" + assume asm:"\env\list(M). arity(\) \ 1 +\<^sub>\ length(env) \ + separation(##M, \x. M, [x] @ env \ \)" + { + fix some + assume "some\list(M)" "length(some) = pred(arity(\))" + moreover + note \\\_\ + moreover from calculation + have "arity(\) \ 1 +\<^sub>\ length(some)" + using le_trans[OF succpred_leI] succpred_leI by simp + moreover from calculation and asm + have "separation(##M, \x. M, [x] @ some \ \)" by blast + ultimately + have "M, some \ sep_body_fm(\)" + using sats_sep_body_fm[of \ "[]" M some] + arity_sats_iff[of \ _ M "[_,_] @ some"] + strong_replacement_cong[of "##M" "\x y. M, Cons(x, Cons(y, some @ _)) \ \" _ ] + by simp + } + with \\\_\ + show "M, [] \ ZF_separation_fm(\)" + using sats_nForall[of "sep_body_fm(\)" ?n] + unfolding ZF_separation_fm_def + by simp +qed + +subsection\The Axiom of Replacement, internalized\ +schematic_goal sats_univalent_fm_auto: + assumes + (* Q_iff_sats:"\a b z env aa bb. nth(a,Cons(z,env)) = aa \ nth(b,Cons(z,env)) = bb \ z\A + \ aa \ A \ bb \ A \ env\ list(A) \ + Q(aa,bb) \ (A, Cons(z,env) \ (Q_fm(a,b)))" \ \using only \ formula\ *) + Q_iff_sats:"\x y z. x \ A \ y \ A \ z\A \ + Q(x,z) \ (A,Cons(z,Cons(y,Cons(x,env))) \ Q1_fm)" + "\x y z. x \ A \ y \ A \ z\A \ + Q(x,y) \ (A,Cons(z,Cons(y,Cons(x,env))) \ Q2_fm)" + and + asms: "nth(i,env) = B" "i \ nat" "env \ list(A)" + shows + "univalent(##A,B,Q) \ A,env \ ?ufm(i)" + unfolding univalent_def + by (insert asms; (rule sep_rules Q_iff_sats | simp)+) + +synthesize_notc "univalent" from_schematic sats_univalent_fm_auto + +lemma univalent_fm_type [TC]: "q1\ formula \ q2\formula \ i\nat \ + univalent_fm(q2,q1,i) \formula" + by (simp add:univalent_fm_def) + +lemma sats_univalent_fm : + assumes + Q_iff_sats:"\x y z. x \ A \ y \ A \ z\A \ + Q(x,z) \ (A,Cons(z,Cons(y,Cons(x,env))) \ Q1_fm)" + "\x y z. x \ A \ y \ A \ z\A \ + Q(x,y) \ (A,Cons(z,Cons(y,Cons(x,env))) \ Q2_fm)" + and + asms: "nth(i,env) = B" "i \ nat" "env \ list(A)" + shows + "(A,env \ univalent_fm(Q1_fm,Q2_fm,i)) \ univalent(##A,B,Q)" + unfolding univalent_fm_def using asms sats_univalent_fm_auto[OF Q_iff_sats] by simp + +definition + swap_vars :: "i\i" where + "swap_vars(\) \ + Exists(Exists(And(Equal(0,3),And(Equal(1,2),iterates(\p. incr_bv(p)`2 , 2, \)))))" + +lemma swap_vars_type[TC] : + "\\formula \ swap_vars(\) \formula" + unfolding swap_vars_def by simp + +lemma sats_swap_vars : + "[x,y] @ env \ list(M) \ \\formula \ + (M, [x,y] @ env \ swap_vars(\)) \ M,[y,x] @ env \ \" + unfolding swap_vars_def + using sats_incr_bv_iff [of _ _ M _ "[y,x]"] by simp + +definition + univalent_Q1 :: "i \ i" where + "univalent_Q1(\) \ incr_bv1(swap_vars(\))" + +definition + univalent_Q2 :: "i \ i" where + "univalent_Q2(\) \ incr_bv(swap_vars(\))`0" + +lemma univalent_Qs_type [TC]: + assumes "\\formula" + shows "univalent_Q1(\) \ formula" "univalent_Q2(\) \ formula" + unfolding univalent_Q1_def univalent_Q2_def using assms by simp_all + +lemma sats_univalent_fm_assm: + assumes + "x \ A" "y \ A" "z\A" "env\ list(A)" "\ \ formula" + shows + "(A, ([x,z] @ env) \ \) \ (A, Cons(z,Cons(y,Cons(x,env))) \ (univalent_Q1(\)))" + "(A, ([x,y] @ env) \ \) \ (A, Cons(z,Cons(y,Cons(x,env))) \ (univalent_Q2(\)))" + unfolding univalent_Q1_def univalent_Q2_def + using + sats_incr_bv_iff[of _ _ A _ "[]"] \ \simplifies iterates of \<^term>\\x. incr_bv(x)`0\\ + sats_incr_bv1_iff[of _ "Cons(x,env)" A z y] + sats_swap_vars assms + by simp_all + +definition + rep_body_fm :: "i \ i" where + "rep_body_fm(p) \ Forall(Implies( + univalent_fm(univalent_Q1(incr_bv(p)`2),univalent_Q2(incr_bv(p)`2),0), + Exists(Forall( + Iff(Member(0,1),Exists(And(Member(0,3),incr_bv(incr_bv(p)`2)`2)))))))" + +lemma rep_body_fm_type [TC]: "p \ formula \ rep_body_fm(p) \ formula" + by (simp add: rep_body_fm_def) + +lemmas ZF_replacement_simps = formula_add_params1[of \ 2 _ M "[_,_]" ] + sats_incr_bv_iff[of _ _ M _ "[]"] \ \simplifies iterates of \<^term>\\x. incr_bv(x)`0\\ + sats_incr_bv_iff[of _ _ M _ "[_,_]"]\ \simplifies \<^term>\\x. incr_bv(x)`2\\ + sats_incr_bv1_iff[of _ _ M] sats_swap_vars for \ M + +lemma sats_rep_body_fm: + assumes + "\ \ formula" "ms\list(M)" "rest\list(M)" + shows + "(M, rest @ ms \ rep_body_fm(\)) \ + strong_replacement(##M,\x y. M, [x,y] @ rest @ ms \ \)" + using assms ZF_replacement_simps + unfolding rep_body_fm_def strong_replacement_def univalent_def + unfolding univalent_fm_def univalent_Q1_def univalent_Q2_def + by simp + +definition + ZF_replacement_fm :: "i \ i" (\\Replacement'(_')\\) where + "ZF_replacement_fm(p) \ Forall^(pred(pred(arity(p))))(rep_body_fm(p))" + +lemma ZF_replacement_fm_type [TC]: "p \ formula \ ZF_replacement_fm(p) \ formula" + by (simp add: ZF_replacement_fm_def) + +lemma sats_ZF_replacement_fm_iff: + assumes + "\\formula" + shows + "(M, [] \ \Replacement(\)\) + \ + (\env\list(M). arity(\) \ 2 +\<^sub>\ length(env) \ + strong_replacement(##M,\x y. M,[x,y] @ env \ \))" +proof (intro iffI ballI impI) + let ?n="pred(pred(arity(\)))" + fix env + assume "M, [] \ ZF_replacement_fm(\)" "arity(\) \ 2 +\<^sub>\ length(env)" "env\list(M)" + moreover from this + have "arity(\) \ succ(succ(length(env)))" by (simp) + moreover from calculation + have "pred(arity(\)) \ succ(length(env))" + using pred_mono[OF _ \arity(\)\succ(_)\] pred_succ_eq by simp + moreover from calculation + obtain some rest where "some\list(M)" "rest\list(M)" + "env = some @ rest" "length(some) = pred(pred(arity(\)))" + using list_split[OF \pred(_) \ _\ \env\_\] by auto + moreover + note \\\_\ + moreover from this + have "arity(\) \ succ(succ(pred(pred(arity(\)))))" + using le_trans[OF succpred_leI] succpred_leI by simp + moreover from calculation + have "M, some \ rep_body_fm(\)" + using sats_nForall[of "rep_body_fm(\)" ?n] + unfolding ZF_replacement_fm_def + by simp + ultimately + show "strong_replacement(##M, \x y. M, [x, y] @ env \ \)" + using sats_rep_body_fm[of \ "[]" M some] + arity_sats_iff[of \ rest M "[_,_] @ some"] + strong_replacement_cong[of "##M" "\x y. M, Cons(x, Cons(y, some @ rest)) \ \" _ ] + by simp +next \ \almost equal to the previous implication\ + let ?n="pred(pred(arity(\)))" + assume asm:"\env\list(M). arity(\) \ 2 +\<^sub>\ length(env) \ + strong_replacement(##M, \x y. M, [x, y] @ env \ \)" + { + fix some + assume "some\list(M)" "length(some) = pred(pred(arity(\)))" + moreover + note \\\_\ + moreover from calculation + have "arity(\) \ 2 +\<^sub>\ length(some)" + using le_trans[OF succpred_leI] succpred_leI by simp + moreover from calculation and asm + have "strong_replacement(##M, \x y. M, [x, y] @ some \ \)" by blast + ultimately + have "M, some \ rep_body_fm(\)" + using sats_rep_body_fm[of \ "[]" M some] + arity_sats_iff[of \ _ M "[_,_] @ some"] + strong_replacement_cong[of "##M" "\x y. M, Cons(x, Cons(y, some @ _)) \ \" _ ] + by simp + } + with \\\_\ + show "M, [] \ ZF_replacement_fm(\)" + using sats_nForall[of "rep_body_fm(\)" ?n] + unfolding ZF_replacement_fm_def + by simp +qed + +definition + ZF_schemes :: "i" where + "ZF_schemes \ {\Separation(p)\ . p \ formula } \ {\Replacement(p)\ . p \ formula }" + +lemma Un_subset_formula [TC]: "A\formula \ B\formula \ A\B \ formula" + by auto + +lemma ZF_schemes_subset_formula [TC]: "ZF_schemes \ formula" + unfolding ZF_schemes_def by auto + +lemma ZF_fin_subset_formula [TC]: "ZF_fin \ formula" + unfolding ZF_fin_def by simp + +definition + ZF :: "i" where + "ZF \ ZF_schemes \ ZF_fin" + +lemma ZF_subset_formula [TC]: "ZF \ formula" + unfolding ZF_def by auto + +definition + ZFC :: "i" where + "ZFC \ ZF \ {\AC\}" + +definition + ZF_minus_P :: "i" where + "ZF_minus_P \ ZF - { \Powerset Ax\ }" + +definition + Zermelo_fms :: "i" (\\Z\\) where + "Zermelo_fms \ ZF_fin \ {\Separation(p)\ . p \ formula }" + +definition + ZC :: "i" where + "ZC \ Zermelo_fms \ {\AC\}" + +lemma ZFC_subset_formula: "ZFC \ formula" + by (simp add:ZFC_def Un_subset_formula) + +txt\Satisfaction of a set of sentences\ +definition + satT :: "[i,i] \ o" ("_ \ _" [36,36] 60) where + "A \ \ \ \\\\. (A,[] \ \)" + +lemma satTI [intro!]: + assumes "\\. \\\ \ A,[] \ \" + shows "A \ \" + using assms unfolding satT_def by simp + +lemma satTD [dest] :"A \ \ \ \\\ \ A,[] \ \" + unfolding satT_def by simp + +lemma satT_mono: "A \ \ \ \ \ \ \ A \ \" + by blast + +lemma satT_Un_iff: "M \ \ \ \ \ M \ \ \ M \ \" by auto + +lemma sats_ZFC_iff_sats_ZF_AC: + "(N \ ZFC) \ (N \ ZF) \ (N, [] \ \AC\)" + unfolding ZFC_def ZF_def by auto + +lemma satT_ZF_imp_satT_Z: "M \ ZF \ M \ \Z\" + unfolding ZF_def ZF_schemes_def Zermelo_fms_def ZF_fin_def by auto + +lemma satT_ZFC_imp_satT_ZC: "M \ ZFC \ M \ ZC" + unfolding ZFC_def ZF_def ZF_schemes_def ZC_def Zermelo_fms_def by auto + +lemma satT_Z_ZF_replacement_imp_satT_ZF: "N \ \Z\ \ N \ {\Replacement(x)\ . x \ formula} \ N \ ZF" + unfolding ZF_def ZF_schemes_def Zermelo_fms_def ZF_fin_def by auto + +lemma satT_ZC_ZF_replacement_imp_satT_ZFC: "N \ ZC \ N \ {\Replacement(x)\ . x \ formula} \ N \ ZFC" + unfolding ZFC_def ZF_def ZF_schemes_def ZC_def Zermelo_fms_def by auto + +lemma ground_repl_fm_sub_ZF: "{\Replacement(ground_repl_fm(\))\ . \ \ formula} \ ZF" + unfolding ZF_def ZF_schemes_def by auto + +lemma ZF_replacement_fms_sub_ZFC: "{\Replacement(\)\ . \ \ formula} \ ZFC" + unfolding ZFC_def ZF_def ZF_schemes_def by auto + +lemma ground_repl_fm_sub_ZFC: "{\Replacement(ground_repl_fm(\))\ . \ \ formula} \ ZFC" + unfolding ZFC_def ZF_def ZF_schemes_def by auto + +lemma ZF_replacement_ground_repl_fm_type: "{\Replacement(ground_repl_fm(\))\ . \ \ formula} \ formula" + by auto + +end diff --git a/thys/Independence_CH/Kappa_Closed_Notions.thy b/thys/Independence_CH/Kappa_Closed_Notions.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Kappa_Closed_Notions.thy @@ -0,0 +1,723 @@ +section\Preservation results for $\kappa$-closed forcing notions\ + +theory Kappa_Closed_Notions + imports + Not_CH +begin + +definition + lerel :: "i\i" where + "lerel(\) \ Memrel(\) \ id(\)" + +lemma lerelI[intro!]: "x\y \ y\\ \ Ord(\) \ \x,y\ \ lerel(\)" + using Ord_trans[of x y \] ltD unfolding lerel_def by auto + +lemma lerelD[dest]: "\x,y\ \ lerel(\) \ Ord(\) \ x\y" + using ltI[THEN leI] Ord_in_Ord unfolding lerel_def by auto + +definition + mono_seqspace :: "[i,i,i] \ i" (\_ \<^sub><\ '(_,_')\ [61] 60) where + "\ \<^sub><\ (P,leq) \ mono_map(\,Memrel(\),P,leq)" + +relativize functional "mono_seqspace" "mono_seqspace_rel" +relationalize "mono_seqspace_rel" "is_mono_seqspace" +synthesize "is_mono_seqspace" from_definition assuming "nonempty" + +context M_ZF_library +begin + +rel_closed for "mono_seqspace" + unfolding mono_seqspace_rel_def mono_map_rel_def + using separation_closed separation_ball separation_imp separation_in + lam_replacement_fst lam_replacement_snd lam_replacement_hcomp lam_replacement_constant + lam_replacement_Pair[THEN[5] lam_replacement_hcomp2] + lam_replacement_apply2[THEN[5] lam_replacement_hcomp2] + by simp_all + +end \ \\<^locale>\M_ZF_library\\ + +abbreviation + mono_seqspace_r (\_ \<^sub><\\<^bsup>_\<^esup> '(_,_')\ [61] 60) where + "\ \<^sub><\\<^bsup>M\<^esup> (P,leq) \ mono_seqspace_rel(M,\,P,leq)" + +abbreviation + mono_seqspace_r_set (\_ \<^sub><\\<^bsup>_\<^esup> '(_,_')\ [61] 60) where + "\ \<^sub><\\<^bsup>M\<^esup> (P,leq) \ mono_seqspace_rel(##M,\,P,leq)" + +lemma mono_seqspaceI[intro!]: + includes mono_map_rules + assumes "f: A\P" "\x y. x\A \ y\A \ x \f`x, f`y\ \ leq" "Ord(A)" + shows "f: A \<^sub><\ (P,leq)" + using ltI[OF _ Ord_in_Ord[of A], THEN [3] assms(2)] assms(1,3) + unfolding mono_seqspace_def by auto + +lemma (in M_ZF_library) mono_seqspace_rel_char: + assumes "M(A)" "M(P)" "M(leq)" + shows "A \<^sub><\\<^bsup>M\<^esup> (P,leq) = {f\A \<^sub><\ (P,leq). M(f)}" + using assms mono_map_rel_char + unfolding mono_seqspace_def mono_seqspace_rel_def by simp + +lemma (in M_ZF_library) mono_seqspace_relI[intro!]: + assumes "f: A\\<^bsup>M\<^esup> P" "\x y. x\A \ y\A \ x \f`x, f`y\ \ leq" + "Ord(A)" "M(A)" "M(P)" "M(leq)" + shows "f: A \<^sub><\\<^bsup>M\<^esup> (P,leq)" + using mono_seqspace_rel_char function_space_rel_char assms by auto + +lemma mono_seqspace_is_fun[dest]: + includes mono_map_rules + shows "j: A \<^sub><\ (P,leq) \ j: A\ P" + unfolding mono_seqspace_def by auto + +lemma mono_map_lt_le_is_mono[dest]: + includes mono_map_rules + assumes "j: A \<^sub><\ (P,leq)" "a\A" "c\A" "a\c" "Ord(A)" "refl(P,leq)" + shows "\j`a,j`c\ \ leq" + using assms mono_map_increasing unfolding mono_seqspace_def refl_def + by (cases "a=c") (auto dest:ltD) + +lemma (in M_ZF_library) mem_mono_seqspace_abs[absolut]: + assumes "M(f)" "M(A)" "M(P)" "M(leq)" + shows "f:A \<^sub><\\<^bsup>M\<^esup> (P,leq) \ f: A \<^sub><\ (P,leq)" + using assms mono_map_rel_char unfolding mono_seqspace_def mono_seqspace_rel_def + by (simp) + +definition + mono_map_lt_le :: "[i,i] \ i" (infixr \\<^sub><\\<^sub>\\ 60) where + "\ \<^sub><\\<^sub>\ \ \ \ \<^sub><\ (\,lerel(\))" + +lemma mono_map_lt_leI[intro!]: + includes mono_map_rules + assumes "f: A\B" "\x y. x\A \ y\A \ x f`x \ f`y" "Ord(A)" "Ord(B)" + shows "f: A \<^sub><\\<^sub>\ B" + using assms + unfolding mono_map_lt_le_def by auto + +\ \Kunen IV.7.13, with “$\kappa$” in place of “$\lambda$”\ +definition + kappa_closed :: "[i,i,i] \ o" (\_-closed'(_,_')\) where + "\-closed(P,leq) \ \\. \<\ \ (\f\\ \<^sub><\ (P,converse(leq)). \q\P. \\\\. \q,f`\\\leq)" + +relativize functional "kappa_closed" "kappa_closed_rel" +relationalize "kappa_closed_rel" "is_kappa_closed" +synthesize "is_kappa_closed" from_definition assuming "nonempty" + +abbreviation + kappa_closed_r (\_-closed\<^bsup>_\<^esup>'(_,_')\ [61] 60) where + "\-closed\<^bsup>M\<^esup>(P,leq) \ kappa_closed_rel(M,\,P,leq)" + +abbreviation + kappa_closed_r_set (\_-closed\<^bsup>_\<^esup>'(_,_')\ [61] 60) where + "\-closed\<^bsup>M\<^esup>(P,leq) \ kappa_closed_rel(##M,\,P,leq)" + +lemma (in forcing_data4) forcing_a_value: + assumes "p \ \0:1\2\ [f_dot, A\<^sup>v, B\<^sup>v]" "a \ A" + "q \ p" "q \ P" "p\P" "f_dot \ M" "A\M" "B\M" + shows "\d\P. \b\B. d \ q \ d \ \0`1 is 2\ [f_dot, a\<^sup>v, b\<^sup>v]" + \ \Old neater version, but harder to use + (without the assumptions on \<^term>\q\):\ + (* "dense_below({q \ P. \b\B. q \ \0`1 is 2\ [f_dot, a\<^sup>v, b\<^sup>v]}, p)" *) +proof - + from assms + have "q \ \0:1\2\ [f_dot, A\<^sup>v, B\<^sup>v]" + using strengthening_lemma[of p "\0:1\2\" q "[f_dot, A\<^sup>v, B\<^sup>v]"] + typed_function_type arity_typed_function_fm + by (auto simp: union_abs2 union_abs1 check_in_M P_in_M) + from \a\A\ \A\M\ + have "a\M" by (auto dest:transM) + from \q\P\ + text\Here we're using countability (via the existence of generic filters) + of \<^term>\M\ as a shortcut, to avoid a further density argument.\ + obtain G where "M_generic(G)" "q\G" + using generic_filter_existence by blast + then + interpret G_generic4_AC _ _ _ _ _ G by unfold_locales + include G_generic1_lemmas + note \q\G\ + moreover + note \q \ \0:1\2\ [f_dot, A\<^sup>v, B\<^sup>v]\ \M_generic(G)\ + moreover + note \q\P\ \f_dot\M\ \B\M\ \A\M\ + moreover from this + have "map(val(P, G), [f_dot, A\<^sup>v, B\<^sup>v]) \ list(M[G])" by simp + moreover from calculation + have "val(P,G,f_dot) : A \\<^bsup>M[G]\<^esup> B" + using truth_lemma[of "\0:1\2\" G "[f_dot, A\<^sup>v, B\<^sup>v]", THEN iffD1] + typed_function_type arity_typed_function_fm valcheck[OF one_in_G one_in_P] + by (auto simp: union_abs2 union_abs1 ext.mem_function_space_rel_abs) + moreover + note \a \ M\ + moreover from calculation and \a\A\ + have "val(P,G,f_dot) ` a \ B" (is "?b \ B") + by (simp add: ext.mem_function_space_rel_abs) + moreover from calculation + have "?b \ M" by (auto dest:transM) + moreover from calculation + have "M[G], map(val(P,G), [f_dot, a\<^sup>v, ?b\<^sup>v]) \ \0`1 is 2\" + by simp + moreover + note \M_generic(G)\ + ultimately + obtain r where "r \ \0`1 is 2\ [f_dot, a\<^sup>v, ?b\<^sup>v]" "r\G" "r\P" + using truth_lemma[of "\0`1 is 2\" G "[f_dot, a\<^sup>v, ?b\<^sup>v]", THEN iffD2] + fun_apply_type arity_fun_apply_fm valcheck[OF one_in_G one_in_P] + by (auto simp: union_abs2 union_abs1 ext.mem_function_space_rel_abs) + moreover from this and \q\G\ + obtain d where "d\q" "d\r" "d\P" by force + moreover + note \f_dot\M\ \a\M\ \?b\B\ \B\M\ + moreover from calculation + have "d \ q \ d \ \0`1 is 2\ [f_dot, a\<^sup>v, ?b\<^sup>v]" + using fun_apply_type arity_fun_apply_fm + strengthening_lemma[of r "\0`1 is 2\" d "[f_dot, a\<^sup>v, ?b\<^sup>v]"] + by (auto dest:transM simp add: union_abs2 union_abs1) + ultimately + show ?thesis by auto +qed + +context G_generic4_AC begin + +context + includes G_generic1_lemmas +begin + +lemma separation_check_snd_aux: + assumes "f_dot\M" "\\M" "\\formula" "arity(\) \ 7" + shows "separation(##M, \r. M, [fst(r), P, leq, \, f_dot, \, snd(r)\<^sup>v] \ \)" +proof - + note types = assms leq_in_M P_in_M one_in_M + let ?f_fm="fst_fm(1,0)" + let ?g_fm="hcomp_fm(check_fm'(6),snd_fm,2,0)" + have "?f_fm \ formula" "arity(?f_fm) \ 7" "?g_fm \ formula" "arity(?g_fm) \ 8" + using ord_simp_union + unfolding hcomp_fm_def check_fm'_def + by (simp_all add:arity) + then + show ?thesis + using separation_sat_after_function assms types + using fst_abs snd_abs types sats_snd_fm sats_check_fm check_abs check_in_M + unfolding hcomp_fm_def check_fm'_def + by simp +qed + +lemma separation_check_fst_snd_aux : + assumes "f_dot\M" "r\M" "\\formula" "arity(\) \ 7" + shows "separation(##M, \p. M, [r, P, leq, \, f_dot, fst(p)\<^sup>v, snd(p)\<^sup>v] \ \)" +proof - + let ?\="\z. [r, P, leq, \, f_dot, fst(z)\<^sup>v, snd(z)\<^sup>v]" + let ?\'="\z. [fst(z)\<^sup>v, P, leq, \, f_dot, r, snd(z)\<^sup>v]" + let ?\=" (\\(\\(\\(\\(\\(\\\\0 = 11\ \ \\1 = 7\ \ \\2 = 8\ \ \\3 = 9\ \ \\4 = 10\ \ \\5 = 6\ \ + (\p. incr_bv(p)`6)^6 (\) \\\\\\\)\)\)\)\)\)" + note types = assms leq_in_M P_in_M one_in_M + let ?f_fm="hcomp_fm(check_fm'(5),fst_fm,1,0)" + let ?g_fm="hcomp_fm(check_fm'(6),snd_fm,2,0)" + have "?f_fm \ formula" "arity(?f_fm) \ 7" "?g_fm \ formula" "arity(?g_fm) \ 8" + using ord_simp_union + unfolding hcomp_fm_def check_fm'_def + by (simp_all add:arity) + moreover from assms + have fm:"?\\formula" by simp + moreover from \\ \ formula\ \arity(\) \ 7\ + have "arity(\) = 0 \ arity(\) = 1 \ arity(\) = 2 \ arity(\) = 3 + \ arity(\) = 4 \ arity(\) = 5 \ arity(\) = 6 \ arity(\) = 7" + unfolding lt_def by auto + with calculation and \\ \ formula\ + have ar:"arity(?\) \ 7" + using arity_incr_bv_lemma by safe (simp_all add: arity ord_simp_union) + moreover from calculation + have sep:"separation(##M,\z. M,?\'(z)\?\)" + using separation_sat_after_function assms types sats_check_fm check_abs check_in_M + fst_abs snd_abs + unfolding hcomp_fm_def check_fm'_def + by simp + moreover + have "?\(z) \ list(M)" if "(##M)(z)" for z + using types that by simp + moreover from calculation and \r \ M\ \\ \ formula\ + have "(M,?\(z) \ \) \ (M,?\'(z)\?\)" if "(##M)(z)" for z + using that types sats_incr_bv_iff[of _ _ M _ "[_,_,_,_,_,_]"] + by simp + ultimately + show ?thesis using separation_cong[THEN iffD1,OF _ sep] + by simp +qed + +lemma separation_leq_and_forces_apply_aux: + assumes "f_dot\M" "B\M" + shows "\n\M. separation(##M, \x. snd(x) \ fst(x) \ + (\b\B. M, [snd(x), P, leq, \, f_dot, (\(n))\<^sup>v, b\<^sup>v] \ forces(\0`1 is 2\ )))" +proof - + have pred_nat_closed: "pred(n)\M" if "n\M" for n + using nat_case_closed that + unfolding pred_def + by auto + have "separation(##M, \z. M, [snd(fst(z)), P, leq, \, f_dot, \, snd(z)\<^sup>v] \ \)" + if "\\formula" "arity(\) \ 7" "\\M" for \ \ + proof - + note types = assms leq_in_M P_in_M one_in_M + let ?f_fm="hcomp_fm(snd_fm,fst_fm,1,0)" + let ?g_fm="hcomp_fm(check_fm'(6),snd_fm,2,0)" + have "?f_fm \ formula" "arity(?f_fm) \ 7" "?g_fm \ formula" "arity(?g_fm) \ 8" + using ord_simp_union + unfolding hcomp_fm_def check_fm'_def + by (simp_all add:arity) + then + show ?thesis + using separation_sat_after_function assms types sats_check_fm check_abs fst_abs snd_abs that + unfolding hcomp_fm_def check_fm'_def + by simp + qed + then + show ?thesis + using P_in_M assms + separation_in lam_replacement_constant lam_replacement_snd lam_replacement_fst + lam_replacement_Pair[THEN[5] lam_replacement_hcomp2] leq_in_M check_in_M pred_nat_closed + arity_forces[of " \0`1 is 2\"] arity_fun_apply_fm[of 0 1 2] ord_simp_union + by(clarify, rule_tac separation_conj,simp_all,rule_tac separation_bex,simp_all) +qed + +lemma separation_ball_leq_and_forces_apply_aux: + assumes "f_dot\M" "p\M" "B\M" + shows "separation + (##M, + \pa. \x\P. x \ p \ + (\y\P. y \ p \ + \x, y\ \ snd(pa) \ + y \ x \ (\b\B. M, [y, P, leq, \, f_dot, (\(fst(pa)))\<^sup>v, b\<^sup>v] \ forces(\0`1 is 2\ ))))" +proof - + have "separation(##M, \z. M, [snd(fst(z)), P, leq, \, f_dot, (\(fst(fst(fst(fst(z))))))\<^sup>v, snd(z)\<^sup>v] \ \)" + if "\\formula" "arity(\) \ 7" for \ + proof - + note types = assms leq_in_M P_in_M one_in_M + let ?f_fm="hcomp_fm(snd_fm,fst_fm,1,0)" + let ?g="\z . (\(fst(fst(fst(fst(z))))))\<^sup>v" + let ?g_fm="hcomp_fm(check_fm'(6),hcomp_fm(big_union_fm,hcomp_fm(fst_fm,hcomp_fm(fst_fm,hcomp_fm(fst_fm,fst_fm)))),2,0)" + let ?h_fm="hcomp_fm(check_fm'(7),snd_fm,3,0)" + have f_fm_facts:"?f_fm \ formula" "arity(?f_fm) \ 6" + using ord_simp_union + unfolding hcomp_fm_def + by (simp_all add:arity) + moreover from types + have "?g_fm \ formula" "arity(?g_fm) \ 7" "?h_fm \ formula" "arity(?h_fm) \ 8" + using ord_simp_union + unfolding hcomp_fm_def check_fm'_def + by (simp_all add:arity) + ultimately + show ?thesis + using separation_sat_after_function3[OF _ _ _ f_fm_facts] check_abs + types assms sats_check_fm that fst_abs snd_abs + unfolding hcomp_fm_def check_fm'_def + by simp + qed + then + show ?thesis + using P_in_M leq_in_M assms + separation_ball separation_imp separation_conj separation_bex separation_in separation_iff' + lam_replacement_constant lam_replacement_identity lam_replacement_hcomp + lam_replacement_fst lam_replacement_snd + lam_replacement_Pair[THEN[5] lam_replacement_hcomp2] + lam_replacement_hcomp[OF + lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_fst] + lam_replacement_snd] + arity_forces[of " \0`1 is 2\"] arity_fun_apply_fm[of 0 1 2] ord_simp_union + separation_in[OF _ lam_replacement_Pair[THEN[5] lam_replacement_hcomp2]] + by simp +qed + +lemma separation_closed_leq_and_forces_eq_check_aux : + assumes "A\M" "r\G" "\ \ M" + shows "(##M)({q\P. \h\A. q \ r \ q \ \0 = 1\ [\, h\<^sup>v]})" +proof - + have "separation(##M, \z. M, [fst(z), P, leq, \, \, snd(z)\<^sup>v] \ \)" if + "\\formula" "arity(\) \ 6" for \ + proof - + let ?f_fm="fst_fm(1,0)" + let ?g_fm="hcomp_fm(check_fm'(6),snd_fm,2,0)" + note types = assms leq_in_M P_in_M one_in_M + moreover + have "?f_fm \ formula" "arity(?f_fm) \ 6" "?g_fm \ formula" "arity(?g_fm) \ 7" + using ord_simp_union + unfolding hcomp_fm_def check_fm'_def + by (simp_all add:arity) + ultimately + show ?thesis + using separation_sat_after_function_1 assms sats_fst_fm that + fst_abs snd_abs types sats_snd_fm sats_check_fm check_abs check_in_M + unfolding hcomp_fm_def check_fm'_def + by simp + qed + then + show ?thesis + using separation_conj separation_in + lam_replacement_constant lam_replacement_fst + lam_replacement_Pair[THEN[5] lam_replacement_hcomp2] + assms leq_in_M G_subset_M[THEN subsetD] generic + arity_forces[of "\0 = 1\",simplified] ord_simp_union + by(rule_tac separation_closed[OF separation_bex],simp_all) +qed + +lemma separation_closed_forces_apply_aux: + assumes "B\M" "f_dot\M" "r\M" + shows "(##M)({\n,b\ \ \ \ B. r \ \0`1 is 2\ [f_dot, n\<^sup>v, b\<^sup>v]})" + using nat_in_M assms check_in_M transitivity[OF _ \B\M\] nat_into_M separation_check_fst_snd_aux + arity_forces[of " \0`1 is 2\"] arity_fun_apply_fm[of 0 1 2] ord_simp_union + unfolding split_def + by simp_all + +\ \Kunen IV.6.9 (3)$\Rightarrow$(2), with general domain.\ +lemma kunen_IV_6_9_function_space_rel_eq: + assumes "\p \. p \ \0:1\2\ [\, A\<^sup>v, B\<^sup>v] \ p\P \ \ \ M \ + \q\P. \h\A \\<^bsup>M\<^esup> B. q \ p \ q \ \0 = 1\ [\, h\<^sup>v]" "A\M" "B\M" + shows + "A \\<^bsup>M\<^esup> B = A \\<^bsup>M[G]\<^esup> B" +proof (intro equalityI; clarsimp simp add: + assms function_space_rel_char ext.function_space_rel_char) + fix f + assume "f \ A \ B" "f \ M[G]" + moreover from this + obtain \ where "val(P,G,\) = f" "\ \ M" + using GenExtD by force + moreover from calculation and \A\M\ \B\M\ + obtain r where "r \ \0:1\2\ [\, A\<^sup>v, B\<^sup>v]" "r\G" + using truth_lemma[of "\0:1\2\" G "[\, A\<^sup>v, B\<^sup>v]"] generic + typed_function_type arity_typed_function_fm valcheck[OF one_in_G one_in_P] + by (auto simp: union_abs2 union_abs1) + moreover from \A\M\ \B\M\ \r\G\ \\ \ M\ + have "{q\P. \h\A \\<^bsup>M\<^esup> B. q \ r \ q \ \0 = 1\ [\, h\<^sup>v]} \ M" (is "?D \ M") + using separation_closed_leq_and_forces_eq_check_aux by auto + moreover from calculation and assms(2-) + have "dense_below(?D, r)" + using strengthening_lemma[of r "\0:1\2\" _ "[\, A\<^sup>v, B\<^sup>v]", THEN assms(1)[of _ \]] + leq_transD generic_dests(1)[of r] + by (auto simp: union_abs2 union_abs1 typed_function_type arity_typed_function_fm) blast + moreover from calculation + obtain q h where "h\A \\<^bsup>M\<^esup> B" "q \ \0 = 1\ [\, h\<^sup>v]" "q \ r" "q\P" "q\G" + using generic_inter_dense_below[of ?D G r, OF _ generic] by blast + note \q \ \0 = 1\ [\, h\<^sup>v]\ \\\M\ \h\A \\<^bsup>M\<^esup> B\ \A\M\ \B\M\ \q\G\ + moreover from this + have "map(val(P, G), [\, h\<^sup>v]) \ list(M[G])" "h\M" + by (auto dest:transM) + ultimately + have "h = f" + using truth_lemma[of "\0=1\" G "[\, h\<^sup>v]"] generic valcheck[OF one_in_G one_in_P] + by (auto simp: ord_simp_union) + with \h\M\ + show "f \ M" by simp +qed + +subsection\$(\omega+1)$-Closed notions preserve countable sequences\ + +\ \Kunen IV.7.15, only for countable sequences\ +lemma succ_omega_closed_imp_no_new_nat_sequences: + assumes "succ(\)-closed\<^bsup>M\<^esup>(P,leq)" "f : \ \ B" "f\M[G]" "B\M" + shows "f\M" +proof - + (* Nice jEdit folding level to read this: 7 *) + txt\The next long block proves that the assumptions of Lemma + @{thm [source] kunen_IV_6_9_function_space_rel_eq} are satisfied.\ + { + fix p f_dot + assume "p \ \0:1\2\ [f_dot, \\<^sup>v, B\<^sup>v]" "p\P" "f_dot\M" + let ?subp="{q\P. q \ p}" + from \p\P\ + have "?subp \ M" + using first_section_closed[of P p "converse(leq)"] leq_in_M P_in_M + by (auto dest:transM) + define S where "S \ \n\nat. + {\q,r\ \ ?subp\?subp. r \ q \ (\b\B. r \ \0`1 is 2\ [f_dot, (\(n))\<^sup>v, b\<^sup>v])}" + (is "S \ \n\nat. ?Y(n)") + define S' where "S' \ \n\nat. + {\q,r\ \ ?subp\?subp. r \ q \ (\b\B. r \ \0`1 is 2\ [f_dot, (pred(n))\<^sup>v, b\<^sup>v])}" + \ \Towards proving \<^term>\S\M\.\ + moreover + have "S = S'" + unfolding S_def S'_def using pred_nat_eq lam_cong by auto + moreover from \B\M\ \?subp\M\ \f_dot\M\ + have "{r \ ?subp. \b\B. r \ \0`1 is 2\ [f_dot, (\(n))\<^sup>v, b\<^sup>v]} \ M" (is "?X(n) \ M") + if "n\\" for n + using that separation_check_snd_aux nat_into_M ord_simp_union + arity_forces[of " \0`1 is 2\"] arity_fun_apply_fm + by(rule_tac separation_closed[OF separation_bex,simplified], simp_all) + moreover + have "?Y(n) = (?subp \ ?X(n)) \ converse(leq)" for n + by (intro equalityI) auto + moreover + note \?subp \ M\ \B\M\ \p\P\ \f_dot\M\ + moreover from calculation + have "n \ \ \ ?Y(n) \ M" for n + using nat_into_M leq_in_M by simp + moreover from calculation + have "S \ M" + using separation_ball_leq_and_forces_apply_aux separation_leq_and_forces_apply_aux + transitivity[OF \p\P\ P_in_M] + unfolding S_def split_def + by(rule_tac lam_replacement_Collect[THEN lam_replacement_imp_lam_closed,simplified], simp_all) + ultimately + have "S' \ M" + by simp + from \p\P\ \f_dot\M\ \p \ \0:1\2\ [f_dot, \\<^sup>v, B\<^sup>v]\ \B\M\ + have exr:"\r\P. r \ q \ (\b\B. r \ \0`1 is 2\ [f_dot, pred(n)\<^sup>v, b\<^sup>v])" + if "q \ p" "q\P" "n\\" for q n + using that forcing_a_value by (auto dest:transM) + have "\q\?subp. \n\\. \r\?subp. \q,r\ \ S'`n" + proof - + { + fix q n + assume "q \ ?subp" "n\\" + moreover from this + have "q \ p" "q \ P" "pred(n) = \n" + using pred_nat_eq by simp_all + moreover from calculation and exr + obtain r where MM:"r \ q" "\b\B. r \ \0`1 is 2\ [f_dot, pred(n)\<^sup>v, b\<^sup>v]" "r\P" + by blast + moreover from calculation \q \ p\ \p \ P\ + have "r \ p" + using leq_transD[of r q p] by auto + ultimately + have "\r\?subp. r \ q \ (\b\B. r \ \0`1 is 2\ [f_dot, (pred(n))\<^sup>v, b\<^sup>v])" + by auto + } + then + show ?thesis + unfolding S'_def by simp + qed + with \p\P\ \?subp \ M\ \S' \ M\ + obtain g where "g \ \ \\<^bsup>M\<^esup> ?subp" "g`0 = p" "\n \ nat. \g`n,g`succ(n)\\S'`succ(n)" + using sequence_DC[simplified] refl_leq[of p] by blast + moreover from this and \?subp \ M\ + have "g : \ \ P" "g \ M" + using fun_weaken_type[of g \ ?subp P] function_space_rel_char by auto + ultimately + have "g : \ \<^sub><\\<^bsup>M\<^esup> (P,converse(leq))" + using decr_succ_decr[of g] leq_preord leq_in_M P_in_M + unfolding S'_def by (auto simp:absolut intro:leI) + moreover from \succ(\)-closed\<^bsup>M\<^esup>(P,leq)\ and this + have "\q\M. q \ P \ (\\\M. \ \ \ \ q \ g ` \)" + using transM[simplified, of g] leq_in_M + mono_seqspace_rel_closed[of \ _ "converse(leq)"] + unfolding kappa_closed_rel_def + by auto + ultimately + obtain r where "r\P" "r\M" "\n\\. r \ g`n" + using nat_into_M by auto + with \g`0 = p\ + have "r \ p" + by blast + let ?h="{\n,b\ \ \ \ B. r \ \0`1 is 2\ [f_dot, n\<^sup>v, b\<^sup>v]}" + have "function(?h)" + proof (rule_tac functionI, rule_tac ccontr, auto simp del: app_Cons) + fix n b b' + assume "n \ \" "b \ b'" "b \ B" "b' \ B" + moreover + assume "r \ \0`1 is 2\ [f_dot, n\<^sup>v, b\<^sup>v]" "r \ \0`1 is 2\ [f_dot, n\<^sup>v, b'\<^sup>v]" + moreover + note \r \ P\ + moreover from this + have "\ r \ r" + by (auto intro!:refl_leq) + moreover + note \f_dot\M\ \B\M\ + ultimately + show False + using forces_neq_apply_imp_incompatible[of r f_dot "n\<^sup>v" b r b'] + transM[of _ B] by (auto dest:transM) + qed + moreover + have "range(?h) \ B" + by auto + moreover + have "domain(?h) = \" + proof - + { + fix n + assume "n \ \" + moreover from this + have 1:"(\(n)) = pred(n)" + using pred_nat_eq by simp + moreover from calculation and \\n \ nat. \g`n,g`succ(n)\\S'`succ(n)\ + obtain b where "g`(succ(n)) \ \0`1 is 2\ [f_dot, n\<^sup>v, b\<^sup>v]" "b\B" + unfolding S'_def by auto + moreover from \B\M\ and calculation + have "b \ M" "n \ M" + by (auto dest:transM) + moreover + note \g : \ \ P\ \\n\\. r \ g`n\ \r\P\ \f_dot\M\ + moreover from calculation + have "r \ \0`1 is 2\ [f_dot, n\<^sup>v, b\<^sup>v]" + using fun_apply_type arity_fun_apply_fm + strengthening_lemma[of "g`succ(n)" "\0`1 is 2\" r "[f_dot, n\<^sup>v, b\<^sup>v]"] + by (simp add: union_abs2 union_abs1) + ultimately + have "\b\B. r \ \0`1 is 2\ [f_dot, n\<^sup>v, b\<^sup>v]" + by auto + } + then + show ?thesis + by force + qed + moreover + have "relation(?h)" + unfolding relation_def by simp + moreover from \f_dot\M\ \r\M\ \B\M\ + have "?h \ M" + using separation_closed_forces_apply_aux by simp + moreover + note \B \ M\ + ultimately + have "?h: \ \\<^bsup>M\<^esup> B" + using function_imp_Pi[THEN fun_weaken_type[of ?h _ "range(?h)" B]] + function_space_rel_char by simp + moreover + note \p \ \0:1\2\ [f_dot, \\<^sup>v, B\<^sup>v]\ \r \ p\ \r\P\ \p\P\ \f_dot\M\ \B\M\ + moreover from this + have "r \ \0:1\2\ [f_dot, \\<^sup>v, B\<^sup>v]" + using strengthening_lemma[of p "\0:1\2\" r "[f_dot, \\<^sup>v, B\<^sup>v]"] + typed_function_type arity_typed_function_fm + by (auto simp: union_abs2 union_abs1) + moreover + note \?h\M\ + moreover from calculation + have "r \ \0 = 1\ [f_dot, ?h\<^sup>v]" + proof (intro definition_of_forcing[THEN iffD2] allI impI, + simp_all add:union_abs2 union_abs1 del:app_Cons) + fix G + let ?f="val(P,G,f_dot)" + assume "M_generic(G) \ r \ G" + moreover from this + interpret g:G_generic1 _ _ _ _ _ G + by unfold_locales simp + note \r\P\ \f_dot\M\ \B\M\ + moreover from this + have "map(val(P, G), [f_dot, \\<^sup>v, B\<^sup>v]) \ list(M[G])" + by simp + moreover from calculation and \r \ \0:1\2\ [f_dot, \\<^sup>v, B\<^sup>v]\ + have "?f : \ \ B" + using truth_lemma[of "\0:1\2\" G "[f_dot, \\<^sup>v, B\<^sup>v]"] one_in_G one_in_P + typed_function_type arity_typed_function_fm valcheck + by (auto simp: union_abs2 union_abs1) + moreover + have "?h`n = ?f`n" if "n \ \" for n + proof - + note \n \ \\ \domain(?h) = \\ + moreover from this + have "n\domain(?h)" + by simp + moreover from this + obtain b where "r \ \0`1 is 2\ [f_dot, n\<^sup>v, b\<^sup>v]" "b\B" + by force + moreover + note \function(?h)\ + moreover from calculation + have "b = ?h`n" + using function_apply_equality by simp + moreover + note \B \ M\ + moreover from calculation + have "?h`n \ M" + by (auto dest:transM) + moreover + note \f_dot \ M\ \r \ P\ \M_generic(G) \ r \ G\ \map(val(P, G), [f_dot, \\<^sup>v, B\<^sup>v]) \ list(M[G])\ + moreover from calculation + have "[?f, n, ?h`n] \ list(M[G])" + using M_subset_MG nat_into_M[of n] one_in_G by (auto dest:transM) + ultimately + show ?thesis + using definition_of_forcing[of r "\0`1 is 2\" "[f_dot, n\<^sup>v, b\<^sup>v]", + THEN iffD1, rule_format, of G]\ \without this line is slower\ + valcheck one_in_G one_in_P nat_into_M + by (auto dest:transM simp add:fun_apply_type + arity_fun_apply_fm union_abs2 union_abs1) + qed + with calculation and \B\M\ \?h: \ \\<^bsup>M\<^esup> B\ + have "?h = ?f" + using function_space_rel_char + by (rule_tac fun_extension[of ?h \ "\_.B" ?f]) auto + ultimately + show "?f = val(P, G, ?h\<^sup>v)" + using valcheck one_in_G one_in_P generic by simp + qed + ultimately + have "\r\P. \h\\ \\<^bsup>M\<^esup> B. r \ p \ r \ \0 = 1\ [f_dot, h\<^sup>v]" + by blast + } + moreover + note \B \ M\ assms + moreover from calculation + have "f : \ \\<^bsup>M\<^esup> B" + using kunen_IV_6_9_function_space_rel_eq function_space_rel_char + ext.mem_function_space_rel_abs by auto + ultimately + show ?thesis + by (auto dest:transM) +qed + +declare mono_seqspace_rel_closed[rule del] + \ \Mysteriously breaks the end of the next proof\ + +lemma succ_omega_closed_imp_no_new_reals: + assumes "succ(\)-closed\<^bsup>M\<^esup>(P,leq)" + shows "\ \\<^bsup>M\<^esup> 2 = \ \\<^bsup>M[G]\<^esup> 2" +proof - + from assms + have "\ \\<^bsup>M[G]\<^esup> 2 \ \ \\<^bsup>M\<^esup> 2" + using succ_omega_closed_imp_no_new_nat_sequences function_space_rel_char + ext.function_space_rel_char Aleph_rel_succ Aleph_rel_zero + by auto + then + show ?thesis + using function_space_rel_transfer by (intro equalityI) auto +qed + +lemma succ_omega_closed_imp_Aleph_1_preserved: + assumes "succ(\)-closed\<^bsup>M\<^esup>(P,leq)" + shows "\\<^bsub>1\<^esub>\<^bsup>M\<^esup> = \\<^bsub>1\<^esub>\<^bsup>M[G]\<^esup>" +proof - + have "\\<^bsub>1\<^esub>\<^bsup>M[G]\<^esup> \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" + proof (rule ccontr) + assume "\ \\<^bsub>1\<^esub>\<^bsup>M[G]\<^esup> \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" + then + have "\\<^bsub>1\<^esub>\<^bsup>M\<^esup> < \\<^bsub>1\<^esub>\<^bsup>M[G]\<^esup>" + \ \Ridiculously complicated proof\ + using Card_rel_is_Ord ext.Card_rel_is_Ord + not_le_iff_lt[THEN iffD1] by auto + then + have "|\\<^bsub>1\<^esub>\<^bsup>M\<^esup>|\<^bsup>M[G]\<^esup> \ \" + using ext.Card_rel_lt_csucc_rel_iff ext.Aleph_rel_zero + ext.Aleph_rel_succ ext.Card_rel_nat + by (auto intro!:ext.lt_csucc_rel_iff[THEN iffD1] + intro:Card_rel_Aleph_rel[THEN Card_rel_is_Ord, of 1]) + then + obtain f where "f \ inj(\\<^bsub>1\<^esub>\<^bsup>M\<^esup>,\)" "f \ M[G]" + using ext.countable_rel_iff_cardinal_rel_le_nat[of "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>", THEN iffD2] + unfolding countable_rel_def lepoll_rel_def + by auto + then + obtain g where "g \ surj\<^bsup>M[G]\<^esup>(\, \\<^bsub>1\<^esub>\<^bsup>M\<^esup>)" + using ext.inj_rel_imp_surj_rel[of f _ \, OF _ zero_lt_Aleph_rel1[THEN ltD]] + by auto + moreover from this + have "g : \ \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "g \ M[G]" + using ext.surj_rel_char surj_is_fun by simp_all + moreover + note \succ(\)-closed\<^bsup>M\<^esup>(P,leq)\ + ultimately + have "g \ surj\<^bsup>M\<^esup>(\, \\<^bsub>1\<^esub>\<^bsup>M\<^esup>)" "g \ M" + using succ_omega_closed_imp_no_new_nat_sequences + mem_surj_abs ext.mem_surj_abs by simp_all + then + show False + using surj_rel_implies_cardinal_rel_le[of g \ "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>"] + Card_rel_nat[THEN Card_rel_cardinal_rel_eq] Card_rel_is_Ord + not_le_iff_lt[THEN iffD2, OF _ _ nat_lt_Aleph_rel1] + by simp + qed + then + show ?thesis + using Aleph_rel_le_Aleph_rel + by (rule_tac le_anti_sym) simp +qed + +end \ \bundle G\_generic1\_lemmas\ + +end \ \\<^locale>\G_generic4_AC\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Names.thy b/thys/Independence_CH/Names.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Names.thy @@ -0,0 +1,634 @@ +section\Names and generic extensions\ + +theory Names + imports + Forcing_Data + FrecR_Arities +begin + +definition + Hv :: "[i,i,i,i]\i" where + "Hv(P,G,x,f) \ { z . y\ domain(x), (\p\P. \y,p\ \ x \ p \ G) \ z=f`y}" + +text\The funcion \<^term>\val\ interprets a name in \<^term>\M\ +according to a (generic) filter \<^term>\G\. Note the definition +in terms of the well-founded recursor.\ + +definition + val :: "[i,i,i]\i" where + "val(P,G,\) \ wfrec(edrel(eclose({\})), \ ,Hv(P,G))" + +definition + GenExt :: "[i,i,i]\i" ("_\<^bsup>_\<^esup>[_]" [71,1]) + where "M\<^bsup>P\<^esup>[G] \ {val(P,G,\). \ \ M}" + +abbreviation (in forcing_notion) + GenExt_at_P :: "i\i\i" ("_[_]" [71,1]) + where "M[G] \ M\<^bsup>P\<^esup>[G]" + +subsection\Values and check-names\ +context forcing_data1 +begin + +definition + Hcheck :: "[i,i] \ i" where + "Hcheck(z,f) \ { \f`y,\\ . y \ z}" + +definition + check :: "i \ i" where + "check(x) \ transrec(x , Hcheck)" + +lemma checkD: + "check(x) = wfrec(Memrel(eclose({x})), x, Hcheck)" + unfolding check_def transrec_def .. + +lemma Hcheck_trancl:"Hcheck(y, restrict(f,Memrel(eclose({x}))-``{y})) + = Hcheck(y, restrict(f,(Memrel(eclose({x}))^+)-``{y}))" + unfolding Hcheck_def + using restrict_trans_eq by simp + +lemma check_trancl: "check(x) = wfrec(rcheck(x), x, Hcheck)" + using checkD wf_eq_trancl Hcheck_trancl unfolding rcheck_def by simp + +lemma rcheck_in_M : "x \ M \ rcheck(x) \ M" + unfolding rcheck_def by (simp flip: setclass_iff) + +lemma aux_def_check: "x \ y \ + wfrec(Memrel(eclose({y})), x, Hcheck) = + wfrec(Memrel(eclose({x})), x, Hcheck)" + by (rule wfrec_eclose_eq,auto simp add: arg_into_eclose eclose_sing) + +lemma def_check : "check(y) = { \check(w),\\ . w \ y}" +proof - + let + ?r="\y. Memrel(eclose({y}))" + have wfr: "\w . wf(?r(w))" + using wf_Memrel .. + then + have "check(y)= Hcheck( y, \x\?r(y) -`` {y}. wfrec(?r(y), x, Hcheck))" + using wfrec[of "?r(y)" y "Hcheck"] checkD by simp + also + have " ... = Hcheck( y, \x\y. wfrec(?r(y), x, Hcheck))" + using under_Memrel_eclose arg_into_eclose by simp + also + have " ... = Hcheck( y, \x\y. check(x))" + using aux_def_check checkD by simp + finally + show ?thesis + using Hcheck_def by simp +qed + +lemma def_checkS : + fixes n + assumes "n \ nat" + shows "check(succ(n)) = check(n) \ {\check(n),\\}" +proof - + have "check(succ(n)) = {\check(i),\\ . i \ succ(n)} " + using def_check by blast + also + have "... = {\check(i),\\ . i \ n} \ {\check(n),\\}" + by blast + also + have "... = check(n) \ {\check(n),\\}" + using def_check[of n,symmetric] by simp + finally + show ?thesis . +qed + +lemma field_Memrel2 : + assumes "x \ M" + shows "field(Memrel(eclose({x}))) \ M" +proof - + have "field(Memrel(eclose({x}))) \ eclose({x})" "eclose({x}) \ M" + using Ordinal.Memrel_type field_rel_subset assms eclose_least[OF trans_M] by auto + then + show ?thesis + using subset_trans by simp +qed + +lemma aux_def_val: + assumes "z \ domain(x)" + shows "wfrec(edrel(eclose({x})),z,Hv(P,G)) = wfrec(edrel(eclose({z})),z,Hv(P,G))" +proof - + let ?r="\x . edrel(eclose({x}))" + have "z\eclose({z})" + using arg_in_eclose_sing . + moreover + have "relation(?r(x))" + using relation_edrel . + moreover + have "wf(?r(x))" + using wf_edrel . + moreover from assms + have "tr_down(?r(x),z) \ eclose({z})" + using tr_edrel_subset by simp + ultimately + have "wfrec(?r(x),z,Hv(P,G)) = wfrec[eclose({z})](?r(x),z,Hv(P,G))" + using wfrec_restr by simp + also from \z\domain(x)\ + have "... = wfrec(?r(z),z,Hv(P,G))" + using restrict_edrel_eq wfrec_restr_eq by simp + finally + show ?thesis . +qed + +text\The next lemma provides the usual recursive expresion for the definition of term\val\.\ + +lemma def_val: "val(P,G,x) = {z . t\domain(x) , (\p\P . \t,p\\x \ p \ G) \ z=val(P,G,t)}" +proof - + let + ?r="\\ . edrel(eclose({\}))" + let + ?f="\z\?r(x)-``{x}. wfrec(?r(x),z,Hv(P,G))" + have "\\. wf(?r(\))" + using wf_edrel by simp + with wfrec [of _ x] + have "val(P,G,x) = Hv(P,G,x,?f)" + using val_def by simp + also + have " ... = Hv(P,G,x,\z\domain(x). wfrec(?r(x),z,Hv(P,G)))" + using dom_under_edrel_eclose by simp + also + have " ... = Hv(P,G,x,\z\domain(x). val(P,G,z))" + using aux_def_val val_def by simp + finally + show ?thesis + using Hv_def by simp +qed + +lemma val_mono : "x\y \ val(P,G,x) \ val(P,G,y)" + by (subst (1 2) def_val, force) + +text\Check-names are the canonical names for elements of the +ground model. Here we show that this is the case.\ + +lemma valcheck : "\ \ G \ \ \ P \ val(P,G,check(y)) = y" +proof (induct rule:eps_induct) + case (1 y) + then show ?case + proof - + have "check(y) = { \check(w), \\ . w \ y}" (is "_ = ?C") + using def_check . + then + have "val(P,G,check(y)) = val(P,G, {\check(w), \\ . w \ y})" + by simp + also + have " ... = {z . t\domain(?C) , (\p\P . \t, p\\?C \ p \ G) \ z=val(P,G,t) }" + using def_val by blast + also + have " ... = {z . t\domain(?C) , (\w\y. t=check(w)) \ z=val(P,G,t) }" + using 1 by simp + also + have " ... = {val(P,G,check(w)) . w\y }" + by force + finally + show "val(P,G,check(y)) = y" + using 1 by simp + qed +qed + +lemma val_of_name : + "val(P,G,{x\A\P. Q(x)}) = {z . t\A , (\p\P . Q(\t,p\) \ p \ G) \ z=val(P,G,t)}" +proof - + let + ?n="{x\A\P. Q(x)}" and + ?r="\\ . edrel(eclose({\}))" + let + ?f="\z\?r(?n)-``{?n}. val(P,G,z)" + have + wfR : "wf(?r(\))" for \ + by (simp add: wf_edrel) + have "domain(?n) \ A" by auto + { fix t + assume H:"t \ domain({x \ A \ P . Q(x)})" + then have "?f ` t = (if t \ ?r(?n)-``{?n} then val(P,G,t) else 0)" + by simp + moreover have "... = val(P,G,t)" + using dom_under_edrel_eclose H if_P by auto + } + then + have Eq1: "t \ domain({x \ A \ P . Q(x)}) \ val(P,G,t) = ?f` t" for t + by simp + have "val(P,G,?n) = {z . t\domain(?n), (\p \ P . \t,p\ \ ?n \ p \ G) \ z=val(P,G,t)}" + by (subst def_val,simp) + also + have "... = {z . t\domain(?n), (\p\P . \t,p\\?n \ p\G) \ z=?f`t}" + unfolding Hv_def + by (auto simp add:Eq1) + also + have "... = {z . t\domain(?n), (\p\P . \t,p\\?n \ p\G) \ z=(if t\?r(?n)-``{?n} then val(P,G,t) else 0)}" + by (simp) + also + have "... = { z . t\domain(?n), (\p\P . \t,p\\?n \ p\G) \ z=val(P,G,t)}" + proof - + have "domain(?n) \ ?r(?n)-``{?n}" + using dom_under_edrel_eclose by simp + then + have "\t\domain(?n). (if t\?r(?n)-``{?n} then val(P,G,t) else 0) = val(P,G,t)" + by auto + then + show "{ z . t\domain(?n), (\p\P . \t,p\\?n \ p\G) \ z=(if t\?r(?n)-``{?n} then val(P,G,t) else 0)} = + { z . t\domain(?n), (\p\P . \t,p\\?n \ p\G) \ z=val(P,G,t)}" + by auto + qed + also + have " ... = { z . t\A, (\p\P . \t,p\\?n \ p\G) \ z=val(P,G,t)}" + by force + finally + show " val(P,G,?n) = { z . t\A, (\p\P . Q(\t,p\) \ p\G) \ z=val(P,G,t)}" + by auto +qed + +lemma val_of_name_alt : + "val(P,G,{x\A\P. Q(x)}) = {z . t\A , (\p\P\G . Q(\t,p\)) \ z=val(P,G,t) }" + using val_of_name by force + +lemma val_only_names: "val(P,F,\) = val(P,F,{x\\. \t\domain(\). \p\P. x=\t,p\})" + (is "_ = val(P,F,?name)") +proof - + have "val(P,F,?name) = {z . t\domain(?name), (\p\P. \t, p\ \ ?name \ p \ F) \ z=val(P,F, t)}" + using def_val by blast + also + have " ... = {val(P,F, t). t\{y\domain(\). \p\P. \y, p\ \ \ \ p \ F}}" + by blast + also + have " ... = {z . t\domain(\), (\p\P. \t, p\ \ \ \ p \ F) \ z=val(P,F, t)}" + by blast + also + have " ... = val(P,F, \)" + using def_val[symmetric] by blast + finally + show ?thesis .. +qed + +lemma val_only_pairs: "val(P,F,\) = val(P,F,{x\\. \t p. x=\t,p\})" +proof + have "val(P,F,\) = val(P,F,{x\\. \t\domain(\). \p\P. x=\t,p\})" (is "_ = val(P,F,?name)") + using val_only_names . + also + have "... \ val(P,F,{x\\. \t p. x=\t,p\})" + using val_mono[of ?name "{x\\. \t p. x=\t,p\}"] by auto + finally + show "val(P,F,\) \ val(P,F,{x\\. \t p. x=\t,p\})" by simp +next + show "val(P,F,{x\\. \t p. x=\t,p\}) \ val(P,F,\)" + using val_mono[of "{x\\. \t p. x=\t,p\}"] by auto +qed + +lemma val_subset_domain_times_range: "val(P,F,\) \ val(P,F,domain(\)\range(\))" + using val_only_pairs[THEN equalityD1] + val_mono[of "{x \ \ . \t p. x = \t, p\}" "domain(\)\range(\)"] by blast + +lemma val_subset_domain_times_P: "val(P,F,\) \ val(P,F,domain(\)\P)" + using val_only_names[of F \] val_mono[of "{x\\. \t\domain(\). \p\P. x=\t,p\}" "domain(\)\P" F] + by auto + +lemma val_of_elem: "\\,p\ \ \ \ p\G \ p\P \ val(P,G,\) \ val(P,G,\)" +proof - + assume "\\,p\ \ \" + then + have "\\domain(\)" + by auto + assume "p\G" "p\P" + with \\\domain(\)\ \\\,p\ \ \\ + have "val(P,G,\) \ {z . t\domain(\) , (\p\P . \t, p\\\ \ p \ G) \ z=val(P,G,t) }" + by auto + then + show ?thesis + by (subst def_val) +qed + +lemma elem_of_val: "x\val(P,G,\) \ \\\domain(\). val(P,G,\) = x" + by (subst (asm) def_val,auto) + +lemma elem_of_val_pair: "x\val(P,G,\) \ \\. \p\G. \\,p\\\ \ val(P,G,\) = x" + by (subst (asm) def_val,auto) + +lemma elem_of_val_pair': + assumes "\\M" "x\val(P,G,\)" + shows "\\\M. \p\G. \\,p\\\ \ val(P,G,\) = x" +proof - + from assms + obtain \ p where "p\G" "\\,p\\\" "val(P,G,\) = x" + using elem_of_val_pair by blast + moreover from this \\\M\ + have "\\M" + using pair_in_M_iff[THEN iffD1, THEN conjunct1, simplified] + transitivity by blast + ultimately + show ?thesis + by blast +qed + + +lemma GenExtD: "x \ M[G] \ \\\M. x = val(P,G,\)" + by (simp add:GenExt_def) + +lemma GenExtI: "x \ M \ val(P,G,x) \ M[G]" + by (auto simp add: GenExt_def) + +lemma Transset_MG : "Transset(M[G])" +proof - + { fix vc y + assume "vc \ M[G]" and "y \ vc" + then + obtain c where "c\M" "val(P,G,c)\M[G]" "y \ val(P,G,c)" + using GenExtD by auto + from \y \ val(P,G,c)\ + obtain \ where "\\domain(c)" "val(P,G,\) = y" + using elem_of_val by blast + with trans_M \c\M\ + have "y \ M[G]" + using domain_trans GenExtI by blast + } + then + show ?thesis + using Transset_def by auto +qed + +lemmas transitivity_MG = Transset_intf[OF Transset_MG] + +text\This lemma can be proved before having \<^term>\check_in_M\. At some point Miguel naïvely +thought that the \<^term>\check_in_M\ could be proved using this argument.\ +lemma check_nat_M : + assumes "n \ nat" + shows "check(n) \ M" + using assms +proof (induct n) + case 0 + then + show ?case + using zero_in_M by (subst def_check,simp) +next + case (succ x) + have "\ \ M" + using one_in_P P_sub_M subsetD by simp + with \check(x)\M\ + have "\check(x),\\ \ M" + using pair_in_M_iff by simp + then + have "{\check(x),\\} \ M" + using singleton_closed by simp + with \check(x)\M\ + have "check(x) \ {\check(x),\\} \ M" + using Un_closed by simp + then + show ?case + using \x\nat\ def_checkS by simp +qed + +lemma def_PHcheck: + assumes + "z\M" "f\M" + shows + "Hcheck(z,f) = Replace(z,PHcheck(##M,\,f))" +proof - + from assms + have "\f`x,\\ \ M" "f`x\M" if "x\z" for x + using pair_in_M_iff one_in_M transitivity that apply_closed by simp_all + then + have "{y . x \ z, y = \f ` x, \\} = {y . x \ z, y = \f ` x, \\ \ y\M \ f`x\M}" + by simp + then + show ?thesis + using \z\M\ \f\M\ transitivity + unfolding Hcheck_def PHcheck_def RepFun_def + by auto +qed + + +(* instance of replacement for hcheck *) +lemma wfrec_Hcheck : + assumes "X\M" + shows "wfrec_replacement(##M,is_Hcheck(##M,\),rcheck(X))" +proof - + let ?f="Exists(And(pair_fm(1,0,2), + is_wfrec_fm(is_Hcheck_fm(8,2,1,0),4,1,0)))" + have "is_Hcheck(##M,\,a,b,c) \ + sats(M,is_Hcheck_fm(8,2,1,0),[c,b,a,d,e,y,x,z,\,rcheck(x)])" + if "a\M" "b\M" "c\M" "d\M" "e\M" "y\M" "x\M" "z\M" + for a b c d e y x z + using that one_in_M \X\M\ rcheck_in_M is_Hcheck_iff_sats zero_in_M + by simp + then + have "sats(M,is_wfrec_fm(is_Hcheck_fm(8,2,1,0),4,1,0), [y,x,z,\,rcheck(X)]) + \ is_wfrec(##M, is_Hcheck(##M,\),rcheck(X), x, y)" + if "x\M" "y\M" "z\M" for x y z + using that sats_is_wfrec_fm \X\M\ rcheck_in_M one_in_M zero_in_M + by simp + moreover from this + have satsf:"sats(M, ?f, [x,z,\,rcheck(X)]) \ + (\y\M. pair(##M,x,y,z) & is_wfrec(##M, is_Hcheck(##M,\),rcheck(X), x, y))" + if "x\M" "z\M" for x z + using that \X\M\ rcheck_in_M one_in_M + by (simp del:pair_abs) + moreover + have artyf:"arity(?f) = 4" + using arity_wfrec_replacement_fm[where p="is_Hcheck_fm(8, 2, 1, 0)" and i=9] + arity_is_Hcheck_fm ord_simp_union + by simp + ultimately + have "strong_replacement(##M,\x z. sats(M,?f,[x,z,\,rcheck(X)]))" + using replacement_ax1(10) artyf \X\M\ rcheck_in_M one_in_M + unfolding replacement_assm_def by simp + then + have "strong_replacement(##M,\x z. + \y\M. pair(##M,x,y,z) & is_wfrec(##M, is_Hcheck(##M,\),rcheck(X), x, y))" + using repl_sats[of M ?f "[\,rcheck(X)]"] satsf by (simp del:pair_abs) + then + show ?thesis + unfolding wfrec_replacement_def by simp +qed + +lemma repl_PHcheck : + assumes "f\M" + shows "strong_replacement(##M,PHcheck(##M,\,f))" +proof - + from \f\M\ + have "strong_replacement(##M,\x y. sats(M,PHcheck_fm(2,3,0,1),[x,y,\,f]))" + using replacement_ax1(11) one_in_M unfolding replacement_assm_def + by (simp add:arity ord_simp_union) + with \f\M\ + show ?thesis + using one_in_M zero_in_M + unfolding strong_replacement_def univalent_def + by simp +qed + +lemma univ_PHcheck : "\ z\M ; f\M \ \ univalent(##M,z,PHcheck(##M,\,f))" + unfolding univalent_def PHcheck_def + by simp + +lemma PHcheck_closed : "\z\M ; f\M ; x\z; PHcheck(##M,\,f,x,y) \ \ (##M)(y)" + unfolding PHcheck_def by simp + +lemma relation2_Hcheck : "relation2(##M,is_Hcheck(##M,\),Hcheck)" +proof - + have "is_Replace(##M,z,PHcheck(##M,\,f),hc) \ hc = Replace(z,PHcheck(##M,\,f))" + if "z\M" "f\M" "hc\M" for z f hc + using that Replace_abs[OF _ _ univ_PHcheck] PHcheck_closed[of z f] + by simp + with def_PHcheck + show ?thesis + unfolding relation2_def is_Hcheck_def Hcheck_def + by simp +qed + +lemma Hcheck_closed : "\y\M. \g\M. function(g) \ Hcheck(y,g)\M" +proof - + have "Replace(y,PHcheck(##M,\,f))\M" if "f\M" "y\M" for f y + using that repl_PHcheck PHcheck_closed[of y f] univ_PHcheck + strong_replacement_closed + by (simp flip: setclass_iff) + then + show ?thesis + using def_PHcheck by auto +qed + +lemma wf_rcheck : "x\M \ wf(rcheck(x))" + unfolding rcheck_def using wf_trancl[OF wf_Memrel] . + +lemma trans_rcheck : "x\M \ trans(rcheck(x))" + unfolding rcheck_def using trans_trancl . + +lemma relation_rcheck : "x\M \ relation(rcheck(x))" + unfolding rcheck_def using relation_trancl . + +lemma check_in_M : "x\M \ check(x) \ M" + unfolding transrec_def + using wfrec_Hcheck[of x] check_trancl wf_rcheck trans_rcheck relation_rcheck rcheck_in_M + Hcheck_closed relation2_Hcheck trans_wfrec_closed[of "rcheck(x)" x "is_Hcheck(##M,\)" Hcheck] + by (simp flip: setclass_iff) + +(* Internalization and absoluteness of rcheck\ *) + +lemma rcheck_abs[Rel] : "\ x\M ; r\M \ \ is_rcheck(##M,x,r) \ r = rcheck(x)" + unfolding rcheck_def is_rcheck_def + using singleton_closed trancl_closed Memrel_closed eclose_closed zero_in_M + by simp + +lemma check_abs[Rel] : + assumes "x\M" "z\M" + shows "is_check(##M,\,x,z) \ z = check(x)" +proof - + have "is_check(##M,\,x,z) \ is_wfrec(##M,is_Hcheck(##M,\),rcheck(x),x,z)" + unfolding is_check_def + using assms rcheck_abs rcheck_in_M zero_in_M + unfolding check_trancl is_check_def + by simp + then + show ?thesis + unfolding check_trancl + using assms wfrec_Hcheck[of x] wf_rcheck trans_rcheck relation_rcheck rcheck_in_M + Hcheck_closed relation2_Hcheck trans_wfrec_abs[of "rcheck(x)" x z "is_Hcheck(##M,\)" Hcheck] + by (simp flip: setclass_iff) +qed + +lemma check_replacement: "{check(x). x\P} \ M" +proof - + have "arity(check_fm(0,2,1)) = 3" + by (simp add:ord_simp_union arity) + then + show ?thesis + using sats_check_fm check_abs P_in_M check_in_M one_in_M transitivity zero_in_M + Replace_relativized_in_M[of "check_fm(0,2,1)" "[\]" _ "is_check(##M,\)" check] + check_fm_type replacement_ax1(12) + by simp +qed + +lemma M_subset_MG : "\ \ G \ M \ M[G]" + using check_in_M one_in_P GenExtI + by (intro subsetI, subst valcheck [of G,symmetric], auto) + +text\The name for the generic filter\ +definition + G_dot :: "i" where + "G_dot \ {\check(p),p\ . p\P}" + +lemma G_dot_in_M : "G_dot \ M" +proof - + let ?is_pcheck = "\x y. \ch\M. is_check(##M,\,x,ch) \ pair(##M,ch,x,y)" + let ?pcheck_fm = "Exists(And(check_fm(1,3,0),pair_fm(0,1,2)))" + have "sats(M,?pcheck_fm,[x,y,\]) \ ?is_pcheck(x,y)" if "x\M" "y\M" for x y + using sats_check_fm that one_in_M zero_in_M by simp + moreover + have "?is_pcheck(x,y) \ y = \check(x),x\" if "x\M" "y\M" for x y + using that check_abs check_in_M by simp + moreover + have "?pcheck_fm\formula" + by simp + moreover + have "arity(?pcheck_fm)=3" + by (simp add:ord_simp_union arity) + moreover + from P_in_M check_in_M pair_in_M_iff P_sub_M + have "\check(p),p\ \ M" if "p\P" for p + using that by auto + ultimately + show ?thesis + unfolding G_dot_def + using one_in_M P_in_M transitivity Replace_relativized_in_M[of ?pcheck_fm "[\]"] + replacement_ax1(13) + by simp +qed + +lemma val_G_dot : + assumes "G \ P" "\ \ G" + shows "val(P,G,G_dot) = G" +proof (intro equalityI subsetI) + fix x + assume "x\val(P,G,G_dot)" + then obtain \ p where "p\G" "\\,p\ \ G_dot" "val(P,G,\) = x" "\ = check(p)" + unfolding G_dot_def using elem_of_val_pair G_dot_in_M + by force + with \\\G\ \G\P\ + show "x \ G" + using valcheck P_sub_M by auto +next + fix p + assume "p\G" + have "\check(q),q\ \ G_dot" if "q\P" for q + unfolding G_dot_def using that by simp + with \p\G\ \G\P\ + have "val(P,G,check(p)) \ val(P,G,G_dot)" + using val_of_elem G_dot_in_M by blast + with \p\G\ \G\P\ \\\G\ + show "p \ val(P,G,G_dot)" + using P_sub_M valcheck by auto +qed + +lemma G_in_Gen_Ext : + assumes "G \ P" "\ \ G" + shows "G \ M[G]" + using assms val_G_dot GenExtI[of _ G] G_dot_in_M + by force + +end \ \\<^locale>\forcing_data1\\ + +locale G_generic1 = forcing_data1 + + fixes G :: "i" + assumes generic : "M_generic(G)" +begin + +lemma zero_in_MG : + "0 \ M[G]" +proof - + have "0 = val(P,G,0)" + using zero_in_M elem_of_val by auto + also + have "... \ M[G]" + using GenExtI zero_in_M by simp + finally + show ?thesis . +qed + +lemma G_nonempty: "G\0" + using generic subset_refl[of P] P_in_M P_dense + unfolding M_generic_def + by auto + +end \ \\<^locale>\G_generic1\\ + +locale G_generic1_AC = G_generic1 + M_ctm1_AC + +end \ No newline at end of file diff --git a/thys/Independence_CH/Not_CH.thy b/thys/Independence_CH/Not_CH.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Not_CH.thy @@ -0,0 +1,684 @@ +section\Model of the negation of the Continuum Hypothesis\ + +theory Not_CH + imports + Cardinal_Preservation +begin + +txt\We are taking advantage that the poset of finite functions is absolute, +and thus we work with the unrelativized \<^term>\Fn\. But it would have been more +appropriate to do the following using the relative \<^term>\Fn_rel\. As it turns +out, the present theory was developed prior to having \<^term>\Fn\ relativized! + +We also note that \<^term>\Fn(\,\\\,2)\ is separative, i.e. each \<^term>\X \ Fn(\,\\\,2)\ +has two incompatible extensions; therefore we may recover part of our previous theorem +@{thm [source] extensions_of_ctms_ZF}. But that result also included the possibility +of not having $\AC$ in the ground model, which would not be sensible in a context +where the cardinality of the continuum is under discussion. It is also the case that +@{thm [source] extensions_of_ctms_ZF} was historically our first formalized result +(with a different proof) that showed the forcing machinery had all of its elements +in place.\ + +abbreviation + Add_subs :: "i \ i" where + "Add_subs(\) \ Fn(\,\\\,2)" + +abbreviation + Add_le :: "i \ i" where + "Add_le(\) \ Fnle(\,\ \ \,2)" + +lemma (in M_aleph) Aleph_rel2_closed[intro,simp]: "M(\\<^bsub>2\<^esub>\<^bsup>M\<^esup>)" + using nat_into_Ord by simp + +locale M_master = M_cohen + M_library_DC + + assumes + UN_lepoll_assumptions: + "M(A) \ M(b) \ M(f) \ M(A') \ separation(M, \y. \x\A'. y = \x, \ i. x\if_range_F_else_F((`)(A), b, f, i)\)" + +subsection\Non-absolute concepts between extensions\ + +locale M_master_sub = M_master + N:M_master N for N + + assumes + M_imp_N: "M(x) \ N(x)" and + Ord_iff: "Ord(x) \ M(x) \ N(x)" + +sublocale M_master_sub \ M_N_Perm + using M_imp_N by unfold_locales + +context M_master_sub +begin + +lemma cardinal_rel_le_cardinal_rel: "M(X) \ |X|\<^bsup>N\<^esup> \ |X|\<^bsup>M\<^esup>" + using M_imp_N N.lepoll_rel_cardinal_rel_le[OF lepoll_rel_transfer Card_rel_is_Ord] + cardinal_rel_eqpoll_rel[THEN eqpoll_rel_sym, THEN eqpoll_rel_imp_lepoll_rel] + by simp + +lemma Aleph_rel_sub_closed: "Ord(\) \ M(\) \ N(\\<^bsub>\\<^esub>\<^bsup>M\<^esup>)" + using Ord_iff[THEN iffD1, OF Card_rel_Aleph_rel[THEN Card_rel_is_Ord]] + by simp + +lemma Card_rel_imp_Card_rel: "Card\<^bsup>N\<^esup>(\) \ M(\) \ Card\<^bsup>M\<^esup>(\)" + using N.Card_rel_is_Ord[of \] M_imp_N Ord_cardinal_rel_le[of \] + cardinal_rel_le_cardinal_rel[of \] le_anti_sym + unfolding Card_rel_def by auto + +lemma csucc_rel_le_csucc_rel: + assumes "Ord(\)" "M(\)" + shows "(\\<^sup>+)\<^bsup>M\<^esup> \ (\\<^sup>+)\<^bsup>N\<^esup>" +proof - + note assms + moreover from this + have "N(L) \ Card\<^bsup>N\<^esup>(L) \ \ < L \ M(L) \ Card\<^bsup>M\<^esup>(L) \ \ < L" + (is "?P(L) \ ?Q(L)") for L + using M_imp_N Ord_iff[THEN iffD2, of L] N.Card_rel_is_Ord lt_Ord + Card_rel_imp_Card_rel by auto + moreover from assms + have "N((\\<^sup>+)\<^bsup>N\<^esup>)" "Card\<^bsup>N\<^esup>((\\<^sup>+)\<^bsup>N\<^esup>)" "\ < (\\<^sup>+)\<^bsup>N\<^esup>" + using N.lt_csucc_rel[of \] N.Card_rel_csucc_rel[of \] M_imp_N by simp_all + ultimately + show ?thesis + using M_imp_N Least_antitone[of _ ?P ?Q] unfolding csucc_rel_def by blast +qed + +lemma Aleph_rel_le_Aleph_rel: "Ord(\) \ M(\) \ \\<^bsub>\\<^esub>\<^bsup>M\<^esup> \ \\<^bsub>\\<^esub>\<^bsup>N\<^esup>" +proof (induct rule:trans_induct3) + case 0 + then + show ?case + using Aleph_rel_zero N.Aleph_rel_zero by simp +next + case (succ x) + then + have "\\<^bsub>x\<^esub>\<^bsup>M\<^esup> \ \\<^bsub>x\<^esub>\<^bsup>N\<^esup>" "Ord(x)" "M(x)" by simp_all + moreover from this + have "(\\<^bsub>x\<^esub>\<^bsup>M\<^esup>\<^sup>+)\<^bsup>M\<^esup> \ (\\<^bsub>x\<^esub>\<^bsup>N\<^esup>\<^sup>+)\<^bsup>M\<^esup>" + using M_imp_N Ord_iff[THEN iffD2, OF N.Card_rel_is_Ord] + by (intro csucc_rel_le_mono) simp_all + moreover from calculation + have "(\\<^bsub>x\<^esub>\<^bsup>N\<^esup>\<^sup>+)\<^bsup>M\<^esup> \ (\\<^bsub>x\<^esub>\<^bsup>N\<^esup>\<^sup>+)\<^bsup>N\<^esup>" + using M_imp_N N.Card_rel_is_Ord Ord_iff[THEN iffD2, OF N.Card_rel_is_Ord] + by (intro csucc_rel_le_csucc_rel) auto + ultimately + show ?case + using M_imp_N Aleph_rel_succ N.Aleph_rel_succ csucc_rel_le_csucc_rel + le_trans by auto +next + case (limit x) + then + show ?case + using M_imp_N Aleph_rel_limit N.Aleph_rel_limit + by simp (blast dest: transM intro!:le_implies_UN_le_UN) +qed + +end \ \\<^locale>\M_master_sub\\ + +lemmas (in M_ZF3_trans) sep_instances = + separation_insnd_ballPair + separation_ifrangeF_body separation_ifrangeF_body2 separation_ifrangeF_body3 + separation_ifrangeF_body4 separation_ifrangeF_body5 separation_ifrangeF_body6 + separation_ifrangeF_body7 separation_cardinal_rel_lesspoll_rel + separation_is_dcwit_body + +lemmas (in M_ZF3_trans) repl_instances = lam_replacement_inj_rel + lam_replacement_cardinal replacement_trans_apply_image + +sublocale M_ZFC3_trans \ M_master "##M" + using replacement_dcwit_repl_body\ \this is another replacement instance\ + by unfold_locales (simp_all add:repl_instances sep_instances del:setclass_iff + add: transrec_replacement_def wfrec_replacement_def dcwit_repl_body_def) + +subsection\Cohen forcing is ccc\ + +context M_ctm3_AC +begin + +lemma ccc_Add_subs_Aleph_2: "ccc\<^bsup>M\<^esup>(Add_subs(\\<^bsub>2\<^esub>\<^bsup>M\<^esup>),Add_le(\\<^bsub>2\<^esub>\<^bsup>M\<^esup>))" +proof - + interpret M_add_reals "##M" "\\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \" + by unfold_locales blast + show ?thesis + using ccc_rel_Fn_nat by fast +qed + +end \ \\<^locale>\M_ctm3_AC\\ + +sublocale G_generic4_AC \ M_master_sub "##M" "##(M[G])" + using M_subset_MG[OF one_in_G] generic Ord_MG_iff + by unfold_locales auto + +lemma (in M_trans) mem_F_bound4: + fixes F A + defines "F \ (`)" + shows "x\F(A,c) \ c \ (range(f) \ domain(A))" + using apply_0 unfolding F_def + by (cases "M(c)", auto simp:F_def) + +lemma (in M_trans) mem_F_bound5: + fixes F A + defines "F \ \_ x. A`x " + shows "x\F(A,c) \ c \ (range(f) \ domain(A))" + using apply_0 unfolding F_def + by (cases "M(c)", auto simp:F_def drSR_Y_def dC_F_def) + +sublocale M_ctm3_AC \ M_replacement_lepoll "##M" "(`)" + using UN_lepoll_assumptions lam_replacement_apply lam_replacement_inj_rel + mem_F_bound4 apply_0 + unfolding lepoll_assumptions_defs +proof (unfold_locales, + rule_tac [3] lam_Least_assumption_general[where U=domain, OF _ mem_F_bound4], simp_all) + fix A i x + assume "A \ M" "x \ M" "x \ A ` i" + then + show "i \ M" + using apply_0[of i A] transM[of _ "domain(A)", simplified] + by force +qed + +context G_generic4_AC begin + +context + includes G_generic1_lemmas +begin + +lemma G_in_MG: "G \ M[G]" + using G_in_Gen_Ext[ OF _ one_in_G, OF _ generic] + by blast + +lemma ccc_preserves_Aleph_succ: + assumes "ccc\<^bsup>M\<^esup>(P,leq)" "Ord(z)" "z \ M" + shows "Card\<^bsup>M[G]\<^esup>(\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>)" +proof (rule ccontr) + assume "\ Card\<^bsup>M[G]\<^esup>(\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>)" + moreover + note \z \ M\ \Ord(z)\ + moreover from this + have "Ord(\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>)" + using Card_rel_is_Ord by fastforce + ultimately + obtain \ f where "\ < \\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>" "f \ surj\<^bsup>M[G]\<^esup>(\, \\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>)" + using ext.lt_surj_rel_empty_imp_Card_rel M_subset_MG[OF one_in_G, OF generic] + by force + moreover from this and \z\M\ \Ord(z)\ + have "\ \ M" "f \ M[G]" + using ext.trans_surj_rel_closed + by (auto dest:transM ext.transM dest!:ltD) + moreover + note \ccc\<^bsup>M\<^esup>(P,leq)\ \z\M\ + ultimately + obtain F where "F:\\Pow\<^bsup>M\<^esup>(\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>)" "\\\\. f`\ \ F`\" "\\\\. |F`\|\<^bsup>M\<^esup> \ \" + "F \ M" + using ccc_fun_approximation_lemma[of \ "\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>" f] + ext.mem_surj_abs[of f \ "\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>"] \Ord(z)\ + surj_is_fun[of f \ "\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>"] by auto + then + have "\ \ \ \ |F`\|\<^bsup>M\<^esup> \ \\<^bsub>0\<^esub>\<^bsup>M\<^esup>" for \ + using Aleph_rel_zero by simp + have "w \ F ` x \ x \ M" for w x + proof - + fix w x + assume "w \ F`x" + then + have "x \ domain(F)" + using apply_0 by auto + with \F:\\Pow\<^bsup>M\<^esup>(\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>)\ + have "x \ \" + using domain_of_fun by simp + with \\ \ M\ + show "x \ M" by (auto dest:transM) + qed + with \\ \ M\ \F:\\Pow\<^bsup>M\<^esup>(\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>)\ \F\M\ + interpret M_cardinal_UN_lepoll "##M" "\\. F`\" \ + using UN_lepoll_assumptions lepoll_assumptions + lam_replacement_apply lam_replacement_inj_rel + proof (unfold_locales, auto dest:transM simp del:if_range_F_else_F_def) + fix f b + assume "b\M" "f\M" + with \F\M\ + show "lam_replacement(##M, \x. \ i. x \ if_range_F_else_F((`)(F), b, f, i))" + using UN_lepoll_assumptions mem_F_bound5 + by (rule_tac lam_Least_assumption_general[where U="domain", OF _ mem_F_bound5]) + simp_all + qed + from \\ < \\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>\ \\ \ M\ assms + have "\ \\<^bsup>M\<^esup> \\<^bsub>z\<^esub>\<^bsup>M\<^esup>" + using + Aleph_rel_zero + cardinal_rel_lt_csucc_rel_iff[of "\\<^bsub>z\<^esub>\<^bsup>M\<^esup>" \] + le_Card_rel_iff[of "\\<^bsub>z\<^esub>\<^bsup>M\<^esup>" \] + Aleph_rel_succ[of z] Card_rel_lt_iff[of \ "\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>"] + lt_Ord[of \ "\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>"] + Card_rel_csucc_rel[of "\\<^bsub>z\<^esub>\<^bsup>M\<^esup>"] + Aleph_rel_closed[of z] + Card_rel_Aleph_rel[THEN Card_rel_is_Ord, OF _ _ Aleph_rel_closed] + by simp + with \\ < \\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>\ \\\\\. |F`\|\<^bsup>M\<^esup> \ \\ \\ \ M\ assms + have "|\\\\. F`\|\<^bsup>M\<^esup> \ \\<^bsub>z\<^esub>\<^bsup>M\<^esup>" + using InfCard_rel_Aleph_rel[of z] Aleph_rel_zero + subset_imp_lepoll_rel[THEN lepoll_rel_imp_cardinal_rel_le, + of "\\\\. F`\" "\\<^bsub>z\<^esub>\<^bsup>M\<^esup>"] Aleph_rel_succ + Aleph_rel_increasing[THEN leI, THEN [2] le_trans, of _ 0 z] + Ord_0_lt_iff[THEN iffD1, of z] + by (cases "0z\M\ \Ord(z)\ + moreover from \\\\\. f`\ \ F`\\ \f \ surj\<^bsup>M[G]\<^esup>(\, \\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>)\ + \\ \ M\ \f \ M[G]\ and this + have "\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup> \ (\\\\. F`\)" + using ext.mem_surj_abs by (force simp add:surj_def) + moreover from \F \ M\ \\ \ M\ + have "(\x\\. F ` x) \ M" + using j.B_replacement\ \NOTE: it didn't require @{thm j.UN_closed} before!\ + by (intro Union_closed[simplified] RepFun_closed[simplified]) + (auto dest:transM) + ultimately + have "\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup> \ \\<^bsub>z\<^esub>\<^bsup>M\<^esup>" + using subset_imp_le_cardinal_rel[of "\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>" "\\\\. F`\"] + le_trans by auto + with assms + show "False" + using Aleph_rel_increasing not_le_iff_lt[of "\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>" "\\<^bsub>z\<^esub>\<^bsup>M\<^esup>"] + Card_rel_Aleph_rel[THEN Card_rel_is_Ord] + by auto +qed + +end \ \bundle G\_generic1\_lemmas\ + +end \ \\<^locale>\G_generic4_AC\\ + +context M_ctm1 +begin + +abbreviation + Add :: "i" where + "Add \ Fn(\, \\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \, 2)" + +end \ \\<^locale>\M_ctm1\\ + +locale add_generic4 = G_generic4_AC "Fn(\, \\<^bsub>2\<^esub>\<^bsup>##M\<^esup> \ \, 2)" "Fnle(\, \\<^bsub>2\<^esub>\<^bsup>##M\<^esup> \ \, 2)" 0 + +sublocale add_generic4 \ cohen_data \ "\\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \" 2 by unfold_locales auto + +context add_generic4 +begin + +notation Leq (infixl "\" 50) +notation Incompatible (infixl "\" 50) +notation GenExt_at_P ("_[_]" [71,1]) + +lemma Add_subs_preserves_Aleph_succ: "Ord(z) \ z\M \ Card\<^bsup>M[G]\<^esup>(\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>)" + using ccc_preserves_Aleph_succ ccc_Add_subs_Aleph_2 + by auto + +lemma Aleph_rel_nats_MG_eq_Aleph_rel_nats_M: + includes G_generic1_lemmas + assumes "z \ \" + shows "\\<^bsub>z\<^esub>\<^bsup>M[G]\<^esup> = \\<^bsub>z\<^esub>\<^bsup>M\<^esup>" + using assms +proof (induct) + case 0 + have "\\<^bsub>0\<^esub>\<^bsup>M[G]\<^esup> = \" + using ext.Aleph_rel_zero . + also + have "\ = \\<^bsub>0\<^esub>\<^bsup>M\<^esup>" + using Aleph_rel_zero by simp + finally + show ?case . +next + case (succ z) + then + have "\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup> \ \\<^bsub>succ(z)\<^esub>\<^bsup>M[G]\<^esup>" + using Aleph_rel_le_Aleph_rel nat_into_M by simp + moreover from \z \ \\ + have "\\<^bsub>z\<^esub>\<^bsup>M\<^esup> \ M[G]" "\\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup> \ M[G]" + using nat_into_M by simp_all + moreover from this and \\\<^bsub>z\<^esub>\<^bsup>M[G]\<^esup> = \\<^bsub>z\<^esub>\<^bsup>M\<^esup>\ \z \ \\ + have "\\<^bsub>succ(z)\<^esub>\<^bsup>M[G]\<^esup> \ \\<^bsub>succ(z)\<^esub>\<^bsup>M\<^esup>" + using ext.Aleph_rel_succ nat_into_M + Add_subs_preserves_Aleph_succ[THEN ext.csucc_rel_le, of z] + Aleph_rel_increasing[of z "succ(z)"] + by simp + ultimately + show ?case using le_anti_sym by blast +qed + +abbreviation + f_G :: "i" (\f\<^bsub>G\<^esub>\) where + "f\<^bsub>G\<^esub> \ \G" + +abbreviation + dom_dense :: "i\i" where + "dom_dense(x) \ { p\Add . x \ domain(p) }" + +(* TODO: write general versions of this for \<^term>\Fn(\,I,J)\ *) +lemma dense_dom_dense: "x \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \ \ dense(dom_dense(x))" +proof + fix p + assume "x \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \" "p \ Add" + show "\d\dom_dense(x). d \ p" + proof (cases "x \ domain(p)") + case True + with \x \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \\ \p \ Add\ + show ?thesis using refl_leq by auto + next + case False + note \p \ Add\ + moreover from this and False and \x \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \\ + have "cons(\x,0\, p) \ Add" + using FiniteFun.consI[of x "\\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \" 0 2 p] + Fn_nat_eq_FiniteFun by auto + moreover from \p \ Add\ + have "x\domain(cons(\x,0\, p))" by simp + ultimately + show ?thesis + by (fastforce del:FnD) + qed +qed + +declare (in M_ctm3_AC) Fn_nat_closed[simplified setclass_iff, simp, intro] +declare (in M_ctm3_AC) Fnle_nat_closed[simp del, rule del, + simplified setclass_iff, simp, intro] +declare (in M_ctm3_AC) cexp_rel_closed[simplified setclass_iff, simp, intro] +declare (in G_generic4_AC) ext.cexp_rel_closed[simplified setclass_iff, simp, intro] + +lemma dom_dense_closed[intro,simp]: "x \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \ \ dom_dense(x) \ M" + using separation_in_domain[of x] nat_into_M + by (rule_tac separation_closed[simplified], blast dest:transM) simp + +lemma domain_f_G: assumes "x \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup>" "y \ \" + shows "\x, y\ \ domain(f\<^bsub>G\<^esub>)" +proof - + from assms + have "dense(dom_dense(\x, y\))" using dense_dom_dense by simp + with assms + obtain p where "p\dom_dense(\x, y\)" "p\G" + using generic[THEN M_generic_denseD, of "dom_dense(\x, y\)"] + by auto + then + show "\x, y\ \ domain(f\<^bsub>G\<^esub>)" by blast +qed + +lemma f_G_funtype: + includes G_generic1_lemmas + shows "f\<^bsub>G\<^esub> : \\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \ \ 2" + using generic domain_f_G + unfolding Pi_def +proof (auto) + show "x \ B \ B \ G \ x \ (\\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \) \ 2" for B x + using Fn_nat_subset_Pow by blast + show "function(f\<^bsub>G\<^esub>)" + using Un_filter_is_function generic + unfolding M_generic_def by fast +qed + +abbreviation + inj_dense :: "i\i\i" where + "inj_dense(w,x) \ + { p\Add . (\n\\. \\w,n\,1\ \ p \ \\x,n\,0\ \ p) }" + +(* TODO write general versions of this for \<^term>\Fn(\,I,J)\ *) +lemma dense_inj_dense: + assumes "w \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup>" "x \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup>" "w \ x" + shows "dense(inj_dense(w,x))" +proof + fix p + assume "p \ Add" + then + obtain n where "\w,n\ \ domain(p)" "\x,n\ \ domain(p)" "n \ \" + proof - + { + assume "\w,n\ \ domain(p) \ \x,n\ \ domain(p)" if "n \ \" for n + then + have "\ \ range(domain(p))" by blast + then + have "\ Finite(p)" + using Finite_range Finite_domain subset_Finite nat_not_Finite + by auto + with \p \ Add\ + have False + using Fn_nat_eq_FiniteFun FiniteFun.dom_subset[of "\\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \" 2] + Fin_into_Finite by auto + } + with that\ \the shape of the goal puts assumptions in this variable\ + show ?thesis by auto + qed + moreover + note \p \ Add\ assms + moreover from calculation + have "cons(\\x,n\,0\, p) \ Add" + using FiniteFun.consI[of "\x,n\" "\\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \" 0 2 p] + Fn_nat_eq_FiniteFun by auto + ultimately + have "cons(\\w,n\,1\, cons(\\x,n\,0\, p) ) \ Add" + using FiniteFun.consI[of "\w,n\" "\\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \" 1 2 "cons(\\x,n\,0\, p)"] + Fn_nat_eq_FiniteFun by auto + with \n \ \\ + show "\d\inj_dense(w,x). d \ p" + using \p \ Add\ by (intro bexI) auto +qed + +lemma inj_dense_closed[intro,simp]: + "w \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ x \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ inj_dense(w,x) \ M" + using transM[OF _ Aleph_rel2_closed] separation_conj separation_bex + lam_replacement_hcomp2[OF _ _ _ _ lam_replacement_Pair] + separation_in lam_replacement_fst lam_replacement_snd lam_replacement_constant + lam_replacement_hcomp[OF lam_replacement_snd lam_replacement_restrict'] + separation_bex separation_conj + by simp + +lemma Aleph_rel2_new_reals: + assumes "w \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup>" "x \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup>" "w \ x" + shows "(\n\\. f\<^bsub>G\<^esub> ` \w, n\) \ (\n\\. f\<^bsub>G\<^esub> ` \x, n\)" +proof - + from assms + have "dense(inj_dense(w,x))" using dense_inj_dense by simp + with assms + obtain p where "p\inj_dense(w,x)" "p\G" + using generic[THEN M_generic_denseD, of "inj_dense(w,x)"] + by blast + then + obtain n where "n \ \" "\\w, n\, 1\ \ p" "\\x, n\, 0\ \ p" + by blast + moreover from this and \p\G\ + have "\\w, n\, 1\ \ f\<^bsub>G\<^esub>" "\\x, n\, 0\ \ f\<^bsub>G\<^esub>" by auto + moreover from calculation + have "f\<^bsub>G\<^esub> ` \w, n\ = 1" "f\<^bsub>G\<^esub> ` \x, n\ = 0" + using f_G_funtype apply_equality + by auto + ultimately + have "(\n\\. f\<^bsub>G\<^esub> ` \w, n\) ` n \ (\n\\. f\<^bsub>G\<^esub> ` \x, n\) ` n" + by simp + then + show ?thesis by fastforce +qed + +definition + h_G :: "i" (\h\<^bsub>G\<^esub>\) where + "h\<^bsub>G\<^esub> \ \\\\\<^bsub>2\<^esub>\<^bsup>M\<^esup>. \n\\. f\<^bsub>G\<^esub>`\\,n\" + +lemma h_G_in_MG[simp]: + includes G_generic1_lemmas + shows "h\<^bsub>G\<^esub> \ M[G]" + using ext.lam_apply_replacement ext.apply_replacement2 + ext.lam_apply_replacement[unfolded lam_replacement_def] + ext.Union_closed[simplified, OF G_in_MG] + \ \The "simplified" here is because of + the \<^term>\setclass\ ocurrences\ + ext.nat_into_M + unfolding h_G_def + by (rule_tac ext.lam_closed[simplified] | + auto dest:transM del:ext.cexp_rel_closed[simplified])+ + +lemma h_G_inj_Aleph_rel2_reals: "h\<^bsub>G\<^esub> \ inj\<^bsup>M[G]\<^esup>(\\<^bsub>2\<^esub>\<^bsup>M\<^esup>, \ \\<^bsup>M[G]\<^esup> 2)" + using Aleph_rel_sub_closed +proof (intro ext.mem_inj_abs[THEN iffD2]) + show "h\<^bsub>G\<^esub> \ inj(\\<^bsub>2\<^esub>\<^bsup>M\<^esup>, \ \\<^bsup>M[G]\<^esup> 2)" + unfolding inj_def + proof (intro ballI CollectI impI) + show "h\<^bsub>G\<^esub> \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \ \\<^bsup>M[G]\<^esup> 2" + using f_G_funtype G_in_MG ext.nat_into_M + unfolding h_G_def + apply (intro lam_type ext.mem_function_space_rel_abs[THEN iffD2], simp_all) + apply (rule_tac ext.lam_closed[simplified], simp_all) + apply (rule ext.apply_replacement2) + apply (auto dest:ext.transM[OF _ Aleph_rel_sub_closed]) + done + fix w x + assume "w \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup>" "x \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup>" "h\<^bsub>G\<^esub> ` w = h\<^bsub>G\<^esub> ` x" + then + show "w = x" + unfolding h_G_def using Aleph_rel2_new_reals by auto + qed +qed simp_all + +lemma Aleph2_extension_le_continuum_rel: + includes G_generic1_lemmas + shows "\\<^bsub>2\<^esub>\<^bsup>M[G]\<^esup> \ 2\<^bsup>\\\<^bsub>0\<^esub>\<^bsup>M[G]\<^esup>,M[G]\<^esup>" +proof - + have "\\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ M[G]" "Ord(\\<^bsub>2\<^esub>\<^bsup>M\<^esup>)" + using Card_rel_is_Ord by auto + moreover from this + have "\\<^bsub>2\<^esub>\<^bsup>M\<^esup> \\<^bsup>M[G]\<^esup> \ \\<^bsup>M[G]\<^esup> 2" + using ext.def_lepoll_rel[of "\\<^bsub>2\<^esub>\<^bsup>M\<^esup>" "\ \\<^bsup>M[G]\<^esup> 2"] + h_G_inj_Aleph_rel2_reals by auto + moreover from calculation + have "\\<^bsub>2\<^esub>\<^bsup>M\<^esup> \\<^bsup>M[G]\<^esup> |\ \\<^bsup>M[G]\<^esup> 2|\<^bsup>M[G]\<^esup>" + using ext.lepoll_rel_imp_lepoll_rel_cardinal_rel by simp + ultimately + have "|\\<^bsub>2\<^esub>\<^bsup>M\<^esup>|\<^bsup>M[G]\<^esup> \ 2\<^bsup>\\\<^bsub>0\<^esub>\<^bsup>M[G]\<^esup>,M[G]\<^esup>" + using ext.lepoll_rel_imp_cardinal_rel_le[of "\\<^bsub>2\<^esub>\<^bsup>M\<^esup>" "\ \\<^bsup>M[G]\<^esup> 2", + OF _ _ ext.function_space_rel_closed] + ext.Aleph_rel_zero Aleph_rel_nats_MG_eq_Aleph_rel_nats_M + unfolding cexp_rel_def by simp + then + show "\\<^bsub>2\<^esub>\<^bsup>M[G]\<^esup> \ 2\<^bsup>\\\<^bsub>0\<^esub>\<^bsup>M[G]\<^esup>,M[G]\<^esup>" + using Aleph_rel_nats_MG_eq_Aleph_rel_nats_M + ext.Card_rel_Aleph_rel[of 2, THEN ext.Card_rel_cardinal_rel_eq] + by simp +qed + +lemma Aleph_rel_lt_continuum_rel: "\\<^bsub>1\<^esub>\<^bsup>M[G]\<^esup> < 2\<^bsup>\\\<^bsub>0\<^esub>\<^bsup>M[G]\<^esup>,M[G]\<^esup>" + using Aleph2_extension_le_continuum_rel + ext.Aleph_rel_increasing[of 1 2] le_trans by auto + +corollary not_CH: "\\<^bsub>1\<^esub>\<^bsup>M[G]\<^esup> \ 2\<^bsup>\\\<^bsub>0\<^esub>\<^bsup>M[G]\<^esup>,M[G]\<^esup>" + using Aleph_rel_lt_continuum_rel by auto + +end \ \\<^locale>\add_generic4\\ + +subsection\Models of fragments of $\ZFC + \neg \CH$\ + +definition + ContHyp :: "o" where + "ContHyp \ \\<^bsub>1\<^esub> = 2\<^bsup>\\\<^bsub>0\<^esub>\<^esup>" + +relativize functional "ContHyp" "ContHyp_rel" +notation ContHyp_rel (\CH\<^bsup>_\<^esup>\) +relationalize "ContHyp_rel" "is_ContHyp" + +context M_master +begin + +is_iff_rel for "ContHyp" + using is_cexp_iff is_Aleph_iff[of 0] is_Aleph_iff[of 1] + unfolding is_ContHyp_def ContHyp_rel_def + by (auto simp del:setclass_iff) (rule rexI[of _ _ M, OF _ nonempty], auto) + +end \ \\<^locale>\M_master\\ + +synthesize "is_ContHyp" from_definition assuming "nonempty" +arity_theorem for "is_ContHyp_fm" + +notation is_ContHyp_fm (\\CH\\) + +theorem ctm_of_not_CH: + assumes + "M \ \" "Transset(M)" "M \ ZC \ {\Replacement(p)\ . p \ overhead}" + "\ \ formula" "M \ { \Replacement(ground_repl_fm(\))\ . \ \ \}" + shows + "\N. + M \ N \ N \ \ \ Transset(N) \ N \ ZC \ {\\\CH\\} \ { \Replacement(\)\ . \ \ \} \ + (\\. Ord(\) \ (\ \ M \ \ \ N))" +proof - + from \M \ ZC \ {\Replacement(p)\ . p \ overhead}\ + interpret M_ZFC4 M + using M_satT_overhead_imp_M_ZF4 by simp + from \Transset(M)\ + interpret M_ZFC4_trans M + using M_satT_imp_M_ZF4 + by unfold_locales + from \M \ \\ + obtain enum where "enum \ bij(\,M)" + using eqpoll_sym unfolding eqpoll_def by blast + then + interpret M_ctm4_AC M enum by unfold_locales + interpret cohen_data \ "\\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ \" 2 by unfold_locales auto + have "Add \ M" "Add_le(\\<^bsub>2\<^esub>\<^bsup>M\<^esup>) \ M" + using nat_into_M Aleph_rel_closed M_nat cartprod_closed Fn_nat_closed Fnle_nat_closed + by simp_all + then + interpret forcing_data1 "Add" "Add_le(\\<^bsub>2\<^esub>\<^bsup>M\<^esup>)" 0 M enum + by unfold_locales simp_all + obtain G where "M_generic(G)" + using generic_filter_existence[OF one_in_P] + by auto + moreover from this + interpret add_generic4 M enum G by unfold_locales + have "\ (\\<^bsub>1\<^esub>\<^bsup>M[G]\<^esup> = 2\<^bsup>\\\<^bsub>0\<^esub>\<^bsup>M[G]\<^esup>,M[G]\<^esup>)" + using not_CH . + then + have "M[G], [] \ \\\CH\\" + using ext.is_ContHyp_iff + by (simp add:ContHyp_rel_def) + then + have "M[G] \ ZC \ {\\\CH\\}" + using ext.M_satT_ZC by auto + moreover + have "Transset(M[G])" using Transset_MG . + moreover + have "M \ M[G]" using M_subset_MG[OF one_in_G] generic by simp + moreover + note \M \ { \Replacement(ground_repl_fm(\))\ . \ \ \}\ \\ \ formula\ + ultimately + show ?thesis + using Ord_MG_iff MG_eqpoll_nat satT_ground_repl_fm_imp_satT_ZF_replacement_fm[of \] + by (rule_tac x="M[G]" in exI, blast) +qed + +lemma ZF_replacement_overhead_sub_ZFC: "{\Replacement(p)\ . p \ overhead} \ ZFC" + using overhead_type unfolding ZFC_def ZF_def ZF_schemes_def by auto + +corollary ctm_ZFC_imp_ctm_not_CH: + assumes + "M \ \" "Transset(M)" "M \ ZFC" + shows + "\N. + M \ N \ N \ \ \ Transset(N) \ N \ ZFC \ {\\\CH\\} \ + (\\. Ord(\) \ (\ \ M \ \ \ N))" +proof- + from assms + have "\N. + M \ N \ + N \ \ \ + Transset(N) \ + N \ ZC \ N \ {\\\CH\\} \ N \ {\Replacement(x)\ . x \ formula} \ (\\. Ord(\) \ \ \ M \ \ \ N)" + using ctm_of_not_CH[of M formula] satT_ZFC_imp_satT_ZC[of M] + satT_mono[OF _ ground_repl_fm_sub_ZFC, of M] + satT_mono[OF _ ZF_replacement_overhead_sub_ZFC, of M] + satT_mono[OF _ ZF_replacement_fms_sub_ZFC, of M] + by (simp add: satT_Un_iff) + then + obtain N where "N \ ZC" "N \ {\\\CH\\}" "N \ {\Replacement(x)\ . x \ formula}" + "M \ N" "N \ \" "Transset(N)" "(\\. Ord(\) \ \ \ M \ \ \ N)" + by auto + moreover from this + have "N \ ZFC" + using satT_ZC_ZF_replacement_imp_satT_ZFC + by auto + moreover from this and \N \ {\\\CH\\}\ + have "N \ ZFC \ {\\\CH\\}" + by auto + ultimately + show ?thesis by auto +qed + +end \ No newline at end of file diff --git a/thys/Independence_CH/Ordinals_In_MG.thy b/thys/Independence_CH/Ordinals_In_MG.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Ordinals_In_MG.thy @@ -0,0 +1,57 @@ +section\Ordinals in generic extensions\ +theory Ordinals_In_MG + imports + Forcing_Theorems +begin + +context G_generic1 +begin + +lemma rank_val: "rank(val(P,G,x)) \ rank(x)" (is "?Q(x)") +proof (induct rule:ed_induction[of ?Q]) + case (1 x) + have "val(P,G,x) = {val(P,G,u). u\{t\domain(x). \p\P . \t,p\\x \ p \ G }}" + using def_val[of G x] by auto + then + have "rank(val(P,G,x)) = (\u\{t\domain(x). \p\P . \t,p\\x \ p \ G }. succ(rank(val(P,G,u))))" + using rank[of "val(P,G,x)"] by simp + moreover + have "succ(rank(val(P,G, y))) \ rank(x)" if "ed(y, x)" for y + using 1[OF that] rank_ed[OF that] by (auto intro:lt_trans1) + moreover from this + have "(\u\{t\domain(x). \p\P . \t,p\\x \ p \ G }. succ(rank(val(P,G,u)))) \ rank(x)" + by (rule_tac UN_least_le) (auto) + ultimately + show ?case + by simp +qed + +lemma Ord_MG_iff: + assumes "Ord(\)" + shows "\ \ M \ \ \ M[G]" +proof + show "\ \ M[G]" if "\ \ M" + using generic[THEN one_in_G, THEN M_subset_MG] that .. +next + assume "\ \ M[G]" + then + obtain x where "x\M" "val(P,G,x) = \" + using GenExtD by auto + then + have "rank(\) \ rank(x)" + using rank_val by blast + with assms + have "\ \ rank(x)" + using rank_of_Ord by simp + then + have "\ \ succ(rank(x))" + using ltD by simp + with \x\M\ + show "\ \ M" + using cons_closed transitivity[of \ "succ(rank(x))"] rank_closed + unfolding succ_def by simp +qed + +end \ \\<^locale>\G_generic1\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Pairing_Axiom.thy b/thys/Independence_CH/Pairing_Axiom.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Pairing_Axiom.thy @@ -0,0 +1,54 @@ +section\The Axiom of Pairing in $M[G]$\ + +theory Pairing_Axiom + imports + Names +begin + +context forcing_data1 +begin + +lemma val_Upair : + "\ \ G \ val(P,G,{\\,\\,\\,\\}) = {val(P,G,\),val(P,G,\)}" + by (insert one_in_P, rule trans, subst def_val,auto) + +lemma pairing_in_MG : + assumes "M_generic(G)" + shows "upair_ax(##M[G])" +proof - + from assms + have types: "\\G" "\\P" "\\M" + using one_in_G one_in_M one_in_P + by simp_all + { + fix x y + note assms types + moreover + assume "x \ M[G]" "y \ M[G]" + moreover from this + obtain \ \ where "val(P,G,\) = x" "val(P,G,\) = y" "\ \ M" "\ \ M" + using GenExtD by blast + moreover from types this + have "\\,\\ \ M" "\\,\\\M" + using pair_in_M_iff by auto + moreover from this + have "{\\,\\,\\,\\} \ M" (is "?\ \ _") + using upair_in_M_iff by simp + moreover from this + have "val(P,G,?\) \ M[G]" + using GenExtI by simp + moreover from calculation + have "{val(P,G,\),val(P,G,\)} \ M[G]" + using val_Upair assms one_in_G by simp + ultimately + have "{x,y} \ M[G]" + by simp + } + then + show ?thesis + unfolding upair_ax_def upair_def by auto +qed + +end \ \\<^locale>\forcing_data1\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Powerset_Axiom.thy b/thys/Independence_CH/Powerset_Axiom.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Powerset_Axiom.thy @@ -0,0 +1,298 @@ +section\The Powerset Axiom in $M[G]$\ +theory Powerset_Axiom + imports Separation_Axiom Pairing_Axiom Union_Axiom +begin + +simple_rename "perm_pow" src "[ss,p,l,o,fs,\]" tgt "[fs,ss,sp,p,l,o,\]" + +lemma Collect_inter_Transset: + assumes + "Transset(M)" "b \ M" + shows + "{x\b . P(x)} = {x\b . P(x)} \ M" + using assms unfolding Transset_def + by (auto) + +context G_generic1 +begin + +lemma name_components_in_M: + assumes "\\,p\\\" "\ \ M" + shows "\\M" "p\M" +proof - + from assms + obtain a where "\ \ a" "p \ a" "a\\\,p\" + unfolding Pair_def by auto + moreover from assms + have "\\,p\\M" + using transitivity by simp + moreover from calculation + have "a\M" + using transitivity by simp + ultimately + show "\\M" "p\M" + using transitivity by simp_all +qed + +lemma sats_fst_snd_in_M: + assumes + "A\M" "B\M" "\ \ formula" "p\M" "l\M" "o\M" "\\M" "arity(\) \ 6" + shows "{\s,q\\A\B . M, [q,p,l,o,s,\] \ \} \ M" (is "?\ \ M") +proof - + let ?\' = "ren(\)`6`7`perm_pow_fn" + from \A\M\ \B\M\ + have "A\B \ M" + using cartprod_closed by simp + from \arity(\) \ 6\ \\\ formula\ + have "?\' \ formula" "arity(?\')\7" + unfolding perm_pow_fn_def + using perm_pow_thm arity_ren ren_tc Nil_type + by auto + with \?\' \ formula\ + have arty: "arity(Exists(Exists(And(pair_fm(0,1,2),?\'))))\5" (is "arity(?\)\5") + using ord_simp_union pred_le + by (auto simp:arity) + { + fix sp + note \A\B \ M\ \A\M\ \B\M\ + moreover + assume "sp \ A\B" + moreover from calculation + have "fst(sp) \ A" "snd(sp) \ B" + using fst_type snd_type by simp_all + ultimately + have "sp \ M" "fst(sp) \ M" "snd(sp) \ M" + using transitivity + by simp_all + note inM = \A\M\ \B\M\ \p\M\ \l\M\ \o\M\ \\\M\ + \sp\M\ \fst(sp)\M\ \snd(sp)\M\ + with arty \sp \ M\ \?\' \ formula\ + have "(M, [sp,p,l,o,\]@[p] \ ?\) \ M,[sp,p,l,o,\] \ ?\" (is "(M,?env0@ _\_) \ _") + using arity_sats_iff[of ?\ "[p]" M ?env0] by auto + also from inM \sp \ A\B\ + have "... \ sats(M,?\',[fst(sp),snd(sp),sp,p,l,o,\])" + by auto + also from inM \\ \ formula\ \arity(\) \ 6\ + have "... \ M, [snd(sp),p,l,o,fst(sp),\] \ \" + (is "sats(_,_,?env1) \ sats(_,_,?env2)") + using sats_iff_sats_ren[of \ 6 7 ?env2 M ?env1 perm_pow_fn] perm_pow_thm + unfolding perm_pow_fn_def by simp + finally + have "(M,[sp,p,l,o,\,p] \ ?\) \ M, [snd(sp),p,l,o,fst(sp),\] \ \" + by simp + } + then + have "?\ = {sp\A\B . sats(M,?\,[sp,p,l,o,\,p])}" + by auto + also from assms \A\B\M\ + have " ... \ M" + proof - + from arty + have "arity(?\) \ 6" + using leI by simp + moreover from \?\' \ formula\ + have "?\ \ formula" + by simp + moreover + note assms \A\B\M\ + ultimately + show "{x \ A\B . M, [x, p, l, o, \, p] \ ?\} \ M" + using separation_ax separation_iff + by simp + qed + finally show ?thesis . +qed + +lemma Pow_inter_MG: + assumes "a\M[G]" + shows "Pow(a) \ M[G] \ M[G]" +proof - + from assms + obtain \ where "\ \ M" "val(P,G, \) = a" + using GenExtD by auto + let ?Q="Pow(domain(\)\P) \ M" + from \\\M\ + have "domain(\)\P \ M" "domain(\) \ M" + using domain_closed cartprod_closed P_in_M + by simp_all + then + have "?Q \ M" + proof - + from power_ax \domain(\)\P \ M\ + obtain Q where "powerset(##M,domain(\)\P,Q)" "Q \ M" + unfolding power_ax_def by auto + moreover from calculation + have "z\Q \ z\M" for z + using transitivity by blast + ultimately + have "Q = {a\Pow(domain(\)\P) . a\M}" + using \domain(\)\P \ M\ powerset_abs[of "domain(\)\P" Q] + by (simp flip: setclass_iff) + also + have " ... = ?Q" + by auto + finally + show ?thesis + using \Q\M\ by simp + qed + let ?\="?Q\{\}" + let ?b="val(P,G,?\)" + from \?Q\M\ + have "?\\M" + using one_in_P P_in_M transitivity + by (simp flip: setclass_iff) + then + have "?b \ M[G]" + using GenExtI by simp + have "Pow(a) \ M[G] \ ?b" + proof + fix c + assume "c \ Pow(a) \ M[G]" + then + obtain \ where "c\M[G]" "\ \ M" "val(P,G,\) = c" + using GenExt_iff by auto + let ?\="{\\,p\ \domain(\)\P . p \ \0 \ 1\ [\,\] }" + have "arity(forces(Member(0,1))) = 6" + using arity_forces_at by auto + with \domain(\) \ M\ \\ \ M\ + have "?\ \ M" + using P_in_M one_in_M leq_in_M sats_fst_snd_in_M + by simp + then + have "?\ \ ?Q" by auto + then + have "val(P,G,?\) \ ?b" + using one_in_G one_in_P generic val_of_elem [of ?\ \ ?\ G] + by auto + have "val(P,G,?\) = c" + proof(intro equalityI subsetI) + fix x + assume "x \ val(P,G,?\)" + then + obtain \ p where 1: "\\,p\\?\" "p\G" "val(P,G,\) = x" + using elem_of_val_pair + by blast + moreover from \\\,p\\?\\ \?\ \ M\ + have "\\M" + using name_components_in_M[of _ _ ?\] by auto + moreover from 1 + have "p \ \0 \ 1\ [\,\]" "p\P" + by simp_all + moreover + note \val(P,G,\) = c\ \\ \ M\ + ultimately + have "M[G], [x, c] \ \0 \ 1\" + using generic definition_of_forcing[where \="\0 \ 1\"] ord_simp_union + by auto + moreover from \\\M\ \\\M\ + have "x\M[G]" + using \val(P,G,\) = x\ GenExtI by blast + ultimately + show "x\c" + using \c\M[G]\ by simp + next + fix x + assume "x \ c" + with \c \ Pow(a) \ M[G]\ + have "x \ a" "c\M[G]" "x\M[G]" + using transitivity_MG by auto + with \val(P,G, \) = a\ + obtain \ where "\\domain(\)" "val(P,G,\) = x" + using elem_of_val by blast + moreover + note \x\c\ \val(P,G,\) = c\ \c\M[G]\ \x\M[G]\ + moreover from calculation + have "val(P,G,\) \ val(P,G,\)" + by simp + moreover from calculation + have "M[G], [x, c] \ \0 \ 1\" + by simp + moreover + have "\\M" + proof - + from \\\domain(\)\ + obtain p where "\\,p\ \ \" + by auto + with \\\M\ + show ?thesis + using name_components_in_M by blast + qed + moreover + note \\ \ M\ + ultimately + obtain p where "p\G" "p \ \0 \ 1\ [\,\]" + using generic truth_lemma[of "\0 \ 1\" "G" "[\,\]" ] ord_simp_union + by auto + moreover from \p\G\ + have "p\P" + using generic by blast + ultimately + have "\\,p\\?\" + using \\\domain(\)\ by simp + with \val(P,G,\) = x\ \p\G\ + show "x\val(P,G,?\)" + using val_of_elem [of _ _ "?\"] by auto + qed + with \val(P,G,?\) \ ?b\ + show "c\?b" + by simp + qed + then + have "Pow(a) \ M[G] = {x\?b . x\a \ x\M[G]}" + by auto + also from \a\M[G]\ + have " ... = {x\?b . ( M[G], [x,a] \ \0 \ 1\ ) \ x\M[G]}" + using Transset_MG by force + also + have " ... = {x\?b . ( M[G], [x,a] \ \0 \ 1\ )} \ M[G]" + by auto + also from \?b\M[G]\ + have " ... = {x\?b . ( M[G], [x,a] \ \0 \ 1\ )}" + using Collect_inter_Transset Transset_MG + by simp + also from \?b\M[G]\ \a\M[G]\ + have " ... \ M[G]" + using Collect_sats_in_MG GenExtI ord_simp_union by (simp add:arity) + finally + show ?thesis . +qed + +end \ \\<^locale>\G_generic1\\ + +sublocale G_generic1 \ ext: M_trivial "##M[G]" + using generic Union_MG pairing_in_MG zero_in_MG transitivity_MG + unfolding M_trivial_def M_trans_def M_trivial_axioms_def + by (simp; blast) + +context G_generic1 begin + +theorem power_in_MG : "power_ax(##(M[G]))" + unfolding power_ax_def +proof (intro rallI, simp only:setclass_iff rex_setclass_is_bex) + (* After simplification, we have to show that for every + a\M[G] there exists some x\M[G] with powerset(##M[G],a,x) + *) + fix a + assume "a \ M[G]" + then + have "(##M[G])(a)" + by simp + have "{x\Pow(a) . x \ M[G]} = Pow(a) \ M[G]" + by auto + also from \a\M[G]\ + have " ... \ M[G]" + using Pow_inter_MG by simp + finally + have "{x\Pow(a) . x \ M[G]} \ M[G]" . + moreover from \a\M[G]\ \{x\Pow(a) . x \ M[G]} \ _\ + have "powerset(##M[G], a, {x\Pow(a) . x \ M[G]})" + using ext.powerset_abs[OF \(##M[G])(a)\] + by simp + ultimately + show "\x\M[G] . powerset(##M[G], a, x)" + by auto +qed + +end \ \\<^locale>\G_generic1\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Proper_Extension.thy b/thys/Independence_CH/Proper_Extension.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Proper_Extension.thy @@ -0,0 +1,79 @@ +section\Separative notions and proper extensions\ +theory Proper_Extension + imports + Names + +begin + +text\The key ingredient to obtain a proper extension is to have +a \<^emph>\separative preorder\:\ + +locale separative_notion = forcing_notion + + assumes separative: "p\P \ \q\P. \r\P. q \ p \ r \ p \ q \ r" +begin + +text\For separative preorders, the complement of every filter is +dense. Hence an $M$-generic filter cannot belong to the ground model.\ + +lemma filter_complement_dense: + assumes "filter(G)" + shows "dense(P - G)" +proof + fix p + assume "p\P" + show "\d\P - G. d \ p" + proof (cases "p\G") + case True + note \p\P\ assms + moreover + obtain q r where "q \ p" "r \ p" "q \ r" "q\P" "r\P" + using separative[OF \p\P\] + by force + with \filter(G)\ + obtain s where "s \ p" "s \ G" "s \ P" + using filter_imp_compat[of G q r] + by auto + then + show ?thesis + by blast + next + case False + with \p\P\ + show ?thesis + using refl_leq unfolding Diff_def by auto + qed +qed + +end \ \\<^locale>\separative_notion\\ + +locale ctm_separative = forcing_data1 + separative_notion +begin + +lemma generic_not_in_M: + assumes "M_generic(G)" + shows "G \ M" +proof + assume "G\M" + then + have "P - G \ M" + using P_in_M Diff_closed by simp + moreover + have "\(\q\G. q \ P - G)" "(P - G) \ P" + unfolding Diff_def by auto + moreover + note assms + ultimately + show "False" + using filter_complement_dense[of G] M_generic_denseD[of G "P-G"] + M_generic_def by simp \ \need to put generic ==> filter in claset\ +qed + +theorem proper_extension: + assumes "M_generic(G)" + shows "M \ M[G]" + using assms G_in_Gen_Ext[of G] one_in_G[of G] generic_not_in_M + by force + +end \ \\<^locale>\ctm_separative\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/ROOT b/thys/Independence_CH/ROOT new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/ROOT @@ -0,0 +1,21 @@ +chapter AFP + +session "Independence_CH" (AFP) = "Transitive_Models" + + description " + The Independence of the Continuum Hypothesis in Isabelle/ZF + + We redeveloped our formalization of forcing in the set theory framework of + Isabelle/ZF. Under the assumption of the existence of a countable + transitive model of ZFC, we construct proper generic extensions + that satisfy the Continuum Hypothesis and its negation. + " + options [timeout=300] + sessions + "Transitive_Models" + theories + "Definitions_Main" + "Demonstrations" + document_files + "root.tex" + "root.bib" + "root.bst" diff --git a/thys/Independence_CH/Replacement_Axiom.thy b/thys/Independence_CH/Replacement_Axiom.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Replacement_Axiom.thy @@ -0,0 +1,393 @@ +section\The Axiom of Replacement in $M[G]$\ +theory Replacement_Axiom + imports + Separation_Axiom +begin + +context forcing_data1 +begin + +bundle sharp_simps1 = snd_abs[simp] fst_abs[simp] fst_closed[simp del, simplified, simp] + snd_closed[simp del, simplified, simp] M_inhabited[simplified, simp] + pair_in_M_iff[simp del, simplified, simp] + +lemma sats_forces_iff_sats_rename_split_fm: + includes sharp_simps1 + assumes + "[\,m,p,P,leq,\,t,\] @ nenv \list(M)" "V \ M" + "\\formula" + shows + "(M, [p, P, leq, \, t, \] @ nenv \ forces(\)) \ + M, [V, \, \, \t,p\, m, P, leq, \] @ nenv \ rename_split_fm(\)" + using assms unfolding rename_split_fm_def + by (simp add:sats_incr_bv_iff[where bvs="[_,_,_,_,_,_]", simplified]) + +lemma sats_body_ground_repl_fm: + includes sharp_simps1 + assumes + "\t p. x=\t,p\" "[x,\,m,P,leq,\] @ nenv \list(M)" + "\\formula" + shows + "(\\\M. \V\M. is_Vset(\a. (##M)(a),\,V) \ \ \ V \ (snd(x) \ \ ([fst(x),\]@nenv))) + \ M, [\, x, m, P, leq, \] @ nenv \ body_ground_repl_fm(\)" +proof - + { + fix \ V t p + assume "\ \ M" "V \ M" "x = \t, p\" "t \ M" "p \ M" + with assms + have "\ \ V \ (M, [p,P,leq,\,t,\] @ nenv \ forces(\)) \ + \ \ V \ (M, [V,\,\,\t, p\,m,P, leq, \] @ nenv \ rename_split_fm(\))" + using sats_forces_iff_sats_rename_split_fm[of \ m p t \, where nenv=nenv and \=\] + by auto + } + note eq = this + show ?thesis + unfolding body_ground_repl_fm_def + apply (insert assms) + apply (rule iff_sats | simp add:nonempty[simplified])+ + using eq + by (auto del: iffI) +qed + +end \ \\<^locale>\forcing_data1\\ + +context G_generic1 +begin + +lemma Replace_sats_in_MG: + assumes + "c\M[G]" "env \ list(M[G])" + "\ \ formula" "arity(\) \ 2 +\<^sub>\ length(env)" + "univalent(##M[G], c, \x v. (M[G] , [x,v]@env \ \) )" + and + ground_replacement: + "\nenv. ground_replacement_assm(M,[P,leq,\] @ nenv, \)" + shows + "{v. x\c, v\M[G] \ (M[G] , [x,v]@env \ \)} \ M[G]" +proof - + let ?R = "\ x v . v\M[G] \ (M[G] , [x,v]@env \ \)" + from \c\M[G]\ + obtain \' where "val(P,G, \') = c" "\' \ M" + using GenExt_def by auto + then + have "domain(\')\P\M" (is "?\\M") + using cartprod_closed P_in_M domain_closed by simp + from \val(P,G, \') = c\ + have "c \ val(P,G,?\)" + using def_val[of G ?\] one_in_P one_in_G[OF generic] elem_of_val + domain_of_prod[OF one_in_P, of "domain(\')"] by force + from \env \ _\ + obtain nenv where "nenv\list(M)" "env = map(val(P,G),nenv)" + using map_val by auto + then + have "length(nenv) = length(env)" by simp + define f where "f(\p) \ \ \. \\M \ (\\\M. \ \ Vset(\) \ + (snd(\p) \ \ ([fst(\p),\] @ nenv)))" (is "_ \ \ \. ?P(\p,\)") for \p + have "f(\p) = (\ \. \\M \ (\\\M. \V\M. is_Vset(##M,\,V) \ \\V \ + (snd(\p) \ \ ([fst(\p),\] @ nenv))))" (is "_ = (\ \. \\M \ ?Q(\p,\))") for \p + unfolding f_def using Vset_abs Vset_closed Ord_Least_cong[of "?P(\p)" "\ \. \\M \ ?Q(\p,\)"] + by (simp, simp del:setclass_iff) + moreover + have "f(\p) \ M" for \p + unfolding f_def using Least_closed'[of "?P(\p)"] by simp + ultimately + have 1:"least(##M,\\. ?Q(\p,\),f(\p))" for \p + using least_abs'[of "\\. \\M \ ?Q(\p,\)" "f(\p)"] least_conj + by (simp flip: setclass_iff) + have "Ord(f(\p))" for \p unfolding f_def by simp + define QQ where "QQ\?Q" + from 1 + have "least(##M,\\. QQ(\p,\),f(\p))" for \p + unfolding QQ_def . + from \arity(\) \ _\ \length(nenv) = _\ + have "arity(\) \ 2 +\<^sub>\ length(nenv)" + by simp + moreover + note assms \nenv\list(M)\ \?\\M\ + moreover + have "\p\?\ \ \t p. \p=\t,p\" for \p + by auto + ultimately + have body:"(M , [\,\p,m,P,leq,\] @ nenv \ body_ground_repl_fm(\)) \ ?Q(\p,\)" + if "\p\?\" "\p\M" "m\M" "\\M" for \ \p m + using that P_in_M leq_in_M one_in_M sats_body_ground_repl_fm[of \p \ m nenv \] by simp + { + fix \p m + assume asm: "\p\M" "\p\?\" "m\M" + note inM = this P_in_M leq_in_M one_in_M \nenv\list(M)\ + with body + have body':"\\. \ \ M \ (\\\M. \V\M. is_Vset(\a. (##M)(a), \, V) \ \ \ V \ + (snd(\p) \ \ ([fst(\p),\] @ nenv))) \ + M, Cons(\, [\p, m, P, leq, \] @ nenv) \ body_ground_repl_fm(\)" by simp + from inM + have "(M , [\p,m,P,leq,\] @ nenv \ ground_repl_fm(\)) \ least(##M, QQ(\p), m)" + using sats_least_fm[OF body', of 1] unfolding QQ_def ground_repl_fm_def + by (simp, simp flip: setclass_iff) + } + then + have "(M, [\p,m,P,leq,\] @ nenv \ ground_repl_fm(\)) \ least(##M, QQ(\p), m)" + if "\p\M" "\p\?\" "m\M" for \p m using that by simp + then + have "univalent(##M, ?\, \\p m. M , [\p,m] @ ([P,leq,\] @ nenv) \ ground_repl_fm(\))" + unfolding univalent_def by (auto intro:unique_least) + moreover from \length(_) = _\ \env \ _\ + have "length([P,leq,\] @ nenv) = 3 +\<^sub>\ length(env)" by simp + moreover from \arity(\) \ 2 +\<^sub>\ length(nenv)\ + \length(_) = length(_)\[symmetric] \nenv\_\ \\\_\ + have "arity(ground_repl_fm(\)) \ 5 +\<^sub>\ length(env)" + using arity_ground_repl_fm[of \] le_trans Un_le by auto + moreover from \\\formula\ + have "ground_repl_fm(\)\formula" by simp + moreover + note inM = P_in_M leq_in_M one_in_M \nenv\list(M)\ \?\\M\ + moreover + note \length(nenv) = length(env)\ + ultimately + obtain Y where "Y\M" + "\m\M. m \ Y \ (\\p\M. \p \ ?\ \ (M, [\p,m] @ ([P,leq,\] @ nenv) \ ground_repl_fm(\)))" + using ground_replacement[of nenv] + unfolding strong_replacement_def ground_replacement_assm_def replacement_assm_def by auto + with \least(_,QQ(_),f(_))\ \f(_) \ M\ \?\\M\ + \_ \ _ \ _ \ (M,_ \ ground_repl_fm(\)) \ least(_,_,_)\ + have "f(\p)\Y" if "\p\?\" for \p + using that transitivity[OF _ \?\\M\] + by (clarsimp, rule_tac x="\x,y\" in bexI, auto) + moreover + have "{y\Y. Ord(y)} \ M" + using \Y\M\ separation_ax sats_ordinal_fm trans_M + separation_cong[of "##M" "\y. sats(M,ordinal_fm(0),[y])" "Ord"] + separation_closed by (simp add:arity) + then + have "\ {y\Y. Ord(y)} \ M" (is "?sup \ M") + using Union_closed by simp + then + have "{x\Vset(?sup). x \ M} \ M" + using Vset_closed by simp + moreover + have "{\} \ M" + using one_in_M singleton_closed by simp + ultimately + have "{x\Vset(?sup). x \ M} \ {\} \ M" (is "?big_name \ M") + using cartprod_closed by simp + then + have "val(P,G,?big_name) \ M[G]" + by (blast intro:GenExtI) + { + fix v x + assume "x\c" + moreover + note \val(P,G,\')=c\ \\'\M\ + moreover + from calculation + obtain \ p where "\\,p\\\'" "val(P,G,\) = x" "p\G" "\\M" + using elem_of_val_pair'[of \' x G] by blast + moreover + assume "v\M[G]" + then + obtain \ where "val(P,G,\) = v" "\\M" + using GenExtD by auto + moreover + assume "sats(M[G], \, [x,v] @ env)" + moreover + note \\\_\ \nenv\_\ \env = _\ \arity(\)\ 2 +\<^sub>\ length(env)\ + ultimately + obtain q where "q\G" "q \ \ ([\,\]@nenv)" + using truth_lemma[OF \\\_\ generic, symmetric, of "[\,\] @ nenv"] + by auto + with \\\,p\\\'\ \\\,q\\?\ \ f(\\,q\)\Y\ + have "f(\\,q\)\Y" + using generic unfolding M_generic_def filter_def by blast + let ?\="succ(rank(\))" + note \\\M\ + moreover from this + have "?\ \ M" + using rank_closed cons_closed by (simp flip: setclass_iff) + moreover + have "\ \ Vset(?\)" + using Vset_Ord_rank_iff by auto + moreover + note \q \ \ ([\,\] @ nenv)\ + ultimately + have "?P(\\,q\,?\)" by (auto simp del: Vset_rank_iff) + moreover + have "(\ \. ?P(\\,q\,\)) = f(\\,q\)" + unfolding f_def by simp + ultimately + obtain \ where "\\M" "\ \ Vset(f(\\,q\))" "q \ \ ([\,\] @ nenv)" + using LeastI[of "\ \. ?P(\\,q\,\)" ?\] by auto + with \q\G\ \\\M\ \nenv\_\ \arity(\)\ 2 +\<^sub>\ length(nenv)\ + have "M[G], map(val(P,G),[\,\] @ nenv) \ \" + using truth_lemma[OF \\\_\ generic, of "[\,\] @ nenv"] by auto + moreover from \x\c\ \c\M[G]\ + have "x\M[G]" using transitivity_MG by simp + moreover + note \M[G],[x,v] @ env\ \\ \env = map(val(P,G),nenv)\ \\\M\ \val(P,G,\)=x\ + \univalent(##M[G],_,_)\ \x\c\ \v\M[G]\ + ultimately + have "v=val(P,G,\)" + using GenExtI[of \ G] unfolding univalent_def by (auto) + from \\ \ Vset(f(\\,q\))\ \Ord(f(_))\ \f(\\,q\)\Y\ + have "\ \ Vset(?sup)" + using Vset_Ord_rank_iff lt_Union_iff[of _ "rank(\)"] by auto + with \\\M\ + have "val(P,G,\) \ val(P,G,?big_name)" + using domain_of_prod[of \ "{\}" "{x\Vset(?sup). x \ M}" ] def_val[of G ?big_name] + one_in_G[OF generic] one_in_P by (auto simp del: Vset_rank_iff) + with \v=val(P,G,\)\ + have "v \ val(P,G,{x\Vset(?sup). x \ M} \ {\})" + by simp + } + then + have "{v. x\c, ?R(x,v)} \ val(P,G,?big_name)" (is "?repl\?big") + by blast + with \?big_name\M\ + have "?repl = {v\?big. \x\c. sats(M[G], \, [x,v] @ env )}" (is "_ = ?rhs") + proof(intro equalityI subsetI) + fix v + assume "v\?repl" + with \?repl\?big\ + obtain x where "x\c" "M[G], [x, v] @ env \ \" "v\?big" + using subsetD by auto + with \univalent(##M[G],_,_)\ \c\M[G]\ + show "v \ ?rhs" + unfolding univalent_def + using transitivity_MG ReplaceI[of "\ x v. \x\c. M[G], [x, v] @ env \ \"] by blast + next + fix v + assume "v\?rhs" + then + obtain x where + "v\val(P,G, ?big_name)" "M[G], [x, v] @ env \ \" "x\c" + by blast + moreover from this \c\M[G]\ + have "v\M[G]" "x\M[G]" + using transitivity_MG GenExtI[OF \?big_name\_\,of G] by auto + moreover from calculation \univalent(##M[G],_,_)\ + have "?R(x,y) \ y = v" for y + unfolding univalent_def by auto + ultimately + show "v\?repl" + using ReplaceI[of ?R x v c] + by blast + qed + moreover + let ?\ = "Exists(And(Member(0,2+\<^sub>\length(env)),\))" + have "v\M[G] \ (\x\c. M[G], [x,v] @ env \ \) \ M[G], [v] @ env @ [c] \ ?\" + "arity(?\) \ 2 +\<^sub>\ length(env)" "?\\formula" + for v + proof - + fix v + assume "v\M[G]" + with \c\M[G]\ + have "nth(length(env)+\<^sub>\1,[v]@env@[c]) = c" + using \env\_\nth_concat[of v c "M[G]" env] + by auto + note inMG= \nth(length(env)+\<^sub>\1,[v]@env@[c]) = c\ \c\M[G]\ \v\M[G]\ \env\_\ + show "(\x\c. M[G], [x,v] @ env \ \) \ M[G], [v] @ env @ [c] \ ?\" + proof + assume "\x\c. M[G], [x, v] @ env \ \" + then obtain x where + "x\c" "M[G], [x, v] @ env \ \" "x\M[G]" + using transitivity_MG[OF _ \c\M[G]\] + by auto + with \\\_\ \arity(\)\2+\<^sub>\length(env)\ inMG + show "M[G], [v] @ env @ [c] \ Exists(And(Member(0, 2 +\<^sub>\ length(env)), \))" + using arity_sats_iff[of \ "[c]" _ "[x,v]@env"] + by auto + next + assume "M[G], [v] @ env @ [c] \ Exists(And(Member(0, 2 +\<^sub>\ length(env)), \))" + with inMG + obtain x where + "x\M[G]" "x\c" "M[G], [x,v]@env@[c] \ \" + by auto + with \\\_\ \arity(\)\2+\<^sub>\length(env)\ inMG + show "\x\c. M[G], [x, v] @ env\ \" + using arity_sats_iff[of \ "[c]" _ "[x,v]@env"] + by auto + qed + next + from \env\_\ \\\_\ + show "arity(?\)\2+\<^sub>\length(env)" + using pred_mono[OF _ \arity(\)\2+\<^sub>\length(env)\] lt_trans[OF _ le_refl] + by (auto simp add:ord_simp_union arity) + next + from \\\_\ + show "?\\formula" by simp + qed + moreover from this + have "{v\?big. \x\c. M[G], [x,v] @ env \ \} = {v\?big. M[G], [v] @ env @ [c] \ ?\}" + using transitivity_MG[OF _ GenExtI, OF _ \?big_name\M\] + by simp + moreover from calculation and \env\_\ \c\_\ \?big\M[G]\ + have "{v\?big. M[G] , [v] @ env @ [c] \ ?\} \ M[G]" + using Collect_sats_in_MG by auto + ultimately + show ?thesis by simp +qed + +theorem strong_replacement_in_MG: + assumes + "\\formula" and "arity(\) \ 2 +\<^sub>\ length(env)" "env \ list(M[G])" + and + ground_replacement: + "\nenv. ground_replacement_assm(M,[P,leq,\] @ nenv, \)" + shows + "strong_replacement(##M[G],\x v. sats(M[G],\,[x,v] @ env))" +proof - + let ?R="\x y . M[G], [x, y] @ env \ \" + { + fix A + let ?Y="{v . x \ A, v\M[G] \ ?R(x,v)}" + assume 1: "(##M[G])(A)" + "\x[##M[G]]. x \ A \ (\y[##M[G]]. \z[##M[G]]. ?R(x,y) \ ?R(x,z) \ y = z)" + then + have "univalent(##M[G], A, ?R)" "A\M[G]" + unfolding univalent_def by simp_all + with assms \A\_\ + have "(##M[G])(?Y)" + using Replace_sats_in_MG ground_replacement + unfolding ground_replacement_assm_def by (auto) + have "b \ ?Y \ (\x[##M[G]]. x \ A \ ?R(x,b))" if "(##M[G])(b)" for b + proof(rule) + from \A\_\ + show "\x[##M[G]]. x \ A \ ?R(x,b)" if "b \ ?Y" + using that transitivity_MG by auto + next + show "b \ ?Y" if "\x[##M[G]]. x \ A \ ?R(x,b)" + proof - + from \(##M[G])(b)\ + have "b\M[G]" by simp + with that + obtain x where "(##M[G])(x)" "x\A" "b\M[G] \ ?R(x,b)" + by blast + moreover from this 1 \(##M[G])(b)\ + have "x\M[G]" "z\M[G] \ ?R(x,z) \ b = z" for z + by auto + ultimately + show ?thesis + using ReplaceI[of "\ x y. y\M[G] \ ?R(x,y)"] by blast + qed + qed + then + have "\b[##M[G]]. b \ ?Y \ (\x[##M[G]]. x \ A \ ?R(x,b))" + by simp + with \(##M[G])(?Y)\ + have " (\Y[##M[G]]. \b[##M[G]]. b \ Y \ (\x[##M[G]]. x \ A \ ?R(x,b)))" + by auto + } + then show ?thesis unfolding strong_replacement_def univalent_def + by auto +qed + +lemma replacement_assm_MG: + assumes + ground_replacement: + "\nenv. ground_replacement_assm(M,[P,leq,\] @ nenv, \)" + shows + "replacement_assm(M[G],env,\)" + using assms strong_replacement_in_MG + unfolding replacement_assm_def by simp + +end \ \\<^locale>\G_generic1\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Replacement_Instances.thy b/thys/Independence_CH/Replacement_Instances.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Replacement_Instances.thy @@ -0,0 +1,1432 @@ +section\More Instances of Replacement\ + +theory Replacement_Instances + imports + Separation_Instances + Transitive_Models.Pointed_DC_Relative +begin + +synthesize "setdiff" from_definition "setdiff" assuming "nonempty" +arity_theorem for "setdiff_fm" + +relationalize "first_rel" "is_first" external +synthesize "first_fm" from_definition "is_first" assuming "nonempty" + +relationalize "minimum_rel" "is_minimum" external +definition is_minimum' where + "is_minimum'(M,R,X,u) \ (M(u) \ u \ X \ (\v[M]. \a[M]. (v \ X \ v \ u \ a \ R) \ pair(M, u, v, a))) \ + (\x[M]. + (M(x) \ x \ X \ (\v[M]. \a[M]. (v \ X \ v \ x \ a \ R) \ pair(M, x, v, a))) \ + (\y[M]. M(y) \ y \ X \ (\v[M]. \a[M]. (v \ X \ v \ y \ a \ R) \ pair(M, y, v, a)) \ y = x)) \ + \ (\x[M]. (M(x) \ x \ X \ (\v[M]. \a[M]. (v \ X \ v \ x \ a \ R) \ pair(M, x, v, a))) \ + (\y[M]. M(y) \ y \ X \ (\v[M]. \a[M]. (v \ X \ v \ y \ a \ R) \ pair(M, y, v, a)) \ y = x)) \ + empty(M, u)" + +synthesize "minimum" from_definition "is_minimum'" assuming "nonempty" +arity_theorem for "minimum_fm" + +lemma composition_fm_type[TC]: "a0 \ \ \ a1 \ \ \ a2 \ \ \ + composition_fm(a0,a1,a2) \ formula" + unfolding composition_fm_def by simp + +arity_theorem for "composition_fm" + +definition is_omega_funspace :: "[i\o,i,i,i]\o" where + "is_omega_funspace(N,B,n,z) \ \o[N]. omega(N,o) \ n\o \ is_funspace(N, n, B, z)" + +synthesize "omega_funspace" from_definition "is_omega_funspace" assuming "nonempty" +arity_theorem for "omega_funspace_fm" + +definition HAleph_wfrec_repl_body where + "HAleph_wfrec_repl_body(N,mesa,x,z) \ \y[N]. + pair(N, x, y, z) \ + (\f[N]. + (\z[N]. + z \ f \ + (\xa[N]. + \y[N]. + \xaa[N]. + \sx[N]. + \r_sx[N]. + \f_r_sx[N]. + pair(N, xa, y, z) \ + pair(N, xa, x, xaa) \ + upair(N, xa, xa, sx) \ + pre_image(N, mesa, sx, r_sx) \ restriction(N, f, r_sx, f_r_sx) \ xaa \ mesa \ is_HAleph(N, xa, f_r_sx, y))) \ + is_HAleph(N, x, f, y))" + +(* MOVE THIS to an appropriate place *) +arity_theorem for "ordinal_fm" +arity_theorem for "is_Limit_fm" +arity_theorem for "empty_fm" +arity_theorem for "fun_apply_fm" + +synthesize "HAleph_wfrec_repl_body" from_definition assuming "nonempty" +arity_theorem for "HAleph_wfrec_repl_body_fm" + +definition dcwit_repl_body where + "dcwit_repl_body(N,mesa,A,a,s,R) \ \x z. \y[N]. pair(N, x, y, z) \ + is_wfrec + (N, \n f. is_nat_case + (N, a, + \m bmfm. + \fm[N]. + \cp[N]. + is_apply(N, f, m, fm) \ + is_Collect(N, A, \x. \fmx[N]. (N(x) \ fmx \ R) \ pair(N, fm, x, fmx), cp) \ + is_apply(N, s, cp, bmfm), + n), + mesa, x, y)" + +manual_schematic for "dcwit_repl_body" assuming "nonempty" + unfolding dcwit_repl_body_def + by (rule iff_sats is_nat_case_iff_sats is_eclose_iff_sats sep_rules | simp)+ + +synthesize "dcwit_repl_body" from_schematic + +definition dcwit_aux_fm where + "dcwit_aux_fm(A,s,R) \ (\\\\4`2 is 0\ \ + (\\\Collect_fm + (succ(succ(succ(succ(succ(succ(succ(succ(succ(succ(A)))))))))), + (\\\\0 \ + succ(succ(succ(succ(succ(succ(succ(succ(succ(succ(succ(succ(R)))))))))))) \ \ + pair_fm(3, 1, 0) \\), + 0) \ + \ succ(succ(succ(succ(succ(succ(succ(succ(succ(succ(s))))))))))`0 is 2\\\)\\)" + +arity_theorem for "dcwit_aux_fm" + +lemma dcwit_aux_fm_type[TC]: "A \ \ \ s \ \ \ R \ \ \ dcwit_aux_fm(A,s,R) \ formula" + by (simp_all add: dcwit_aux_fm_def) + +definition is_nat_case_dcwit_aux_fm where + "is_nat_case_dcwit_aux_fm(A,a,s,R) \ is_nat_case_fm + (succ(succ(succ(succ(succ(succ(a)))))),dcwit_aux_fm(A,s,R), + 2, 0)" + +lemma is_nat_case_dcwit_aux_fm_type[TC]: "A \ \ \ a \ \ \ s \ \ \ R \ \ \ is_nat_case_dcwit_aux_fm(A,a,s,R) \ formula" + by (simp_all add: is_nat_case_dcwit_aux_fm_def) + +manual_arity for "is_nat_case_dcwit_aux_fm" + unfolding is_nat_case_dcwit_aux_fm_def + by (rule arity_dcwit_aux_fm[THEN [6] arity_is_nat_case_fm]) simp_all + +manual_arity for "dcwit_repl_body_fm" + using arity_is_nat_case_dcwit_aux_fm[THEN [6] arity_is_wfrec_fm] + unfolding dcwit_repl_body_fm_def is_nat_case_dcwit_aux_fm_def dcwit_aux_fm_def + by (auto simp add: arity(1-33)) + +lemma arity_dcwit_repl_body: "arity(dcwit_repl_body_fm(6,5,4,3,2,0,1)) = 7" + by (simp_all add: FOL_arities arity_dcwit_repl_body_fm ord_simp_union) + +definition fst2_snd2 + where "fst2_snd2(x) \ \fst(fst(x)), snd(snd(x))\" + +relativize functional "fst2_snd2" "fst2_snd2_rel" +relationalize "fst2_snd2_rel" "is_fst2_snd2" + +lemma (in M_trivial) fst2_snd2_abs: + assumes "M(x)" "M(res)" + shows "is_fst2_snd2(M, x, res) \ res = fst2_snd2(x)" + unfolding is_fst2_snd2_def fst2_snd2_def + using fst_rel_abs[symmetric] snd_rel_abs[symmetric] fst_abs snd_abs assms + by simp + +synthesize "is_fst2_snd2" from_definition assuming "nonempty" +arity_theorem for "is_fst2_snd2_fm" + +definition sndfst_fst2_snd2 + where "sndfst_fst2_snd2(x) \ \snd(fst(x)), fst(fst(x)), snd(snd(x))\" + +relativize functional "sndfst_fst2_snd2" "sndfst_fst2_snd2_rel" +relationalize "sndfst_fst2_snd2_rel" "is_sndfst_fst2_snd2" +synthesize "is_sndfst_fst2_snd2" from_definition assuming "nonempty" +arity_theorem for "is_sndfst_fst2_snd2_fm" + +definition RepFun_body :: "i \ i \ i"where + "RepFun_body(u,v) \ {{\v, x\} . x \ u}" + +relativize functional "RepFun_body" "RepFun_body_rel" +relationalize "RepFun_body_rel" "is_RepFun_body" +synthesize "is_RepFun_body" from_definition assuming "nonempty" +arity_theorem for "is_RepFun_body_fm" + +lemma arity_body_repfun: + "arity((\\\cons_fm(0, 3, 2) \ pair_fm(5, 1, 0) \\)) = 5" + using arity_cons_fm arity_pair_fm pred_Un_distrib union_abs1 FOL_arities + by auto + +lemma arity_RepFun: "arity(is_RepFun_body_fm(0, 1, 2)) = 3" + unfolding is_RepFun_body_fm_def + using arity_Replace_fm[OF _ _ _ _ arity_body_repfun] arity_fst_fm arity_snd_fm arity_empty_fm + pred_Un_distrib union_abs2 union_abs1 FOL_arities + by simp + +definition order_eq_map where + "order_eq_map(M,A,r,a,z) \ \x[M]. \g[M]. \mx[M]. \par[M]. + ordinal(M,x) & pair(M,a,x,z) & membership(M,x,mx) & + pred_set(M,A,a,r,par) & order_isomorphism(M,par,r,x,mx,g)" + +synthesize "order_eq_map" from_definition assuming "nonempty" +arity_theorem for "is_ord_iso_fm" +arity_theorem for "order_eq_map_fm" + +(* Banach *) +synthesize "is_banach_functor" from_definition assuming "nonempty" +arity_theorem for "is_banach_functor_fm" + +definition banach_body_iterates where + "banach_body_iterates(M,X,Y,f,g,W,n,x,z) \ +\y[M]. + pair(M, x, y, z) \ + (\fa[M]. + (\z[M]. + z \ fa \ + (\xa[M]. + \y[M]. + \xaa[M]. + \sx[M]. + \r_sx[M]. + \f_r_sx[M]. \sn[M]. \msn[M]. successor(M,n,sn) \ + membership(M,sn,msn) \ + pair(M, xa, y, z) \ + pair(M, xa, x, xaa) \ + upair(M, xa, xa, sx) \ + pre_image(M, msn, sx, r_sx) \ + restriction(M, fa, r_sx, f_r_sx) \ + xaa \ msn \ + (empty(M, xa) \ y = W) \ + (\m[M]. + successor(M, m, xa) \ + (\gm[M]. + is_apply(M, f_r_sx, m, gm) \ is_banach_functor(M, X, Y, f, g, gm, y))) \ + (is_quasinat(M, xa) \ empty(M, y)))) \ + (empty(M, x) \ y = W) \ + (\m[M]. + successor(M, m, x) \ + (\gm[M]. is_apply(M, fa, m, gm) \ is_banach_functor(M, X, Y, f, g, gm, y))) \ + (is_quasinat(M, x) \ empty(M, y)))" + +synthesize "is_quasinat" from_definition assuming "nonempty" +arity_theorem for "is_quasinat_fm" + +synthesize "banach_body_iterates" from_definition assuming "nonempty" +arity_theorem for "banach_body_iterates_fm" + +definition banach_is_iterates_body where + "banach_is_iterates_body(M,X,Y,f,g,W,n,y) \ \om[M]. omega(M,om) \ n \ om \ + (\sn[M]. + \msn[M]. + successor(M, n, sn) \ + membership(M, sn, msn) \ + (\fa[M]. + (\z[M]. + z \ fa \ + (\x[M]. + \y[M]. + \xa[M]. + \sx[M]. + \r_sx[M]. + \f_r_sx[M]. + pair(M, x, y, z) \ + pair(M, x, n, xa) \ + upair(M, x, x, sx) \ + pre_image(M, msn, sx, r_sx) \ + restriction(M, fa, r_sx, f_r_sx) \ + xa \ msn \ + (empty(M, x) \ y = W) \ + (\m[M]. + successor(M, m, x) \ + (\gm[M]. + fun_apply(M, f_r_sx, m, gm) \ is_banach_functor(M, X, Y, f, g, gm, y))) \ + (is_quasinat(M, x) \ empty(M, y)))) \ + (empty(M, n) \ y = W) \ + (\m[M]. + successor(M, m, n) \ + (\gm[M]. fun_apply(M, fa, m, gm) \ is_banach_functor(M, X, Y, f, g, gm, y))) \ + (is_quasinat(M, n) \ empty(M, y))))" + +synthesize "banach_is_iterates_body" from_definition assuming "nonempty" +arity_theorem for "banach_is_iterates_body_fm" + +(* (##M)(f) \ strong_replacement(##M, \x y. y = \x, transrec(x, \a g. f ` (g `` a))\) *) + +definition trans_apply_image where + "trans_apply_image(f) \ \a g. f ` (g `` a)" + +relativize functional "trans_apply_image" "trans_apply_image_rel" +relationalize "trans_apply_image" "is_trans_apply_image" + +(* MOVE THIS to an appropriate place *) +schematic_goal arity_is_recfun_fm[arity]: + "p \ formula \ a \ \ \ z \ \ \ r \ \ \ arity(is_recfun_fm(p, a, z ,r)) = ?ar" + unfolding is_recfun_fm_def + by (simp add:arity) (* clean simpset from arities, use correct attrib *) + (* Don't know why it doesn't use the theorem at \<^file>\Arities\ *) +schematic_goal arity_is_wfrec_fm[arity]: + "p \ formula \ a \ \ \ z \ \ \ r \ \ \ arity(is_wfrec_fm(p, a, z ,r)) = ?ar" + unfolding is_wfrec_fm_def + by (simp add:arity) +schematic_goal arity_is_transrec_fm[arity]: + "p \ formula \ a \ \ \ z \ \ \ arity(is_transrec_fm(p, a, z)) = ?ar" + unfolding is_transrec_fm_def + by (simp add:arity) + +synthesize "is_trans_apply_image" from_definition assuming "nonempty" +arity_theorem for "is_trans_apply_image_fm" + + +definition transrec_apply_image_body where + "transrec_apply_image_body(M,f,mesa,x,z) \ \y[M]. pair(M, x, y, z) \ + (\fa[M]. + (\z[M]. + z \ fa \ + (\xa[M]. + \y[M]. + \xaa[M]. + \sx[M]. + \r_sx[M]. + \f_r_sx[M]. + pair(M, xa, y, z) \ + pair(M, xa, x, xaa) \ + upair(M, xa, xa, sx) \ + pre_image(M, mesa, sx, r_sx) \ + restriction(M, fa, r_sx, f_r_sx) \ + xaa \ mesa \ is_trans_apply_image(M, f, xa, f_r_sx, y))) \ + is_trans_apply_image(M, f, x, fa, y))" + +synthesize "transrec_apply_image_body" from_definition assuming "nonempty" +arity_theorem for "transrec_apply_image_body_fm" + +definition is_trans_apply_image_body where + "is_trans_apply_image_body(M,f,\,a,w) \ \z[M]. pair(M,a,z,w) \ a\\ \ (\sa[M]. + \esa[M]. + \mesa[M]. + upair(M, a, a, sa) \ + is_eclose(M, sa, esa) \ + membership(M, esa, mesa) \ + (\fa[M]. + (\z[M]. + z \ fa \ + (\x[M]. + \y[M]. + \xa[M]. + \sx[M]. + \r_sx[M]. + \f_r_sx[M]. + pair(M, x, y, z) \ + pair(M, x, a, xa) \ + upair(M, x, x, sx) \ + pre_image(M, mesa, sx, r_sx) \ + restriction(M, fa, r_sx, f_r_sx) \ + xa \ mesa \ is_trans_apply_image(M, f, x, f_r_sx, y))) \ + is_trans_apply_image(M, f, a, fa, z)))" + +manual_schematic "is_trans_apply_image_body_schematic" for "is_trans_apply_image_body"assuming "nonempty" + unfolding is_trans_apply_image_body_def + by (rule sep_rules is_eclose_iff_sats is_trans_apply_image_iff_sats | simp)+ + +synthesize "is_trans_apply_image_body" from_schematic "is_trans_apply_image_body_schematic" +arity_theorem for "is_trans_apply_image_body_fm" + +synthesize "is_converse" from_definition assuming "nonempty" +arity_theorem for "is_converse_fm" + +definition replacement_is_omega_funspace_fm where "replacement_is_omega_funspace_fm \ omega_funspace_fm(2,0,1)" +definition replacement_HAleph_wfrec_repl_body_fm where "replacement_HAleph_wfrec_repl_body_fm \ HAleph_wfrec_repl_body_fm(2,0,1)" +definition replacement_is_fst2_snd2_fm where "replacement_is_fst2_snd2_fm \ is_fst2_snd2_fm(0,1)" +definition replacement_is_sndfst_fst2_snd2_fm where "replacement_is_sndfst_fst2_snd2_fm \ is_sndfst_fst2_snd2_fm(0,1)" +definition replacement_is_order_eq_map_fm where "replacement_is_order_eq_map_fm \ order_eq_map_fm(2,3,0,1)" +definition replacement_transrec_apply_image_body_fm where "replacement_transrec_apply_image_body_fm \ transrec_apply_image_body_fm(3,2,0,1)" +definition banach_replacement_iterates_fm where "banach_replacement_iterates_fm \ banach_is_iterates_body_fm(6,5,4,3,2,0,1)" +definition replacement_is_trans_apply_image_fm where "replacement_is_trans_apply_image_fm \ is_trans_apply_image_body_fm(3,2,0,1)" +definition banach_iterates_fm where "banach_iterates_fm \ banach_body_iterates_fm(7,6,5,4,3,2,0,1)" +definition replacement_dcwit_repl_body_fm where "replacement_dcwit_repl_body_fm \ dcwit_repl_body_fm(6,5,4,3,2,0,1)" + +locale M_ZF2 = M_ZF1 + + assumes + replacement_ax2: + "replacement_assm(M,env,replacement_is_omega_funspace_fm)" + "replacement_assm(M,env,replacement_HAleph_wfrec_repl_body_fm)" + "replacement_assm(M,env,replacement_is_fst2_snd2_fm)" + "replacement_assm(M,env,replacement_is_sndfst_fst2_snd2_fm)" + "replacement_assm(M,env,replacement_is_order_eq_map_fm)" + "replacement_assm(M,env,replacement_transrec_apply_image_body_fm)" + "replacement_assm(M,env,banach_replacement_iterates_fm)" + "replacement_assm(M,env,replacement_is_trans_apply_image_fm)" + "replacement_assm(M,env,banach_iterates_fm)" + "replacement_assm(M,env,replacement_dcwit_repl_body_fm)" + and + Lambda_in_M_replacement2: + "replacement_assm(M,env,Lambda_in_M_fm(fst_fm(0,1),0))" + "replacement_assm(M,env,Lambda_in_M_fm(domain_fm(0,1),0))" + "replacement_assm(M,env,Lambda_in_M_fm(snd_fm(0,1),0))" + "replacement_assm(M,env,Lambda_in_M_fm(big_union_fm(0,1),0))" + "replacement_assm(M,env,Lambda_in_M_fm(is_cardinal_fm(0,1),0))" + "replacement_assm(M,env,Lambda_in_M_fm(is_converse_fm(0,1),0))" + and + LambdaPair_in_M_replacement2: + "replacement_assm(M,env,LambdaPair_in_M_fm(image_fm(0,1,2),0))" + "replacement_assm(M,env,LambdaPair_in_M_fm(setdiff_fm(0,1,2),0))" + "replacement_assm(M,env,LambdaPair_in_M_fm(minimum_fm(0,1,2),0))" + "replacement_assm(M,env,LambdaPair_in_M_fm(upair_fm(0,1,2),0))" + "replacement_assm(M,env,LambdaPair_in_M_fm(is_RepFun_body_fm(0,1,2),0))" + "replacement_assm(M,env,LambdaPair_in_M_fm(composition_fm(0,1,2),0))" + +definition instances2_fms where "instances2_fms \ + { replacement_is_omega_funspace_fm, + replacement_HAleph_wfrec_repl_body_fm, + replacement_is_fst2_snd2_fm, + replacement_is_sndfst_fst2_snd2_fm, + replacement_is_order_eq_map_fm, + replacement_transrec_apply_image_body_fm, + banach_replacement_iterates_fm, + replacement_is_trans_apply_image_fm, + banach_iterates_fm, + replacement_dcwit_repl_body_fm, + Lambda_in_M_fm(fst_fm(0,1),0), + Lambda_in_M_fm(domain_fm(0,1),0), + Lambda_in_M_fm(snd_fm(0,1),0), + Lambda_in_M_fm(big_union_fm(0,1),0), + Lambda_in_M_fm(is_cardinal_fm(0,1),0), + Lambda_in_M_fm(is_converse_fm(0,1),0), + LambdaPair_in_M_fm(image_fm(0,1,2),0), + LambdaPair_in_M_fm(setdiff_fm(0,1,2),0), + LambdaPair_in_M_fm(minimum_fm(0,1,2),0), + LambdaPair_in_M_fm(upair_fm(0,1,2),0), + LambdaPair_in_M_fm(is_RepFun_body_fm(0,1,2),0), + LambdaPair_in_M_fm(composition_fm(0,1,2),0) }" + +txt\This set has 22 internalized formulas.\ + +lemmas replacement_instances2_defs = + replacement_is_omega_funspace_fm_def + replacement_HAleph_wfrec_repl_body_fm_def + replacement_is_fst2_snd2_fm_def + replacement_is_sndfst_fst2_snd2_fm_def + replacement_is_order_eq_map_fm_def + replacement_transrec_apply_image_body_fm_def + banach_replacement_iterates_fm_def + replacement_is_trans_apply_image_fm_def + banach_iterates_fm_def + replacement_dcwit_repl_body_fm_def + +declare (in M_ZF2) replacement_instances2_defs [simp] + +lemma instances2_fms_type[TC]: "instances2_fms \ formula" + unfolding replacement_instances2_defs instances2_fms_def + by (simp del:Lambda_in_M_fm_def) + +locale M_ZF2_trans = M_ZF1_trans + M_ZF2 + +locale M_ZFC2 = M_ZFC1 + M_ZF2 + +locale M_ZFC2_trans = M_ZFC1_trans + M_ZF2_trans + M_ZFC2 + +lemma (in M_ZF2_trans) lam_replacement_domain : "lam_replacement(##M, domain)" + using lam_replacement_iff_lam_closed[THEN iffD2,of domain] + Lambda_in_M[where \="domain_fm(0,1)" and env="[]"] domain_type domain_abs + Lambda_in_M_replacement2(2) + arity_domain_fm[of 0 1] ord_simp_union transitivity domain_closed + by simp + +lemma (in M_ZF2_trans) lam_replacement_converse : "lam_replacement(##M, converse)" + using lam_replacement_iff_lam_closed[THEN iffD2,of converse] nonempty + Lambda_in_M[where \="is_converse_fm(0,1)" and env="[]"] + is_converse_fm_type converse_abs + arity_is_converse_fm[of 0 1] ord_simp_union transitivity converse_closed + Lambda_in_M_replacement2(6) + by simp + +lemma (in M_ZF2_trans) lam_replacement_fst : "lam_replacement(##M, fst)" + using lam_replacement_iff_lam_closed[THEN iffD2,of fst] + Lambda_in_M[where \="fst_fm(0,1)" and env="[]"] + fst_iff_sats[symmetric] fst_abs fst_type + arity_fst_fm[of 0 1] ord_simp_union transitivity fst_closed + Lambda_in_M_replacement2(1) + by simp + +lemma (in M_ZF2_trans) lam_replacement_snd : "lam_replacement(##M, snd)" + using lam_replacement_iff_lam_closed[THEN iffD2,of snd] + Lambda_in_M[where \="snd_fm(0,1)" and env="[]"] + snd_iff_sats[symmetric] snd_abs snd_type + arity_snd_fm[of 0 1] ord_simp_union transitivity snd_closed + Lambda_in_M_replacement2(3) + by simp + +lemma (in M_ZF2_trans) lam_replacement_Union : "lam_replacement(##M, Union)" + using lam_replacement_iff_lam_closed[THEN iffD2,of Union] + Lambda_in_M[where \="big_union_fm(0,1)" and env="[]"] Union_abs + union_fm_def big_union_iff_sats[symmetric] + arity_big_union_fm[of 0 1] ord_simp_union transitivity Union_closed + Lambda_in_M_replacement2(4) + by simp + +lemma (in M_ZF2_trans) lam_replacement_image: + "lam_replacement(##M, \p. fst(p) `` snd(p))" + using lam_replacement2_in_ctm[where \="image_fm(0,1,2)" and env="[]"] + image_type image_iff_sats image_abs + arity_image_fm[of 0 1 2] ord_simp_union transitivity image_closed fst_snd_closed + LambdaPair_in_M_replacement2(1) + by simp + +lemma (in M_ZF2_trans) lam_replacement_Diff: + "lam_replacement(##M, \p. fst(p) - snd(p))" + using lam_replacement2_in_ctm[where \="setdiff_fm(0,1,2)" and env="[]"] + setdiff_fm_type setdiff_iff_sats setdiff_abs + arity_setdiff_fm[of 0 1 2] ord_simp_union transitivity Diff_closed fst_snd_closed + nonempty LambdaPair_in_M_replacement2(2) + by simp + +lemma is_minimum_eq : + "M(R) \ M(X) \ M(u) \ is_minimum(M,R,X,u) \ is_minimum'(M,R,X,u)" + unfolding is_minimum_def is_minimum'_def is_The_def is_first_def by simp + +context M_trivial +begin + +lemma first_closed: + "M(B) \ M(r) \ first(u,r,B) \ M(u)" + using transM[OF first_is_elem] by simp + +is_iff_rel for "first" + unfolding is_first_def first_rel_def by auto + +is_iff_rel for "minimum" + unfolding is_minimum_def minimum_rel_def + using is_first_iff The_abs nonempty + by force + +end \ \\<^locale>\M_trivial\\ + +lemma (in M_ZF2_trans) lam_replacement_minimum: + "lam_replacement(##M, \p. minimum(fst(p), snd(p)))" + using lam_replacement2_in_ctm[where \="minimum_fm(0,1,2)" and env="[]"] + minimum_iff_sats[symmetric] is_minimum_iff minimum_abs is_minimum_eq + arity_minimum_fm[of 0 1 2] ord_simp_union minimum_fm_type + minimum_closed zero_in_M LambdaPair_in_M_replacement2(3) + by simp + +lemma (in M_ZF2_trans) lam_replacement_Upair: "lam_replacement(##M, \p. Upair(fst(p), snd(p)))" + using lam_replacement2_in_ctm[where \="upair_fm(0,1,2)" and env="[]" and f="Upair"] + Upair_closed upair_type upair_iff_sats Upair_eq_cons + arity_upair_fm[of 0 1 2,simplified] ord_simp_union LambdaPair_in_M_replacement2(4) + by simp + +lemma (in M_ZF2_trans) lam_replacement_comp: + "lam_replacement(##M, \p. comp(fst(p), snd(p)))" + using lam_replacement2_in_ctm[where \="composition_fm(0,1,2)" and env="[]" and f="comp"] + comp_closed composition_fm_type composition_iff_sats + arity_composition_fm[of 0 1 2] ord_simp_union LambdaPair_in_M_replacement2(6) + by simp + +lemma (in M_ZF2_trans) omega_funspace_abs: + "B\M \ n\M \ z\M \ is_omega_funspace(##M,B,n,z) \ n\\ \ is_funspace(##M,n,B,z)" + unfolding is_omega_funspace_def using nat_in_M by simp + +lemma (in M_ZF2_trans) replacement_is_omega_funspace: + "B\M \ strong_replacement(##M, is_omega_funspace(##M,B))" + using strong_replacement_rel_in_ctm[where \="omega_funspace_fm(2,0,1)" and env="[B]"] + zero_in_M arity_omega_funspace_fm ord_simp_union replacement_ax2(1) + by simp + +lemma (in M_ZF2_trans) replacement_omega_funspace: + "b\M\strong_replacement(##M, \n z. n\\ \ is_funspace(##M,n,b,z))" + using strong_replacement_cong[THEN iffD2,OF _ replacement_is_omega_funspace[of b]] + omega_funspace_abs[of b] setclass_iff[THEN iffD1] + by (simp del:setclass_iff) + +lemma (in M_ZF2_trans) replacement_HAleph_wfrec_repl_body: + "B\M \ strong_replacement(##M, HAleph_wfrec_repl_body(##M,B))" + using strong_replacement_rel_in_ctm[where \="HAleph_wfrec_repl_body_fm(2,0,1)" and env="[B]"] + zero_in_M arity_HAleph_wfrec_repl_body_fm replacement_ax2(2) ord_simp_union + by simp + +lemma (in M_ZF2_trans) HAleph_wfrec_repl: + "(##M)(sa) \ + (##M)(esa) \ + (##M)(mesa) \ + strong_replacement + (##M, + \x z. \y[##M]. + pair(##M, x, y, z) \ + (\f[##M]. + (\z[##M]. + z \ f \ + (\xa[##M]. + \y[##M]. + \xaa[##M]. + \sx[##M]. + \r_sx[##M]. + \f_r_sx[##M]. + pair(##M, xa, y, z) \ + pair(##M, xa, x, xaa) \ + upair(##M, xa, xa, sx) \ + pre_image(##M, mesa, sx, r_sx) \ restriction(##M, f, r_sx, f_r_sx) \ xaa \ mesa \ is_HAleph(##M, xa, f_r_sx, y))) \ + is_HAleph(##M, x, f, y)))" + using replacement_HAleph_wfrec_repl_body unfolding HAleph_wfrec_repl_body_def by simp + +lemma dcwit_replacement:"Ord(na) \ + N(na) \ + N(A) \ + N(a) \ + N(s) \ + N(R) \ + transrec_replacement + (N, \n f ntc. + is_nat_case + (N, a, + \m bmfm. + \fm[N]. \cp[N]. + is_apply(N, f, m, fm) \ + is_Collect(N, A, \x. \fmx[N]. (N(x) \ fmx \ R) \ pair(N, fm, x, fmx), cp) \ + is_apply(N, s, cp, bmfm), + n, ntc),na)" + unfolding transrec_replacement_def wfrec_replacement_def oops + +lemma (in M_ZF2_trans) replacement_dcwit_repl_body: + "(##M)(mesa) \ (##M)(A) \ (##M)(a) \ (##M)(s) \ (##M)(R) \ + strong_replacement(##M, dcwit_repl_body(##M,mesa,A,a,s,R))" + using strong_replacement_rel_in_ctm[where \="dcwit_repl_body_fm(6,5,4,3,2,0,1)" + and env="[R,s,a,A,mesa]" and f="dcwit_repl_body(##M,mesa,A,a,s,R)"] + zero_in_M arity_dcwit_repl_body replacement_ax2(10) + by simp + +lemma (in M_ZF2_trans) dcwit_repl: + "(##M)(sa) \ + (##M)(esa) \ + (##M)(mesa) \ (##M)(A) \ (##M)(a) \ (##M)(s) \ (##M)(R) \ + strong_replacement + ((##M), \x z. \y[(##M)]. pair((##M), x, y, z) \ + is_wfrec + ((##M), \n f. is_nat_case + ((##M), a, + \m bmfm. + \fm[(##M)]. + \cp[(##M)]. + is_apply((##M), f, m, fm) \ + is_Collect((##M), A, \x. \fmx[(##M)]. ((##M)(x) \ fmx \ R) \ pair((##M), fm, x, fmx), cp) \ + is_apply((##M), s, cp, bmfm), + n), + mesa, x, y))" + using replacement_dcwit_repl_body unfolding dcwit_repl_body_def by simp + +lemma (in M_ZF2_trans) replacement_fst2_snd2: "strong_replacement(##M, \x y. y = \fst(fst(x)), snd(snd(x))\)" + using strong_replacement_in_ctm[where \="is_fst2_snd2_fm(0,1)" and env="[]"] + zero_in_M fst_snd_closed pair_in_M_iff + arity_is_fst2_snd2_fm ord_simp_union fst2_snd2_abs replacement_ax2(3) + unfolding fst2_snd2_def + by simp + +lemma (in M_trivial) sndfst_fst2_snd2_abs: + assumes "M(x)" "M(res)" + shows "is_sndfst_fst2_snd2(M, x, res) \ res = sndfst_fst2_snd2(x)" + unfolding is_sndfst_fst2_snd2_def sndfst_fst2_snd2_def + using fst_rel_abs[symmetric] snd_rel_abs[symmetric] fst_abs snd_abs assms + by simp + +lemma (in M_ZF2_trans) replacement_sndfst_fst2_snd2: + "strong_replacement(##M, \x y. y = \snd(fst(x)), fst(fst(x)), snd(snd(x))\)" + using strong_replacement_in_ctm[where \="is_sndfst_fst2_snd2_fm(0,1)" and env="[]"] + zero_in_M fst_snd_closed pair_in_M_iff + arity_is_sndfst_fst2_snd2_fm ord_simp_union sndfst_fst2_snd2_abs replacement_ax2(4) + unfolding sndfst_fst2_snd2_def + by simp + +lemmas (in M_ZF2_trans) M_replacement_ZF_instances = lam_replacement_domain + lam_replacement_fst lam_replacement_snd lam_replacement_Union + lam_replacement_Upair lam_replacement_image + lam_replacement_Diff lam_replacement_converse + replacement_fst2_snd2 replacement_sndfst_fst2_snd2 + lam_replacement_comp + +lemmas (in M_ZF2_trans) M_separation_ZF_instances = separation_fstsnd_in_sndsnd + separation_sndfst_eq_fstsnd + +sublocale M_ZF2_trans \ M_replacement "##M" + using M_replacement_ZF_instances M_separation_ZF_instances + by unfold_locales simp + +lemma (in M_ZF1_trans) separation_is_dcwit_body: + assumes "(##M)(A)" "(##M)(a)" "(##M)(g)" "(##M)(R)" + shows "separation(##M,is_dcwit_body(##M, A, a, g, R))" + using assms separation_in_ctm[where env="[A,a,g,R]" and \="is_dcwit_body_fm(1,2,3,4,0)", + OF _ _ _ is_dcwit_body_iff_sats[symmetric], + of "\_.A" "\_.a" "\_.g" "\_.R" "\x. x"] + nonempty arity_is_dcwit_body_fm is_dcwit_body_fm_type + by (simp add:ord_simp_union) + +lemma (in M_trivial) RepFun_body_abs: + assumes "M(u)" "M(v)" "M(res)" + shows "is_RepFun_body(M, u, v, res) \ res = RepFun_body(u,v)" + unfolding is_RepFun_body_def RepFun_body_def + using fst_rel_abs[symmetric] snd_rel_abs[symmetric] fst_abs snd_abs assms + Replace_abs[where P="\xa a. a = {\v, xa\}" and A="u"] + univalent_triv transM[of _ u] + by auto + +lemma (in M_ZF2_trans) RepFun_SigFun_closed: "x \ M \ z \ M \ {{\z, x\} . x \ x} \ M" + using lam_replacement_sing_const_id lam_replacement_imp_strong_replacement RepFun_closed + transitivity singleton_in_M_iff pair_in_M_iff + by simp + +lemma (in M_ZF2_trans) replacement_RepFun_body: + "lam_replacement(##M, \p . {{\snd(p), x\} . x \ fst(p)})" + using lam_replacement2_in_ctm[where \="is_RepFun_body_fm(0,1,2)" and env="[]" and f="\p q . {{\q, x\} . x \ p}"] + RepFun_SigFun_closed[OF fst_snd_closed[THEN conjunct1,simplified] fst_snd_closed[THEN conjunct2,simplified]] + arity_RepFun ord_simp_union transitivity zero_in_M RepFun_body_def RepFun_body_abs RepFun_SigFun_closed + LambdaPair_in_M_replacement2(5) + by simp + +sublocale M_ZF2_trans \ M_replacement_extra "##M" + by unfold_locales (simp_all add: replacement_RepFun_body + lam_replacement_minimum del:setclass_iff) + +sublocale M_ZF2_trans \ M_Perm "##M" + using separation_PiP_rel separation_injP_rel separation_surjP_rel + lam_replacement_imp_strong_replacement[OF + lam_replacement_Sigfun[OF lam_replacement_constant]] + Pi_replacement1 unfolding Sigfun_def + by unfold_locales simp_all + +lemma (in M_ZF2_trans) replacement_is_order_eq_map: + "A\M \ r\M \ strong_replacement(##M, order_eq_map(##M,A,r))" + using strong_replacement_rel_in_ctm[where \="order_eq_map_fm(2,3,0,1)" and env="[A,r]" and f="order_eq_map(##M,A,r)"] + order_eq_map_iff_sats[where env="[_,_,A,r]"] zero_in_M fst_snd_closed pair_in_M_iff + arity_order_eq_map_fm ord_simp_union replacement_ax2(5) + by simp + +lemma (in M_ZF2_trans) banach_iterates: + assumes "X\M" "Y\M" "f\M" "g\M" "W\M" + shows "iterates_replacement(##M, is_banach_functor(##M,X,Y,f,g), W)" +proof - + have "strong_replacement(##M, \ x z . banach_body_iterates(##M,X,Y,f,g,W,n,x,z))" if "n\\" for n + using assms that arity_banach_body_iterates_fm ord_simp_union nat_into_M + strong_replacement_rel_in_ctm[where \="banach_body_iterates_fm(7,6,5,4,3,2,0,1)" + and env="[n,W,g,f,Y,X]"] replacement_ax2(9) + by simp + then + show ?thesis + using assms nat_into_M Memrel_closed + unfolding iterates_replacement_def wfrec_replacement_def is_wfrec_def M_is_recfun_def + is_nat_case_def iterates_MH_def banach_body_iterates_def + by simp +qed + +lemma (in M_ZF2_trans) banach_replacement_iterates: + assumes "X\M" "Y\M" "f\M" "g\M" "W\M" + shows "strong_replacement(##M, \n y. n\\ \ is_iterates(##M,is_banach_functor(##M,X, Y, f, g),W,n,y))" +proof - + have "strong_replacement(##M, \ n z . banach_is_iterates_body(##M,X,Y,f,g,W,n,z))" + using assms arity_banach_is_iterates_body_fm ord_simp_union nat_into_M + strong_replacement_rel_in_ctm[where \="banach_is_iterates_body_fm(6,5,4,3,2,0,1)" + and env="[W,g,f,Y,X]"] replacement_ax2(7) + by simp + then + show ?thesis + using assms nat_in_M + unfolding is_iterates_def wfrec_replacement_def is_wfrec_def M_is_recfun_def + is_nat_case_def iterates_MH_def banach_is_iterates_body_def + by simp +qed + +lemma (in M_ZF2_trans) banach_replacement: + assumes "(##M)(X)" "(##M)(Y)" "(##M)(f)" "(##M)(g)" + shows "strong_replacement(##M, \n y. n\nat \ y = banach_functor(X, Y, f, g)^n (0))" + using iterates_abs[OF banach_iterates banach_functor_abs,of X Y f g] + assms banach_functor_closed zero_in_M + strong_replacement_cong[THEN iffD1,OF _ banach_replacement_iterates[of X Y f g 0]] + by simp + +lemma (in M_ZF2_trans) lam_replacement_cardinal : "lam_replacement(##M, cardinal_rel(##M))" + using lam_replacement_iff_lam_closed[THEN iffD2,of "cardinal_rel(##M)"] + cardinal_rel_closed is_cardinal_iff + Lambda_in_M[where \="is_cardinal_fm(0,1)" and env="[]",OF is_cardinal_fm_type[of 0 1]] + arity_is_cardinal_fm[of 0 1] ord_simp_union cardinal_rel_closed transitivity zero_in_M + Lambda_in_M_replacement2(5) + by simp_all + +lemma (in M_basic) rel2_trans_apply: + "M(f) \ relation2(M,is_trans_apply_image(M,f),trans_apply_image(f))" + unfolding is_trans_apply_image_def trans_apply_image_def relation2_def + by auto + +lemma (in M_basic) apply_image_closed: + shows "M(f) \ \x[M]. \g[M]. function(g) \ M(trans_apply_image(f, x, g))" + unfolding trans_apply_image_def by simp + +lemma (in M_basic) apply_image_closed': + shows "M(f) \ \x[M]. \g[M]. M(trans_apply_image(f, x, g))" + unfolding trans_apply_image_def by simp + +lemma (in M_ZF2_trans) replacement_transrec_apply_image_body : + "(##M)(f) \ (##M)(mesa) \ strong_replacement(##M,transrec_apply_image_body(##M,f,mesa))" + using strong_replacement_rel_in_ctm[where \="transrec_apply_image_body_fm(3,2,0,1)" and env="[mesa,f]"] + zero_in_M arity_transrec_apply_image_body_fm ord_simp_union + replacement_ax2(6) + by simp + +lemma (in M_ZF2_trans) transrec_replacement_apply_image: + assumes "(##M)(f)" "(##M)(\)" + shows "transrec_replacement(##M, is_trans_apply_image(##M, f), \)" + unfolding transrec_replacement_def wfrec_replacement_def is_wfrec_def M_is_recfun_def + using replacement_transrec_apply_image_body[unfolded transrec_apply_image_body_def] assms + Memrel_closed singleton_closed eclose_closed + by simp + +lemma (in M_ZF2_trans) rec_trans_apply_image_abs: + assumes "(##M)(f)" "(##M)(x)" "(##M)(y)" "Ord(x)" + shows "is_transrec(##M,is_trans_apply_image(##M, f),x,y) \ y = transrec(x,trans_apply_image(f))" + using transrec_abs[OF transrec_replacement_apply_image rel2_trans_apply] assms apply_image_closed + by simp + +lemma (in M_ZF2_trans) replacement_is_trans_apply_image: + "(##M)(f) \ (##M)(\) \ strong_replacement(##M, \ x z . + \y[##M]. pair(##M,x,y,z) \ x\\ \ (is_transrec(##M,is_trans_apply_image(##M, f),x,y)))" + unfolding is_transrec_def is_wfrec_def M_is_recfun_def + apply(rule_tac strong_replacement_cong[ + where P="\ x z. M,[x,z,\,f] \ is_trans_apply_image_body_fm(3,2,0,1)",THEN iffD1]) + apply(rule_tac is_trans_apply_image_body_iff_sats[symmetric,unfolded is_trans_apply_image_body_def,where env="[_,_,\,f]"]) + apply(simp_all add:zero_in_M) + apply(rule_tac replacement_ax2(8)[unfolded replacement_assm_def, rule_format, where env="[\,f]",simplified]) + apply(simp_all add: arity_is_trans_apply_image_body_fm is_trans_apply_image_body_fm_type ord_simp_union) + done + +lemma (in M_ZF2_trans) trans_apply_abs: + "(##M)(f) \ (##M)(\) \ Ord(\) \ (##M)(x) \ (##M)(z) \ + (x\\ \ z = \x, transrec(x, \a g. f ` (g `` a)) \) \ + (\y[##M]. pair(##M,x,y,z) \ x\\ \ (is_transrec(##M,is_trans_apply_image(##M, f),x,y)))" + using rec_trans_apply_image_abs Ord_in_Ord + transrec_closed[OF transrec_replacement_apply_image rel2_trans_apply,of f,simplified] + apply_image_closed'[of f] + unfolding trans_apply_image_def + by auto + +lemma (in M_ZF2_trans) replacement_trans_apply_image: + "(##M)(f) \ (##M)(\) \ Ord(\) \ + strong_replacement(##M, \x y. x\\ \ y = \x, transrec(x, \a g. f ` (g `` a))\)" + using strong_replacement_cong[THEN iffD1,OF _ replacement_is_trans_apply_image,simplified] + trans_apply_abs Ord_in_Ord + by simp + +definition ifrFb_body where + "ifrFb_body(M,b,f,x,i) \ x \ + (if b = 0 then if i \ range(f) then + if M(converse(f) ` i) then converse(f) ` i else 0 else 0 else if M(i) then i else 0)" + +relativize functional "ifrFb_body" "ifrFb_body_rel" +relationalize "ifrFb_body_rel" "is_ifrFb_body" + +synthesize "is_ifrFb_body" from_definition assuming "nonempty" +arity_theorem for "is_ifrFb_body_fm" + +definition ifrangeF_body :: "[i\o,i,i,i,i] \ o" where + "ifrangeF_body(M,A,b,f) \ \y. \x\A. y = \x,\ i. ifrFb_body(M,b,f,x,i)\" + +relativize functional "ifrangeF_body" "ifrangeF_body_rel" +relationalize "ifrangeF_body_rel" "is_ifrangeF_body" + +synthesize "is_ifrangeF_body" from_definition assuming "nonempty" +arity_theorem for "is_ifrangeF_body_fm" + +lemma (in M_Z_trans) separation_is_ifrangeF_body: + "(##M)(A) \ (##M)(r) \ (##M)(s) \ separation(##M, is_ifrangeF_body(##M,A,r,s))" + using separation_in_ctm[where \="is_ifrangeF_body_fm(1,2,3,0)" and env="[A,r,s]"] + zero_in_M arity_is_ifrangeF_body_fm ord_simp_union is_ifrangeF_body_fm_type + by simp + +lemma (in M_basic) is_ifrFb_body_closed: "M(r) \ M(s) \ is_ifrFb_body(M, r, s, x, i) \ M(i)" + unfolding ifrangeF_body_def is_ifrangeF_body_def is_ifrFb_body_def If_abs + by (cases "i\range(s)"; cases "r=0"; auto dest:transM) + +lemma (in M_ZF1_trans) ifrangeF_body_abs: + assumes "(##M)(A)" "(##M)(r)" "(##M)(s)" "(##M)(x)" + shows "is_ifrangeF_body(##M,A,r,s,x) \ ifrangeF_body(##M,A,r,s,x)" +proof - + { + fix a + assume "a\M" + with assms + have "(\ i. i\ M \ is_ifrFb_body(##M, r, s, z, i))= (\ i. is_ifrFb_body(##M, r, s, z, i))" for z + using is_ifrFb_body_closed[of r s z] + by (rule_tac Least_cong[of "\i. i\M \ is_ifrFb_body(##M,r,s,z,i)"]) auto + moreover + have "(\ i. is_ifrFb_body(##M, r, s, z, i))= (\ i. ifrFb_body(##M, r, s, z, i))" for z + proof (rule_tac Least_cong[of "\i. is_ifrFb_body(##M,r,s,z,i)" "\i. ifrFb_body(##M,r,s,z,i)"]) + fix y + from assms \a\M\ + show "is_ifrFb_body(##M, r, s, z, y) \ ifrFb_body(##M, r, s, z, y)" + using If_abs apply_0 + unfolding ifrFb_body_def is_ifrFb_body_def + by (cases "y\M"; cases "y\range(s)"; cases "converse(s)`y \ M"; + auto dest:transM split del: split_if del:iffI) + (auto simp flip:setclass_iff; (force simp only:setclass_iff))+ + qed + moreover from \a\M\ + have "least(##M, \i. i \ M \ is_ifrFb_body(##M, r, s, z, i), a) + \ a = (\ i. i\ M \ is_ifrFb_body(##M, r, s, z,i))" for z + using If_abs least_abs'[of "\i. (##M)(i) \ is_ifrFb_body(##M,r,s,z,i)" a] + by simp + ultimately + have "least(##M, \i. i \ M \ is_ifrFb_body(##M, r, s, z, i), a) + \ a = (\ i. ifrFb_body(##M, r, s, z,i))" for z + by simp + } + with assms + show ?thesis + using pair_in_M_iff apply_closed zero_in_M transitivity[of _ A] + unfolding ifrangeF_body_def is_ifrangeF_body_def + by (auto dest:transM) +qed + +lemma (in M_ZF1_trans) separation_ifrangeF_body: + "(##M)(A) \ (##M)(b) \ (##M)(f) \ separation + (##M, \y. \x\A. y = \x, \ i. x \ if_range_F_else_F(\x. if (##M)(x) then x else 0, b, f, i)\)" + using separation_is_ifrangeF_body ifrangeF_body_abs + separation_cong[where P="is_ifrangeF_body(##M,A,b,f)" and M="##M",THEN iffD1] + unfolding ifrangeF_body_def if_range_F_def if_range_F_else_F_def ifrFb_body_def + by simp + +(* (##M)(A) \ (##M)(G) \ (##M)(b) \ (##M)(f) \ + separation(##M, + \y. \x\A. y = \x, \ i. x \ if_range_F_else_F(\a. if (##M)(a) then G`a else 0, b, f, i)\) *) + +definition ifrFb_body2 where + "ifrFb_body2(M,G,b,f,x,i) \ x \ + (if b = 0 then if i \ range(f) then + if M(converse(f) ` i) then G`(converse(f) ` i) else 0 else 0 else if M(i) then G`i else 0)" + +relativize functional "ifrFb_body2" "ifrFb_body2_rel" +relationalize "ifrFb_body2_rel" "is_ifrFb_body2" + +synthesize "is_ifrFb_body2" from_definition assuming "nonempty" +arity_theorem for "is_ifrFb_body2_fm" + +definition ifrangeF_body2 :: "[i\o,i,i,i,i,i] \ o" where + "ifrangeF_body2(M,A,G,b,f) \ \y. \x\A. y = \x,\ i. ifrFb_body2(M,G,b,f,x,i)\" + +relativize functional "ifrangeF_body2" "ifrangeF_body2_rel" +relationalize "ifrangeF_body2_rel" "is_ifrangeF_body2" + +synthesize "is_ifrangeF_body2" from_definition assuming "nonempty" +arity_theorem for "is_ifrangeF_body2_fm" + +lemma (in M_Z_trans) separation_is_ifrangeF_body2: + "(##M)(A) \ (##M)(G) \ (##M)(r) \ (##M)(s) \ separation(##M, is_ifrangeF_body2(##M,A,G,r,s))" + using separation_in_ctm[where \="is_ifrangeF_body2_fm(1,2,3,4,0)" and env="[A,G,r,s]"] + zero_in_M arity_is_ifrangeF_body2_fm ord_simp_union is_ifrangeF_body2_fm_type + by simp + +lemma (in M_basic) is_ifrFb_body2_closed: "M(G) \ M(r) \ M(s) \ is_ifrFb_body2(M, G, r, s, x, i) \ M(i)" + unfolding ifrangeF_body2_def is_ifrangeF_body2_def is_ifrFb_body2_def If_abs + by (cases "i\range(s)"; cases "r=0"; auto dest:transM) + +lemma (in M_ZF1_trans) ifrangeF_body2_abs: + assumes "(##M)(A)" "(##M)(G)" "(##M)(r)" "(##M)(s)" "(##M)(x)" + shows "is_ifrangeF_body2(##M,A,G,r,s,x) \ ifrangeF_body2(##M,A,G,r,s,x)" +proof - + { + fix a + assume "a\M" + with assms + have "(\ i. i\ M \ is_ifrFb_body2(##M, G, r, s, z, i))= (\ i. is_ifrFb_body2(##M, G, r, s, z, i))" for z + using is_ifrFb_body2_closed[of G r s z] + by (rule_tac Least_cong[of "\i. i\M \ is_ifrFb_body2(##M,G,r,s,z,i)"]) auto + moreover + have "(\ i. is_ifrFb_body2(##M, G, r, s, z, i))= (\ i. ifrFb_body2(##M, G, r, s, z, i))" for z + proof (rule_tac Least_cong[of "\i. is_ifrFb_body2(##M,G,r,s,z,i)" "\i. ifrFb_body2(##M,G,r,s,z,i)"]) + fix y + from assms \a\M\ + show "is_ifrFb_body2(##M, G, r, s, z, y) \ ifrFb_body2(##M, G, r, s, z, y)" + using If_abs apply_0 + unfolding ifrFb_body2_def is_ifrFb_body2_def + by (cases "y\M"; cases "y\range(s)"; cases "converse(s)`y \ M"; + auto dest:transM split del: split_if del:iffI) + (auto simp flip:setclass_iff; (force simp only:setclass_iff))+ + qed + moreover from \a\M\ + have "least(##M, \i. i \ M \ is_ifrFb_body2(##M, G, r, s, z, i), a) + \ a = (\ i. i\ M \ is_ifrFb_body2(##M, G, r, s, z,i))" for z + using If_abs least_abs'[of "\i. (##M)(i) \ is_ifrFb_body2(##M,G,r,s,z,i)" a] + by simp + ultimately + have "least(##M, \i. i \ M \ is_ifrFb_body2(##M, G, r, s, z, i), a) + \ a = (\ i. ifrFb_body2(##M, G, r, s, z,i))" for z + by simp + } + with assms + show ?thesis + using pair_in_M_iff apply_closed zero_in_M transitivity[of _ A] + unfolding ifrangeF_body2_def is_ifrangeF_body2_def + by (auto dest:transM) +qed + +lemma (in M_ZF1_trans) separation_ifrangeF_body2: + "(##M)(A) \ (##M)(G) \ (##M)(b) \ (##M)(f) \ + separation + (##M, + \y. \x\A. + y = + \x, \ i. x \ + if_range_F_else_F(\a. if (##M)(a) then G ` a else 0, b, f, i)\)" + using separation_is_ifrangeF_body2 ifrangeF_body2_abs + separation_cong[where P="is_ifrangeF_body2(##M,A,G,b,f)" and M="##M",THEN iffD1] + unfolding ifrangeF_body2_def if_range_F_def if_range_F_else_F_def ifrFb_body2_def + by simp + +(* (##M)(A) \ (##M)(b) \ (##M)(f) \ (##M)(F) \ + separation(##M, + \y. \x\A. y = \x, \ i. x \ if_range_F_else_F(\a. if (##M)(a) then F -`` {a} else 0, b, f, i)\) *) + +definition ifrFb_body3 where + "ifrFb_body3(M,G,b,f,x,i) \ x \ + (if b = 0 then if i \ range(f) then + if M(converse(f) ` i) then G-``{converse(f) ` i} else 0 else 0 else if M(i) then G-``{i} else 0)" + +relativize functional "ifrFb_body3" "ifrFb_body3_rel" +relationalize "ifrFb_body3_rel" "is_ifrFb_body3" + +synthesize "is_ifrFb_body3" from_definition assuming "nonempty" +arity_theorem for "is_ifrFb_body3_fm" + +definition ifrangeF_body3 :: "[i\o,i,i,i,i,i] \ o" where + "ifrangeF_body3(M,A,G,b,f) \ \y. \x\A. y = \x,\ i. ifrFb_body3(M,G,b,f,x,i)\" + +relativize functional "ifrangeF_body3" "ifrangeF_body3_rel" +relationalize "ifrangeF_body3_rel" "is_ifrangeF_body3" + +synthesize "is_ifrangeF_body3" from_definition assuming "nonempty" +arity_theorem for "is_ifrangeF_body3_fm" + +lemma (in M_Z_trans) separation_is_ifrangeF_body3: + "(##M)(A) \ (##M)(G) \ (##M)(r) \ (##M)(s) \ separation(##M, is_ifrangeF_body3(##M,A,G,r,s))" + using separation_in_ctm[where \="is_ifrangeF_body3_fm(1,2,3,4,0)" and env="[A,G,r,s]"] + zero_in_M arity_is_ifrangeF_body3_fm ord_simp_union is_ifrangeF_body3_fm_type + by simp + +lemma (in M_basic) is_ifrFb_body3_closed: "M(G) \ M(r) \ M(s) \ is_ifrFb_body3(M, G, r, s, x, i) \ M(i)" + unfolding ifrangeF_body3_def is_ifrangeF_body3_def is_ifrFb_body3_def If_abs + by (cases "i\range(s)"; cases "r=0"; auto dest:transM) + +lemma (in M_ZF1_trans) ifrangeF_body3_abs: + assumes "(##M)(A)" "(##M)(G)" "(##M)(r)" "(##M)(s)" "(##M)(x)" + shows "is_ifrangeF_body3(##M,A,G,r,s,x) \ ifrangeF_body3(##M,A,G,r,s,x)" +proof - + { + fix a + assume "a\M" + with assms + have "(\ i. i\ M \ is_ifrFb_body3(##M, G, r, s, z, i))= (\ i. is_ifrFb_body3(##M, G, r, s, z, i))" for z + using is_ifrFb_body3_closed[of G r s z] + by (rule_tac Least_cong[of "\i. i\M \ is_ifrFb_body3(##M,G,r,s,z,i)"]) auto + moreover + have "(\ i. is_ifrFb_body3(##M, G, r, s, z, i))= (\ i. ifrFb_body3(##M, G, r, s, z, i))" for z + proof (rule_tac Least_cong[of "\i. is_ifrFb_body3(##M,G,r,s,z,i)" "\i. ifrFb_body3(##M,G,r,s,z,i)"]) + fix y + from assms \a\M\ + show "is_ifrFb_body3(##M, G, r, s, z, y) \ ifrFb_body3(##M, G, r, s, z, y)" + using If_abs apply_0 + unfolding ifrFb_body3_def is_ifrFb_body3_def + by (cases "y\M"; cases "y\range(s)"; cases "converse(s)`y \ M"; + auto dest:transM split del: split_if del:iffI) + (auto simp flip:setclass_iff; (force simp only:setclass_iff))+ + qed + moreover from \a\M\ + have "least(##M, \i. i \ M \ is_ifrFb_body3(##M, G, r, s, z, i), a) + \ a = (\ i. i\ M \ is_ifrFb_body3(##M, G, r, s, z,i))" for z + using If_abs least_abs'[of "\i. (##M)(i) \ is_ifrFb_body3(##M,G,r,s,z,i)" a] + by simp + ultimately + have "least(##M, \i. i \ M \ is_ifrFb_body3(##M, G, r, s, z, i), a) + \ a = (\ i. ifrFb_body3(##M, G, r, s, z,i))" for z + by simp + } + with assms + show ?thesis + using pair_in_M_iff apply_closed zero_in_M transitivity[of _ A] + unfolding ifrangeF_body3_def is_ifrangeF_body3_def + by (auto dest:transM) +qed + +lemma (in M_ZF1_trans) separation_ifrangeF_body3: + "(##M)(A) \ (##M)(G) \ (##M)(b) \ (##M)(f) \ + separation + (##M, + \y. \x\A. + y = + \x, \ i. x \ + if_range_F_else_F(\a. if (##M)(a) then G-``{a} else 0, b, f, i)\)" + using separation_is_ifrangeF_body3 ifrangeF_body3_abs + separation_cong[where P="is_ifrangeF_body3(##M,A,G,b,f)" and M="##M",THEN iffD1] + unfolding ifrangeF_body3_def if_range_F_def if_range_F_else_F_def ifrFb_body3_def + by simp + +(* (##M)(A) \ (##M)(b) \ (##M)(f) \ (##M)(A') \ + separation(##M, \y. \x\A'. y = \x, \ i. x \ if_range_F_else_F((`)(A), b, f, i)\) *) + +definition ifrFb_body4 where + "ifrFb_body4(G,b,f,x,i) \ x \ + (if b = 0 then if i \ range(f) then G`(converse(f) ` i) else 0 else G`i)" + +relativize functional "ifrFb_body4" "ifrFb_body4_rel" +relationalize "ifrFb_body4_rel" "is_ifrFb_body4" + +synthesize "is_ifrFb_body4" from_definition assuming "nonempty" +arity_theorem for "is_ifrFb_body4_fm" + +definition ifrangeF_body4 :: "[i\o,i,i,i,i,i] \ o" where + "ifrangeF_body4(M,A,G,b,f) \ \y. \x\A. y = \x,\ i. ifrFb_body4(G,b,f,x,i)\" + +relativize functional "ifrangeF_body4" "ifrangeF_body4_rel" +relationalize "ifrangeF_body4_rel" "is_ifrangeF_body4" + +synthesize "is_ifrangeF_body4" from_definition assuming "nonempty" +arity_theorem for "is_ifrangeF_body4_fm" + +lemma (in M_Z_trans) separation_is_ifrangeF_body4: + "(##M)(A) \ (##M)(G) \ (##M)(r) \ (##M)(s) \ separation(##M, is_ifrangeF_body4(##M,A,G,r,s))" + using separation_in_ctm[where \="is_ifrangeF_body4_fm(1,2,3,4,0)" and env="[A,G,r,s]"] + zero_in_M arity_is_ifrangeF_body4_fm ord_simp_union is_ifrangeF_body4_fm_type + by simp + +lemma (in M_basic) is_ifrFb_body4_closed: "M(G) \ M(r) \ M(s) \ is_ifrFb_body4(M, G, r, s, x, i) \ M(i)" + using If_abs + unfolding ifrangeF_body4_def is_ifrangeF_body4_def is_ifrFb_body4_def fun_apply_def + by (cases "i\range(s)"; cases "r=0"; auto dest:transM) + +lemma (in M_ZF1_trans) ifrangeF_body4_abs: + assumes "(##M)(A)" "(##M)(G)" "(##M)(r)" "(##M)(s)" "(##M)(x)" + shows "is_ifrangeF_body4(##M,A,G,r,s,x) \ ifrangeF_body4(##M,A,G,r,s,x)" +proof - + { + fix a + assume "a\M" + with assms + have "(\ i. i\ M \ is_ifrFb_body4(##M, G, r, s, z, i))= (\ i. is_ifrFb_body4(##M, G, r, s, z, i))" for z + using is_ifrFb_body4_closed[of G r s z] + by (rule_tac Least_cong[of "\i. i\M \ is_ifrFb_body4(##M,G,r,s,z,i)"]) auto + moreover + have "(\ i. is_ifrFb_body4(##M, G, r, s, z, i))= (\ i. ifrFb_body4(G, r, s, z, i))" if "z\M" for z + proof (rule_tac Least_cong[of "\i. is_ifrFb_body4(##M,G,r,s,z,i)" "\i. ifrFb_body4(G,r,s,z,i)"]) + fix y + from assms \a\M\ \z\M\ + show "is_ifrFb_body4(##M, G, r, s, z, y) \ ifrFb_body4(G, r, s, z, y)" + using If_abs apply_0 + unfolding ifrFb_body4_def is_ifrFb_body4_def + apply (cases "y\M"; cases "y\range(s)"; cases "r=0"; cases "y\domain(G)"; + auto dest:transM split del: split_if del:iffI) + by (auto simp flip:setclass_iff; (force simp only: fun_apply_def setclass_iff)) + (auto simp flip:setclass_iff simp: fun_apply_def ) + qed + moreover from \a\M\ + have "least(##M, \i. i \ M \ is_ifrFb_body4(##M, G, r, s, z, i), a) + \ a = (\ i. i\ M \ is_ifrFb_body4(##M, G, r, s, z,i))" for z + using If_abs least_abs'[of "\i. (##M)(i) \ is_ifrFb_body4(##M,G,r,s,z,i)" a] + by simp + ultimately + have "z\M \ least(##M, \i. i \ M \ is_ifrFb_body4(##M, G, r, s, z, i), a) + \ a = (\ i. ifrFb_body4(G, r, s, z,i))" for z + by simp + } + with assms + show ?thesis + using pair_in_M_iff apply_closed zero_in_M transitivity[of _ A] + unfolding ifrangeF_body4_def is_ifrangeF_body4_def + by (auto dest:transM) +qed + +lemma (in M_ZF1_trans) separation_ifrangeF_body4: + "(##M)(A) \ (##M)(G) \ (##M)(b) \ (##M)(f) \ + separation(##M, \y. \x\A. y = \x, \ i. x \ if_range_F_else_F((`)(G), b, f, i)\)" + using separation_is_ifrangeF_body4 ifrangeF_body4_abs + separation_cong[where P="is_ifrangeF_body4(##M,A,G,b,f)" and M="##M",THEN iffD1] + unfolding ifrangeF_body4_def if_range_F_def if_range_F_else_F_def ifrFb_body4_def + by simp + +(* (##M)(G) \ (##M)(A) \ + separation(##M, + \y. \x\A. y = \x, \ i. x \ if_range_F_else_F(\x. {xa \ G . x \ xa}, b, f, i)\) *) + +definition ifrFb_body5 where + "ifrFb_body5(G,b,f,x,i) \ x \ + (if b = 0 then if i \ range(f) then {xa \ G . converse(f) ` i \ xa} else 0 else {xa \ G . i \ xa})" + +relativize functional "ifrFb_body5" "ifrFb_body5_rel" +relationalize "ifrFb_body5_rel" "is_ifrFb_body5" + +synthesize "is_ifrFb_body5" from_definition assuming "nonempty" +arity_theorem for "is_ifrFb_body5_fm" + +definition ifrangeF_body5 :: "[i\o,i,i,i,i,i] \ o" where + "ifrangeF_body5(M,A,G,b,f) \ \y. \x\A. y = \x,\ i. ifrFb_body5(G,b,f,x,i)\" + +relativize functional "ifrangeF_body5" "ifrangeF_body5_rel" +relationalize "ifrangeF_body5_rel" "is_ifrangeF_body5" + +synthesize "is_ifrangeF_body5" from_definition assuming "nonempty" +arity_theorem for "is_ifrangeF_body5_fm" + +lemma (in M_Z_trans) separation_is_ifrangeF_body5: + "(##M)(A) \ (##M)(G) \ (##M)(r) \ (##M)(s) \ separation(##M, is_ifrangeF_body5(##M,A,G,r,s))" + using separation_in_ctm[where \="is_ifrangeF_body5_fm(1,2,3,4,0)" and env="[A,G,r,s]"] + zero_in_M arity_is_ifrangeF_body5_fm ord_simp_union is_ifrangeF_body5_fm_type + by simp + +lemma (in M_basic) is_ifrFb_body5_closed: "M(G) \ M(r) \ M(s) \ is_ifrFb_body5(M, G, r, s, x, i) \ M(i)" + using If_abs + unfolding ifrangeF_body5_def is_ifrangeF_body5_def is_ifrFb_body5_def fun_apply_def + by (cases "i\range(s)"; cases "r=0"; auto dest:transM) + +lemma (in M_ZF1_trans) ifrangeF_body5_abs: + assumes "(##M)(A)" "(##M)(G)" "(##M)(r)" "(##M)(s)" "(##M)(x)" + shows "is_ifrangeF_body5(##M,A,G,r,s,x) \ ifrangeF_body5(##M,A,G,r,s,x)" +proof - + { + fix a + assume "a\M" + with assms + have "(\ i. i\ M \ is_ifrFb_body5(##M, G, r, s, z, i))= (\ i. is_ifrFb_body5(##M, G, r, s, z, i))" for z + using is_ifrFb_body5_closed[of G r s z] + by (rule_tac Least_cong[of "\i. i\M \ is_ifrFb_body5(##M,G,r,s,z,i)"]) auto + moreover + have "(\ i. is_ifrFb_body5(##M, G, r, s, z, i))= (\ i. ifrFb_body5(G, r, s, z, i))" if "z\M" for z + proof (rule_tac Least_cong[of "\i. is_ifrFb_body5(##M,G,r,s,z,i)" "\i. ifrFb_body5(G,r,s,z,i)"]) + fix y + from assms \a\M\ \z\M\ + show "is_ifrFb_body5(##M, G, r, s, z, y) \ ifrFb_body5(G, r, s, z, y)" + using If_abs apply_0 separation_in_constant separation_in_rev + unfolding ifrFb_body5_def is_ifrFb_body5_def + apply (cases "y\M"; cases "y\range(s)"; cases "r=0"; cases "y\domain(G)"; + auto dest:transM split del: split_if del:iffI) + apply (auto simp flip:setclass_iff; (force simp only: fun_apply_def setclass_iff)) + apply (auto simp flip:setclass_iff simp: fun_apply_def) + apply (auto dest:transM) + done + qed + moreover from \a\M\ + have "least(##M, \i. i \ M \ is_ifrFb_body5(##M, G, r, s, z, i), a) + \ a = (\ i. i\ M \ is_ifrFb_body5(##M, G, r, s, z,i))" for z + using If_abs least_abs'[of "\i. (##M)(i) \ is_ifrFb_body5(##M,G,r,s,z,i)" a] + by simp + ultimately + have "z\M \ least(##M, \i. i \ M \ is_ifrFb_body5(##M, G, r, s, z, i), a) + \ a = (\ i. ifrFb_body5(G, r, s, z,i))" for z + by simp + } + with assms + show ?thesis + using pair_in_M_iff apply_closed zero_in_M transitivity[of _ A] + unfolding ifrangeF_body5_def is_ifrangeF_body5_def + by (auto dest:transM) +qed + +lemma (in M_ZF1_trans) separation_ifrangeF_body5: + "(##M)(A) \ (##M)(G) \ (##M)(b) \ (##M)(f) \ + separation(##M, \y. \x\A. y = \x, \ i. x \ if_range_F_else_F(\x. {xa \ G . x \ xa}, b, f, i)\)" + using separation_is_ifrangeF_body5 ifrangeF_body5_abs + separation_cong[where P="is_ifrangeF_body5(##M,A,G,b,f)" and M="##M",THEN iffD1] + unfolding ifrangeF_body5_def if_range_F_def if_range_F_else_F_def ifrFb_body5_def + by simp + +(* (##M)(A) \ (##M)(b) \ (##M)(f) \ + separation(##M, + \y. \x\A'. y = \x, \ i. x \ if_range_F_else_F(\a. {p \ A . domain(p) = a}, b, f, i)\) *) + +definition ifrFb_body6 where + "ifrFb_body6(G,b,f,x,i) \ x \ + (if b = 0 then if i \ range(f) then {p\G . domain(p) = converse(f) ` i} else 0 else {p\G . domain(p) = i})" + +relativize functional "ifrFb_body6" "ifrFb_body6_rel" +relationalize "ifrFb_body6_rel" "is_ifrFb_body6" + +synthesize "is_ifrFb_body6" from_definition assuming "nonempty" +arity_theorem for "is_ifrFb_body6_fm" + +definition ifrangeF_body6 :: "[i\o,i,i,i,i,i] \ o" where + "ifrangeF_body6(M,A,G,b,f) \ \y. \x\A. y = \x,\ i. ifrFb_body6(G,b,f,x,i)\" + +relativize functional "ifrangeF_body6" "ifrangeF_body6_rel" +relationalize "ifrangeF_body6_rel" "is_ifrangeF_body6" + +synthesize "is_ifrangeF_body6" from_definition assuming "nonempty" +arity_theorem for "is_ifrangeF_body6_fm" + +lemma (in M_Z_trans) separation_is_ifrangeF_body6: + "(##M)(A) \ (##M)(G) \ (##M)(r) \ (##M)(s) \ separation(##M, is_ifrangeF_body6(##M,A,G,r,s))" + using separation_in_ctm[where \="is_ifrangeF_body6_fm(1,2,3,4,0)" and env="[A,G,r,s]"] + zero_in_M arity_is_ifrangeF_body6_fm ord_simp_union is_ifrangeF_body6_fm_type + by simp + +lemma (in M_basic) ifrFb_body6_closed: "M(G) \ M(r) \ M(s) \ ifrFb_body6(G, r, s, x, i) \ M(i) \ ifrFb_body6(G, r, s, x, i)" + using If_abs + unfolding ifrangeF_body6_def is_ifrangeF_body6_def ifrFb_body6_def fun_apply_def + by (cases "i\range(s)"; cases "r=0"; auto dest:transM) + +lemma (in M_basic) is_ifrFb_body6_closed: "M(G) \ M(r) \ M(s) \ is_ifrFb_body6(M, G, r, s, x, i) \ M(i)" + using If_abs + unfolding ifrangeF_body6_def is_ifrangeF_body6_def is_ifrFb_body6_def fun_apply_def + by (cases "i\range(s)"; cases "r=0"; auto dest:transM) + +lemma (in M_ZF2_trans) ifrangeF_body6_abs: + assumes "(##M)(A)" "(##M)(G)" "(##M)(r)" "(##M)(s)" "(##M)(x)" + shows "is_ifrangeF_body6(##M,A,G,r,s,x) \ ifrangeF_body6(##M,A,G,r,s,x)" +proof - + { + fix a + assume "a\M" + with assms + have "(\ i. i\ M \ is_ifrFb_body6(##M, G, r, s, z, i))= (\ i. is_ifrFb_body6(##M, G, r, s, z, i))" for z + using is_ifrFb_body6_closed[of G r s z] + by (rule_tac Least_cong[of "\i. i\M \ is_ifrFb_body6(##M,G,r,s,z,i)"]) auto + moreover + have "(\ i. i\M \ is_ifrFb_body6(##M, G, r, s, z, i))= (\ i. i\M \ ifrFb_body6(G, r, s, z, i))" if "z\M" for z + proof (rule_tac Least_cong[of "\i. i\M \ is_ifrFb_body6(##M,G,r,s,z,i)" "\i. i\M \ ifrFb_body6(G,r,s,z,i)"]) + fix y + from assms \a\M\ \z\M\ + show "y\M \ is_ifrFb_body6(##M, G, r, s, z, y) \ y\M \ ifrFb_body6(G, r, s, z, y)" + using If_abs apply_0 separation_in_constant transitivity[of _ G] + separation_closed converse_closed apply_closed range_closed zero_in_M + separation_cong[OF eq_commute,THEN iffD1,OF domain_eq_separation] + unfolding ifrFb_body6_def is_ifrFb_body6_def + by auto + qed + moreover from \a\M\ + have "least(##M, \i. i \ M \ is_ifrFb_body6(##M, G, r, s, z, i), a) + \ a = (\ i. i\ M \ is_ifrFb_body6(##M, G, r, s, z,i))" for z + using If_abs least_abs'[of "\i. (##M)(i) \ is_ifrFb_body6(##M,G,r,s,z,i)" a] + by simp + ultimately + have "z\M \ least(##M, \i. i \ M \ is_ifrFb_body6(##M, G, r, s, z, i), a) + \ a = (\ i. ifrFb_body6(G, r, s, z,i))" for z + using Least_cong[OF ifrFb_body6_closed[of G r s]] assms + by simp + } + with assms + show ?thesis + using pair_in_M_iff apply_closed zero_in_M transitivity[of _ A] + unfolding ifrangeF_body6_def is_ifrangeF_body6_def + by (auto dest:transM) +qed + +lemma (in M_ZF2_trans) separation_ifrangeF_body6: + "(##M)(A) \ (##M)(G) \ (##M)(b) \ (##M)(f) \ + separation(##M, + \y. \x\A. y = \x, \ i. x \ if_range_F_else_F(\a. {p \ G . domain(p) = a}, b, f, i)\)" + using separation_is_ifrangeF_body6 ifrangeF_body6_abs + separation_cong[where P="is_ifrangeF_body6(##M,A,G,b,f)" and M="##M",THEN iffD1] + unfolding ifrangeF_body6_def if_range_F_def if_range_F_else_F_def ifrFb_body6_def + by simp + + + +(* (##M)(A) \ (##M)(f) \ (##M)(b) \ (##M)(D) \ (##M)(r') \ (##M)(A') \ + separation(##M, + \y. \x\A'. y = \x, \ i. x \ if_range_F_else_F(drSR_Y(r', D, A), b, f, i)\) *) + +definition ifrFb_body7 where + "ifrFb_body7(B,D,A,b,f,x,i) \ x \ + (if b = 0 then if i \ range(f) then + {d \ D . \r\A. restrict(r, B) = converse(f) ` i \ d = domain(r)} else 0 + else {d \ D . \r\A. restrict(r, B) = i \ d = domain(r)})" + +relativize functional "ifrFb_body7" "ifrFb_body7_rel" +relationalize "ifrFb_body7_rel" "is_ifrFb_body7" + +synthesize "is_ifrFb_body7" from_definition assuming "nonempty" +arity_theorem for "is_ifrFb_body7_fm" + +definition ifrangeF_body7 :: "[i\o,i,i,i,i,i,i,i] \ o" where + "ifrangeF_body7(M,A,B,D,G,b,f) \ \y. \x\A. y = \x,\ i. ifrFb_body7(B,D,G,b,f,x,i)\" + +relativize functional "ifrangeF_body7" "ifrangeF_body7_rel" +relationalize "ifrangeF_body7_rel" "is_ifrangeF_body7" + +synthesize "is_ifrangeF_body7" from_definition assuming "nonempty" +arity_theorem for "is_ifrangeF_body7_fm" + +lemma (in M_Z_trans) separation_is_ifrangeF_body7: + "(##M)(A) \ (##M)(B) \ (##M)(D) \ (##M)(G) \ (##M)(r) \ (##M)(s) \ separation(##M, is_ifrangeF_body7(##M,A,B,D,G,r,s))" + using separation_in_ctm[where \="is_ifrangeF_body7_fm(1,2,3,4,5,6,0)" and env="[A,B,D,G,r,s]"] + zero_in_M arity_is_ifrangeF_body7_fm ord_simp_union is_ifrangeF_body7_fm_type + by simp + +lemma (in M_basic) ifrFb_body7_closed: "M(B) \ M(D) \ M(G) \ M(r) \ M(s) \ + ifrFb_body7(B,D,G, r, s, x, i) \ M(i) \ ifrFb_body7(B,D,G, r, s, x, i)" + using If_abs + unfolding ifrangeF_body7_def is_ifrangeF_body7_def ifrFb_body7_def fun_apply_def + by (cases "i\range(s)"; cases "r=0"; auto dest:transM) + +lemma (in M_basic) is_ifrFb_body7_closed: "M(B) \ M(D) \ M(G) \ M(r) \ M(s) \ + is_ifrFb_body7(M, B,D,G, r, s, x, i) \ M(i)" + using If_abs + unfolding ifrangeF_body7_def is_ifrangeF_body7_def is_ifrFb_body7_def fun_apply_def + by (cases "i\range(s)"; cases "r=0"; auto dest:transM) + +lemma (in M_ZF2_trans) ifrangeF_body7_abs: + assumes "(##M)(A)" "(##M)(B)" "(##M)(D)" "(##M)(G)" "(##M)(r)" "(##M)(s)" "(##M)(x)" + shows "is_ifrangeF_body7(##M,A,B,D,G,r,s,x) \ ifrangeF_body7(##M,A,B,D,G,r,s,x)" +proof - + from assms + have sep_dr: "y\M \ separation(##M, \d . \r\M . r\G\ y = restrict(r, B) \ d = domain(r))" for y + by(rule_tac separation_cong[where P'="\d . \r\ M . r\G \ y = restrict(r, B) \ d = domain(r)",THEN iffD1,OF _ + separation_restrict_eq_dom_eq[rule_format,of G B y]],auto simp:transitivity[of _ G]) + + from assms + have sep_dr'': "y\M \ separation(##M, \d . \r\M. r \ G \ d = domain(r) \ converse(s) ` y = restrict(r, B))" for y + apply(rule_tac separation_cong[where P'="\d . \r\ M . r\G \ d = domain(r) \ converse(s) ` y = restrict(r, B)",THEN iffD1,OF _ separation_restrict_eq_dom_eq[rule_format,of G B "converse(s) ` y "]]) + by(auto simp:transitivity[of _ G] apply_closed[simplified] converse_closed[simplified]) + from assms + have sep_dr':"separation(##M, \x. \r\M. r \ G \ x = domain(r) \ 0 = restrict(r, B))" + apply(rule_tac separation_cong[where P'="\d . \r\ M . r\G \ d = domain(r) \ 0 = restrict(r, B)",THEN iffD1,OF _ separation_restrict_eq_dom_eq[rule_format,of G B 0]]) + by(auto simp:transitivity[of _ G] zero_in_M) + { + fix a + assume "a\M" + with assms + have "(\ i. i\ M \ is_ifrFb_body7(##M, B,D,G, r, s, z, i))= (\ i. is_ifrFb_body7(##M,B,D, G, r, s, z, i))" for z + using is_ifrFb_body7_closed[of B D G r s z] + by (rule_tac Least_cong[of "\i. i\M \ is_ifrFb_body7(##M,B,D,G,r,s,z,i)"]) auto + moreover from this + have "(\ i. i\M \ is_ifrFb_body7(##M, B,D,G, r, s, z, i))= (\ i. i\M \ ifrFb_body7(B,D,G, r, s, z, i))" if "z\M" for z + proof (rule_tac Least_cong[of "\i. i\M \ is_ifrFb_body7(##M,B,D,G,r,s,z,i)" "\i. i\M \ ifrFb_body7(B,D,G,r,s,z,i)"]) + from assms \a\M\ \z\M\ + have "is_ifrFb_body7(##M, B,D,G, r, s, z, y) \ ifrFb_body7(B,D,G, r, s, z, y)" if "y\M" for y + using If_abs apply_0 + separation_closed converse_closed apply_closed range_closed zero_in_M + separation_restrict_eq_dom_eq + transitivity[of _ D] transitivity[of _ G] that sep_dr sep_dr' sep_dr'' + unfolding ifrFb_body7_def is_ifrFb_body7_def + by auto + then + show " y \ M \ is_ifrFb_body7(##M, B, D, G, r, s, z, y) \ y \ M \ ifrFb_body7(B, D, G, r, s, z, y)" for y + using conj_cong + by simp + qed + moreover from \a\M\ + have "least(##M, \i. i \ M \ is_ifrFb_body7(##M, B,D,G, r, s, z, i), a) + \ a = (\ i. i\ M \ is_ifrFb_body7(##M,B,D,G, r, s, z,i))" for z + using If_abs least_abs'[of "\i. (##M)(i) \ is_ifrFb_body7(##M,B,D,G,r,s,z,i)" a] + by simp + ultimately + have "z\M \ least(##M, \i. i \ M \ is_ifrFb_body7(##M,B,D,G, r, s, z, i), a) + \ a = (\ i. ifrFb_body7(B,D,G, r, s, z,i))" for z + using Least_cong[OF ifrFb_body7_closed[of B D G r s]] assms + by simp + } + with assms + show ?thesis + using pair_in_M_iff apply_closed zero_in_M transitivity[of _ A] + unfolding ifrangeF_body7_def is_ifrangeF_body7_def + by (auto dest:transM) +qed + +lemma (in M_ZF2_trans) separation_ifrangeF_body7: + "(##M)(A) \ (##M)(B) \ (##M)(D) \ (##M)(G) \ (##M)(b) \ (##M)(f) \ + separation(##M, + \y. \x\A. y = \x, \ i. x \ if_range_F_else_F(drSR_Y(B, D, G), b, f, i)\)" + using separation_is_ifrangeF_body7 ifrangeF_body7_abs drSR_Y_equality + separation_cong[where P="is_ifrangeF_body7(##M,A,B,D,G,b,f)" and M="##M",THEN iffD1] + unfolding ifrangeF_body7_def if_range_F_def if_range_F_else_F_def ifrFb_body7_def + by simp + +end \ No newline at end of file diff --git a/thys/Independence_CH/Separation_Axiom.thy b/thys/Independence_CH/Separation_Axiom.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Separation_Axiom.thy @@ -0,0 +1,399 @@ +section\The Axiom of Separation in $M[G]$\ +theory Separation_Axiom + imports Forcing_Theorems Separation_Rename +begin + +context G_generic1 +begin + +lemma map_val : + assumes "env\list(M[G])" + shows "\nenv\list(M). env = map(val(P,G),nenv)" + using assms + proof(induct env) + case Nil + have "map(val(P,G),Nil) = Nil" by simp + then show ?case by force + next + case (Cons a l) + then obtain a' l' where + "l' \ list(M)" "l=map(val(P,G),l')" "a = val(P,G,a')" + "Cons(a,l) = map(val(P,G),Cons(a',l'))" "Cons(a',l') \ list(M)" + using \a\M[G]\ GenExtD + by force + then show ?case by force +qed + +lemma Collect_sats_in_MG : + assumes + "c\M[G]" + "\ \ formula" "env\list(M[G])" "arity(\) \ 1 +\<^sub>\ length(env)" + shows + "{x\c. (M[G], [x] @ env \ \)}\ M[G]" +proof - + from \c\M[G]\ + obtain \ where "\ \ M" "val(P,G, \) = c" + using GenExt_def by auto + let ?\="\\ 0 \ (1 +\<^sub>\ length(env)) \ \ \ \" and ?Pl1="[P,leq,\]" + let ?new_form="sep_ren(length(env),forces(?\))" + let ?\="Exists(Exists(And(pair_fm(0,1,2),?new_form)))" + note phi = \\\formula\ \arity(\) \ 1 +\<^sub>\ length(env)\ + then + have "?\\formula" by simp + with \env\_\ phi + have "arity(?\) \ 2+\<^sub>\length(env) " + using ord_simp_union leI FOL_arities by simp + with \env\list(_)\ phi + have "arity(forces(?\)) \ 6 +\<^sub>\ length(env)" + using arity_forces_le by simp + then + have "arity(forces(?\)) \ 7 +\<^sub>\ length(env)" + using ord_simp_union arity_forces leI by simp + with \arity(forces(?\)) \7 +\<^sub>\ _\ \env \ _\ \\ \ formula\ + have "arity(?new_form) \ 7 +\<^sub>\ length(env)" "?new_form \ formula" + using arity_rensep[OF definability[of "?\"]] definability[of "?\"] type_rensep + by auto + then + have "pred(pred(arity(?new_form))) \ 5 +\<^sub>\ length(env)" "?\\formula" + unfolding pair_fm_def upair_fm_def + using ord_simp_union length_type[OF \env\list(M[G])\] + pred_mono[OF _ pred_mono[OF _ \arity(?new_form) \ _\]] + by auto + with \arity(?new_form) \ _\ \?new_form \ formula\ + have "arity(?\) \ 5 +\<^sub>\ length(env)" + unfolding pair_fm_def upair_fm_def + using ord_simp_union arity_forces + by (auto simp:arity) + from \\\formula\ + have "forces(?\) \ formula" + using definability by simp + from \\\M\ P_in_M + have "domain(\)\M" "domain(\) \ P \ M" + by (simp_all flip:setclass_iff) + from \env \ _\ + obtain nenv where "nenv\list(M)" "env = map(val(P,G),nenv)" "length(nenv) = length(env)" + using map_val by auto + from \arity(\) \ _\ \env\_\ \\\_\ + have "arity(\) \ 2+\<^sub>\ length(env)" + using le_trans[OF \arity(\)\_\] add_le_mono[of 1 2,OF _ le_refl] + by auto + with \nenv\_\ \env\_\ \\\M\ \\\_\ \length(nenv) = length(env)\ + have "arity(?\) \ length([\] @ nenv @ [\])" for \ + using union_abs2[OF \arity(\) \ 2+\<^sub>\ _\] ord_simp_union FOL_arities + by simp + note in_M = \\\M\ \domain(\) \ P \ M\ P_in_M one_in_M leq_in_M + { + fix u + assume "u \ domain(\) \ P" "u \ M" + with in_M \?new_form \ formula\ \?\\formula\ \nenv \ _\ + have Eq1: "(M, [u] @ ?Pl1 @ [\] @ nenv \ ?\) \ + (\\\M. \p\P. u =\\,p\ \ + (M, [\,p,u]@?Pl1@[\] @ nenv \ ?new_form))" + by (auto simp add: transitivity) + have Eq3: "\\M \ p\P \ + (M, [\,p,u]@?Pl1@[\]@nenv \ ?new_form) \ + (\F. M_generic(F) \ p \ F \ (M[F], map(val(P,F), [\] @ nenv@[\]) \ ?\))" + for \ p + proof - + fix p \ + assume "\ \ M" "p\P" + then + have "p\M" using P_in_M by (simp add: transitivity) + note in_M' = in_M \\ \ M\ \p\M\ \u \ domain(\) \ P\ \u \ M\ \nenv\_\ + then + have "[\,u] \ list(M)" by simp + let ?env="[p]@?Pl1@[\] @ nenv @ [\,u]" + let ?new_env=" [\,p,u,P,leq,\,\] @ nenv" + let ?\="Exists(Exists(And(pair_fm(0,1,2),?new_form)))" + have "[\, p, u, \, leq, \, \] \ list(M)" + using in_M' by simp + have "?\ \ formula" "forces(?\)\ formula" + using phi by simp_all + from in_M' + have "?Pl1 \ list(M)" by simp + from in_M' have "?env \ list(M)" by simp + have Eq1': "?new_env \ list(M)" using in_M' by simp + then + have "(M, [\,p,u]@?Pl1@[\] @ nenv \ ?new_form) \ (M, ?new_env \ ?new_form)" + by simp + from in_M' \env \ _\ Eq1' \length(nenv) = length(env)\ + \arity(forces(?\)) \ 7 +\<^sub>\ length(env)\ \forces(?\)\ formula\ + \[\, p, u, \, leq, \, \] \ list(M)\ + have "... \ M, ?env \ forces(?\)" + using sepren_action[of "forces(?\)" "nenv",OF _ _ \nenv\list(M)\] + by simp + also from in_M' + have "... \ M, ([p,P, leq, \,\]@nenv@ [\])@[u] \ forces(?\)" + using app_assoc by simp + also + from in_M' \env\_\ phi \length(nenv) = length(env)\ + \arity(forces(?\)) \ 6 +\<^sub>\ length(env)\ \forces(?\)\formula\ + have "... \ M, [p,P, leq, \,\]@ nenv @ [\] \ forces(?\)" + by (rule_tac arity_sats_iff,auto) + also + from \arity(forces(?\)) \ 6 +\<^sub>\ length(env)\ \forces(?\)\formula\ in_M' phi + have " ... \ (\F. M_generic(F) \ p \ F \ + M[F], map(val(P,F), [\] @ nenv @ [\]) \ ?\)" + proof (intro iffI) + assume a1: "M, [p,P, leq, \,\] @ nenv @ [\] \ forces(?\)" + note \arity(\)\ 1+\<^sub>\_\ + with \nenv\_\ \arity(?\) \ length([\] @ nenv @ [\])\ \env\_\ + have "p \ P \ ?\\formula \ [\,\] \ list(M) \ + M, [p,P, leq, \] @ [\]@ nenv@[\] \ forces(?\) \ + \G. M_generic(G) \ p \ G \ M[G], map(val(P,G), [\] @ nenv @[\]) \ ?\" + using definition_of_forcing[where \="\\ 0 \ (1 +\<^sub>\ length(env)) \ \ \ \"] + by auto + then + show "\F. M_generic(F) \ p \ F \ + M[F], map(val(P,F), [\] @ nenv @ [\]) \ ?\" + using \?\\formula\ \p\P\ a1 \\\M\ \\\M\ by simp + next + assume "\F. M_generic(F) \ p \ F \ + M[F], map(val(P,F), [\] @ nenv @[\]) \ ?\" + with \?\\formula\ \p\P\ in_M' + \arity(?\) \ length([\] @ nenv @ [\])\ + show "M, [p, P, leq, \,\] @ nenv @ [\] \ forces(?\)" + using definition_of_forcing[where \="\\ 0 \ (1 +\<^sub>\ length(env)) \ \ \ \", + THEN iffD2] by auto + qed + finally + show "(M, [\,p,u]@?Pl1@[\]@nenv \ ?new_form) \ (\F. M_generic(F) \ p \ F \ + M[F], map(val(P,F), [\] @ nenv @ [\]) \ ?\)" + by simp + qed + with Eq1 + have "(M, [u] @ ?Pl1 @ [\] @ nenv \ ?\) \ + (\\\M. \p\P. u =\\,p\ \ + (\F. M_generic(F) \ p \ F \ M[F], map(val(P,F), [\] @ nenv @ [\]) \ ?\))" + by auto + } + then + have Equivalence: "u\ domain(\) \ P \ u \ M \ + (M, [u] @ ?Pl1 @ [\] @ nenv \ ?\) \ + (\\\M. \p\P. u =\\,p\ \ + (\F. M_generic(F) \ p \ F \ M[F], map(val(P,F), [\] @ nenv @[\]) \ ?\))" + for u + by simp + moreover from \env = _\ \\\M\ \nenv\list(M)\ + have map_nenv:"map(val(P,G), nenv@[\]) = env @ [val(P,G,\)]" + using map_app_distrib append1_eq_iff by auto + ultimately + have aux:"(\\\M. \p\P. u =\\,p\ \ (p\G \ M[G], [val(P,G,\)] @ env @ [val(P,G,\)] \ ?\))" + (is "(\\\M. \p\P. _ ( _ \ _, ?vals(\) \ _))") + if "u \ domain(\) \ P" "u \ M" "M, [u]@ ?Pl1 @[\] @ nenv \ ?\" for u + using Equivalence[THEN iffD1, OF that] generic by force + moreover + have "\\M \ val(P,G,\)\M[G]" for \ + using GenExt_def by auto + moreover + have "\\ M \ [val(P,G, \)] @ env @ [val(P,G, \)] \ list(M[G])" for \ + proof - + from \\\M\ + have "val(P,G,\)\ M[G]" using GenExtI by simp + moreover + assume "\ \ M" + moreover + note \env \ list(M[G])\ + ultimately + show ?thesis + using GenExtI by simp + qed + ultimately + have "(\\\M. \p\P. u=\\,p\ \ (p\G \ val(P,G,\)\nth(1 +\<^sub>\ length(env),[val(P,G, \)] @ env @ [val(P,G, \)]) + \ (M[G], ?vals(\) \ \)))" + if "u \ domain(\) \ P" "u \ M" "M, [u] @ ?Pl1 @[\] @ nenv \ ?\" for u + using aux[OF that] by simp + moreover from \env \ _\ \\\M\ + have nth:"nth(1 +\<^sub>\ length(env),[val(P,G, \)] @ env @ [val(P,G, \)]) = val(P,G,\)" + if "\\M" for \ + using nth_concat[of "val(P,G,\)" "val(P,G,\)" "M[G]"] using that GenExtI by simp + ultimately + have "(\\\M. \p\P. u=\\,p\ \ (p\G \ val(P,G,\)\val(P,G,\) \ (M[G],?vals(\) \ \)))" + if "u \ domain(\) \ P" "u \ M" "M, [u] @ ?Pl1 @[\] @ nenv \ ?\" for u + using that \\\M\ \env \ _\ by simp + with \domain(\)\P\M\ + have "\u\domain(\)\P . (M, [u] @ ?Pl1 @[\] @ nenv \ ?\) \ (\\\M. \p\P. u =\\,p\ \ + (p \ G \ val(P,G, \)\val(P,G, \) \ (M[G],?vals(\) \ \)))" + by (simp add:transitivity) + then + have "{u\domain(\)\P . (M,[u] @ ?Pl1 @[\] @ nenv \ ?\) } \ + {u\domain(\)\P . \\\M. \p\P. u =\\,p\ \ + (p \ G \ val(P,G, \)\val(P,G, \) \ (M[G], ?vals(\) \ \))}" + (is "?n\?m") + by auto + with val_mono + have first_incl: "val(P,G,?n) \ val(P,G,?m)" + by simp + note \val(P,G,\) = c\ (* from the assumptions *) + with \?\\formula\ \arity(?\) \ _\ in_M \nenv \ _\ \env \ _\ \length(nenv) = _\ + have "?n\M" + using separation_ax leI separation_iff by auto + from generic + have "filter(G)" "G\P" + unfolding M_generic_def filter_def by simp_all + from \val(P,G,\) = c\ + have "val(P,G,?m) = + {z . t\domain(\) , (\q\P . + (\\\M. \p\P. \t,q\ = \\, p\ \ + (p \ G \ val(P,G, \) \ c \ (M[G], [val(P,G, \)] @ env @ [c] \ \)) \ q \ G)) \ + z=val(P,G,t)}" + using val_of_name by auto + also + have "... = {z . t\domain(\) , (\q\P. + val(P,G, t) \ c \ (M[G], [val(P,G, t)] @ env @ [c] \ \) \ q \ G) \ z=val(P,G,t)}" + proof - + have "t\M \ + (\q\P. (\\\M. \p\P. \t,q\ = \\, p\ \ + (p \ G \ val(P,G, \) \ c \ (M[G], [val(P,G, \)] @ env @ [c] \ \)) \ q \ G)) + \ + (\q\P. val(P,G, t) \ c \ ( M[G], [val(P,G, t)]@env@[c]\ \ ) \ q \ G)" for t + by auto + then show ?thesis using \domain(\)\M\ by (auto simp add:transitivity) + qed + also + have "... = {x\c . \q\P. x \ c \ (M[G], [x] @ env @ [c] \ \) \ q \ G}" + proof + show "... \ {x\c . \q\P. x \ c \ (M[G], [x] @ env @ [c] \ \) \ q \ G}" + by auto + next + (* Now we show the other inclusion: + {x .. x\c , \q\P. x \ c \ (M[G], [x, w, c] \ \) \ q \ G} + \ + {val(P,G,t)..t\domain(\),\q\P.val(P,G,t)\c\(M[G], [val(P,G,t),w] \ \)\q\G} + *) + { + fix x + assume "x\{x\c . \q\P. x \ c \ (M[G], [x] @ env @ [c] \ \) \ q \ G}" + then + have "\q\P. x \ c \ (M[G], [x] @ env @ [c] \ \) \ q \ G" + by simp + with \val(P,G,\) = c\ + have "\q\P. \t\domain(\). val(P,G,t) =x \ (M[G], [val(P,G,t)] @ env @ [c] \ \) \ q \ G" + using elem_of_val by auto + } + then + show " {x\c . \q\P. x \ c \ (M[G], [x] @ env @ [c] \ \) \ q \ G} \ ..." + by force + qed + also + have " ... = {x\c. (M[G], [x] @ env @ [c] \ \)}" + using \G\P\ G_nonempty by force + finally + have val_m: "val(P,G,?m) = {x\c. (M[G], [x] @ env @ [c] \ \)}" by simp + have "val(P,G,?m) \ val(P,G,?n)" + proof + fix x + assume "x \ val(P,G,?m)" + with val_m + have Eq4: "x \ {x\c. (M[G], [x] @ env @ [c] \ \)}" by simp + with \val(P,G,\) = c\ + have "x \ val(P,G,\)" by simp + then + have "\\. \q\G. \\,q\\\ \ val(P,G,\) =x" + using elem_of_val_pair by auto + then obtain \ q where + "\\,q\\\" "q\G" "val(P,G,\)=x" by auto + from \\\,q\\\\ + have "\\M" + using domain_trans[OF trans_M \\\_\] by auto + with \\\M\ \nenv \ _\ \env = _\ + have "[val(P,G,\), val(P,G,\)] @ env \list(M[G])" + using GenExt_def by auto + with Eq4 \val(P,G,\)=x\ \val(P,G,\) = c\ \x \ val(P,G,\)\ nth \\\M\ + have Eq5: "M[G], [val(P,G,\)] @ env @[val(P,G,\)] \ And(Member(0,1 +\<^sub>\ length(env)),\)" + by auto + (* Recall ?\ = And(Member(0,1 +\<^sub>\ length(env)),\) *) + with \\\M\ \\\M\ Eq5 \M_generic(G)\ \\\formula\ \nenv \ _ \ \env = _ \ map_nenv + \arity(?\) \ length([\] @ nenv @ [\])\ + have "(\r\G. M, [r,P,leq,\,\] @ nenv @[\] \ forces(?\))" + using truth_lemma[of "\\ 0 \ (1 +\<^sub>\ length(env)) \ \ \ \"] + by auto + then obtain r where (* I can't "obtain" this directly *) + "r\G" "M, [r,P,leq,\,\] @ nenv @ [\] \ forces(?\)" by auto + with \filter(G)\ and \q\G\ obtain p where + "p\G" "p\q" "p\r" + unfolding filter_def compat_in_def by force + with \r\G\ \q\G\ \G\P\ + have "p\P" "r\P" "q\P" "p\M" + using P_in_M by (auto simp add:transitivity) + with \\\formula\ \\\M\ \\\M\ \p\r\ \nenv \ _\ \arity(?\) \ length([\] @ nenv @ [\])\ + \M, [r,P,leq,\,\] @ nenv @ [\] \ forces(?\)\ \env\_\ + have "M, [p,P,leq,\,\] @ nenv @ [\] \ forces(?\)" + using strengthening_lemma + by simp + with \p\P\ \\\formula\ \\\M\ \\\M\ \nenv \ _\ \arity(?\) \ length([\] @ nenv @ [\])\ + have "\F. M_generic(F) \ p \ F \ + M[F], map(val(P,F), [\] @ nenv @[\]) \ ?\" + using definition_of_forcing[where \="\\ 0 \ (1 +\<^sub>\ length(env)) \ \ \ \"] + by simp + with \p\P\ \\\M\ + have Eq6: "\\'\M. \p'\P. \\,p\ = <\',p'> \ (\F. M_generic(F) \ p' \ F \ + M[F], map(val(P,F), [\'] @ nenv @ [\]) \ ?\)" by auto + from \\\M\ \\\,q\\\\ + have "\\,q\ \ M" by (simp add:transitivity) + from \\\,q\\\\ \\\M\ \p\P\ \p\M\ + have "\\,p\\M" "\\,p\\domain(\)\P" + using pair_in_M_iff by auto + with \\\M\ Eq6 \p\P\ + have "M, [\\,p\] @ ?Pl1 @ [\] @ nenv \ ?\" + using Equivalence by auto + with \\\,p\\domain(\)\P\ + have "\\,p\\?n" by simp + with \p\G\ \p\P\ + have "val(P,G,\)\val(P,G,?n)" + using val_of_elem[of \ p] by simp + with \val(P,G,\)=x\ + show "x\val(P,G,?n)" by simp + qed (* proof of "val(P,G,?m) \ val(P,G,?n)" *) + with val_m first_incl + have "val(P,G,?n) = {x\c. (M[G], [x] @ env @ [c] \ \)}" by auto + also + have " ... = {x\c. (M[G], [x] @ env \ \)}" + proof - + { + fix x + assume "x\c" + moreover from assms + have "c\M[G]" + unfolding GenExt_def by auto + moreover from this and \x\c\ + have "x\M[G]" + using transitivity_MG + by simp + ultimately + have "(M[G], ([x] @ env) @[c] \ \) \ (M[G], [x] @ env \ \)" + using phi \env \ _\ by (rule_tac arity_sats_iff, simp_all) (* Enhance this *) + } + then show ?thesis by auto + qed + finally + show "{x\c. (M[G], [x] @ env \ \)}\ M[G]" + using \?n\M\ GenExt_def by force +qed + +theorem separation_in_MG: + assumes + "\\formula" and "arity(\) \ 1 +\<^sub>\ length(env)" and "env\list(M[G])" + shows + "separation(##M[G],\x. (M[G], [x] @ env \ \))" +proof - + { + fix c + assume "c\M[G]" + moreover from \env \ _\ + obtain nenv where "nenv\list(M)" + "env = map(val(P,G),nenv)" "length(env) = length(nenv)" + using GenExt_def map_val[of env] by auto + moreover note \\ \ _\ \arity(\) \ _\ \env \ _\ + ultimately + have Eq1: "{x\c. (M[G], [x] @ env \ \)} \ M[G]" + using Collect_sats_in_MG by auto + } + then + show ?thesis + using separation_iff rev_bexI unfolding is_Collect_def by force +qed + +end \ \\<^locale>\G_generic1\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Separation_Instances.thy b/thys/Independence_CH/Separation_Instances.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Separation_Instances.thy @@ -0,0 +1,234 @@ +subsection\More Instances of Separation\ + +theory Separation_Instances + imports + Names +begin + +text\The following instances are mostly the same repetitive task; and we just +copied and pasted, tweaking some lemmas if needed (for example, we might have +needed to use some closedness results). +\ + +definition radd_body :: "[i,i,i] \ o" where + "radd_body(R,S) \ \z. (\x y. z = \Inl(x), Inr(y)\) \ + (\x' x. z = \Inl(x'), Inl(x)\ \ \x', x\ \ R) \ + (\y' y. z = \Inr(y'), Inr(y)\ \ \y', y\ \ S)" + +relativize functional "radd_body" "radd_body_rel" +relationalize "radd_body_rel" "is_radd_body" + +synthesize "is_radd_body" from_definition +arity_theorem for "is_radd_body_fm" + +lemma (in M_ZF1_trans) radd_body_abs: + assumes "(##M)(R)" "(##M)(S)" "(##M)(x)" + shows "is_radd_body(##M,R,S,x) \ radd_body(R,S,x)" + using assms pair_in_M_iff Inl_in_M_iff Inr_in_M_iff + unfolding radd_body_def is_radd_body_def + by (auto) + +lemma (in M_ZF1_trans) separation_radd_body: + "(##M)(R) \ (##M)(S) \ separation + (##M, \z. (\x y. z = \Inl(x), Inr(y)\) \ + (\x' x. z = \Inl(x'), Inl(x)\ \ \x', x\ \ R) \ + (\y' y. z = \Inr(y'), Inr(y)\ \ \y', y\ \ S))" + using separation_in_ctm[where \="is_radd_body_fm(1,2,0)" and env="[R,S]"] + is_radd_body_def arity_is_radd_body_fm ord_simp_union is_radd_body_fm_type radd_body_abs + unfolding radd_body_def + by simp + +definition rmult_body :: "[i,i,i] \ o" where + "rmult_body(b,d) \ \z. \x' y' x y. z = \\x', y'\, x, y\ \ (\x', x\ \ +b \ x' = x \ \y', y\ \ d)" + +relativize functional "rmult_body" "rmult_body_rel" +relationalize "rmult_body_rel" "is_rmult_body" + +synthesize "is_rmult_body" from_definition +arity_theorem for "is_rmult_body_fm" + +lemma (in M_ZF1_trans) rmult_body_abs: + assumes "(##M)(b)" "(##M)(d)" "(##M)(x)" + shows "is_rmult_body(##M,b,d,x) \ rmult_body(b,d,x)" + using assms pair_in_M_iff apply_closed + unfolding rmult_body_def is_rmult_body_def + by (auto) + +lemma (in M_ZF1_trans) separation_rmult_body: + "(##M)(b) \ (##M)(d) \ separation + (##M, \z. \x' y' x y. z = \\x', y'\, x, y\ \ (\x', x\ \ b \ x' = x \ \y', y\ \ d))" + using separation_in_ctm[where \="is_rmult_body_fm(1,2,0)" and env="[b,d]"] + is_rmult_body_def arity_is_rmult_body_fm ord_simp_union is_rmult_body_fm_type rmult_body_abs + unfolding rmult_body_def + by simp + +lemma (in M_replacement) separation_well_ord: + "(M)(f) \ (M)(r) \ (M)(A) \ separation + (M, \x. x \ A \ (\y[M]. \p[M]. is_apply(M, f, x, y) \ pair(M, y, x, p) \ p \ r))" + using separation_imp separation_in lam_replacement_identity lam_replacement_constant + lam_replacement_apply[of f] lam_replacement_Pair[THEN [5] lam_replacement_hcomp2] + by simp + +definition is_obase_body :: "[i\o,i,i,i] \ o" where + "is_obase_body(N,A,r,x) \ x \ A \ + \ (\y[N]. + \g[N]. + ordinal(N, y) \ + (\my[N]. + \pxr[N]. + membership(N, y, my) \ + pred_set(N, A, x, r, pxr) \ + order_isomorphism(N, pxr, r, y, my, g)))" + +synthesize "is_obase_body" from_definition +arity_theorem for "is_obase_body_fm" + +lemma (in M_ZF1_trans) separation_is_obase: + "(##M)(f) \ (##M)(r) \ (##M)(A) \ separation + (##M, \x. x \ A \ + \ (\y[##M]. + \g[##M]. + ordinal(##M, y) \ + (\my[##M]. + \pxr[##M]. + membership(##M, y, my) \ + pred_set(##M, A, x, r, pxr) \ + order_isomorphism(##M, pxr, r, y, my, g))))" + using separation_in_ctm[where \="is_obase_body_fm(1,2,0)" and env="[A,r]"] + is_obase_body_def arity_is_obase_body_fm ord_simp_union is_obase_body_fm_type + by simp + +definition is_obase_equals :: "[i\o,i,i,i] \ o" where + "is_obase_equals(N,A,r,a) \ \x[N]. + \g[N]. + \mx[N]. + \par[N]. + ordinal(N, x) \ + membership(N, x, mx) \ + pred_set(N, A, a, r, par) \ order_isomorphism(N, par, r, x, mx, g)" + +synthesize "is_obase_equals" from_definition +arity_theorem for "is_obase_equals_fm" + +lemma (in M_ZF1_trans) separation_obase_equals: + "(##M)(f) \ (##M)(r) \ (##M)(A) \ separation + (##M, \a. \x[##M]. + \g[##M]. + \mx[##M]. + \par[##M]. + ordinal(##M, x) \ + membership(##M, x, mx) \ + pred_set(##M, A, a, r, par) \ order_isomorphism(##M, par, r, x, mx, g))" + using separation_in_ctm[where \="is_obase_equals_fm(1,2,0)" and env="[A,r]"] + is_obase_equals_def arity_is_obase_equals_fm ord_simp_union is_obase_equals_fm_type + by simp + +synthesize "PiP_rel" from_definition assuming "nonempty" +arity_theorem for "PiP_rel_fm" + +lemma (in M_ZF1_trans) separation_PiP_rel: + "(##M)(A) \ separation(##M, PiP_rel(##M,A))" + using separation_in_ctm[where env="[A]" and \="PiP_rel_fm(1,0)"] + nonempty PiP_rel_iff_sats[symmetric] arity_PiP_rel_fm PiP_rel_fm_type + by(simp_all add: ord_simp_union) + +synthesize "injP_rel" from_definition assuming "nonempty" +arity_theorem for "injP_rel_fm" + +lemma (in M_ZF1_trans) separation_injP_rel: + "(##M)(A) \ separation(##M, injP_rel(##M,A))" + using separation_in_ctm[where env="[A]" and \="injP_rel_fm(1,0)"] + nonempty injP_rel_iff_sats[symmetric] arity_injP_rel_fm injP_rel_fm_type + by(simp_all add: ord_simp_union) + +synthesize "surjP_rel" from_definition assuming "nonempty" +arity_theorem for "surjP_rel_fm" + +lemma (in M_ZF1_trans) separation_surjP_rel: + "(##M)(A) \ (##M)(B) \ separation(##M, surjP_rel(##M,A,B))" + using separation_in_ctm[where env="[A,B]" and \="surjP_rel_fm(1,2,0)"] + nonempty surjP_rel_iff_sats[symmetric] arity_surjP_rel_fm surjP_rel_fm_type + by(simp_all add: ord_simp_union) + +synthesize "cons_like_rel" from_definition assuming "nonempty" +arity_theorem for "cons_like_rel_fm" + +lemma (in M_ZF1_trans) separation_cons_like_rel: + "separation(##M, cons_like_rel(##M))" + using separation_in_ctm[where env="[]" and \="cons_like_rel_fm(0)"] + nonempty cons_like_rel_iff_sats[symmetric] arity_cons_like_rel_fm cons_like_rel_fm_type + by simp + +lemma (in M_ZF1_trans) separation_is_function: + "separation(##M, is_function(##M))" + using separation_in_ctm[where env="[]" and \="function_fm(0)"] arity_function_fm + by simp + +(* Instances in M_replacement*) + +definition fstsnd_in_sndsnd :: "[i] \ o" where + "fstsnd_in_sndsnd \ \x. fst(snd(x)) \ snd(snd(x))" +relativize "fstsnd_in_sndsnd" "is_fstsnd_in_sndsnd" +synthesize "is_fstsnd_in_sndsnd" from_definition assuming "nonempty" +arity_theorem for "is_fstsnd_in_sndsnd_fm" + +lemma (in M_ZF1_trans) fstsnd_in_sndsnd_abs: + assumes "(##M)(x)" + shows "is_fstsnd_in_sndsnd(##M,x) \ fstsnd_in_sndsnd(x)" + using assms pair_in_M_iff fst_abs snd_abs fst_snd_closed + unfolding fstsnd_in_sndsnd_def is_fstsnd_in_sndsnd_def + by auto + +lemma (in M_ZF1_trans) separation_fstsnd_in_sndsnd: + "separation(##M, \x. fst(snd(x)) \ snd(snd(x)))" + using separation_in_ctm[where env="[]" and \="is_fstsnd_in_sndsnd_fm(0)" and Q=fstsnd_in_sndsnd] + nonempty fstsnd_in_sndsnd_abs arity_is_fstsnd_in_sndsnd_fm + unfolding fstsnd_in_sndsnd_def + by simp + +definition sndfst_eq_fstsnd :: "[i] \ o" where + "sndfst_eq_fstsnd \ \x. snd(fst(x)) = fst(snd(x))" +relativize "sndfst_eq_fstsnd" "is_sndfst_eq_fstsnd" +synthesize "is_sndfst_eq_fstsnd" from_definition assuming "nonempty" +arity_theorem for "is_sndfst_eq_fstsnd_fm" + +lemma (in M_ZF1_trans) sndfst_eq_fstsnd_abs: + assumes "(##M)(x)" + shows "is_sndfst_eq_fstsnd(##M,x) \ sndfst_eq_fstsnd(x)" + using assms pair_in_M_iff fst_abs snd_abs fst_snd_closed + unfolding sndfst_eq_fstsnd_def is_sndfst_eq_fstsnd_def + by auto + +lemma (in M_ZF1_trans) separation_sndfst_eq_fstsnd: + "separation(##M, \x. snd(fst(x)) = fst(snd(x)))" + using separation_in_ctm[where env="[]" and \="is_sndfst_eq_fstsnd_fm(0)" and Q=sndfst_eq_fstsnd] + nonempty sndfst_eq_fstsnd_abs arity_is_sndfst_eq_fstsnd_fm + unfolding sndfst_eq_fstsnd_def + by simp + +(* "M(G) \ M(Q) \ separation(M, \p. \x\G. x \ snd(p) \ (\s\fst(p). \s, x\ \ Q))" *) +definition insnd_ballPair :: "[i,i,i] \ o" where + "insnd_ballPair(B,A) \ \p. \x\B. x \ snd(p) \ (\s\fst(p). \s, x\ \ A)" + +relativize "insnd_ballPair" "is_insnd_ballPair" +synthesize "is_insnd_ballPair" from_definition assuming "nonempty" +arity_theorem for "is_insnd_ballPair_fm" + +lemma (in M_ZF1_trans) insnd_ballPair_abs: + assumes "(##M)(B)" "(##M)(A)" "(##M)(x)" + shows "is_insnd_ballPair(##M,B,A,x) \ insnd_ballPair(B,A,x)" + using assms pair_in_M_iff fst_abs snd_abs fst_snd_closed + transM[of _ B] transM[of _ "snd(x)"] transM[of _ "fst(x)"] + unfolding insnd_ballPair_def is_insnd_ballPair_def + by (auto) + +lemma (in M_ZF1_trans) separation_insnd_ballPair: + "(##M)(B) \ (##M)(A) \ separation(##M, \p. \x\B. x \ snd(p) \ (\s\fst(p). \s, x\ \ A))" + using insnd_ballPair_abs nonempty + separation_in_ctm[where \="is_insnd_ballPair_fm(2,1,0)" and env="[A,B]"] + arity_is_insnd_ballPair_fm ord_simp_union is_insnd_ballPair_fm_type + unfolding insnd_ballPair_def + by simp + +end \ No newline at end of file diff --git a/thys/Independence_CH/Separation_Rename.thy b/thys/Independence_CH/Separation_Rename.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Separation_Rename.thy @@ -0,0 +1,519 @@ +section\Auxiliary renamings for Separation\ +theory Separation_Rename + imports + Interface +begin + +lemmas apply_fun = apply_iff[THEN iffD1] + +lemma nth_concat : "[p,t] \ list(A) \ env\ list(A) \ nth(1 +\<^sub>\ length(env),[p]@ env @ [t]) = t" + by(auto simp add:nth_append) + +lemma nth_concat2 : "env\ list(A) \ nth(length(env),env @ [p,t]) = p" + by(auto simp add:nth_append) + +lemma nth_concat3 : "env\ list(A) \ u = nth(succ(length(env)), env @ [pi, u])" + by(auto simp add:nth_append) + +definition + sep_var :: "i \ i" where + "sep_var(n) \ {\0,1\,\1,3\,\2,4\,\3,5\,\4,0\,\5+\<^sub>\n,6\,\6+\<^sub>\n,2\}" + +definition + sep_env :: "i \ i" where + "sep_env(n) \ \ i \ (5+\<^sub>\n)-5 . i+\<^sub>\2" + +definition weak :: "[i, i] \ i" where + "weak(n,m) \ {i+\<^sub>\m . i \ n}" + +lemma weakD : + assumes "n \ nat" "k\nat" "x \ weak(n,k)" + shows "\ i \ n . x = i+\<^sub>\k" + using assms unfolding weak_def by blast + +lemma weak_equal : + assumes "n\nat" "m\nat" + shows "weak(n,m) = (m+\<^sub>\n) - m" +proof - + have "weak(n,m)\(m+\<^sub>\n)-m" + proof(intro subsetI) + fix x + assume "x\weak(n,m)" + with assms + obtain i where + "i\n" "x=i+\<^sub>\m" + using weakD by blast + then + have "m\i+\<^sub>\m" "im\nat\ \n\nat\ ltI[OF \i\n\] by simp_all + then + have "\i+\<^sub>\mn\nat\ \i\n\] \m\nat\ by simp + with \x=i+\<^sub>\m\ + have "x\m" + using ltI \m\nat\ by auto + moreover + from assms \x=i+\<^sub>\m\ \i + have "x\n" + using add_lt_mono1[OF \i \n\nat\] by simp + ultimately + show "x\(m+\<^sub>\n)-m" + using ltD DiffI by simp + qed + moreover + have "(m+\<^sub>\n)-m\weak(n,m)" + proof (intro subsetI) + fix x + assume "x\(m+\<^sub>\n)-m" + then + have "x\m+\<^sub>\n" "x\m" + using DiffD1[of x "n+\<^sub>\m" m] DiffD2[of x "n+\<^sub>\m" m] by simp_all + then + have "x\n" "x\nat" + using ltI in_n_in_nat[OF add_type[of m n]] by simp_all + then + obtain i where + "m+\<^sub>\n = succ(x+\<^sub>\i)" + using less_iff_succ_add[OF \x\nat\,of "m+\<^sub>\n"] add_type by auto + then + have "x+\<^sub>\i\n" using succ_le_iff by simp + with \x\m\ + have "\xm\nat\ \x\nat\ + have "m\x" using not_lt_iff_le by simp + with \x\n\ \n\nat\ + have "x-\<^sub>\m\n-\<^sub>\m" + using diff_mono[OF \x\nat\ _ \m\nat\] by simp + have "m+\<^sub>\n-\<^sub>\m = n" using diff_cancel2 \m\nat\ \n\nat\ by simp + with \x-\<^sub>\m\n-\<^sub>\m\ \x\nat\ + have "x-\<^sub>\m \ n" "x=x-\<^sub>\m+\<^sub>\m" + using ltD add_diff_inverse2[OF \m\x\] by simp_all + then + show "x\weak(n,m)" + unfolding weak_def by auto + qed + ultimately + show ?thesis by auto +qed + +lemma weak_zero: + shows "weak(0,n) = 0" + unfolding weak_def by simp + +lemma weakening_diff : + assumes "n \ nat" + shows "weak(n,7) - weak(n,5) \ {5+\<^sub>\n, 6+\<^sub>\n}" + unfolding weak_def using assms +proof(auto) + { + fix i + assume "i\n" "succ(succ(natify(i)))\n" "\w\n. succ(succ(natify(i))) \ natify(w)" + then + have "in\nat\ by simp + from \n\nat\ \i\n\ \succ(succ(natify(i)))\n\ + have "i\nat" "succ(succ(i))\n" using in_n_in_nat by simp_all + from \i + have "succ(i)\n" using succ_leI by simp + with \n\nat\ + consider (a) "succ(i) = n" | (b) "succ(i) < n" + using leD by auto + then have "succ(i) = n" + proof cases + case a + then show ?thesis . + next + case b + then + have "succ(succ(i))\n" using succ_leI by simp + with \n\nat\ + consider (a) "succ(succ(i)) = n" | (b) "succ(succ(i)) < n" + using leD by auto + then have "succ(i) = n" + proof cases + case a + with \succ(succ(i))\n\ show ?thesis by blast + next + case b + then + have "succ(succ(i))\n" using ltD by simp + with \i\nat\ + have "succ(succ(natify(i))) \ natify(succ(succ(i)))" + using \\w\n. succ(succ(natify(i))) \ natify(w)\ by auto + then + have "False" using \i\nat\ by auto + then show ?thesis by blast + qed + then show ?thesis . + qed + with \i\nat\ have "succ(natify(i)) = n" by simp + } + then + show "n \ nat \ + succ(succ(natify(y))) \ n \ + \x\n. succ(succ(natify(y))) \ natify(x) \ + y \ n \ succ(natify(y)) = n" for y + by blast +qed + +lemma in_add_del : + assumes "x\j+\<^sub>\n" "n\nat" "j\nat" + shows "x < j \ x \ weak(n,j)" +proof (cases "xnat" "j+\<^sub>\n\nat" + using in_n_in_nat[OF _ \x\j+\<^sub>\n\] assms by simp_all + then + have "j \ x" "x < j+\<^sub>\n" + using not_lt_iff_le False \j\nat\ \n\nat\ ltI[OF \x\j+\<^sub>\n\] by auto + then + have "x-\<^sub>\j < (j +\<^sub>\ n) -\<^sub>\ j" "x = j +\<^sub>\ (x -\<^sub>\j)" + using diff_mono \x\nat\ \j+\<^sub>\n\nat\ \j\nat\ \n\nat\ + add_diff_inverse[OF \j\x\] by simp_all + then + have "x-\<^sub>\j < n" "x = (x -\<^sub>\j ) +\<^sub>\ j" + using diff_add_inverse \n\nat\ add_commute by simp_all + then + have "x-\<^sub>\j \n" using ltD by simp + then + have "x \ weak(n,j)" + unfolding weak_def + using \x= (x-\<^sub>\j) +\<^sub>\j\ RepFunI[OF \x-\<^sub>\j\n\] add_commute by force + then show ?thesis .. +qed + + +lemma sep_env_action: + assumes + "[t,p,u,P,leq,o,pi] \ list(M)" + "env \ list(M)" + shows "\ i . i \ weak(length(env),5) \ + nth(sep_env(length(env))`i,[t,p,u,P,leq,o,pi]@env) = nth(i,[p,P,leq,o,t] @ env @ [pi,u])" +proof - + from assms + have A: "5+\<^sub>\length(env)\nat" "[p, P, leq, o, t] \list(M)" + by simp_all + let ?f="sep_env(length(env))" + have EQ: "weak(length(env),5) = 5+\<^sub>\length(env) - 5" + using weak_equal length_type[OF \env\list(M)\] by simp + let ?tgt="[t,p,u,P,leq,o,pi]@env" + let ?src="[p,P,leq,o,t] @ env @ [pi,u]" + have "nth(?f`i,[t,p,u,P,leq,o,pi]@env) = nth(i,[p,P,leq,o,t] @ env @ [pi,u])" + if "i \ (5+\<^sub>\length(env)-5)" for i + proof - + from that + have 2: "i \ 5+\<^sub>\length(env)" "i \ 5" "i \ nat" "i-\<^sub>\5\nat" "i+\<^sub>\2\nat" + using in_n_in_nat[OF \5+\<^sub>\length(env)\nat\] by simp_all + then + have 3: "\ i < 5" using ltD by force + then + have "5 \ i" "2 \ 5" + using not_lt_iff_le \i\nat\ by simp_all + then have "2 \ i" using le_trans[OF \2\5\] by simp + from A \i \ 5+\<^sub>\length(env)\ + have "i < 5+\<^sub>\length(env)" using ltI by simp + with \i\nat\ \2\i\ A + have C:"i+\<^sub>\2 < 7+\<^sub>\length(env)" by simp + with that + have B: "?f`i = i+\<^sub>\2" unfolding sep_env_def by simp + from 3 assms(1) \i\nat\ + have "\ i+\<^sub>\2 < 7" using not_lt_iff_le add_le_mono by simp + from \i < 5+\<^sub>\length(env)\ 3 \i\nat\ + have "i-\<^sub>\5 < 5+\<^sub>\length(env) -\<^sub>\ 5" + using diff_mono[of i "5+\<^sub>\length(env)" 5,OF _ _ _ \i < 5+\<^sub>\length(env)\] + not_lt_iff_le[THEN iffD1] by force + with assms(2) + have "i-\<^sub>\5 < length(env)" using diff_add_inverse length_type by simp + have "nth(i,?src) =nth(i-\<^sub>\5,env@[pi,u])" + using nth_append[OF A(2) \i\nat\] 3 by simp + also + have "... = nth(i-\<^sub>\5, env)" + using nth_append[OF \env \list(M)\ \i-\<^sub>\5\nat\] \i-\<^sub>\5 < length(env)\ by simp + also + have "... = nth(i+\<^sub>\2, ?tgt)" + using nth_append[OF assms(1) \i+\<^sub>\2\nat\] \\ i+\<^sub>\2 <7\ by simp + ultimately + have "nth(i,?src) = nth(?f`i,?tgt)" + using B by simp + then show ?thesis using that by simp + qed + then show ?thesis using EQ by force +qed + +lemma sep_env_type : + assumes "n \ nat" + shows "sep_env(n) : (5+\<^sub>\n)-5 \ (7+\<^sub>\n)-7" +proof - + let ?h="sep_env(n)" + from \n\nat\ + have "(5+\<^sub>\n)+\<^sub>\2 = 7+\<^sub>\n" "7+\<^sub>\n\nat" "5+\<^sub>\n\nat" by simp_all + have + D: "sep_env(n)`x \ (7+\<^sub>\n)-7" if "x \ (5+\<^sub>\n)-5" for x + proof - + from \x\5+\<^sub>\n-5\ + have "?h`x = x+\<^sub>\2" "x<5+\<^sub>\n" "x\nat" + unfolding sep_env_def using ltI in_n_in_nat[OF \5+\<^sub>\n\nat\] by simp_all + then + have "x+\<^sub>\2 < 7+\<^sub>\n" by simp + then + have "x+\<^sub>\2 \ 7+\<^sub>\n" using ltD by simp + from \x\5+\<^sub>\n-5\ + have "x\5" by simp + then have "\x<5" using ltD by blast + then have "5\x" using not_lt_iff_le \x\nat\ by simp + then have "7\x+\<^sub>\2" using add_le_mono \x\nat\ by simp + then have "\x+\<^sub>\2<7" using not_lt_iff_le \x\nat\ by simp + then have "x+\<^sub>\2 \ 7" using ltI \x\nat\ by force + with \x+\<^sub>\2 \ 7+\<^sub>\n\ show ?thesis using \?h`x = x+\<^sub>\2\ DiffI by simp + qed + then show ?thesis unfolding sep_env_def using lam_type by simp +qed + +lemma sep_var_fin_type : + assumes "n \ nat" + shows "sep_var(n) : 7+\<^sub>\n -||> 7+\<^sub>\n" + unfolding sep_var_def + using consI ltD emptyI by force + +lemma sep_var_domain : + assumes "n \ nat" + shows "domain(sep_var(n)) = 7+\<^sub>\n - weak(n,5)" +proof - + let ?A="weak(n,5)" + have A:"domain(sep_var(n)) \ (7+\<^sub>\n)" + unfolding sep_var_def + by(auto simp add: le_natE) + have C: "x=5+\<^sub>\n \ x=6+\<^sub>\n \ x \ 4" if "x\domain(sep_var(n))" for x + using that unfolding sep_var_def by auto + have D : "x\7" if "x\7+\<^sub>\n" for x + using that \n\nat\ ltI by simp + have "\ 5+\<^sub>\n < 5+\<^sub>\n" using \n\nat\ lt_irrefl[of _ False] by force + have "\ 6+\<^sub>\n < 5+\<^sub>\n" using \n\nat\ by force + have R: "x < 5+\<^sub>\n" if "x\?A" for x + proof - + from that + obtain i where + "i\i" + unfolding weak_def + using ltI \n\nat\ RepFun_iff by force + with \n\nat\ + have "5+\<^sub>\i < 5+\<^sub>\n" using add_lt_mono2 by simp + with \x=5+\<^sub>\i\ + show "x < 5+\<^sub>\n" by simp + qed + then + have 1:"x\?A" if "\x <5+\<^sub>\n" for x using that by blast + have "5+\<^sub>\n \ ?A" "6+\<^sub>\n\?A" + proof - + show "5+\<^sub>\n \ ?A" using 1 \\5+\<^sub>\n<5+\<^sub>\n\ by blast + with 1 show "6+\<^sub>\n \ ?A" using \\6+\<^sub>\n<5+\<^sub>\n\ by blast + qed + then + have E:"x\?A" if "x\domain(sep_var(n))" for x + unfolding weak_def + using C that by force + then + have F: "domain(sep_var(n)) \ 7+\<^sub>\n - ?A" using A by auto + from assms + have "x<7 \ x\weak(n,7)" if "x\7+\<^sub>\n" for x + using in_add_del[OF \x\7+\<^sub>\n\] by simp + moreover + { + fix x + assume asm:"x\7+\<^sub>\n" "x\?A" "x\weak(n,7)" + then + have "x\domain(sep_var(n))" + proof - + from \n\nat\ + have "weak(n,7)-weak(n,5)\{n+\<^sub>\5,n+\<^sub>\6}" + using weakening_diff by simp + with \x\?A\ asm + have "x\{n+\<^sub>\5,n+\<^sub>\6}" using subsetD DiffI by blast + then + show ?thesis unfolding sep_var_def by simp + qed + } + moreover + { + fix x + assume asm:"x\7+\<^sub>\n" "x\?A" "x<7" + then have "x\domain(sep_var(n))" + proof (cases "2 \ n") + case True + moreover + have "0n\nat\ \2\n\] lt_imp_0_lt by auto + ultimately + have "x<5" + using \x<7\ \x\?A\ \n\nat\ in_n_in_nat + unfolding weak_def + by (clarsimp simp add:not_lt_iff_le, auto simp add:lt_def) + then + show ?thesis unfolding sep_var_def + by (clarsimp simp add:not_lt_iff_le, auto simp add:lt_def) + next + case False + then + show ?thesis + proof (cases "n=0") + case True + then show ?thesis + unfolding sep_var_def using ltD asm \n\nat\ by auto + next + case False + then + have "n < 2" using \n\nat\ not_lt_iff_le \\ 2 \ n\ by force + then + have "\ n <1" using \n\0\ by simp + then + have "n=1" using not_lt_iff_le \n<2\ le_iff by auto + then show ?thesis + using \x\?A\ + unfolding weak_def sep_var_def + using ltD asm \n\nat\ by force + qed + qed + } + ultimately + have "w\domain(sep_var(n))" if "w\ 7+\<^sub>\n - ?A" for w + using that by blast + then + have "7+\<^sub>\n - ?A \ domain(sep_var(n))" by blast + with F + show ?thesis by auto +qed + +lemma sep_var_type : + assumes "n \ nat" + shows "sep_var(n) : (7+\<^sub>\n)-weak(n,5) \ 7+\<^sub>\n" + using FiniteFun_is_fun[OF sep_var_fin_type[OF \n\nat\]] + sep_var_domain[OF \n\nat\] by simp + +lemma sep_var_action : + assumes + "[t,p,u,P,leq,o,pi] \ list(M)" + "env \ list(M)" + shows "\ i . i \ (7+\<^sub>\length(env)) - weak(length(env),5) \ + nth(sep_var(length(env))`i,[t,p,u,P,leq,o,pi]@env) = nth(i,[p,P,leq,o,t] @ env @ [pi,u])" + using assms +proof (subst sep_var_domain[OF length_type[OF \env\list(M)\],symmetric],auto) + fix i y + assume "\i, y\ \ sep_var(length(env))" + with assms + show "nth(sep_var(length(env)) ` i, + Cons(t, Cons(p, Cons(u, Cons(P, Cons(leq, Cons(o, Cons(pi, env)))))))) = + nth(i, Cons(p, Cons(P, Cons(leq, Cons(o, Cons(t, env @ [pi, u]))))))" + using apply_fun[OF sep_var_type] assms + unfolding sep_var_def + using nth_concat2[OF \env\list(M)\] nth_concat3[OF \env\list(M)\,symmetric] + by force +qed + +definition + rensep :: "i \ i" where + "rensep(n) \ union_fun(sep_var(n),sep_env(n),7+\<^sub>\n-weak(n,5),weak(n,5))" + +lemma rensep_aux : + assumes "n\nat" + shows "(7+\<^sub>\n-weak(n,5)) \ weak(n,5) = 7+\<^sub>\n" "7+\<^sub>\n \ ( 7 +\<^sub>\ n - 7) = 7+\<^sub>\n" +proof - + from \n\nat\ + have "weak(n,5) = n+\<^sub>\5-5" + using weak_equal by simp + with \n\nat\ + show "(7+\<^sub>\n-weak(n,5)) \ weak(n,5) = 7+\<^sub>\n" "7+\<^sub>\n \ ( 7 +\<^sub>\ n - 7) = 7+\<^sub>\n" + using Diff_partition le_imp_subset by auto +qed + +lemma rensep_type : + assumes "n\nat" + shows "rensep(n) \ 7+\<^sub>\n \ 7+\<^sub>\n" +proof - + from \n\nat\ + have "rensep(n) \ (7+\<^sub>\n-weak(n,5)) \ weak(n,5) \ 7+\<^sub>\n \ (7+\<^sub>\n - 7)" + unfolding rensep_def + using union_fun_type sep_var_type \n\nat\ sep_env_type weak_equal + by force + then + show ?thesis using rensep_aux \n\nat\ by auto +qed + +lemma rensep_action : + assumes "[t,p,u,P,leq,o,pi] @ env \ list(M)" + shows "\ i . i < 7+\<^sub>\length(env) \ nth(rensep(length(env))`i,[t,p,u,P,leq,o,pi]@env) = nth(i,[p,P,leq,o,t] @ env @ [pi,u])" +proof - + let ?tgt="[t,p,u,P,leq,o,pi]@env" + let ?src="[p,P,leq,o,t] @ env @ [pi,u]" + let ?m="7 +\<^sub>\ length(env) - weak(length(env),5)" + let ?p="weak(length(env),5)" + let ?f="sep_var(length(env))" + let ?g="sep_env(length(env))" + let ?n="length(env)" + from assms + have 1 : "[t,p,u,P,leq,o,pi] \ list(M)" " env \ list(M)" + "?src \ list(M)" "?tgt \ list(M)" + "7+\<^sub>\?n = (7+\<^sub>\?n-weak(?n,5)) \ weak(?n,5)" + " length(?src) = (7+\<^sub>\?n-weak(?n,5)) \ weak(?n,5)" + using Diff_partition le_imp_subset rensep_aux by auto + then + have "nth(i, ?src) = nth(union_fun(?f, ?g, ?m, ?p) ` i, ?tgt)" if "i < 7+\<^sub>\length(env)" for i + proof - + from \i<7+\<^sub>\?n\ + have "i \ (7+\<^sub>\?n-weak(?n,5)) \ weak(?n,5)" + using ltD by simp + then show ?thesis + unfolding rensep_def using + union_fun_action[OF \?src\list(M)\ \?tgt\list(M)\ \length(?src) = (7+\<^sub>\?n-weak(?n,5)) \ weak(?n,5)\ + sep_var_action[OF \[t,p,u,P,leq,o,pi] \ list(M)\ \env\list(M)\] + sep_env_action[OF \[t,p,u,P,leq,o,pi] \ list(M)\ \env\list(M)\] + ] that + by simp + qed + then show ?thesis unfolding rensep_def by simp +qed + +definition sep_ren :: "[i,i] \ i" where + "sep_ren(n,\) \ ren(\)`(7+\<^sub>\n)`(7+\<^sub>\n)`rensep(n)" + +lemma arity_rensep: assumes "\\formula" "env \ list(M)" + "arity(\) \ 7+\<^sub>\length(env)" +shows "arity(sep_ren(length(env),\)) \ 7+\<^sub>\length(env)" + unfolding sep_ren_def + using arity_ren rensep_type assms + by simp + +lemma type_rensep [TC]: + assumes "\\formula" "env\list(M)" + shows "sep_ren(length(env),\) \ formula" + unfolding sep_ren_def + using ren_tc rensep_type assms + by simp + +lemma sepren_action: + assumes "arity(\) \ 7 +\<^sub>\ length(env)" + "[t,p,u,P,leq,o,pi] \ list(M)" + "env\list(M)" + "\\formula" + shows "sats(M, sep_ren(length(env),\),[t,p,u,P,leq,o,pi] @ env) \ sats(M, \,[p,P,leq,o,t] @ env @ [pi,u])" +proof - + from assms + have 1: "[t, p, u, P, leq, o, pi] @ env \ list(M)" + by simp_all + then + have 2: "[p,P,leq,o,t] @ env @ [pi,u] \ list(M)" + using app_type by simp + show ?thesis + unfolding sep_ren_def + using sats_iff_sats_ren[OF \\\formula\ + add_type[of 7 "length(env)"] + add_type[of 7 "length(env)"] + 2 1 + rensep_type[OF length_type[OF \env\list(M)\]] + \arity(\) \ 7 +\<^sub>\ length(env)\] + rensep_action[OF 1,rule_format,symmetric] + by simp +qed + +end \ No newline at end of file diff --git a/thys/Independence_CH/Succession_Poset.thy b/thys/Independence_CH/Succession_Poset.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Succession_Poset.thy @@ -0,0 +1,240 @@ +section\A poset of successions\ + +theory Succession_Poset + imports + Replacement_Instances + Proper_Extension +begin + +text\In this theory we define a separative poset. Its underlying set is the +set of finite binary sequences (that is, with codomain $2={0,1}$); +of course, one can see that set as +the set \<^term>\\-||>2\ or equivalently as the set of partial functions +\<^term>\Fn(\,\,2)\, i.e. the set of partial functions bounded by \<^term>\\\. + +The order relation of the poset is that of being less defined as functions +(cf. \<^term>\Fnlerel(A\<^bsup><\\<^esup>)\), so it could be surprising that we have not used +\<^term>\Fn(\,\,2)\ for the set. The only reason why we keep this alternative +definition is because we can prove \<^term>\A\<^bsup><\\<^esup> \ M\ (and therefore +\<^term>\Fnlerel(A\<^bsup><\\<^esup>) \ M\) using only one instance of replacement.\ + +sublocale M_ZF2_trans \ M_seqspace "##M" + by (unfold_locales, simp add:replacement_omega_funspace) + +definition seq_upd :: "i \ i \ i" where + "seq_upd(f,a) \ \ j \ succ(domain(f)) . if j < domain(f) then f`j else a" + +lemma seq_upd_succ_type : + assumes "n\nat" "f\n\A" "a\A" + shows "seq_upd(f,a)\ succ(n) \ A" +proof - + from assms + have equ: "domain(f) = n" + using domain_of_fun by simp + { + fix j + assume "j\succ(domain(f))" + with equ \n\_\ + have "j\n" + using ltI by auto + with \n\_\ + consider (lt) "j A" + proof cases + case lt + with \f\_\ + show ?thesis + using apply_type ltD[OF lt] by simp + next + case eq + with \a\_\ + show ?thesis + by auto + qed + } + with equ + show ?thesis + unfolding seq_upd_def + using lam_type[of "succ(domain(f))"] + by auto +qed + +lemma seq_upd_type : + assumes "f\A\<^bsup><\\<^esup>" "a\A" + shows "seq_upd(f,a) \ A\<^bsup><\\<^esup>" +proof - + from \f\_\ + obtain y where "y\nat" "f\y\A" + unfolding seqspace_def by blast + with \a\A\ + have "seq_upd(f,a)\succ(y)\A" + using seq_upd_succ_type by simp + with \y\_\ + show ?thesis + unfolding seqspace_def by auto +qed + +lemma seq_upd_apply_domain [simp]: + assumes "f:n\A" "n\nat" + shows "seq_upd(f,a)`n = a" + unfolding seq_upd_def using assms domain_of_fun by auto + +lemma zero_in_seqspace : + shows "0 \ A\<^bsup><\\<^esup>" + unfolding seqspace_def + by force + +definition + seqlerel :: "i \ i" where + "seqlerel(A) \ Fnlerel(A\<^bsup><\\<^esup>)" + +definition + seqle :: "i" where + "seqle \ seqlerel(2)" + +lemma seqleI[intro!]: + "\f,g\ \ 2\<^bsup><\\<^esup>\2\<^bsup><\\<^esup> \ g \ f \ \f,g\ \ seqle" + unfolding seqle_def seqlerel_def seqspace_def Rrel_def Fnlerel_def + by blast + +lemma seqleD[dest!]: + "z \ seqle \ \x y. \x,y\ \ 2\<^bsup><\\<^esup>\2\<^bsup><\\<^esup> \ y \ x \ z = \x,y\" + unfolding Rrel_def seqle_def seqlerel_def Fnlerel_def + by blast + +lemma upd_leI : + assumes "f\2\<^bsup><\\<^esup>" "a\2" + shows "\seq_upd(f,a),f\\seqle" (is "\?f,_\\_") +proof + show " \?f, f\ \ 2\<^bsup><\\<^esup> \ 2\<^bsup><\\<^esup>" + using assms seq_upd_type by auto +next + show "f \ seq_upd(f,a)" + proof + fix x + assume "x \ f" + moreover from \f \ 2\<^bsup><\\<^esup>\ + obtain n where "n\nat" "f : n \ 2" + by blast + moreover from calculation + obtain y where "y\n" "x=\y,f`y\" + using Pi_memberD[of f n "\_ . 2"] + by blast + moreover from \f:n\2\ + have "domain(f) = n" + using domain_of_fun by simp + ultimately + show "x \ seq_upd(f,a)" + unfolding seq_upd_def lam_def + by (auto intro:ltI) + qed +qed + +lemma preorder_on_seqle: "preorder_on(2\<^bsup><\\<^esup>,seqle)" + unfolding preorder_on_def refl_def trans_on_def by blast + +lemma zero_seqle_max: "x\2\<^bsup><\\<^esup> \ \x,0\ \ seqle" + using zero_in_seqspace + by auto + +interpretation sp:forcing_notion "2\<^bsup><\\<^esup>" "seqle" "0" + using preorder_on_seqle zero_seqle_max zero_in_seqspace + by unfold_locales simp_all + +notation sp.Leq (infixl "\s" 50) +notation sp.Incompatible (infixl "\s" 50) +notation sp.GenExt_at_P ("_\<^bsup>s\<^esup>[_]" [71,1]) + +lemma seqspace_separative: + assumes "f\2\<^bsup><\\<^esup>" + shows "seq_upd(f,0) \s seq_upd(f,1)" (is "?f \s ?g") +proof + assume "sp.compat(?f, ?g)" + then + obtain h where "h \ 2\<^bsup><\\<^esup>" "?f \ h" "?g \ h" + by blast + moreover from \f\_\ + obtain y where "y\nat" "f:y\2" + by blast + moreover from this + have "?f: succ(y) \ 2" "?g: succ(y) \ 2" + using seq_upd_succ_type by blast+ + moreover from this + have "\y,?f`y\ \ ?f" "\y,?g`y\ \ ?g" + using apply_Pair by auto + ultimately + have "\y,0\ \ h" "\y,1\ \ h" + by auto + moreover from \h \ 2\<^bsup><\\<^esup>\ + obtain n where "n\nat" "h:n\2" + by blast + ultimately + show "False" + using fun_is_function[of h n "\_. 2"] + unfolding seqspace_def function_def by auto +qed + +definition seqleR_fm :: "i \ i" where + "seqleR_fm(fg) \ Exists(Exists(And(pair_fm(0,1,fg+\<^sub>\2),subset_fm(1,0))))" + +lemma type_seqleR_fm : "fg \ nat \ seqleR_fm(fg) \ formula" + unfolding seqleR_fm_def + by simp + +arity_theorem for "seqleR_fm" + +lemma (in M_ctm1) seqleR_fm_sats : + assumes "fg\nat" "env\list(M)" + shows "(M, env \ seqleR_fm(fg)) \ (\f[##M]. \g[##M]. pair(##M,f,g,nth(fg,env)) \ f \ g)" + unfolding seqleR_fm_def + using assms trans_M sats_subset_fm pair_iff_sats + by auto + +locale M_ctm2 = M_ctm1 + M_ZF2_trans + +locale M_ctm2_AC = M_ctm2 + M_ZFC2_trans + +locale forcing_data2 = forcing_data1 + M_ctm2 + +context M_ctm2 +begin + +lemma seqle_in_M: "seqle \ M" + using arity_seqleR_fm seqleR_fm_sats type_seqleR_fm + cartprod_closed seqspace_closed nat_into_M nat_in_M pair_in_M_iff + unfolding seqle_def seqlerel_def Rrel_def Fnlerel_def + by (rule_tac Collect_in_M[of "seqleR_fm(0)" "[]"],auto) + +subsection\Cohen extension is proper\ + +interpretation ctm_separative "2\<^bsup><\\<^esup>" seqle 0 +proof (unfold_locales) + fix f + let ?q="seq_upd(f,0)" and ?r="seq_upd(f,1)" + assume "f \ 2\<^bsup><\\<^esup>" + then + have "?q \s f \ ?r \s f \ ?q \s ?r" + using upd_leI seqspace_separative by auto + moreover from calculation + have "?q \ 2\<^bsup><\\<^esup>" "?r \ 2\<^bsup><\\<^esup>" + using seq_upd_type[of f 2] by auto + ultimately + show "\q\2\<^bsup><\\<^esup>. \r\2\<^bsup><\\<^esup>. q \s f \ r \s f \ q \s r" + by (rule_tac bexI)+ \ \why the heck auto-tools don't solve this?\ +next + show "2\<^bsup><\\<^esup> \ M" + using nat_into_M seqspace_closed by simp +next + show "seqle \ M" + using seqle_in_M . +qed + +lemma cohen_extension_is_proper: "\G. M_generic(G) \ M \ M\<^bsup>2\<^bsup><\\<^esup>\<^esup>[G]" + using proper_extension generic_filter_existence zero_in_seqspace + by force + +end \ \\<^locale>\M_ctm2\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/Union_Axiom.thy b/thys/Independence_CH/Union_Axiom.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/Union_Axiom.thy @@ -0,0 +1,201 @@ +section\The Axiom of Unions in $M[G]$\ +theory Union_Axiom + imports Names +begin + +definition Union_name_body :: "[i,i,i,i] \ o" where + "Union_name_body(P,leq,\,x) \ \ \ . \q\P . \r\P . + \\,q\ \ \ \ \fst(x),r\ \ \ \ \snd(x),r\ \ leq \ \snd(x),q\ \ leq" + +relativize relational "Union_name_body" "is_Union_name_body" +reldb_add functional "Union_name_body" "is_Union_name_body" +synthesize "is_Union_name_body" from_definition assuming "nonempty" +arity_theorem for "is_Union_name_body_fm" + +definition Union_name :: "[i,i,i] \ i" where + "Union_name(P,leq,\) \ {u \ domain(\(domain(\))) \ P . Union_name_body(P,leq,\,u)}" + +relativize functional "Union_name" "Union_name_rel" +relativize relational "Union_name" "is_Union_name" +synthesize "is_Union_name" from_definition assuming "nonempty" +arity_theorem for "is_Union_name_fm" + +context M_basic +begin + +is_iff_rel for "Union_name" + using transM[OF _ cartprod_closed] domain_closed Union_closed + unfolding is_Union_name_def Union_name_rel_def + by simp + +lemma Union_name_body_iff: + assumes "M(x)" "M(leq)" "M(P)" "M(\)" + shows "is_Union_name_body(M, P, leq, \, x) \ Union_name_body(P, leq, \, x)" +proof - + from \M(\)\ + have "M(\)" if "\\,q\\\" for \ q + using transM[of _ \] transM[of _ "\\,q\"] that + unfolding Pair_def + by auto + then + show ?thesis + using assms transM[OF _ cartprod_closed] pair_abs fst_abs snd_abs + unfolding is_Union_name_body_def Union_name_body_def + by auto +qed + +lemma Union_name_abs : + assumes "M(P)" "M(leq)" "M(\)" + shows "Union_name_rel(M,P,leq,\) = Union_name(P,leq,\)" + using assms transM[OF _ cartprod_closed] Union_name_body_iff + unfolding Union_name_rel_def Union_name_def + by auto + +end \ \\<^locale>\M_basic\\ + +context forcing_data1 +begin + +lemma Union_name_closed : + assumes "\ \ M" + shows "Union_name(P,leq,\) \ M" +proof - + let ?\="is_Union_name_body_fm(3,2,1,0)" + let ?P="\ x . M,[x,\,leq,P] \ ?\" + let ?Q="Union_name_body(P,leq,\)" + from \\\M\ + have "domain(\(domain(\)))\M" (is "?d \ _") + using domain_closed Union_closed by simp + then + have "?d \ P \ M" + using cartprod_closed P_in_M by simp + note types = leq_in_M P_in_M assms \?d\P \ M\ \?d\M\ + moreover + have "arity(?\)\4" + using arity_is_Union_name_body_fm ord_simp_union by simp + moreover from calculation + have "separation(##M,?P)" + using separation_ax by simp + moreover from calculation + have closed:"{ u \ ?d \ P . ?P(u) } \ M" + using separation_iff by force + moreover from calculation + have "?P(x)\ x\M \ ?Q(x)" if "x\?d\P" for x + proof - + note calculation that + moreover from this + have "x = \fst(x),snd(x)\" "x\M" "fst(x) \ M" "snd(x) \ M" + using Pair_fst_snd_eq transitivity[of x "?d\P"] fst_snd_closed + by simp_all + ultimately + show "?P(x) \ x\M \ ?Q(x)" + using types zero_in_M is_Union_name_body_iff_sats Union_name_body_iff + by simp + qed + with \?d \ P \ M\ types + have "Union_name_rel(##M,P,leq,\) \ M" + unfolding Union_name_rel_def + using transitivity[OF _ \?d\P\_\] Collect_cong closed Union_name_body_iff + by simp + ultimately + show ?thesis + using Union_name_abs + by simp +qed + +lemma Union_MG_Eq : + assumes "a \ M[G]" and "a = val(P,G,\)" and "filter(G)" and "\ \ M" + shows "\ a = val(P,G,Union_name(P,leq,\))" +proof (intro equalityI subsetI) + fix x + assume "x \ \ a" + with \a=_\ + have "x\ \ (val(P,G,\))" + by simp + then + obtain i where "i \ val(P,G,\)" "x \ i" + by blast + with \\ \ M\ + obtain \ q where "q \ G" "\\,q\ \ \" "val(P,G,\) = i" "\ \ M" + using elem_of_val_pair domain_trans[OF trans_M] + by blast + moreover from this \x \ i\ + obtain \ r where "r \ G" "\\,r\ \ \" "val(P,G,\) = x" "\ \ M" + using elem_of_val_pair domain_trans[OF trans_M] by blast + moreover from calculation + have "\ \ domain(\(domain(\)))" + by auto + moreover from calculation \filter(G)\ + obtain p where "p \ G" "\p,r\ \ leq" "\p,q\ \ leq" "p \ P" "r \ P" "q \ P" + using low_bound_filter filterD by blast + moreover from this + have "p \ M" "q\M" "r\M" + using P_in_M by (auto dest:transM) + moreover from calculation + have "\\,p\ \ Union_name(P,leq,\)" + unfolding Union_name_def Union_name_body_def + by auto + moreover from this \p\P\ \p\G\ + have "val(P,G,\) \ val(P,G,Union_name(P,leq,\))" + using val_of_elem by simp + ultimately + show "x \ val(P,G,Union_name(P,leq,\))" + by simp +next + fix x + assume "x \ (val(P,G,Union_name(P,leq,\)))" + moreover + note \filter(G)\ \a=val(P,G,\)\ + moreover from calculation + obtain \ p where "p \ G" "\\,p\ \ Union_name(P,leq,\)" "val(P,G,\) = x" + using elem_of_val_pair by blast + moreover from calculation + have "p\P" + using filterD by simp + moreover from calculation + obtain \ q r where "\\,q\ \ \" "\\,r\ \ \" "\p,r\ \ leq" "\p,q\ \ leq" "r\P" "q\P" + unfolding Union_name_def Union_name_body_def + by force + moreover from calculation + have "r \ G" "q \ G" + using filter_leqD by auto + moreover from this \\\,r\ \ \\ \\\,q\\\\ \q\P\ \r\P\ + have "val(P,G,\) \ val(P,G,\)" "val(P,G,\) \ val(P,G,\)" + using val_of_elem by simp+ + moreover from this + have "val(P,G,\) \ \ val(P,G,\)" + by blast + ultimately + show "x \ \ a" + by simp +qed + +lemma union_in_MG : + assumes "filter(G)" + shows "Union_ax(##M[G])" + unfolding Union_ax_def +proof(clarsimp) + fix a + assume "a \ M[G]" + moreover + note \filter(G)\ + moreover from calculation + interpret mgtrans : M_trans "##M[G]" + using transitivity_MG by (unfold_locales; auto) + from calculation + obtain \ where "\ \ M" "a=val(P,G,\)" + using GenExtD by blast + moreover from this + have "val(P,G,Union_name(P,leq,\)) \ M[G]" + using GenExtI Union_name_closed P_in_M leq_in_M by simp + ultimately + show "\z\M[G] . big_union(##M[G],a,z)" + using Union_MG_Eq by auto +qed + +theorem Union_MG : "M_generic(G) \ Union_ax(##M[G])" + by (simp add:M_generic_def union_in_MG) + +end \ \\<^locale>\forcing_data1\\ + +end \ No newline at end of file diff --git a/thys/Independence_CH/ZF_Trans_Interpretations.thy b/thys/Independence_CH/ZF_Trans_Interpretations.thy new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/ZF_Trans_Interpretations.thy @@ -0,0 +1,622 @@ +section\Further instances of axiom-schemes\ + +theory ZF_Trans_Interpretations + imports + Internal_ZFC_Axioms + Succession_Poset +begin + +locale M_ZF3 = M_ZF2 + + assumes + replacement_ax3: + "replacement_assm(M,env,replacement_is_order_body_fm)" + "replacement_assm(M,env,wfrec_replacement_order_pred_fm)" + "replacement_assm(M,env,replacement_is_jump_cardinal_body_fm)" + "replacement_assm(M,env,replacement_is_aleph_fm)" + and + LambdaPair_in_M_replacement3: + "replacement_assm(M,env,LambdaPair_in_M_fm(is_inj_fm(0,1,2),0))" + +definition instances3_fms where "instances3_fms \ + { replacement_is_order_body_fm, + wfrec_replacement_order_pred_fm, + replacement_is_jump_cardinal_body_fm, + replacement_is_aleph_fm, + LambdaPair_in_M_fm(is_inj_fm(0,1,2),0) }" + +txt\This set has 5 internalized formulas.\ + +lemmas replacement_instances3_defs = + replacement_is_order_body_fm_def wfrec_replacement_order_pred_fm_def + replacement_is_jump_cardinal_body_fm_def + replacement_is_aleph_fm_def + +declare (in M_ZF3) replacement_instances3_defs [simp] + +locale M_ZF3_trans = M_ZF2_trans + M_ZF3 + +locale M_ZFC3 = M_ZFC2 + M_ZF3 + +locale M_ZFC3_trans = M_ZFC2_trans + M_ZF3_trans + +locale M_ctm3 = M_ctm2 + M_ZF3_trans + M_ZF2_trans + +locale M_ctm3_AC = M_ctm3 + M_ctm1_AC + M_ZFC3_trans + +locale forcing_data3 = forcing_data2 + M_ctm3_AC + +lemmas (in M_ZF2_trans) separation_instances = + separation_well_ord + separation_obase_equals separation_is_obase + separation_PiP_rel separation_surjP_rel + separation_radd_body separation_rmult_body + +lemma (in M_ZF3_trans) lam_replacement_inj_rel: + shows + "lam_replacement(##M, \p. inj\<^bsup>##M\<^esup>(fst(p),snd(p)))" + using lam_replacement_iff_lam_closed[THEN iffD2,of "\p. inj\<^bsup>M\<^esup>(fst(p),snd(p))"] + LambdaPair_in_M[where \="is_inj_fm(0,1,2)" and is_f="is_inj(##M)" and env="[]",OF + is_inj_fm_type _ is_inj_iff_sats[symmetric] inj_rel_iff,simplified] + arity_is_inj_fm[of 0 1 2] ord_simp_union transitivity fst_snd_closed + inj_rel_closed zero_in_M LambdaPair_in_M_replacement3 + by simp + +arity_theorem for "is_transitive_fm" +arity_theorem for "is_linear_fm" +arity_theorem for "is_wellfounded_on_fm" +arity_theorem for "is_well_ord_fm" + +arity_theorem for "pred_set_fm" +arity_theorem for "image_fm" +definition omap_wfrec_body where + "omap_wfrec_body(A,r) \ (\\\image_fm(2, 0, 1) \ + pred_set_fm + (succ(succ(succ(succ(succ(succ(succ(succ(succ(A))))))))), 3, + succ(succ(succ(succ(succ(succ(succ(succ(succ(r))))))))), 0) \\)" + +lemma type_omap_wfrec_body_fm :"A\nat \ r\nat \ omap_wfrec_body(A,r)\formula" + unfolding omap_wfrec_body_def by simp + +lemma arity_aux : "A\nat \ r\nat \ arity(omap_wfrec_body(A,r)) = (9+\<^sub>\A) \ (9+\<^sub>\r)" + unfolding omap_wfrec_body_def + using arity_image_fm arity_pred_set_fm pred_Un_distrib union_abs2[of 3] union_abs1 + by (simp add:FOL_arities, auto simp add:Un_assoc[symmetric] union_abs1) + +lemma arity_omap_wfrec: "A\nat \ r\nat \ + arity(is_wfrec_fm(omap_wfrec_body(A,r),succ(succ(succ(r))), 1, 0)) = + (4+\<^sub>\A) \ (4+\<^sub>\r)" + using Arities.arity_is_wfrec_fm[OF _ _ _ _ _ arity_aux,of A r "3+\<^sub>\r" 1 0] pred_Un_distrib + union_abs1 union_abs2 type_omap_wfrec_body_fm + by auto + +lemma arity_isordermap: "A\nat \ r\nat \d\nat\ + arity(is_ordermap_fm(A,r,d)) = succ(d) \ (succ(A) \ succ(r))" + unfolding is_ordermap_fm_def + using arity_lambda_fm[where i="(4+\<^sub>\A) \ (4+\<^sub>\r)",OF _ _ _ _ arity_omap_wfrec, + unfolded omap_wfrec_body_def] pred_Un_distrib union_abs1 + by auto + +lemma arity_is_ordertype: "A\nat \ r\nat \d\nat\ + arity(is_ordertype_fm(A,r,d)) = succ(d) \ (succ(A) \ succ(r))" + unfolding is_ordertype_fm_def + using arity_isordermap arity_image_fm pred_Un_distrib FOL_arities + by auto + +arity_theorem for "is_order_body_fm" + +lemma arity_is_order_body: "arity(is_order_body_fm(2,0,1)) = 3" + using arity_is_order_body_fm arity_is_ordertype ord_simp_union + by (simp add:FOL_arities) + +lemma (in M_ZF3_trans) replacement_is_order_body: + "X\M \ strong_replacement(##M, is_order_body(##M,X))" + apply(rule_tac strong_replacement_cong[ + where P="\ x f. M,[x,f,X] \ is_order_body_fm(2,0,1)",THEN iffD1]) + apply(rule_tac is_order_body_iff_sats[where env="[_,_,X]",symmetric]) + apply(simp_all add:zero_in_M) + apply(rule_tac replacement_ax3(1)[unfolded replacement_assm_def, rule_format, where env="[X]",simplified]) + apply(simp_all add: arity_is_order_body ) + done + +lemma (in M_pre_cardinal_arith) is_order_body_abs : + "M(X) \ M(x) \ M(z) \ is_order_body(M, X, x, z) \ + M(z) \ M(x) \ x\Pow_rel(M,X\X) \ well_ord(X, x) \ z = ordertype(X, x)" + using well_ord_abs is_well_ord_iff_wellordered is_ordertype_iff' ordertype_rel_abs + well_ord_is_linear subset_abs Pow_rel_char + unfolding is_order_body_def + by simp + + +definition H_order_pred where + "H_order_pred(A,r) \ \x f . f `` Order.pred(A, x, r)" + +relationalize "H_order_pred" "is_H_order_pred" + +lemma (in M_basic) H_order_pred_abs : + "M(A) \ M(r) \ M(x) \ M(f) \ M(z) \ + is_H_order_pred(M,A,r,x,f,z) \ z = H_order_pred(A,r,x,f)" + unfolding is_H_order_pred_def H_order_pred_def + by simp + +synthesize "is_H_order_pred" from_definition assuming "nonempty" + +lemma (in M_ZF3_trans) wfrec_replacement_order_pred: + "A\M \ r\M \ wfrec_replacement(##M, \x g z. is_H_order_pred(##M,A,r,x,g,z) , r)" + unfolding wfrec_replacement_def is_wfrec_def M_is_recfun_def is_H_order_pred_def + apply(rule_tac strong_replacement_cong[ + where P="\ x f. M,[x,f,r,A] \ order_pred_wfrec_body_fm(3,2,1,0)",THEN iffD1]) + apply(subst order_pred_wfrec_body_def[symmetric]) + apply(rule_tac order_pred_wfrec_body_iff_sats[where env="[_,_,r,A]",symmetric]) + apply(simp_all add:zero_in_M) + apply(rule_tac replacement_ax3(2)[unfolded replacement_assm_def, rule_format, where env="[r,A]",simplified]) + apply(simp_all add: arity_order_pred_wfrec_body_fm ord_simp_union) + done + +lemma (in M_ZF3_trans) wfrec_replacement_order_pred': + "A\M \ r\M \ wfrec_replacement(##M, \x g z. z = H_order_pred(A,r,x,g) , r)" + using wfrec_replacement_cong[OF H_order_pred_abs[of A r,rule_format] refl,THEN iffD1, + OF _ _ _ _ _ wfrec_replacement_order_pred[of A r]] + by simp + +sublocale M_ZF3_trans \ M_pre_cardinal_arith "##M" + using separation_instances wfrec_replacement_order_pred'[unfolded H_order_pred_def] + replacement_is_order_eq_map[unfolded order_eq_map_def] banach_replacement + by unfold_locales simp_all + +lemma (in M_ZF3_trans) replacement_ordertype: + "X\M \ strong_replacement(##M, \x z. z \ M \ x \ M \ x \ Pow\<^bsup>M\<^esup>(X \ X) \ well_ord(X, x) \ z = ordertype(X, x))" + using strong_replacement_cong[THEN iffD1,OF _ replacement_is_order_body,simplified] is_order_body_abs + unfolding is_order_body_def + by simp + +lemma arity_is_jump_cardinal_body: "arity(is_jump_cardinal_body'_fm(0,1)) = 2" + unfolding is_jump_cardinal_body'_fm_def + using arity_is_ordertype arity_is_well_ord_fm arity_is_Pow_fm arity_cartprod_fm + arity_Replace_fm[where i=5] ord_simp_union FOL_arities + by simp + +lemma (in M_ZF3_trans) replacement_is_jump_cardinal_body: + "strong_replacement(##M, is_jump_cardinal_body'(##M))" + apply(rule_tac strong_replacement_cong[ + where P="\ x f. M,[x,f] \ is_jump_cardinal_body'_fm(0,1)",THEN iffD1]) + apply(rule_tac is_jump_cardinal_body'_iff_sats[where env="[_,_]",symmetric]) + apply(simp_all add:zero_in_M) + apply(rule_tac replacement_ax3(3)[unfolded replacement_assm_def, rule_format, where env="[]",simplified]) + apply(simp_all add: arity_is_jump_cardinal_body ) + done + +lemma (in M_pre_cardinal_arith) univalent_aux2: "M(X) \ univalent(M,Pow_rel(M,X\X), + \r z. M(z) \ M(r) \ is_well_ord(M, X, r) \ is_ordertype(M, X, r, z))" + using is_well_ord_iff_wellordered + is_ordertype_iff[of _ X] + trans_on_subset[OF well_ord_is_trans_on] + well_ord_is_wf[THEN wf_on_subset_A] mem_Pow_rel_abs + unfolding univalent_def + by (simp) + +lemma (in M_pre_cardinal_arith) is_jump_cardinal_body_abs : + "M(X) \ M(c) \ is_jump_cardinal_body'(M, X, c) \ c = jump_cardinal_body'_rel(M,X)" + using well_ord_abs is_well_ord_iff_wellordered is_ordertype_iff' ordertype_rel_abs + well_ord_is_linear subset_abs Pow_rel_iff Replace_abs[of "Pow_rel(M,X\X)",OF _ _ + univalent_aux2] + unfolding is_jump_cardinal_body'_def jump_cardinal_body'_rel_def + by simp + +lemma (in M_ZF3_trans) replacement_jump_cardinal_body: + "strong_replacement(##M, \x z. z \ M \ x \ M \ z = jump_cardinal_body(##M, x))" + using strong_replacement_cong[THEN iffD1,OF _ replacement_is_jump_cardinal_body,simplified] + jump_cardinal_body_eq is_jump_cardinal_body_abs + by simp + +sublocale M_ZF3_trans \ M_pre_aleph "##M" + using replacement_ordertype replacement_jump_cardinal_body HAleph_wfrec_repl + by unfold_locales (simp_all add: transrec_replacement_def + wfrec_replacement_def is_wfrec_def M_is_recfun_def flip:setclass_iff) + +arity_theorem intermediate for "is_HAleph_fm" +lemma arity_is_HAleph_fm: "arity(is_HAleph_fm(2, 1, 0)) = 3" + using arity_fun_apply_fm[of "11" 0 1,simplified] + arity_is_HAleph_fm' arity_ordinal_fm arity_is_If_fm + arity_empty_fm arity_is_Limit_fm + arity_is_If_fm + arity_is_Limit_fm arity_empty_fm + arity_Replace_fm[where i="12" and v=10 and n=3] + pred_Un_distrib ord_simp_union + by (simp add:FOL_arities) + +lemma arity_is_Aleph: "arity(is_Aleph_fm(0, 1)) = 2" + unfolding is_Aleph_fm_def + using arity_transrec_fm[OF _ _ _ _ arity_is_HAleph_fm] ord_simp_union + by simp + +lemma (in M_ZF3_trans) replacement_is_aleph: + "strong_replacement(##M, \x y. Ord(x) \ is_Aleph(##M,x,y))" + apply(rule_tac strong_replacement_cong[ + where P="\ x y. M,[x,y] \ And(ordinal_fm(0),is_Aleph_fm(0,1))",THEN iffD1]) + apply (auto simp add: ordinal_iff_sats[where env="[_,_]",symmetric]) + apply(rule_tac is_Aleph_iff_sats[where env="[_,_]",THEN iffD2],simp_all add:zero_in_M) + apply(rule_tac is_Aleph_iff_sats[where env="[_,_]",THEN iffD1],simp_all add:zero_in_M) + apply(rule_tac replacement_ax3(4)[unfolded replacement_assm_def, rule_format, where env="[]",simplified]) + apply(simp_all add:arity_is_Aleph FOL_arities arity_ordinal_fm ord_simp_union is_Aleph_fm_type) + done + +lemma (in M_ZF3_trans) replacement_aleph_rel: + shows "strong_replacement(##M, \x y. Ord(x) \ y = \\<^bsub>x\<^esub>\<^bsup>M\<^esup>)" + using strong_replacement_cong[THEN iffD2,OF _ replacement_is_aleph,where P1="\x y . Ord(x) \ y=Aleph_rel(##M,x)"] + is_Aleph_iff + by auto + +sublocale M_ZF3_trans \ M_aleph "##M" + by (unfold_locales,simp add: replacement_aleph_rel) + +sublocale M_ZF2_trans \ M_FiniteFun "##M" + using separation_cons_like_rel separation_is_function + by unfold_locales simp + +sublocale M_ZFC1_trans \ M_AC "##M" + using choice_ax by (unfold_locales, simp_all) + +sublocale M_ZFC3_trans \ M_cardinal_AC "##M" .. + +(* TopLevel *) + +lemma (in M_ZF2_trans) separation_cardinal_rel_lesspoll_rel: + "(##M)(\) \ separation(##M, \x. x \\<^bsup>M\<^esup> \)" + using separation_in_ctm[where \="( \0 \ 1\ )" and env="[\]"] + is_lesspoll_iff nonempty + arity_is_cardinal_fm arity_is_lesspoll_fm arity_is_bij_fm ord_simp_union + by (simp add:FOL_arities) + +sublocale M_ZFC3_trans \ M_library "##M" + using separation_cardinal_rel_lesspoll_rel + by unfold_locales simp_all + +locale M_ZF4 = M_ZF3 + + assumes + ground_replacements4: + "ground_replacement_assm(M,env,replacement_is_order_body_fm)" + "ground_replacement_assm(M,env,wfrec_replacement_order_pred_fm)" + "ground_replacement_assm(M,env,replacement_is_jump_cardinal_body_fm)" + "ground_replacement_assm(M,env,replacement_is_aleph_fm)" + "ground_replacement_assm(M,env,LambdaPair_in_M_fm(is_inj_fm(0,1,2),0))" + "ground_replacement_assm(M,env,wfrec_Hfrc_at_fm)" + "ground_replacement_assm(M,env,list_repl1_intf_fm)" + "ground_replacement_assm(M,env,list_repl2_intf_fm)" + "ground_replacement_assm(M,env,formula_repl2_intf_fm)" + "ground_replacement_assm(M,env,eclose_repl2_intf_fm)" + "ground_replacement_assm(M,env,powapply_repl_fm)" + "ground_replacement_assm(M,env,phrank_repl_fm)" + "ground_replacement_assm(M,env,wfrec_rank_fm)" + "ground_replacement_assm(M,env,trans_repl_HVFrom_fm)" + "ground_replacement_assm(M,env,wfrec_Hcheck_fm)" + "ground_replacement_assm(M,env,repl_PHcheck_fm)" + "ground_replacement_assm(M,env,check_replacement_fm)" + "ground_replacement_assm(M,env,G_dot_in_M_fm)" + "ground_replacement_assm(M,env,repl_opname_check_fm)" + "ground_replacement_assm(M,env,tl_repl_intf_fm)" + "ground_replacement_assm(M,env,formula_repl1_intf_fm)" + "ground_replacement_assm(M,env,eclose_repl1_intf_fm)" + "ground_replacement_assm(M,env,replacement_is_omega_funspace_fm)" + "ground_replacement_assm(M,env,replacement_HAleph_wfrec_repl_body_fm)" + "ground_replacement_assm(M,env,replacement_is_fst2_snd2_fm)" + "ground_replacement_assm(M,env,replacement_is_sndfst_fst2_snd2_fm)" + "ground_replacement_assm(M,env,replacement_is_order_eq_map_fm)" + "ground_replacement_assm(M,env,replacement_transrec_apply_image_body_fm)" + "ground_replacement_assm(M,env,banach_replacement_iterates_fm)" + "ground_replacement_assm(M,env,replacement_is_trans_apply_image_fm)" + "ground_replacement_assm(M,env,banach_iterates_fm)" + "ground_replacement_assm(M,env,dcwit_repl_body_fm(6,5,4,3,2,0,1))" + "ground_replacement_assm(M,env,Lambda_in_M_fm(fst_fm(0,1),0))" + "ground_replacement_assm(M,env,Lambda_in_M_fm(big_union_fm(0,1),0))" + "ground_replacement_assm(M,env,Lambda_in_M_fm(is_cardinal_fm(0,1),0))" + "ground_replacement_assm(M,env,Lambda_in_M_fm(snd_fm(0,1),0))" + "ground_replacement_assm(M,env,LambdaPair_in_M_fm(image_fm(0,1,2),0))" + "ground_replacement_assm(M,env,LambdaPair_in_M_fm(setdiff_fm(0,1,2),0))" + "ground_replacement_assm(M,env,LambdaPair_in_M_fm(minimum_fm(0,1,2),0))" + "ground_replacement_assm(M,env,LambdaPair_in_M_fm(upair_fm(0,1,2),0))" + "ground_replacement_assm(M,env,LambdaPair_in_M_fm(is_RepFun_body_fm(0,1,2),0))" + "ground_replacement_assm(M,env,LambdaPair_in_M_fm(composition_fm(0,1,2),0))" + "ground_replacement_assm(M,env,Lambda_in_M_fm(is_converse_fm(0,1),0))" + "ground_replacement_assm(M,env,Lambda_in_M_fm(domain_fm(0,1),0))" + +definition instances4_fms where "instances4_fms \ + { ground_repl_fm(replacement_is_order_body_fm), + ground_repl_fm(wfrec_replacement_order_pred_fm), + ground_repl_fm(replacement_is_jump_cardinal_body_fm), + ground_repl_fm(replacement_is_aleph_fm), + ground_repl_fm(LambdaPair_in_M_fm(is_inj_fm(0,1,2),0)), + ground_repl_fm(wfrec_Hfrc_at_fm), + ground_repl_fm(list_repl1_intf_fm), + ground_repl_fm(list_repl2_intf_fm), + ground_repl_fm(formula_repl2_intf_fm), + ground_repl_fm(eclose_repl2_intf_fm), + ground_repl_fm(powapply_repl_fm), + ground_repl_fm(phrank_repl_fm), + ground_repl_fm(wfrec_rank_fm), + ground_repl_fm(trans_repl_HVFrom_fm), + ground_repl_fm(wfrec_Hcheck_fm), + ground_repl_fm(repl_PHcheck_fm), + ground_repl_fm(check_replacement_fm), + ground_repl_fm(G_dot_in_M_fm), + ground_repl_fm(repl_opname_check_fm), + ground_repl_fm(tl_repl_intf_fm), + ground_repl_fm(formula_repl1_intf_fm), + ground_repl_fm(eclose_repl1_intf_fm), + ground_repl_fm(replacement_is_omega_funspace_fm), + ground_repl_fm(replacement_HAleph_wfrec_repl_body_fm), + ground_repl_fm(replacement_is_fst2_snd2_fm), + ground_repl_fm(replacement_is_sndfst_fst2_snd2_fm), + ground_repl_fm(replacement_is_order_eq_map_fm), + ground_repl_fm(replacement_transrec_apply_image_body_fm), + ground_repl_fm(banach_replacement_iterates_fm), + ground_repl_fm(replacement_is_trans_apply_image_fm), + ground_repl_fm(banach_iterates_fm), + ground_repl_fm(dcwit_repl_body_fm(6,5,4,3,2,0,1)), + ground_repl_fm(Lambda_in_M_fm(fst_fm(0,1),0)), + ground_repl_fm(Lambda_in_M_fm(big_union_fm(0,1),0)), + ground_repl_fm(Lambda_in_M_fm(is_cardinal_fm(0,1),0)), + ground_repl_fm(Lambda_in_M_fm(snd_fm(0,1),0)), + ground_repl_fm(LambdaPair_in_M_fm(image_fm(0,1,2),0)), + ground_repl_fm(LambdaPair_in_M_fm(setdiff_fm(0,1,2),0)), + ground_repl_fm(LambdaPair_in_M_fm(minimum_fm(0,1,2),0)), + ground_repl_fm(LambdaPair_in_M_fm(upair_fm(0,1,2),0)), + ground_repl_fm(LambdaPair_in_M_fm(is_RepFun_body_fm(0,1,2),0)), + ground_repl_fm(LambdaPair_in_M_fm(composition_fm(0,1,2),0)), + ground_repl_fm(Lambda_in_M_fm(is_converse_fm(0,1),0)), + ground_repl_fm(Lambda_in_M_fm(domain_fm(0,1),0)) }" + +txt\This set has 44 internalized formulas, corresponding to the total count +of previous replacement instances.\ + +definition overhead where + "overhead \ instances1_fms \ instances2_fms \ instances3_fms \ instances4_fms" + +txt\Hence, the “overhead” to force $\CH$ and its negation consists +of 88 replacement instances.\ + +lemma instances3_fms_type[TC] : "instances3_fms \ formula" + unfolding instances3_fms_def replacement_is_order_body_fm_def + wfrec_replacement_order_pred_fm_def replacement_is_jump_cardinal_body_fm_def + replacement_is_aleph_fm_def + by (auto simp del: Lambda_in_M_fm_def + ccc_fun_closed_lemma_aux2_fm_def ccc_fun_closed_lemma_fm_def) + +lemma overhead_type: "overhead \ formula" + using instances1_fms_type instances2_fms_type + unfolding overhead_def instances3_fms_def instances4_fms_def + replacement_instances1_defs replacement_instances2_defs replacement_instances3_defs + using ground_repl_fm_type Lambda_in_M_fm_type + by (auto simp del: Lambda_in_M_fm_def + ccc_fun_closed_lemma_aux2_fm_def ccc_fun_closed_lemma_fm_def) + +locale M_ZF4_trans = M_ZF3_trans + M_ZF4 + +locale M_ZFC4 = M_ZFC3 + M_ZF4 + +locale M_ZFC4_trans = M_ZFC3_trans + M_ZF4_trans + +locale M_ctm4 = M_ctm3 + M_ZF4_trans + +locale M_ctm4_AC = M_ctm4 + M_ctm1_AC + M_ZFC4_trans + +locale forcing_data4 = forcing_data3 + M_ctm4_AC + +lemma M_satT_imp_M_ZF2: "(M \ ZF) \ M_ZF2(M)" +proof - + assume "M \ ZF" + then + have fin: "upair_ax(##M)" "Union_ax(##M)" "power_ax(##M)" + "extensionality(##M)" "foundation_ax(##M)" "infinity_ax(##M)" + unfolding ZF_def ZF_fin_def ZFC_fm_defs satT_def + using ZFC_fm_sats[of M] by simp_all + { + fix \ env + assume "\ \ formula" "env\list(M)" + moreover from \M \ ZF\ + have "\p\formula. (M, [] \ (ZF_separation_fm(p)))" + "\p\formula. (M, [] \ (ZF_replacement_fm(p)))" + unfolding ZF_def ZF_schemes_def by auto + moreover from calculation + have "arity(\) \ succ(length(env)) \ separation(##M, \x. (M, Cons(x, env) \ \))" + "arity(\) \ succ(succ(length(env))) \ strong_replacement(##M,\x y. sats(M,\,Cons(x,Cons(y, env))))" + using sats_ZF_separation_fm_iff sats_ZF_replacement_fm_iff by simp_all + } + with fin + show "M_ZF2(M)" + by unfold_locales (simp_all add:replacement_assm_def ground_replacement_assm_def) +qed + +lemma M_satT_imp_M_ZFC2: + shows "(M \ ZFC) \ M_ZFC2(M)" +proof - + have "(M \ ZF) \ choice_ax(##M) \ M_ZFC2(M)" + using M_satT_imp_M_ZF2[of M] unfolding M_ZF2_def M_ZFC1_def M_ZFC2_def + M_ZC_basic_def M_ZF1_def M_AC_def by auto + then + show ?thesis + unfolding ZFC_def by auto +qed + +lemma M_satT_instances12_imp_M_ZF2: + assumes "(M \ \Z\ \ {\Replacement(p)\ . p \ instances1_fms \ instances2_fms})" + shows "M_ZF2(M)" +proof - + from assms + have fin: "upair_ax(##M)" "Union_ax(##M)" "power_ax(##M)" + "extensionality(##M)" "foundation_ax(##M)" "infinity_ax(##M)" + unfolding ZF_fin_def Zermelo_fms_def ZFC_fm_defs satT_def + using ZFC_fm_sats[of M] by simp_all + moreover + { + fix \ env + from \M \ \Z\ \ {\Replacement(p)\ . p \ instances1_fms \ instances2_fms}\ + have "\p\formula. (M, [] \ (ZF_separation_fm(p)))" + unfolding Zermelo_fms_def ZF_def instances1_fms_def + instances2_fms_def by auto + moreover + assume "\ \ formula" "env\list(M)" + ultimately + have "arity(\) \ succ(length(env)) \ separation(##M, \x. (M, Cons(x, env) \ \))" + using sats_ZF_separation_fm_iff by simp_all + } + moreover + { + fix \ env + assume "\ \ instances1_fms \ instances2_fms" "env\list(M)" + moreover from this and \M \ \Z\ \ {\Replacement(p)\ . p \ instances1_fms \ instances2_fms}\ + have "M, [] \ \Replacement(\)\" by auto + ultimately + have "arity(\) \ succ(succ(length(env))) \ strong_replacement(##M,\x y. sats(M,\,Cons(x,Cons(y, env))))" + using sats_ZF_replacement_fm_iff[of \] instances1_fms_type instances2_fms_type by auto + } + ultimately + show ?thesis + unfolding instances1_fms_def instances2_fms_def + by unfold_locales (simp_all add:replacement_assm_def ground_replacement_assm_def) +qed + +lemma (in M_Z_basic) M_satT_Zermelo_fms: "M \ \Z\" + using upair_ax Union_ax power_ax extensionality foundation_ax + infinity_ax separation_ax sats_ZF_separation_fm_iff + unfolding Zermelo_fms_def ZF_fin_def + by auto + +lemma (in M_ZFC1) M_satT_ZC: "M \ ZC" + using upair_ax Union_ax power_ax extensionality foundation_ax + infinity_ax separation_ax sats_ZF_separation_fm_iff choice_ax + unfolding ZC_def Zermelo_fms_def ZF_fin_def + by auto + +locale M_ZF = M_Z_basic + + assumes + replacement_ax:"replacement_assm(M,env,\)" + +lemma M_satT_imp_M_ZF: " M \ ZF \ M_ZF(M)" +proof - + assume "M \ ZF" + then + have fin: "upair_ax(##M)" "Union_ax(##M)" "power_ax(##M)" + "extensionality(##M)" "foundation_ax(##M)" "infinity_ax(##M)" + unfolding ZF_def ZF_fin_def ZFC_fm_defs satT_def + using ZFC_fm_sats[of M] by simp_all + { + fix \ env + assume "\ \ formula" "env\list(M)" + moreover from \M \ ZF\ + have "\p\formula. (M, [] \ (ZF_separation_fm(p)))" + "\p\formula. (M, [] \ (ZF_replacement_fm(p)))" + unfolding ZF_def ZF_schemes_def by auto + moreover from calculation + have "arity(\) \ succ(length(env)) \ separation(##M, \x. (M, Cons(x, env) \ \))" + "arity(\) \ succ(succ(length(env))) \ strong_replacement(##M,\x y. sats(M,\,Cons(x,Cons(y, env))))" + using sats_ZF_separation_fm_iff sats_ZF_replacement_fm_iff by simp_all + } + with fin + show "M_ZF(M)" + unfolding M_ZF_def M_Z_basic_def M_ZF_axioms_def replacement_assm_def by simp +qed + +lemma (in M_ZF) M_satT_ZF: "M \ ZF" + using upair_ax Union_ax power_ax extensionality foundation_ax + infinity_ax separation_ax sats_ZF_separation_fm_iff + replacement_ax sats_ZF_replacement_fm_iff + unfolding ZF_def ZF_schemes_def ZF_fin_def replacement_assm_def + by auto + +lemma M_ZF_iff_M_satT: "M_ZF(M) \ (M \ ZF)" + using M_ZF.M_satT_ZF M_satT_imp_M_ZF + by auto + +locale M_ZFC = M_ZF + M_ZC_basic + +lemma M_ZFC_iff_M_satT: + notes iff_trans[trans] + shows "M_ZFC(M) \ (M \ ZFC)" +proof - + have "M_ZFC(M) \ (M \ ZF) \ choice_ax(##M)" + using M_ZF_iff_M_satT + unfolding M_ZFC_def M_ZC_basic_def M_AC_def M_ZF_def by auto + also + have " \ \ M \ ZFC" + unfolding ZFC_def by auto + ultimately + show ?thesis by simp +qed + +lemma M_satT_imp_M_ZF4: "(M \ ZF) \ M_ZF4(M)" +proof + assume "M \ ZF" + then + have fin: "upair_ax(##M)" "Union_ax(##M)" "power_ax(##M)" + "extensionality(##M)" "foundation_ax(##M)" "infinity_ax(##M)" + unfolding ZF_def ZF_fin_def ZFC_fm_defs satT_def + using ZFC_fm_sats[of M] by simp_all + { + fix \ env + assume "\ \ formula" "env\list(M)" + moreover from \M \ ZF\ + have "\p\formula. (M, [] \ (ZF_separation_fm(p)))" + "\p\formula. (M, [] \ (ZF_replacement_fm(p)))" + unfolding ZF_def ZF_schemes_def by auto + moreover from calculation + have "arity(\) \ succ(length(env)) \ separation(##M, \x. (M, Cons(x, env) \ \))" + "arity(\) \ succ(succ(length(env))) \ strong_replacement(##M,\x y. sats(M,\,Cons(x,Cons(y, env))))" + using sats_ZF_separation_fm_iff sats_ZF_replacement_fm_iff by simp_all + } + with fin + show "M_ZF4(M)" + by unfold_locales (simp_all add:replacement_assm_def ground_replacement_assm_def) +qed + +lemma M_satT_imp_M_ZFC4: + shows "(M \ ZFC) \ M_ZFC4(M)" +proof - + have "(M \ ZF) \ choice_ax(##M) \ M_ZFC4(M)" + using M_satT_imp_M_ZF4[of M] unfolding M_ZF4_def M_ZFC1_def M_ZFC4_def + M_ZF3_def M_ZFC3_def M_ZF2_def M_ZFC2_def + M_ZC_basic_def M_ZF1_def M_AC_def by auto + then + show ?thesis + unfolding ZFC_def by auto +qed + +lemma M_satT_overhead_imp_M_ZF4: + "(M \ ZC \ {\Replacement(p)\ . p \ overhead}) \ M_ZFC4(M)" +proof + assume "M \ ZC \ {\Replacement(p)\ . p \ overhead}" + then + have fin: "upair_ax(##M)" "Union_ax(##M)" "power_ax(##M)" "choice_ax(##M)" + "extensionality(##M)" "foundation_ax(##M)" "infinity_ax(##M)" + unfolding ZC_def ZF_fin_def Zermelo_fms_def ZFC_fm_defs satT_def + using ZFC_fm_sats[of M] by simp_all + moreover + { + fix \ env + from \M \ ZC \ {\Replacement(p)\ . p \ overhead}\ + have "\p\formula. (M, [] \ (ZF_separation_fm(p)))" + unfolding ZC_def Zermelo_fms_def ZF_def overhead_def instances1_fms_def + instances2_fms_def instances3_fms_def instances4_fms_def by auto + moreover + assume "\ \ formula" "env\list(M)" + ultimately + have "arity(\) \ succ(length(env)) \ separation(##M, \x. (M, Cons(x, env) \ \))" + using sats_ZF_separation_fm_iff by simp_all + } + moreover + { + fix \ env + assume "\ \ overhead" "env\list(M)" + moreover from this and \M \ ZC \ {\Replacement(p)\ . p \ overhead}\ + have "M, [] \ \Replacement(\)\" by auto + ultimately + have "arity(\) \ succ(succ(length(env))) \ strong_replacement(##M,\x y. sats(M,\,Cons(x,Cons(y, env))))" + using sats_ZF_replacement_fm_iff[of \] overhead_type by auto + } + ultimately + show "M_ZFC4(M)" + unfolding overhead_def instances1_fms_def + instances2_fms_def instances3_fms_def instances4_fms_def + by unfold_locales (simp_all add:replacement_assm_def ground_replacement_assm_def) +qed + +end \ No newline at end of file diff --git a/thys/Independence_CH/document/root.bib b/thys/Independence_CH/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/document/root.bib @@ -0,0 +1,75 @@ +@article{DBLP:journals/jar/PaulsonG96, + author = {Lawrence C. Paulson and + Krzysztof Grabczewski}, + title = {Mechanizing Set Theory}, + journal = {J. Autom. Reasoning}, + volume = {17}, + number = {3}, + pages = {291--323}, + year = {1996}, + xurl = {https://doi.org/10.1007/BF00283132}, + doi = {10.1007/BF00283132}, + timestamp = {Sat, 20 May 2017 00:22:31 +0200}, + biburl = {https://dblp.org/rec/bib/journals/jar/PaulsonG96}, + bibsource = {dblp computer science bibliography, https://dblp.org} +} + +@inproceedings{2018arXiv180705174G, + author = {Gunther, Emmanuel and Pagano, Miguel and S{\'a}nchez Terraf, Pedro}, + title = {First Steps Towards a Formalization of Forcing}, + booktitle = {Proceedings of the 13th Workshop on Logical and Semantic Frameworks + with Applications, {LSFA} 2018, Fortaleza, Brazil, September 26-28, + 2018}, + pages = {119--136}, + year = {2018}, + url = {https://doi.org/10.1016/j.entcs.2019.07.008}, + doi = {10.1016/j.entcs.2019.07.008}, + timestamp = {Wed, 05 Feb 2020 13:47:23 +0100}, + biburl = {https://dblp.org/rec/journals/entcs/GuntherPT19.bib}, + bibsource = {dblp computer science bibliography, https://dblp.org} +} + + +@ARTICLE{2019arXiv190103313G, + author = {Gunther, Emmanuel and Pagano, Miguel and S{\'a}nchez Terraf, Pedro}, + title = "{Mechanization of Separation in Generic Extensions}", + journal = {arXiv e-prints}, + keywords = {Computer Science - Logic in Computer Science, Mathematics - Logic, 03B35 (Primary) 03E40, 03B70, 68T15 (Secondary), F.4.1}, + year = 2019, + month = Jan, + eid = {arXiv:1901.03313}, + volume = {1901.03313}, +archivePrefix = {arXiv}, + eprint = {1901.03313}, + primaryClass = {cs.LO}, + adsurl = {https://ui.adsabs.harvard.edu/\#abs/2019arXiv190103313G}, + adsnote = {Provided by the SAO/NASA Astrophysics Data System}, + abstract = {We mechanize, in the proof assistant Isabelle, a proof of the axiom-scheme of Separation in generic extensions of models of set theory by using the fundamental theorems of forcing. We also formalize the satisfaction of the axioms of Extensionality, Foundation, Union, and Powerset. The axiom of Infinity is likewise treated, under additional assumptions on the ground model. In order to achieve these goals, we extended Paulson's library on constructibility with renaming of variables for internalized formulas, improved results on definitions by recursion on well-founded relations, and sharpened hypotheses in his development of relativization and absoluteness.} +} + +@inproceedings{2020arXiv200109715G, + author = {Gunther, Emmanuel and Pagano, Miguel and S{\'a}nchez Terraf, Pedro}, + title = "{Formalization of Forcing in Isabelle/ZF}", + isbn = {978-3-662-45488-6}, + booktitle = {Automated Reasoning. 10th International Joint Conference, IJCAR 2020, Paris, France, July 1--4, 2020, Proceedings, Part II}, + volume = 12167, + series = {Lecture Notes in Artificial Intelligence}, + editor = {Peltier, Nicolas and Sofronie-Stokkermans, Viorica}, + publisher = {Springer International Publishing}, + doi = {10.1007/978-3-030-51054-1}, + pages = {221--235}, + journal = {arXiv e-prints}, + keywords = {Computer Science - Logic in Computer Science, Mathematics - Logic, 03B35 (Primary) 03E40, 03B70, 68T15 (Secondary), F.4.1}, + year = 2020, + eid = {arXiv:2001.09715}, +archivePrefix = {arXiv}, + eprint = {2001.09715}, + primaryClass = {cs.LO}, + adsurl = {https://ui.adsabs.harvard.edu/abs/2020arXiv200109715G}, + abstract = {We formalize the theory of forcing in the set theory framework of +Isabelle/ZF. Under the assumption of the existence of a countable +transitive model of $\mathit{ZFC}$, we construct a proper generic extension and show +that the latter also satisfies $\mathit{ZFC}$. In doing so, we remodularized +Paulson's ZF-Constructibility library.}, + adsnote = {Provided by the SAO/NASA Astrophysics Data System} +} diff --git a/thys/Independence_CH/document/root.bst b/thys/Independence_CH/document/root.bst new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/document/root.bst @@ -0,0 +1,1440 @@ +%% +%% by pedro +%% Based on file `model1b-num-names.bst' +%% +%% +%% +ENTRY + { address + author + booktitle + chapter + edition + editor + howpublished + institution + journal + key + month + note + number + organization + pages + publisher + school + series + title + type + volume + year + } + {} + { label extra.label sort.label short.list } +INTEGERS { output.state before.all mid.sentence after.sentence after.block } +FUNCTION {init.state.consts} +{ #0 'before.all := + #1 'mid.sentence := + #2 'after.sentence := + #3 'after.block := +} +STRINGS { s t} +FUNCTION {output.nonnull} +{ 's := + output.state mid.sentence = + { ", " * write$ } + { output.state after.block = + { add.period$ write$ + newline$ + "\newblock " write$ + } + { output.state before.all = + 'write$ + { add.period$ " " * write$ } + if$ + } + if$ + mid.sentence 'output.state := + } + if$ + s +} +FUNCTION {output} +{ duplicate$ empty$ + 'pop$ + 'output.nonnull + if$ +} +FUNCTION {output.check} +{ 't := + duplicate$ empty$ + { pop$ "empty " t * " in " * cite$ * warning$ } + 'output.nonnull + if$ +} +FUNCTION {fin.entry} +{ add.period$ + write$ + newline$ +} + +FUNCTION {new.block} +{ output.state before.all = + 'skip$ + { after.block 'output.state := } + if$ +} +FUNCTION {new.sentence} +{ output.state after.block = + 'skip$ + { output.state before.all = + 'skip$ + { after.sentence 'output.state := } + if$ + } + if$ +} +FUNCTION {add.blank} +{ " " * before.all 'output.state := +} + +FUNCTION {date.block} +{ + skip$ +} + +FUNCTION {not} +{ { #0 } + { #1 } + if$ +} +FUNCTION {and} +{ 'skip$ + { pop$ #0 } + if$ +} +FUNCTION {or} +{ { pop$ #1 } + 'skip$ + if$ +} +FUNCTION {new.block.checkb} +{ empty$ + swap$ empty$ + and + 'skip$ + 'new.block + if$ +} +FUNCTION {field.or.null} +{ duplicate$ empty$ + { pop$ "" } + 'skip$ + if$ +} +FUNCTION {emphasize} +{ duplicate$ empty$ + { pop$ "" } + { "\textit{" swap$ * "}" * } + if$ +} +%% by pedro +FUNCTION {slanted} +{ duplicate$ empty$ + { pop$ "" } + { "\textsl{" swap$ * "}" * } + if$ +} +FUNCTION {smallcaps} +{ duplicate$ empty$ + { pop$ "" } + { "\textsc{" swap$ * "}" * } + if$ +} +FUNCTION {bold} +{ duplicate$ empty$ + { pop$ "" } + { "\textbf{" swap$ * "}" * } + if$ +} + + +FUNCTION {tie.or.space.prefix} +{ duplicate$ text.length$ #3 < + { "~" } + { " " } + if$ + swap$ +} + +FUNCTION {capitalize} +{ "u" change.case$ "t" change.case$ } + +FUNCTION {space.word} +{ " " swap$ * " " * } + % Here are the language-specific definitions for explicit words. + % Each function has a name bbl.xxx where xxx is the English word. + % The language selected here is ENGLISH +FUNCTION {bbl.and} +{ "and"} + +FUNCTION {bbl.etal} +{ "et~al." } + +FUNCTION {bbl.editors} +{ "eds." } + +FUNCTION {bbl.editor} +{ "ed." } + +FUNCTION {bbl.edby} +{ "edited by" } + +FUNCTION {bbl.edition} +{ "edition" } + +FUNCTION {bbl.volume} +{ "volume" } + +FUNCTION {bbl.of} +{ "of" } + +FUNCTION {bbl.number} +{ "number" } + +FUNCTION {bbl.nr} +{ "no." } + +FUNCTION {bbl.in} +{ "in" } + +FUNCTION {bbl.pages} +{ "pp." } + +FUNCTION {bbl.page} +{ "p." } + +FUNCTION {bbl.chapter} +{ "chapter" } + +FUNCTION {bbl.techrep} +{ "Technical Report" } + +FUNCTION {bbl.mthesis} +{ "Master's thesis" } + +FUNCTION {bbl.phdthesis} +{ "Ph.D. thesis" } + +MACRO {jan} {"January"} + +MACRO {feb} {"February"} + +MACRO {mar} {"March"} + +MACRO {apr} {"April"} + +MACRO {may} {"May"} + +MACRO {jun} {"June"} + +MACRO {jul} {"July"} + +MACRO {aug} {"August"} + +MACRO {sep} {"September"} + +MACRO {oct} {"October"} + +MACRO {nov} {"November"} + +MACRO {dec} {"December"} + +MACRO {acmcs} {"ACM Comput. Surv."} + +MACRO {acta} {"Acta Inf."} + +MACRO {cacm} {"Commun. ACM"} + +MACRO {ibmjrd} {"IBM J. Res. Dev."} + +MACRO {ibmsj} {"IBM Syst.~J."} + +MACRO {ieeese} {"IEEE Trans. Software Eng."} + +MACRO {ieeetc} {"IEEE Trans. Comput."} + +MACRO {ieeetcad} + {"IEEE Trans. Comput. Aid. Des."} + +MACRO {ipl} {"Inf. Process. Lett."} + +MACRO {jacm} {"J.~ACM"} + +MACRO {jcss} {"J.~Comput. Syst. Sci."} + +MACRO {scp} {"Sci. Comput. Program."} + +MACRO {sicomp} {"SIAM J. Comput."} + +MACRO {tocs} {"ACM Trans. Comput. Syst."} + +MACRO {tods} {"ACM Trans. Database Syst."} + +MACRO {tog} {"ACM Trans. Graphic."} + +MACRO {toms} {"ACM Trans. Math. Software"} + +MACRO {toois} {"ACM Trans. Office Inf. Syst."} + +MACRO {toplas} {"ACM Trans. Progr. Lang. Syst."} + +MACRO {tcs} {"Theor. Comput. Sci."} + +FUNCTION {bibinfo.check} +{ swap$ + duplicate$ missing$ + { + pop$ pop$ + "" + } + { duplicate$ empty$ + { + swap$ pop$ + } + { swap$ + "\bibinfo{" swap$ * "}{" * swap$ * "}" * + } + if$ + } + if$ +} +FUNCTION {bibinfo.warn} +{ swap$ + duplicate$ missing$ + { + swap$ "missing " swap$ * " in " * cite$ * warning$ pop$ + "" + } + { duplicate$ empty$ + { + swap$ "empty " swap$ * " in " * cite$ * warning$ + } + { swap$ + pop$ + } + if$ + } + if$ +} +STRINGS { bibinfo} +INTEGERS { nameptr namesleft numnames } + +FUNCTION {format.names} +{ 'bibinfo := + duplicate$ empty$ 'skip$ { + 's := + "" 't := + #1 'nameptr := + s num.names$ 'numnames := + numnames 'namesleft := + { namesleft #0 > } + { s nameptr + "{f{.}.~}{vv~}{ll}{, jj}" + format.name$ + bibinfo bibinfo.check + 't := + nameptr #1 > + { + namesleft #1 > + { ", " * t * } + { + "," * + s nameptr "{ll}" format.name$ duplicate$ "others" = + { 't := } + { pop$ } + if$ + t "others" = + { + " " * bbl.etal * + } + { " " * t * } + if$ + } + if$ + } + 't + if$ + nameptr #1 + 'nameptr := + namesleft #1 - 'namesleft := + } + while$ + } if$ +} +FUNCTION {format.names.ed} +{ + format.names +} +FUNCTION {format.key} +{ empty$ + { key field.or.null } + { "" } + if$ +} + +FUNCTION {format.authors} +{ author "author" format.names smallcaps +} +FUNCTION {get.bbl.editor} +{ editor num.names$ #1 > 'bbl.editors 'bbl.editor if$ } + +FUNCTION {format.editors} +{ editor "editor" format.names duplicate$ empty$ 'skip$ + { + " " * + get.bbl.editor + capitalize + "(" swap$ * ")" * + * + } + if$ +} +FUNCTION {format.note} +{ + note empty$ + { "" } + { note #1 #1 substring$ + duplicate$ "{" = + 'skip$ + { output.state mid.sentence = + { "l" } + { "u" } + if$ + change.case$ + } + if$ + note #2 global.max$ substring$ * "note" bibinfo.check + } + if$ +} + +FUNCTION {format.title} +{ title + duplicate$ empty$ 'skip$ + { "t" change.case$ } + if$ + "title" bibinfo.check +} +FUNCTION {format.full.names} +{'s := + "" 't := + #1 'nameptr := + s num.names$ 'numnames := + numnames 'namesleft := + { namesleft #0 > } + { s nameptr + "{vv~}{ll}" format.name$ + 't := + nameptr #1 > + { + namesleft #1 > + { ", " * t * } + { + s nameptr "{ll}" format.name$ duplicate$ "others" = + { 't := } + { pop$ } + if$ + t "others" = + { + " " * bbl.etal * + } + { + bbl.and + space.word * t * + } + if$ + } + if$ + } + 't + if$ + nameptr #1 + 'nameptr := + namesleft #1 - 'namesleft := + } + while$ +} + +FUNCTION {author.editor.key.full} +{ author empty$ + { editor empty$ + { key empty$ + { cite$ #1 #3 substring$ } + 'key + if$ + } + { editor format.full.names } + if$ + } + { author format.full.names } + if$ +} + +FUNCTION {author.key.full} +{ author empty$ + { key empty$ + { cite$ #1 #3 substring$ } + 'key + if$ + } + { author format.full.names } + if$ +} + +FUNCTION {editor.key.full} +{ editor empty$ + { key empty$ + { cite$ #1 #3 substring$ } + 'key + if$ + } + { editor format.full.names } + if$ +} + +FUNCTION {make.full.names} +{ type$ "book" = + type$ "inbook" = + or + 'author.editor.key.full + { type$ "proceedings" = + 'editor.key.full + 'author.key.full + if$ + } + if$ +} + +FUNCTION {output.bibitem} +{ newline$ + "\bibitem[{" write$ + label write$ + ")" make.full.names duplicate$ short.list = + { pop$ } + { * } + if$ + "}]{" * write$ + cite$ write$ + "}" write$ + newline$ + "" + before.all 'output.state := +} + +FUNCTION {n.dashify} +{ + 't := + "" + { t empty$ not } + { t #1 #1 substring$ "-" = + { t #1 #2 substring$ "--" = not + { "--" * + t #2 global.max$ substring$ 't := + } + { { t #1 #1 substring$ "-" = } + { "-" * + t #2 global.max$ substring$ 't := + } + while$ + } + if$ + } + { t #1 #1 substring$ * + t #2 global.max$ substring$ 't := + } + if$ + } + while$ +} + +FUNCTION {word.in} +{ bbl.in + ":" * + " " * } + +FUNCTION {format.date} +{ year "year" bibinfo.check duplicate$ empty$ + { + "empty year in " cite$ * "; set to ????" * warning$ + pop$ "????" + } + 'skip$ + if$ + % extra.label * + %% by pedro + " (" swap$ * ")" * +} +FUNCTION{format.year} +{ year "year" bibinfo.check duplicate$ empty$ + { "empty year in " cite$ * + "; set to ????" * + warning$ + pop$ "????" + } + { + } + if$ + % extra.label * + " (" swap$ * ")" * +} +FUNCTION {format.btitle} +{ title "title" bibinfo.check + duplicate$ empty$ 'skip$ + { + } + if$ + %% by pedro + "``" swap$ * "''" * +} +FUNCTION {either.or.check} +{ empty$ + 'pop$ + { "can't use both " swap$ * " fields in " * cite$ * warning$ } + if$ +} +FUNCTION {format.bvolume} +{ volume empty$ + { "" } + %% by pedro + { series "series" bibinfo.check + duplicate$ empty$ 'pop$ + { %slanted + } + if$ + "volume and number" number either.or.check + volume tie.or.space.prefix + "volume" bibinfo.check + bold + * * + } + if$ +} +FUNCTION {format.number.series} +{ volume empty$ + { number empty$ + { series field.or.null } + { series empty$ + { number "number" bibinfo.check } + { output.state mid.sentence = + { bbl.number } + { bbl.number capitalize } + if$ + number tie.or.space.prefix "number" bibinfo.check * * + bbl.in space.word * + series "series" bibinfo.check * + } + if$ + } + if$ + } + { "" } + if$ +} + +FUNCTION {format.edition} +{ edition duplicate$ empty$ 'skip$ + { + output.state mid.sentence = + { "l" } + { "t" } + if$ change.case$ + "edition" bibinfo.check + " " * bbl.edition * + } + if$ +} + +INTEGERS { multiresult } +FUNCTION {multi.page.check} +{ 't := + #0 'multiresult := + { multiresult not + t empty$ not + and + } + { t #1 #1 substring$ + duplicate$ "-" = + swap$ duplicate$ "," = + swap$ "+" = + or or + { #1 'multiresult := } + { t #2 global.max$ substring$ 't := } + if$ + } + while$ + multiresult +} +FUNCTION {format.pages} +{ pages duplicate$ empty$ 'skip$ + { duplicate$ multi.page.check + { + bbl.pages swap$ + n.dashify + } + { + bbl.page swap$ + } + if$ + tie.or.space.prefix + "pages" bibinfo.check + * * + } + if$ +} + +FUNCTION {format.pages.simple} +{ pages duplicate$ empty$ 'skip$ + { duplicate$ multi.page.check + { +% bbl.pages swap$ + n.dashify + } + { +% bbl.page swap$ + } + if$ + tie.or.space.prefix + "pages" bibinfo.check + * + } + if$ +} +FUNCTION {format.journal.pages} +{ pages duplicate$ empty$ 'pop$ + { swap$ duplicate$ empty$ + { pop$ pop$ format.pages } + { + ": " * + swap$ + n.dashify + "pages" bibinfo.check + * + } + if$ + } + if$ +} +FUNCTION {format.vol.num.pages} +{ volume field.or.null + duplicate$ empty$ 'skip$ + { + "volume" bibinfo.check + } + if$ + %% by pedro + bold + pages duplicate$ empty$ 'pop$ + { swap$ duplicate$ empty$ + { pop$ pop$ format.pages } + { + ": " * + swap$ + n.dashify + "pages" bibinfo.check + * + } + if$ + } + if$ + format.year * +} + +FUNCTION {format.chapter.pages} +{ chapter empty$ + { "" } + { type empty$ + { bbl.chapter } + { type "l" change.case$ + "type" bibinfo.check + } + if$ + chapter tie.or.space.prefix + "chapter" bibinfo.check + * * + } + if$ +} + +FUNCTION {format.booktitle} +{ + booktitle "booktitle" bibinfo.check +} +FUNCTION {format.in.ed.booktitle} +{ format.booktitle duplicate$ empty$ 'skip$ + { + editor "editor" format.names.ed duplicate$ empty$ 'pop$ + { + " " * + get.bbl.editor + capitalize + "(" swap$ * "), " * + * swap$ + * } + if$ + word.in swap$ * + } + if$ +} +FUNCTION {format.thesis.type} +{ type duplicate$ empty$ + 'pop$ + { swap$ pop$ + "t" change.case$ "type" bibinfo.check + } + if$ +} +FUNCTION {format.tr.number} +{ number "number" bibinfo.check + type duplicate$ empty$ + { pop$ bbl.techrep } + 'skip$ + if$ + "type" bibinfo.check + swap$ duplicate$ empty$ + { pop$ "t" change.case$ } + { tie.or.space.prefix * * } + if$ +} +FUNCTION {format.article.crossref} +{ + word.in + " \cite{" * crossref * "}" * +} +FUNCTION {format.book.crossref} +{ volume duplicate$ empty$ + { "empty volume in " cite$ * "'s crossref of " * crossref * warning$ + pop$ word.in + } + { bbl.volume + swap$ tie.or.space.prefix "volume" bibinfo.check * * bbl.of space.word * + } + if$ + " \cite{" * crossref * "}" * +} +FUNCTION {format.incoll.inproc.crossref} +{ + word.in + " \cite{" * crossref * "}" * +} +FUNCTION {format.org.or.pub} +{ 't := + "" + address empty$ t empty$ and + 'skip$ + { + t empty$ + { address "address" bibinfo.check * + } + { t * + address empty$ + 'skip$ + { ", " * address "address" bibinfo.check * } + if$ + } + if$ + } + if$ +} +FUNCTION {format.publisher.address} +{ publisher "publisher" bibinfo.check format.org.or.pub +} +FUNCTION {format.publisher.address.year} +{ publisher "publisher" bibinfo.check format.org.or.pub + format.journal.pages + format.year * +} + +FUNCTION {school.address.year} +{ school "school" bibinfo.warn + address empty$ + 'skip$ + { ", " * address "address" bibinfo.check * } + if$ + format.year * +} + +FUNCTION {format.publisher.address.pages} +{ publisher "publisher" bibinfo.check format.org.or.pub + format.year * + +} + +FUNCTION {format.organization.address} +{ organization "organization" bibinfo.check format.org.or.pub +} + +FUNCTION {format.organization.address.year} +{ organization "organization" bibinfo.check format.org.or.pub + format.journal.pages + format.year * +} + +FUNCTION {article} +{ "%Type = Article" write$ + output.bibitem + format.authors "author" output.check + author format.key output + format.title "title" output.check + crossref missing$ + { + journal + "journal" bibinfo.check + %% by pedro + emphasize + "journal" output.check + add.blank + format.vol.num.pages output + } + { format.article.crossref output.nonnull + } + if$ +% format.journal.pages + new.sentence + format.note output + fin.entry +} +FUNCTION {book} +{ "%Type = Book" write$ + output.bibitem + author empty$ + { format.editors "author and editor" output.check + editor format.key output + } + { format.authors output.nonnull + crossref missing$ + { "author and editor" editor either.or.check } + 'skip$ + if$ + } + if$ + format.btitle "title" output.check + crossref missing$ + { %% by pedro + format.bvolume output + format.number.series output + % format.bvolume output + format.publisher.address.year output + } + { + format.book.crossref output.nonnull + } + if$ + format.edition output + % format.date "year" output.check + new.sentence + format.note output + fin.entry +} +FUNCTION {booklet} +{ "%Type = Booklet" write$ + output.bibitem + format.authors output + author format.key output + format.title "title" output.check + howpublished "howpublished" bibinfo.check output + address "address" bibinfo.check output + format.date "year" output.check + new.sentence + format.note output + fin.entry +} + +FUNCTION {inbook} +{ "%Type = Inbook" write$ + output.bibitem + author empty$ + { format.editors "author and editor" output.check + editor format.key output + } + { format.authors output.nonnull + format.title "title" output.check + crossref missing$ + { "author and editor" editor either.or.check } + 'skip$ + if$ + } + if$ + format.btitle "title" output.check + crossref missing$ + { + format.bvolume output + format.number.series output + format.publisher.address output + format.pages "pages" output.check + format.edition output + format.date "year" output.check + } + { + format.book.crossref output.nonnull + } + if$ +% format.edition output +% format.pages "pages" output.check + new.sentence + format.note output + fin.entry +} + +FUNCTION {incollection} +{ "%Type = Incollection" write$ + output.bibitem + format.authors "author" output.check + author format.key output + format.title "title" output.check + crossref missing$ + { format.in.ed.booktitle "booktitle" output.check + format.bvolume output + format.number.series output + format.pages "pages" output.check + % format.publisher.address output + % format.date "year" output.check + format.publisher.address.year output + format.edition output + } + { format.incoll.inproc.crossref output.nonnull + } + if$ +% format.pages "pages" output.check + new.sentence + format.note output + fin.entry +} +FUNCTION {inproceedings} +{ "%Type = Inproceedings" write$ + output.bibitem + format.authors "author" output.check + author format.key output + format.title "title" output.check + crossref missing$ + { + journal + "journal" bibinfo.check + "journal" output.check + format.in.ed.booktitle "booktitle" output.check + format.bvolume output + format.number.series output + publisher empty$ + { %format.organization.address output + format.organization.address.year output +% format.journal.pages + } + { organization "organization" bibinfo.check output + format.publisher.address.year output + % format.date "year" output.check +% format.journal.pages + } + if$ + } + { format.incoll.inproc.crossref output.nonnull + format.journal.pages + } + if$ +% format.pages.simple "pages" output.check +%%% La que sigue la muevo adentro del "if" +% format.journal.pages + new.sentence + format.note output + fin.entry +} +FUNCTION {conference} { inproceedings } +FUNCTION {manual} +{ "%Type = Manual" write$ + output.bibitem + format.authors output + author format.key output + format.btitle "title" output.check + organization "organization" bibinfo.check output + address "address" bibinfo.check output + format.edition output + format.date "year" output.check + new.sentence + format.note output + fin.entry +} + +FUNCTION {mastersthesis} +{ "%Type = Masterthesis" write$ + output.bibitem + format.authors "author" output.check + author format.key output + format.btitle + "title" output.check + bbl.mthesis format.thesis.type output.nonnull +% school "school" bibinfo.warn output +% address "address" bibinfo.check output +% format.date "year" output.check + school.address.year output + new.sentence + format.note output + fin.entry +} + +FUNCTION {misc} +{ "%Type = Misc" write$ + output.bibitem + format.authors output + author format.key output + format.title output + howpublished "howpublished" bibinfo.check output + format.date "year" output.check + new.sentence + format.note output + fin.entry +} +FUNCTION {phdthesis} +{ "%Type = Phdthesis" write$ + output.bibitem + format.authors "author" output.check + author format.key output + format.btitle + "title" output.check + bbl.phdthesis format.thesis.type output.nonnull +% school "school" bibinfo.warn output +% address "address" bibinfo.check output +% format.date "year" output.check + school.address.year output + new.sentence + format.note output + fin.entry +} + +FUNCTION {proceedings} +{ "%Type = Proceedings" write$ + output.bibitem + format.editors output + editor format.key output + format.btitle "title" output.check + format.bvolume output + format.number.series output + publisher empty$ + { format.organization.address output } + { organization "organization" bibinfo.check output + format.publisher.address output + } + if$ + format.date "year" output.check + new.sentence + format.note output + fin.entry +} + +FUNCTION {techreport} +{ "%Type = Techreport" write$ + output.bibitem + format.authors "author" output.check + author format.key output + format.btitle + "title" output.check + format.tr.number output.nonnull + institution "institution" bibinfo.warn output + address "address" bibinfo.check output + format.date "year" output.check + new.sentence + format.note output + fin.entry +} + +FUNCTION {unpublished} +{ "%Type = Unpublished" write$ + output.bibitem + format.authors "author" output.check + author format.key output + format.title "title" output.check + format.date "year" output.check + new.sentence + format.note "note" output.check + fin.entry +} + +FUNCTION {default.type} { misc } +READ +FUNCTION {sortify} +{ purify$ + "l" change.case$ +} +INTEGERS { len } +FUNCTION {chop.word} +{ 's := + 'len := + s #1 len substring$ = + { s len #1 + global.max$ substring$ } + 's + if$ +} +FUNCTION {format.lab.names} +{ 's := + "" 't := + s #1 "{vv~}{ll}" format.name$ + s num.names$ duplicate$ + #2 > + { pop$ + " " * bbl.etal * + } + { #2 < + 'skip$ + { s #2 "{ff }{vv }{ll}{ jj}" format.name$ "others" = + { + " " * bbl.etal * + } + { bbl.and space.word * s #2 "{vv~}{ll}" format.name$ + * } + if$ + } + if$ + } + if$ +} + +FUNCTION {author.key.label} +{ author empty$ + { key empty$ + { cite$ #1 #3 substring$ } + 'key + if$ + } + { author format.lab.names } + if$ +} + +FUNCTION {author.editor.key.label} +{ author empty$ + { editor empty$ + { key empty$ + { cite$ #1 #3 substring$ } + 'key + if$ + } + { editor format.lab.names } + if$ + } + { author format.lab.names } + if$ +} + +FUNCTION {editor.key.label} +{ editor empty$ + { key empty$ + { cite$ #1 #3 substring$ } + 'key + if$ + } + { editor format.lab.names } + if$ +} + +FUNCTION {calc.short.authors} +{ type$ "book" = + type$ "inbook" = + or + 'author.editor.key.label + { type$ "proceedings" = + 'editor.key.label + 'author.key.label + if$ + } + if$ + 'short.list := +} + +FUNCTION {calc.label} +{ calc.short.authors + short.list + "(" + * + year duplicate$ empty$ + { pop$ "????" } + { purify$ #-1 #4 substring$ } + if$ + * + 'label := +} + +FUNCTION {sort.format.names} +{ 's := + #1 'nameptr := + "" + s num.names$ 'numnames := + numnames 'namesleft := + { namesleft #0 > } + { s nameptr + "{ll{ }}{ f{ }}{ jj{ }}" + format.name$ 't := + nameptr #1 > + { + " " * + namesleft #1 = t "others" = and + { "zzzzz" * } + { t sortify * } + if$ + } + { t sortify * } + if$ + nameptr #1 + 'nameptr := + namesleft #1 - 'namesleft := + } + while$ +} + +FUNCTION {sort.format.title} +{ 't := + "A " #2 + "An " #3 + "The " #4 t chop.word + chop.word + chop.word + sortify + #1 global.max$ substring$ +} +FUNCTION {author.sort} +{ author empty$ + { key empty$ + { "to sort, need author or key in " cite$ * warning$ + "" + } + { key sortify } + if$ + } + { author sort.format.names } + if$ +} +FUNCTION {author.editor.sort} +{ author empty$ + { editor empty$ + { key empty$ + { "to sort, need author, editor, or key in " cite$ * warning$ + "" + } + { key sortify } + if$ + } + { editor sort.format.names } + if$ + } + { author sort.format.names } + if$ +} +FUNCTION {editor.sort} +{ editor empty$ + { key empty$ + { "to sort, need editor or key in " cite$ * warning$ + "" + } + { key sortify } + if$ + } + { editor sort.format.names } + if$ +} +FUNCTION {presort} +{ calc.label + label sortify + " " + * + type$ "book" = + type$ "inbook" = + or + 'author.editor.sort + { type$ "proceedings" = + 'editor.sort + 'author.sort + if$ + } + if$ + #1 entry.max$ substring$ + 'sort.label := + sort.label + * + " " + * + title field.or.null + sort.format.title + * + #1 entry.max$ substring$ + 'sort.key$ := +} + +ITERATE {presort} +SORT +STRINGS { last.label next.extra } +INTEGERS { last.extra.num number.label } +FUNCTION {initialize.extra.label.stuff} +{ #0 int.to.chr$ 'last.label := + "" 'next.extra := + #0 'last.extra.num := + #0 'number.label := +} +FUNCTION {forward.pass} +{ last.label label = + { last.extra.num #1 + 'last.extra.num := + last.extra.num int.to.chr$ 'extra.label := + } + { "a" chr.to.int$ 'last.extra.num := + "" 'extra.label := + label 'last.label := + } + if$ + number.label #1 + 'number.label := +} +FUNCTION {reverse.pass} +{ next.extra "b" = + { "a" 'extra.label := } + 'skip$ + if$ + extra.label 'next.extra := + extra.label + duplicate$ empty$ + 'skip$ + { "{\natexlab{" swap$ * "}}" * } + if$ + 'extra.label := + label extra.label * 'label := +} +EXECUTE {initialize.extra.label.stuff} +ITERATE {forward.pass} +REVERSE {reverse.pass} +FUNCTION {bib.sort.order} +{ sort.label + " " + * + year field.or.null sortify + * + " " + * + title field.or.null + sort.format.title + * + #1 entry.max$ substring$ + 'sort.key$ := +} +ITERATE {bib.sort.order} +SORT +FUNCTION {begin.bib} +{ preamble$ empty$ + 'skip$ + { preamble$ write$ newline$ } + if$ + "\begin{small}\begin{thebibliography}{" number.label int.to.str$ * "}" * + write$ newline$ + "\expandafter\ifx\csname natexlab\endcsname\relax\def\natexlab#1{#1}\fi" + write$ newline$ + "\providecommand{\bibinfo}[2]{#2}" + write$ newline$ + "\ifx\xfnm\relax \def\xfnm[#1]{\unskip,\space#1}\fi" + write$ newline$ +} +EXECUTE {begin.bib} +EXECUTE {init.state.consts} +ITERATE {call.type$} +FUNCTION {end.bib} +{ newline$ + "\end{thebibliography}\end{small}" write$ newline$ +} +EXECUTE {end.bib} diff --git a/thys/Independence_CH/document/root.tex b/thys/Independence_CH/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/Independence_CH/document/root.tex @@ -0,0 +1,124 @@ +\documentclass[11pt,a4paper,english]{article} +\usepackage{isabelle,isabellesym} +\usepackage[numbers]{natbib} +\usepackage{babel} + +\usepackage{relsize} +\DeclareRobustCommand{\isactrlbsub}{\emph\bgroup\math{}\sb\bgroup\mbox\bgroup\isaspacing\itshape\smaller} +\DeclareRobustCommand{\isactrlesub}{\egroup\egroup\endmath\egroup} +\DeclareRobustCommand{\isactrlbsup}{\emph\bgroup\math{}\sp\bgroup\mbox\bgroup\isaspacing\itshape\smaller} +\DeclareRobustCommand{\isactrlesup}{\egroup\egroup\endmath\egroup} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage{eurosym} + %for \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \, \, \, \, + %\ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} +\newcommand{\forces}{\Vdash} +\newcommand{\dom}{\mathsf{dom}} +\renewcommand{\isacharunderscorekeyword}{\mbox{\_}} +\renewcommand{\isacharunderscore}{\mbox{\_}} +\renewcommand{\isasymtturnstile}{\isamath{\Vdash}} +\renewcommand{\isacharminus}{-} +\newcommand{\session}[1]{\textit{#1}} +\newcommand{\theory}[1]{\texttt{#1}} +\newcommand{\axiomas}[1]{\mathit{#1}} +\newcommand{\ZFC}{\axiomas{ZFC}} +\newcommand{\ZF}{\axiomas{ZF}} +\newcommand{\AC}{\axiomas{AC}} +\newcommand{\CH}{\axiomas{CH}} +\newcommand{\calV}{\mathcal{V}} + +\begin{document} + +\title{The Independence of the Continuum Hypothesis in Isabelle/ZF} +\author{Emmanuel Gunther\thanks{Universidad Nacional de C\'ordoba. + Facultad de Matem\'atica, Astronom\'{\i}a, F\'{\i}sica y + Computaci\'on.} + \and + Miguel Pagano\footnotemark[1] + \and + Pedro S\'anchez Terraf\footnotemark[1] \thanks{Centro de Investigaci\'on y Estudios de Matem\'atica + (CIEM-FaMAF), Conicet. C\'ordoba. Argentina. + Supported by Secyt-UNC project 33620180100465CB.} + \and + Mat\'{\i}as Steinberg\footnotemark[1] +} +\maketitle + +\begin{abstract} + We redeveloped our formalization of forcing in the set theory framework of + Isabelle/ZF. Under the assumption of the existence of a countable + transitive model of $\ZFC$, we construct proper generic extensions + that satisfy the Continuum Hypothesis and its negation. +\end{abstract} + + +\tableofcontents + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +\section{Introduction} +We formalize the theory of forcing. We work on top of the Isabelle/ZF +framework developed by \citet{DBLP:journals/jar/PaulsonG96}. Our +mechanization is described in more detail in our papers +\cite{2018arXiv180705174G} (LSFA 2018), \cite{2019arXiv190103313G}, +and \cite{2020arXiv200109715G} (IJCAR 2020). + +The main entry point of the present session is +\theory{Definitions\_Main.thy} (Section~\ref{sec:main-definitions}), +in which a path from fundamental set +theoretic concepts formalized in Isabelle reaching to our main theorems +is expounded. Cross-references to major milestones are provided there. + +In order to provide evidence for the correctness of several of our relativized +definitions, we needed to assume the Axiom of Choice ($\AC$) during the +aforementioned theory. Nevertheless, the whole of our development is +independent of $\AC$, and the theory \theory{CH.thy} already provides +all of our results and does not import that axiom. + +\subsection*{Release notes} +\label{sec:release-notes} + +Previous versions of this development can be found at +\url{https://cs.famaf.unc.edu.ar/~pedro/forcing/}. + +% generated text of all theories +\input{session} + +% optional bibliography +\bibliographystyle{root} +\bibliography{root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,668 +1,669 @@ ADS_Functor AI_Planning_Languages_Semantics AODV AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Actuarial_Mathematics Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BD_Security_Compositional BNF_CC BNF_Operations BTree Banach_Steinhaus Belief_Revision Bell_Numbers_Spivey BenOr_Kozen_Reif Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Blue_Eyes Bondy Boolean_Expression_Checkers Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CISC-Kernel CRDT CSP_RefTK CYK CZH_Elementary_Categories CZH_Foundations CZH_Universal_Constructions CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Cauchy Cayley_Hamilton Certification_Monads Chandy_Lamport Chord_Segments Circus Clean ClockSynchInst Closest_Pair_Points CoCon CoSMeDis CoSMed CofGroups Coinductive Coinductive_Languages Collections Combinatorics_Words Combinatorics_Words_Graph_Lemma Combinatorics_Words_Lyndon Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Bounded_Operators Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Conditional_Simplification Conditional_Transfer_Rule Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers CoreC++ Core_DOM Core_SC_DOM Correctness_Algebras Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties Cubic_Quartic_Equations DFS_Framework DOM_Components DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Design_Theory Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos Dominance_CHK DynamicArchitectures Dynamic_Tables E_Transcendental Echelon_Form EdmondsKarp_Maxflow Efficient-Mergesort Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Equivalence_Relation_Enumeration Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Eval_FO Example-Submission Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Axiomatic FOL_Harrison FOL_Seq_Calc1 FOL_Seq_Calc2 Factor_Algebraic_Polynomial Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF Finitely_Generated_Abelian_Groups First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Foundation_of_geometry Fourier FO_Theory_Rewriting Free-Boolean-Algebra Free-Groups Fresh_Identifiers FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC GaleStewart_Games Gale_Shapley Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Grothendieck_Schemes Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Hahn_Jordan_Decomposition Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL Hyperdual IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IMP_Compiler IP_Addresses Imperative_Insertion_Sort Impossible_Geometry Incompleteness Incredible_Proof_Machine +Independence_CH Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpolation_Polynomials_HOL_Algebra Interpreter_Optimizations Interval_Arithmetic_Word32 Intro_Dest_Elim Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Irrationals_From_THEBOOK IsaGeoCoq Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Kleene_Algebra Knights_Tour Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option Lifting_the_Exponent LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Logging_Independent_Anonymity Lowe_Ontological_Argument Lower_Semicontinuous Lp LP_Duality Lucas_Theorem MDP-Algorithms MDP-Rewards MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Median_Method Median_Of_Medians_Selection Menger Mereology Mersenne_Primes Metalogic_ProofChecker MiniML MiniSail Minimal_SSA Minkowskis_Theorem Minsky_Machines Modal_Logics_for_NTS Modular_Assembly_Kit_Security Modular_arithmetic_LLL_and_HNF_algorithms Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Myhill-Nerode Name_Carrying_Type_Inference Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers Nominal2 Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Octonions OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker PAL PCF PLM POPLmark-deBruijn PSemigroupsConvolution Padic_Ints Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Poincare_Bendixson Poincare_Disc Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Program-Conflict-Analysis Progress_Tracking Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem Public_Announcement_Logic QHLProver QR_Decomposition Quantales Quasi_Borel_Spaces Quaternions Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Real_Power Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Registers Regression_Test_Selection Regular-Sets Regular_Algebras Regular_Tree_Relations Relation_Algebra Relational-Incorrectness-Logic Relational_Disjoint_Set_Forests Relational_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rep_Fin_Groups ResiduatedTransitionSystem Residuated_Lattices Resolution_FOL Rewriting_Z Ribbon_Proofs Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Roth_Arithmetic_Progressions Routing Roy_Floyd_Warshall SATSolverVerification SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Saturation_Framework Saturation_Framework_Extensions Schutz_Spacetime Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL SequentInvertibility Shadow_DOM Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Firewall Simplex Simplicial_complexes_and_boolean_functions SimplifiedOntologicalArgument Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Sort_Encodings Source_Coding_Theorem SpecCheck Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares Sunflowers SuperCalc Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Szemeredi_Regularity Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Three_Circles Timed_Automata Topological_Semantics Topology TortoiseHare Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Transitive_Models Treaps Tree-Automata Tree_Decomposition Triangle Trie Twelvefold_Way Tycon Types_Tableaus_and_Goedels_God Types_To_Sets_Extension UPF UPF_Firewall UTP Universal_Hash_Families Universal_Turing_Machine UpDown_Scheme Valuation Van_Emde_Boas_Trees Van_der_Waerden VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves Virtual_Substitution VolpanoSmith VYDRA_MDL WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Weighted_Path_Order Well_Quasi_Orders Wetzels_Problem Winding_Number_Eval Word_Lib WorkerWrapper X86_Semantics XML Youngs_Inequality ZFC_in_HOL Zeta_3_Irrational Zeta_Function pGCL