diff --git a/thys/Delta_System_Lemma/Cardinal_Library.thy b/thys/Delta_System_Lemma/Cardinal_Library.thy --- a/thys/Delta_System_Lemma/Cardinal_Library.thy +++ b/thys/Delta_System_Lemma/Cardinal_Library.thy @@ -1,691 +1,680 @@ section\Cardinal Arithmetic under Choice\label{sec:cardinal-lib}\ theory Cardinal_Library imports ZF_Library ZF.Cardinal_AC begin text\This theory includes results on cardinalities that depend on $\AC$\ subsection\Results on cardinal exponentiation\ text\Non trivial instances of cardinal exponentiation require that the relevant function spaces are well-ordered, hence this implies a strong use of choice.\ lemma cexp_eqpoll_cong: assumes "A \ A'" "B \ B'" shows "A\<^bsup>\B\<^esup> = A'\<^bsup>\B'\<^esup>" unfolding cexp_def using cardinal_eqpoll_iff function_space_eqpoll_cong assms by simp lemma cexp_cexp_cmult: "(\\<^bsup>\\1\<^esup>)\<^bsup>\\2\<^esup> = \\<^bsup>\\2 \ \1\<^esup>" proof - have "(\\<^bsup>\\1\<^esup>)\<^bsup>\\2\<^esup> = (\1 \ \)\<^bsup>\\2\<^esup>" using cardinal_eqpoll by (intro cexp_eqpoll_cong) (simp_all add:cexp_def) also have " \ = \\<^bsup>\\2 \ \1\<^esup>" unfolding cexp_def using curry_eqpoll cardinal_cong by blast also have " \ = \\<^bsup>\\2 \ \1\<^esup>" using cardinal_eqpoll[THEN eqpoll_sym] unfolding cmult_def by (intro cexp_eqpoll_cong) (simp) finally show ?thesis . qed lemma cardinal_Pow: "|Pow(X)| = 2\<^bsup>\X\<^esup>" \ \Perhaps it's better with |X|\ using cardinal_eqpoll_iff[THEN iffD2, OF Pow_eqpoll_function_space] unfolding cexp_def by simp lemma cantor_cexp: assumes "Card(\)" shows "\ < 2\<^bsup>\\\<^esup>" using assms Card_is_Ord Card_cexp proof (intro not_le_iff_lt[THEN iffD1] notI) assume "2\<^bsup>\\\<^esup> \ \" then have "|Pow(\)| \ \" using cardinal_Pow by simp with assms have "Pow(\) \ \" using cardinal_eqpoll_iff Card_le_imp_lepoll Card_cardinal_eq by auto then obtain g where "g \ inj(Pow(\), \)" by blast then show "False" using cantor_inj by simp qed simp lemma cexp_left_mono: assumes "\1 \ \2" shows "\1\<^bsup>\\\<^esup> \ \2\<^bsup>\\\<^esup>" (* \ \short, unreadable proof: \ unfolding cexp_def using subset_imp_lepoll[THEN lepoll_imp_cardinal_le] assms le_subset_iff[THEN iffD1, OF assms] Pi_weaken_type[of _ _ "\_. \1" "\_. \2"] by auto *) proof - from assms have "\1 \ \2" using le_subset_iff by simp then have "\ \ \1 \ \ \ \2" using Pi_weaken_type by auto then show ?thesis unfolding cexp_def using lepoll_imp_cardinal_le subset_imp_lepoll by simp qed lemma cantor_cexp': assumes "2 \ \" "Card(\)" shows "\ < \\<^bsup>\\\<^esup>" using cexp_left_mono assms cantor_cexp lt_trans2 by blast lemma InfCard_cexp: assumes "2 \ \" "InfCard(\)" shows "InfCard(\\<^bsup>\\\<^esup>)" using assms cantor_cexp'[THEN leI] le_trans Card_cexp unfolding InfCard_def by auto lemmas InfCard_cexp' = InfCard_cexp[OF nats_le_InfCard, simplified] \ \\<^term>\InfCard(\) \ InfCard(\) \ InfCard(\\<^bsup>\\\<^esup>)\\ subsection\Miscellaneous\ lemma cardinal_RepFun_le: "|{f(a) . a\A}| \ |A|" proof - have "(\x\A. f(x)) \ surj(A, {f(a) . a\A})" unfolding surj_def using lam_funtype by auto then show ?thesis using surj_implies_cardinal_le by blast qed lemma subset_imp_le_cardinal: "A \ B \ |A| \ |B|" using subset_imp_lepoll[THEN lepoll_imp_cardinal_le] . lemma lt_cardinal_imp_not_subset: "|A| < |B| \ \ B \ A" using subset_imp_le_cardinal le_imp_not_lt by blast lemma cardinal_lt_csucc_iff: "Card(K) \ |K'| < K\<^sup>+ \ |K'| \ K" by (simp add: Card_lt_csucc_iff) lemma cardinal_UN_le_nat: "(\i. i\\ \ |X(i)| \ \) \ |\i\\. X(i)| \ \" by (simp add: cardinal_UN_le InfCard_nat) lemma lepoll_imp_cardinal_UN_le: notes [dest] = InfCard_is_Card Card_is_Ord assumes "InfCard(K)" "J \ K" "\i. i\J \ |X(i)| \ K" shows "|\i\J. X(i)| \ K" proof - from \J \ K\ obtain f where "f \ inj(J,K)" by blast define Y where "Y(k) \ if k\range(f) then X(converse(f)`k) else 0" for k have "i\J \ f`i \ K" for i using inj_is_fun[OF \f \ inj(J,K)\] by auto have "(\i\J. X(i)) \ (\i\K. Y(i))" proof (standard, elim UN_E) fix x i assume "i\J" "x\X(i)" with \f \ inj(J,K)\ \i\J \ f`i \ K\ have "x \ Y(f`i)" "f`i \ K" unfolding Y_def using inj_is_fun[OF \f \ inj(J,K)\] right_inverse apply_rangeI by auto then show "x \ (\i\K. Y(i))" by auto qed then have "|\i\J. X(i)| \ |\i\K. Y(i)|" unfolding Y_def using subset_imp_le_cardinal by simp with assms \\i. i\J \ f`i \ K\ show "|\i\J. X(i)| \ K" using inj_converse_fun[OF \f \ inj(J,K)\] unfolding Y_def by (rule_tac le_trans[OF _ cardinal_UN_le]) (auto intro:Ord_0_le)+ qed \ \For backwards compatibility\ lemmas leqpoll_imp_cardinal_UN_le = lepoll_imp_cardinal_UN_le lemma cardinal_lt_csucc_iff': includes Ord_dests assumes "Card(\)" shows "\ < |X| \ \\<^sup>+ \ |X|" using assms cardinal_lt_csucc_iff[of \ X] Card_csucc[of \] not_le_iff_lt[of "\\<^sup>+" "|X|"] not_le_iff_lt[of "|X|" \] by blast lemma lepoll_imp_subset_bij: "X \ Y \ (\Z. Z \ Y \ Z \ X)" proof assume "X \ Y" then obtain j where "j \ inj(X,Y)" by blast then have "range(j) \ Y" "j \ bij(X,range(j))" using inj_bij_range inj_is_fun range_fun_subset_codomain by blast+ then show "\Z. Z \ Y \ Z \ X" using eqpoll_sym unfolding eqpoll_def by force next assume "\Z. Z \ Y \ Z \ X" then obtain Z f where "f \ bij(Z,X)" "Z \ Y" unfolding eqpoll_def by force then have "converse(f) \ inj(X,Y)" using bij_is_inj inj_weaken_type bij_converse_bij by blast then show "X \ Y" by blast qed text\The following result proves to be very useful when combining \<^term>\cardinal\ and \<^term>\eqpoll\ in a calculation.\ lemma cardinal_Card_eqpoll_iff: "Card(\) \ |X| = \ \ X \ \" using Card_cardinal_eq[of \] cardinal_eqpoll_iff[of X \] by auto \ \Compare @{thm "le_Card_iff"}\ lemma lepoll_imp_lepoll_cardinal: assumes "X \ Y" shows "X \ |Y|" using assms cardinal_Card_eqpoll_iff[of "|Y|" Y] lepoll_eq_trans[of _ _ "|Y|"] by simp lemma lepoll_Un: assumes "InfCard(\)" "A \ \" "B \ \" shows "A \ B \ \" proof - have "A \ B \ sum(A,B)" using Un_lepoll_sum . moreover note assms moreover from this have "|sum(A,B)| \ \ \ \" using sum_lepoll_mono[of A \ B \] lepoll_imp_cardinal_le unfolding cadd_def by auto ultimately show ?thesis using InfCard_cdouble_eq Card_cardinal_eq InfCard_is_Card Card_le_imp_lepoll[of "sum(A,B)" \] lepoll_trans[of "A\B"] by auto qed lemma cardinal_Un_le: assumes "InfCard(\)" "|A| \ \" "|B| \ \" shows "|A \ B| \ \" using assms lepoll_Un le_Card_iff InfCard_is_Card by auto text\This is the unconditional version under choice of @{thm Cardinal.Finite_cardinal_iff}.\ lemma Finite_cardinal_iff': "Finite(|i|) \ Finite(i)" using cardinal_eqpoll_iff eqpoll_imp_Finite_iff by fastforce lemma cardinal_subset_of_Card: assumes "Card(\)" "a \ \" shows "|a| < \ \ |a| = \" proof - from assms have "|a| < |\| \ |a| = |\|" using subset_imp_le_cardinal le_iff by simp with assms show ?thesis using Card_cardinal_eq by simp qed lemma cardinal_cases: includes Ord_dests shows "Card(\) \ |X| < \ \ \ |X| \ \" using not_le_iff_lt by auto subsection\Countable and uncountable sets\ -\ \Kunen's Definition I.10.5\ -definition - countable :: "i\o" where - "countable(X) \ X \ \" - -lemma countableI[intro]: "X \ \ \ countable(X)" - unfolding countable_def by simp - -lemma countableD[dest]: "countable(X) \ X \ \" - unfolding countable_def by simp - lemma countable_iff_cardinal_le_nat: "countable(X) \ |X| \ \" using le_Card_iff[of \ X] Card_nat unfolding countable_def by simp lemma lepoll_countable: "X \ Y \ countable(Y) \ countable(X)" using lepoll_trans[of X Y] by blast \ \Next lemma can be proved without using AC\ lemma surj_countable: "countable(X) \ f \ surj(X,Y) \ countable(Y)" using surj_implies_cardinal_le[of f X Y, THEN le_trans] countable_iff_cardinal_le_nat by simp lemma Finite_imp_countable: "Finite(X) \ countable(X)" unfolding Finite_def by (auto intro:InfCard_nat nats_le_InfCard[of _ \, THEN le_imp_lepoll] dest!:eq_lepoll_trans[of X _ \]) lemma countable_imp_countable_UN: assumes "countable(J)" "\i. i\J \ countable(X(i))" shows "countable(\i\J. X(i))" using assms lepoll_imp_cardinal_UN_le[of \ J X] InfCard_nat countable_iff_cardinal_le_nat by auto lemma countable_union_countable: assumes "\x. x \ C \ countable(x)" "countable(C)" shows "countable(\C)" using assms countable_imp_countable_UN[of C "\x. x"] by simp abbreviation uncountable :: "i\o" where "uncountable(X) \ \ countable(X)" lemma uncountable_iff_nat_lt_cardinal: "uncountable(X) \ \ < |X|" using countable_iff_cardinal_le_nat not_le_iff_lt by simp lemma uncountable_not_empty: "uncountable(X) \ X \ 0" using empty_lepollI by auto lemma uncountable_imp_Infinite: "uncountable(X) \ Infinite(X)" using uncountable_iff_nat_lt_cardinal[of X] lepoll_nat_imp_Infinite[of X] cardinal_le_imp_lepoll[of \ X] leI by simp lemma uncountable_not_subset_countable: assumes "countable(X)" "uncountable(Y)" shows "\ (Y \ X)" using assms lepoll_trans subset_imp_lepoll[of Y X] by blast subsection\Results on Alephs\ lemma nat_lt_Aleph1: "\ < \\<^bsub>1\<^esub>" by (simp add: Aleph_def lt_csucc) lemma zero_lt_Aleph1: "0 < \\<^bsub>1\<^esub>" by (rule lt_trans[of _ "\"], auto simp add: ltI nat_lt_Aleph1) lemma le_aleph1_nat: "Card(k) \ k<\\<^bsub>1\<^esub> \ k \ \" by (simp add: Aleph_def Card_lt_csucc_iff Card_nat) lemma Aleph_succ: "\\<^bsub>succ(\)\<^esub> = \\<^bsub>\\<^esub>\<^sup>+" unfolding Aleph_def by simp lemma lesspoll_aleph_plus_one: assumes "Ord(\)" shows "d \ \\<^bsub>succ(\)\<^esub> \ d \ \\<^bsub>\\<^esub>" using assms lesspoll_csucc Aleph_succ Card_is_Ord by simp lemma cardinal_Aleph [simp]: "Ord(\) \ |\\<^bsub>\\<^esub>| = \\<^bsub>\\<^esub>" using Card_cardinal_eq by simp \ \Could be proved without using AC\ lemma Aleph_lesspoll_increasing: includes Aleph_intros shows "a < b \ \\<^bsub>a\<^esub> \ \\<^bsub>b\<^esub>" using cardinal_lt_iff_lesspoll[of "\\<^bsub>a\<^esub>" "\\<^bsub>b\<^esub>"] Card_cardinal_eq[of "\\<^bsub>b\<^esub>"] lt_Ord lt_Ord2 Card_Aleph[THEN Card_is_Ord] by auto lemma uncountable_iff_subset_eqpoll_Aleph1: includes Ord_dests notes Aleph_zero_eq_nat[simp] Card_nat[simp] Aleph_succ[simp] shows "uncountable(X) \ (\S. S \ X \ S \ \\<^bsub>1\<^esub>)" proof assume "uncountable(X)" then have "\\<^bsub>1\<^esub> \ X" using uncountable_iff_nat_lt_cardinal cardinal_lt_csucc_iff' cardinal_le_imp_lepoll by force then obtain S where "S \ X" "S \ \\<^bsub>1\<^esub>" using lepoll_imp_subset_bij by auto then show "\S. S \ X \ S \ \\<^bsub>1\<^esub>" using cardinal_cong Card_csucc[of \] Card_cardinal_eq by auto next assume "\S. S \ X \ S \ \\<^bsub>1\<^esub>" then have "\\<^bsub>1\<^esub> \ X" using subset_imp_lepoll[THEN [2] eq_lepoll_trans, of "\\<^bsub>1\<^esub>" _ X, OF eqpoll_sym] by auto then show "uncountable(X)" using Aleph_lesspoll_increasing[of 0 1, THEN [2] lesspoll_trans1, of "\\<^bsub>1\<^esub>"] lepoll_trans[of "\\<^bsub>1\<^esub>" X \] by auto qed lemma lt_Aleph_imp_cardinal_UN_le_nat: "function(G) \ domain(G) \ \ \ \n\domain(G). |G`n|<\\<^bsub>1\<^esub> \ |\n\domain(G). G`n|\\" proof - assume "function(G)" let ?N="domain(G)" and ?R="\n\domain(G). G`n" assume "?N \ \" assume Eq1: "\n\?N. |G`n|<\\<^bsub>1\<^esub>" { fix n assume "n\?N" with Eq1 have "|G`n| \ \" using le_aleph1_nat by simp } then have "n\?N \ |G`n| \ \" for n . with \?N \ \\ show ?thesis using InfCard_nat lepoll_imp_cardinal_UN_le by simp qed lemma Aleph1_eq_cardinal_vimage: "f:\\<^bsub>1\<^esub>\\ \ \n\\. |f-``{n}| = \\<^bsub>1\<^esub>" proof - assume "f:\\<^bsub>1\<^esub>\\" then have "function(f)" "domain(f) = \\<^bsub>1\<^esub>" "range(f)\\" by (simp_all add: domain_of_fun fun_is_function range_fun_subset_codomain) let ?G="\n\range(f). f-``{n}" from \f:\\<^bsub>1\<^esub>\\\ have "range(f) \ \" by (simp add: range_fun_subset_codomain) then have "domain(?G) \ \" using subset_imp_lepoll by simp have "function(?G)" by (simp add:function_lam) from \f:\\<^bsub>1\<^esub>\\\ have "n\\ \ f-``{n} \ \\<^bsub>1\<^esub>" for n using Pi_vimage_subset by simp with \range(f) \ \\ have "\\<^bsub>1\<^esub> = (\n\range(f). f-``{n})" proof (intro equalityI, intro subsetI) fix x assume "x \ \\<^bsub>1\<^esub>" with \f:\\<^bsub>1\<^esub>\\\ \function(f)\ \domain(f) = \\<^bsub>1\<^esub>\ have "x \ f-``{f`x}" "f`x \ range(f)" using function_apply_Pair vimage_iff apply_rangeI by simp_all then show "x \ (\n\range(f). f-``{n})" by auto qed auto { assume "\n\range(f). |f-``{n}| < \\<^bsub>1\<^esub>" then have "\n\domain(?G). |?G`n| < \\<^bsub>1\<^esub>" using zero_lt_Aleph1 by (auto) with \function(?G)\ \domain(?G) \ \\ have "|\n\domain(?G). ?G`n|\\" using lt_Aleph_imp_cardinal_UN_le_nat by blast then have "|\n\range(f). f-``{n}|\\" by simp with \\\<^bsub>1\<^esub> = _\ have "|\\<^bsub>1\<^esub>| \ \" by simp then have "\\<^bsub>1\<^esub> \ \" using Card_Aleph Card_cardinal_eq by simp then have "False" using nat_lt_Aleph1 by (blast dest:lt_trans2) } with \range(f)\\\ obtain n where "n\\" "\(|f -`` {n}| < \\<^bsub>1\<^esub>)" by blast moreover from this have "\\<^bsub>1\<^esub> \ |f-``{n}|" using not_lt_iff_le Card_is_Ord by auto moreover note \n\\ \ f-``{n} \ \\<^bsub>1\<^esub>\ ultimately show ?thesis using subset_imp_le_cardinal[THEN le_anti_sym, of _ "\\<^bsub>1\<^esub>"] Card_Aleph Card_cardinal_eq by auto qed \ \There is some asymmetry between assumptions and conclusion (\<^term>\(\)\ versus \<^term>\cardinal\)\ lemma eqpoll_Aleph1_cardinal_vimage: assumes "X \ \\<^bsub>1\<^esub>" "f : X \ \" shows "\n\\. |f-``{n}| = \\<^bsub>1\<^esub>" proof - from assms obtain g where "g\bij(\\<^bsub>1\<^esub>,X)" using eqpoll_sym by blast with \f : X \ \\ have "f O g : \\<^bsub>1\<^esub> \ \" "converse(g) \ bij(X, \\<^bsub>1\<^esub>)" using bij_is_fun comp_fun bij_converse_bij by blast+ then obtain n where "n\\" "|(f O g)-``{n}| = \\<^bsub>1\<^esub>" using Aleph1_eq_cardinal_vimage by auto then have "\\<^bsub>1\<^esub> = |converse(g) `` (f -``{n})|" using image_comp converse_comp unfolding vimage_def by simp also from \converse(g) \ bij(X, \\<^bsub>1\<^esub>)\ \f: X\ \\ have "\ = |f -``{n}|" using range_of_subset_eqpoll[of "converse(g)" X _ "f -``{n}"] bij_is_inj cardinal_cong bij_is_fun eqpoll_sym Pi_vimage_subset by fastforce finally show ?thesis using \n\\\ by auto qed subsection\Applications of transfinite recursive constructions\ text\The next lemma is an application of recursive constructions. It works under the assumption that whenever the already constructed subsequence is small enough, another element can be added.\ lemma bounded_cardinal_selection: includes Ord_dests assumes "\X. |X| < \ \ X \ G \ \a\G. \s\X. Q(s,a)" "b\G" "Card(\)" shows "\S. S : \ \ G \ (\\ \ \. \\ \ \. \<\ \ Q(S`\,S`\))" proof - let ?cdlt\="{X\Pow(G) . |X|<\}" \ \``cardinal less than \<^term>\\\''\ and ?inQ="\Y.{a\G. \s\Y. Q(s,a)}" from assms have "\Y \ ?cdlt\. \a. a \ ?inQ(Y)" by blast then have "\f. f \ Pi(?cdlt\,?inQ)" using AC_ball_Pi[of ?cdlt\ ?inQ] by simp then obtain f where f_type:"f \ Pi(?cdlt\,?inQ)" by auto moreover define Cb where "Cb \ \_\Pow(G)-?cdlt\. b" moreover from \b\G\ have "Cb \ Pow(G)-?cdlt\ \ G" unfolding Cb_def by simp moreover note \Card(\)\ ultimately have "f \ Cb : (\x\Pow(G). ?inQ(x) \ G)" using fun_Pi_disjoint_Un[ of f ?cdlt\ ?inQ Cb "Pow(G)-?cdlt\" "\_.G"] Diff_partition[of "{X\Pow(G). |X|<\}" "Pow(G)", OF Collect_subset] by auto moreover have "?inQ(x) \ G = G" for x by auto ultimately have "f \ Cb : Pow(G) \ G" by simp define S where "S\\\\\. rec_constr(f \ Cb, \)" from \f \ Cb: Pow(G) \ G\ \Card(\)\ have "S : \ \ G" using Ord_in_Ord unfolding S_def by (intro lam_type rec_constr_type) auto moreover have "\\\\. \\\\. \ < \ \ Q(S ` \, S ` \)" proof (intro ballI impI) fix \ \ assume "\\\" with \Card(\)\ have "{rec_constr(f \ Cb, x) . x\\} = {S`x . x \ \}" using Ord_trans[OF _ _ Card_is_Ord, of _ \ \] unfolding S_def by auto moreover from \\\\\ \S : \ \ G\ \Card(\)\ have "{S`x . x \ \} \ G" using Ord_trans[OF _ _ Card_is_Ord, of _ \ \] apply_type[of S \ "\_. G"] by auto moreover from \Card(\)\ \\\\\ have "|{S`x . x \ \}| < \" using cardinal_RepFun_le[of \] Ord_in_Ord lt_trans1[of "|{S`x . x \ \}|" "|\|" \] Card_lt_iff[THEN iffD2, of \ \, OF _ _ ltI] by force moreover have "\x\\. Q(S`x, f ` {S`x . x \ \})" proof - from calculation and f_type have "f ` {S`x . x \ \} \ {a\G. \x\\. Q(S`x,a)}" using apply_type[of f ?cdlt\ ?inQ "{S`x . x \ \}"] by blast then show ?thesis by simp qed moreover assume "\\\" "\ < \" moreover note \\\\\ \Cb \ Pow(G)-?cdlt\ \ G\ ultimately show "Q(S ` \, S ` \)" using fun_disjoint_apply1[of "{S`x . x \ \}" Cb f] domain_of_fun[of Cb] ltD[of \ \] by (subst (2) S_def, auto) (subst rec_constr_unfold, auto) qed ultimately show ?thesis by blast qed text\The following basic result can, in turn, be proved by a bounded-cardinal selection.\ lemma Infinite_iff_lepoll_nat: "Infinite(X) \ \ \ X" proof assume "Infinite(X)" then obtain b where "b\X" using Infinite_not_empty by auto { fix Y assume "|Y| < \" then have "Finite(Y)" using Finite_cardinal_iff' ltD nat_into_Finite by blast with \Infinite(X)\ have "X \ Y" by auto } with \b\X\ obtain S where "S : \ \ X" "\\\\. \\\\. \ < \ \ S`\ \ S`\" using bounded_cardinal_selection[of \ X "\x y. x\y"] Card_nat by blast moreover from this have "\ \ \ \ \ \ \ \ \\\ \ S`\ \ S`\" for \ \ by (rule_tac lt_neq_symmetry[of "\" "\\ \. S`\ \ S`\"]) auto ultimately show "\ \ X" unfolding lepoll_def inj_def by blast qed (intro lepoll_nat_imp_Infinite) lemma Infinite_InfCard_cardinal: "Infinite(X) \ InfCard(|X|)" using lepoll_eq_trans eqpoll_sym lepoll_nat_imp_Infinite Infinite_iff_lepoll_nat Inf_Card_is_InfCard cardinal_eqpoll by simp lemma Finite_to_one_surj_imp_cardinal_eq: assumes "F \ Finite_to_one(X,Y) \ surj(X,Y)" "Infinite(X)" shows "|Y| = |X|" proof - from \F \ Finite_to_one(X,Y) \ surj(X,Y)\ have "X = (\y\Y. {x\X . F`x = y})" using apply_type by fastforce show ?thesis proof (cases "Finite(Y)") case True with \X = (\y\Y. {x\X . F`x = y})\ and assms show ?thesis using Finite_RepFun[THEN [2] Finite_Union, of Y "\y. {x\X . F`x = y}"] by auto next case False moreover from this have "Y \ |Y|" using cardinal_eqpoll eqpoll_sym eqpoll_imp_lepoll by simp moreover note assms moreover from calculation have "y \ Y \ |{x\X . F`x = y}| \ |Y|" for y using Infinite_imp_nats_lepoll[THEN lepoll_imp_cardinal_le, of Y "|{x\X . F`x = y}|"] cardinal_idem by auto ultimately have "|\y\Y. {x\X . F`x = y}| \ |Y|" using lepoll_imp_cardinal_UN_le[of "|Y|" Y] Infinite_InfCard_cardinal[of Y] by simp moreover from \F \ Finite_to_one(X,Y) \ surj(X,Y)\ have "|Y| \ |X|" using surj_implies_cardinal_le by auto moreover note \X = (\y\Y. {x\X . F`x = y})\ ultimately show ?thesis using le_anti_sym by auto qed qed lemma cardinal_map_Un: assumes "Infinite(X)" "Finite(b)" shows "|{a \ b . a \ X}| = |X|" proof - have "(\a\X. a \ b) \ Finite_to_one(X,{a \ b . a \ X})" "(\a\X. a \ b) \ surj(X,{a \ b . a \ X})" unfolding surj_def proof fix d have "Finite({a \ X . a \ b = d})" (is "Finite(?Y(b,d))") using \Finite(b)\ proof (induct arbitrary:d) case 0 have "{a \ X . a \ 0 = d} = (if d\X then {d} else 0)" by auto then show ?case by simp next case (cons c b) from \c \ b\ have "?Y(cons(c,b),d) \ (if c\d then ?Y(b,d) \ ?Y(b,d-{c}) else 0)" by auto with cons show ?case using subset_Finite by simp qed moreover assume "d \ {x \ b . x \ X}" ultimately show "Finite({a \ X . (\x\X. x \ b) ` a = d})" by simp qed (auto intro:lam_funtype) with assms show ?thesis using Finite_to_one_surj_imp_cardinal_eq by fast qed end \ No newline at end of file diff --git a/thys/Delta_System_Lemma/ZF_Library.thy b/thys/Delta_System_Lemma/ZF_Library.thy --- a/thys/Delta_System_Lemma/ZF_Library.thy +++ b/thys/Delta_System_Lemma/ZF_Library.thy @@ -1,1064 +1,1079 @@ section\Library of basic $\ZF$ results\label{sec:zf-lib}\ theory ZF_Library imports "ZF-Constructible.Normal" begin text\This theory gathers basic ``combinatorial'' results that can be proved in $\ZF$ (that is, without using the Axiom of Choice $\AC$). \ text\We begin by setting up math-friendly notation.\ no_notation oadd (infixl \++\ 65) no_notation sum (infixr \+\ 65) notation oadd (infixl \+\ 65) notation nat (\\\) notation csucc (\_\<^sup>+\ [90]) no_notation Aleph (\\_\ [90] 90) notation Aleph (\\\<^bsub>_\<^esub>\) syntax "_ge" :: "[i,i] \ o" (infixl \\\ 50) translations "x \ y" \ "y \ x" subsection\Some minimal arithmetic/ordinal stuff\ lemma Un_leD1 : "i \ j \ k \ Ord(i) \ Ord(j) \ Ord(k) \ i \ k" by (rule Un_least_lt_iff[THEN iffD1[THEN conjunct1]],simp_all) lemma Un_leD2 : "i \ j \ k \ Ord(i) \ Ord(j) \ Ord(k) \ j \ k" by (rule Un_least_lt_iff[THEN iffD1[THEN conjunct2]],simp_all) lemma Un_memD1: "i \ j \ k \ Ord(i) \ Ord(j) \ Ord(k) \ i \ k" by (drule ltI, assumption, drule leI, rule Un_least_lt_iff[THEN iffD1[THEN conjunct1]],simp_all) lemma Un_memD2 : "i \ j \ k \ Ord(i) \ Ord(j) \ Ord(k) \ j \ k" by (drule ltI, assumption, drule leI, rule Un_least_lt_iff[THEN iffD1[THEN conjunct2]],simp_all) text\This lemma allows to apply arithmetic simprocs to ordinal addition\ lemma nat_oadd_add[simp]: assumes "m \ \" "n \ \" shows "n + m = n #+ m" using assms by induct simp_all lemma Ord_has_max_imp_succ: assumes "Ord(\)" "\ \ \" "\\\\. \ \ \" shows "\ = succ(\)" using assms Ord_trans[of _ \ \] unfolding lt_def by (intro equalityI subsetI) auto lemma Least_antitone: assumes "Ord(j)" "P(j)" "\i. P(i) \ Q(i)" shows "(\ i. Q(i)) \ (\ i. P(i))" using assms LeastI2[of P j Q] Least_le by simp lemma Least_set_antitone: "Ord(j) \ j\A \ A \ B \ (\ i. i\B) \ (\ i. i\A)" using subset_iff by (auto intro:Least_antitone) lemma le_neq_imp_lt: "x\y \ x\y \ xStrict upper bound of a set of ordinals.\ definition str_bound :: "i\i" where "str_bound(A) \ \a\A. succ(a)" lemma str_bound_type [TC]: "\a\A. Ord(a) \ Ord(str_bound(A))" unfolding str_bound_def by auto lemma str_bound_lt: "\a\A. Ord(a) \ \a\A. a < str_bound(A)" unfolding str_bound_def using str_bound_type by (blast intro:ltI) lemma naturals_lt_nat[intro]: "n \ \ \ n < \" unfolding lt_def by simp text\The next two lemmas are handy when one is constructing some object recursively. The first handles injectivity (of recursively constructed sequences of sets), while the second is helpful for establishing a symmetry argument.\ lemma Int_eq_zero_imp_not_eq: assumes "\x y. x\D \ y \ D \ x \ y \ A(x) \ A(y) = 0" "\x. x\D \ A(x) \ 0" "a\D" "b\D" "a\b" shows "A(a) \ A(b)" using assms by fastforce lemma lt_neq_symmetry: assumes "\\ \. \ \ \ \ \ \ \ \ \ < \ \ Q(\,\)" "\\ \. Q(\,\) \ Q(\,\)" "\ \ \" "\ \ \" "\ \ \" "Ord(\)" shows "Q(\,\)" proof - from assms consider "\<\" | "\<\" using Ord_linear_lt[of \ \ thesis] Ord_in_Ord[of \] by auto then show ?thesis by cases (auto simp add:assms) qed lemma cardinal_succ_not_0: "|A| = succ(n) \ A \ 0" by auto lemma Ord_eq_Collect_lt: "i<\ \ {j\\. j \almost the same proof as @{thm nat_eq_Collect_lt}\ apply (rule equalityI) apply (blast dest: ltD) apply (auto simp add: Ord_mem_iff_lt) apply (rule Ord_trans ltI[OF _ lt_Ord]; auto simp add:lt_def dest:ltD)+ done subsection\Manipulation of function spaces\ definition Finite_to_one :: "[i,i] \ i" where "Finite_to_one(X,Y) \ {f:X\Y. \y\Y. Finite({x\X . f`x = y})}" lemma Finite_to_oneI[intro]: assumes "f:X\Y" "\y. y\Y \ Finite({x\X . f`x = y})" shows "f \ Finite_to_one(X,Y)" using assms unfolding Finite_to_one_def by simp lemma Finite_to_oneD[dest]: "f \ Finite_to_one(X,Y) \ f:X\Y" "f \ Finite_to_one(X,Y) \ y\Y \ Finite({x\X . f`x = y})" unfolding Finite_to_one_def by simp_all lemma subset_Diff_Un: "X \ A \ A = (A - X) \ X " by auto lemma Diff_bij: assumes "\A\F. X \ A" shows "(\A\F. A-X) \ bij(F, {A-X. A\F})" using assms unfolding bij_def inj_def surj_def by (auto intro:lam_type, subst subset_Diff_Un[of X]) auto lemma function_space_nonempty: assumes "b\B" shows "(\x\A. b) : A \ B" using assms lam_type by force lemma vimage_lam: "(\x\A. f(x)) -`` B = { x\A . f(x) \ B }" using lam_funtype[of A f, THEN [2] domain_type] lam_funtype[of A f, THEN [2] apply_equality] lamI[of _ A f] by auto blast lemma range_fun_subset_codomain: assumes "h:B \ C" shows "range(h) \ C" unfolding range_def domain_def converse_def using range_type[OF _ assms] by auto lemma Pi_rangeD: assumes "f\Pi(A,B)" "b \ range(f)" shows "\a\A. f`a = b" using assms apply_equality[OF _ assms(1), of _ b] domain_type[OF _ assms(1)] by auto lemma Pi_range_eq: "f \ Pi(A,B) \ range(f) = {f ` x . x \ A}" using Pi_rangeD[of f A B] apply_rangeI[of f A B] by blast lemma Pi_vimage_subset : "f \ Pi(A,B) \ f-``C \ A" unfolding Pi_def by auto lemma apply_in_codomain_Ord: assumes "Ord(\)" "\\0" "f: A \ \" shows "f`x\\" proof (cases "x\A") case True from assms \x\A\ show ?thesis using domain_of_fun apply_rangeI by simp next case False from assms \x\A\ show ?thesis using apply_0 Ord_0_lt ltD domain_of_fun by auto qed lemma range_eq_image: assumes "f:A\B" shows "range(f) = f``A" proof show "f `` A \ range(f)" unfolding image_def by blast { fix x assume "x\range(f)" with assms have "x\f``A" using domain_of_fun[of f A "\_. B"] by auto } then show "range(f) \ f `` A" .. qed lemma Image_sub_codomain: "f:A\B \ f``C \ B" using image_subset fun_is_rel[of _ _ "\_ . B"] by force lemma inj_to_Image: assumes "f:A\B" "f \ inj(A,B)" shows "f \ inj(A,f``A)" using assms inj_inj_range range_eq_image by force lemma inj_imp_surj: fixes f b notes inj_is_fun[dest] defines [simp]: "ifx(x) \ if x\range(f) then converse(f)`x else b" assumes "f \ inj(B,A)" "b\B" shows "(\x\A. ifx(x)) \ surj(A,B)" proof - from assms have "converse(f) \ surj(range(f),B)" "range(f) \ A" "converse(f) : range(f) \ B" using inj_converse_surj range_fun_subset_codomain surj_is_fun by blast+ with \b\B\ show "(\x\A. ifx(x)) \ surj(A,B)" unfolding surj_def proof (intro CollectI lam_type ballI; elim CollectE) fix y assume "y \ B" "\y\B. \x\range(f). converse(f) ` x = y" with \range(f) \ A\ show "\x\A. (\x\A. ifx(x)) ` x = y" by (drule_tac bspec, auto) qed simp qed lemma fun_Pi_disjoint_Un: assumes "f \ Pi(A,B)" "g \ Pi(C,D)" "A \ C = 0" shows "f \ g \ Pi(A \ C, \x. B(x) \ D(x))" using assms by (simp add: Pi_iff extension Un_rls) (unfold function_def, blast) lemma Un_restrict_decomposition: assumes "f \ Pi(A,B)" shows "f = restrict(f, A \ C) \ restrict(f, A - C)" using assms proof (rule fun_extension) from assms have "restrict(f,A\C) \ restrict(f,A-C) \ Pi(A\C \ (A-C), \x. B(x)\B(x))" using restrict_type2[of f A B] by (rule_tac fun_Pi_disjoint_Un) force+ moreover have "(A \ C) \ (A - C) = A" by auto ultimately show "restrict(f, A \ C) \ restrict(f, A - C) \ Pi(A, B)" by simp next fix x assume "x \ A" with assms show "f ` x = (restrict(f, A \ C) \ restrict(f, A - C)) ` x" using restrict fun_disjoint_apply1[of _ "restrict(f,_)"] fun_disjoint_apply2[of _ "restrict(f,_)"] domain_restrict[of f] apply_0 domain_of_fun by (cases "x\C") simp_all qed lemma restrict_eq_imp_Un_into_Pi: assumes "f \ Pi(A,B)" "g \ Pi(C,D)" "restrict(f, A \ C) = restrict(g, A \ C)" shows "f \ g \ Pi(A \ C, \x. B(x) \ D(x))" proof - note assms moreover from this have "x \ g \ x \ restrict(g, A \ C)" for x using restrict_subset[of g "A \ C"] by auto moreover from calculation have "x \ f \ x \ restrict(f, A - C) \ x \ restrict(g, A \ C)" for x by (subst (asm) Un_restrict_decomposition[of f A B "C"]) auto ultimately have "f \ g = restrict(f, A - C) \ g" using restrict_subset[of g "A \ C"] by (subst Un_restrict_decomposition[of f A B "C"]) auto moreover have "A - C \ C = A \ C" by auto moreover note assms ultimately show ?thesis using fun_Pi_disjoint_Un[OF restrict_type2[of f A B "A-C"], of g C D] by auto qed lemma restrict_eq_imp_Un_into_Pi': assumes "f \ Pi(A,B)" "g \ Pi(C,D)" "restrict(f, domain(f) \ domain(g)) = restrict(g, domain(f) \ domain(g))" shows "f \ g \ Pi(A \ C, \x. B(x) \ D(x))" using assms domain_of_fun restrict_eq_imp_Un_into_Pi by simp lemma restrict_subset_Sigma: "f \ Sigma(C,B) \ restrict(f,A) \ Sigma(A\C, B)" by (auto simp add: restrict_def) subsection\Finite sets\ lemma Replace_sing1: "\ (\a. P(d,a)) \ (\y y'. P(d,y) \ P(d,y') \ y=y') \ \ \a. {y . x \ {d}, P(x,y)} = {a}" by blast \ \Not really necessary\ lemma Replace_sing2: assumes "\a. \ P(d,a)" shows "{y . x \ {d}, P(x,y)} = 0" using assms by auto lemma Replace_sing3: assumes "\c e. c \ e \ P(d,c) \ P(d,e)" shows "{y . x \ {d}, P(x,y)} = 0" proof - { fix z { assume "\y. P(d, y) \ y = z" with assms have "False" by auto } then have "z \ {y . x \ {d}, P(x,y)}" using Replace_iff by auto } then show ?thesis by (intro equalityI subsetI) simp_all qed lemma Replace_Un: "{b . a \ A \ B, Q(a, b)} = {b . a \ A, Q(a, b)} \ {b . a \ B, Q(a, b)}" by (intro equalityI subsetI) (auto simp add:Replace_iff) lemma Replace_subset_sing: "\z. {y . x \ {d}, P(x,y)} \ {z}" proof - consider (1) "(\a. P(d,a)) \ (\y y'. P(d,y) \ P(d,y') \ y=y')" | (2) "\a. \ P(d,a)" | (3) "\c e. c \ e \ P(d,c) \ P(d,e)" by auto then show "\z. {y . x \ {d}, P(x,y)} \ {z}" proof (cases) case 1 then show ?thesis using Replace_sing1[of P d] by auto next case 2 then show ?thesis by auto next case 3 then show ?thesis using Replace_sing3[of P d] by auto qed qed lemma Finite_Replace: "Finite(A) \ Finite(Replace(A,Q))" proof (induct rule:Finite_induct) case 0 then show ?case by simp next case (cons x B) moreover have "{b . a \ cons(x, B), Q(a, b)} = {b . a \ B, Q(a, b)} \ {b . a \ {x}, Q(a, b)}" using Replace_Un unfolding cons_def by auto moreover obtain d where "{b . a \ {x}, Q(a, b)} \ {d}" using Replace_subset_sing[of _ Q] by blast moreover from this have "Finite({b . a \ {x}, Q(a, b)})" using subset_Finite by simp ultimately show ?case using subset_Finite by simp qed lemma Finite_domain: "Finite(A) \ Finite(domain(A))" using Finite_Replace unfolding domain_def by auto lemma Finite_converse: "Finite(A) \ Finite(converse(A))" using Finite_Replace unfolding converse_def by auto lemma Finite_range: "Finite(A) \ Finite(range(A))" using Finite_domain Finite_converse unfolding range_def by blast lemma Finite_Sigma: "Finite(A) \ \x. Finite(B(x)) \ Finite(Sigma(A,B))" unfolding Sigma_def using Finite_RepFun Finite_Union by simp lemma Finite_Pi: "Finite(A) \ \x. Finite(B(x)) \ Finite(Pi(A,B))" using Finite_Sigma Finite_Pow subset_Finite[of "Pi(A,B)" "Pow(Sigma(A,B))"] unfolding Pi_def by auto subsection\Basic results on equipollence, cardinality and related concepts\ lemma lepollD[dest]: "A \ B \ \f. f \ inj(A, B)" unfolding lepoll_def . lemma lepollI[intro]: "f \ inj(A, B) \ A \ B" unfolding lepoll_def by blast lemma eqpollD[dest]: "A \ B \ \f. f \ bij(A, B)" unfolding eqpoll_def . declare bij_imp_eqpoll[intro] lemma range_of_subset_eqpoll: assumes "f \ inj(X,Y)" "S \ X" shows "S \ f `` S" using assms restrict_bij by blast text\I thank Miguel Pagano for this proof.\ lemma function_space_eqpoll_cong: assumes "A \ A'" "B \ B'" shows "A \ B \ A' \ B'" proof - from assms(1)[THEN eqpoll_sym] assms(2) obtain f g where "f \ bij(A',A)" "g \ bij(B,B')" by blast then have "converse(g) : B' \ B" "converse(f): A \ A'" using bij_converse_bij bij_is_fun by auto show ?thesis unfolding eqpoll_def proof (intro exI fg_imp_bijective, rule_tac [1-2] lam_type) fix F assume "F: A \ B" with \f \ bij(A',A)\ \g \ bij(B,B')\ show "g O F O f : A' \ B'" using bij_is_fun comp_fun by blast next fix F assume "F: A' \ B'" with \converse(g) : B' \ B\ \converse(f): A \ A'\ show "converse(g) O F O converse(f) : A \ B" using comp_fun by blast next from \f\_\ \g\_\ \converse(f)\_\ \converse(g)\_\ have "(\x. x \ A' \ B' \ converse(g) O x O converse(f) \ A \ B)" using bij_is_fun comp_fun by blast then have "(\x\A \ B. g O x O f) O (\x\A' \ B'. converse(g) O x O converse(f)) = (\x\A' \ B' . (g O converse(g)) O x O (converse(f) O f))" using lam_cong comp_assoc comp_lam[of "A' \ B'" ] by auto also have "... = (\x\A' \ B' . id(B') O x O (id(A')))" using left_comp_inverse[OF bij_is_inj[OF \f\_\]] right_comp_inverse[OF bij_is_surj[OF \g\_\]] by auto also have "... = (\x\A' \ B' . x)" using left_comp_id[OF fun_is_rel] right_comp_id[OF fun_is_rel] lam_cong by auto also have "... = id(A'\B')" unfolding id_def by simp finally show "(\x\A -> B. g O x O f) O (\x\A' -> B'. converse(g) O x O converse(f)) = id(A' -> B')" . next from \f\_\ \g\_\ have "(\x. x \ A \ B \ g O x O f \ A' \ B')" using bij_is_fun comp_fun by blast then have "(\x\A' -> B'. converse(g) O x O converse(f)) O (\x\A -> B. g O x O f) = (\x\A \ B . (converse(g) O g) O x O (f O converse(f)))" using comp_lam comp_assoc by auto also have "... = (\x\A \ B . id(B) O x O (id(A)))" using right_comp_inverse[OF bij_is_surj[OF \f\_\]] left_comp_inverse[OF bij_is_inj[OF \g\_\]] lam_cong by auto also have "... = (\x\A \ B . x)" using left_comp_id[OF fun_is_rel] right_comp_id[OF fun_is_rel] lam_cong by auto also have "... = id(A\B)" unfolding id_def by simp finally show "(\x\A' \ B'. converse(g) O x O converse(f)) O (\x\A -> B. g O x O f) = id(A -> B)" . qed qed lemma curry_eqpoll: fixes d \1 \2 \ shows "\1 \ \2 \ \ \ \1 \ \2 \ \" unfolding eqpoll_def proof (intro exI, rule lam_bijective, rule_tac [1-2] lam_type, rule_tac [2] lam_type) fix f z assume "f : \1 \ \2 \ \" "z \ \1 \ \2" then show "f`fst(z)`snd(z) \ \" by simp next fix f x y assume "f : \1 \ \2 \ \" "x\\1" "y\\2" then show "f`\x,y\ \ \" by simp next \ \one composition is the identity:\ fix f assume "f : \1 \ \2 \ \" then show "(\x\\1 \ \2. (\x\\1. \xa\\2. f ` \x, xa\) ` fst(x) ` snd(x)) = f" by (auto intro:fun_extension) qed simp \ \the other composition follows automatically\ lemma Pow_eqpoll_function_space: fixes d X notes bool_of_o_def [simp] defines [simp]:"d(A) \ (\x\X. bool_of_o(x\A))" \ \the witnessing map for the thesis:\ shows "Pow(X) \ X \ 2" unfolding eqpoll_def proof (intro exI, rule lam_bijective) \ \We give explicit mutual inverses\ fix A assume "A\Pow(X)" then show "d(A) : X \ 2" using lam_type[of _ "\x. bool_of_o(x\A)" "\_. 2"] by force from \A\Pow(X)\ show "{y\X. d(A)`y = 1} = A" by (auto) next fix f assume "f: X\2" then show "d({y \ X . f ` y = 1}) = f" using apply_type[OF \f: X\2\] by (force intro:fun_extension) qed blast lemma cantor_inj: "f \ inj(Pow(A),A)" using inj_imp_surj[OF _ Pow_bottom] cantor_surj by blast definition cexp :: "[i,i] \ i" ("_\<^bsup>\_\<^esup>" [76,1] 75) where "\\<^bsup>\\\<^esup> \ |\ \ \|" lemma Card_cexp: "Card(\\<^bsup>\\\<^esup>)" unfolding cexp_def Card_cardinal by simp lemma eq_csucc_ord: "Ord(i) \ i\<^sup>+ = |i|\<^sup>+" using Card_lt_iff Least_cong unfolding csucc_def by auto text\I thank Miguel Pagano for this proof.\ lemma lesspoll_csucc: assumes "Ord(\)" shows "d \ \\<^sup>+ \ d \ \" proof assume "d \ \\<^sup>+" moreover note Card_is_Ord \Ord(\)\ moreover from calculation have "\ < \\<^sup>+" "Card(\\<^sup>+)" using Ord_cardinal_eqpoll csucc_basic by simp_all moreover from calculation have "d \ |\|\<^sup>+" "Card(|\|)" "d \ |d|" using eq_csucc_ord[of \] lesspoll_imp_eqpoll eqpoll_sym by simp_all moreover from calculation have "|d| < |\|\<^sup>+" using lesspoll_cardinal_lt csucc_basic by simp moreover from calculation have "|d| \ |\|" using Card_lt_csucc_iff le_imp_lepoll by simp moreover from calculation have "|d| \ \" using lepoll_eq_trans Ord_cardinal_eqpoll by simp ultimately show "d \ \" using eq_lepoll_trans by simp next from \Ord(\)\ have "\ < \\<^sup>+" "Card(\\<^sup>+)" using csucc_basic by simp_all moreover assume "d \ \" ultimately have "d \ \\<^sup>+" using le_imp_lepoll leI lepoll_trans by simp moreover from \d \ \\ \Ord(\)\ have "\\<^sup>+ \ \" if "d \ \\<^sup>+" using eqpoll_sym[OF that] eq_lepoll_trans[OF _ \d\\\] by simp moreover from calculation \Card(_)\ have "\ d \ \\<^sup>+" using lesspoll_irrefl lesspoll_trans1 lt_Card_imp_lesspoll[OF _ \\ <_\] by auto ultimately show "d \ \\<^sup>+" unfolding lesspoll_def by simp qed abbreviation Infinite :: "i\o" where "Infinite(X) \ \ Finite(X)" lemma Infinite_not_empty: "Infinite(X) \ X \ 0" using empty_lepollI by auto lemma Infinite_imp_nats_lepoll: assumes "Infinite(X)" "n \ \" shows "n \ X" using \n \ \\ proof (induct) case 0 then show ?case using empty_lepollI by simp next case (succ x) show ?case proof - from \Infinite(X)\ and \x \ \\ have "\ (x \ X)" using eqpoll_sym unfolding Finite_def by auto with \x \ X\ obtain f where "f \ inj(x,X)" "f \ surj(x,X)" unfolding bij_def eqpoll_def by auto moreover from this obtain b where "b \ X" "\a\x. f`a \ b" using inj_is_fun unfolding surj_def by auto ultimately have "f \ inj(x,X-{b})" unfolding inj_def by (auto intro:Pi_type) then have "cons(\x, b\, f) \ inj(succ(x), cons(b, X - {b}))" using inj_extend[of f x "X-{b}" x b] unfolding succ_def by (auto dest:mem_irrefl) moreover from \b\X\ have "cons(b, X - {b}) = X" by auto ultimately show "succ(x) \ X" by auto qed qed lemma zero_lesspoll: assumes "0<\" shows "0 \ \" using assms eqpoll_0_iff[THEN iffD1, of \] eqpoll_sym unfolding lesspoll_def lepoll_def by (auto simp add:inj_def) lemma lepoll_nat_imp_Infinite: "\ \ X \ Infinite(X)" proof (rule ccontr, simp) assume "\ \ X" "Finite(X)" moreover from this obtain n where "X \ n" "n \ \" unfolding Finite_def by auto moreover from calculation have "\ \ n" using lepoll_eq_trans by simp ultimately show "False" using lepoll_nat_imp_Finite nat_not_Finite by simp qed lemma InfCard_imp_Infinite: "InfCard(\) \ Infinite(\)" using le_imp_lepoll[THEN lepoll_nat_imp_Infinite, of \] unfolding InfCard_def by simp lemma lt_surj_empty_imp_Card: assumes "Ord(\)" "\\. \ < \ \ surj(\,\) = 0" shows "Card(\)" proof - { assume "|\| < \" with assms have "False" using LeastI[of "\i. i \ \" \, OF eqpoll_refl] Least_le[of "\i. i \ \" "|\|", OF Ord_cardinal_eqpoll] unfolding Card_def cardinal_def eqpoll_def bij_def by simp } with assms show ?thesis using Ord_cardinal_le[of \] not_lt_imp_le[of "|\|" \] le_anti_sym unfolding Card_def by auto qed subsection\Morphisms of binary relations\ text\The main case of interest is in the case of partial orders.\ lemma mono_map_mono: assumes "f \ mono_map(A,r,B,s)" "B \ C" shows "f \ mono_map(A,r,C,s)" unfolding mono_map_def proof (intro CollectI ballI impI) from \f \ mono_map(A,_,B,_)\ have "f: A \ B" using mono_map_is_fun by simp with \B\C\ show "f: A \ C" using fun_weaken_type by simp fix x y assume "x\A" "y\A" "\x,y\ \ r" moreover from this and \f: A \ B\ have "f`x \ B" "f`y \ B" using apply_type by simp_all moreover note \f \ mono_map(_,r,_,s)\ ultimately show "\f ` x, f ` y\ \ s" unfolding mono_map_def by blast qed lemma ordertype_zero_imp_zero: "ordertype(A,r) = 0 \ A = 0" using ordermap_type[of A r] by (cases "A=0") auto lemma mono_map_increasing: "j\mono_map(A,r,B,s) \ a\A \ c\A \ \a,c\\r \ \j`a,j`c\\s" unfolding mono_map_def by simp lemma linear_mono_map_reflects: assumes "linear(\,r)" "trans[\](s)" "irrefl(\,s)" "f\mono_map(\,r,\,s)" "x\\" "y\\" "\f`x,f`y\\s" shows "\x,y\\r" proof - from \f\mono_map(_,_,_,_)\ have preserves:"x\\ \ y\\ \ \x,y\\r \ \f`x,f`y\\s" for x y unfolding mono_map_def by blast { assume "\x, y\ \ r" "x\\" "y\\" moreover note \\f`x,f`y\\s\ and \linear(\,r)\ moreover from calculation have "y = x \ \y,x\\r" unfolding linear_def by blast moreover note preserves [of y x] ultimately have "y = x \ \f`y, f`x\\ s" by blast moreover from \f\mono_map(_,_,\,_)\ \x\\\ \y\\\ have "f`x\\" "f`y\\" using apply_type[OF mono_map_is_fun] by simp_all moreover note \\f`x,f`y\\s\ \trans[\](s)\ \irrefl(\,s)\ ultimately have "False" using trans_onD[of \ s "f`x" "f`y" "f`x"] irreflE by blast } with assms show "\x,y\\r" by blast qed lemma irrefl_Memrel: "irrefl(x, Memrel(x))" unfolding irrefl_def using mem_irrefl by auto lemmas Memrel_mono_map_reflects = linear_mono_map_reflects [OF well_ord_is_linear[OF well_ord_Memrel] well_ord_is_trans_on[OF well_ord_Memrel] irrefl_Memrel] \ \Same proof as Paulson's @{thm mono_map_is_inj}\ lemma mono_map_is_inj': "\ linear(A,r); irrefl(B,s); f \ mono_map(A,r,B,s) \ \ f \ inj(A,B)" unfolding irrefl_def mono_map_def inj_def using linearE by (clarify, rename_tac x w) (erule_tac x=w and y=x in linearE, assumption+, (force intro: apply_type)+) lemma mono_map_imp_ord_iso_image: assumes "linear(\,r)" "trans[\](s)" "irrefl(\,s)" "f\mono_map(\,r,\,s)" shows "f \ ord_iso(\,r,f``\,s)" unfolding ord_iso_def proof (intro CollectI ballI iffI) \ \Enough to show it's bijective and preserves both ways\ from assms have "f \ inj(\,\)" using mono_map_is_inj' by blast moreover from \f \ mono_map(_,_,_,_)\ have "f \ surj(\, f``\)" unfolding mono_map_def using surj_image by auto ultimately show "f \ bij(\, f``\)" unfolding bij_def using inj_is_fun inj_to_Image by simp from \f\mono_map(_,_,_,_)\ show "x\\ \ y\\ \ \x,y\\r \ \f`x,f`y\\s" for x y unfolding mono_map_def by blast with assms show "\f`x,f`y\\s \ x\\ \ y\\ \ \x,y\\r" for x y using linear_mono_map_reflects by blast qed text\We introduce the following notation for strictly increasing maps between ordinals.\ abbreviation mono_map_Memrel :: "[i,i] \ i" (infixr \\\<^sub><\ 60) where "\ \\<^sub>< \ \ mono_map(\,Memrel(\),\,Memrel(\))" lemma mono_map_imp_ord_iso_Memrel: assumes "Ord(\)" "Ord(\)" "f:\ \\<^sub>< \" shows "f \ ord_iso(\,Memrel(\),f``\,Memrel(\))" using assms mono_map_imp_ord_iso_image[OF well_ord_is_linear[OF well_ord_Memrel] well_ord_is_trans_on[OF well_ord_Memrel] irrefl_Memrel] by blast lemma mono_map_ordertype_image': assumes "X\\" "Ord(\)" "Ord(\)" "f \ mono_map(X,Memrel(\),\,Memrel(\))" shows "ordertype(f``X,Memrel(\)) = ordertype(X,Memrel(\))" using assms mono_map_is_fun[of f X _ \] ordertype_eq mono_map_imp_ord_iso_image[OF well_ord_is_linear[OF well_ord_Memrel, THEN linear_subset] well_ord_is_trans_on[OF well_ord_Memrel] irrefl_Memrel, of \ X \ f] well_ord_subset[OF well_ord_Memrel] Image_sub_codomain[of f X \ X] by auto lemma mono_map_ordertype_image: assumes "Ord(\)" "Ord(\)" "f:\ \\<^sub>< \" shows "ordertype(f``\,Memrel(\)) = \" using assms mono_map_is_fun ordertype_Memrel ordertype_eq[of f \ "Memrel(\)"] mono_map_imp_ord_iso_Memrel well_ord_subset[OF well_ord_Memrel] Image_sub_codomain[of _ \] by auto lemma apply_in_image: "f:A\B \ a\A \ f`a \ f``A" using range_eq_image apply_rangeI[of f] by simp lemma Image_subset_Ord_imp_lt: assumes "Ord(\)" "h``A \ \" "x\domain(h)" "x\A" "function(h)" shows "h`x < \" using assms unfolding domain_def using imageI ltI function_apply_equality by auto lemma ordermap_le_arg: assumes "X\\" "x\X" "Ord(\)" shows "x\X \ ordermap(X,Memrel(\))`x\x" proof (induct rule:Ord_induct[OF subsetD, OF assms]) case (1 x) have "wf[X](Memrel(\))" using wf_imp_wf_on[OF wf_Memrel] . with 1 have "ordermap(X,Memrel(\))`x = {ordermap(X,Memrel(\))`y . y\{y\X . y\x \ y\\}}" using ordermap_unfold Ord_trans[of _ x \] by auto also from assms have "... = {ordermap(X,Memrel(\))`y . y\{y\X . y\x}}" using Ord_trans[of _ x \] Ord_in_Ord by blast finally have ordm:"ordermap(X,Memrel(\))`x = {ordermap(X,Memrel(\))`y . y\{y\X . y\x}}" . from 1 have "y\x \ y\X \ ordermap(X,Memrel(\))`y \ y" for y by simp with \x\\\ and \Ord(\)\ have "y\x \ y\X \ ordermap(X,Memrel(\))`y \ x" for y using ltI[OF _ Ord_in_Ord[of \ x]] lt_trans1 ltD by blast with ordm have "ordermap(X,Memrel(\))`x \ x" by auto with \x\X\ assms show ?case using subset_imp_le Ord_in_Ord[of \ x] Ord_ordermap well_ord_subset[OF well_ord_Memrel, of \] by force qed lemma subset_imp_ordertype_le: assumes "X\\" "Ord(\)" shows "ordertype(X,Memrel(\))\\" proof - { fix x assume "x\X" with assms have "ordermap(X,Memrel(\))`x \ x" using ordermap_le_arg by simp with \x\X\ and assms have "ordermap(X,Memrel(\))`x \ \" (is "?y \ _") using ltD[of ?y "succ(x)"] Ord_trans[of ?y x \] by auto } then have "ordertype(X, Memrel(\)) \ \" using ordertype_unfold[of X] by auto with assms show ?thesis using subset_imp_le Ord_ordertype[OF well_ord_subset, OF well_ord_Memrel] by simp qed lemma mono_map_imp_le: assumes "f\mono_map(\, Memrel(\),\, Memrel(\))" "Ord(\)" "Ord(\)" shows "\\\" proof - from assms have "f \ \\, Memrel(\)\ \ \f``\, Memrel(\)\" using mono_map_imp_ord_iso_Memrel by simp then have "converse(f) \ \f``\, Memrel(\)\ \ \\, Memrel(\)\" using ord_iso_sym by simp with \Ord(\)\ have "\ = ordertype(f``\,Memrel(\))" using ordertype_eq well_ord_Memrel ordertype_Memrel by auto also from assms have "ordertype(f``\,Memrel(\)) \ \" using subset_imp_ordertype_le mono_map_is_fun[of f] Image_sub_codomain[of f] by force finally show ?thesis . qed \ \\<^term>\Ord(A) \ f \ mono_map(A, Memrel(A), B, Memrel(Aa)) \ f \ inj(A, B)\\ lemmas Memrel_mono_map_is_inj = mono_map_is_inj [OF well_ord_is_linear[OF well_ord_Memrel] wf_imp_wf_on[OF wf_Memrel]] lemma mono_mapI: assumes "f: A\B" "\x y. x\A \ y\A \ \x,y\\r \ \f`x,f`y\\s" shows "f \ mono_map(A,r,B,s)" unfolding mono_map_def using assms by simp lemmas mono_mapD = mono_map_is_fun mono_map_increasing bundle mono_map_rules = mono_mapI[intro!] mono_map_is_fun[dest] mono_mapD[dest] lemma nats_le_InfCard: assumes "n \ \" "InfCard(\)" shows "n \ \" using assms Ord_is_Transset le_trans[of n \ \, OF le_subset_iff[THEN iffD2]] unfolding InfCard_def Transset_def by simp lemma nat_into_InfCard: assumes "n \ \" "InfCard(\)" shows "n \ \" using assms le_imp_subset[of \ \] unfolding InfCard_def by auto subsection\Alephs are infinite cardinals\ lemma Aleph_zero_eq_nat: "\\<^bsub>0\<^esub> = \" unfolding Aleph_def by simp lemma InfCard_Aleph: notes Aleph_zero_eq_nat[simp] assumes "Ord(\)" shows "InfCard(\\<^bsub>\\<^esub>)" proof - have "\ (\\<^bsub>\\<^esub> \ \)" proof (cases "\=0") case True then show ?thesis using mem_irrefl by auto next case False with \Ord(\)\ have "\ \ \\<^bsub>\\<^esub>" using Ord_0_lt[of \] ltD by (auto dest:Aleph_increasing) then show ?thesis using foundation by blast qed with \Ord(\)\ have "\ (|\\<^bsub>\\<^esub>| \ \)" using Card_cardinal_eq by auto then have "\ Finite(\\<^bsub>\\<^esub>)" by auto with \Ord(\)\ show ?thesis using Inf_Card_is_InfCard by simp qed +text\Most properties of cardinals depend on $\AC$, even for the countable. +Here we just state the definition of this concept, and most proofs will +appear after assuming Choice.\ +\ \Kunen's Definition I.10.5\ +definition + countable :: "i\o" where + "countable(X) \ X \ \" + +lemma countableI[intro]: "X \ \ \ countable(X)" + unfolding countable_def by simp + +lemma countableD[dest]: "countable(X) \ X \ \" + unfolding countable_def by simp + + text\A \<^emph>\delta system\ is family of sets with a common pairwise intersection. We will work with this notion in Section~\ref{sec:dsl}, but we state the definition here in order to have it available in a choiceless context.\ definition delta_system :: "i \ o" where "delta_system(D) \ \r. \A\D. \B\D. A \ B \ A \ B = r" lemma delta_systemI[intro]: assumes "\A\D. \B\D. A \ B \ A \ B = r" shows "delta_system(D)" using assms unfolding delta_system_def by simp lemma delta_systemD[dest]: "delta_system(D) \ \r. \A\D. \B\D. A \ B \ A \ B = r" unfolding delta_system_def by simp text\Hence, pairwise intersections equal the intersection of the whole family.\ lemma delta_system_root_eq_Inter: assumes "delta_system(D)" shows "\A\D. \B\D. A \ B \ A \ B = \D" proof (clarify, intro equalityI, auto) fix A' B' x C assume hyp:"A'\D" "B'\ D" "A'\B'" "x\A'" "x\B'" "C\D" with assms obtain r where delta:"\A\D. \B\D. A \ B \ A \ B = r" by auto show "x \ C" proof (cases "C=A'") case True with hyp and assms show ?thesis by simp next case False moreover note hyp moreover from calculation and delta have "r = C \ A'" "A' \ B' = r" "x\r" by auto ultimately show ?thesis by simp qed qed lemmas Limit_Aleph = InfCard_Aleph[THEN InfCard_is_Limit] lemmas Aleph_cont = Normal_imp_cont[OF Normal_Aleph] lemmas Aleph_sup = Normal_Union[OF _ _ Normal_Aleph] bundle Ord_dests = Limit_is_Ord[dest] Card_is_Ord[dest] bundle Aleph_dests = Aleph_cont[dest] Aleph_sup[dest] bundle Aleph_intros = Aleph_increasing[intro!] bundle Aleph_mem_dests = Aleph_increasing[OF ltI, THEN ltD, dest] subsection\Transfinite recursive constructions\ definition rec_constr :: "[i,i] \ i" where "rec_constr(f,\) \ transrec(\,\a g. f`(g``a))" text\The function \<^term>\rec_constr\ allows to perform \<^emph>\recursive constructions\: given a choice function on the powerset of some set, a transfinite sequence is created by successively choosing some new element. The next result explains its use.\ lemma rec_constr_unfold: "rec_constr(f,\) = f`({rec_constr(f,\). \\\})" using def_transrec[OF rec_constr_def, of f \] image_lam by simp lemma rec_constr_type: assumes "f:Pow(G)\ G" "Ord(\)" shows "rec_constr(f,\) \ G" using assms(2,1) by (induct rule:trans_induct) (subst rec_constr_unfold, rule apply_type[of f "Pow(G)" "\_. G"], auto) end \ No newline at end of file diff --git a/thys/Independence_CH/CH.thy b/thys/Independence_CH/CH.thy --- a/thys/Independence_CH/CH.thy +++ b/thys/Independence_CH/CH.thy @@ -1,447 +1,442 @@ 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 ?rnf="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 "?rnf \ M" "f\\\Coll" using nat_in_M image_closed Pi_iff by simp_all moreover from calculation have 1:"\d\?rnf. d \ h \ d \ g" if "h \ ?rnf" "g \ ?rnf" for h g proof - from calculation have "?rnf={f`x . x\\}" using image_function[of f \] Pi_iff domain_of_fun by auto from \?rnf=_\ 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 \?rnf=_\ have "f`(m\n)\?rnf" "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 "?rnf \ (\\<^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 "\ ?rnf \ (\\<^bsub>1\<^esub>\<^bsup>M\<^esup> \\<^bsup>##M\<^esup> (nat \\<^bsup>M\<^esup> 2))" using pfun_Un_filter_closed'[OF \?rnf\_\ 1] \?rnf\M\ by simp moreover from calculation have "\?rnf \\<^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 "\?rnf"] by auto moreover from calculation have "\?rnf\Coll" unfolding Fn_rel_def by simp moreover from calculation have "\?rnf \ 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\\ sublocale M_ZFC3_trans \ M_library "##M" by unfold_locales (simp_all add:repl_instances sep_instances del:setclass_iff add: transrec_replacement_def wfrec_replacement_def) locale collapse_CH = G_generic4_AC_CH "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_CH \ forcing_notion "Coll" "Colleq" 0 using zero_lt_Aleph_rel1 by unfold_locales context collapse_CH begin notation Leq (infixl "\" 50) notation Incompatible (infixl "\" 50) 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 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 - have "(\n\\. 0) \ \ \\<^bsup>M\<^esup> 2" using function_space_rel_nonempty[of 0 2 \] by auto with assms have "dense(dom_dense(x))" "x\M" using dense_dom_dense InfCard_rel_Aleph_rel[of 1] transitivity[OF _ Aleph_rel_closed[of 1,THEN setclass_iff[THEN iffD1]]] unfolding dense_def by auto with assms obtain p where "p\dom_dense(x)" "p\G" using M_generic_denseD[of "dom_dense(x)"] by auto then show "x \ domain(f\<^bsub>G\<^esub>)" by blast qed 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!: ext.transM simp add:G_in_MG) 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!: M_genericD) 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 + by fast ultimately show ?thesis using generic domain_f_G Pi_iff by auto qed abbreviation surj_dense :: "i\i" where "surj_dense(x) \ { p\Coll . x \ range(p) }" 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 lepoll_rel_refl InfCard_rel_Aleph_rel unfolding dense_def by auto with \x \ \ \\<^bsup>M\<^esup> 2\ obtain p where "p\surj_dense(x)" "p\G" using 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],simp_all) show "f\<^bsub>G\<^esub> \ surj(\\<^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 reals_sub_image_f_G succ_omega_closed_Coll succ_omega_closed_imp_no_new_reals unfolding surj_def by auto qed 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] 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_CH\\ subsection\Models of fragments of $\ZFC + \CH$\ theorem ctm_of_CH: assumes "M \ \" "Transset(M)" "M \ ZC \ {\Replacement(p)\ . p \ overhead_CH}" "\ \ 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_CH}\ interpret M_ZFC4 M - using M_satT_overhead_imp_M_ZF4 unfolding overhead_CH_def by auto - from \M \ ZC \ {\Replacement(p)\ . p \ overhead_CH}\ - have "M \ \Z\ \ {\Replacement(p)\ . p \ instances1_fms \ instances2_fms \ - instances_ground_fms \ {replacement_dcwit_repl_body_fm}}" - unfolding overhead_CH_def overhead_def ZC_def - by auto - with \Transset(M)\ + using M_satT_overhead_imp_M_ZF4 unfolding overhead_CH_def overhead_notCH_def by auto + from \M \ ZC \ {\Replacement(p)\ . p \ overhead_CH}\ \Transset(M)\ interpret M_ZF_ground_CH_trans M using M_satT_imp_M_ZF_ground_CH_trans - by simp + unfolding ZC_def by auto from \M \ \\ obtain enum where "enum \ bij(\,M)" using eqpoll_sym unfolding eqpoll_def by blast then interpret M_ctm3_AC_CH 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_CH 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_CH_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 --- a/thys/Independence_CH/Cardinal_Preservation.thy +++ b/thys/Independence_CH/Cardinal_Preservation.thy @@ -1,495 +1,488 @@ section\Preservation of cardinals in generic extensions\ theory Cardinal_Preservation imports Forcing_Main begin context forcing_data1 begin lemma antichain_abs' [absolut]: "\ A\M \ \ antichain\<^bsup>M\<^esup>(P,leq,A) \ antichain(P,leq,A)" unfolding antichain_rel_def antichain_def compat_def using transitivity[of _ A] by (auto simp add:absolut) 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 by (auto dest:transitivity; drule_tac bspec; auto dest:transitivity) qed notation check (\_\<^sup>v\ [101] 100) end \ \\<^locale>\forcing_data1\\ 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 locale G_generic4_AC_CH = G_generic4_AC + M_ZFC3_ground_CH_trans 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(G),[f,a,b'\<^sup>v]) \ \0`1 is 2\" using truth_lemma[of "\0`1 is 2\" "[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(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\\" ] 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 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 \<^term>\setclass\)\ 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(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\ 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" 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 separation_in lam_replacement_constant lam_replacement_identity - lam_replacement_Pair[THEN[5] lam_replacement_hcomp2] + lam_replacement_product separation_conj arity_fun_apply_fm union_abs1 by simp_all 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\ ))" + 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),hcomp_fm(fst_fm,fst_fm),2,0)" + let ?g_fm="hcomp_fm(check_fm(6),fst_fm,2,0)" note assms moreover 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 by (simp_all add:arity) ultimately show ?thesis using separation_sat_after_function assms that sats_fst_fm snd_abs sats_snd_fm sats_check_fm check_abs fst_abs unfolding hcomp_fm_def by simp qed - then + with assms 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] assms + using lam_replacement_imp_lam_closed separation_conj separation_in + lam_replacement_product lam_replacement_constant transitivity[of _ B] + lam_replacement_snd lam_replacement_Collect' ccc_fun_closed_lemma_aux 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\ ))" + have "separation(##M, \z. M, [snd(z), P, leq, \, f_dot, fst(fst(z))\<^sup>v, snd(fst(z))\<^sup>v] \ forces(\0`1 is 2\ ))" proof - 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 ?g_fm="hcomp_fm(check_fm(6),hcomp_fm(fst_fm,fst_fm),2,0)" let ?h_fm="hcomp_fm(check_fm(7),hcomp_fm(snd_fm,fst_fm),3,0)" note assms moreover 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 by (simp_all add:arity) ultimately show ?thesis using separation_sat_after_function3 assms sats_check_fm check_abs fst_abs snd_abs unfolding hcomp_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\ ))" + have 1:"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 assms moreover 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 by (simp_all add:arity) ultimately show ?thesis using separation_sat_after_function that fst_abs snd_abs sats_check_fm check_abs unfolding hcomp_fm_def by simp qed moreover note assms 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] + using lam_replacement_imp_lam_closed lam_replacement_Collect' transitivity[of _ A] + lam_replacement_constant lam_replacement_identity lam_replacement_snd + lam_replacement_product separation_conj separation_in separation_bex separation_iff' 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(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 G_subset_M truth_lemma[of "\0:1\2\" "[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:transitivity) moreover note \B\M\ \f = val(G,f_dot)\ moreover from calculation have "a\M" "val(G, f_dot)`a\M" by (auto dest:transitivity) 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 lam_replacement_imp_strong_replacement[OF lam_replacement_Sigfun[OF lam_replacement_constant]] Pi_replacement1 transitivity[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(P,leq,range(q))" using Pi_range_eq[of _ _ "\_ . P"] unfolding antichain_def compat_in_def by auto moreover from this and \q\M\ have "antichain\<^bsup>M\<^esup>(P,leq,range(q))" by (simp add:absolut del:P_in_M) 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\<^bsup>M\<^esup>(P,leq,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 --- a/thys/Independence_CH/Choice_Axiom.thy +++ b/thys/Independence_CH/Choice_Axiom.thy @@ -1,326 +1,326 @@ section\The Axiom of Choice in $M[G]$\ theory Choice_Axiom imports Powerset_Axiom Extensionality_Axiom Foundation_Axiom Replacement_Axiom Infinity_Axiom begin definition upair_name :: "i \ i \ i \ i" where "upair_name(\,\,on) \ Upair(\\,on\,\\,on\)" definition opair_name :: "i \ i \ i \ i" where "opair_name(\,\,on) \ upair_name(upair_name(\,\,on),upair_name(\,\,on),on)" 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 lemma (in M_ZF1_trans) 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 locale G_generic2 = G_generic1 + forcing_data2 context G_generic2 begin lemma val_upair_name : "val(G,upair_name(\,\,\)) = {val(G,\),val(G,\)}" unfolding upair_name_def using val_Upair Upair_eq_cons generic one_in_G by simp lemma val_opair_name : "val(G,opair_name(\,\,\)) = \val(G,\),val(G,\)\" unfolding opair_name_def Pair_def using val_upair_name by simp lemma val_RepFun_one: "val(G,{\f(x),\\ . x\a}) = {val(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(G,{\f(x),\\ . x \ a}) = val(G,{t \ ?A \ P . ?Q(t)})" by simp also have "... = {z . t \ ?A , (\p\P\G . ?Q(\t,p\)) \ z= val(G,t)}" using val_of_name_alt by simp also from \\\P\G\ have "... = {val(G,t) . t \ ?A }" by force also have "... = {val(G,f(x)) . x \ a}" by auto finally show ?thesis by simp qed end\ \\<^locale>\G_generic2\\ 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 lemma (in M_replacement) upair_name_lam_replacement : "M(z) \ lam_replacement(M,\x . upair_name(fst(x),snd(x),z))" using lam_replacement_Upair[THEN [5] lam_replacement_hcomp2] - lam_replacement_Pair[THEN [5] lam_replacement_hcomp2] + lam_replacement_product lam_replacement_fst lam_replacement_snd lam_replacement_constant unfolding upair_name_def by simp lemma (in forcing_data2) repl_opname_check : assumes "A\M" "f\M" shows "{opair_name(check(x),f`x,\). x\A}\M" using assms lam_replacement_constant check_lam_replacement lam_replacement_identity upair_name_lam_replacement[THEN [5] lam_replacement_hcomp2] lam_replacement_apply2[THEN [5] lam_replacement_hcomp2] lam_replacement_imp_strong_replacement_aux transitivity RepFun_closed upair_name_closed apply_closed unfolding opair_name_def by simp theorem (in G_generic2) choice_in_MG: assumes "choice_ax(##M)" shows "choice_ax(##M[G])" proof - { fix a assume "a\M[G]" then obtain \ where "\\M" "val(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(G,?f_dot)" from \?g\M\ \?f_dot = ?g\{\}\ have "?f_dot\M" using cartprod_closed singleton_closed by simp then have "f \ M[G]" unfolding f_def by (blast intro:GenExtI) have "f = {val(G,opair_name(check(\),s`\,\)) . \\\}" unfolding f_def using val_RepFun_one by simp also have "... = {\\,val(G,s`\)\ . \\\}" using val_opair_name val_check generic one_in_G one_in_P by simp finally have "f = {\\,val(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(G,\) = a\ obtain \ where "\\domain(\)" "val(G,\) = y" using elem_of_val[of y _ \] by blast with \s\surj(\,domain(\))\ obtain \ where "\\\" "s`\ = \" unfolding surj_def by auto with \val(G,\) = y\ have "val(G,s`\) = y" by simp with \f = {\\,val(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 locale G_generic2_AC = G_generic1_AC + G_generic2 + M_ctm2_AC sublocale G_generic2_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 --- a/thys/Independence_CH/Cohen_Posets_Relative.thy +++ b/thys/Independence_CH/Cohen_Posets_Relative.thy @@ -1,552 +1,547 @@ 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\\ abbreviation inj_dense :: "[i,i,i,i]\i" where "inj_dense(I,J,w,x) \ { p\Fn(\,I\\,J) . (\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 \ I" "x \ I" "w \ x" "p\Fn(\,I\\,J)" "0\J" "1\J" shows "\d\inj_dense(I,J,w,x). \d ,p\ \ Fnle(\,I\\,J)" proof - 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 \ _\ have False using Fn_nat_eq_FiniteFun FiniteFun.dom_subset[of "I\\" J] 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 \ _\ assms moreover from calculation have "cons(\\x,n\,0\, p) \ Fn(\,I\\,J)" using FiniteFun.consI[of "\x,n\" "I\\" 0 J p] Fn_nat_eq_FiniteFun by auto ultimately have "cons(\\w,n\,1\, cons(\\x,n\,0\, p) ) \ Fn(\,I\\,J)" using FiniteFun.consI[of "\w,n\" "I \ \" 1 J "cons(\\x,n\,0\, p)"] Fn_nat_eq_FiniteFun by auto with \n \ \\ show ?thesis using \p \ _\ by (intro bexI) auto qed 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)\)" 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 lam_replacement_minimum 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 lam_replacement_minimum 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 lam_replacement_minimum 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] lam_replacement_minimum 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 --- a/thys/Independence_CH/Definitions_Main.thy +++ b/thys/Independence_CH/Definitions_Main.thy @@ -1,628 +1,630 @@ 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 \<^session>\Transitive_Models\ session and the present one. 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 on which our development is based, might jump directly to Section~\ref{sec:relative-arith}, which treats relative cardinal arithmetic as implemented in \<^session>\Transitive_Models\. But in case one wants to dive deeper, the following sections treat some basic concepts of 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 on \<^term>\\\ 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) *) text\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,[])))" .. text\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 . text\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}\ text\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))) *) text\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\ text\“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) *) text\Definitions for the other connectives and the internal existential quantifier are also provided. For instance, negation:\ thm Neg_def text\@{thm [display] Neg_def}\ (* \\p\ \ \\(p \ p)\ *) 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)) *) text\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)*) text\as well as the satisfaction of an arbitrary set of sentences.\ thm satT_def text\@{thm [display] satT_def}\ (* A \ \ \ \\\\. A, [] \ \ *) text\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\label{sec:relative-arith}\ text\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 text\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 text\Analogous to 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 . text\Under appropriate hypotheses (this time, from the locale \<^term>\M_ZF_library\), \<^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_ZF_library.is_ContHyp_iff is_ContHyp_iff_CH[unfolded ContHyp_def] text\@{thm [display] M_ZF_library.is_ContHyp_iff is_ContHyp_iff_CH[unfolded ContHyp_def]}\ (* M_ZF_library(M) \ is_ContHyp(M) \ CH\<^bsup>M\<^esup> is_ContHyp(\) \ \\<^bsub>1\<^esub> = 2\<^bsup>\\\<^bsub>0\<^esub>\<^esup> *) text\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}\ text\Our first milestone was to obtain a proper extension using forcing. Its 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) *) text\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) *) -text\These results can be strengthened by enumerating five finite sets of +text\These results can be strengthened by enumerating six 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\, and \<^term>\instances_ground_fms\ which are -then collected into the $35$-element set \<^term>\overhead\. +through \<^term>\instances4_fms\, \<^term>\instances_ground_fms\, and +\<^term>\instances_ground_notCH_fms\, +which are then collected into the $33$-element set \<^term>\overhead_notCH\. For example, we have:\ thm instances1_fms_def text\@{thm [display] instances1_fms_def}\ (* instances1_fms \ -{ list_repl1_intf_fm, list_repl2_intf_fm, formula_repl1_intf_fm, formula_repl2_intf_fm, - eclose_repl1_intf_fm, eclose_repl2_intf_fm, wfrec_rank_fm, trans_repl_HVFrom_fm, - tl_repl_intf_fm } +{ list_repl1_intf_fm, list_repl2_intf_fm, formula_repl1_intf_fm, + formula_repl2_intf_fm, eclose_repl1_intf_fm, eclose_repl2_intf_fm, + wfrec_rank_fm, trans_repl_HVFrom_fm, tl_repl_intf_fm } *) -thm overhead_def -text\@{thm [display] overhead_def}\ +thm overhead_def overhead_notCH_def +text\@{thm [display] overhead_def overhead_notCH_def overhead_CH_def}\ (* -overhead \ instances1_fms \ instances2_fms \ instances3_fms \ - instances4_fms \ instances_ground_fms + overhead \ instances1_fms \ instances2_fms \ instances_ground_fms + + overhead_notCH \ overhead \ + instances3_fms \ instances4_fms \ instances_ground_notCH_fms *) -text\One further instance is needed to force $\CH$:\ - +text\One further instance is needed to force $\CH$, with a total count +of $32$ instances:\ thm overhead_CH_def text\@{thm [display] overhead_CH_def}\ (* -overhead_CH \ overhead \ {replacement_dcwit_repl_body_fm} + overhead_CH \ overhead_notCH \ {replacement_dcwit_repl_body_fm} *) thm extensions_of_ctms text\@{thm [display] extensions_of_ctms}\ (* +M \ \ \ Transset(M) \ -M \ \Z\ \ - {\Replacement(p)\ . - p \ instances1_fms \ instances2_fms \ instances_ground_fms} \ +M \ \Z\ \ {\Replacement(p)\ . p \ overhead} \ \ \ 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} \ +M \ ZC \ {\Replacement(p)\ . p \ overhead_notCH} \ \ \ 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_CH} \ \ \ formula \ M \ {\Replacement(ground_repl_fm(\))\ . \ \ \} \ \N. M \ N \ N \ \ \ Transset(N) \ N \ ZC \ {\CH\} \ {\Replacement(\)\ . \ \ \} \ (\\. Ord(\) \ \ \ M \ \ \ N) *) text\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/Fm_Definitions.thy b/thys/Independence_CH/Fm_Definitions.thy --- a/thys/Independence_CH/Fm_Definitions.thy +++ b/thys/Independence_CH/Fm_Definitions.thy @@ -1,1021 +1,1011 @@ section\Concepts involved in instances of Replacement\ theory Fm_Definitions imports Transitive_Models.Renaming_Auto Transitive_Models.Aleph_Relative FrecR_Arities begin (* Really, I have no idea why is this needed again. At the end of the imported theories, notation works just fine. *) no_notation Aleph (\\_\ [90] 90) text\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] synthesize "setdiff" from_definition "setdiff" assuming "nonempty" arity_theorem for "setdiff_fm" synthesize "is_converse" from_definition assuming "nonempty" arity_theorem for "is_converse_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 is_lambda_iff_sats[iff_sats]: assumes is_F_iff_sats: "!!a0 a1 a2. [|a0\Aa; a1\Aa; a2\Aa|] ==> is_F(a1, a0) \ sats(Aa, is_F_fm, Cons(a0,Cons(a1,Cons(a2,env))))" shows "nth(A, env) = Ab \ nth(r, env) = ra \ A \ nat \ r \ nat \ env \ list(Aa) \ is_lambda(##Aa, Ab, is_F, ra) \ Aa, env \ lambda_fm(is_F_fm,A, r)" using sats_lambda_fm[OF assms, of A r] by simp \ \same as @{thm sats_is_wfrec_fm}, but changing length assumptions to \<^term>\0\ being in the model\ lemma sats_is_wfrec_fm': assumes MH_iff_sats: "!!a0 a1 a2 a3 a4. [|a0\A; a1\A; a2\A; a3\A; a4\A|] ==> MH(a2, a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))" shows "[|x \ nat; y \ nat; z \ nat; env \ list(A); 0 \ A|] ==> sats(A, is_wfrec_fm(p,x,y,z), env) \ is_wfrec(##A, MH, nth(x,env), nth(y,env), nth(z,env))" using MH_iff_sats [THEN iff_sym] nth_closed sats_is_recfun_fm by (simp add: is_wfrec_fm_def is_wfrec_def) blast lemma is_wfrec_iff_sats'[iff_sats]: assumes MH_iff_sats: "!!a0 a1 a2 a3 a4. [|a0\Aa; a1\Aa; a2\Aa; a3\Aa; a4\Aa|] ==> MH(a2, a1, a0) \ sats(Aa, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))" "nth(x, env) = xx" "nth(y, env) = yy" "nth(z, env) = zz" "x \ nat" "y \ nat" "z \ nat" "env \ list(Aa)" "0 \ Aa" shows "is_wfrec(##Aa, MH, xx, yy, zz) \ Aa, env \ is_wfrec_fm(p,x,y,z)" using assms(2-4) sats_is_wfrec_fm'[OF assms(1,5-9)] by simp lemma is_wfrec_on_iff_sats[iff_sats]: assumes MH_iff_sats: "!!a0 a1 a2 a3 a4. [|a0\Aa; a1\Aa; a2\Aa; a3\Aa; a4\Aa|] ==> MH(a2, a1, a0) \ sats(Aa, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))" shows "nth(x, env) = xx \ nth(y, env) = yy \ nth(z, env) = zz \ x \ nat \ y \ nat \ z \ nat \ env \ list(Aa) \ 0 \ Aa \ is_wfrec_on(##Aa, MH, aa,xx, yy, zz) \ Aa, env \ is_wfrec_fm(p,x,y,z)" using assms sats_is_wfrec_fm'[OF assms] unfolding is_wfrec_on_def by simp text\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(o,x,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(o,x,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(o,x,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 _\\) \ \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 for every $q\leqslant p$ both $q\forces \sigma \in \tau$ and $q\forces \sigma \in \theta$ hold for all $\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 exist $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 synthesize "is_ordermap" from_definition assuming "nonempty" synthesize "is_ordertype" from_definition assuming "nonempty" -definition is_order_body - where "is_order_body(M,X,x,z) \ - is_well_ord(M,X,x) \ is_ordertype(M,X,x,z)" - synthesize "is_order_body" from_definition assuming "nonempty" arity_theorem for "is_order_body_fm" 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 -lemma arity_is_order_body: "arity(is_order_body_fm(2,0,1)) = 3" +lemma arity_is_order_body: "arity(is_order_body_fm(0,1)) = 2" 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" -synthesize "is_jump_cardinal_body_rel" from_definition assuming "nonempty" -arity_theorem for "is_jump_cardinal_body_rel_fm" -synthesize "is_jump_cardinal_body'_rel" from_definition assuming "nonempty" -arity_theorem for "is_jump_cardinal_body'_rel_fm" -definition replacement_is_order_body_fm where "replacement_is_order_body_fm \ is_order_body_fm(2,0,1)" +definition replacement_is_order_body_fm where "replacement_is_order_body_fm \ (\\\is_order_body_fm(1, 0) \ \\1,0\ is 2 \\\)" 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'_rel_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 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 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 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(\\1 is 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 +end \ No newline at end of file diff --git a/thys/Independence_CH/Forces_Definition.thy b/thys/Independence_CH/Forces_Definition.thy --- a/thys/Independence_CH/Forces_Definition.thy +++ b/thys/Independence_CH/Forces_Definition.thy @@ -1,854 +1,852 @@ 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 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 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] 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 (eq) "n1 \ domain(b) \ domain(c) \ (n2 = b \ n2 =c)" | (mem) "n1 = b \ n2 \ domain(c)" unfolding frecR_def by (auto simp add:components_simp) then show ?thesis proof cases case eq then have "n1 \ eclose(b) \ n1 \ eclose(c)" using Un_iff in_dom_in_eclose by auto with eq 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 mem 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\ 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 \X\M\ forcerel_in_M Hfrc_at_iff_sats[of concl:M P leq a b c 8 9 2 1 0] 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 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 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 ZF_ground_replacements(1) 1 artyf \X\M\ forcerel_in_M unfolding replacement_assm_def wfrec_Hfrc_at_fm_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 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 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 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 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 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 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) \ (q \ \ env) \ (q \ \ env))" unfolding forces_def using sats_is_leq_fm_auto assms sats_ren_forces_nand 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) \ (q \ \ env)))" 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. p \ \ ([x] @ env))" unfolding forces_def using assms sats_ren_forces_forall 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 --- a/thys/Independence_CH/Forcing_Data.thy +++ b/thys/Independence_CH/Forcing_Data.thy @@ -1,161 +1,163 @@ 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 ZF_Trans_Interpretations begin no_notation Aleph (\\_\ [90] 90) subsection\A forcing locale and generic filters\ text\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 + M_ZF_ground_trans + assumes P_in_M: "P \ M" and leq_in_M: "leq \ M" locale forcing_data2 = forcing_data1 + M_ctm2 locale forcing_data3 = forcing_data2 + M_ctm3_AC locale forcing_data4 = forcing_data3 + M_ctm4_AC context forcing_data1 begin 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)" declare iff_trans [trans] +lemma M_generic_imp_filter[dest]: "M_generic(G) \ filter(G)" + unfolding M_generic_def by blast + 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 declare P_in_M [simp,intro] declare one_in_M [simp,intro] declare leq_in_M [simp,intro] declare one_in_P [intro] end \ \\<^locale>\forcing_data1\\ locale G_generic1 = forcing_data1 + fixes G :: "i" assumes generic : "M_generic(G)" begin lemma G_nonempty: "G\0" using generic subset_refl[of P] P_dense unfolding M_generic_def by auto lemma M_genericD [dest]: "x\G \ x\P" using generic - unfolding M_generic_def by (blast dest:filterD) + by (blast dest:filterD) lemma M_generic_leqD [dest]: "p\G \ q\P \ p\q \ q\G" using generic - unfolding M_generic_def by (blast dest:filter_leqD) + by (blast dest:filter_leqD) lemma M_generic_compatD [dest]: "p\G \ r\G \ \q\G. q\p \ q\r" using generic - unfolding M_generic_def by (blast dest:low_bound_filter) + by (blast dest:low_bound_filter) lemma M_generic_denseD [dest]: "dense(D) \ D\P \ D\M \ \q\G. q\D" using generic unfolding M_generic_def by blast lemma G_subset_P: "G\P" - using generic - unfolding M_generic_def filter_def by simp + using generic by auto lemma one_in_G : "\ \ G" proof - have "increasing(G)" using generic unfolding M_generic_def filter_def by simp then show ?thesis using G_nonempty one_max unfolding increasing_def by blast qed lemma G_subset_M: "G \ M" using generic transitivity[OF _ P_in_M] 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/Forcing_Main.thy b/thys/Independence_CH/Forcing_Main.thy --- a/thys/Independence_CH/Forcing_Main.thy +++ b/thys/Independence_CH/Forcing_Main.thy @@ -1,169 +1,170 @@ section\The main theorem\ theory Forcing_Main imports Ordinals_In_MG Choice_Axiom Succession_Poset begin subsection\The generic extension is countable\ lemma (in forcing_data1) surj_nat_MG : "\f. f \ surj(\,M[G])" proof - let ?f="\n\\. val(G,enum`n)" have "x \ \ \ val(G, enum ` x)\ M[G]" for x using GenExtI bij_is_fun[OF M_countable] by simp then have "?f: \ \ M[G]" using lam_type[of \ "\n. val(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$\ 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 \ instances_ground_fms}" + "M \ \" "Transset(M)" + "M \ \Z\ \ {\Replacement(p)\ . p \ overhead}" "\ \ 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\ \ _\ \Transset(M)\ interpret M_ZF_ground_trans M using M_satT_imp_M_ZF_ground_trans by simp 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 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[G]" using cohen_extension_is_proper by blast text\Recall that \<^term>\M[G]\ denotes the generic extension of \<^term>\M\ using the poset of sequences \<^term>\2\<^bsup><\\<^esup>\.\ then interpret G_generic2 "2\<^bsup><\\<^esup>" seqle 0 _ enum G by unfold_locales interpret MG: M_Z_basic "M[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[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[G])" using choice_in_MG by simp then show "M[G], [] \ \AC\" using ZF_choice_auto sats_ZFC_iff_sats_ZF_AC unfolding ZF_choice_fm_def by simp qed moreover note \M \ M[G]\ \M \ { \Replacement(ground_repl_fm(\))\ . \ \ \}\ \\ \ formula\ moreover have "Transset(M[G])" using Transset_MG . moreover have "M \ M[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[G]" in exI, auto) qed -lemma ZF_replacement_instances12_sub_ZF: "{\Replacement(p)\ . p \ instances1_fms \ instances2_fms \ instances_ground_fms} \ ZF" +lemma ZF_replacement_instances12_sub_ZF: "{\Replacement(p)\ . p \ overhead} \ ZF" using instances1_fms_type instances2_fms_type instances_ground_fms_type - unfolding ZF_def ZF_schemes_def by auto + unfolding overhead_def 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 --- a/thys/Independence_CH/Forcing_Notions.thy +++ b/thys/Independence_CH/Forcing_Notions.thy @@ -1,429 +1,432 @@ 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 hide_const (open) Order.pred 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 Incompatible_imp_not_eq: "\ p \ q; p\P; q\P \\ p \ q" using refl_leq 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_subset_notion[dest]: "filter(G) \ G \ P" + by (auto dest:filterD) + 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 --- a/thys/Independence_CH/Forcing_Theorems.thy +++ b/thys/Independence_CH/Forcing_Theorems.thy @@ -1,1528 +1,1528 @@ 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 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 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 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 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 pair_in_M_iff 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 pair_in_M_iff 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(G,\) \ \\. \p\G. \\,p\\\ \ val(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(G,\) = x \ x\val(G,\)" by (subst def_val, auto) lemma GenExt_iff: "x\M[G] \ (\\\M. x = val(G,\))" unfolding GenExt_def by simp end end context G_generic1 begin subsection\Kunen 2013, Lemma IV.2.29\ lemma generic_inter_dense_below: assumes "D\M" "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(2), of s] leq_transD[of _ s r] by blast qed qed have "?D\P" by auto let ?d_fm="\\\compat_in_fm(1, 2, 3, 0) \ \ \0 \ 4\\" from \p\G\ have "p\M" using G_subset_M subsetD by simp moreover have "?d_fm\formula" by simp moreover have "arity(?d_fm) = 5" by (auto simp add: arity) moreover from \D\M\ \p\M\ 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 zero_in_M by simp moreover from \p\M\ 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 by simp ultimately have "?D\M" using Collect_in_M[of ?d_fm "[P,leq,p,D]"] \D\M\ by simp note asm = \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 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 (in forcing_data1) 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 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 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 assms Collect_in_M[of ?\ "[P,leq,\,\]"] by simp qed (* Lemma IV.2.40(a), membership *) lemma IV240a_mem: assumes "p\G" "\\M" "\\M" "p forces\<^sub>a (\ \ \)" "\q \. q\P \ q\G \ \\domain(\) \ q forces\<^sub>a (\ = \) \ val(G,\) = val(G,\)" (* inductive hypothesis *) shows "val(G,\)\val(G,\)" proof (intro elem_of_valI) let ?D="{q\P. \\. \r. r\P \ \\,r\ \ \ \ q\r \ q forces\<^sub>a (\ = \)}" from \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 \p\G\ ultimately obtain q where "q\G" "q\?D" using generic_inter_dense_below[of ?D p] 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(G,\) = val(G,\)" by blast+ ultimately show "\ \. \p\P. p \ G \ \\, p\ \ \ \ val(G, \) = val(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: includes some_rules assumes "p\G" "p forces\<^sub>a (\ = \)" and IH:"\q \. q\P \ q\G \ \\domain(\) \ domain(\) \ (q forces\<^sub>a (\ \ \) \ val(G,\) \ val(G,\)) \ (q forces\<^sub>a (\ \ \) \ val(G,\) \ val(G,\))" (* Strong enough for this case: *) (* IH:"\q \. q\P \ \\domain(\) \ q forces\<^sub>a (\ \ \) \ val(G,\) \ val(G,\)" *) shows "val(G,\) \ val(G,\)" proof fix x assume "x\val(G,\)" then obtain \ r where "\\,r\\\" "r\G" "val(G,\)=x" by blast moreover from this and \p\G\ obtain q where "q\G" "q\p" "q\r" by force moreover from this and \p\G\ have "q\P" "p\P" by blast+ moreover from calculation have "q forces\<^sub>a (\ \ \)" using forces_memI by auto moreover note \p forces\<^sub>a (\ = \)\ ultimately have "q forces\<^sub>a (\ \ \)" using def_forces_eq by auto with \q\P\ \q\G\ IH[of q \] \\\,r\\\\ \val(G,\) = x\ show "x\val(G,\)" by blast qed (* Lemma IV.2.40(a), equality, second inclusion--- COPY-PASTE *) lemma IV240a_eq_2nd_incl: includes some_rules assumes "p\G" "p forces\<^sub>a (\ = \)" and IH:"\q \. q\P \ q\G \ \\domain(\) \ domain(\) \ (q forces\<^sub>a (\ \ \) \ val(G,\) \ val(G,\)) \ (q forces\<^sub>a (\ \ \) \ val(G,\) \ val(G,\))" shows "val(G,\) \ val(G,\)" proof fix x assume "x\val(G,\)" then obtain \ r where "\\,r\\\" "r\G" "val(G,\)=x" by blast moreover from this and \p\G\ obtain q where "q\G" "q\p" "q\r" by force moreover from this and \p\G\ have "q\P" "p\P" by blast+ moreover from calculation have "q forces\<^sub>a (\ \ \)" using forces_memI by auto moreover note \p forces\<^sub>a (\ = \)\ ultimately have "q forces\<^sub>a (\ \ \)" using def_forces_eq by auto with \q\P\ \q\G\ IH[of q \] \\\,r\\\\ \val(G,\) = x\ show "x\val(G,\)" by blast qed (* Lemma IV.2.40(a), equality, second inclusion--- COPY-PASTE *) lemma IV240a_eq: includes some_rules assumes "p\G" "p forces\<^sub>a (\ = \)" and IH:"\q \. q\P \ q\G \ \\domain(\) \ domain(\) \ (q forces\<^sub>a (\ \ \) \ val(G,\) \ val(G,\)) \ (q forces\<^sub>a (\ \ \) \ val(G,\) \ val(G,\))" shows "val(G,\) = val(G,\)" using IV240a_eq_1st_incl[OF assms] IV240a_eq_2nd_incl[OF assms] IH by blast subsection\Induction on names\ lemma (in forcing_data1) 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 (in forcing_data1) 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 (in forcing_data1) 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: shows "(\\M \ \\M \ (\p\G. p forces\<^sub>a (\ = \) \ val(G,\) = val(G,\))) \ (\\M \ \\M \ (\p\G. p forces\<^sub>a (\ \ \) \ val(G,\) \ val(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(G, \) = val(G, \)" if "q\P" "q\G" for q \ using that domain_closed[of \] transitivity by auto ultimately show "\p\G. p forces\<^sub>a (\ \ \) \ val(G,\) \ val(G,\)" using IV240a_mem domain_closed transitivity by simp next fix \ \ assume "\\M" "\\M" and d:"\ \ domain(\) \ domain(\) \ ?R(\,\) \ ?R(\,\)" for \ moreover from this have IH':"(q forces\<^sub>a (\ \ \) \ val(G, \) \ val(G, \)) \ (q forces\<^sub>a (\ \ \) \ val(G, \) \ val(G, \))" if "\ \ domain(\) \ domain(\)" "q\G" for q \ proof - from d that have A:"?R(\,\)" "?R(\,\)" by auto from \\\_\ \\\M\ \q\G\ \\\_\ show ?thesis using transitivity[of \] domain_closed A[rule_format,of q] by auto qed show "\p\G. p forces\<^sub>a (\ = \) \ val(G,\) = val(G,\)" using IV240a_eq[OF _ _ IH'] by simp qed subsection\Lemma IV.2.40(b)\ (* Lemma IV.2.40(b), membership *) lemma IV240b_mem: includes some_rules assumes "val(G,\)\val(G,\)" "\\M" "\\M" and IH:"\\. \\domain(\) \ val(G,\) = val(G,\) \ \p\G. p forces\<^sub>a (\ = \)" (* inductive hypothesis *) shows "\p\G. p forces\<^sub>a (\ \ \)" proof - from \val(G,\)\val(G,\)\ obtain \ r where "r\G" "\\,r\\\" "val(G,\) = val(G,\)" by auto moreover from this and IH obtain p' where "p'\G" "p' forces\<^sub>a (\ = \)" by blast ultimately obtain p where "p\r" "p\p'" "p\G" "p forces\<^sub>a (\ = \)" using M_generic_compatD strengthening_eq[of p'] M_genericD by auto 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 \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 \ \\<^locale>\G_generic1\\ context forcing_data1 begin 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 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 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 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 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 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 assms forces_nmem Collect_in_M[of ?\ "[P,leq,\,\]"] using abs2 fty fsats2 farit assms forces_nmem domain_closed Un_closed Collect_in_M[of ?\ "[P,leq,\,\]"] by simp_all qed end \ \\<^locale>\forcing_data1\\ context G_generic1 begin (* Lemma IV.2.40(b), equality *) lemma IV240b_eq: includes some_rules assumes "val(G,\) = val(G,\)" "\\M" "\\M" and IH:"\\. \\domain(\)\domain(\) \ (val(G,\)\val(G,\) \ (\q\G. q forces\<^sub>a (\ \ \))) \ (val(G,\)\val(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\?D . d \ p" by blast qed moreover have "?D \ P" by auto ultimately obtain p where "p\G" "p\?D" using M_generic_denseD[of ?D] 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(G,\)\val(G,\)" using IV240a[of \ \] transitivity[OF _ domain_closed[simplified]] by force moreover note \val(G,\) = _\ ultimately obtain q where "q\G" "q forces\<^sub>a (\ \ \)" using IH[OF \\\_\] by auto moreover from this and \p\G\ obtain r where "r\P" "r\p" "r\q" by blast 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(G,\)\val(G,\)" using IV240a[of \ \] transitivity[OF _ domain_closed[simplified]] by force moreover note \val(G,\) = _\ ultimately obtain q where "q\G" "q forces\<^sub>a (\ \ \)" using IH[OF \\\_\] by auto moreover from this and \p\G\ obtain r where "r\P" "r\p" "r\q" by blast 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: "(\\M\\\M\val(G,\) = val(G,\) \ (\p\G. p forces\<^sub>a (\ = \))) \ (\\M\\\M\val(G,\) \ val(G,\) \ (\p\G. p forces\<^sub>a (\ \ \)))" (is "?Q(\,\) \ ?R(\,\)") proof (intro forces_induction) fix \ \ p assume "\\domain(\) \ ?Q(\, \)" for \ then 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(G, \) \ val(G, \) \ (\q\G. q forces\<^sub>a (\ \ \))) \ (val(G, \) \ val(G, \) \ (\q\G. q forces\<^sub>a (\ \ \)))" for \ using domain_trans[OF trans_M] by blast ultimately show "?Q(\,\)" using IV240b_eq by auto qed lemma truth_lemma_mem: assumes "env\list(M)" "n\nat" "m\nat" "np\G. p \ Member(n,m) env) \ M[G], map(val(G),env) \ Member(n,m)" using assms IV240a[of "nth(n,env)" "nth(m,env)"] IV240b[of "nth(n,env)" "nth(m,env)"] M_genericD 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)" "n\nat" "m\nat" "np\G. p \ Equal(n,m) env) \ M[G], map(val(G),env) \ Equal(n,m)" using assms IV240a(1)[of "nth(n,env)" "nth(m,env)"] IV240b(1)[of "nth(n,env)" "nth(m,env)"] M_genericD Forces_Equal[of _ "nth(n,env)" "nth(m,env)" env n m] map_val_in_MG by auto end \ \\<^locale>\G_generic1\\ 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\ context forcing_data1 begin 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 end context G_generic1 begin lemma truth_lemma_Neg: assumes "\\formula" "env\list(M)" "arity(\)\length(env)" and IH: "(\p\G. p \ \ env) \ M[G], map(val(G),env) \ \" shows "(\p\G. p \ Neg(\) env) \ M[G], map(val(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(G),env) \ Neg(\))" moreover note assms moreover from calculation have "M[G], map(val(G),env) \ \" "p\P" using map_val_in_MG by auto with IH obtain r where "r \ \ env" "r\G" "r\P" by blast moreover from this and \p\G\ obtain q where "q\p" "q\r" "q\G" "q\P" "q\M" using transitivity[OF _ P_in_M] by blast moreover from calculation have "q \ \ env" using strengthening_lemma by simp with assms \p \ _ _\ \q\p\ \q\M\ \p\P\ \q\P\ show "False" using Forces_Neg by auto next assume "M[G], map(val(G),env) \ Neg(\)" with assms have "\ (M[G], map(val(G),env) \ \)" using map_val_in_MG by simp let ?D="{p\P. (p \ \ env) \ (p \ Neg(\) env)}" from assms have "?D \ M" using separation_disj separation_closed separation_forces by simp moreover have "?D \ P" by auto moreover have "dense(?D)" proof fix q assume "q\P" with assms show "\d\{p \ P . (p \ \ env) \ (p \ Neg(\) env)}. d\ q" using refl_leq Forces_Neg by (cases "q \ Neg(\) env", auto) qed 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(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)" and IH: "(\p\G. p \ \ env) \ M[G], map(val(G),env) \ \" "(\p\G. p \ \ env) \ M[G], map(val(G),env) \ \" shows "(\p\G. (p \ And(\,\) env)) \ M[G] , map(val(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(G),env) \ And(\,\)" using Forces_And[of _ _ \ \] map_val_in_MG M_genericD by auto next assume "M[G], map(val(G),env) \ And(\,\)" moreover note assms moreover from calculation obtain q r where "q \ \ env" "r \ \ env" "q\G" "r\G" "r\P" "q\P" using map_val_in_MG Forces_And[OF M_genericD assms(1-5)] M_genericD by auto moreover from calculation obtain p where "p\q" "p\r" "p\G" by auto moreover from calculation have "(p \ \ env) \ (p \ \ env)" (* can't solve as separate goals *) using strengthening_lemma[OF M_genericD] by force ultimately show "\p\G. (p \ And(\,\) env)" using Forces_And[OF M_genericD assms(1-5)] by auto qed end 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) context forcing_data1 begin 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 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 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 separation_ax[of "?\" "[P,leq,\]@env"] separation_cong[of "##M" "\y. (M, [y,P,leq,\]@env \?\)"] by simp qed end context G_generic1 begin lemma truth_lemma: assumes "\\formula" "env\list(M)" "arity(\)\length(env)" shows "(\p\G. p \ \ env) \ M[G], map(val(G),env) \ \" using assms proof (induct arbitrary:env) case (Member x y) then show ?case using truth_lemma_mem[OF \env\list(M)\ \x\nat\ \y\nat\] arities_at_aux by simp next case (Equal x y) then show ?case using truth_lemma_eq[OF \env\list(M)\ \x\nat\ \y\nat\] arities_at_aux by simp next case (Nand \ \) then 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 \) then show ?case proof (intro iffI) assume "\p\G. (p \ Forall(\) env)" then 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)"] show "M[G], map(val(G),env) \ Forall(\)" using pred_le2 map_val_in_MG by (auto iff:GenExt_iff) next assume "M[G], map(val(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" note \arity(Forall(\)) \ length(env)\ \\\formula\ \env\list(M)\ moreover from this have ar\:"arity(\)\succ(length(env))" using pred_le2 by simp moreover from calculation have "?D1\M" using Collect_forces by simp moreover from \env\list(M)\ \\\formula\ have "?D2\M" using truth_lemma'[of \] separation_closed ar\ 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 generic 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-) obtain p where "p\G" "p\P" "p \ \ ([b] @ env)" using pred_le2 map_val_in_MG M_genericD by (auto iff:GenExt_iff) moreover note \d\G\ ultimately obtain q where "q\G" "q\P" "q\d" "q\p" using M_genericD by force 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 end context forcing_data1 begin 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(G),env) \ \)" proof (intro iffI allI impI, elim conjE) fix G assume "(p \ \ env)" "M_generic(G)" "p \ G" moreover from this interpret G_generic1 P leq \ M enum G by (unfold_locales,simp) from calculation assms show "M[G], map(val(G),env) \ \" using truth_lemma[of \] by auto next assume 1: "\G.(M_generic(G)\ p\G) \ M[G] , map(val(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 this interpret G_generic1 P leq \ M enum G by (unfold_locales,simp) from calculation 2 \p\P\ have "p\G" - unfolding M_generic_def using filter_leqD by simp + using filter_leqD by auto moreover note 1 ultimately have "M[G], map(val(G),env) \ \" by simp moreover note assms moreover from calculation obtain s where "s\G" "(s \ \ env)" using truth_lemma[of \] by blast moreover from this \r\G\ obtain q where "q\G" "q\s" "q\r" "s\P" "q\P" by blast ultimately have "\q\P. q\r \ (q \ \ env)" using strengthening_lemma[of s] by auto } 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/Interface.thy b/thys/Independence_CH/Interface.thy --- a/thys/Independence_CH/Interface.thy +++ b/thys/Independence_CH/Interface.thy @@ -1,1717 +1,1668 @@ 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 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,list_repl1_intf_fm)" "replacement_assm(M,env,list_repl2_intf_fm)" "replacement_assm(M,env,formula_repl1_intf_fm)" "replacement_assm(M,env,formula_repl2_intf_fm)" "replacement_assm(M,env,eclose_repl1_intf_fm)" "replacement_assm(M,env,eclose_repl2_intf_fm)" "replacement_assm(M,env,wfrec_rank_fm)" "replacement_assm(M,env,trans_repl_HVFrom_fm)" "replacement_assm(M,env,tl_repl_intf_fm)" definition instances1_fms where "instances1_fms \ { list_repl1_intf_fm, list_repl2_intf_fm, formula_repl1_intf_fm, formula_repl2_intf_fm, eclose_repl1_intf_fm, eclose_repl2_intf_fm, wfrec_rank_fm, trans_repl_HVFrom_fm, tl_repl_intf_fm }" text\This set has 9 internalized formulas.\ lemmas replacement_instances1_defs = list_repl1_intf_fm_def list_repl2_intf_fm_def formula_repl1_intf_fm_def formula_repl2_intf_fm_def eclose_repl1_intf_fm_def eclose_repl2_intf_fm_def wfrec_rank_fm_def trans_repl_HVFrom_fm_def tl_repl_intf_fm_def lemma instances1_fms_type[TC]: "instances1_fms \ formula" using Lambda_in_M_fm_type 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 ImageSep where "ImageSep(N,B,r,y) \ (\p[N]. p\r \ (\x[N]. x\B \ pair(N,x,y,p)))" synthesize "ImageSep" from_definition assuming "nonempty" arity_theorem for "ImageSep_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 "ImageSep_fm(1,2,0)" "[A,B]" "ImageSep(##M,A,B)"] ImageSep_iff_sats[of 1 "[_,A,B]" _ 2 _ 0 _ M] arity_ImageSep_fm ImageSep_fm_type ord_simp_union zero_in_M unfolding ImageSep_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 pair_in_M_iff[simplified] cons_closed[simplified] unfolding strong_replacement_def univalent_def apply (clarsimp, rename_tac A) 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]) 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\\ context M_ZF1_trans begin lemma 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 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" arity_theorem for "formula_functor_fm" 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" arity_theorem for "big_union_fm" context M_ZF1_trans begin lemma 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(1)[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 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 lemma 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(3)[unfolded replacement_assm_def] by simp lemma 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(9)[unfolded replacement_assm_def] by simp lemma 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(5)[unfolded replacement_assm_def] ord_simp_union by simp lemma 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(2)[unfolded replacement_assm_def] repl_sats[of M ?f "[A,0,nat]"] by simp qed lemma 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 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(6)[unfolded replacement_assm_def] by simp qed end \ \\<^locale>\M_ZF1_trans\\ 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\.\ 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" declare is_Hrank_fm_def[fm_definitions add] context M_ZF1_trans begin lemma 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(7)[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 lemma 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(8)[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 end \ \\<^locale>\M_ZF1_trans\\ 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 definition separation_assm_fm :: "[i,i,i] \ i" where "separation_assm_fm(A,x,f_fm) \ (\\ (\\ \\0 \ A +\<^sub>\ 2\ \ \\\0,1\ is x+\<^sub>\ 2 \ \ f_fm \\\)\)" lemma separation_assm_fm_type[TC]: "A \ \ \ y \ \ \ f_fm \ formula \ separation_assm_fm(A, y,f_fm) \ formula" unfolding separation_assm_fm_def by simp lemma arity_separation_assm_fm : "A \ \ \ x \ \ \ f_fm \ formula \ arity(separation_assm_fm(A, x, f_fm)) = succ(A) \ succ(x) \ pred(pred(arity(f_fm)))" using pred_Un_distrib unfolding separation_assm_fm_def by (auto simp add:arity) definition separation_assm_bin_fm where "separation_assm_bin_fm(A,y,f_fm) \ (\\(\\(\\(\\(\(\\3 \ A +\<^sub>\ 4\ \ \\3,2\ is y +\<^sub>\ 4\\ ) \ \f_fm \ \ \fst(3) is 0 \ \ \snd(3) is 1\\\\ ) \)\)\)\) " lemma separation_assm_bin_fm_type[TC]: "A \ \ \ y \ \ \ f_fm \ formula \ separation_assm_bin_fm(A, y,f_fm) \ formula" unfolding separation_assm_bin_fm_def by simp lemma arity_separation_assm_bin_fm : "A \ \ \ x \ \ \ f_fm \ formula \ arity(separation_assm_bin_fm(A, x, f_fm)) = succ(A) \ succ(x) \ (pred^4(arity(f_fm)))" using pred_Un_distrib unfolding separation_assm_bin_fm_def by (auto simp add:arity) context M_Z_trans begin lemma separation_assm_sats : assumes f_fm: "\ \ formula" and f_ar: "arity(\) = 2" and fsats: "\env x y. env\list(M) \ x\M \ y\M \ (M,[x,y]@env \ \) \ is_f(x,y)" and fabs: "\x y. x\M \ y\M \ is_f(x,y) \ y = f(x)" and fclosed: "\x. x\M \ f(x) \ M" and "A\M" shows "separation(##M, \y. \x \ M . x\A \ y = \x, f(x)\)" proof - let ?\'="separation_assm_fm(1,0,\)" let ?p="\y. \x\M . x\A \ y = \x, f(x)\" from f_fm have "?\'\formula" by simp moreover from this f_ar f_fm have "arity(?\') = 2" using arity_separation_assm_fm[of 1 0 \] ord_simp_union by simp moreover from \A\M\ calculation have "separation(##M,\y . M,[y,A] \ ?\')" using separation_ax by auto moreover have "y\M \ (M,[y,A] \ ?\') \ ?p(y)" for y using assms transitivity[OF _ \A\M\] unfolding separation_assm_fm_def by auto ultimately show ?thesis by(rule_tac separation_cong[THEN iffD1],auto) qed lemma separation_assm_bin_sats : assumes f_fm: "\ \ formula" and f_ar: "arity(\) = 3" and fsats: "\env x z y. env\list(M) \ x\M \ z\M \ y\M \ (M,[x,z,y]@env \ \) \ is_f(x,z,y)" and fabs: "\x z y. x\M \ z\M \ y\M \ is_f(x,z,y) \ y = f(x,z)" and fclosed: "\x z . x\M \ z\M \ f(x,z) \ M" and "A\M" shows "separation(##M, \y. \x \ M . x\A \ y = \x, f(fst(x),snd(x))\)" proof - let ?\'="separation_assm_bin_fm(1,0,\)" let ?p="\y. \x\M . x\A \ y = \x, f(fst(x),snd(x))\" from f_fm have "?\'\formula" by simp moreover from this f_ar f_fm have "arity(?\') = 2" using arity_separation_assm_bin_fm[of 1 0 \] ord_simp_union by simp moreover from \A\M\ calculation have "separation(##M,\y . M,[y,A] \ ?\')" using separation_ax by auto moreover have "y\M \ (M,[y,A] \ ?\') \ ?p(y)" for y using assms transitivity[OF _ \A\M\] pair_in_M_iff fst_abs snd_abs fst_closed snd_closed unfolding separation_assm_bin_fm_def by auto ultimately show ?thesis by(rule_tac separation_cong[THEN iffD1],auto) qed -lemma separation_domain: "A\M \ - separation(##M, \y. \x \ M . x\A \ y = \x, domain(x)\)" - using separation_assm_sats[of "domain_fm(0,1)"] arity_domain_fm ord_simp_union - domain_closed[simplified] - by simp - -lemma lam_replacement_domain: "lam_replacement(##M, domain)" - using lam_replacement_domain' separation_domain transM by simp - lemma separation_Union: "A\M \ separation(##M, \y. \x \ M . x\A \ y = \x, Union(x)\)" using separation_assm_sats[of "big_union_fm(0,1)"] arity_big_union_fm ord_simp_union Union_closed[simplified] by simp lemma lam_replacement_Union: "lam_replacement(##M, Union)" using lam_replacement_Union' separation_Union transM by simp lemma separation_fst: "A\M \ separation(##M, \y. \x \ M . x\A \ y = \x, fst(x)\)" using separation_assm_sats[of "fst_fm(0,1)"] arity_fst_fm ord_simp_union fst_closed fst_abs by simp lemma lam_replacement_fst: "lam_replacement(##M, fst)" using lam_replacement_fst' separation_fst transM by simp lemma separation_snd: "A\M \ separation(##M, \y. \x \ M . x\A \ y = \x, snd(x)\)" using separation_assm_sats[of "snd_fm(0,1)"] arity_snd_fm ord_simp_union snd_closed[simplified] snd_abs by simp lemma lam_replacement_snd: "lam_replacement(##M, snd)" using lam_replacement_snd' separation_snd transM by simp -lemma separation_converse: "A\M \ - separation(##M, \y. \x \ M . x\A \ y = \x, converse(x)\)" - using separation_assm_sats[of "is_converse_fm(0,1)"] arity_is_converse_fm ord_simp_union - converse_closed[simplified] nonempty - by simp - -lemma lam_replacement_converse: "lam_replacement(##M, converse)" - using lam_replacement_converse' separation_converse transM by simp - - text\Binary lambda-replacements\ -lemma separation_Upair: "A\M \ - separation(##M, \y. \x\M. x \ A \ y = \x, Upair(fst(x), snd(x))\)" - using arity_upair_fm ord_simp_union - nonempty Upair_closed upair_abs - by(rule_tac separation_assm_bin_sats[of "upair_fm(0,1,2)"],auto) - -lemma lam_replacement_Upair: "lam_replacement(##M, \x . Upair(fst(x),snd(x)))" - using lam_replacement_Upair' separation_Upair - by simp - -lemma separation_Diff: "A\M \ - separation(##M, \y. \x\M. x \ A \ y = \x, Diff(fst(x), snd(x))\)" - using arity_setdiff_fm ord_simp_union - nonempty Diff_closed setdiff_abs - by (rule_tac separation_assm_bin_sats[of "setdiff_fm(0,1,2)"],auto) - -lemma lam_replacement_Diff: "lam_replacement(##M, \x . Diff(fst(x),snd(x)))" - using lam_replacement_Diff'' separation_Diff - by simp - lemma separation_Image: "A\M \ separation(##M, \y. \x\M. x \ A \ y = \x, fst(x) `` snd(x)\)" using arity_image_fm ord_simp_union nonempty image_closed image_abs by (rule_tac separation_assm_bin_sats[of "image_fm(0,1,2)"],auto) lemma lam_replacement_Image: "lam_replacement(##M, \x . fst(x) `` snd(x))" using lam_replacement_Image' separation_Image by simp -lemma separation_comp: "A\M \ - separation(##M, \y. \x\M. x \ A \ y = \x, comp(fst(x), snd(x))\)" - using arity_composition_fm ord_simp_union - nonempty comp_closed composition_abs - by (rule_tac separation_assm_bin_sats[of "composition_fm(0,1,2)"],auto) - -lemma lam_replacement_comp: "lam_replacement(##M, \x . comp(fst(x),snd(x)))" - using lam_replacement_comp'' separation_comp - by simp - lemma separation_middle_del: "A\M \ separation(##M, \y. \x\M. x \ A \ y = \x, middle_del(fst(x), snd(x))\)" using arity_is_middle_del_fm ord_simp_union nonempty fst_abs snd_abs fst_closed snd_closed pair_in_M_iff by (rule_tac separation_assm_bin_sats[of "is_middle_del_fm(0,1,2)"], auto simp:is_middle_del_def middle_del_def) lemma lam_replacement_middle_del: "lam_replacement(##M, \r . middle_del(fst(r),snd(r)))" using lam_replacement_middle_del' separation_middle_del by simp lemma separation_prodRepl: "A\M \ separation(##M, \y. \x\M. x \ A \ y = \x, prodRepl(fst(x), snd(x))\)" using arity_is_prodRepl_fm ord_simp_union nonempty fst_abs snd_abs fst_closed snd_closed pair_in_M_iff by (rule_tac separation_assm_bin_sats[of "is_prodRepl_fm(0,1,2)"], auto simp:is_prodRepl_def prodRepl_def) lemma lam_replacement_prodRepl: "lam_replacement(##M, \r . prodRepl(fst(r),snd(r)))" using lam_replacement_prodRepl' separation_prodRepl by simp end \ \\<^locale>\M_Z_trans\\ 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\\ context M_Z_trans begin lemma (in M_basic) is_minimum_equivalence : "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 lemma separation_minimum: "A\M \ separation(##M, \y. \x\M. x \ A \ y = \x, minimum(fst(x), snd(x))\)" using arity_minimum_fm ord_simp_union is_minimum_iff minimum_abs is_minimum_equivalence nonempty minimum_closed minimum_abs by (rule_tac separation_assm_bin_sats[of "minimum_fm(0,1,2)"], auto) lemma lam_replacement_minimum: "lam_replacement(##M, \x . minimum(fst(x),snd(x)))" using lam_replacement_minimum' separation_minimum by simp end \ \\<^locale>\M_Z_trans\\ end \ No newline at end of file diff --git a/thys/Independence_CH/Kappa_Closed_Notions.thy b/thys/Independence_CH/Kappa_Closed_Notions.thy --- a/thys/Independence_CH/Kappa_Closed_Notions.thy +++ b/thys/Independence_CH/Kappa_Closed_Notions.thy @@ -1,732 +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_product 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) from \a\A\ \A\M\ have "a\M" by (auto dest:transitivity) 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( G), [f_dot, A\<^sup>v, B\<^sup>v]) \ list(M[G])" by simp moreover from calculation have "val(G,f_dot) : A \\<^bsup>M[G]\<^esup> B" using truth_lemma[of "\0:1\2\" "[f_dot, A\<^sup>v, B\<^sup>v]", THEN iffD1] typed_function_type arity_typed_function_fm val_check[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(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:transitivity) moreover from calculation have "M[G], map(val(G), [f_dot, a\<^sup>v, ?b\<^sup>v]) \ \0`1 is 2\" by simp 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\" "[f_dot, a\<^sup>v, ?b\<^sup>v]", THEN iffD2] fun_apply_type arity_fun_apply_fm val_check[OF one_in_G one_in_P] G_subset_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:transitivity simp add: union_abs2 union_abs1) ultimately show ?thesis by auto qed locale M_master_CH = M_master + M_library_DC sublocale M_ZFC3_ground_CH_trans \ M_master_CH "##M" using replacement_dcwit_repl_body by unfold_locales (simp_all add:sep_instances del:setclass_iff add: transrec_replacement_def wfrec_replacement_def dcwit_repl_body_def) context G_generic4_AC_CH 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 - let ?f_fm="fst_fm(1,0)" let ?g_fm="hcomp_fm(check_fm(6),snd_fm,2,0)" note assms moreover have "?f_fm \ formula" "arity(?f_fm) \ 7" "?g_fm \ formula" "arity(?g_fm) \ 8" using ord_simp_union unfolding hcomp_fm_def by (simp_all add:arity) ultimately show ?thesis using separation_sat_after_function using fst_abs snd_abs sats_snd_fm sats_check_fm check_abs unfolding hcomp_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 (\) \\\\\\\)\)\)\)\)\)" let ?f_fm="hcomp_fm(check_fm(5),fst_fm,1,0)" let ?g_fm="hcomp_fm(check_fm(6),snd_fm,2,0)" note assms moreover have "?f_fm \ formula" "arity(?f_fm) \ 7" "?g_fm \ formula" "arity(?g_fm) \ 8" using ord_simp_union unfolding hcomp_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 sats_check_fm check_abs fst_abs snd_abs unfolding hcomp_fm_def by simp moreover from assms have "?\(z) \ list(M)" if "(##M)(z)" for z using that by simp moreover from calculation and \r \ M\ \\ \ formula\ have "(M,?\(z) \ \) \ (M,?\'(z)\?\)" if "(##M)(z)" for z using that 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 - let ?f_fm="hcomp_fm(snd_fm,fst_fm,1,0)" let ?g_fm="hcomp_fm(check_fm(6),snd_fm,2,0)" note assms moreover have "?f_fm \ formula" "arity(?f_fm) \ 7" "?g_fm \ formula" "arity(?g_fm) \ 8" using ord_simp_union unfolding hcomp_fm_def by (simp_all add:arity) ultimately show ?thesis using separation_sat_after_function sats_check_fm check_abs fst_abs snd_abs that unfolding hcomp_fm_def by simp qed with assms show ?thesis using separation_in lam_replacement_constant lam_replacement_snd lam_replacement_fst - lam_replacement_Pair[THEN[5] lam_replacement_hcomp2] pred_nat_closed + lam_replacement_product 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: +lemma separation_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\ ))))" + (##M, \p . snd(snd(p)) \ fst(snd(p)) \ + (\b\B. M, [snd(snd(p)), P, leq, \, f_dot, (\fst(p))\<^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] \ \)" + have "separation(##M, \z. M, [snd(snd(fst(z))), P, leq, \, f_dot, (\fst(fst(z)))\<^sup>v, snd(z)\<^sup>v] \ \)" if "\\formula" "arity(\) \ 7" for \ proof - - 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 ?f_fm="hcomp_fm(snd_fm,hcomp_fm(snd_fm,fst_fm),1,0)" + let ?g="\z . (\(fst(fst(z))))\<^sup>v" + let ?g_fm="hcomp_fm(check_fm(6),hcomp_fm(big_union_fm,hcomp_fm(fst_fm,fst_fm)),2,0)" let ?h_fm="hcomp_fm(check_fm(7),snd_fm,3,0)" note assms moreover 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 assms have "?g_fm \ formula" "arity(?g_fm) \ 7" "?h_fm \ formula" "arity(?h_fm) \ 8" using ord_simp_union unfolding hcomp_fm_def by (simp_all add:arity) ultimately show ?thesis using separation_sat_after_function3[OF _ _ _ f_fm_facts] check_abs - sats_check_fm that fst_abs snd_abs + sats_check_fm that fst_abs snd_abs sats_fst_fm sats_snd_fm unfolding hcomp_fm_def by simp qed with assms show ?thesis using - separation_ball separation_imp separation_conj separation_bex separation_in separation_iff' - lam_replacement_constant lam_replacement_identity lam_replacement_hcomp + separation_conj separation_bex + lam_replacement_constant 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]] + separation_in[OF _ lam_replacement_product] 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 assms moreover have "?f_fm \ formula" "arity(?f_fm) \ 6" "?g_fm \ formula" "arity(?g_fm) \ 7" using ord_simp_union unfolding hcomp_fm_def by (simp_all add:arity) ultimately show ?thesis using separation_sat_after_function_1 sats_fst_fm that fst_abs snd_abs sats_snd_fm sats_check_fm check_abs unfolding hcomp_fm_def by simp qed with assms show ?thesis - using separation_conj separation_in - lam_replacement_constant lam_replacement_fst - lam_replacement_Pair[THEN[5] lam_replacement_hcomp2] - G_subset_M[THEN subsetD] + using separation_conj separation_in G_subset_M[THEN subsetD] + lam_replacement_constant lam_replacement_fst lam_replacement_product 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 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(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\" "[\, A\<^sup>v, B\<^sup>v]"] typed_function_type arity_typed_function_fm val_check[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 r] 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(G), [\, h\<^sup>v]) \ list(M[G])" "h\M" by (auto dest:transitivity) ultimately have "h = f" using truth_lemma[of "\0=1\" "[\, h\<^sup>v]"] val_check[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 *) text\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)"] by (auto dest:transitivity) 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 by simp moreover from calculation have "S \ M" - using separation_ball_leq_and_forces_apply_aux separation_leq_and_forces_apply_aux + using separation_leq_and_forces_apply_aux separation_leq_and_forces_apply_aux' transitivity[OF \p\P\] unfolding S_def split_def - by(rule_tac lam_replacement_Collect[THEN lam_replacement_imp_lam_closed,simplified], simp_all) + 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:transitivity) 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 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 transitivity[simplified, of g] 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'] transitivity[of _ B] by (auto dest:transitivity) 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:transitivity) 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 H let ?f="val(H,f_dot)" assume "M_generic(H) \ r \ H" moreover from this interpret g:G_generic1 _ _ _ _ _ H by unfold_locales simp note \r\P\ \f_dot\M\ \B\M\ moreover from calculation have "map(val(H), [f_dot, \\<^sup>v, B\<^sup>v]) \ list(M[H])" "r\H" by simp_all moreover from calculation and \r\H\ and \r \ \0:1\2\ [f_dot, \\<^sup>v, B\<^sup>v]\ have "?f : \ \ B" using g.truth_lemma[of "\0:1\2\" "[f_dot, \\<^sup>v, B\<^sup>v]",THEN iffD1] g.one_in_G one_in_P typed_function_type arity_typed_function_fm val_check 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:transitivity) moreover note \f_dot \ M\ \r \ P\ \M_generic(H) \ r \ H\ \map(val(H), [f_dot, \\<^sup>v, B\<^sup>v]) \ list(M[H])\ moreover from calculation have "[?f, n, ?h`n] \ list(M[H])" using M_subset_MG nat_into_M[of n] g.one_in_G by (auto dest:transitivity) 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 H]\ \without this line is slower\ val_check g.one_in_G one_in_P nat_into_M by (auto dest:transitivity 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(H, ?h\<^sup>v)" using val_check g.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:transitivity) 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/Not_CH.thy b/thys/Independence_CH/Not_CH.thy --- a/thys/Independence_CH/Not_CH.thy +++ b/thys/Independence_CH/Not_CH.thy @@ -1,600 +1,593 @@ section\Model of the negation of the Continuum Hypothesis\ theory Not_CH imports Cardinal_Preservation begin text\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 + 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\ sublocale M_master \ M_Pi_replacement by unfold_locales -\ \Note that in L. 132 we interpret \<^locale>\M_master\ from \<^locale>\M_ZFC3_trans\, so -the assumption \<^locale>\M_aleph\ for \<^term>\N\ is misleading.\ locale M_master_sub = M_master + N:M_aleph N for N + assumes M_imp_N: "M(x) \ N(x)" and Ord_iff: "Ord(x) \ M(x) \ N(x)" - (* TODO: update ground replacement assms in M_ZF4: those stemming from - M_DC, M_cardinal_library, and M_seqspace should no longer be needed - (5 total). At least M_cardinal_library is needed for CH and M_library for Not_CH. *) 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_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 separation_cdltgamma separation_cdeqgamma lemmas (in M_ZF3_trans) repl_instances = lam_replacement_inj_rel -sublocale M_ZFC3_ground_trans \ M_master "##M" +sublocale M_ZFC3_ground_notCH_trans \ M_master "##M" using replacement_trans_apply_image by unfold_locales (simp_all add:repl_instances sep_instances del:setclass_iff add: transrec_replacement_def wfrec_replacement_def) sublocale M_ZFC3_trans \ M_Pi_replacement "##M" by unfold_locales 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 lam_replacement_minimum 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 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] 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>)\ \\ \ M\ show "x \ M" using domain_of_fun 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 lam_replacement_minimum 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\ \Ord(z)\ \z\M\ have "\ \\<^bsup>M\<^esup> \\<^bsub>z\<^esub>\<^bsup>M\<^esup>" using 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>"] Card_rel_Aleph_rel[THEN Card_rel_is_Ord] 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!\ + using j.B_replacement 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) 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 show ?case by(rule trans[OF ext.Aleph_rel_zero Aleph_rel_zero[symmetric]]) 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) }" 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 "Add = Fn\<^bsup>M\<^esup>(\,\\<^bsub>2\<^esub>\<^bsup>M\<^esup>\\,2)" using Fn_nat_abs by auto moreover from this have "Fnle(\,\\<^bsub>2\<^esub>\<^bsup>M\<^esup>\\,2) = Fnle\<^bsup>M\<^esup>(\,\\<^bsub>2\<^esub>\<^bsup>M\<^esup>\\,2)" unfolding Fnle_rel_def Fnle_def by auto moreover from calculation assms have "dense(dom_dense(\x, y\))" using dense_dom_dense[of "\x,y\" "\\<^bsub>2\<^esub>\<^bsup>M\<^esup>\\" \ 2] InfCard_rel_nat unfolding dense_def by auto with assms obtain p where "p\dom_dense(\x, y\)" "p\G" using 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 Pi_iff Un_filter_is_function generic subset_trans[OF filter_subset_notion Fn_nat_subset_Pow] - unfolding M_generic_def by force lemma inj_dense_closed[intro,simp]: "w \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ x \ \\<^bsub>2\<^esub>\<^bsup>M\<^esup> \ inj_dense(\\<^bsub>2\<^esub>\<^bsup>M\<^esup>,2,w,x) \ M" using transM[OF _ Aleph_rel2_closed] separation_conj separation_bex - lam_replacement_Pair[THEN [5] lam_replacement_hcomp2] + lam_replacement_product 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 - have "0\2" by auto with assms have "dense(inj_dense(\\<^bsub>2\<^esub>\<^bsup>M\<^esup>,2,w,x))" unfolding dense_def using dense_inj_dense by auto with assms obtain p where "p\inj_dense(\\<^bsub>2\<^esub>\<^bsup>M\<^esup>,2,w,x)" "p\G" using M_generic_denseD[of "inj_dense(\\<^bsub>2\<^esub>\<^bsup>M\<^esup>,2,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.curry_closed[unfolded curry_def] G_in_MG unfolding h_G_def by simp 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 f_G_funtype G_in_MG Aleph_rel_sub_closed ext.curry_rel_exp[unfolded curry_def] ext.curry_closed[unfolded curry_def] ext.mem_function_space_rel_abs by (intro ext.mem_inj_abs[THEN iffD2],simp_all) (auto simp: inj_def h_G_def dest:Aleph_rel2_new_reals) 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[G]\<^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 Aleph_rel_nats_MG_eq_Aleph_rel_nats_M by auto moreover from calculation have "\\<^bsub>2\<^esub>\<^bsup>M[G]\<^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[G]\<^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[G]\<^esup>" "\ \\<^bsup>M[G]\<^esup> 2", OF _ _ ext.function_space_rel_closed] ext.Aleph_rel_zero 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 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_ZF_library 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_ZF_library\\ 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}" + "M \ \" "Transset(M)" "M \ ZC \ {\Replacement(p)\ . p \ overhead_notCH}" "\ \ 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}\ + from \M \ ZC \ {\Replacement(p)\ . p \ overhead_notCH}\ interpret M_ZFC4 M - using M_satT_overhead_imp_M_ZF4 by simp - from \M \ ZC \ {\Replacement(p)\ . p \ overhead}\ - have "M \ \Z\ \ {\Replacement(p)\ . p \ instances1_fms \ instances2_fms \ instances_ground_fms }" - unfolding overhead_def ZC_def - by auto - with \Transset(M)\ - interpret M_ZF_ground_trans M - using M_satT_imp_M_ZF_ground_trans - by simp + using M_satT_overhead_imp_M_ZF4 unfolding overhead_notCH_def by force + from \M \ ZC \ {\Replacement(p)\ . p \ overhead_notCH}\ \Transset(M)\ + interpret M_ZF_ground_notCH_trans M + using M_satT_imp_M_ZF_ground_notCH_trans + unfolding ZC_def by auto 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 +lemma ZF_replacement_overhead_notCH_sub_ZFC: "{\Replacement(p)\ . p \ overhead_notCH} \ ZFC" + using overhead_notCH_type unfolding ZFC_def ZF_def ZF_schemes_def by auto + lemma ZF_replacement_overhead_CH_sub_ZFC: "{\Replacement(p)\ . p \ overhead_CH} \ ZFC" using overhead_CH_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_overhead_notCH_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/Proper_Extension.thy b/thys/Independence_CH/Proper_Extension.thy --- a/thys/Independence_CH/Proper_Extension.thy +++ b/thys/Independence_CH/Proper_Extension.thy @@ -1,83 +1,83 @@ 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_data2 + separative_notion begin context fixes G assumes generic: "M_generic(G)" begin interpretation G_generic1 P leq \ M enum G by unfold_locales (simp add:generic) lemma generic_not_in_M: shows "G \ M" proof assume "G\M" then have "P - G \ M" using Diff_closed by simp moreover have "\(\q\G. q \ P - G)" "(P - G) \ P" unfolding Diff_def by auto moreover note generic ultimately show "False" using filter_complement_dense[of G] M_generic_denseD[of "P-G"] - M_generic_def[of G] by simp (*TODO: put generic ==> filter in claset ?*) + by auto qed theorem proper_extension: "M \ M[G]" using generic G_in_Gen_Ext one_in_G generic_not_in_M by force end end \ \\<^locale>\ctm_separative\\ end \ No newline at end of file diff --git a/thys/Independence_CH/Replacement_Axiom.thy b/thys/Independence_CH/Replacement_Axiom.thy --- a/thys/Independence_CH/Replacement_Axiom.thy +++ b/thys/Independence_CH/Replacement_Axiom.thy @@ -1,306 +1,306 @@ 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_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(\)" unfolding body_ground_repl_fm_def rename_split_fm_def by ((insert assms,rule iff_sats | simp add:nonempty[simplified])+, insert sats_incr_bv_iff[where bvs="[_,_,_,_,_,_]", simplified],auto del: iffI) 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(G, \') = c" "\' \ M" using GenExt_def by auto then have "domain(\')\P\M" (is "?\\M") using cartprod_closed domain_closed by simp from \val(G, \') = c\ have "c \ val(G,?\)" using def_val[of G ?\] elem_of_val[of _ G \'] one_in_G domain_of_prod[OF one_in_P, of "domain(\')"] by (force del:M_genericD) from \env \ _\ obtain nenv where "nenv\list(M)" "env = map(val(G),nenv)" using map_val by auto then have "length(nenv) = length(env)" by simp with \arity(\) \ _\ have "arity(\) \ 2 +\<^sub>\ length(nenv)" 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 note inM = \nenv\list(M)\ \?\\M\ moreover have "f(\p) \ M" "Ord(f(\p))" for \p unfolding f_def using Least_closed'[of "?P(\p)"] by simp_all 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) define QQ where "QQ\?Q" from 1 have "least(##M,\\. QQ(\p,\),f(\p))" for \p unfolding QQ_def . have body:"(M, [\p,m,P,leq,\] @ nenv \ ground_repl_fm(\)) \ least(##M, QQ(\p), m)" if "\p\M" "\p\?\" "m\M" for \p m proof - note inM that moreover from this assms 1 have "(M , [\,\p,m,P,leq,\] @ nenv \ body_ground_repl_fm(\)) \ ?Q(\p,\)" if "\\M" for \ using that sats_body_ground_repl_fm[of \p \ m nenv \] by auto moreover from calculation have body:"\\. \ \ M \ (\\\M. \V\M. is_Vset(\a. a\M, \, V) \ \ \ V \ (snd(\p) \ \ ([fst(\p),\] @ nenv))) \ M, Cons(\, [\p, m, P, leq, \] @ nenv) \ body_ground_repl_fm(\)" by simp ultimately show "(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) qed 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 \length(nenv) = length(env)\ inM 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\ body have "f(\p)\Y" if "\p\?\" for \p using that transitivity[OF _ \?\\M\] by (clarsimp,rename_tac \ p \p, rule_tac x="\\,p\" in bexI, auto) from \Y\M\ have "\ {y\Y. Ord(y)} \ M" (is "?sup \ M") using separation_Ord separation_closed Union_closed by simp then have "{x\Vset(?sup). x \ M} \ {\} \ M" (is "?big_name \ M") using Vset_closed cartprod_closed singleton_closed by simp then have "val(G,?big_name) \ M[G]" by (blast intro:GenExtI) have "{v. x\c, ?R(x,v)} \ val(G,?big_name)" (is "?repl\?big") proof(intro subsetI) fix v assume "v\?repl" moreover from this obtain x where "x\c" "M[G], [x, v] @ env \ \" "v\M[G]" by auto moreover note \val(G,\')=c\ \\'\M\ moreover from calculation obtain \ p where "\\,p\\\'" "val(G,\) = x" "p\G" "\\M" using elem_of_val_pair' by blast moreover from this \v\M[G]\ obtain \ where "val(G,\) = v" "\\M" using GenExtD by (force del:M_genericD) moreover note \\\_\ \nenv\_\ \env = _\ \arity(\)\ 2 +\<^sub>\ length(env)\ ultimately obtain q where "q\G" "q \ \ ([\,\]@nenv)" "q\P" using truth_lemma[OF \\\_\,of "[\,\] @ nenv"] by auto with \\\,p\\\'\ \\\,q\\?\ \ f(\\,q\)\Y\ have "f(\\,q\)\Y" - using generic unfolding M_generic_def filter_def by blast + using generic by blast let ?\="succ(rank(\))" note \\\M\ moreover from this have "?\ \ M" "\ \ Vset(?\)" using rank_closed cons_closed Vset_Ord_rank_iff by (simp_all flip: setclass_iff) 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(G),[\,\] @ nenv) \ \" using truth_lemma[OF \\\_\, 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(G),nenv)\ \\\M\ \val(G,\)=x\ \univalent(##M[G],_,_)\ \x\c\ \v\M[G]\ ultimately have "v=val(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(G,\) \ val(G,?big_name)" using domain_of_prod[of \ "{\}" "{x\Vset(?sup). x \ M}" ] def_val[of G ?big_name] one_in_G one_in_P by (auto simp del: Vset_rank_iff) with \v=val(G,\)\ show "v \ val(G,?big_name)" by simp qed from \?big_name\M\ have "?repl = {v\?big. \x\c. 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(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 ?\ = "(\\\\0 \ 2 +\<^sub>\ length(env) \ \ \\\)" from \\\_\ have "?\\formula" "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) moreover from \\\_\ \arity(\)\2+\<^sub>\length(env)\ \c\M[G]\ \env\_\ have "(\x\c. M[G], [x,v] @ env \ \) \ M[G], [v] @ env @ [c] \ ?\" if "v\M[G]" for v using that nth_concat transitivity_MG[OF _ \c\M[G]\] arity_sats_iff[of \ "[c]" _ "[_,v]@env"] by auto 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. 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)" "univalent(##M[G], A, ?R)" with assms have "(##M[G])(?Y)" using Replace_sats_in_MG ground_replacement 1 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 \(##M[G])(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 unfolding univalent_def 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 by simp 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 --- a/thys/Independence_CH/Replacement_Instances.thy +++ b/thys/Independence_CH/Replacement_Instances.thy @@ -1,1352 +1,1367 @@ section\More Instances of Replacement\ theory Replacement_Instances imports Separation_Instances Transitive_Models.Pointed_DC_Relative begin 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) \ (\g[N]. (\u[N]. u \ g \ (\a[N]. \y[N]. \ax[N]. \sx[N]. \r_sx[N]. \f_r_sx[N]. pair(N, a, y, u) \ pair(N, a, x, ax) \ upair(N, a, a, sx) \ pre_image(N, mesa, sx, r_sx) \ restriction(N, g, r_sx, f_r_sx) \ ax \ mesa \ is_HAleph(N, a, f_r_sx, y))) \ is_HAleph(N, x, g, 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 snd_rel_abs 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 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)))" synthesize "is_trans_apply_image_body" from_definition assuming "nonempty" arity_theorem for "is_trans_apply_image_body_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 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_HAleph_wfrec_repl_body_fm)" "replacement_assm(M,env,replacement_is_order_eq_map_fm)" - "replacement_assm(M,env,banach_iterates_fm)" definition instances2_fms where "instances2_fms \ { replacement_HAleph_wfrec_repl_body_fm, - replacement_is_order_eq_map_fm, - banach_iterates_fm }" + replacement_is_order_eq_map_fm }" text\This set has 12 internalized formulas.\ lemmas replacement_instances2_defs = replacement_HAleph_wfrec_repl_body_fm_def replacement_is_order_eq_map_fm_def - banach_iterates_fm_def + (* banach_iterates_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 text\The following instances are needed only on the ground model. The first one corresponds to the recursive definition of forces for atomic formulas; the next two corresponds to \<^term>\PHcheck\; the following is used to get a generic filter using some form of choice.\ locale M_ZF_ground = M_ZF2 + assumes ZF_ground_replacements: "replacement_assm(M,env,wfrec_Hfrc_at_fm)" "replacement_assm(M,env,wfrec_Hcheck_fm)" "replacement_assm(M,env,Lambda_in_M_fm(check_fm(2,0,1),1))" - "replacement_assm(M,env,replacement_transrec_apply_image_body_fm)" - "replacement_assm(M,env,replacement_is_trans_apply_image_fm)" locale M_ZF_ground_trans = M_ZF2_trans + M_ZF_ground definition instances_ground_fms where "instances_ground_fms \ { wfrec_Hfrc_at_fm, wfrec_Hcheck_fm, - Lambda_in_M_fm(check_fm(2,0,1),1), - replacement_transrec_apply_image_body_fm, - replacement_is_trans_apply_image_fm }" + Lambda_in_M_fm(check_fm(2,0,1),1) }" lemmas replacement_instances_ground_defs = wfrec_Hfrc_at_fm_def wfrec_Hcheck_fm_def - replacement_transrec_apply_image_body_fm_def - replacement_is_trans_apply_image_fm_def declare (in M_ZF_ground) replacement_instances_ground_defs [simp] lemma instances_ground_fms_type[TC]: "instances_ground_fms \ formula" using Lambda_in_M_fm_type unfolding instances_ground_fms_def replacement_instances_ground_defs by simp -locale M_ZF_ground_CH = M_ZF_ground + +locale M_ZF_ground_notCH = M_ZF_ground + + assumes + ZF_ground_notCH_replacements: + "replacement_assm(M,env,replacement_transrec_apply_image_body_fm)" + "replacement_assm(M,env,replacement_is_trans_apply_image_fm)" + +definition instances_ground_notCH_fms where "instances_ground_notCH_fms \ + { replacement_transrec_apply_image_body_fm, + replacement_is_trans_apply_image_fm }" + +lemma instances_ground_notCH_fms_type[TC]: "instances_ground_notCH_fms \ formula" + unfolding instances_ground_notCH_fms_def replacement_transrec_apply_image_body_fm_def + replacement_is_trans_apply_image_fm_def + by simp + +declare (in M_ZF_ground_notCH) replacement_transrec_apply_image_body_fm_def[simp] + replacement_is_trans_apply_image_fm_def[simp] + +locale M_ZF_ground_notCH_trans = M_ZF_ground_trans + M_ZF_ground_notCH + +locale M_ZF_ground_CH = M_ZF_ground_notCH + assumes dcwit_replacement: "replacement_assm(M,env,replacement_dcwit_repl_body_fm)" declare (in M_ZF_ground_CH) replacement_dcwit_repl_body_fm_def [simp] -locale M_ZF_ground_CH_trans = M_ZF_ground_trans + M_ZF_ground_CH +locale M_ZF_ground_CH_trans = M_ZF_ground_notCH_trans + M_ZF_ground_CH locale M_ZFC2 = M_ZFC1 + M_ZF2 locale M_ZFC2_trans = M_ZFC1_trans + M_ZF2_trans + M_ZFC2 locale M_ctm1 = M_ZF1_trans + fixes enum assumes M_countable: "enum\bij(nat,M)" locale M_ctm1_AC = M_ctm1 + M_ZFC1_trans locale M_ctm2 = M_ctm1 + M_ZF_ground_trans locale M_ctm2_AC = M_ctm2 + M_ZFC2_trans context M_ZF2_trans begin lemma 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(1) ord_simp_union by simp lemma 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 (in M_ZF_ground_CH_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 dcwit_replacement unfolding replacement_dcwit_repl_body_fm_def by simp lemma (in M_ZF_ground_CH_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 -lemmas 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 +lemmas M_replacement_ZF_instances = lam_replacement_fst lam_replacement_snd + lam_replacement_Union lam_replacement_Image lam_replacement_middle_del lam_replacement_prodRepl - lam_replacement_comp lemmas M_separation_ZF_instances = separation_fstsnd_in_sndsnd separation_sndfst_eq_fstsnd 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) end sublocale M_ZF2_trans \ M_replacement "##M" using M_replacement_ZF_instances M_separation_ZF_instances by unfold_locales simp context M_ZF2_trans begin lemma separation_Pow_rel: "A\M \ separation(##M, \y. \x \ M . x\A \ y = \x, Pow\<^bsup>##M\<^esup>(x)\)" using separation_assm_sats[of "is_Pow_fm(0,1)"] arity_is_Pow_fm ord_simp_union Pow_rel_closed nonempty Pow_rel_iff by simp lemma strong_replacement_Powapply_rel: "f\M \ strong_replacement(##M, \x y. y = Powapply\<^bsup>##M\<^esup>(f,x))" using Powapply_rel_replacement separation_Pow_rel transM by simp end sublocale M_ZF2_trans \ M_Vfrom "##M" using power_ax strong_replacement_Powapply_rel phrank_repl trans_repl_HVFrom wfrec_rank by unfold_locales auto 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 sublocale M_ZF2_trans \ M_pre_seqspace "##M" by unfold_locales context M_ZF2_trans begin lemma separation_inj_rel: "A\M \ separation(##M, \y. \x\M. x \ A \ y = \x, inj_rel(##M,fst(x), snd(x))\)" using arity_is_inj_fm ord_simp_union nonempty inj_rel_closed[simplified] inj_rel_iff[simplified] by (rule_tac separation_assm_bin_sats[of "is_inj_fm(0,1,2)"]) (simp_all add:setclass_def) lemma lam_replacement_inj_rel: "lam_replacement(##M, \x . inj_rel(##M,fst(x),snd(x)))" using lam_replacement_inj_rel' separation_inj_rel by simp lemma 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(2) by simp +(* + +These lemmas were required for the original proof of Schröder-Bernstein. + lemma 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(3) 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 separation_banach_functor_iterates: assumes "X\M" "Y\M" "f\M" "g\M" "A\M" shows "separation(##M, \b. \x\A. x \ \ \ b = banach_functor(X, Y, f, g)^x (0))" proof - have " (\xa\M. xa \ A \ xa \ \ \ banach_is_iterates_body(##M, X, Y, f, g, 0, xa, x)) \ (\n\A. n \ \ \ banach_is_iterates_body(##M, X, Y, f, g, 0, n, x))" if "x\M" for x using assms nat_into_M nat_in_M transM[of _ A] transM[of _ \] that by auto then have "separation(##M, \ z . \n\A . n\\ \ banach_is_iterates_body(##M,X,Y,f,g,0,n,z))" using assms nat_into_M nat_in_M arity_banach_is_iterates_body_fm[of 6 5 4 3 2 0 1] ord_simp_union separation_in_ctm[where \="(\\ \\0\7\ \ \\0\8\ \ banach_is_iterates_body_fm(6,5,4,3,2,0,1) \\\)" and env="[0,g,f,Y,X,A,\]"] by (simp add:arity_Exists arity_And) moreover from assms have "(\x\A. x \ \ \ is_iterates(##M,is_banach_functor(##M,X, Y, f, g),0,x,z)) \ (\n\A . n\\ \ banach_is_iterates_body(##M,X,Y,f,g,0,n,z))" if "z\M" for z using nat_in_M nat_into_M transM[of _ A] transM[of _ \] unfolding is_iterates_def wfrec_replacement_def is_wfrec_def M_is_recfun_def is_nat_case_def iterates_MH_def banach_body_iterates_def banach_is_iterates_body_def by simp moreover from assms have "(\x\A. x \ \ \ is_iterates(##M,is_banach_functor(##M,X, Y, f, g),0,x,z)) \ (\x\A. x \ \ \ z = banach_functor(X, Y, f, g)^x (0))" if "z\M" for z using transM[of _ A] nat_in_M nat_into_M that iterates_abs[OF banach_iterates banach_functor_abs] banach_functor_closed by auto ultimately show ?thesis by(rule_tac separation_cong[THEN iffD2],auto) qed lemma banach_replacement: assumes "X\M" "Y\M" "f\M" "g\M" shows "strong_replacement(##M, \n y. n\nat \ y = banach_functor(X, Y, f, g)^n (0))" using assms banach_repl_iter' separation_banach_functor_iterates by simp +*) + 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]. M(trans_apply_image(f, x, g))" unfolding trans_apply_image_def by simp end \ \\<^locale>\M_ZF2_trans\\ -context M_ZF_ground_trans +context M_ZF_ground_notCH_trans begin lemma 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 - ZF_ground_replacements(4) + ZF_ground_notCH_replacements(1) by simp lemma transrec_replacement_apply_image: assumes "(##M)(f)" "(##M)(\)" shows "transrec_replacement(##M, is_trans_apply_image(##M, f), \)" using replacement_transrec_apply_image_body[unfolded transrec_apply_image_body_def] assms Memrel_closed singleton_closed eclose_closed unfolding transrec_replacement_def wfrec_replacement_def is_wfrec_def M_is_recfun_def by simp lemma 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 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 ZF_ground_replacements(5)[unfolded replacement_assm_def, rule_format, where env="[\,f]",simplified]) + apply(rule_tac ZF_ground_notCH_replacements(2)[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 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 unfolding trans_apply_image_def by auto lemma 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 -end \ \\<^locale>\M_ZF_ground_trans\\ +end \ \\<^locale>\M_ZF_ground_notCH_trans\\ 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 by(rule_tac separation_cong[THEN iffD1,OF _ separation_restrict_eq_dom_eq[rule_format,of G B "converse(s) ` y "]], auto simp:transitivity[of _ G] apply_closed[simplified] converse_closed[simplified]) { 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 transitivity[of _ D] transitivity[of _ G] that 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 definition omfunspace :: "[i,i] \ o" where "omfunspace(B) \ \z. \x. \n\\. z\x \ x = n\B" relativize functional "omfunspace" "omfunspace_rel" relationalize "omfunspace_rel" "is_omfunspace" synthesize "is_omfunspace" from_definition assuming "nonempty" arity_theorem for "is_omfunspace_fm" context M_pre_seqspace begin is_iff_rel for "omfunspace" using is_function_space_iff unfolding omfunspace_rel_def is_omfunspace_def by (simp add:absolut) end \ \\<^locale>\M_pre_seqspace\\ context M_ZF2_trans begin lemma separation_omfunspace: assumes "(##M)(B)" shows "separation(##M, \z. \x[##M]. \n[##M]. n \ \ \ z \ x \ x = n \\<^bsup>M\<^esup> B)" using assms separation_in_ctm[where env="[B]" and \="is_omfunspace_fm(1,0)" and Q="is_omfunspace(##M,B)"] nonempty is_omfunspace_iff[of B, THEN separation_cong, of "##M"] arity_is_omfunspace_fm is_omfunspace_fm_type unfolding omfunspace_rel_def by (auto simp add:ord_simp_union) end \ \\<^locale>\M_ZF2_trans\\ sublocale M_ZF2_trans \ M_seqspace "##M" using separation_omfunspace by unfold_locales definition cdltgamma :: "[i,i] \ o" where "cdltgamma(\) \ \Z . |Z| < \" relativize functional "cdltgamma" "cdltgamma_rel" relationalize "cdltgamma_rel" "is_cdltgamma" synthesize "is_cdltgamma" from_definition assuming "nonempty" arity_theorem for "is_cdltgamma_fm" definition cdeqgamma :: "[i] \ o" where "cdeqgamma \ \Z . |fst(Z)| = snd(Z)" relativize functional "cdeqgamma" "cdeqgamma_rel" relationalize "cdeqgamma_rel" "is_cdeqgamma" synthesize "is_cdeqgamma" from_definition assuming "nonempty" arity_theorem for "is_cdeqgamma_fm" context M_Perm begin is_iff_rel for "cdltgamma" using is_cardinal_iff unfolding cdltgamma_rel_def is_cdltgamma_def by (simp add:absolut) is_iff_rel for "cdeqgamma" using is_cardinal_iff fst_rel_abs snd_rel_abs unfolding cdeqgamma_rel_def is_cdeqgamma_def by (auto simp add:absolut) lemma is_cdeqgamma_iff_split: "M(Z) \ cdeqgamma_rel(M, Z) \ (\\x,y\. |x|\<^bsup>M\<^esup> = y)(Z)" using fst_rel_abs snd_rel_abs unfolding cdeqgamma_rel_def split_def by simp end context M_ZF2_trans begin lemma separation_cdltgamma: assumes "(##M)(\)" shows "separation(##M, \Z . cardinal_rel(##M,Z) < \)" using assms separation_in_ctm[where env="[\]" and \="is_cdltgamma_fm(1,0)" and Q="cdltgamma_rel(##M,\)"] nonempty is_cdltgamma_iff[of \] arity_is_cdltgamma_fm is_cdltgamma_fm_type unfolding cdltgamma_rel_def by (auto simp add:ord_simp_union) lemma separation_cdeqgamma: shows "separation(##M, \Z. (\\x,y\ . cardinal_rel(##M,x) = y)(Z))" using separation_in_ctm[where env="[]" and \="is_cdeqgamma_fm(0)" and Q="cdeqgamma_rel(##M)"] is_cdeqgamma_iff_split nonempty is_cdeqgamma_iff arity_is_cdeqgamma_fm is_cdeqgamma_fm_type separation_cong[OF is_cdeqgamma_iff_split, of "##M"] unfolding cdeqgamma_rel_def by (simp add:ord_simp_union) end \ \\<^locale>\M_ZF2_trans\\ end \ No newline at end of file diff --git a/thys/Independence_CH/Separation_Axiom.thy b/thys/Independence_CH/Separation_Axiom.thy --- a/thys/Independence_CH/Separation_Axiom.thy +++ b/thys/Independence_CH/Separation_Axiom.thy @@ -1,291 +1,291 @@ 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(G),nenv)" using assms proof(induct env) case Nil have "map(val(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(G),l')" "a = val(G,a')" "Cons(a,l) = map(val(G),Cons(a',l'))" "Cons(a',l') \ list(M)" using GenExtD by force then show ?case by force qed lemma Collect_sats_in_MG : assumes "A\M[G]" "\ \ formula" "env\list(M[G])" "arity(\) \ 1 +\<^sub>\ length(env)" shows "{x \ A . (M[G], [x] @ env \ \)} \ M[G]" proof - from \A\M[G]\ obtain \ where "\ \ M" "val(G, \) = A" using GenExt_def by auto then have "domain(\)\M" "domain(\) \ P \ M" using cartprod_closed[of _ P,simplified] by (simp_all flip:setclass_iff) let ?\="\\ 0 \ (1 +\<^sub>\ length(env)) \ \ \ \" let ?new_form="sep_ren(length(env),forces(?\))" let ?\="(\\(\\\\\0,1\ is 2 \ \ ?new_form \ \)\)" note phi = \\\formula\ \arity(\) \ 1 +\<^sub>\ length(env)\ then have "?\\formula" "forces(?\) \ formula" "arity(\) \ 2+\<^sub>\ length(env)" using definability le_trans[OF \arity(\)\_\] add_le_mono[of 1 2,OF _ le_refl] by simp_all 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" "?\\formula" using arity_rensep[OF definability[of "?\"]] by auto then have "arity(?\) \ 5 +\<^sub>\ length(env)" using ord_simp_union arity_forces pred_mono[OF _ pred_mono[OF _ \arity(?new_form) \ _\]] by (auto simp:arity) from \env \ _\ obtain nenv where "nenv\list(M)" "env = map(val(G),nenv)" "length(nenv) = length(env)" using map_val by auto from phi \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\ have Equivalence: " (M, [u,P,leq,\,\] @ nenv \ ?\) \ (\\\M. \p\P. u =\\,p\ \ (\F. M_generic(F) \ p \ F \ M[F], map(val(F), [\] @ nenv @[\]) \ ?\))" if "u \ domain(\) \ P" for u proof - from \u \ domain(\) \ P\ \domain(\) \ P \ M\ have "u\M" by (simp add:transitivity) have "(M, [\,p,u,P,leq,\,\]@nenv \ ?new_form) \ (\F. M_generic(F) \ p \ F \ (M[F], map(val(F), [\] @ nenv@[\]) \ ?\))" if "\\M" "p\P" for \ p proof - from \p\P\ have "p\M" by (simp add: transitivity) let ?env="[p,P,leq,\,\] @ nenv @ [\,u]" let ?new_env=" [\,p,u,P,leq,\,\] @ nenv" note types = in_M \\ \ M\ \p\M\ \u \ domain(\) \ P\ \u \ M\ \nenv\_\ then have tyenv:"?env \ list(M)" "?new_env \ list(M)" by simp_all from types have eq_env:"[p, P, leq, \] @ ([\] @ nenv @ [\,u]) = ([p, P, leq, \] @ ([\] @ nenv @ [\])) @ [u]" using app_assoc by simp then have "(M, [\,p,u,P,leq,\,\] @ nenv \ ?new_form) \ (M, ?new_env \ ?new_form)" by simp from tyenv \length(nenv) = length(env)\ \arity(forces(?\)) \ 7 +\<^sub>\ length(env)\ \forces(?\) \ formula\ have "... \ p \ ?\ ([\] @ nenv @ [\,u])" using sepren_action[of "forces(?\)" "nenv",OF _ _ \nenv\list(M)\] by simp also from types phi \env\_\ \length(nenv) = length(env)\ \arity(forces(?\)) \ 6 +\<^sub>\ length(env)\ have "... \ p \ ?\ ([\] @ nenv @ [\])" by (subst eq_env,rule_tac arity_sats_iff,auto) also from types phi \p\P\ \arity(forces(?\)) \ 6 +\<^sub>\ length(env)\ \arity(?\) \ length([\] @ nenv @ [\])\ have " ... \ (\F . M_generic(F) \ p \ F \ M[F], map(val(F), [\] @ nenv @ [\]) \ ?\)" using definition_of_forcing[where \="\\ 0 \ (1 +\<^sub>\ length(env)) \ \ \ \"] by auto finally show ?thesis by simp qed with in_M \?new_form \ formula\ \?\\formula\ \nenv \ _\ \u \ domain(\)\P\ show ?thesis by (auto simp add: transitivity) qed moreover from \env = _\ \\\M\ \nenv\list(M)\ have map_nenv:"map(val(G), nenv @ [\]) = env @ [val(G,\)]" using map_app_distrib append1_eq_iff by auto ultimately have aux:"(\\\M. \p\P. u =\\,p\ \ (p\G \ M[G], [val(G,\)] @ env @ [val(G,\)] \ ?\))" (is "(\\\M. \p\P. _ ( _ \ M[G] , ?vals(\) \ _))") if "u \ domain(\) \ P" "M, [u,P,leq,\,\] @ nenv \ ?\" for u using Equivalence[THEN iffD1, OF that] generic by force moreover have "[val(G, \)] @ env @ [val(G, \)] \ list(M[G])" if "\\M" for \ using \\\M\ \env \ list(M[G])\ GenExtI that by force ultimately have "(\\\M. \p\P. u=\\,p\ \ (p\G \ val(G,\)\nth(1 +\<^sub>\ length(env),[val(G, \)] @ env @ [val(G, \)]) \ (M[G], ?vals(\) \ \)))" if "u \ domain(\) \ P" "M, [u,P,leq,\,\] @ nenv \ ?\" for u using aux[OF that] by simp moreover from \env \ _\ \\\M\ have nth:"nth(1 +\<^sub>\ length(env),[val(G, \)] @ env @ [val(G, \)]) = val(G,\)" if "\\M" for \ using nth_concat[of "val(G,\)" "val(G,\)" "M[G]"] that GenExtI by simp ultimately have "(\\\M. \p\P. u=\\,p\ \ (p\G \ val(G,\)\val(G,\) \ (M[G],?vals(\) \ \)))" if "u \ domain(\) \ P" "M, [u,P,leq,\,\] @ nenv \ ?\" for u using that \\\M\ \env \ _\ by simp with \domain(\)\P\M\ have "\u\domain(\)\P . (M, [u,P,leq,\,\] @ nenv \ ?\) \ (\\\M. \p\P. u =\\,p\ \ (p \ G \ val(G, \)\val(G, \) \ (M[G],?vals(\) \ \)))" by (simp add:transitivity) then have "{u\domain(\)\P . (M,[u,P,leq,\,\] @ nenv \ ?\) } \ {u\domain(\)\P . \\\M. \p\P. u =\\,p\ \ (p \ G \ val(G, \)\val(G, \) \ (M[G], ?vals(\) \ \))}" (is "?n\?m") by auto then have first_incl: "val(G,?n) \ val(G,?m)" using val_mono by simp note \val(G,\) = A\ (* 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 + by auto from \val(G,\) = A\ have "val(G,?m) = {z . t\domain(\) , (\q\P . (\\\M. \p\P. \t,q\ = \\, p\ \ (p \ G \ val(G, \) \ A \ (M[G], [val(G, \)] @ env @ [A] \ \)) \ q \ G)) \ z=val(G,t)}" using val_of_name by auto also have "... = {z . t\domain(\) , (\q\P. val(G, t) \ A \ (M[G], [val(G, t)] @ env @ [A] \ \) \ q \ G) \ z=val(G,t)}" using \domain(\)\M\ by (auto simp add:transitivity) also have "... = {x\A . \q\P. x \ A \ (M[G], [x] @ env @ [A] \ \) \ q \ G}" proof(intro equalityI, auto) (* Now we show the other inclusion: {x .. x \ A , \q\P. x \ A \ (M[G], [x, w, c] \ \) \ q \ G} \ {val(G,t)..t\domain(\),\q\P.val(G,t)\ A\(M[G], [val(G,t),w] \ \)\q\G} *) { fix x q assume "M[G], Cons(x, env @ [A]) \ \" "x\A" "q \ P" "q \ G" from this \val(G,\) = A\ show "x \ {y . x \ domain(\), val(G, x) \ A \ (M[G], Cons(val(G, x), env @ [A]) \ \) \ (\q\P. q \ G) \ y = val(G, x)}" using elem_of_val by force } qed also have " ... = {x \ A. (M[G], [x] @ env @ [A] \ \)}" using \G\P\ G_nonempty by force finally have val_m: "val(G,?m) = {x \ A. (M[G], [x] @ env @ [A] \ \)}" by simp have "val(G,?m) \ val(G,?n)" proof fix x assume "x \ val(G,?m)" with val_m have "x \ {x \ A. (M[G], [x] @ env @ [A] \ \)}" by simp with \val(G,\) = A\ have "x \ val(G,\)" by simp then obtain \ q where "\\,q\\\" "q\G" "val(G,\)=x" "\\M" using elem_of_val_pair domain_trans[OF trans_M \\\_\] by force with \\\M\ \nenv \ _\ \env = _\ have "[val(G,\), val(G,\)] @ env \ list(M[G])" "[\] @ nenv @ [\]\list(M)" using GenExt_def by auto with \val(G,\)=x\ \val(G,\) = A\ \x \ val(G,\)\ nth \\\M\ \x\ {x \ A . _}\ have "M[G], [val(G,\)] @ env @ [val(G,\)] \ \\ 0 \ (1 +\<^sub>\ length(env)) \ \ \ \" by auto \ \Recall \<^term>\?\ = And(Member(0,1 +\<^sub>\ length(env)),\)\\ with \[_] @ nenv @ [_] \ _ \ map_nenv \arity(?\) \ length(_)\ \length(nenv) = _\ obtain r where "r\G" "r \ ?\ ([\] @ nenv @ [\])" using truth_lemma[OF \?\\_\,of "[\] @ nenv @ [\]"] 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 transitivity[OF _ P_in_M] subsetD by simp_all with \\\formula\ \\\M\ \\\M\ \p\r\ \nenv \ _\ \arity(?\) \ length(_)\ \r \ ?\ _\ \env\_\ have "p \ ?\ ([\] @ nenv @ [\])" using strengthening_lemma by simp with \p\P\ \\\formula\ \\\M\ \\\M\ \nenv \ _\ \arity(?\) \ length(_)\ have "\F. M_generic(F) \ p \ F \ M[F], map(val(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(F), [\'] @ nenv @ [\]) \ ?\)" by auto from \\\M\ \\\,q\\\\ \\\M\ \p\P\ \p\M\ have "\\,q\ \ M" "\\,p\\M" "\\,p\\domain(\)\P" using pair_in_M_iff transitivity by auto with \\\M\ Eq6 \p\P\ have "M, [\\,p\,P,leq,\,\] @ nenv \ ?\" using Equivalence by auto with \\\,p\\domain(\)\P\ have "\\,p\\?n" by simp with \p\G\ \p\P\ have "val(G,\)\val(G,?n)" using val_of_elem[of \ p] by simp with \val(G,\)=x\ show "x\val(G,?n)" by simp qed (* proof of "val(G,?m) \ val(G,?n)" *) with val_m first_incl have "val(G,?n) = {x \ A. (M[G], [x] @ env @ [A] \ \)}" by auto also from \A\_\ phi \env \ _\ have " ... = {x \ A. (M[G], [x] @ env \ \)}" using arity_sats_iff[where env="[_]@env"] transitivity_MG by auto finally show "{x \ A. (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 A assume "A\M[G]" moreover from \env \ _\ obtain nenv where "nenv\list(M)""env = map(val(G),nenv)" "length(env) = length(nenv)" using GenExt_def map_val[of env] by auto moreover note \\ \ _\ \arity(\) \ _\ \env \ _\ ultimately have "{x \ A . (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 --- a/thys/Independence_CH/Separation_Instances.thy +++ b/thys/Independence_CH/Separation_Instances.thy @@ -1,212 +1,211 @@ subsection\More Instances of Separation\ theory Separation_Instances imports Interface 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 closure 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" - 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)" + "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_replacement) separation_well_ord: +lemma (in M_replacement) separation_well_ord_iso: "(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] + lam_replacement_apply[of f] lam_replacement_product 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" 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" synthesize "PiP_rel" from_definition assuming "nonempty" arity_theorem for "PiP_rel_fm" synthesize "injP_rel" from_definition assuming "nonempty" arity_theorem for "injP_rel_fm" synthesize "surjP_rel" from_definition assuming "nonempty" arity_theorem for "surjP_rel_fm" context M_ZF1_trans begin lemma 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 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 lemma 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 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 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 lemma 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 lemma 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) lemma 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) lemma 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) lemma separation_is_function: "separation(##M, is_function(##M))" using separation_in_ctm[where env="[]" and \="function_fm(0)"] arity_function_fm by simp end \ \\<^locale>\M_ZF1_trans\\ (* 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" 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" context M_ZF1_trans begin lemma 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 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 lemma 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 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 end \ \\<^locale>\M_ZF1_trans\\ end \ No newline at end of file diff --git a/thys/Independence_CH/Union_Axiom.thy b/thys/Independence_CH/Union_Axiom.thy --- a/thys/Independence_CH/Union_Axiom.thy +++ b/thys/Independence_CH/Union_Axiom.thy @@ -1,134 +1,133 @@ 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) \ \ \\domain(\) . \q\P . \r\P . \\,q\ \ \ \ \fst(x),r\ \ \ \ \snd(x),r\ \ leq \ \snd(x),q\ \ leq" definition Union_name :: "[i,i,i] \ i" where "Union_name(P,leq,\) \ {u \ domain(\(domain(\))) \ P . Union_name_body(P,leq,\,u)}" context forcing_data1 begin lemma Union_name_closed : assumes "\ \ M" shows "Union_name(P,leq,\) \ M" proof - let ?Q="Union_name_body(P,leq,\)" note lr_fst2 = lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_fst] and lr_fst3 = lam_replacement_hcomp[OF lr_fst2] lam_replacement_hcomp[OF lr_fst2 lr_fst2] note \\\M\ moreover from this have "domain(\(domain(\)))\M" (is "?d \ _") using domain_closed Union_closed by simp moreover from this have "?d \ P \ M" using cartprod_closed by simp note types = assms \?d\P \ M\ \?d\M\ ultimately show ?thesis using domain_closed pair_in_M_iff fst_closed snd_closed separation_closed lam_replacement_constant lam_replacement_hcomp - lam_replacement_fst lam_replacement_snd - lam_replacement_Pair[THEN [5] lam_replacement_hcomp2] + lam_replacement_fst lam_replacement_snd lam_replacement_product separation_bex separation_conj separation_in lr_fst2 lr_fst3 lam_replacement_hcomp[OF lr_fst3(1) lam_replacement_snd] unfolding Union_name_body_def Union_name_def by simp qed lemma Union_MG_Eq : assumes "a \ M[G]" and "a = val(G,\)" and "filter(G)" and "\ \ M" shows "\ a = val(G,Union_name(P,leq,\))" proof (intro equalityI subsetI) fix x assume "x \ \ a" with \a=_\ have "x \ \ (val(G,\))" by simp then obtain i where "i \ val(G,\)" "x \ i" by blast with \\ \ M\ obtain \ q where "q \ G" "\\,q\ \ \" "val(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(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" by (auto dest:transitivity) 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(G,\) \ val(G,Union_name(P,leq,\))" using val_of_elem by simp ultimately show "x \ val(G,Union_name(P,leq,\))" by simp next fix x assume "x \ (val(G,Union_name(P,leq,\)))" moreover note \filter(G)\ \a=val(G,\)\ moreover from calculation obtain \ p where "p \ G" "\\,p\ \ Union_name(P,leq,\)" "val(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 auto moreover from calculation have "r \ G" "q \ G" using filter_leqD by auto moreover from this \\\,r\ \ \\ \\\,q\\\\ \q\P\ \r\P\ have "val(G,\) \ val(G,\)" "val(G,\) \ val(G,\)" using val_of_elem by simp+ ultimately show "x \ \ a" by blast 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(G,\)" using GenExtD by blast moreover from this have "val(G,Union_name(P,leq,\)) \ M[G]" using GenExtI Union_name_closed 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) + by (auto simp: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 --- a/thys/Independence_CH/ZF_Trans_Interpretations.thy +++ b/thys/Independence_CH/ZF_Trans_Interpretations.thy @@ -1,665 +1,647 @@ section\Further instances of axiom-schemes\ theory ZF_Trans_Interpretations imports Internal_ZFC_Axioms Replacement_Instances 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)" definition instances3_fms where "instances3_fms \ { replacement_is_order_body_fm, - wfrec_replacement_order_pred_fm, - replacement_is_jump_cardinal_body_fm }" + wfrec_replacement_order_pred_fm }" 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_ZF3_ground = M_ZF3 + M_ZF_ground - -locale M_ZF3_ground_trans = M_ZF3_trans + M_ZF3_ground + M_ZF_ground_trans - -locale M_ZFC3_ground = M_ZFC3 + M_ZF3_ground +locale M_ZF3_ground_notCH = M_ZF3 + M_ZF_ground_notCH -locale M_ZFC3_ground_trans = M_ZFC3_trans + M_ZF3_ground_trans +locale M_ZF3_ground_notCH_trans = M_ZF3_trans + M_ZF3_ground_notCH + M_ZF_ground_notCH_trans -locale M_ZFC3_ground_CH_trans = M_ZFC3_ground_trans + M_ZF_ground_CH_trans +locale M_ZFC3_ground_notCH = M_ZFC3 + M_ZF3_ground_notCH -locale M_ctm3 = M_ctm2 + M_ZF3_ground_trans +locale M_ZFC3_ground_notCH_trans = M_ZFC3_trans + M_ZFC3_ground_notCH + M_ZF3_ground_notCH_trans -locale M_ctm3_AC = M_ctm3 + M_ctm1_AC + M_ZFC3_ground_trans +locale M_ZFC3_ground_CH_trans = M_ZFC3_ground_notCH_trans + M_ZF_ground_CH_trans + +locale M_ctm3 = M_ctm2 + M_ZF3_ground_notCH_trans + +locale M_ctm3_AC = M_ctm3 + M_ctm1_AC + M_ZFC3_ground_notCH_trans locale M_ctm3_AC_CH = M_ctm3_AC + M_ZFC3_ground_CH_trans lemmas (in M_ZF2_trans) separation_instances = - separation_well_ord + separation_well_ord_iso separation_obase_equals separation_is_obase separation_PiP_rel separation_surjP_rel separation_radd_body separation_rmult_body -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) \\)" + "omap_wfrec_body(A,r) \ (\\\image_fm(2, 0, 1) \ pred_set_fm(A #+ 9, 3, r #+ 9, 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(1,0)) = 2" + using arity_is_order_body_fm arity_is_ordertype ord_simp_union + by (simp add:FOL_arities) -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))" + "strong_replacement(##M, \x z . \y[##M]. is_order_body(##M,x,y) \ z = \x,y\)" 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 ) + where P="\ x f. M,[x,f] \ (\\ \is_order_body_fm(1,0) \ pair_fm(1,0,2) \\)",THEN iffD1]) + apply(simp add: is_order_body_iff_sats[where env="[_,_]",symmetric]) + apply(simp_all add:zero_in_M ) + apply(rule_tac replacement_ax3(1)[unfolded replacement_assm_def, rule_format, where env="[]",simplified]) + apply(simp_all add:arity_is_order_body arity pred_Un_distrib ord_simp_union) done 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 + replacement_is_order_eq_map[unfolded order_eq_map_def] by unfold_locales simp_all -lemma arity_is_jump_cardinal_body: "arity(is_jump_cardinal_body_rel_fm(0,2,3)) = 4" - unfolding is_jump_cardinal_body_rel_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 arity_is_jump_cardinal_body': "arity(is_jump_cardinal_body'_rel_fm(0,1)) = 2" - unfolding is_jump_cardinal_body'_rel_fm_def - using arity_is_jump_cardinal_body arity_is_Pow_fm arity_cartprod_fm - ord_simp_union FOL_arities - by simp - -lemma (in M_ZF3_trans) replacement_is_jump_cardinal_body: - "strong_replacement(##M, is_jump_cardinal_body'_rel(##M))" - apply(rule_tac strong_replacement_cong[ - where P="\ x f. M,[x,f] \ is_jump_cardinal_body'_rel_fm(0,1)",THEN iffD1]) - apply(rule_tac is_jump_cardinal_body'_rel_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) +definition is_well_ord_fst_snd where + "is_well_ord_fst_snd(A,x) \ (\a[A]. \b[A]. is_well_ord(A,a,b) \ is_snd(A, x, b) \ is_fst(A, x, a))" -lemma (in M_pre_cardinal_arith) is_jump_cardinal_body_abs : - "M(X) \ M(c) \ is_jump_cardinal_body'_rel(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,simplified] - unfolding is_jump_cardinal_body'_rel_def jump_cardinal_body'_rel_def - is_jump_cardinal_body_rel_def jump_cardinal_body_rel_def - by simp +synthesize "is_well_ord_fst_snd" from_definition assuming "nonempty" +arity_theorem for "is_well_ord_fst_snd_fm" -lemma (in M_ZF3_trans) replacement_jump_cardinal_body: - "strong_replacement(##M, \x z. z = jump_cardinal_body'_rel(##M,x))" - using strong_replacement_cong[THEN iffD1,OF _ replacement_is_jump_cardinal_body,simplified] - is_jump_cardinal_body_abs - by simp +lemma (in M_ZF3_trans) separation_well_ord: "separation(##M, \x. is_well_ord(##M,fst(x), snd(x)))" + using arity_is_well_ord_fst_snd_fm is_well_ord_iff_sats[symmetric] nonempty + fst_closed snd_closed fst_abs snd_abs + separation_in_ctm[where env="[]" and \="is_well_ord_fst_snd_fm(0)"] + by(simp_all add: is_well_ord_fst_snd_def) + sublocale M_ZF3_trans \ M_pre_aleph "##M" - using replacement_jump_cardinal_body[unfolded jump_cardinal_body'_rel_def] - HAleph_wfrec_repl replacement_is_order_body[unfolded is_order_body_def] + using HAleph_wfrec_repl replacement_is_order_body + separation_well_ord separation_Pow_rel 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]: "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 definition bex_Aleph_rel :: "[i\o,i,i] \ o" where "bex_Aleph_rel(M,x) \ \y. \z\x. y = \\<^bsub>z\<^esub>\<^bsup>M\<^esup>" relationalize "bex_Aleph_rel" "is_bex_Aleph" schematic_goal sats_is_bex_Aleph_fm_auto: "a \ nat \ c \ nat \ env \ list(A) \ a < length(env) \ c < length(env) \ 0 \ A \ is_bex_Aleph(##A, nth(a, env), nth(c, env)) \ A, env \ ?fm(a, c)" unfolding is_bex_Aleph_def by (rule iff_sats | simp)+ synthesize_notc "is_bex_Aleph" from_schematic lemma is_bex_Aleph_fm_type [TC]: "x \ \ \ z \ \ \ is_bex_Aleph_fm(x, z) \ formula" unfolding is_bex_Aleph_fm_def by simp lemma sats_is_bex_Aleph_fm: "x \ \ \ z \ \ \ x < length(env) \ z < length(env) \ env \ list(Aa) \ 0 \ Aa \ (Aa, env \ is_bex_Aleph_fm(x, z)) \ is_bex_Aleph(##Aa,nth(x, env), nth(z, env))" using sats_is_bex_Aleph_fm_auto unfolding is_bex_Aleph_def is_bex_Aleph_fm_def by simp lemma is_bex_Aleph_iff_sats [iff_sats]: "nth(x, env) = xa \ nth(z, env) = za \ x \ \ \ z \ \ \ x < length(env) \ z < length(env) \ env \ list(Aa) \ 0 \ Aa \ is_bex_Aleph(##Aa, xa, za) \ Aa, env \ is_bex_Aleph_fm(x, z)" using sats_is_bex_Aleph_fm by simp arity_theorem for "is_bex_Aleph_fm" lemma (in M_ZF1_trans) separation_is_bex_Aleph: assumes "(##M)(A)" shows "separation(##M,is_bex_Aleph(##M, A))" using assms separation_in_ctm[where env="[A]" and \="is_bex_Aleph_fm(1,0)", OF _ _ _ is_bex_Aleph_iff_sats[symmetric], of "\_.A"] nonempty arity_is_bex_Aleph_fm is_bex_Aleph_fm_type by (simp add:ord_simp_union) lemma (in M_pre_aleph) bex_Aleph_rel_abs: assumes "Ord(u)" "M(u)" "M(v)" shows "is_bex_Aleph(M, u, v) \ bex_Aleph_rel(M,u,v)" unfolding is_bex_Aleph_def bex_Aleph_rel_def using assms is_Aleph_iff transM[of _ u] Ord_in_Ord by simp lemma (in M_ZF3_trans) separation_bex_Aleph_rel: "Ord(x) \ (##M)(x) \ separation(##M, bex_Aleph_rel(##M,x))" using separation_is_bex_Aleph bex_Aleph_rel_abs separation_cong[where P="is_bex_Aleph(##M,x)" and M="##M",THEN iffD1] unfolding bex_Aleph_rel_def by simp sublocale M_ZF3_trans \ M_aleph "##M" using separation_bex_Aleph_rel[unfolded bex_Aleph_rel_def] by unfold_locales sublocale M_ZF2_trans \ M_FiniteFun "##M" using separation_is_function separation_omfunspace 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" using lam_replacement_minimum by unfold_locales simp (* 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 lam_replacement_minimum 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,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,wfrec_rank_fm)" "ground_replacement_assm(M,env,trans_repl_HVFrom_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_HAleph_wfrec_repl_body_fm)" "ground_replacement_assm(M,env,replacement_is_order_eq_map_fm)" - "ground_replacement_assm(M,env,banach_iterates_fm)" 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(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(wfrec_rank_fm), ground_repl_fm(trans_repl_HVFrom_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_HAleph_wfrec_repl_body_fm), - ground_repl_fm(replacement_is_order_eq_map_fm), - ground_repl_fm(banach_iterates_fm) }" + ground_repl_fm(replacement_is_order_eq_map_fm) }" -text\This set has 15 internalized formulas, corresponding to the total -count of previous replacement instances (apart from those 5 in -\<^term>\instances_ground_fms\, and \<^term>\replacement_dcwit_repl_body_fm\).\ +text\This set has $13$ internalized formulas, corresponding to the total +count of previous replacement instances (apart from those $5$ in +\<^term>\instances_ground_fms\ and \<^term>\instances_ground_notCH_fms\, +and \<^term>\replacement_dcwit_repl_body_fm\).\ definition overhead where - "overhead \ instances1_fms \ instances2_fms \ instances3_fms \ - instances4_fms \ instances_ground_fms" + "overhead \ instances1_fms \ instances2_fms \ instances_ground_fms" + +definition overhead_notCH where + "overhead_notCH \ overhead \ instances3_fms \ + instances4_fms \ instances_ground_notCH_fms" definition overhead_CH where - "overhead_CH \ overhead \ { replacement_dcwit_repl_body_fm }" + "overhead_CH \ overhead_notCH \ { replacement_dcwit_repl_body_fm }" -text\Hence, the “overhead” to force $\neg\CH$ consists -of 35 replacement instances, and one further instance is needed to +text\Hence, the “overhead” to create a proper extension of a ctm by forcing +consists of $16$ replacement instances. To force $\neg\CH$, +31 instances are need, and one further instance is required to force $\CH$.\ 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 + wfrec_replacement_order_pred_fm_def replacement_is_aleph_fm_def by (auto simp del: Lambda_in_M_fm_def) lemma overhead_type: "overhead \ formula" using instances1_fms_type instances2_fms_type instances_ground_fms_type - unfolding overhead_def instances3_fms_def instances4_fms_def - replacement_instances1_defs replacement_instances2_defs replacement_instances3_defs - replacement_instances_ground_defs - using ground_repl_fm_type Lambda_in_M_fm_type - by (auto simp del: Lambda_in_M_fm_def) + unfolding overhead_def replacement_instances1_defs replacement_instances2_defs + by simp + +lemma overhead_notCH_type: "overhead_notCH \ formula" + using overhead_type + unfolding overhead_notCH_def replacement_transrec_apply_image_body_fm_def + replacement_is_trans_apply_image_fm_def instances_ground_notCH_fms_def + instances3_fms_def instances4_fms_def + by (auto simp: replacement_instances1_defs replacement_instances2_defs + replacement_instances3_defs simp del: Lambda_in_M_fm_def) lemma overhead_CH_type: "overhead_CH \ formula" - using overhead_type unfolding overhead_CH_def replacement_dcwit_repl_body_fm_def + using overhead_notCH_type unfolding overhead_CH_def replacement_dcwit_repl_body_fm_def by auto 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 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 theorem M_satT_imp_M_ZF_ground_trans: - assumes "Transset(M)" "M \ \Z\ \ {\Replacement(p)\ . p \ instances1_fms \ instances2_fms \ instances_ground_fms}" + assumes "Transset(M)" "M \ \Z\ \ {\Replacement(p)\ . p \ overhead}" shows "M_ZF_ground_trans(M)" proof - from \M \ \Z\ \ _\ have "M \ \Z\ \ {\Replacement(p)\ . p \ instances1_fms \ instances2_fms}" "M \ {\Replacement(p)\ . p \ instances_ground_fms }" - by auto + unfolding overhead_def by auto then 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 { fix \ env assume "\ \ instances_ground_fms" "env\list(M)" moreover from this and \M \ {\Replacement(p)\ . p \ instances_ground_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 \] instances_ground_fms_type by auto } then show ?thesis unfolding instances_ground_fms_def by unfold_locales (simp_all add:replacement_assm_def) qed +theorem M_satT_imp_M_ZF_ground_notCH_trans: + assumes + "Transset(M)" + "M \ \Z\ \ {\Replacement(p)\ . p \ overhead_notCH}" + shows "M_ZF_ground_notCH_trans(M)" +proof - + from assms + interpret M_ZF_ground_trans M + using M_satT_imp_M_ZF_ground_trans unfolding overhead_notCH_def by force + { + fix \ env + assume "\ \ instances_ground_notCH_fms" "env\list(M)" + moreover from this and assms + have "M, [] \ \Replacement(\)\" + unfolding overhead_notCH_def 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 \] instances_ground_notCH_fms_type by auto + } + then + show ?thesis + by unfold_locales (simp_all add:replacement_assm_def instances_ground_notCH_fms_def) +qed + theorem M_satT_imp_M_ZF_ground_CH_trans: assumes "Transset(M)" - "M \ \Z\ \ {\Replacement(p)\ . p \ instances1_fms \ instances2_fms - \ instances_ground_fms \ {replacement_dcwit_repl_body_fm}}" + "M \ \Z\ \ {\Replacement(p)\ . p \ overhead_CH }" shows "M_ZF_ground_CH_trans(M)" proof - from assms - interpret M_ZF_ground_trans M - using M_satT_imp_M_ZF_ground_trans by auto + interpret M_ZF_ground_notCH_trans M + using M_satT_imp_M_ZF_ground_notCH_trans unfolding overhead_CH_def by auto { fix env assume "env \ list(M)" moreover from assms - have "M, [] \ \Replacement(replacement_dcwit_repl_body_fm)\" by auto + have "M, [] \ \Replacement(replacement_dcwit_repl_body_fm)\" + unfolding overhead_CH_def by auto ultimately have "arity(replacement_dcwit_repl_body_fm) \ succ(succ(length(env))) \ strong_replacement(##M,\x y. sats(M,replacement_dcwit_repl_body_fm,Cons(x,Cons(y, env))))" using sats_ZF_replacement_fm_iff[of replacement_dcwit_repl_body_fm] by (auto simp:replacement_dcwit_repl_body_fm_def) } then show ?thesis by unfold_locales (simp_all add: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)" + "(M \ ZC \ {\Replacement(p)\ . p \ overhead_notCH}) \ M_ZFC4(M)" proof - assume "M \ ZC \ {\Replacement(p)\ . p \ overhead}" + assume "M \ ZC \ {\Replacement(p)\ . p \ overhead_notCH}" 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}\ + from \M \ ZC \ {\Replacement(p)\ . p \ overhead_notCH}\ 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 + unfolding ZC_def Zermelo_fms_def ZF_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}\ + assume "\ \ overhead_notCH" "env\list(M)" + moreover from this and \M \ ZC \ {\Replacement(p)\ . p \ overhead_notCH}\ 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 + using sats_ZF_replacement_fm_iff[of \] overhead_notCH_type by auto } ultimately show "M_ZFC4(M)" - unfolding overhead_def instances1_fms_def + unfolding overhead_def overhead_notCH_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/Transitive_Models/CardinalArith_Relative.thy b/thys/Transitive_Models/CardinalArith_Relative.thy --- a/thys/Transitive_Models/CardinalArith_Relative.thy +++ b/thys/Transitive_Models/CardinalArith_Relative.thy @@ -1,1603 +1,1611 @@ section\Relative, Choice-less Cardinal Arithmetic\ theory CardinalArith_Relative imports Cardinal_Relative begin relativize functional "rvimage" "rvimage_rel" external relationalize "rvimage_rel" "is_rvimage" definition csquare_lam :: "i\i" where "csquare_lam(K) \ \\x,y\\K\K. \x \ y, x, y\" \ \Can't do the next thing because split is a missing HOC\ (* relativize functional "csquare_lam" "csquare_lam_rel" *) -relativize_tm " snd(x), fst(x), snd(x)>" "is_csquare_lam_body" +relativize_tm "\fst(x) \ snd(x), fst(x), snd(x)\" "is_csquare_lam_body" definition is_csquare_lam :: "[i\o,i,i]\o" where "is_csquare_lam(M,K,l) \ \K2[M]. cartprod(M,K,K,K2) \ is_lambda(M,K2,is_csquare_lam_body(M),l)" definition jump_cardinal_body :: "[i,i] \ i" where "jump_cardinal_body(U,X) \ {z . r \ U, well_ord(X, r) \ z = ordertype(X, r)}" +lemma jump_cardinal_body_char : + "jump_cardinal_body(U,X) = {ordertype(X, r) . r \ {r\U . well_ord(X, r)}}" + unfolding jump_cardinal_body_def + by auto + +\ \We internalize the relativized version of this particular instance.\ definition jump_cardinal_body' where "jump_cardinal_body'(x) \ jump_cardinal_body(Pow(x\x),x)" lemma (in M_cardinals) csquare_lam_closed[intro,simp]: "M(K) \ M(csquare_lam(K))" using csquare_lam_replacement unfolding csquare_lam_def by (rule lam_closed) (auto dest:transM) locale M_pre_cardinal_arith = M_cardinals + assumes wfrec_pred_replacement:"M(A) \ M(r) \ wfrec_replacement(M, \x f z. z = f `` Order.pred(A, x, r), r)" relativize_tm "\x' y' x y. z = \\x', y'\, x, y\ \ (\x', x\ \ r \ x' = x \ \y', y\ \ s)" "is_rmultP" relativize functional "rmult" "rmult_rel" external relationalize "rmult_rel" "is_rmult" lemma (in M_trivial) rmultP_abs [absolut]: "\ M(r); M(s); M(z) \ \ is_rmultP(M,s,r,z) \ (\x' y' x y. z = \\x', y'\, x, y\ \ (\x', x\ \ r \ x' = x \ \y', y\ \ s))" unfolding is_rmultP_def by (auto dest:transM) definition is_csquare_rel :: "[i\o,i,i]\o" where "is_csquare_rel(M,K,cs) \ \K2[M]. \la[M]. \memK[M]. \rmKK[M]. \rmKK2[M]. cartprod(M,K,K,K2) \ is_csquare_lam(M,K,la) \ membership(M,K,memK) \ is_rmult(M,K,memK,K,memK,rmKK) \ is_rmult(M,K,memK,K2,rmKK,rmKK2) \ is_rvimage(M,K2,la,rmKK2,cs)" context M_basic begin lemma rvimage_abs[absolut]: assumes "M(A)" "M(f)" "M(r)" "M(z)" shows "is_rvimage(M,A,f,r,z) \ z = rvimage(A,f,r)" using assms transM[OF _ \M(A)\] unfolding is_rvimage_def rvimage_def by auto lemma rmult_abs [absolut]: "\ M(A); M(r); M(B); M(s); M(z) \ \ is_rmult(M,A,r,B,s,z) \ z=rmult(A,r,B,s)" using rmultP_abs transM[of _ "(A \ B) \ A \ B"] unfolding is_rmultP_def is_rmult_def rmult_def by (auto del: iffI) lemma csquare_lam_body_abs[absolut]: "M(x) \ M(z) \ is_csquare_lam_body(M,x,z) \ z = snd(x), fst(x), snd(x)>" unfolding is_csquare_lam_body_def by (simp add:absolut) lemma csquare_lam_abs[absolut]: "M(K) \ M(l) \ is_csquare_lam(M,K,l) \ l = (\x\K\K. \fst(x) \ snd(x), fst(x), snd(x)\)" unfolding is_csquare_lam_def using lambda_abs2[of "K\K" "is_csquare_lam_body(M)" "\x. \fst(x) \ snd(x), fst(x), snd(x)\"] unfolding Relation1_def by (simp add:absolut) lemma csquare_lam_eq_lam:"csquare_lam(K) = (\z\K\K. snd(z), fst(z), snd(z)>)" proof - have "(\\x,y\\K \ K. \x \ y, x, y\)`z = (\z\K\K. snd(z), fst(z), snd(z)>)`z" if "z\K\K" for z using that by auto then show ?thesis unfolding csquare_lam_def by simp qed end \ \\<^locale>\M_basic\\ context M_pre_cardinal_arith begin lemma csquare_rel_closed[intro,simp]: "M(K) \ M(csquare_rel(K))" using csquare_lam_replacement unfolding csquare_rel_def by (intro rvimage_closed lam_closed) (auto dest:transM) -(* Ugly proof ahead, please enhance *) lemma csquare_rel_abs[absolut]: "\ M(K); M(cs)\ \ is_csquare_rel(M,K,cs) \ cs = csquare_rel(K)" + using csquare_lam_closed[unfolded csquare_lam_eq_lam] unfolding is_csquare_rel_def csquare_rel_def - using csquare_lam_closed[unfolded csquare_lam_eq_lam] by (simp add:absolut csquare_lam_eq_lam[unfolded csquare_lam_def]) end \ \\<^locale>\M_pre_cardinal_arith\\ -(************* Discipline for csucc ****************) +subsection\Discipline for \<^term>\csucc\\ + relativize functional "csucc" "csucc_rel" external relationalize "csucc_rel" "is_csucc" synthesize "is_csucc" from_definition assuming "nonempty" arity_theorem for "is_csucc_fm" abbreviation csucc_r :: "[i,i\o] \ i" (\'(_\<^sup>+')\<^bsup>_\<^esup>\) where "csucc_r(x,M) \ csucc_rel(M,x)" abbreviation csucc_r_set :: "[i,i] \ i" (\'(_\<^sup>+')\<^bsup>_\<^esup>\) where "csucc_r_set(x,M) \ csucc_rel(##M,x)" context M_Perm begin rel_closed for "csucc" using Least_closed'[of "\ L. M(L) \ Card\<^bsup>M\<^esup>(L) \ K < L"] unfolding csucc_rel_def by simp is_iff_rel for "csucc" using least_abs'[of "\ L. M(L) \ Card\<^bsup>M\<^esup>(L) \ K < L" res] is_Card_iff unfolding is_csucc_def csucc_rel_def by (simp add:absolut) end \ \\<^locale>\M_Perm\\ notation csucc_rel (\csucc\<^bsup>_\<^esup>'(_')\) -(*************** end Discipline *********************) - context M_cardinals begin lemma Card_rel_Union [simp,intro,TC]: assumes A: "\x. x\A \ Card\<^bsup>M\<^esup>(x)" and types:"M(A)" shows "Card\<^bsup>M\<^esup>(\(A))" proof (rule Card_relI) show "Ord(\A)" using A by (simp add: Card_rel_is_Ord types transM) next fix j assume j: "j < \A" moreover from this have "M(j)" unfolding lt_def by (auto simp add:types dest:transM) from j have "\c\A. j \ c \ Card\<^bsup>M\<^esup>(c)" using A types unfolding lt_def by (simp) then obtain c where c: "c\A" "j < c" "Card\<^bsup>M\<^esup>(c)" "M(c)" using Card_rel_is_Ord types unfolding lt_def by (auto dest:transM) with \M(j)\ have jls: "j \\<^bsup>M\<^esup> c" by (simp add: lt_Card_rel_imp_lesspoll_rel types) { assume eqp: "j \\<^bsup>M\<^esup> \A" have "c \\<^bsup>M\<^esup> \A" using c by (blast intro: subset_imp_lepoll_rel types) also from types \M(j)\ have "... \\<^bsup>M\<^esup> j" by (rule_tac eqpoll_rel_sym [OF eqp]) (simp_all add:types) also have "... \\<^bsup>M\<^esup> c" by (rule jls) finally have "c \\<^bsup>M\<^esup> c" by (simp_all add:\M(c)\ \M(j)\ types) with \M(c)\ have False by (auto dest:lesspoll_rel_irrefl) } thus "\ j \\<^bsup>M\<^esup> \A" by blast qed (simp_all add:types) -(* -lemma Card_UN: "(!!x. x \ A ==> Card(K(x))) ==> Card(\x\A. K(x))" - by blast - - -lemma Card_OUN [simp,intro,TC]: - "(!!x. x \ A ==> Card(K(x))) ==> Card(\xWe are not relativizing directly the results involving unions of +(ordinal) indexed families (namely, @{thm [source] Card_UN} and +@{thm [source] Card_OUN}, respectively) because of their higher order nature.\ lemma in_Card_imp_lesspoll: "[| Card\<^bsup>M\<^esup>(K); b \ K; M(K); M(b) |] ==> b \\<^bsup>M\<^esup> K" apply (unfold lesspoll_rel_def) apply (simp add: Card_rel_iff_initial) apply (fast intro!: le_imp_lepoll_rel ltI leI) done subsection\Cardinal addition\ text\Note (Paulson): Could omit proving the algebraic laws for cardinal addition and multiplication. On finite cardinals these operations coincide with addition and multiplication of natural numbers; on infinite cardinals they coincide with union (maximum). Either way we get most laws for free.\ subsubsection\Cardinal addition is commutative\ lemma sum_commute_eqpoll_rel: "M(A) \ M(B) \ A+B \\<^bsup>M\<^esup> B+A" proof (simp add: def_eqpoll_rel, rule rexI) show "(\z\A+B. case(Inr,Inl,z)) \ bij(A+B, B+A)" by (auto intro: lam_bijective [where d = "case(Inr,Inl)"]) assume "M(A)" "M(B)" then show "M(\z\A + B. case(Inr, Inl, z))" using case_replacement1 by (rule_tac lam_closed) (auto dest:transM) qed lemma cadd_rel_commute: "M(i) \ M(j) \ i \\<^bsup>M\<^esup> j = j \\<^bsup>M\<^esup> i" apply (unfold cadd_rel_def) apply (auto intro: sum_commute_eqpoll_rel [THEN cardinal_rel_cong]) done subsubsection\Cardinal addition is associative\ lemma sum_assoc_eqpoll_rel: "M(A) \ M(B) \ M(C) \ (A+B)+C \\<^bsup>M\<^esup> A+(B+C)" apply (simp add: def_eqpoll_rel) apply (rule rexI) apply (rule sum_assoc_bij) using case_replacement2 by (rule_tac lam_closed) (auto dest:transM) text\Unconditional version requires AC\ lemma well_ord_cadd_rel_assoc: assumes i: "well_ord(i,ri)" and j: "well_ord(j,rj)" and k: "well_ord(k,rk)" and types: "M(i)" "M(ri)" "M(j)" "M(rj)" "M(k)" "M(rk)" shows "(i \\<^bsup>M\<^esup> j) \\<^bsup>M\<^esup> k = i \\<^bsup>M\<^esup> (j \\<^bsup>M\<^esup> k)" proof (simp add: assms cadd_rel_def, rule cardinal_rel_cong) from types have "|i + j|\<^bsup>M\<^esup> + k \\<^bsup>M\<^esup> (i + j) + k" by (auto intro!: sum_eqpoll_rel_cong well_ord_cardinal_rel_eqpoll_rel eqpoll_rel_refl well_ord_radd i j) also have "... \\<^bsup>M\<^esup> i + (j + k)" by (rule sum_assoc_eqpoll_rel) (simp_all add:types) also have "... \\<^bsup>M\<^esup> i + |j + k|\<^bsup>M\<^esup>" proof (auto intro!: sum_eqpoll_rel_cong intro:eqpoll_rel_refl simp add:types) from types have "|j + k|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> j + k" using well_ord_cardinal_rel_eqpoll_rel[OF well_ord_radd, OF j k] by (simp) with types show "j + k \\<^bsup>M\<^esup> |j + k|\<^bsup>M\<^esup>" using eqpoll_rel_sym by simp qed finally show "|i + j|\<^bsup>M\<^esup> + k \\<^bsup>M\<^esup> i + |j + k|\<^bsup>M\<^esup>" by (simp_all add:types) qed (simp_all add:types) subsubsection\0 is the identity for addition\ lemma case_id_eq: "x\sum(A,B) \ case(\z . z, \z. z ,x) = snd(x)" unfolding case_def cond_def by (auto simp:Inl_def Inr_def) lemma lam_case_id: "(\z\0 + A. case(\x. x, \y. y, z)) = (\z\0 + A . snd(z))" using case_id_eq by simp lemma sum_0_eqpoll_rel: "M(A) \ 0+A \\<^bsup>M\<^esup> A" apply (simp add:def_eqpoll_rel) apply (rule rexI) apply (rule bij_0_sum,subst lam_case_id) using lam_replacement_snd[unfolded lam_replacement_def] by (rule lam_closed) (auto simp add:case_def cond_def Inr_def dest:transM) lemma cadd_rel_0 [simp]: "Card\<^bsup>M\<^esup>(K) \ M(K) \ 0 \\<^bsup>M\<^esup> K = K" apply (simp add: cadd_rel_def) apply (simp add: sum_0_eqpoll_rel [THEN cardinal_rel_cong] Card_rel_cardinal_rel_eq) done subsubsection\Addition by another cardinal\ lemma sum_lepoll_rel_self: "M(A) \ M(B) \ A \\<^bsup>M\<^esup> A+B" proof (simp add: def_lepoll_rel, rule rexI) show "(\x\A. Inl (x)) \ inj(A, A + B)" by (simp add: inj_def) assume "M(A)" "M(B)" then show "M(\x\A. Inl(x))" using Inl_replacement1 transM[OF _ \M(A)\] by (rule_tac lam_closed) (auto simp add: Inl_def) qed -(*Could probably weaken the premises to well_ord(K,r), or removing using AC*) - lemma cadd_rel_le_self: assumes K: "Card\<^bsup>M\<^esup>(K)" and L: "Ord(L)" and types:"M(K)" "M(L)" shows "K \ (K \\<^bsup>M\<^esup> L)" proof (simp add:types cadd_rel_def) have "K \ |K|\<^bsup>M\<^esup>" by (rule Card_rel_cardinal_rel_le [OF K]) (simp add:types) moreover have "|K|\<^bsup>M\<^esup> \ |K + L|\<^bsup>M\<^esup>" using K L by (blast intro: well_ord_lepoll_rel_imp_cardinal_rel_le sum_lepoll_rel_self well_ord_radd well_ord_Memrel Card_rel_is_Ord types) ultimately show "K \ |K + L|\<^bsup>M\<^esup>" by (blast intro: le_trans) qed subsubsection\Monotonicity of addition\ lemma sum_lepoll_rel_mono: "[| A \\<^bsup>M\<^esup> C; B \\<^bsup>M\<^esup> D; M(A); M(B); M(C); M(D) |] ==> A + B \\<^bsup>M\<^esup> C + D" apply (simp add: def_lepoll_rel) apply (elim rexE) apply (rule_tac x = "\z\A+B. case (%w. Inl(f`w), %y. Inr(fa`y), z)" in rexI) apply (rule_tac d = "case (%w. Inl(converse(f) `w), %y. Inr(converse(fa) ` y))" in lam_injective) apply (typecheck add: inj_is_fun, auto) apply (rule_tac lam_closed, auto dest:transM intro:case_replacement4) done lemma cadd_rel_le_mono: "[| K' \ K; L' \ L;M(K');M(K);M(L');M(L) |] ==> (K' \\<^bsup>M\<^esup> L') \ (K \\<^bsup>M\<^esup> L)" apply (unfold cadd_rel_def) apply (safe dest!: le_subset_iff [THEN iffD1]) apply (rule well_ord_lepoll_rel_imp_cardinal_rel_le) apply (blast intro: well_ord_radd well_ord_Memrel) apply (auto intro: sum_lepoll_rel_mono subset_imp_lepoll_rel) done subsubsection\Addition of finite cardinals is "ordinary" addition\ lemma sum_succ_eqpoll_rel: "M(A) \ M(B) \ succ(A)+B \\<^bsup>M\<^esup> succ(A+B)" apply (simp add:def_eqpoll_rel) apply (rule rexI) apply (rule_tac c = "%z. if z=Inl (A) then A+B else z" and d = "%z. if z=A+B then Inl (A) else z" in lam_bijective) apply simp_all apply (blast dest: sym [THEN eq_imp_not_mem] elim: mem_irrefl)+ apply(rule_tac lam_closed, auto dest:transM intro:if_then_range_replacement2) done -(*Pulling the succ(...) outside the |...| requires m, n \ nat *) -(*Unconditional version requires AC*) lemma cadd_succ_lemma: assumes "Ord(m)" "Ord(n)" and types: "M(m)" "M(n)" shows "succ(m) \\<^bsup>M\<^esup> n = |succ(m \\<^bsup>M\<^esup> n)|\<^bsup>M\<^esup>" using types proof (simp add: cadd_rel_def) have [intro]: "m + n \\<^bsup>M\<^esup> |m + n|\<^bsup>M\<^esup>" using assms by (blast intro: eqpoll_rel_sym well_ord_cardinal_rel_eqpoll_rel well_ord_radd well_ord_Memrel) have "|succ(m) + n|\<^bsup>M\<^esup> = |succ(m + n)|\<^bsup>M\<^esup>" by (rule sum_succ_eqpoll_rel [THEN cardinal_rel_cong]) (simp_all add:types) also have "... = |succ(|m + n|\<^bsup>M\<^esup>)|\<^bsup>M\<^esup>" by (blast intro: succ_eqpoll_rel_cong cardinal_rel_cong types) finally show "|succ(m) + n|\<^bsup>M\<^esup> = |succ(|m + n|\<^bsup>M\<^esup>)|\<^bsup>M\<^esup>" . qed lemma nat_cadd_rel_eq_add: assumes m: "m \ nat" and [simp]: "n \ nat" shows"m \\<^bsup>M\<^esup> n = m +\<^sub>\ n" using m proof (induct m) case 0 thus ?case using transM[OF _ M_nat] by (auto simp add: nat_into_Card_rel) next case (succ m) thus ?case using transM[OF _ M_nat] by (simp add: cadd_succ_lemma nat_into_Card_rel Card_rel_cardinal_rel_eq) qed subsection\Cardinal multiplication\ subsubsection\Cardinal multiplication is commutative\ lemma prod_commute_eqpoll_rel: "M(A) \ M(B) \ A*B \\<^bsup>M\<^esup> B*A" apply (simp add: def_eqpoll_rel) apply (rule rexI) apply (rule_tac c = "%." and d = "%." in lam_bijective, auto) apply(rule_tac lam_closed, auto intro:swap_replacement dest:transM) done lemma cmult_rel_commute: "M(i) \ M(j) \ i \\<^bsup>M\<^esup> j = j \\<^bsup>M\<^esup> i" apply (unfold cmult_rel_def) apply (rule prod_commute_eqpoll_rel [THEN cardinal_rel_cong], simp_all) done subsubsection\Cardinal multiplication is associative\ lemma prod_assoc_eqpoll_rel: "M(A) \ M(B) \ M(C) \ (A*B)*C \\<^bsup>M\<^esup> A*(B*C)" apply (simp add: def_eqpoll_rel) apply (rule rexI) apply (rule prod_assoc_bij) apply(rule_tac lam_closed, auto intro:assoc_replacement dest:transM) done - text\Unconditional version requires AC\ lemma well_ord_cmult_rel_assoc: assumes i: "well_ord(i,ri)" and j: "well_ord(j,rj)" and k: "well_ord(k,rk)" and types: "M(i)" "M(ri)" "M(j)" "M(rj)" "M(k)" "M(rk)" shows "(i \\<^bsup>M\<^esup> j) \\<^bsup>M\<^esup> k = i \\<^bsup>M\<^esup> (j \\<^bsup>M\<^esup> k)" proof (simp add: assms cmult_rel_def, rule cardinal_rel_cong) have "|i * j|\<^bsup>M\<^esup> * k \\<^bsup>M\<^esup> (i * j) * k" by (auto intro!: prod_eqpoll_rel_cong well_ord_cardinal_rel_eqpoll_rel eqpoll_rel_refl well_ord_rmult i j simp add:types) also have "... \\<^bsup>M\<^esup> i * (j * k)" by (rule prod_assoc_eqpoll_rel, simp_all add:types) also have "... \\<^bsup>M\<^esup> i * |j * k|\<^bsup>M\<^esup>" by (blast intro: prod_eqpoll_rel_cong well_ord_cardinal_rel_eqpoll_rel eqpoll_rel_refl well_ord_rmult j k eqpoll_rel_sym types) finally show "|i * j|\<^bsup>M\<^esup> * k \\<^bsup>M\<^esup> i * |j * k|\<^bsup>M\<^esup>" by (simp add:types) qed (simp_all add:types) subsubsection\Cardinal multiplication distributes over addition\ lemma sum_prod_distrib_eqpoll_rel: "M(A) \ M(B) \ M(C) \ (A+B)*C \\<^bsup>M\<^esup> (A*C)+(B*C)" apply (simp add: def_eqpoll_rel) apply (rule rexI) apply (rule sum_prod_distrib_bij) apply(rule_tac lam_closed, auto intro:case_replacement5 dest:transM) done - lemma well_ord_cadd_cmult_distrib: assumes i: "well_ord(i,ri)" and j: "well_ord(j,rj)" and k: "well_ord(k,rk)" and types: "M(i)" "M(ri)" "M(j)" "M(rj)" "M(k)" "M(rk)" shows "(i \\<^bsup>M\<^esup> j) \\<^bsup>M\<^esup> k = (i \\<^bsup>M\<^esup> k) \\<^bsup>M\<^esup> (j \\<^bsup>M\<^esup> k)" proof (simp add: assms cadd_rel_def cmult_rel_def, rule cardinal_rel_cong) have "|i + j|\<^bsup>M\<^esup> * k \\<^bsup>M\<^esup> (i + j) * k" by (blast intro: prod_eqpoll_rel_cong well_ord_cardinal_rel_eqpoll_rel eqpoll_rel_refl well_ord_radd i j types) also have "... \\<^bsup>M\<^esup> i * k + j * k" by (rule sum_prod_distrib_eqpoll_rel) (simp_all add:types) also have "... \\<^bsup>M\<^esup> |i * k|\<^bsup>M\<^esup> + |j * k|\<^bsup>M\<^esup>" by (blast intro: sum_eqpoll_rel_cong well_ord_cardinal_rel_eqpoll_rel well_ord_rmult i j k eqpoll_rel_sym types) finally show "|i + j|\<^bsup>M\<^esup> * k \\<^bsup>M\<^esup> |i * k|\<^bsup>M\<^esup> + |j * k|\<^bsup>M\<^esup>" by (simp add:types) qed (simp_all add:types) subsubsection\Multiplication by 0 yields 0\ lemma prod_0_eqpoll_rel: "M(A) \ 0*A \\<^bsup>M\<^esup> 0" apply (simp add: def_eqpoll_rel) apply (rule rexI) apply (rule lam_bijective, auto) done lemma cmult_rel_0 [simp]: "M(i) \ 0 \\<^bsup>M\<^esup> i = 0" by (simp add: cmult_rel_def prod_0_eqpoll_rel [THEN cardinal_rel_cong]) subsubsection\1 is the identity for multiplication\ lemma prod_singleton_eqpoll_rel: "M(x) \ M(A) \ {x}*A \\<^bsup>M\<^esup> A" apply (simp add: def_eqpoll_rel) apply (rule rexI) apply (rule singleton_prod_bij [THEN bij_converse_bij]) apply (rule converse_closed) apply(rule_tac lam_closed, auto intro:prepend_replacement dest:transM) done lemma cmult_rel_1 [simp]: "Card\<^bsup>M\<^esup>(K) \ M(K) \ 1 \\<^bsup>M\<^esup> K = K" apply (simp add: cmult_rel_def succ_def) apply (simp add: prod_singleton_eqpoll_rel[THEN cardinal_rel_cong] Card_rel_cardinal_rel_eq) done subsection\Some inequalities for multiplication\ lemma prod_square_lepoll_rel: "M(A) \ A \\<^bsup>M\<^esup> A*A" apply (simp add:def_lepoll_rel inj_def) apply (rule_tac x = "\x\A. " in rexI, simp) apply(rule_tac lam_closed, auto intro:id_replacement dest:transM) done -(*Could probably weaken the premise to well_ord(K,r), or remove using AC*) lemma cmult_rel_square_le: "Card\<^bsup>M\<^esup>(K) \ M(K) \ K \ K \\<^bsup>M\<^esup> K" apply (unfold cmult_rel_def) apply (rule le_trans) apply (rule_tac [2] well_ord_lepoll_rel_imp_cardinal_rel_le) apply (rule_tac [3] prod_square_lepoll_rel) apply (simp add: le_refl Card_rel_is_Ord Card_rel_cardinal_rel_eq) apply (blast intro: well_ord_rmult well_ord_Memrel Card_rel_is_Ord) apply simp_all done subsubsection\Multiplication by a non-zero cardinal\ lemma prod_lepoll_rel_self: "b \ B \ M(b) \ M(B) \ M(A) \ A \\<^bsup>M\<^esup> A*B" apply (simp add: def_lepoll_rel inj_def) apply (rule_tac x = "\x\A. " in rexI, simp) apply(rule_tac lam_closed, auto intro:pospend_replacement dest:transM) done -(*Could probably weaken the premises to well_ord(K,r), or removing using AC*) lemma cmult_rel_le_self: "[| Card\<^bsup>M\<^esup>(K); Ord(L); 0 K \ (K \\<^bsup>M\<^esup> L)" apply (unfold cmult_rel_def) apply (rule le_trans [OF Card_rel_cardinal_rel_le well_ord_lepoll_rel_imp_cardinal_rel_le]) apply assumption apply simp apply (blast intro: well_ord_rmult well_ord_Memrel Card_rel_is_Ord) apply (auto intro: prod_lepoll_rel_self ltD) done subsubsection\Monotonicity of multiplication\ lemma prod_lepoll_rel_mono: "[| A \\<^bsup>M\<^esup> C; B \\<^bsup>M\<^esup> D; M(A); M(B); M(C); M(D)|] ==> A * B \\<^bsup>M\<^esup> C * D" apply (simp add:def_lepoll_rel) apply (elim rexE) apply (rule_tac x = "lam :A*B. " in rexI) apply (rule_tac d = "%. " in lam_injective) apply (typecheck add: inj_is_fun, auto) apply(rule_tac lam_closed, auto intro:prod_fun_replacement dest:transM) done lemma cmult_rel_le_mono: "[| K' \ K; L' \ L;M(K');M(K);M(L');M(L) |] ==> (K' \\<^bsup>M\<^esup> L') \ (K \\<^bsup>M\<^esup> L)" apply (unfold cmult_rel_def) apply (safe dest!: le_subset_iff [THEN iffD1]) apply (rule well_ord_lepoll_rel_imp_cardinal_rel_le) apply (blast intro: well_ord_rmult well_ord_Memrel) apply (auto intro: prod_lepoll_rel_mono subset_imp_lepoll_rel) done subsection\Multiplication of finite cardinals is "ordinary" multiplication\ lemma prod_succ_eqpoll_rel: "M(A) \ M(B) \ succ(A)*B \\<^bsup>M\<^esup> B + A*B" apply (simp add: def_eqpoll_rel) apply (rule rexI) apply (rule_tac c = "\p. if fst(p)=A then Inl (snd(p)) else Inr (p)" and d = "case (%y. , %z. z)" in lam_bijective) apply safe apply (simp_all add: succI2 if_type mem_imp_not_eq) apply(rule_tac lam_closed, auto intro:Inl_replacement2 dest:transM) done -(*Unconditional version requires AC*) lemma cmult_rel_succ_lemma: "[| Ord(m); Ord(n) ; M(m); M(n) |] ==> succ(m) \\<^bsup>M\<^esup> n = n \\<^bsup>M\<^esup> (m \\<^bsup>M\<^esup> n)" apply (simp add: cmult_rel_def cadd_rel_def) apply (rule prod_succ_eqpoll_rel [THEN cardinal_rel_cong, THEN trans], simp_all) apply (rule cardinal_rel_cong [symmetric], simp_all) apply (rule sum_eqpoll_rel_cong [OF eqpoll_rel_refl well_ord_cardinal_rel_eqpoll_rel], assumption) apply (blast intro: well_ord_rmult well_ord_Memrel) apply simp_all done lemma nat_cmult_rel_eq_mult: "[| m \ nat; n \ nat |] ==> m \\<^bsup>M\<^esup> n = m#*n" using transM[OF _ M_nat] apply (induct_tac m) apply (simp_all add: cmult_rel_succ_lemma nat_cadd_rel_eq_add) done lemma cmult_rel_2: "Card\<^bsup>M\<^esup>(n) \ M(n) \ 2 \\<^bsup>M\<^esup> n = n \\<^bsup>M\<^esup> n" by (simp add: cmult_rel_succ_lemma Card_rel_is_Ord cadd_rel_commute [of _ 0]) lemma sum_lepoll_rel_prod: assumes C: "2 \\<^bsup>M\<^esup> C" and types:"M(C)" "M(B)" shows "B+B \\<^bsup>M\<^esup> C*B" proof - have "B+B \\<^bsup>M\<^esup> 2*B" by (simp add: sum_eq_2_times types) also have "... \\<^bsup>M\<^esup> C*B" by (blast intro: prod_lepoll_rel_mono lepoll_rel_refl C types) finally show "B+B \\<^bsup>M\<^esup> C*B" by (simp_all add:types) qed lemma lepoll_imp_sum_lepoll_prod: "[| A \\<^bsup>M\<^esup> B; 2 \\<^bsup>M\<^esup> A; M(A) ;M(B) |] ==> A+B \\<^bsup>M\<^esup> A*B" by (blast intro: sum_lepoll_rel_mono sum_lepoll_rel_prod lepoll_rel_trans lepoll_rel_refl) end \ \\<^locale>\M_cardinals\\ subsection\Infinite Cardinals are Limit Ordinals\ -(*This proof is modelled upon one assuming nat<=A, with injection - \z\cons(u,A). if z=u then 0 else if z \ nat then succ(z) else z - and inverse %y. if y \ nat then nat_case(u, %z. z, y) else y. \ - If f \ inj(nat,A) then range(f) behaves like the natural numbers.*) - - context M_pre_cardinal_arith begin lemma nat_cons_lepoll_rel: "nat \\<^bsup>M\<^esup> A \ M(A) \ M(u) ==> cons(u,A) \\<^bsup>M\<^esup> A" apply (simp add: def_lepoll_rel) apply (erule rexE) apply (rule_tac x = "\z\cons (u,A). if z=u then f`0 else if z \ range (f) then f`succ (converse (f) `z) else z" in rexI) apply (rule_tac d = "%y. if y \ range(f) then nat_case (u, %z. f`z, converse(f) `y) else y" in lam_injective) apply (fast intro!: if_type apply_type intro: inj_is_fun inj_converse_fun) apply (simp add: inj_is_fun [THEN apply_rangeI] inj_converse_fun [THEN apply_rangeI] inj_converse_fun [THEN apply_funtype]) proof - fix f assume "M(A)" "M(f)" "M(u)" then show "M(\z\cons(u, A). if z = u then f ` 0 else if z \ range(f) then f ` succ(converse(f) ` z) else z)" using if_then_range_replacement transM[OF _ \M(A)\] by (rule_tac lam_closed, auto) qed lemma nat_cons_eqpoll_rel: "nat \\<^bsup>M\<^esup> A ==> M(A) \ M(u) \ cons(u,A) \\<^bsup>M\<^esup> A" apply (erule nat_cons_lepoll_rel [THEN eqpoll_relI], assumption+) apply (rule subset_consI [THEN subset_imp_lepoll_rel], simp_all) done lemma nat_succ_eqpoll_rel: "nat \ A ==> M(A) \ succ(A) \\<^bsup>M\<^esup> A" apply (unfold succ_def) apply (erule subset_imp_lepoll_rel [THEN nat_cons_eqpoll_rel], simp_all) done lemma InfCard_rel_nat: "InfCard\<^bsup>M\<^esup>(nat)" apply (simp add: InfCard_rel_def) apply (blast intro: Card_rel_nat Card_rel_is_Ord) done lemma InfCard_rel_is_Card_rel: "M(K) \ InfCard\<^bsup>M\<^esup>(K) \ Card\<^bsup>M\<^esup>(K)" apply (simp add: InfCard_rel_def) done lemma InfCard_rel_Un: "[| InfCard\<^bsup>M\<^esup>(K); Card\<^bsup>M\<^esup>(L); M(K); M(L) |] ==> InfCard\<^bsup>M\<^esup>(K \ L)" apply (simp add: InfCard_rel_def) apply (simp add: Card_rel_Un Un_upper1_le [THEN [2] le_trans] Card_rel_is_Ord) done lemma InfCard_rel_is_Limit: "InfCard\<^bsup>M\<^esup>(K) ==> M(K) \ Limit(K)" apply (simp add: InfCard_rel_def) apply (erule conjE) apply (frule Card_rel_is_Ord, assumption) apply (rule ltI [THEN non_succ_LimitI]) apply (erule le_imp_subset [THEN subsetD]) apply (safe dest!: Limit_nat [THEN Limit_le_succD]) apply (unfold Card_rel_def) apply (drule trans) apply (erule le_imp_subset [THEN nat_succ_eqpoll_rel, THEN cardinal_rel_cong], simp_all) apply (erule Ord_cardinal_rel_le [THEN lt_trans2, THEN lt_irrefl], assumption) apply (rule le_eqI) prefer 2 apply (rule Ord_cardinal_rel, assumption+) done end \ \\<^locale>\M_pre_cardinal_arith\\ -(*** An infinite cardinal equals its square (Kunen, Thm 10.12, page 29) ***) - lemma (in M_ordertype) ordertype_abs[absolut]: assumes "wellordered(M,A,r)" "M(A)" "M(r)" "M(i)" shows "otype(M,A,r,i) \ i = ordertype(A,r)" proof - have eq:"j = k" if "g \ \A, r\ \ \j, Memrel(j)\" "Ord(j)" "h \ \A, r\ \ \k, Memrel(k)\" "Ord(k)" for j k g h proof - from that have "h O converse(g) \ \j, Memrel(j)\ \ \k, Memrel(k)\" using ord_iso_sym ord_iso_trans by blast with \Ord(j)\ \Ord(k)\ show "j=k" using Ord_iso_implies_eq[of j k "h O converse(g)"] by simp qed have otypeE: "\ h . h \ \A, r\ \ \k, Memrel(k)\" if "otype(M, A, r, k)" "M(k)" for k proof - note that assms moreover from this obtain h where "omap(M, A, r, h)" "M(h)" using that omap_exists[of A r] by auto moreover from calculation have "h \ \A, r\ \ \k, Memrel(k)\" using that omap_ord_iso obase_equals by simp then show ?thesis .. qed show ?thesis proof(intro iffI) note assms moreover assume "otype(M, A, r, i)" moreover from calculation obtain f j where "M(f)" "Ord(j)" "f \ \A, r\ \ \j, Memrel(j)\" using ordertype_exists[of A r] by auto ultimately show "i = ordertype(A, r)" using Ord_otype[OF _ well_ord_is_trans_on] eq[of f j _ i] ordertypes_are_absolute otypeE[of i] by auto next note assms moreover assume "i = ordertype(A, r)" moreover from calculation obtain j where "otype(M,A,r,j)" "M(j)" using otype_exists by auto ultimately show "otype(M, A, r, i)" using Ord_ordertype eq[of _ i _ j] Ord_otype[OF _ well_ord_is_trans_on] otypeE[of j] ordertype_ord_iso[of A r] by auto qed qed lemma (in M_ordertype) ordertype_closed[intro,simp]: "\ wellordered(M,A,r);M(A);M(r)\ \ M(ordertype(A,r))" using ordertype_exists ordertypes_are_absolute by blast -(* -This apparent duplication of definitions is needed because in ZF-Constructible -pairs are in their absolute version and this breaks the synthesis of formulas. -*) +\ \This apparent duplication of definitions is needed because in \<^session>\ZF-Constructible\ +pairs are in their absolute version and this breaks the synthesis of formulas.\ relationalize "transitive_rel" "is_transitive" external synthesize "is_transitive" from_definition assuming "nonempty" arity_theorem for "is_transitive_fm" lemma (in M_trivial) is_transitive_iff_transitive_rel: "M(A)\ M(r) \ transitive_rel(M, A, r) \ is_transitive(M,A, r)" unfolding transitive_rel_def is_transitive_def by simp relationalize "linear_rel" "is_linear" external synthesize "is_linear" from_definition assuming "nonempty" arity_theorem for "is_linear_fm" reldb_add relational "transitive_on" "transitive_rel" reldb_add relational "linear_on" "linear_rel" reldb_add functional "transitive_on" "transitive_rel" reldb_add functional "linear_on" "linear_rel" lemma (in M_trivial) is_linear_iff_linear_rel: "M(A)\ M(r) \ is_linear(M,A, r) \ linear_rel(M, A, r)" unfolding linear_rel_def is_linear_def by simp relationalize "wellfounded_on" "is_wellfounded_on" external synthesize "is_wellfounded_on" from_definition assuming "nonempty" arity_theorem for "is_wellfounded_on_fm" lemma (in M_trivial) is_wellfounded_on_iff_wellfounded_on: "M(A)\ M(r) \ is_wellfounded_on(M,A, r) \ wellfounded_on(M, A, r)" unfolding wellfounded_on_def is_wellfounded_on_def by simp definition is_well_ord :: "[i=>o,i,i]=>o" where \ \linear and wellfounded on \A\\ "is_well_ord(M,A,r) == is_transitive(M,A,r) \ is_linear(M,A,r) \ is_wellfounded_on(M,A,r)" lemma (in M_trivial) is_well_ord_iff_wellordered: "M(A)\ M(r) \ is_well_ord(M,A, r) \ wellordered(M, A, r)" using is_wellfounded_on_iff_wellfounded_on is_linear_iff_linear_rel is_transitive_iff_transitive_rel unfolding wellordered_def is_well_ord_def by simp synthesize "is_well_ord" from_definition assuming "nonempty" arity_theorem for "is_well_ord_fm" reldb_add relational "well_ord" "is_well_ord" reldb_add functional "well_ord" "well_ord" \ \One keyword (functional or relational) means going from an absolute term to that kind of term\ reldb_add relational "Order.pred" "pred_set" \ \The following form (twice the same argument) is only correct when an "\_abs" theorem is available\ reldb_add functional "Order.pred" "Order.pred" + +\ \Two keywords denote origin and destination, respectively\ (* -\ \Two keywords denote origin and destination, respectively\ reldb_add functional relational "Ord" "ordinal" *) relativize functional "ord_iso" "ord_iso_rel" external - \ \The following corresponds to "relativize functional relational"\ +\ \The following corresponds to "relativize functional relational"\ relationalize "ord_iso_rel" "is_ord_iso" context M_pre_cardinal_arith begin is_iff_rel for "ord_iso" using bij_rel_iff unfolding is_ord_iso_def ord_iso_rel_def by simp rel_closed for "ord_iso" using ord_iso_separation unfolding ord_iso_rel_def by simp end \ \\<^locale>\M_pre_cardinal_arith\\ synthesize "is_ord_iso" from_definition assuming "nonempty" lemma trans_on_iff_trans: "trans[A](r) \ trans(r \ A\A)" unfolding trans_on_def trans_def by auto lemma trans_on_subset: "trans[A](r) \ B \ A \ trans[B](r)" unfolding trans_on_def by auto lemma relation_Int: "relation(r \ B\B)" unfolding relation_def by auto text\Discipline for \<^term>\ordermap\\ relativize functional "ordermap" "ordermap_rel" external relationalize "ordermap_rel" "is_ordermap" lemma wfrec_on_pred_eq: assumes "r \ Pow(A\A)" shows "wfrec[A](r, x, \x f. f `` Order.pred(A, x, r)) = wfrec(r, x, \x f. f `` Order.pred(A, x, r))" proof - from \r \ Pow(A\A)\ have "r \ A\A = r" by auto then show ?thesis unfolding wfrec_on_def by simp qed context M_pre_cardinal_arith begin lemma wfrec_on_pred_closed: assumes "wf[A](r)" "trans[A](r)" "r \ Pow\<^bsup>M\<^esup>(A\A)" "M(A)" "x \ A" shows "M(wfrec(r, x, \x f. f `` Order.pred(A, x, r)))" proof - note assms moreover from this have "M(r)" "M(x)" using transM[of _ "Pow\<^bsup>M\<^esup>(A\A)"] transM[of _ A] by simp_all moreover from calculation have "wfrec[A](r, x, \x f. f `` Order.pred(A, x, r)) = wfrec(r, x, \x f. f `` Order.pred(A, x, r))" using wfrec_on_pred_eq Pow_rel_char by simp moreover from calculation have "M(wfrec(r, x, \x f. f `` Order.pred(A, x, r)))" using wfrec_pred_replacement wf_on_imp_wf trans_on_imp_trans subset_Sigma_imp_relation Pow_rel_char by (rule_tac MH="\x f b. \a[M]. image(M, f, a, b) \ pred_set(M, A, x, r, a)" in trans_wfrec_closed) (auto dest:transM simp:relation2_def) ultimately show ?thesis by simp qed lemma wfrec_on_pred_closed': assumes "wf[A](r)" "trans[A](r)" "r \ Pow\<^bsup>M\<^esup>(A\A)" "M(A)" "x \ A" shows "M(wfrec[A](r, x, \x f. f `` Order.pred(A, x, r)))" using assms wfrec_on_pred_closed wfrec_on_pred_eq Pow_rel_char by simp lemma ordermap_rel_closed[intro,simp]: assumes "wf[A](r)" "trans[A](r)" "r \ Pow\<^bsup>M\<^esup>(A\A)" "M(A)" shows "M(ordermap_rel(M, A, r))" proof - from assms have "r \ A\A = r" "r\Pow(A\A)" "M(r)" using Pow_rel_char by auto with assms have "wf(r)" "trans(r)" "relation(r)" unfolding wf_on_def using trans_on_iff_trans relation_def by auto with \M(A)\ \M(r)\ - have 1:"\ x z . M(x) \ M(z) \ + have 1:" (\y[M]. pair(M, x, y, z) \ is_wfrec(M, \x f z. z = f `` Order.pred(A, x, r), r, x, y)) \ - z = x f. f `` Order.pred(A, x, r))>" - using trans_wfrec_abs[of r,where + z = \x,wfrec(r,x,\x f. f `` Order.pred(A, x, r))\" if "M(x)" "M(z)" for x z + using that trans_wfrec_abs[of r,where H="\x f. f `` Order.pred(A, x, r)" and MH="\x f z . z= f `` Order.pred(A, x, r)",simplified] wfrec_pred_replacement unfolding relation2_def by auto with \M(A)\ \M(r)\ - have "strong_replacement(M,\x z. z = x f. f `` Order.pred(A, x, r))>)" + have "strong_replacement(M,\x z. z = \x,wfrec(r,x,\x f. f `` Order.pred(A, x, r))\)" using strong_replacement_cong[of M,OF 1,THEN iffD1,OF _ _ wfrec_pred_replacement[unfolded wfrec_replacement_def]] by simp with \M(A)\ \M(r)\ \r \ Pow(A\A)\ assms show ?thesis unfolding ordermap_rel_def using lam_cong[OF refl wfrec_on_pred_eq] wfrec_on_pred_closed lam_closed by auto qed lemma is_ordermap_iff: assumes "r \ Pow\<^bsup>M\<^esup>(A\A)" "wf[A](r)" "trans[A](r)" "M(A)" "M(res)" shows "is_ordermap(M, A, r, res) \ res = ordermap_rel(M, A, r)" proof - from \r \ _\ \M(A)\ - have "r \ A\A = r" "M(r)" "r \ Pow(A\A)" "\ x y . \r \ x\A" + have "r \ A\A = r" "M(r)" "r \ Pow(A\A)" "\ x y . \x,y\\r \ x\A" using Pow_rel_char by auto with assms - have 1:"wf(r)" "trans(r)" "relation(r)" + have "wf(r)" "trans(r)" "relation(r)" unfolding wf_on_def using trans_on_iff_trans relation_def by auto with \r\Pow(A\A)\ show ?thesis - using ordermap_rel_closed[of r A] assms wfrec_on_pred_closed wfrec_pred_replacement 1 + using ordermap_rel_closed[of r A] assms wfrec_on_pred_closed wfrec_pred_replacement unfolding is_ordermap_def ordermap_rel_def apply (rule_tac lambda_abs2,simp_all add:Relation1_def,clarify) - apply (rule trans_wfrec_on_abs) - apply (auto dest:transM simp add: relation_Int relation2_def) + apply (rule trans_wfrec_on_abs) + apply (auto dest:transM simp add: relation_Int relation2_def) by(rule_tac wfrec_on_pred_closed'[of A r],auto) qed end \ \\<^locale>\M_pre_cardinal_arith\\ text\Discipline for \<^term>\ordertype\\ relativize functional "ordertype" "ordertype_rel" external relationalize "ordertype_rel" "is_ordertype" definition is_order_body - where "is_order_body(M,X,r,z) \ \A[M]. cartprod(M,X,X,A) \ subset(M,r,A) \ M(z) \ M(r) \ - wellordered(M,X,r) \ is_ordertype(M,X,r,z)" + where "is_order_body(M,p,z) \ \X[M].\r[M].\A[M]. is_fst(M,p,X) \ is_snd(M,p,r) \ + cartprod(M,X,X,A) \ subset(M,r,A) \ M(z) \ M(r) \ is_well_ord(M,X,r) \ is_ordertype(M,X,r,z)" context M_pre_cardinal_arith begin lemma is_ordertype_iff: - assumes "r \ Pow\<^bsup>M\<^esup>(A\A)" "wf[A](r)" "trans[A](r)" + assumes "r \ Pow\<^bsup>M\<^esup>(A\A)" "well_ord(A,r)" shows "M(A) \ M(res) \ is_ordertype(M, A, r, res) \ res = ordertype_rel(M, A, r)" using assms is_ordermap_iff[of r A] trans_on_iff_trans ordermap_rel_closed[of A r] - unfolding is_ordertype_def ordertype_rel_def wf_on_def by simp - -lemma is_ordertype_iff': - assumes "r \ Pow\<^bsup>M\<^esup>(A\A)" "well_ord(A,r)" - shows "M(A) \ M(res) \ is_ordertype(M, A, r, res) \ res = ordertype_rel(M, A, r)" - using assms is_ordertype_iff Pow_rel_char - unfolding well_ord_def part_ord_def tot_ord_def by simp + unfolding is_ordertype_def ordertype_rel_def wf_on_def well_ord_def part_ord_def tot_ord_def + by simp lemma ordertype_rel_abs: assumes "wellordered(M,X,r)" "M(X)" "M(r)" shows "ordertype_rel(M,X,r) = ordertype(X,r)" using assms ordertypes_are_absolute[of X r] unfolding ordertype_def ordertype_rel_def ordermap_rel_def ordermap_def by simp 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) \ x\Pow\<^bsup>M\<^esup>(X\X) \ well_ord(X, x) \ z = ordertype(X, x)" - using is_ordertype_iff' ordertype_rel_abs well_ord_is_linear Pow_rel_char + "M(Xr) \ M(z) \ is_order_body(M, Xr, z) \ + snd(Xr)\Pow\<^bsup>M\<^esup>(fst(Xr)\fst(Xr)) \ well_ord(fst(Xr), snd(Xr)) \ z = ordertype(fst(Xr), snd(Xr))" + using is_ordertype_iff ordertype_rel_abs well_ord_is_linear Pow_rel_char + is_well_ord_iff_wellordered snd_abs fst_abs unfolding is_order_body_def by simp lemma well_ord_restr: "well_ord(X, r) \ well_ord(X, r \ X\X)" proof - have "r \ X\X \ X\X = r \ X\X" by auto moreover assume "well_ord(X, r)" ultimately show ?thesis unfolding well_ord_def tot_ord_def part_ord_def linear_def irrefl_def wf_on_def by simp_all (simp only: trans_on_def, blast) qed lemma ordertype_restr_eq : assumes "well_ord(X,r)" shows "ordertype(X, r) = ordertype(X, r \ X\X)" using ordermap_restr_eq assms unfolding ordertype_def by simp lemma ordertype_rel_closed[intro,simp]: assumes "well_ord(A,r)" "r \ Pow\<^bsup>M\<^esup>(A\A)" "M(A)" shows "M(ordertype_rel(M,A,r))" using assms Pow_rel_char unfolding well_ord_def tot_ord_def part_ord_def ordertype_rel_def by simp end \ \\<^locale>\M_pre_cardinal_arith\\ relativize functional "jump_cardinal_body" "jump_cardinal_body_rel" relationalize "jump_cardinal_body_rel" "is_jump_cardinal_body_rel" relativize functional "jump_cardinal_body'" "jump_cardinal_body'_rel" relationalize "jump_cardinal_body'_rel" "is_jump_cardinal_body'_rel" -\ \NOTE: not quite the same as \<^term>\jump_cardinal\, - note \<^term>\Pow(X*X)\.\ +text\Notice that this is not quite the same as \<^term>\jump_cardinal\: +observe \<^term>\Pow(X*X)\.\ + definition jump_cardinal' :: "i\i" where "jump_cardinal'(K) \ \X\Pow(K). {z. r \ Pow(X*X), z = jump_cardinal_body(Pow(X*X),X)}" relativize functional "jump_cardinal" "jump_cardinal_rel" external relativize functional "trans_on" "transitivie_rel" external relationalize "jump_cardinal_rel" "is_jump_cardinal" lemma (in M_ordertype) ordermap_closed[intro,simp]: assumes "wellordered(M,A,r)" and types:"M(A)" "M(r)" shows "M(ordermap(A,r))" proof - note assms moreover from this obtain i f where "Ord(i)" "f \ ord_iso(A, r, i, Memrel(i))" "M(i)" "M(f)" using ordertype_exists by blast moreover from calculation have "i = ordertype(A,r)" using ordertypes_are_absolute by force moreover from calculation have "ordermap(A,r) \ ord_iso(A, r, i, Memrel(i))" using ordertype_ord_iso by simp ultimately have "f = ordermap(A,r)" using well_ord_iso_unique by fastforce with \M(f)\ show ?thesis by simp qed context M_pre_cardinal_arith begin lemma def_jump_cardinal_body: "M(X) \ M(U) \ jump_cardinal_body(U,X) = {z . r \ U, M(z) \ well_ord(X, r) \ z = ordertype(X, r)}" unfolding jump_cardinal_body_def using ordertype_closed by(rule_tac Replace_cong,auto dest:transM) lemma jump_cardinal_body_abs: shows "M(A) \ jump_cardinal_body_rel(M,Pow\<^bsup>M\<^esup>(A\A),A) = jump_cardinal_body(Pow\<^bsup>M\<^esup>(A\A),A)" unfolding jump_cardinal_body_rel_def using def_jump_cardinal_body ordertype_rel_abs transM[of _ "Pow\<^bsup>M\<^esup>(A\A)"] by simp end \ \\<^locale>\M_pre_cardinal_arith\\ locale M_cardinal_arith = M_pre_cardinal_arith + assumes - ordertype_replacement_ax : - "M(X) \ strong_replacement(M,\ x z . is_well_ord(M,X, x) \ is_ordertype(M,X,x,z))" + is_ordertype_replacement : + "strong_replacement(M,\ x z . \y[M]. is_order_body(M,x,y) \ z=\x,y\)" and - strong_replacement_jc_body : - "strong_replacement(M,\ x z . z = jump_cardinal_body_rel(M,Pow\<^bsup>M\<^esup>(x\x),x))" + pow_rel_separation : "\A[M]. separation(M, \y. \x[M]. x \ A \ y = \x, Pow\<^bsup>M\<^esup>(x)\)" + and + separation_is_well_ord : "separation(M, \x . is_well_ord(M,fst(x),snd(x)))" begin -lemma jump_cardinal_closed_aux1: - assumes "M(X)" - shows "M(jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X),X))" +lemmas Pow_rel_replacement = lam_replacement_Pow_rel[OF pow_rel_separation] + +lemma ordtype_replacement : + "strong_replacement(M , \x z . (snd(x) \ Pow\<^bsup>M\<^esup>(fst(x)\fst(x)) \ well_ord(fst(x),snd(x))) \ + z =\x, ordertype(fst(x),snd(x))\)" + using strong_replacement_cong[THEN iffD1,OF _ is_ordertype_replacement] is_order_body_abs + by simp + +lemma separation_well_ord : "separation(M, \x . well_ord(fst(x),snd(x)))" + using separation_cong[THEN iffD1] separation_is_well_ord is_well_ord_iff_wellordered well_ord_abs + by simp + +lemma jump_cardinal_body_lam_replacement : + shows "lam_replacement(M, \X .jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X),X))" and + "M(X) \ M(jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X), X))" proof - - note \M(X)\ - moreover from this - have "univalent(M,Pow\<^bsup>M\<^esup>(X\X), - \r z. M(z) \ r\Pow\<^bsup>M\<^esup>(X\X) \ well_ord(X, r) \ is_ordertype(M, X, r, z))" - using - 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 + define WO where "WO \ \X . {r\Pow\<^bsup>M\<^esup>(X\X) . well_ord(X,r)}" + define lam_jump_cardinal_body where + "lam_jump_cardinal_body \ \X . \r\WO(X) . ordertype(X,r)" + have "lam_jump_cardinal_body(X) = {\r,ordertype(X,r)\ . r \ WO(X)}" for X + unfolding lam_jump_cardinal_body_def WO_def lam_def + by simp + then + have "jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X),X) = {snd(p) . p \ lam_jump_cardinal_body(X)}" if "M(X)" for X + unfolding jump_cardinal_body_def WO_def lam_def + by force + moreover + have "lam_replacement(M, \x. {y . r \ Pow\<^bsup>M\<^esup>(x \ x), (r \ Pow\<^bsup>M\<^esup>(x \ x) \ well_ord(x, r)) \ y = ordertype(x, r)})" + (is "lam_replacement(M,?q)") + using Pow_rel_replacement[THEN [2] lam_replacement_hcomp] + lam_replacement_CartProd lam_replacement_identity lam_replacement_fst lam_replacement_snd + fst_closed[OF transM] snd_closed[OF transM] + separation_well_ord separation_in separation_conj transM[of _ "Pow\<^bsup>M\<^esup>(_\_)"] + strong_lam_replacement_imp_lam_replacement_RepFun[OF ordtype_replacement, where g="\X. Pow\<^bsup>M\<^esup>(X \ X)"] + by force + moreover + have "M(?q(x))" if "M(x)" for x + using that Pow_rel_replacement[THEN [2] lam_replacement_hcomp] + lam_replacement_CartProd lam_replacement_identity + ordertype_closed transM[of _ "Pow\<^bsup>M\<^esup>(_\_)"] + fst_closed[OF transM] snd_closed[OF transM] + by(rule_tac strong_lam_replacement_imp_strong_replacement[OF ordtype_replacement,THEN strong_replacement_closed], + auto) + moreover + have "?q(x)={snd(p) . p \ (\r\WO(x). ordertype(x, r))}" if "M(x)" for x + using that + unfolding lam_def WO_def + by force + moreover from calculation + have "?q(X) = jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X), X)" if "M(X)" for X + using that + unfolding lam_jump_cardinal_body_def by simp moreover from calculation - have "jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X),X) = - jump_cardinal_body_rel(M,Pow\<^bsup>M\<^esup>(X \ X),X)" - using jump_cardinal_body_abs - by simp + have "M(jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X), X))" if "M(X)" for X + using that by simp moreover from calculation - have "jump_cardinal_body_rel(M,Pow\<^bsup>M\<^esup>(X \ X),X) = - {z . r \ Pow\<^bsup>M\<^esup>(X \ X), M(z) \ M(r) \ well_ord(X, r) \ is_ordertype(M,X, r,z)}" - using ordertype_rel_closed transM[of _ "Pow\<^bsup>M\<^esup>(X\X)"] is_ordertype_iff' - unfolding jump_cardinal_body_rel_def + have 1:"lam_replacement(M, \x . {snd(p) . p \ lam_jump_cardinal_body(x)})" + using lam_replacement_cong by auto - moreover from calculation - have "M({z . r \ Pow\<^bsup>M\<^esup>(X \ X), M(z) \ M(r) \ well_ord(X, r) \ is_ordertype(M,X, r,z)})" - using ordertype_replacement_ax is_well_ord_iff_wellordered - by(rule_tac strong_replacement_closed,auto) ultimately - show ?thesis by simp + show "lam_replacement(M, \X .jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X),X))" and + "M(X) \ M(jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X), X))" + using lam_replacement_imp_strong_replacement[OF lam_replacement_snd] + lam_replacement_cong[OF 1,unfolded lam_jump_cardinal_body_def] + by simp_all qed -lemma jump_cardinal_body_closed: +lemmas jump_cardinal_body_closed = jump_cardinal_body_lam_replacement(2) + +lemma jump_cardinal_closed: assumes "M(K)" shows "M({jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X),X) . X \ Pow\<^bsup>M\<^esup>(K)})" -proof - - have "univalent(M,X,\ x z . z = jump_cardinal_body(Pow\<^bsup>M\<^esup>(x \ x),x))" if "M(X)" for X - using transM[of _ X] jump_cardinal_closed_aux1 - by auto - then - show ?thesis - using assms jump_cardinal_closed_aux1 strong_replacement_jc_body - transM[of _ "Pow\<^bsup>M\<^esup>(K)"] jump_cardinal_body_abs - by(rule_tac RepFun_closed,auto) -qed + using assms jump_cardinal_body_lam_replacement lam_replacement_imp_strong_replacement + transM[of _ "Pow\<^bsup>M\<^esup>(K)"] RepFun_closed + by auto end \ \\<^locale>\M_cardinal_arith\\ context M_pre_cardinal_arith begin (*A general fact about ordermap*) lemma ordermap_eqpoll_pred: "[| well_ord(A,r); x \ A ; M(A);M(r);M(x)|] ==> ordermap(A,r)`x \\<^bsup>M\<^esup> Order.pred(A,x,r)" apply (simp add: def_eqpoll_rel) apply (rule rexI) - apply (simp add: ordermap_eq_image well_ord_is_wf) - apply (erule ordermap_bij [THEN bij_is_inj, THEN restrict_bij, + apply (simp add: ordermap_eq_image well_ord_is_wf) + apply (erule ordermap_bij [THEN bij_is_inj, THEN restrict_bij, THEN bij_converse_bij]) - apply (rule pred_subset, simp) + apply (rule pred_subset, simp) done text\Kunen: "each \<^term>\\x,y\ \ K \ K\ has no more than \<^term>\z \ z\ predecessors..." (page 29)\ lemma ordermap_csquare_le: assumes K: "Limit(K)" and x: "x K, csquare_rel(K)) ` \x,y\|\<^bsup>M\<^esup> \ |succ(succ(x \ y))|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> |succ(succ(x \ y))|\<^bsup>M\<^esup>" using types proof (simp add: cmult_rel_def, rule_tac well_ord_lepoll_rel_imp_cardinal_rel_le) let ?z="succ(x \ y)" show "well_ord(|succ(?z)|\<^bsup>M\<^esup> \ |succ(?z)|\<^bsup>M\<^esup>, rmult(|succ(?z)|\<^bsup>M\<^esup>, Memrel(|succ(?z)|\<^bsup>M\<^esup>), |succ(?z)|\<^bsup>M\<^esup>, Memrel(|succ(?z)|\<^bsup>M\<^esup>)))" by (blast intro: well_ord_Memrel well_ord_rmult types) next let ?z="succ(x \ y)" have zK: "?z K, csquare_rel(K)))" using well_ord_csquare Limit_is_Ord by fastforce then have "ordermap(K \ K, csquare_rel(K)) ` \x,y\ \\<^bsup>M\<^esup> ordermap(K \ K, csquare_rel(K)) ` \?z,?z\" by (blast intro: ordermap_z_lt leI le_imp_lepoll_rel K x y types) also have "... \\<^bsup>M\<^esup> Order.pred(K \ K, \?z,?z\, csquare_rel(K))" proof (rule ordermap_eqpoll_pred) show "well_ord(K \ K, csquare_rel(K))" using K by (rule Limit_is_Ord [THEN well_ord_csquare]) next show "\?z, ?z\ \ K \ K" using zK by (blast intro: ltD) qed (simp_all add:types) also have "... \\<^bsup>M\<^esup> succ(?z) \ succ(?z)" using zK by (rule_tac pred_csquare_subset [THEN subset_imp_lepoll_rel]) (simp_all add:types) also have "... \\<^bsup>M\<^esup> |succ(?z)|\<^bsup>M\<^esup> \ |succ(?z)|\<^bsup>M\<^esup>" using oz by (blast intro: prod_eqpoll_rel_cong Ord_cardinal_rel_eqpoll_rel eqpoll_rel_sym types) finally show "ordermap(K \ K, csquare_rel(K)) ` \x,y\ \\<^bsup>M\<^esup> |succ(?z)|\<^bsup>M\<^esup> \ |succ(?z)|\<^bsup>M\<^esup>" by (simp_all add:types Mom) from Mom show "M(ordermap(K \ K, csquare_rel(K)) ` \x, y\)" by (simp_all add:types) qed (simp_all add:types) text\Kunen: "... so the order type is \\\ K"\ lemma ordertype_csquare_le_M: assumes IK: "InfCard\<^bsup>M\<^esup>(K)" and eq: "\y. y\K \ InfCard\<^bsup>M\<^esup>(y) \ M(y) \ y \\<^bsup>M\<^esup> y = y" \ \Note the weakened hypothesis @{thm eq}\ and types: "M(K)" shows "ordertype(K*K, csquare_rel(K)) \ K" proof - have CK: "Card\<^bsup>M\<^esup>(K)" using IK by (rule_tac InfCard_rel_is_Card_rel) (simp_all add:types) hence OK: "Ord(K)" by (rule Card_rel_is_Ord) (simp_all add:types) moreover have "Ord(ordertype(K \ K, csquare_rel(K)))" using OK by (rule well_ord_csquare [THEN Ord_ordertype]) ultimately show ?thesis proof (rule all_lt_imp_le) fix i assume i:"i < ordertype(K \ K, csquare_rel(K))" hence Oi: "Ord(i)" by (elim ltE) obtain x y where x: "x \ K" and y: "y \ K" and ieq: "i = ordermap(K \ K, csquare_rel(K)) ` \x,y\" using i by (auto simp add: ordertype_unfold elim: ltE) hence xy: "Ord(x)" "Ord(y)" "x < K" "y < K" using OK by (blast intro: Ord_in_Ord ltI)+ hence ou: "Ord(x \ y)" by (simp) from OK types have "M(ordertype(K \ K, csquare_rel(K)))" using well_ord_csquare by fastforce with i x y types have types': "M(K)" "M(i)" "M(x)" "M(y)" using types by (auto dest:transM ltD) show "i < K" proof (rule Card_rel_lt_imp_lt [OF _ Oi CK]) have "|i|\<^bsup>M\<^esup> \ |succ(succ(x \ y))|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> |succ(succ(x \ y))|\<^bsup>M\<^esup>" using IK xy by (auto simp add: ieq types intro: InfCard_rel_is_Limit [THEN ordermap_csquare_le] types') moreover have "|succ(succ(x \ y))|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> |succ(succ(x \ y))|\<^bsup>M\<^esup> < K" proof (cases rule: Ord_linear2 [OF ou Ord_nat]) assume "x \ y < nat" hence "|succ(succ(x \ y))|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> |succ(succ(x \ y))|\<^bsup>M\<^esup> \ nat" by (simp add: lt_def nat_cmult_rel_eq_mult nat_succI nat_into_Card_rel [THEN Card_rel_cardinal_rel_eq] types') also have "... \ K" using IK by (simp add: InfCard_rel_def le_imp_subset types) finally show "|succ(succ(x \ y))|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> |succ(succ(x \ y))|\<^bsup>M\<^esup> < K" by (simp add: ltI OK) next assume natxy: "nat \ x \ y" hence seq: "|succ(succ(x \ y))|\<^bsup>M\<^esup> = |x \ y|\<^bsup>M\<^esup>" using xy by (simp add: le_imp_subset nat_succ_eqpoll_rel [THEN cardinal_rel_cong] le_succ_iff types') also have "... < K" using xy by (simp add: Un_least_lt Ord_cardinal_rel_le [THEN lt_trans1] types') finally have "|succ(succ(x \ y))|\<^bsup>M\<^esup> < K" . moreover have "InfCard\<^bsup>M\<^esup>(|succ(succ(x \ y))|\<^bsup>M\<^esup>)" using xy natxy by (simp add: seq InfCard_rel_def nat_le_cardinal_rel types') ultimately show ?thesis by (simp add: eq ltD types') qed ultimately show "|i|\<^bsup>M\<^esup> < K" by (blast intro: lt_trans1) qed (simp_all add:types') qed qed (*Main result: Kunen's Theorem 10.12*) lemma InfCard_rel_csquare_eq: assumes IK: "InfCard\<^bsup>M\<^esup>(K)" and types: "M(K)" shows "K \\<^bsup>M\<^esup> K = K" proof - have OK: "Ord(K)" using IK by (simp add: Card_rel_is_Ord InfCard_rel_is_Card_rel types) from OK assms show "K \\<^bsup>M\<^esup> K = K" proof (induct rule: trans_induct) case (step i) note types = \M(K)\ \M(i)\ show "i \\<^bsup>M\<^esup> i = i" proof (rule le_anti_sym) from step types have Mot:"M(ordertype(i \ i, csquare_rel(i)))" "M(ordermap(i \ i, csquare_rel(i)))" using well_ord_csquare Limit_is_Ord by simp_all then have "|i \ i|\<^bsup>M\<^esup> = |ordertype(i \ i, csquare_rel(i))|\<^bsup>M\<^esup>" by (rule_tac cardinal_rel_cong, simp_all add: step.hyps well_ord_csquare [THEN ordermap_bij, THEN bij_imp_eqpoll_rel] types) with Mot have "i \\<^bsup>M\<^esup> i \ ordertype(i \ i, csquare_rel(i))" by (simp add: step.hyps cmult_rel_def Ord_cardinal_rel_le well_ord_csquare [THEN Ord_ordertype] types) moreover have "ordertype(i \ i, csquare_rel(i)) \ i" using step by (rule_tac ordertype_csquare_le_M) (simp add: types) ultimately show "i \\<^bsup>M\<^esup> i \ i" by (rule le_trans) next show "i \ i \\<^bsup>M\<^esup> i" using step by (blast intro: cmult_rel_square_le InfCard_rel_is_Card_rel) qed qed qed (*Corollary for arbitrary well-ordered sets (all sets, assuming AC)*) lemma well_ord_InfCard_rel_square_eq: assumes r: "well_ord(A,r)" and I: "InfCard\<^bsup>M\<^esup>(|A|\<^bsup>M\<^esup>)" and types: "M(A)" "M(r)" shows "A \ A \\<^bsup>M\<^esup> A" proof - have "A \ A \\<^bsup>M\<^esup> |A|\<^bsup>M\<^esup> \ |A|\<^bsup>M\<^esup>" by (blast intro: prod_eqpoll_rel_cong well_ord_cardinal_rel_eqpoll_rel eqpoll_rel_sym r types) also have "... \\<^bsup>M\<^esup> A" proof (rule well_ord_cardinal_rel_eqE [OF _ r]) show "well_ord(|A|\<^bsup>M\<^esup> \ |A|\<^bsup>M\<^esup>, rmult(|A|\<^bsup>M\<^esup>, Memrel(|A|\<^bsup>M\<^esup>), |A|\<^bsup>M\<^esup>, Memrel(|A|\<^bsup>M\<^esup>)))" by (blast intro: well_ord_rmult well_ord_Memrel r types) next show "||A|\<^bsup>M\<^esup> \ |A|\<^bsup>M\<^esup>|\<^bsup>M\<^esup> = |A|\<^bsup>M\<^esup>" using InfCard_rel_csquare_eq I by (simp add: cmult_rel_def types) qed (simp_all add:types) finally show ?thesis by (simp_all add:types) qed lemma InfCard_rel_square_eqpoll: assumes "InfCard\<^bsup>M\<^esup>(K)" and types:"M(K)" shows "K \ K \\<^bsup>M\<^esup> K" using assms apply (rule_tac well_ord_InfCard_rel_square_eq) - apply (erule InfCard_rel_is_Card_rel [THEN Card_rel_is_Ord, THEN well_ord_Memrel]) - apply (simp_all add: InfCard_rel_is_Card_rel [THEN Card_rel_cardinal_rel_eq] types) + apply (erule InfCard_rel_is_Card_rel [THEN Card_rel_is_Ord, THEN well_ord_Memrel]) + apply (simp_all add: InfCard_rel_is_Card_rel [THEN Card_rel_cardinal_rel_eq] types) done lemma Inf_Card_rel_is_InfCard_rel: "[| Card\<^bsup>M\<^esup>(i); ~ Finite_rel(M,i) ; M(i) |] ==> InfCard\<^bsup>M\<^esup>(i)" by (simp add: InfCard_rel_def Card_rel_is_Ord [THEN nat_le_infinite_Ord]) subsubsection\Toward's Kunen's Corollary 10.13 (1)\ lemma InfCard_rel_le_cmult_rel_eq: "[| InfCard\<^bsup>M\<^esup>(K); L \ K; 0 K \\<^bsup>M\<^esup> L = K" apply (rule le_anti_sym) - prefer 2 - apply (erule ltE, blast intro: cmult_rel_le_self InfCard_rel_is_Card_rel) + prefer 2 + apply (erule ltE, blast intro: cmult_rel_le_self InfCard_rel_is_Card_rel) apply (frule InfCard_rel_is_Card_rel [THEN Card_rel_is_Ord, THEN le_refl]) prefer 3 - apply (rule cmult_rel_le_mono [THEN le_trans], assumption+) - apply (simp_all add: InfCard_rel_csquare_eq) + apply (rule cmult_rel_le_mono [THEN le_trans], assumption+) + apply (simp_all add: InfCard_rel_csquare_eq) done (*Corollary 10.13 (1), for cardinal multiplication*) lemma InfCard_rel_cmult_rel_eq: "[| InfCard\<^bsup>M\<^esup>(K); InfCard\<^bsup>M\<^esup>(L); M(K) ; M(L) |] ==> K \\<^bsup>M\<^esup> L = K \ L" apply (rule_tac i = K and j = L in Ord_linear_le) - apply (typecheck add: InfCard_rel_is_Card_rel Card_rel_is_Ord) - apply (rule cmult_rel_commute [THEN ssubst]) prefer 3 - apply (rule Un_commute [THEN ssubst]) - apply (simp_all add: InfCard_rel_is_Limit [THEN Limit_has_0] InfCard_rel_le_cmult_rel_eq + apply (typecheck add: InfCard_rel_is_Card_rel Card_rel_is_Ord) + apply (rule cmult_rel_commute [THEN ssubst]) prefer 3 + apply (rule Un_commute [THEN ssubst]) + apply (simp_all add: InfCard_rel_is_Limit [THEN Limit_has_0] InfCard_rel_le_cmult_rel_eq subset_Un_iff2 [THEN iffD1] le_imp_subset) done lemma InfCard_rel_cdouble_eq: "InfCard\<^bsup>M\<^esup>(K) \ M(K) \ K \\<^bsup>M\<^esup> K = K" apply (simp add: cmult_rel_2 [symmetric] InfCard_rel_is_Card_rel cmult_rel_commute) apply (simp add: InfCard_rel_le_cmult_rel_eq InfCard_rel_is_Limit Limit_has_0 Limit_has_succ) done (*Corollary 10.13 (1), for cardinal addition*) lemma InfCard_rel_le_cadd_rel_eq: "[| InfCard\<^bsup>M\<^esup>(K); L \ K ; M(K) ; M(L)|] ==> K \\<^bsup>M\<^esup> L = K" apply (rule le_anti_sym) - prefer 2 - apply (erule ltE, blast intro: cadd_rel_le_self InfCard_rel_is_Card_rel) + prefer 2 + apply (erule ltE, blast intro: cadd_rel_le_self InfCard_rel_is_Card_rel) apply (frule InfCard_rel_is_Card_rel [THEN Card_rel_is_Ord, THEN le_refl]) prefer 3 - apply (rule cadd_rel_le_mono [THEN le_trans], assumption+) - apply (simp_all add: InfCard_rel_cdouble_eq) + apply (rule cadd_rel_le_mono [THEN le_trans], assumption+) + apply (simp_all add: InfCard_rel_cdouble_eq) done lemma InfCard_rel_cadd_rel_eq: "[| InfCard\<^bsup>M\<^esup>(K); InfCard\<^bsup>M\<^esup>(L); M(K) ; M(L) |] ==> K \\<^bsup>M\<^esup> L = K \ L" apply (rule_tac i = K and j = L in Ord_linear_le) - apply (typecheck add: InfCard_rel_is_Card_rel Card_rel_is_Ord) - apply (rule cadd_rel_commute [THEN ssubst]) prefer 3 - apply (rule Un_commute [THEN ssubst]) - apply (simp_all add: InfCard_rel_le_cadd_rel_eq subset_Un_iff2 [THEN iffD1] le_imp_subset) + apply (typecheck add: InfCard_rel_is_Card_rel Card_rel_is_Ord) + apply (rule cadd_rel_commute [THEN ssubst]) prefer 3 + apply (rule Un_commute [THEN ssubst]) + apply (simp_all add: InfCard_rel_le_cadd_rel_eq subset_Un_iff2 [THEN iffD1] le_imp_subset) done (*The other part, Corollary 10.13 (2), refers to the cardinality of the set of all n-tuples of elements of K. A better version for the Isabelle theory might be InfCard(K) ==> |list(K)| = K. *) end \ \\<^locale>\M_pre_cardinal_arith\\ subsection\For Every Cardinal Number There Exists A Greater One\ text\This result is Kunen's Theorem 10.16, which would be trivial using AC\ locale M_cardinal_arith_jump = M_cardinal_arith + M_ordertype begin lemma def_jump_cardinal_rel_aux: "X \ Pow\<^bsup>M\<^esup>(K) \ M(K) \ {z . r \ Pow\<^bsup>M\<^esup>(X \ X), M(z) \ well_ord(X, r) \ z = ordertype(X, r)} = {z . r \ Pow\<^bsup>M\<^esup>(K \ K), M(z) \ well_ord(X, r) \ z = ordertype(X, r)}" proof(rule,auto simp:Pow_rel_char dest:transM) let ?L="{z . r \ Pow\<^bsup>M\<^esup>(X \ X), M(z) \ well_ord(X, r) \ z = ordertype(X, r)}" let ?R="{z . r \ Pow\<^bsup>M\<^esup>(K \ K), M(z) \ well_ord(X, r) \ z = ordertype(X, r)}" show "ordertype(X, r) \ {y . x \ {x \ Pow(X \ X) . M(x)}, M(y) \ well_ord(X, x) \ y = ordertype(X, x)}" if "M(K)" "M(r)" "r\K\K" "X\K" "M(X)" "well_ord(X,r)" for r proof - from that have "ordertype(X,r) = ordertype(X,r\X\X)" "(r\X\X)\X\X" "M(r\X\X)" "well_ord(X,r\X\X)" "wellordered(M,X,r\X\X)" using well_ord_restr ordertype_restr_eq by auto moreover from this have "ordertype(X,r\X\X) \ ?L" using that Pow_rel_char ReplaceI[of "\ z r . M(z) \ well_ord(X, r) \ z = ordertype(X, r)" "ordertype(X,r\X\X)"] by auto ultimately show ?thesis using Pow_rel_char by auto qed qed lemma def_jump_cardinal_rel_aux2: assumes "X \ Pow\<^bsup>M\<^esup>(K)" "M(K)" shows "jump_cardinal_body(Pow\<^bsup>M\<^esup>(K\K),X) = jump_cardinal_body(Pow\<^bsup>M\<^esup>(X\X),X)" proof - from assms have "jump_cardinal_body(Pow\<^bsup>M\<^esup>(K\K),X) = {z . r \Pow\<^bsup>M\<^esup>(K\K), M(z) \ well_ord(X, r) \ z = ordertype(X, r)}" using def_jump_cardinal_body transM[of _ "Pow\<^bsup>M\<^esup>(K)"] by simp also from assms have "... = {z . r \Pow\<^bsup>M\<^esup>(X\X), M(z) \ well_ord(X, r) \ z = ordertype(X, r)}" using def_jump_cardinal_rel_aux transM[of _ "Pow\<^bsup>M\<^esup>(K)"] by simp also from assms have "... = jump_cardinal_body(Pow\<^bsup>M\<^esup>(X\X),X)" using def_jump_cardinal_body transM[of _ "Pow\<^bsup>M\<^esup>(K)"] by auto finally show ?thesis . qed lemma def_jump_cardinal_rel: assumes "M(K)" shows "jump_cardinal_rel(M,K) = (\{jump_cardinal_body(Pow\<^bsup>M\<^esup>(K*K),X) . X\ Pow\<^bsup>M\<^esup>(K)})" proof - have "jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X),X) = {z . r \ Pow\<^bsup>M\<^esup>(X \ X), M(z) \ well_ord(X, r) \ z=ordertype(X, r)}" if "M(X)" for X using that ordertype_closed transM[of _ "Pow\<^bsup>M\<^esup>(X\X)"] unfolding jump_cardinal_body_def by(intro equalityI subsetI,auto) then have "M({z . r \ Pow\<^bsup>M\<^esup>(X \ X), M(z) \ well_ord(X, r) \ z = ordertype(X, r)})" if "M(Y)" "X \ Pow\<^bsup>M\<^esup>(Y)" for Y X - using jump_cardinal_closed_aux1[of X] that transM[of _ "Pow\<^bsup>M\<^esup>(Y)"] + using jump_cardinal_body_closed[of X] that transM[of _ "Pow\<^bsup>M\<^esup>(Y)"] ordertype_closed transM[of _ "Pow\<^bsup>M\<^esup>(X\X)"] unfolding jump_cardinal_body_def by auto moreover from \M(K)\ - have "R \ Pow\<^bsup>M\<^esup>(X \ X) \ X \ Pow\<^bsup>M\<^esup>(K) \ R \ Pow\<^bsup>M\<^esup>(K \ K)" - for X R using mem_Pow_rel_abs transM[OF _ Pow_rel_closed, of R "X\X"] + have "R \ Pow\<^bsup>M\<^esup>(X \ X) \ X \ Pow\<^bsup>M\<^esup>(K) \ R \ Pow\<^bsup>M\<^esup>(K \ K)" for X R + using mem_Pow_rel_abs transM[OF _ Pow_rel_closed, of R "X\X"] transM[OF _ Pow_rel_closed, of X K] by auto ultimately show ?thesis using assms is_ordertype_iff ordertype_rel_abs transM[of _ "Pow\<^bsup>M\<^esup>(K)"] transM[of _ "Pow\<^bsup>M\<^esup>(K\K)"] def_jump_cardinal_rel_aux jump_cardinal_body_closed ordertype_closed unfolding jump_cardinal_rel_def jump_cardinal_body_def - apply(intro equalityI subsetI,auto dest:transM[of _ "Pow\<^bsup>M\<^esup>(K)"]) + apply(intro equalityI subsetI,auto dest:transM[of _ "Pow\<^bsup>M\<^esup>(K)"],rename_tac y r) by(rule_tac x="{z . xa \ Pow\<^bsup>M\<^esup>(K \ K), M(z) \ well_ord(y, xa) \ z = ordertype(y, xa)}" in bexI,auto) qed lemma Ord_jump_cardinal_rel: "M(K) \ Ord(jump_cardinal_rel(M,K))" apply (unfold def_jump_cardinal_rel jump_cardinal_body_def ) apply (rule Ord_is_Transset [THEN [2] OrdI]) - prefer 2 apply (blast intro!: Ord_ordertype) + prefer 2 apply (blast intro!: Ord_ordertype) apply (unfold Transset_def) apply (safe del: subsetI) apply (subst ordertype_pred_unfold, simp, safe) apply (rule UN_I) - apply (rule_tac [2] ReplaceI) - prefer 4 apply (blast intro: well_ord_subset elim!: predE, simp_all) - prefer 2 apply (blast intro: well_ord_subset elim!: predE) + apply (rule_tac [2] ReplaceI) + prefer 4 apply (blast intro: well_ord_subset elim!: predE, simp_all) + prefer 2 apply (blast intro: well_ord_subset elim!: predE) proof - fix X r xb assume "M(K)" "X \ Pow\<^bsup>M\<^esup>(K)" "r \ Pow\<^bsup>M\<^esup>(K \ K)" "well_ord(X, r)" "xb \ X" moreover from this have "M(X)" "M(r)" using cartprod_closed trans_Pow_rel_closed by auto moreover from this have "M(xb)" using transM[OF \xb\X\] by simp ultimately show "Order.pred(X, xb, r) \ Pow\<^bsup>M\<^esup>(K)" using def_Pow_rel by (auto dest:predE) qed declare conj_cong [cong del] \ \incompatible with some of the proofs of the original theory\ lemma jump_cardinal_rel_iff_old: "M(i) \ M(K) \ i \ jump_cardinal_rel(M,K) \ (\r[M]. \X[M]. r \ K*K & X \ K & well_ord(X,r) & i = ordertype(X,r))" apply (unfold def_jump_cardinal_rel jump_cardinal_body_def) apply (auto del: subsetI) - apply (rename_tac y r) - apply (rule_tac x=r in rexI, intro conjI) prefer 2 - apply (rule_tac x=y in rexI, intro conjI) - apply (auto dest:mem_Pow_rel transM) - apply (rule_tac A=r in rev_subsetD, assumption) - defer - apply (rename_tac r y) - apply (rule_tac x=y in bexI) - apply (rule_tac x=r in ReplaceI, auto) + apply (rename_tac y r) + apply (rule_tac x=r in rexI, intro conjI) prefer 2 + apply (rule_tac x=y in rexI, intro conjI) + apply (auto dest:mem_Pow_rel transM) + apply (rule_tac A=r in rev_subsetD, assumption) + defer + apply (rename_tac r y) + apply (rule_tac x=y in bexI) + apply (rule_tac x=r in ReplaceI, auto) using def_Pow_rel - apply (force+)[2] + apply (force+)[2] apply (rule_tac A=r in rev_subsetD, assumption) using mem_Pow_rel[THEN conjunct1] apply auto done (*The easy part of Theorem 10.16: jump_cardinal_rel(K) exceeds K*) lemma K_lt_jump_cardinal_rel: "Ord(K) ==> M(K) \ K < jump_cardinal_rel(M,K)" apply (rule Ord_jump_cardinal_rel [THEN [2] ltI]) - apply (rule jump_cardinal_rel_iff_old [THEN iffD2], assumption+) - apply (rule_tac x="Memrel(K)" in rexI) - apply (rule_tac x=K in rexI) - apply (simp add: ordertype_Memrel well_ord_Memrel) + apply (rule jump_cardinal_rel_iff_old [THEN iffD2], assumption+) + apply (rule_tac x="Memrel(K)" in rexI) + apply (rule_tac x=K in rexI) + apply (simp add: ordertype_Memrel well_ord_Memrel) using Memrel_closed - apply (simp_all add: Memrel_def subset_iff) + apply (simp_all add: Memrel_def subset_iff) done lemma def_jump_cardinal_rel': assumes "M(K)" shows "jump_cardinal_rel(M,K) = - (\X\Pow\<^bsup>M\<^esup>(K). {z. r \ Pow\<^bsup>M\<^esup>(X*X), well_ord(X,r) & z = ordertype(X,r)})" + (\X\Pow\<^bsup>M\<^esup>(K). {z. r \ Pow\<^bsup>M\<^esup>(X\X), well_ord(X,r) \ z = ordertype(X,r)})" proof - from assms have "jump_cardinal_rel(M, K) = (\X\Pow\<^bsup>M\<^esup>(K). jump_cardinal_body(Pow\<^bsup>M\<^esup>(K \ K), X))" using def_jump_cardinal_rel by simp also from \M(K)\ have " ... = (\X\Pow\<^bsup>M\<^esup>(K). jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X), X))" using def_jump_cardinal_rel_aux2 by simp finally show ?thesis unfolding jump_cardinal_body_def by simp qed rel_closed for "jump_cardinal" proof - fix K assume "M(K)" moreover from this have "jump_cardinal_rel(M, K) = \{jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X), X) . X\Pow\<^bsup>M\<^esup>(K)}" using def_jump_cardinal_rel' unfolding jump_cardinal_body_def[symmetric] by simp moreover from calculation have "M((\{jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X), X) . X\Pow\<^bsup>M\<^esup>(K)}))" - using def_jump_cardinal_rel_aux jump_cardinal_body_closed + using def_jump_cardinal_rel_aux jump_cardinal_closed by simp ultimately show "M(jump_cardinal_rel(M,K))" by simp qed (*The proof by contradiction: the bijection f yields a wellordering of X whose ordertype is jump_cardinal_rel(K). *) lemma Card_rel_jump_cardinal_rel_lemma: "[| well_ord(X,r); r \ K * K; X \ K; f \ bij(ordertype(X,r), jump_cardinal_rel(M,K)); M(X); M(r); M(K); M(f) |] ==> jump_cardinal_rel(M,K) \ jump_cardinal_rel(M,K)" apply (subgoal_tac "f O ordermap (X,r) \ bij (X, jump_cardinal_rel (M,K))") - prefer 2 apply (blast intro: comp_bij ordermap_bij) + prefer 2 apply (blast intro: comp_bij ordermap_bij) apply (rule jump_cardinal_rel_iff_old [THEN iffD2],simp+) apply (intro rexI conjI) - apply (rule subset_trans [OF rvimage_type Sigma_mono], assumption+) - apply (erule bij_is_inj [THEN well_ord_rvimage]) - apply (rule Ord_jump_cardinal_rel [THEN well_ord_Memrel]) - apply (simp_all add: well_ord_Memrel [THEN [2] bij_ordertype_vimage] + apply (rule subset_trans [OF rvimage_type Sigma_mono], assumption+) + apply (erule bij_is_inj [THEN well_ord_rvimage]) + apply (rule Ord_jump_cardinal_rel [THEN well_ord_Memrel]) + apply (simp_all add: well_ord_Memrel [THEN [2] bij_ordertype_vimage] ordertype_Memrel Ord_jump_cardinal_rel) done (*The hard part of Theorem 10.16: jump_cardinal_rel(K) is itself a cardinal*) lemma Card_rel_jump_cardinal_rel: "M(K) \ Card_rel(M,jump_cardinal_rel(M,K))" apply (rule Ord_jump_cardinal_rel [THEN Card_relI]) - apply (simp_all add: def_eqpoll_rel) + apply (simp_all add: def_eqpoll_rel) apply (drule_tac i1=j in jump_cardinal_rel_iff_old [THEN iffD1, OF _ _ ltD, of _ K], safe) apply (blast intro: Card_rel_jump_cardinal_rel_lemma [THEN mem_irrefl]) done subsection\Basic Properties of Successor Cardinals\ lemma csucc_rel_basic: "Ord(K) ==> M(K) \ Card_rel(M,csucc_rel(M,K)) & K < csucc_rel(M,K)" apply (unfold csucc_rel_def) apply (rule LeastI[of "\i. M(i) \ Card_rel(M,i) \ K < i", THEN conjunct2]) - apply (blast intro: Card_rel_jump_cardinal_rel K_lt_jump_cardinal_rel Ord_jump_cardinal_rel)+ + apply (blast intro: Card_rel_jump_cardinal_rel K_lt_jump_cardinal_rel Ord_jump_cardinal_rel)+ done lemmas Card_rel_csucc_rel = csucc_rel_basic [THEN conjunct1] lemmas lt_csucc_rel = csucc_rel_basic [THEN conjunct2] lemma Ord_0_lt_csucc_rel: "Ord(K) ==> M(K) \ 0 < csucc_rel(M,K)" by (blast intro: Ord_0_le lt_csucc_rel lt_trans1) lemma csucc_rel_le: "[| Card_rel(M,L); K csucc_rel(M,K) \ L" apply (unfold csucc_rel_def) apply (rule Least_le) - apply (blast intro: Card_rel_is_Ord)+ + apply (blast intro: Card_rel_is_Ord)+ done lemma lt_csucc_rel_iff: "[| Ord(i); Card_rel(M,K); M(K); M(i)|] ==> i < csucc_rel(M,K) \ |i|\<^bsup>M\<^esup> \ K" apply (rule iffI) - apply (rule_tac [2] Card_rel_lt_imp_lt) - apply (erule_tac [2] lt_trans1) - apply (simp_all add: lt_csucc_rel Card_rel_csucc_rel Card_rel_is_Ord) + apply (rule_tac [2] Card_rel_lt_imp_lt) + apply (erule_tac [2] lt_trans1) + apply (simp_all add: lt_csucc_rel Card_rel_csucc_rel Card_rel_is_Ord) apply (rule notI [THEN not_lt_imp_le]) - apply (rule Card_rel_cardinal_rel [THEN csucc_rel_le, THEN lt_trans1, THEN lt_irrefl], simp_all+) - apply (rule Ord_cardinal_rel_le [THEN lt_trans1]) - apply (simp_all add: Card_rel_is_Ord) + apply (rule Card_rel_cardinal_rel [THEN csucc_rel_le, THEN lt_trans1, THEN lt_irrefl], simp_all+) + apply (rule Ord_cardinal_rel_le [THEN lt_trans1]) + apply (simp_all add: Card_rel_is_Ord) done lemma Card_rel_lt_csucc_rel_iff: "[| Card_rel(M,K'); Card_rel(M,K); M(K'); M(K) |] ==> K' < csucc_rel(M,K) \ K' \ K" by (simp add: lt_csucc_rel_iff Card_rel_cardinal_rel_eq Card_rel_is_Ord) lemma InfCard_rel_csucc_rel: "InfCard_rel(M,K) \ M(K) ==> InfCard_rel(M,csucc_rel(M,K))" by (simp add: InfCard_rel_def Card_rel_csucc_rel Card_rel_is_Ord lt_csucc_rel [THEN leI, THEN [2] le_trans]) subsubsection\Theorems by Krzysztof Grabczewski, proofs by lcp\ lemma nat_sum_eqpoll_rel_sum: assumes m: "m \ nat" and n: "n \ nat" shows "m + n \\<^bsup>M\<^esup> m +\<^sub>\ n" proof - have "m + n \\<^bsup>M\<^esup> |m+n|\<^bsup>M\<^esup>" using m n by (blast intro: nat_implies_well_ord well_ord_radd well_ord_cardinal_rel_eqpoll_rel eqpoll_rel_sym) also have "... = m +\<^sub>\ n" using m n by (simp add: nat_cadd_rel_eq_add [symmetric] cadd_rel_def transM[OF _ M_nat]) finally show ?thesis . qed lemma Ord_nat_subset_into_Card_rel: "[| Ord(i); i \ nat |] ==> Card\<^bsup>M\<^esup>(i)" by (blast dest: Ord_subset_natD intro: Card_rel_nat nat_into_Card_rel) end \ \\<^locale>\M_cardinal_arith_jump\\ end diff --git a/thys/Transitive_Models/Cardinal_AC_Relative.thy b/thys/Transitive_Models/Cardinal_AC_Relative.thy --- a/thys/Transitive_Models/Cardinal_AC_Relative.thy +++ b/thys/Transitive_Models/Cardinal_AC_Relative.thy @@ -1,441 +1,436 @@ section\Relative, Cardinal Arithmetic Using AC\ theory Cardinal_AC_Relative imports CardinalArith_Relative begin locale M_AC = fixes M assumes choice_ax: "choice_ax(M)" locale M_cardinal_AC = M_cardinal_arith + M_AC + assumes lam_replacement_minimum:"lam_replacement(M, \p. minimum(fst(p),snd(p)))" begin lemma lam_replacement_minimum_vimage: "M(f) \ M(r) \ lam_replacement(M, \x. minimum(r, f -`` {x}))" using lam_replacement_minimum lam_replacement_vimage_sing_fun lam_replacement_constant - by (rule_tac lam_replacement_hcomp2[of _ _ minimum]) - (force intro: lam_replacement_identity)+ + lam_replacement_identity lam_replacement_hcomp2[of _ _ minimum] + by simp lemmas surj_imp_inj_replacement4 = lam_replacement_minimum_vimage[unfolded lam_replacement_def] -lemma lam_replacement_min: "M(f) \ M(r) \ lam_replacement(M, \x . minimum(r, f -`` {x}))" - using lam_replacement_hcomp2[OF lam_replacement_constant[of r] lam_replacement_vimage_sing_fun] - lam_replacement_minimum - by simp - lemmas surj_imp_inj_replacement = surj_imp_inj_replacement1 surj_imp_inj_replacement2 surj_imp_inj_replacement4 lam_replacement_vimage_sing_fun[THEN lam_replacement_imp_strong_replacement] lemma well_ord_surj_imp_lepoll_rel: assumes "well_ord(A,r)" "h \ surj(A,B)" and types:"M(A)" "M(r)" "M(h)" "M(B)" shows "B \\<^bsup>M\<^esup> A" proof - note eq=vimage_fun_sing[OF surj_is_fun[OF \h\_\]] from assms have "(\b\B. minimum(r, {a\A. h`a=b})) \ inj(B,A)" (is "?f\_") using well_ord_surj_imp_inj_inverse assms(1,2) by simp with assms have "M(?f`b)" if "b\B" for b using apply_type[OF inj_is_fun[OF \?f\_\]] that transM[OF _ \M(A)\] by simp with assms have "M(?f)" using lam_closed surj_imp_inj_replacement4 eq by auto with \?f\_\ assms have "?f \ inj\<^bsup>M\<^esup>(B,A)" using mem_inj_abs by simp with \M(?f)\ show ?thesis unfolding lepoll_rel_def by auto qed lemma surj_imp_well_ord_M: assumes wos: "well_ord(A,r)" "h \ surj(A,B)" and types: "M(A)" "M(r)" "M(h)" "M(B)" shows "\s[M]. well_ord(B,s)" using assms lepoll_rel_well_ord well_ord_surj_imp_lepoll_rel by fast lemma choice_ax_well_ord: "M(S) \ \r[M]. well_ord(S,r)" using choice_ax well_ord_Memrel[THEN surj_imp_well_ord_M] unfolding choice_ax_def by auto lemma Finite_cardinal_rel_Finite: assumes "Finite(|i|\<^bsup>M\<^esup>)" "M(i)" shows "Finite(i)" proof - note assms moreover from this obtain r where "M(r)" "well_ord(i,r)" using choice_ax_well_ord by auto moreover from calculation have "|i|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> i" using well_ord_cardinal_rel_eqpoll_rel by auto ultimately show ?thesis using eqpoll_rel_imp_Finite by auto qed end \ \\<^locale>\M_cardinal_AC\\ locale M_Pi_assumptions_choice = M_Pi_assumptions + M_cardinal_AC + assumes B_replacement: "strong_replacement(M, \x y. y = B(x))" and \ \The next one should be derivable from (some variant) of B\_replacement. Proving both instances each time seems inconvenient.\ minimum_replacement: "M(r) \ strong_replacement(M, \x y. y = \x, minimum(r, B(x))\)" begin lemma AC_M: assumes "a \ A" "\x. x \ A \ \y. y \ B(x)" shows "\z[M]. z \ Pi\<^bsup>M\<^esup>(A, B)" proof - have "M(\x\A. B(x))" using assms family_union_closed Pi_assumptions B_replacement by simp then obtain r where "well_ord(\x\A. B(x),r)" "M(r)" using choice_ax_well_ord by blast let ?f="\x\A. minimum(r,B(x))" have "M(minimum(r, B(x)))" if "x\A" for x proof - from \well_ord(_,r)\ \x\A\ have "well_ord(B(x),r)" using well_ord_subset UN_upper by simp with assms \x\A\ \M(r)\ show ?thesis using Pi_assumptions by blast qed with assms and \M(r)\ have "M(?f)" using Pi_assumptions minimum_replacement lam_closed by simp moreover from assms and calculation have "?f \ Pi\<^bsup>M\<^esup>(A,B)" using lam_type[OF minimum_in, OF \well_ord(\x\A. B(x),r)\, of A B] Pi_rel_char by auto ultimately show ?thesis by blast qed lemma AC_Pi_rel: assumes "\x. x \ A \ \y. y \ B(x)" shows "\z[M]. z \ Pi\<^bsup>M\<^esup>(A, B)" proof (cases "A=0") interpret Pi0:M_Pi_assumptions_0 using Pi_assumptions by unfold_locales auto case True then show ?thesis using assms by simp next case False then obtain a where "a \ A" by auto \ \It is noteworthy that without obtaining an element of \<^term>\A\, the final step won't work\ with assms show ?thesis by (blast intro!: AC_M) qed end \ \\<^locale>\M_Pi_assumptions_choice\\ context M_cardinal_AC begin subsection\Strengthened Forms of Existing Theorems on Cardinals\ lemma cardinal_rel_eqpoll_rel: "M(A) \ |A|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> A" apply (rule choice_ax_well_ord [THEN rexE]) apply (auto intro:well_ord_cardinal_rel_eqpoll_rel) done lemmas cardinal_rel_idem = cardinal_rel_eqpoll_rel [THEN cardinal_rel_cong, simp] lemma cardinal_rel_eqE: "|X|\<^bsup>M\<^esup> = |Y|\<^bsup>M\<^esup> ==> M(X) \ M(Y) \ X \\<^bsup>M\<^esup> Y" apply (rule choice_ax_well_ord [THEN rexE], assumption) apply (rule choice_ax_well_ord [THEN rexE, of Y], assumption) apply (rule well_ord_cardinal_rel_eqE, assumption+) done lemma cardinal_rel_eqpoll_rel_iff: "M(X) \ M(Y) \ |X|\<^bsup>M\<^esup> = |Y|\<^bsup>M\<^esup> \ X \\<^bsup>M\<^esup> Y" by (blast intro: cardinal_rel_cong cardinal_rel_eqE) lemma cardinal_rel_disjoint_Un: "[| |A|\<^bsup>M\<^esup>=|B|\<^bsup>M\<^esup>; |C|\<^bsup>M\<^esup>=|D|\<^bsup>M\<^esup>; A \ C = 0; B \ D = 0; M(A); M(B); M(C); M(D)|] ==> |A \ C|\<^bsup>M\<^esup> = |B \ D|\<^bsup>M\<^esup>" by (simp add: cardinal_rel_eqpoll_rel_iff eqpoll_rel_disjoint_Un) lemma lepoll_rel_imp_cardinal_rel_le: "A \\<^bsup>M\<^esup> B ==> M(A) \ M(B) \ |A|\<^bsup>M\<^esup> \ |B|\<^bsup>M\<^esup>" apply (rule choice_ax_well_ord [THEN rexE]) prefer 2 apply (erule well_ord_lepoll_rel_imp_cardinal_rel_le, assumption+) done lemma cadd_rel_assoc: "\M(i); M(j); M(k)\ \ (i \\<^bsup>M\<^esup> j) \\<^bsup>M\<^esup> k = i \\<^bsup>M\<^esup> (j \\<^bsup>M\<^esup> k)" apply (rule choice_ax_well_ord [THEN rexE]) prefer 2 apply (rule choice_ax_well_ord [THEN rexE]) prefer 2 apply (rule choice_ax_well_ord [THEN rexE]) prefer 2 apply (rule well_ord_cadd_rel_assoc, assumption+) done lemma cmult_rel_assoc: "\M(i); M(j); M(k)\ \ (i \\<^bsup>M\<^esup> j) \\<^bsup>M\<^esup> k = i \\<^bsup>M\<^esup> (j \\<^bsup>M\<^esup> k)" apply (rule choice_ax_well_ord [THEN rexE]) prefer 2 apply (rule choice_ax_well_ord [THEN rexE]) prefer 2 apply (rule choice_ax_well_ord [THEN rexE]) prefer 2 apply (rule well_ord_cmult_rel_assoc, assumption+) done lemma cadd_cmult_distrib: "\M(i); M(j); M(k)\ \ (i \\<^bsup>M\<^esup> j) \\<^bsup>M\<^esup> k = (i \\<^bsup>M\<^esup> k) \\<^bsup>M\<^esup> (j \\<^bsup>M\<^esup> k)" apply (rule choice_ax_well_ord [THEN rexE]) prefer 2 apply (rule choice_ax_well_ord [THEN rexE]) prefer 2 apply (rule choice_ax_well_ord [THEN rexE]) prefer 2 apply (rule well_ord_cadd_cmult_distrib, assumption+) done lemma InfCard_rel_square_eq: "InfCard\<^bsup>M\<^esup>(|A|\<^bsup>M\<^esup>) \ M(A) \ A\A \\<^bsup>M\<^esup> A" apply (rule choice_ax_well_ord [THEN rexE]) prefer 2 apply (erule well_ord_InfCard_rel_square_eq, assumption, simp_all) done subsection \The relationship between cardinality and le-pollence\ lemma Card_rel_le_imp_lepoll_rel: assumes "|A|\<^bsup>M\<^esup> \ |B|\<^bsup>M\<^esup>" and types: "M(A)" "M(B)" shows "A \\<^bsup>M\<^esup> B" proof - have "A \\<^bsup>M\<^esup> |A|\<^bsup>M\<^esup>" by (rule cardinal_rel_eqpoll_rel [THEN eqpoll_rel_sym], simp_all add:types) also have "... \\<^bsup>M\<^esup> |B|\<^bsup>M\<^esup>" by (rule le_imp_subset [THEN subset_imp_lepoll_rel]) (rule assms, simp_all add:types) also have "... \\<^bsup>M\<^esup> B" by (rule cardinal_rel_eqpoll_rel, simp_all add:types) finally show ?thesis by (simp_all add:types) qed lemma le_Card_rel_iff: "Card\<^bsup>M\<^esup>(K) ==> M(K) \ M(A) \ |A|\<^bsup>M\<^esup> \ K \ A \\<^bsup>M\<^esup> K" apply (erule Card_rel_cardinal_rel_eq [THEN subst], assumption, rule iffI, erule Card_rel_le_imp_lepoll_rel, assumption+) apply (erule lepoll_rel_imp_cardinal_rel_le, assumption+) done lemma cardinal_rel_0_iff_0 [simp]: "M(A) \ |A|\<^bsup>M\<^esup> = 0 \ A = 0" using cardinal_rel_0 eqpoll_rel_0_iff [THEN iffD1] cardinal_rel_eqpoll_rel_iff [THEN iffD1, OF _ nonempty] by auto lemma cardinal_rel_lt_iff_lesspoll_rel: assumes i: "Ord(i)" and types: "M(i)" "M(A)" shows "i < |A|\<^bsup>M\<^esup> \ i \\<^bsup>M\<^esup> A" proof assume "i < |A|\<^bsup>M\<^esup>" hence "i \\<^bsup>M\<^esup> |A|\<^bsup>M\<^esup>" by (blast intro: lt_Card_rel_imp_lesspoll_rel types) also have "... \\<^bsup>M\<^esup> A" by (rule cardinal_rel_eqpoll_rel) (simp_all add:types) finally show "i \\<^bsup>M\<^esup> A" by (simp_all add:types) next assume "i \\<^bsup>M\<^esup> A" also have "... \\<^bsup>M\<^esup> |A|\<^bsup>M\<^esup>" by (blast intro: cardinal_rel_eqpoll_rel eqpoll_rel_sym types) finally have "i \\<^bsup>M\<^esup> |A|\<^bsup>M\<^esup>" by (simp_all add:types) thus "i < |A|\<^bsup>M\<^esup>" using i types by (force intro: cardinal_rel_lt_imp_lt lesspoll_rel_cardinal_rel_lt) qed lemma cardinal_rel_le_imp_lepoll_rel: " i \ |A|\<^bsup>M\<^esup> ==> M(i) \ M(A) \i \\<^bsup>M\<^esup> A" by (blast intro: lt_Ord Card_rel_le_imp_lepoll_rel Ord_cardinal_rel_le le_trans) subsection\Other Applications of AC\ text\We have an example of instantiating a locale involving higher order variables inside a proof, by using the assumptions of the first order, active locale.\ lemma surj_rel_implies_inj_rel: assumes f: "f \ surj\<^bsup>M\<^esup>(X,Y)" and types: "M(f)" "M(X)" "M(Y)" shows "\g[M]. g \ inj\<^bsup>M\<^esup>(Y,X)" proof - from types interpret M_Pi_assumptions_choice _ Y "\y. f-``{y}" by unfold_locales (auto intro:surj_imp_inj_replacement dest:transM) from f AC_Pi_rel obtain z where z: "z \ Pi\<^bsup>M\<^esup>(Y, \y. f -`` {y})" \ \In this and the following ported result, it is not clear how uniformly are "\_char" theorems to be used\ using surj_rel_char by (auto simp add: surj_def types) (fast dest: apply_Pair) show ?thesis proof show "z \ inj\<^bsup>M\<^esup>(Y, X)" "M(z)" using z surj_is_fun[of f X Y] f Pi_rel_char by (auto dest: apply_type Pi_memberD intro: apply_equality Pi_type f_imp_injective simp add:types mem_surj_abs) qed qed text\Kunen's Lemma 10.20\ lemma surj_rel_implies_cardinal_rel_le: assumes f: "f \ surj\<^bsup>M\<^esup>(X,Y)" and types:"M(f)" "M(X)" "M(Y)" shows "|Y|\<^bsup>M\<^esup> \ |X|\<^bsup>M\<^esup>" proof (rule lepoll_rel_imp_cardinal_rel_le) from f [THEN surj_rel_implies_inj_rel] obtain g where "g \ inj\<^bsup>M\<^esup>(Y,X)" by (blast intro:types) then show "Y \\<^bsup>M\<^esup> X" using inj_rel_char by (auto simp add: def_lepoll_rel types) qed (simp_all add:types) end \ \\<^locale>\M_cardinal_AC\\ text\The set-theoretic universe.\ abbreviation Universe :: "i\o" (\\\) where "\(x) \ True" lemma separation_absolute: "separation(\, P)" unfolding separation_def by (rule rallI, rule_tac x="{x\_ . P(x)}" in rexI) auto lemma univalent_absolute: assumes "univalent(\, A, P)" "P(x, b)" "x \ A" shows "P(x, y) \ y = b" using assms unfolding univalent_def by force lemma replacement_absolute: "strong_replacement(\, P)" unfolding strong_replacement_def proof (intro rallI impI) fix A assume "univalent(\, A, P)" then show "\Y[\]. \b[\]. b \ Y \ (\x[\]. x \ A \ P(x, b))" by (rule_tac x="{y. x\A , P(x,y)}" in rexI) (auto dest:univalent_absolute[of _ P]) qed lemma Union_ax_absolute: "Union_ax(\)" unfolding Union_ax_def big_union_def by (auto intro:rexI[of _ "\_"]) lemma upair_ax_absolute: "upair_ax(\)" unfolding upair_ax_def upair_def rall_def rex_def by (auto) lemma power_ax_absolute:"power_ax(\)" proof - { fix x have "\y[\]. y \ Pow(x) \ (\z[\]. z \ y \ z \ x)" by auto } then show "power_ax(\)" unfolding power_ax_def powerset_def subset_def by blast qed locale M_cardinal_UN = M_Pi_assumptions_choice _ K X for K X + assumes \ \The next assumption is required by @{thm Least_closed}\ X_witness_in_M: "w \ X(x) \ M(x)" and lam_m_replacement:"M(f) \ strong_replacement(M, \x y. y = \x, \ i. x \ X(i), f ` (\ i. x \ X(i)) ` x\)" and inj_replacement: "M(x) \ strong_replacement(M, \y z. y \ inj\<^bsup>M\<^esup>(X(x), K) \ z = {\x, y\})" "strong_replacement(M, \x y. y = inj\<^bsup>M\<^esup>(X(x), K))" "strong_replacement(M, \x z. z = Sigfun(x, \i. inj\<^bsup>M\<^esup>(X(i), K)))" "M(r) \ strong_replacement(M, \x y. y = \x, minimum(r, inj\<^bsup>M\<^esup>(X(x), K))\)" begin lemma UN_closed: "M(\i\K. X(i))" using family_union_closed B_replacement Pi_assumptions by simp text\Kunen's Lemma 10.21\ lemma cardinal_rel_UN_le: assumes K: "InfCard\<^bsup>M\<^esup>(K)" shows "(\i. i\K \ |X(i)|\<^bsup>M\<^esup> \ K) \ |\i\K. X(i)|\<^bsup>M\<^esup> \ K" proof (simp add: K InfCard_rel_is_Card_rel le_Card_rel_iff Pi_assumptions) have "M(f) \ M(\x\(\x\K. X(x)). \\ i. x \ X(i), f ` (\ i. x \ X(i)) ` x\)" for f using lam_m_replacement X_witness_in_M Least_closed' Pi_assumptions UN_closed by (rule_tac lam_closed) (auto dest:transM) note types = this Pi_assumptions UN_closed have [intro]: "Ord(K)" by (blast intro: InfCard_rel_is_Card_rel Card_rel_is_Ord K types) interpret pii:M_Pi_assumptions_choice _ K "\i. inj\<^bsup>M\<^esup>(X(i), K)" using inj_replacement Pi_assumptions transM[of _ K] by unfold_locales (simp_all del:mem_inj_abs) assume asm:"\i. i\K \ X(i) \\<^bsup>M\<^esup> K" then have "\i. i\K \ M(inj\<^bsup>M\<^esup>(X(i), K))" by (auto simp add: types) interpret V:M_N_Perm M "\" using separation_absolute replacement_absolute Union_ax_absolute power_ax_absolute upair_ax_absolute by unfold_locales auto note bad_simps[simp del] = V.N.Forall_in_M_iff V.N.Equal_in_M_iff V.N.nonempty have abs:"inj_rel(\,x,y) = inj(x,y)" for x y using V.N.inj_rel_char by simp from asm have "\i. i\K \ \f[M]. f \ inj\<^bsup>M\<^esup>(X(i), K)" by (simp add: types def_lepoll_rel) then obtain f where "f \ (\i\K. inj\<^bsup>M\<^esup>(X(i), K))" "M(f)" using pii.AC_Pi_rel pii.Pi_rel_char by auto with abs have f:"f \ (\i\K. inj(X(i), K))" using Pi_weaken_type[OF _ V.inj_rel_transfer, of f K X "\_. K"] Pi_assumptions by simp { fix z assume z: "z \ (\i\K. X(i))" then obtain i where i: "i \ K" "Ord(i)" "z \ X(i)" by (blast intro: Ord_in_Ord [of K]) hence "(\ i. z \ X(i)) \ i" by (fast intro: Least_le) hence "(\ i. z \ X(i)) < K" by (best intro: lt_trans1 ltI i) hence "(\ i. z \ X(i)) \ K" and "z \ X(\ i. z \ X(i))" by (auto intro: LeastI ltD i) } note mems = this have "(\i\K. X(i)) \\<^bsup>M\<^esup> K \ K" proof (simp add:types def_lepoll_rel) show "\f[M]. f \ inj(\x\K. X(x), K \ K)" apply (rule rexI) apply (rule_tac c = "\z. \\ i. z \ X(i), f ` (\ i. z \ X(i)) ` z\" and d = "\\i,j\. converse (f`i) ` j" in lam_injective) apply (force intro: f inj_is_fun mems apply_type Perm.left_inverse)+ apply (simp add:types \M(f)\) done qed also have "... \\<^bsup>M\<^esup> K" by (simp add: K InfCard_rel_square_eq InfCard_rel_is_Card_rel Card_rel_cardinal_rel_eq types) finally have "(\i\K. X(i)) \\<^bsup>M\<^esup> K" by (simp_all add:types) then show ?thesis by (simp add: K InfCard_rel_is_Card_rel le_Card_rel_iff types) qed end \ \\<^locale>\M_cardinal_UN\\ end \ No newline at end of file diff --git a/thys/Transitive_Models/Cardinal_Library_Relative.thy b/thys/Transitive_Models/Cardinal_Library_Relative.thy --- a/thys/Transitive_Models/Cardinal_Library_Relative.thy +++ b/thys/Transitive_Models/Cardinal_Library_Relative.thy @@ -1,1258 +1,1241 @@ section\Cardinal Arithmetic under Choice\label{sec:cardinal-lib-rel}\ theory Cardinal_Library_Relative imports Replacement_Lepoll begin locale M_library = M_ZF_library + M_cardinal_AC + assumes separation_cardinal_rel_lesspoll_rel: "M(\) \ separation(M, \x . x \\<^bsup>M\<^esup> \)" begin declare eqpoll_rel_refl [simp] subsection\Miscellaneous\ lemma cardinal_rel_RepFun_apply_le: assumes "S \ A\B" "M(S)" "M(A)" "M(B)" shows "|{S`a . a\A}|\<^bsup>M\<^esup> \ |A|\<^bsup>M\<^esup>" proof - note assms moreover from this have "{S ` a . a \ A} = S``A" using image_eq_UN RepFun_def UN_iff by force moreover from calculation have "M(\x\A. S ` x)" "M({S ` a . a \ A})" - using lam_closed[of "\ x. S`x"] apply_type[OF \S\_\] - transM[OF _ \M(B)\] image_closed + using lam_closed[of "\ x. S`x"] apply_type[OF \S\_\] transM[OF _ \M(B)\] by auto moreover from assms this have "(\x\A. S`x) \ surj_rel(M,A, {S`a . a\A})" using mem_surj_abs lam_funtype[of A "\x . S`x"] unfolding surj_def by auto ultimately show ?thesis using surj_rel_char surj_rel_implies_cardinal_rel_le by simp qed (* TODO: Check if we can use this lemma to prove the previous one and not the other way around *) lemma cardinal_rel_RepFun_le: assumes lrf:"lam_replacement(M,f)" and f_closed:"\x[M]. M(f(x))" and "M(X)" shows "|{f(x) . x \ X}|\<^bsup>M\<^esup> \ |X|\<^bsup>M\<^esup>" using \M(X)\ f_closed cardinal_rel_RepFun_apply_le[OF lam_funtype, of X _, OF lrf[THEN [2] lam_replacement_iff_lam_closed[THEN iffD1, THEN rspec]]] lrf[THEN lam_replacement_imp_strong_replacement] by simp (auto simp flip:setclass_iff intro!:RepFun_closed dest:transM) lemma subset_imp_le_cardinal_rel: "A \ B \ M(A) \ M(B) \ |A|\<^bsup>M\<^esup> \ |B|\<^bsup>M\<^esup>" using subset_imp_lepoll_rel[THEN lepoll_rel_imp_cardinal_rel_le] . lemma lt_cardinal_rel_imp_not_subset: "|A|\<^bsup>M\<^esup> < |B|\<^bsup>M\<^esup> \ M(A) \ M(B) \ \ B \ A" using subset_imp_le_cardinal_rel le_imp_not_lt by blast lemma cardinal_rel_lt_csucc_rel_iff: "Card_rel(M,K) \ M(K) \ M(K') \ |K'|\<^bsup>M\<^esup> < (K\<^sup>+)\<^bsup>M\<^esup> \ |K'|\<^bsup>M\<^esup> \ K" by (simp add: Card_rel_lt_csucc_rel_iff) end \ \\<^locale>\M_library\\ locale M_cardinal_UN_nat = M_cardinal_UN _ \ X for X begin lemma cardinal_rel_UN_le_nat: assumes "\i. i\\ \ |X(i)|\<^bsup>M\<^esup> \ \" shows "|\i\\. X(i)|\<^bsup>M\<^esup> \ \" proof - from assms show ?thesis by (simp add: cardinal_rel_UN_le InfCard_rel_nat) qed end \ \\<^locale>\M_cardinal_UN_nat\\ locale M_cardinal_UN_inj = M_library + j:M_cardinal_UN _ J + y:M_cardinal_UN _ K "\k. if k\range(f) then X(converse(f)`k) else 0" for J K f + assumes f_inj: "f \ inj_rel(M,J,K)" begin lemma inj_rel_imp_cardinal_rel_UN_le: notes [dest] = InfCard_is_Card Card_is_Ord fixes Y defines "Y(k) \ if k\range(f) then X(converse(f)`k) else 0" assumes "InfCard\<^bsup>M\<^esup>(K)" "\i. i\J \ |X(i)|\<^bsup>M\<^esup> \ K" shows "|\i\J. X(i)|\<^bsup>M\<^esup> \ K" proof - have "M(K)" "M(J)" "\w x. w \ X(x) \ M(x)" using y.Pi_assumptions j.Pi_assumptions j.X_witness_in_M by simp_all then have "M(f)" using inj_rel_char f_inj by simp note inM = \M(f)\ \M(K)\ \M(J)\ \\w x. w \ X(x) \ M(x)\ have "i\J \ f`i \ K" for i using inj_rel_is_fun[OF f_inj] apply_type function_space_rel_char by (auto simp add:inM) have "(\i\J. X(i)) \ (\i\K. Y(i))" proof (standard, elim UN_E) fix x i assume "i\J" "x\X(i)" with \i\J \ f`i \ K\ have "x \ Y(f`i)" "f`i \ K" unfolding Y_def using inj_is_fun right_inverse f_inj by (auto simp add:inM Y_def intro: apply_rangeI) then show "x \ (\i\K. Y(i))" by auto qed then have "|\i\J. X(i)|\<^bsup>M\<^esup> \ |\i\K. Y(i)|\<^bsup>M\<^esup>" using subset_imp_le_cardinal_rel j.UN_closed y.UN_closed unfolding Y_def by (simp add:inM) moreover note assms \\i. i\J \ f`i \ K\ inM moreover from this have "k\range(f) \ converse(f)`k \ J" for k using inj_rel_converse_fun[OF f_inj] range_fun_subset_codomain function_space_rel_char by simp ultimately show "|\i\J. X(i)|\<^bsup>M\<^esup> \ K" using InfCard_rel_is_Card_rel[THEN Card_rel_is_Ord,THEN Ord_0_le, of K] by (rule_tac le_trans[OF _ y.cardinal_rel_UN_le]) (auto intro:Ord_0_le simp:Y_def)+ qed end \ \\<^locale>\M_cardinal_UN_inj\\ locale M_cardinal_UN_lepoll = M_library + M_replacement_lepoll _ "\_. X" + j:M_cardinal_UN _ J for J begin -(* FIXME: this "LEQpoll" should be "LEPOLL"; same correction in Delta System *) -lemma leqpoll_rel_imp_cardinal_rel_UN_le: +lemma lepoll_rel_imp_cardinal_rel_UN_le: notes [dest] = InfCard_is_Card Card_is_Ord assumes "InfCard\<^bsup>M\<^esup>(K)" "J \\<^bsup>M\<^esup> K" "\i. i\J \ |X(i)|\<^bsup>M\<^esup> \ K" "M(K)" shows "|\i\J. X(i)|\<^bsup>M\<^esup> \ K" proof - from \J \\<^bsup>M\<^esup> K\ obtain f where "f \ inj_rel(M,J,K)" "M(f)" by blast moreover let ?Y="\k. if k\range(f) then X(converse(f)`k) else 0" note \M(K)\ moreover from calculation have "k \ range(f) \ converse(f)`k \ J" for k using mem_inj_rel[THEN inj_converse_fun, THEN apply_type] j.Pi_assumptions by blast moreover from \M(f)\ have "w \ ?Y(x) \ M(x)" for w x by (cases "x\range(f)") (auto dest:transM) moreover from calculation interpret M_Pi_assumptions_choice _ K ?Y using j.Pi_assumptions lepoll_assumptions proof (unfold_locales, auto dest:transM) show "strong_replacement(M, \y z. False)" unfolding strong_replacement_def by auto qed from calculation interpret M_cardinal_UN_inj _ _ _ _ f using lepoll_assumptions by unfold_locales auto from assms show ?thesis using inj_rel_imp_cardinal_rel_UN_le by simp qed end \ \\<^locale>\M_cardinal_UN_lepoll\\ context M_library begin lemma cardinal_rel_lt_csucc_rel_iff': includes Ord_dests assumes "Card_rel(M,\)" and types:"M(\)" "M(X)" shows "\ < |X|\<^bsup>M\<^esup> \ (\\<^sup>+)\<^bsup>M\<^esup> \ |X|\<^bsup>M\<^esup>" using assms cardinal_rel_lt_csucc_rel_iff[of \ X] Card_rel_csucc_rel[of \] not_le_iff_lt[of "(\\<^sup>+)\<^bsup>M\<^esup>" "|X|\<^bsup>M\<^esup>"] not_le_iff_lt[of "|X|\<^bsup>M\<^esup>" \] by blast lemma lepoll_rel_imp_subset_bij_rel: assumes "M(X)" "M(Y)" shows "X \\<^bsup>M\<^esup> Y \ (\Z[M]. Z \ Y \ Z \\<^bsup>M\<^esup> X)" proof assume "X \\<^bsup>M\<^esup> Y" then obtain j where "j \ inj_rel(M,X,Y)" by blast with assms have "range(j) \ Y" "j \ bij_rel(M,X,range(j))" "M(range(j))" "M(j)" using inj_rel_bij_rel_range inj_rel_char inj_rel_is_fun[THEN range_fun_subset_codomain,of j X Y] by auto with assms have "range(j) \ Y" "X \\<^bsup>M\<^esup> range(j)" unfolding eqpoll_rel_def by auto with assms \M(j)\ show "\Z[M]. Z \ Y \ Z \\<^bsup>M\<^esup> X" using eqpoll_rel_sym[OF \X \\<^bsup>M\<^esup> range(j)\] by auto next assume "\Z[M]. Z \ Y \ Z \\<^bsup>M\<^esup> X" then obtain Z f where "f \ bij_rel(M,Z,X)" "Z \ Y" "M(Z)" "M(f)" unfolding eqpoll_rel_def by blast with assms have "converse(f) \ inj_rel(M,X,Y)" "M(converse(f))" using inj_rel_weaken_type[OF bij_rel_converse_bij_rel[THEN bij_rel_is_inj_rel],of f Z X Y] by auto then show "X \\<^bsup>M\<^esup> Y" unfolding lepoll_rel_def by auto qed text\The following result proves to be very useful when combining \<^term>\cardinal_rel\ and \<^term>\eqpoll_rel\ in a calculation.\ lemma cardinal_rel_Card_rel_eqpoll_rel_iff: "Card_rel(M,\) \ M(\) \ M(X) \ |X|\<^bsup>M\<^esup> = \ \ X \\<^bsup>M\<^esup> \" using Card_rel_cardinal_rel_eq[of \] cardinal_rel_eqpoll_rel_iff[of X \] by auto lemma lepoll_rel_imp_lepoll_rel_cardinal_rel: assumes"X \\<^bsup>M\<^esup> Y" "M(X)" "M(Y)" shows "X \\<^bsup>M\<^esup> |Y|\<^bsup>M\<^esup>" using assms cardinal_rel_Card_rel_eqpoll_rel_iff[of "|Y|\<^bsup>M\<^esup>" Y] Card_rel_cardinal_rel lepoll_rel_eq_trans[of _ _ "|Y|\<^bsup>M\<^esup>"] by simp lemma lepoll_rel_Un: assumes "InfCard_rel(M,\)" "A \\<^bsup>M\<^esup> \" "B \\<^bsup>M\<^esup> \" "M(A)" "M(B)" "M(\)" shows "A \ B \\<^bsup>M\<^esup> \" proof - from assms have "A \ B \\<^bsup>M\<^esup> sum(A,B)" using Un_lepoll_rel_sum by simp moreover note assms moreover from this have "|sum(A,B)|\<^bsup>M\<^esup> \ \ \\<^bsup>M\<^esup> \" using sum_lepoll_rel_mono[of A \ B \] lepoll_rel_imp_cardinal_rel_le unfolding cadd_rel_def by auto ultimately show ?thesis using InfCard_rel_cdouble_eq Card_rel_cardinal_rel_eq InfCard_rel_is_Card_rel Card_rel_le_imp_lepoll_rel[of "sum(A,B)" \] lepoll_rel_trans[of "A\B"] by auto qed lemma cardinal_rel_Un_le: assumes "InfCard_rel(M,\)" "|A|\<^bsup>M\<^esup> \ \" "|B|\<^bsup>M\<^esup> \ \" "M(\)" "M(A)" "M(B)" shows "|A \ B|\<^bsup>M\<^esup> \ \" using assms lepoll_rel_Un le_Card_rel_iff InfCard_rel_is_Card_rel by auto lemma Finite_cardinal_rel_iff': "M(i) \ Finite(|i|\<^bsup>M\<^esup>) \ Finite(i)" using eqpoll_rel_imp_Finite_iff[OF cardinal_rel_eqpoll_rel] by auto lemma cardinal_rel_subset_of_Card_rel: assumes "Card_rel(M,\)" "a \ \" "M(a)" "M(\)" shows "|a|\<^bsup>M\<^esup> < \ \ |a|\<^bsup>M\<^esup> = \" proof - from assms have "|a|\<^bsup>M\<^esup> < |\|\<^bsup>M\<^esup> \ |a|\<^bsup>M\<^esup> = |\|\<^bsup>M\<^esup>" using subset_imp_le_cardinal_rel[THEN le_iff[THEN iffD1]] by simp with assms show ?thesis using Card_rel_cardinal_rel_eq by auto qed lemma cardinal_rel_cases: includes Ord_dests assumes "M(\)" "M(X)" shows "Card_rel(M,\) \ |X|\<^bsup>M\<^esup> < \ \ \ |X|\<^bsup>M\<^esup> \ \" using assms not_le_iff_lt Card_rel_is_Ord Ord_cardinal_rel by auto end \ \\<^locale>\M_library\\ subsection\Countable and uncountable sets\ -definition (* FIXME: From Cardinal_Library, on the context of AC *) - countable :: "i\o" where - "countable(X) \ X \ \" - relativize functional "countable" "countable_rel" external relationalize "countable_rel" "is_countable" notation countable_rel (\countable\<^bsup>_\<^esup>'(_')\) abbreviation countable_r_set :: "[i,i]\o" (\countable\<^bsup>_\<^esup>'(_')\) where "countable\<^bsup>M\<^esup>(i) \ countable_rel(##M,i)" context M_library begin lemma countableI[intro]: "X \\<^bsup>M\<^esup> \ \ countable_rel(M,X)" unfolding countable_rel_def by simp lemma countableD[dest]: "countable_rel(M,X) \ X \\<^bsup>M\<^esup> \" unfolding countable_rel_def by simp lemma countable_rel_iff_cardinal_rel_le_nat: "M(X) \ countable_rel(M,X) \ |X|\<^bsup>M\<^esup> \ \" using le_Card_rel_iff[of \ X] Card_rel_nat unfolding countable_rel_def by simp lemma lepoll_rel_countable_rel: "X \\<^bsup>M\<^esup> Y \ countable_rel(M,Y) \ M(X) \ M(Y) \ countable_rel(M,X)" using lepoll_rel_trans[of X Y] by blast \ \Next lemma can be proved without using AC\ lemma surj_rel_countable_rel: "countable_rel(M,X) \ f \ surj_rel(M,X,Y) \ M(X) \ M(Y) \ M(f) \ countable_rel(M,Y)" using surj_rel_implies_cardinal_rel_le[of f X Y, THEN le_trans] countable_rel_iff_cardinal_rel_le_nat by simp lemma Finite_imp_countable_rel: "Finite_rel(M,X) \ M(X) \ countable_rel(M,X)" unfolding Finite_rel_def by (auto intro:InfCard_rel_nat nats_le_InfCard_rel[of _ \, THEN le_imp_lepoll_rel] dest!:eq_lepoll_rel_trans[of X _ \] ) end \ \\<^locale>\M_library\\ lemma (in M_cardinal_UN_lepoll) countable_rel_imp_countable_rel_UN: assumes "countable_rel(M,J)" "\i. i\J \ countable_rel(M,X(i))" shows "countable_rel(M,\i\J. X(i))" - using assms leqpoll_rel_imp_cardinal_rel_UN_le[of \] InfCard_rel_nat + using assms lepoll_rel_imp_cardinal_rel_UN_le[of \] InfCard_rel_nat InfCard_rel_is_Card_rel j.UN_closed countable_rel_iff_cardinal_rel_le_nat j.Pi_assumptions Card_rel_le_imp_lepoll_rel[of J \] Card_rel_cardinal_rel_eq[of \] by auto locale M_cardinal_library = M_library + M_replacement + assumes lam_replacement_inj_rel:"lam_replacement(M, \x. inj\<^bsup>M\<^esup>(fst(x),snd(x)))" and cardinal_lib_assms1: "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)\)" and cardinal_lib_assms2: "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)\)" and cardinal_lib_assms3: "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)\)" and cardinal_rel_separation : "separation(M, \\x,y\. cardinal_rel(M,x) = y)" and separation_cardinal_rel_lt : "M(\) \ separation(M, \Z . cardinal_rel(M,Z) < \)" begin -lemma cdlt_assms: "M(G) \ M(Q) \ separation(M, \p. \x\G. x \ snd(p) \ (\s\fst(p). \s, x\ \ Q))" - using lam_replacement_snd lam_replacement_fst lam_replacement_hcomp separation_in - lam_replacement_Pair[THEN [5] lam_replacement_hcomp2] - by (rule_tac separation_ball,simp_all,rule_tac separation_iff',auto) - (rule_tac separation_all,auto simp:lam_replacement_constant) - -lemma cdlt_assms': "M(x) \ M(Q) \ separation(M, \a . \s\x. \s, a\ \ Q)" - using separation_in[OF _ - lam_replacement_hcomp2[OF _ _ _ _ lam_replacement_Pair] _ - lam_replacement_constant] - separation_ball lam_replacement_hcomp lam_replacement_fst lam_replacement_snd +lemma cdlt_assms: "M(x) \ M(Q) \ separation(M, \a . \s\x. \s, a\ \ Q)" + using separation_in[OF _ lam_replacement_product] separation_ball + lam_replacement_fst lam_replacement_snd lam_replacement_constant by simp_all lemma countable_rel_union_countable_rel: assumes "\x. x \ C \ countable_rel(M,x)" "countable_rel(M,C)" "M(C)" shows "countable_rel(M,\C)" proof - have "x \ (if M(i) then i else 0) \ M(i)" for x i by (cases "M(i)") auto then interpret M_replacement_lepoll M "\_ x. if M(x) then x else 0" using lam_replacement_if[OF lam_replacement_identity lam_replacement_constant[OF nonempty], where b=M] lam_replacement_inj_rel lam_replacement_minimum proof(unfold_locales,auto simp add: separation_def) fix b f assume "M(b)" "M(f)" show "lam_replacement(M, \x. \ i. x \ if_range_F_else_F(\x. if M(x) then x else 0, b, f, i))" proof (cases "b=0") case True with \M(f)\ show ?thesis using cardinal_lib_assms1 by (simp_all; rule_tac lam_Least_assumption_ifM_b0)+ next case False with \M(f)\ \M(b)\ show ?thesis using cardinal_lib_assms1 separation_Ord lam_replacement_minimum by (rule_tac lam_Least_assumption_ifM_bnot0) auto qed qed note \M(C)\ moreover have "w \ (if M(x) then x else 0) \ M(x)" for w x by (cases "M(x)") auto ultimately interpret M_cardinal_UN_lepoll _ "\c. if M(c) then c else 0" C using lepoll_assumptions lam_replacement_minimum by unfold_locales auto have "(if M(i) then i else 0) = i" if "i\C" for i using transM[OF _ \M(C)\] that by simp then show ?thesis using assms countable_rel_imp_countable_rel_UN by simp qed end \ \\<^locale>\M_cardinal_library\\ abbreviation uncountable_rel :: "[i\o,i]\o" where "uncountable_rel(M,X) \ \ countable_rel(M,X)" context M_library begin lemma uncountable_rel_iff_nat_lt_cardinal_rel: "M(X) \ uncountable_rel(M,X) \ \ < |X|\<^bsup>M\<^esup>" using countable_rel_iff_cardinal_rel_le_nat not_le_iff_lt by simp lemma uncountable_rel_not_empty: "uncountable_rel(M,X) \ X \ 0" using empty_lepoll_relI by auto lemma uncountable_rel_imp_Infinite: "uncountable_rel(M,X) \ M(X) \ Infinite(X)" using uncountable_rel_iff_nat_lt_cardinal_rel[of X] lepoll_rel_nat_imp_Infinite[of X] cardinal_rel_le_imp_lepoll_rel[of \ X] leI by simp lemma uncountable_rel_not_subset_countable_rel: assumes "countable_rel(M,X)" "uncountable_rel(M,Y)" "M(X)" "M(Y)" shows "\ (Y \ X)" using assms lepoll_rel_trans subset_imp_lepoll_rel[of Y X] by blast subsection\Results on Aleph\_rels\ lemma nat_lt_Aleph_rel1: "\ < \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" by (simp add: Aleph_rel_succ Aleph_rel_zero lt_csucc_rel) lemma zero_lt_Aleph_rel1: "0 < \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" by (rule lt_trans[of _ "\"], auto simp add: ltI nat_lt_Aleph_rel1) lemma le_Aleph_rel1_nat: "M(k) \ Card_rel(M,k) \ k<\\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ k \ \" by (simp add: Aleph_rel_succ Aleph_rel_zero Card_rel_lt_csucc_rel_iff Card_rel_nat) lemma lesspoll_rel_Aleph_rel_succ: assumes "Ord(\)" and types:"M(\)" "M(d)" shows "d \\<^bsup>M\<^esup> \\<^bsub>succ(\)\<^esub>\<^bsup>M\<^esup> \ d \\<^bsup>M\<^esup> \\<^bsub>\\<^esub>\<^bsup>M\<^esup>" using assms Aleph_rel_succ Card_rel_is_Ord Ord_Aleph_rel lesspoll_rel_csucc_rel by simp lemma cardinal_rel_Aleph_rel [simp]: "Ord(\) \ M(\) \ |\\<^bsub>\\<^esub>\<^bsup>M\<^esup>|\<^bsup>M\<^esup> = \\<^bsub>\\<^esub>\<^bsup>M\<^esup>" using Card_rel_cardinal_rel_eq by simp \ \Could be proved without using AC\ lemma Aleph_rel_lesspoll_rel_increasing: includes Aleph_rel_intros assumes "M(b)" "M(a)" shows "a < b \ \\<^bsub>a\<^esub>\<^bsup>M\<^esup> \\<^bsup>M\<^esup> \\<^bsub>b\<^esub>\<^bsup>M\<^esup>" using assms cardinal_rel_lt_iff_lesspoll_rel[of "\\<^bsub>a\<^esub>\<^bsup>M\<^esup>" "\\<^bsub>b\<^esub>\<^bsup>M\<^esup>"] Aleph_rel_increasing[of a b] Card_rel_cardinal_rel_eq[of "\\<^bsub>b\<^esub>"] lt_Ord lt_Ord2 Card_rel_Aleph_rel[THEN Card_rel_is_Ord] by auto lemma uncountable_rel_iff_subset_eqpoll_rel_Aleph_rel1: includes Ord_dests assumes "M(X)" notes Aleph_rel_zero[simp] Card_rel_nat[simp] Aleph_rel_succ[simp] shows "uncountable_rel(M,X) \ (\S[M]. S \ X \ S \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>)" proof assume "uncountable_rel(M,X)" with \M(X)\ have "\\<^bsub>1\<^esub>\<^bsup>M\<^esup> \\<^bsup>M\<^esup> X" using uncountable_rel_iff_nat_lt_cardinal_rel cardinal_rel_lt_csucc_rel_iff' cardinal_rel_le_imp_lepoll_rel by auto with \M(X)\ obtain S where "M(S)" "S \ X" "S \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" using lepoll_rel_imp_subset_bij_rel by auto then show "\S[M]. S \ X \ S \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" using cardinal_rel_cong Card_rel_csucc_rel[of \] Card_rel_cardinal_rel_eq by auto next note Aleph_rel_lesspoll_rel_increasing[of 1 0,simplified] assume "\S[M]. S \ X \ S \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" moreover have eq:"\\<^bsub>1\<^esub>\<^bsup>M\<^esup> = (\\<^sup>+)\<^bsup>M\<^esup>" by auto moreover from calculation \M(X)\ have A:"(\\<^sup>+)\<^bsup>M\<^esup> \\<^bsup>M\<^esup> X" using subset_imp_lepoll_rel[THEN [2] eq_lepoll_rel_trans, of "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" _ X, OF eqpoll_rel_sym] by auto with \M(X)\ show "uncountable_rel(M,X)" using lesspoll_rel_trans1[OF lepoll_rel_trans[OF A _] \\ \\<^bsup>M\<^esup> (\\<^sup>+)\<^bsup>M\<^esup>\] lesspoll_rel_not_refl by auto qed lemma UN_if_zero: "M(K) \ (\x\K. if M(x) then G ` x else 0) =(\x\K. G ` x)" using transM[of _ K] by auto lemma mem_F_bound1: fixes F G defines "F \ \_ x. if M(x) then G`x else 0" shows "x\F(A,c) \ c \ (range(f) \ domain(G) )" using apply_0 unfolding F_def by (cases "M(c)", auto simp:F_def drSR_Y_def dC_F_def) end \ \\<^locale>\M_library\\ context M_cardinal_library begin lemma lt_Aleph_rel_imp_cardinal_rel_UN_le_nat: "function(G) \ domain(G) \\<^bsup>M\<^esup> \ \ \n\domain(G). |G`n|\<^bsup>M\<^esup><\\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ M(G) \ |\n\domain(G). G`n|\<^bsup>M\<^esup>\\" proof - assume "M(G)" moreover from this have "x \ (if M(i) then G ` i else 0) \ M(i)" for x i by (cases "M(i)") auto moreover have "separation(M, M)" unfolding separation_def by auto ultimately interpret M_replacement_lepoll M "\_ x. if M(x) then G`x else 0" using lam_replacement_inj_rel cardinal_lib_assms2 mem_F_bound1[of _ _ G] lam_if_then_replacement_apply lam_replacement_minimum by (unfold_locales, simp_all) (rule lam_Least_assumption_general[where U="\_. domain(G)"], auto) note \M(G)\ moreover have "w \ (if M(x) then G ` x else 0) \ M(x)" for w x by (cases "M(x)") auto ultimately interpret M_cardinal_UN_lepoll _ "\n. if M(n) then G`n else 0" "domain(G)" using lepoll_assumptions1[where S="domain(G)",unfolded lepoll_assumptions1_def] cardinal_lib_assms2 lepoll_assumptions by (unfold_locales, auto) assume "function(G)" let ?N="domain(G)" and ?R="\n\domain(G). G`n" assume "?N \\<^bsup>M\<^esup> \" assume Eq1: "\n\?N. |G`n|\<^bsup>M\<^esup><\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" { fix n assume "n\?N" with Eq1 \M(G)\ have "|G`n|\<^bsup>M\<^esup> \ \" - using le_Aleph_rel1_nat[of "|G ` n|\<^bsup>M\<^esup>"] leqpoll_rel_imp_cardinal_rel_UN_le + using le_Aleph_rel1_nat[of "|G ` n|\<^bsup>M\<^esup>"] lepoll_rel_imp_cardinal_rel_UN_le UN_if_zero[of "domain(G)" G] by (auto dest:transM) } then have "n\?N \ |G`n|\<^bsup>M\<^esup> \ \" for n . moreover note \?N \\<^bsup>M\<^esup> \\ \M(G)\ moreover have "(if M(i) then G ` i else 0) \ G ` i" for i by auto with \M(G)\ have "|if M(i) then G ` i else 0|\<^bsup>M\<^esup> \ |G ` i|\<^bsup>M\<^esup>" for i proof(cases "M(i)") case True with \M(G)\ show ?thesis using Ord_cardinal_rel[OF apply_closed] by simp next case False then have "i\domain(G)" using transM[OF _ domain_closed[OF \M(G)\]] by auto then show ?thesis using Ord_cardinal_rel[OF apply_closed] apply_0 by simp qed ultimately show ?thesis - using InfCard_rel_nat leqpoll_rel_imp_cardinal_rel_UN_le[of \] + using InfCard_rel_nat lepoll_rel_imp_cardinal_rel_UN_le[of \] UN_if_zero[of "domain(G)" G] le_trans[of "|if M(_) then G ` _ else 0|\<^bsup>M\<^esup>" "|G ` _|\<^bsup>M\<^esup>" \] by auto blast qed lemma Aleph_rel1_eq_cardinal_rel_vimage: "f:\\<^bsub>1\<^esub>\<^bsup>M\<^esup>\\<^bsup>M\<^esup>\ \ \n\\. |f-``{n}|\<^bsup>M\<^esup> = \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" proof - assume "f:\\<^bsub>1\<^esub>\<^bsup>M\<^esup>\\<^bsup>M\<^esup>\" then have "function(f)" "domain(f) = \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "range(f)\\" "f\\\<^bsub>1\<^esub>\<^bsup>M\<^esup>\\" "M(f)" using mem_function_space_rel[OF \f\_\] domain_of_fun fun_is_function range_fun_subset_codomain function_space_rel_char by auto let ?G="\n\range(f). f-``{n}" from \f:\\<^bsub>1\<^esub>\<^bsup>M\<^esup>\\\ have "range(f) \ \" "domain(?G) = range(f)" using range_fun_subset_codomain by simp_all from \f:\\<^bsub>1\<^esub>\<^bsup>M\<^esup>\\\ \M(f)\ \range(f) \ \\ have "M(f-``{n})" if "n \ range(f)" for n using that transM[of _ \] by auto with \M(f)\ \range(f) \ \\ have "domain(?G) \\<^bsup>M\<^esup> \" "M(?G)" using subset_imp_lepoll_rel lam_closed[of "\x . f-``{x}"] cardinal_lib_assms4 by simp_all have "function(?G)" by (simp add:function_lam) from \f:\\<^bsub>1\<^esub>\<^bsup>M\<^esup>\\\ have "n\\ \ f-``{n} \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" for n using Pi_vimage_subset by simp with \range(f) \ \\ have "\\<^bsub>1\<^esub>\<^bsup>M\<^esup> = (\n\range(f). f-``{n})" proof (intro equalityI, intro subsetI) fix x assume "x \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" with \f:\\<^bsub>1\<^esub>\<^bsup>M\<^esup>\\\ \function(f)\ \domain(f) = \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ have "x \ f-``{f`x}" "f`x \ range(f)" using function_apply_Pair vimage_iff apply_rangeI by simp_all then show "x \ (\n\range(f). f-``{n})" by auto qed auto { assume "\n\range(f). |f-``{n}|\<^bsup>M\<^esup> < \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" then have "\n\domain(?G). |?G`n|\<^bsup>M\<^esup> < \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" using zero_lt_Aleph_rel1 by (auto) with \function(?G)\ \domain(?G) \\<^bsup>M\<^esup> \\ \M(?G)\ have "|\n\domain(?G). ?G`n|\<^bsup>M\<^esup>\\" using lt_Aleph_rel_imp_cardinal_rel_UN_le_nat[of ?G] by auto then have "|\n\range(f). f-``{n}|\<^bsup>M\<^esup>\\" by simp with \\\<^bsub>1\<^esub>\<^bsup>M\<^esup> = _\ have "|\\<^bsub>1\<^esub>\<^bsup>M\<^esup>|\<^bsup>M\<^esup> \ \" by auto then have "\\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ \" using Card_rel_Aleph_rel Card_rel_cardinal_rel_eq by auto then have "False" using nat_lt_Aleph_rel1 by (blast dest:lt_trans2) } with \range(f)\\\ \M(f)\ obtain n where "n\\" "\(|f -`` {n}|\<^bsup>M\<^esup> < \\<^bsub>1\<^esub>\<^bsup>M\<^esup>)" "M(f -`` {n})" using nat_into_M by auto moreover from this have "\\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ |f-``{n}|\<^bsup>M\<^esup>" using not_lt_iff_le Card_rel_is_Ord by simp moreover note \n\\ \ f-``{n} \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ ultimately show ?thesis using subset_imp_le_cardinal_rel[THEN le_anti_sym, of _ "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>"] Card_rel_Aleph_rel Card_rel_cardinal_rel_eq by auto qed \ \There is some asymmetry between assumptions and conclusion (\<^term>\eqpoll_rel\ versus \<^term>\cardinal_rel\)\ lemma eqpoll_rel_Aleph_rel1_cardinal_rel_vimage: assumes "Z \\<^bsup>M\<^esup> (\\<^bsub>1\<^esub>\<^bsup>M\<^esup>)" "f \ Z \\<^bsup>M\<^esup> \" "M(Z)" shows "\n\\. |f-``{n}|\<^bsup>M\<^esup> = \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" proof - have "M(1)" "M(\)" by simp_all then have "M(\\<^bsub>1\<^esub>\<^bsup>M\<^esup>)" by simp with assms \M(1)\ obtain g where A:"g\bij_rel(M,\\<^bsub>1\<^esub>\<^bsup>M\<^esup>,Z)" "M(g)" using eqpoll_rel_sym unfolding eqpoll_rel_def by blast with \f : Z \\<^bsup>M\<^esup> \\ assms have "M(f)" "converse(g) \ bij_rel(M,Z, \\<^bsub>1\<^esub>\<^bsup>M\<^esup>)" "f\Z\\" "g\ \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\Z" using bij_rel_is_fun_rel bij_rel_converse_bij_rel bij_rel_char function_space_rel_char by simp_all with \g\bij_rel(M,\\<^bsub>1\<^esub>\<^bsup>M\<^esup>,Z)\ \M(g)\ have "f O g : \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \\<^bsup>M\<^esup> \" "M(converse(g))" using comp_fun[OF _ \f\ Z\_\,of g] function_space_rel_char by simp_all then obtain n where "n\\" "|(f O g)-``{n}|\<^bsup>M\<^esup> = \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" using Aleph_rel1_eq_cardinal_rel_vimage by auto with \M(f)\ \M(converse(g))\ have "M(converse(g) `` (f -`` {n}))" "f -`` {n} \ Z" using image_comp converse_comp Pi_iff[THEN iffD1,OF \f\Z\\\] vimage_subset unfolding vimage_def using transM[OF \n\\\] by auto from \n\\\ \|(f O g)-``{n}|\<^bsup>M\<^esup> = \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ have "\\<^bsub>1\<^esub>\<^bsup>M\<^esup> = |converse(g) `` (f -``{n})|\<^bsup>M\<^esup>" using image_comp converse_comp unfolding vimage_def by auto also from \converse(g) \ bij_rel(M,Z, \\<^bsub>1\<^esub>\<^bsup>M\<^esup>)\ \f: Z\\<^bsup>M\<^esup> \\ \M(Z)\ \M(f)\ \M(\\<^bsub>1\<^esub>\<^bsup>M\<^esup>)\ \M(converse(g) `` (f -`` {n}))\ have "\ = |f -``{n}|\<^bsup>M\<^esup>" using range_of_subset_eqpoll_rel[of "converse(g)" Z _ "f -``{n}", OF bij_rel_is_inj_rel[OF \converse(g)\_\] \f -`` {n} \ Z\] cardinal_rel_cong vimage_closed[OF singleton_closed[OF transM[OF \n\\\]],of f] by auto finally show ?thesis using \n\_\ by auto qed end \ \\<^locale>\M_cardinal_library\\ subsection\Applications of transfinite recursive constructions\ locale M_cardinal_library_extra = M_cardinal_library + assumes replacement_trans_apply_image: "M(f) \ M(\) \ Ord(\) \ strong_replacement(M, \x y. x\\ \ y = \x, transrec(x, \a g. f ` (g `` a))\)" begin lemma rec_constr_type_rel: assumes "f:Pow_rel(M,G)\\<^bsup>M\<^esup> G" "Ord(\)" "M(G)" shows "M(\) \ rec_constr(f,\) \ G" using assms(2) proof(induct rule:trans_induct) case (step \) with assms have "{rec_constr(f, x) . x \ \} = {y . x \ \, y=rec_constr(f, x)}" (is "_ = ?Y") "M(f)" using transM[OF _ \M(\)\] function_space_rel_char Ord_in_Ord by auto moreover from assms this step \M(\)\ \Ord(\)\ have "M({y . x \ \, y=})" (is "M(?Z)") using strong_replacement_closed[OF replacement_trans_apply_image,of f \ \,OF _ _ _ _ univalent_conjI2[where P="\x _ . x\\",OF univalent_triv]] transM[OF _ \M(\)\] transM[OF step(2) \M(G)\] Ord_in_Ord unfolding rec_constr_def by auto moreover from assms this step \M(\)\ \Ord(\)\ have "?Y = {snd(y) . y\?Z}" proof(intro equalityI, auto) fix u assume "u\\" with assms this step \M(\)\ \Ord(\)\ have " \ ?Z" "rec_constr(f, u) = snd()" by auto then show "\x\{y . x \ \, y = \x, rec_constr(f, x)\}. rec_constr(f, u) = snd(x)" using bexI[of _ u] by force qed moreover from \M(?Z)\ \?Y = _\ have "M(?Y)" using RepFun_closed[OF lam_replacement_imp_strong_replacement[OF lam_replacement_snd] \M(?Z)\] fst_snd_closed[THEN conjunct2] transM[OF _ \M(?Z)\] by simp moreover from assms step have "{rec_constr(f, x) . x \ \} \ Pow(G)" (is "?X\_") using transM[OF _ \M(\)\] function_space_rel_char by auto moreover from assms calculation step have "M(?X)" by simp moreover from calculation \M(G)\ have "?X\Pow_rel(M,G)" using Pow_rel_char by simp ultimately have "f`?X \ G" using assms apply_type[OF mem_function_space_rel[of f],of "Pow_rel(M,G)" G ?X] by auto then show ?case by (subst rec_constr_unfold,simp) qed lemma rec_constr_closed : assumes "f:Pow_rel(M,G)\\<^bsup>M\<^esup> G" "Ord(\)" "M(G)" "M(\)" shows "M(rec_constr(f,\))" using transM[OF rec_constr_type_rel \M(G)\] assms by auto lemma lambda_rec_constr_closed : assumes "Ord(\)" "M(\)" "M(f)" "f:Pow_rel(M,G)\\<^bsup>M\<^esup> G" "M(G)" shows "M(\\\\ . rec_constr(f,\))" using lam_closed2[OF replacement_trans_apply_image,unfolded rec_constr_def[symmetric],of f \] rec_constr_type_rel[OF \f\_\ Ord_in_Ord[of \]] transM[OF _ \M(G)\] assms by simp text\The next lemma is an application of recursive constructions. It works under the assumption that whenever the already constructed subsequence is small enough, another element can be added.\ lemma bounded_cardinal_rel_selection: includes Ord_dests assumes "\Z. |Z|\<^bsup>M\<^esup> < \ \ Z \ G \ M(Z) \ \a\G. \s\Z. \Q" "b\G" "Card_rel(M,\)" "M(G)" "M(Q)" "M(\)" shows "\S[M]. S : \ \\<^bsup>M\<^esup> G \ (\\ \ \. \\ \ \. \<\ \ ,S`\>\Q)" proof - from assms have "M(x) \ M({a \ G . \s\x. \s, a\ \ Q})" for x using separation_orthogonal by simp let ?cdlt\="{Z\Pow_rel(M,G) . |Z|\<^bsup>M\<^esup><\}" \ \“cardinal\_rel less than \<^term>\\\”\ and ?inQ="\Y.{a\G. \s\Y. \Q}" from \M(G)\ \Card_rel(M,\)\ \M(\)\ have "M(?cdlt\)" "Ord(\)" using separation_cardinal_rel_lt[OF \M(\)\] Card_rel_is_Ord by simp_all from assms have H:"\a. a \ ?inQ(Y)" if "Y\?cdlt\" for Y proof - { fix Y assume "Y\?cdlt\" then have A:"Y\Pow_rel(M,G)" "|Y|\<^bsup>M\<^esup><\" by simp_all then have "Y\G" "M(Y)" using Pow_rel_char[OF \M(G)\] by simp_all with A obtain a where "a\G" "\s\Y. \Q" using assms(1) by force with \M(G)\ have "\a. a \ ?inQ(Y)" by auto } then show ?thesis using that by simp qed then have "\f[M]. f \ Pi_rel(M,?cdlt\,?inQ) \ f \ Pi(?cdlt\,?inQ)" proof - from \\x. M(x) \ M({a \ G . \s\x. \s, a\ \ Q})\ \M(G)\ have "x \ {Z \ Pow\<^bsup>M\<^esup>(G) . |Z|\<^bsup>M\<^esup> < \} \ M({a \ G . \s\x. \s, a\ \ Q})" for x by (auto dest:transM) with\M(G)\ \\x. M(x) \ M({a \ G . \s\x. \s, a\ \ Q})\ \M(Q)\ \M(?cdlt\)\ interpret M_Pi_assumptions_choice M ?cdlt\ ?inQ - using cdlt_assms[where Q=Q] lam_replacement_Collect_ball_Pair[THEN - lam_replacement_imp_strong_replacement] surj_imp_inj_replacement3 - lam_replacement_hcomp2[OF lam_replacement_constant - lam_replacement_Collect_ball_Pair _ _ lam_replacement_minimum, - unfolded lam_replacement_def] - lam_replacement_hcomp lam_replacement_Sigfun[OF - lam_replacement_Collect_ball_Pair, of G Q, THEN - lam_replacement_imp_strong_replacement] separation_orthogonal + using cdlt_assms[where Q=Q] surj_imp_inj_replacement3 + lam_replacement_Collect_ball_Pair[THEN lam_replacement_imp_strong_replacement] + lam_replacement_minimum[THEN [5] lam_replacement_hcomp2,OF + lam_replacement_constant lam_replacement_Collect_ball_Pair,unfolded lam_replacement_def] + lam_replacement_Sigfun[OF lam_replacement_Collect_ball_Pair,THEN + lam_replacement_imp_strong_replacement] by unfold_locales (blast dest: transM, auto dest:transM) show ?thesis using AC_Pi_rel Pi_rel_char H by auto qed then obtain f where f_type:"f \ Pi_rel(M,?cdlt\,?inQ)" "f \ Pi(?cdlt\,?inQ)" and "M(f)" by auto moreover define Cb where "Cb \ \_\Pow_rel(M,G)-?cdlt\. b" moreover from \b\G\ \M(?cdlt\)\ \M(G)\ have "Cb \ Pow_rel(M,G)-?cdlt\ \ G" "M(Cb)" - using lam_closed[of "\_.b" "Pow_rel(M,G)-?cdlt\"] - tag_replacement transM[OF \b\G\] + using lam_closed[of "\_.b" "Pow_rel(M,G)-?cdlt\"] tag_replacement transM[OF \b\G\] unfolding Cb_def by auto moreover note \Card_rel(M,\)\ ultimately have "f \ Cb : (\x\Pow_rel(M,G). ?inQ(x) \ G)" using fun_Pi_disjoint_Un[ of f ?cdlt\ ?inQ Cb "Pow_rel(M,G)-?cdlt\" "\_.G"] Diff_partition[of "{Z\Pow_rel(M,G). |Z|\<^bsup>M\<^esup><\}" "Pow_rel(M,G)", OF Collect_subset] by auto moreover have "?inQ(x) \ G = G" for x by auto moreover from calculation have "f \ Cb : Pow_rel(M,G) \ G" using function_space_rel_char by simp ultimately have "f \ Cb : Pow_rel(M,G) \\<^bsup>M\<^esup> G" using function_space_rel_char \M(f)\ \M(Cb)\ Pow_rel_closed \M(G)\ by auto define S where "S\\\\\. rec_constr(f \ Cb, \)" from \f \ Cb: Pow_rel(M,G) \\<^bsup>M\<^esup> G\ \Card_rel(M,\)\ \M(\)\ \M(G)\ have "S : \ \ G" "M(f \ Cb)" unfolding S_def using Ord_in_Ord[OF Card_rel_is_Ord] rec_constr_type_rel lam_type transM[OF _ \M(\)\] function_space_rel_char by auto moreover from \f\Cb \ _\\<^bsup>M\<^esup> G\ \Card_rel(M,\)\ \M(\)\ \M(G)\ \M(f \ Cb)\ \Ord(\)\ have "M(S)" unfolding S_def using lambda_rec_constr_closed by simp moreover have "\\\\. \\\\. \ < \ \ \S`\, S`\\ \ Q" proof (intro ballI impI) fix \ \ assume "\\\" with \Card_rel(M,\)\ \M(S)\ \M(\)\ have "\\\" "M(S``\)" "M(\)" "{S`x . x \ \} = {restrict(S,\)`x . x \ \}" using transM[OF \\\\\ \M(\)\] image_closed Card_rel_is_Ord OrdmemD by auto with \\\_\ \Card_rel(M,\)\ \M(\)\ have "{rec_constr(f \ Cb, x) . x\\} = {S`x . x \ \}" using Ord_trans[OF _ _ Card_rel_is_Ord, of _ \ \] unfolding S_def by auto moreover from \\\\\ \S : \ \ G\ \Card_rel(M,\)\ \M(\)\ \M(S``\)\ have "{S`x . x \ \} \ G" "M({S`x . x \ \})" using Ord_trans[OF _ _ Card_rel_is_Ord, of _ \ \] apply_type[of S \ "\_. G"] by(auto,simp add:image_fun_subset[OF \S\_\ \\\_\]) moreover from \Card_rel(M,\)\ \\\\\ \S\_\ \\\\\ \M(S)\ \M(\)\ \M(G)\ \M(\)\ have "|{S`x . x \ \}|\<^bsup>M\<^esup> < \" using \{S`x . x\\} = {restrict(S,\)`x . x\\}\[symmetric] cardinal_rel_RepFun_apply_le[of "restrict(S,\)" \ G, OF restrict_type2[of S \ "\_.G" \] restrict_closed] Ord_in_Ord Ord_cardinal_rel lt_trans1[of "|{S`x . x \ \}|\<^bsup>M\<^esup>" "|\|\<^bsup>M\<^esup>" \] Card_rel_lt_iff[THEN iffD2, of \ \, OF _ _ _ _ ltI] Card_rel_is_Ord by auto moreover have "\x\\. \}> \ Q" proof - from calculation and f_type have "f ` {S`x . x \ \} \ {a\G. \x\\. \Q}" using apply_type[of f ?cdlt\ ?inQ "{S`x . x \ \}"] Pow_rel_char[OF \M(G)\] by simp then show ?thesis by simp qed moreover assume "\\\" "\ < \" moreover from this have "\\\" using ltD by simp moreover note \\\\\ \Cb \ Pow_rel(M,G)-?cdlt\ \ G\ ultimately show "\S`\, S`\\\Q" using fun_disjoint_apply1[of "{S`x . x \ \}" Cb f] domain_of_fun[of Cb] ltD[of \ \] by (subst (2) S_def, auto) (subst rec_constr_unfold, auto) qed moreover note \M(G)\ \M(\)\ ultimately show ?thesis using function_space_rel_char by auto qed text\The following basic result can, in turn, be proved by a bounded-cardinal\_rel selection.\ lemma Infinite_iff_lepoll_rel_nat: "M(Z) \ Infinite(Z) \ \ \\<^bsup>M\<^esup> Z" proof define Distinct where "Distinct = { \ Z\Z . x\y}" have "Distinct = {xy \ Z\Z . \a b. xy = \a, b\ \ a \ b}" (is "_=?A") unfolding Distinct_def by auto moreover assume "Infinite(Z)" "M(Z)" moreover from calculation have "M(Distinct)" using separation_dist by simp from \Infinite(Z)\ \M(Z)\ obtain b where "b\Z" using Infinite_not_empty by auto { fix Y assume "|Y|\<^bsup>M\<^esup> < \" "M(Y)" then have "Finite(Y)" using Finite_cardinal_rel_iff' ltD nat_into_Finite by auto with \Infinite(Z)\ have "Z \ Y" by auto } moreover have "(\W. M(W) \ |W|\<^bsup>M\<^esup> < \ \ W \ Z \ \a\Z. \s\W. \Distinct)" proof - fix W assume "M(W)" "|W|\<^bsup>M\<^esup> < \" "W \ Z" moreover from this have "Finite_rel(M,W)" using cardinal_rel_closed[OF \M(W)\] Card_rel_nat lt_Card_rel_imp_lesspoll_rel[of \,simplified,OF _ \|W|\<^bsup>M\<^esup> < \\] lesspoll_rel_nat_is_Finite_rel[of W] eqpoll_rel_imp_lepoll_rel eqpoll_rel_sym[OF cardinal_rel_eqpoll_rel,of W] lesspoll_rel_trans1[of W "|W|\<^bsup>M\<^esup>" \] by auto moreover from calculation have "\Z\W" using equalityI \Infinite(Z)\ by auto moreover from calculation show "\a\Z. \s\W. \Distinct" unfolding Distinct_def by auto qed moreover from \b\Z\ \M(Z)\ \M(Distinct)\ this obtain S where "S : \ \\<^bsup>M\<^esup> Z" "M(S)" "\\\\. \\\\. \ < \ \ ,S`\> \ Distinct" using bounded_cardinal_rel_selection[OF _ \b\Z\ Card_rel_nat,of Distinct] by blast moreover from this have "\ \ \ \ \ \ \ \ \\\ \ S`\ \ S`\" for \ \ unfolding Distinct_def by (rule_tac lt_neq_symmetry[of "\" "\\ \. S`\ \ S`\"]) auto moreover from this \S\_\ \M(Z)\ have "S\inj(\,Z)" using function_space_rel_char unfolding inj_def by auto ultimately show "\ \\<^bsup>M\<^esup> Z" unfolding lepoll_rel_def using inj_rel_char \M(Z)\ by auto next assume "\ \\<^bsup>M\<^esup> Z" "M(Z)" then show "Infinite(Z)" using lepoll_rel_nat_imp_Infinite by simp qed lemma Infinite_InfCard_rel_cardinal_rel: "Infinite(Z) \ M(Z) \ InfCard_rel(M,|Z|\<^bsup>M\<^esup>)" using lepoll_rel_eq_trans eqpoll_rel_sym lepoll_rel_nat_imp_Infinite Infinite_iff_lepoll_rel_nat Inf_Card_rel_is_InfCard_rel cardinal_rel_eqpoll_rel by simp lemma (in M_trans) mem_F_bound2: fixes F A defines "F \ \_ x. if M(x) then A-``{x} else 0" shows "x\F(A,c) \ c \ (range(f) \ range(A))" using apply_0 unfolding F_def by (cases "M(c)", auto simp:F_def drSR_Y_def dC_F_def) lemma Finite_to_one_rel_surj_rel_imp_cardinal_rel_eq: assumes "F \ Finite_to_one_rel(M,Z,Y) \ surj_rel(M,Z,Y)" "Infinite(Z)" "M(Z)" "M(Y)" shows "|Y|\<^bsup>M\<^esup> = |Z|\<^bsup>M\<^esup>" proof - have sep_true: "separation(M, M)" unfolding separation_def by auto note \M(Z)\ \M(Y)\ moreover from this assms have "M(F)" "F \ Z \ Y" unfolding Finite_to_one_rel_def using function_space_rel_char by simp_all moreover from this have "x \ (if M(i) then F -`` {i} else 0) \ M(i)" for x i by (cases "M(i)") auto moreover from calculation interpret M_replacement_lepoll M "\_ x. if M(x) then F-``{x} else 0" using lam_replacement_inj_rel mem_F_bound2 cardinal_lib_assms3 lam_replacement_vimage_sing_fun lam_replacement_minimum lam_replacement_if[OF _ lam_replacement_constant[OF nonempty],where b=M] sep_true by (unfold_locales, simp_all) (rule lam_Least_assumption_general[where U="\_. range(F)"], auto) have "w \ (if M(y) then F-``{y} else 0) \ M(y)" for w y by (cases "M(y)") auto moreover from \F\_\_\ have 0:"Finite(F-``{y})" if "y\Y" for y unfolding Finite_to_one_rel_def using vimage_fun_sing \F\Z\Y\ transM[OF that \M(Y)\] transM[OF _ \M(Z)\] that by simp ultimately interpret M_cardinal_UN_lepoll _ "\y. if M(y) then F-``{y} else 0" Y using cardinal_lib_assms3 lepoll_assumptions by unfold_locales (auto dest:transM simp del:mem_inj_abs) from \F\Z\Y\ have "Z = (\y\Y. {x\Z . F`x = y})" using apply_type by auto then show ?thesis proof (cases "Finite(Y)") case True with \Z = (\y\Y. {x\Z . F`x = y})\ and assms and \F\Z\Y\ show ?thesis using Finite_RepFun[THEN [2] Finite_Union, of Y "\y. F-``{y}"] 0 vimage_fun_sing[OF \F\Z\Y\] by simp next case False moreover from this \M(Y)\ have "Y \\<^bsup>M\<^esup> |Y|\<^bsup>M\<^esup>" using cardinal_rel_eqpoll_rel eqpoll_rel_sym eqpoll_rel_imp_lepoll_rel by auto moreover note assms moreover from \F\_\_\ have "Finite({x\Z . F`x = y})" "M(F-``{y})" if "y\Y" for y unfolding Finite_to_one_rel_def using transM[OF that \M(Y)\] transM[OF _ \M(Z)\] vimage_fun_sing[OF \F\Z\Y\] that by simp_all moreover from calculation have "|{x\Z . F`x = y}|\<^bsup>M\<^esup> \ \" if "y\Y" for y using Finite_cardinal_rel_in_nat that transM[OF that \M(Y)\] vimage_fun_sing[OF \F\Z\Y\] that by simp moreover from calculation have "|{x\Z . F`x = y}|\<^bsup>M\<^esup> \ |Y|\<^bsup>M\<^esup>" if "y\Y" for y using Infinite_imp_nats_lepoll_rel[THEN lepoll_rel_imp_cardinal_rel_le, of _ "|{x\Z . F`x = y}|\<^bsup>M\<^esup>"] that cardinal_rel_idem transM[OF that \M(Y)\] vimage_fun_sing[OF \F\Z\Y\] by auto ultimately have "|\y\Y. {x\Z . F`x = y}|\<^bsup>M\<^esup> \ |Y|\<^bsup>M\<^esup>" - using leqpoll_rel_imp_cardinal_rel_UN_le + using lepoll_rel_imp_cardinal_rel_UN_le Infinite_InfCard_rel_cardinal_rel[of Y] vimage_fun_sing[OF \F\Z\Y\] by(auto simp add:transM[OF _ \M(Y)\]) moreover from \F \ Finite_to_one_rel(M,Z,Y) \ surj_rel(M,Z,Y)\ \M(Z)\ \M(F)\ \M(Y)\ have "|Y|\<^bsup>M\<^esup> \ |Z|\<^bsup>M\<^esup>" using surj_rel_implies_cardinal_rel_le by auto moreover note \Z = (\y\Y. {x\Z . F`x = y})\ ultimately show ?thesis using le_anti_sym by auto qed qed lemma cardinal_rel_map_Un: assumes "Infinite(X)" "Finite(b)" "M(X)" "M(b)" shows "|{a \ b . a \ X}|\<^bsup>M\<^esup> = |X|\<^bsup>M\<^esup>" proof - have "(\a\X. a \ b) \ Finite_to_one_rel(M,X,{a \ b . a \ X})" "(\a\X. a \ b) \ surj_rel(M,X,{a \ b . a \ X})" "M({a \ b . a \ X})" unfolding def_surj_rel proof fix d have "Finite({a \ X . a \ b = d})" (is "Finite(?Y(b,d))") using \Finite(b)\ proof (induct arbitrary:d) case 0 have "{a \ X . a \ 0 = d} = (if d\X then {d} else 0)" by auto then show ?case by simp next case (cons c b) from \c \ b\ have "?Y(cons(c,b),d) \ (if c\d then ?Y(b,d) \ ?Y(b,d-{c}) else 0)" by auto with cons show ?case using subset_Finite by simp qed moreover assume "d \ {x \ b . x \ X}" ultimately show "Finite({a \ X . M(a) \ (\x\X. x \ b) ` a = d})" using subset_Finite[of "{a \ X . M(a) \ (\x\X. x \ b) ` a = d}" "{a \ X . (\x\X. x \ b) ` a = d}"] by auto next note \M(X)\ \M(b)\ moreover show "M(\a\X. a \ b)" using lam_closed[of "\ x . x\b",OF _ \M(X)\] Un_closed[OF transM[OF _ \M(X)\] \M(b)\] tag_union_replacement[OF \M(b)\] by simp moreover from this have "{a \ b . a \ X} = (\x\X. x \ b) `` X" using image_lam by simp with calculation show "M({a \ b . a \ X})" by auto moreover from calculation show "(\a\X. a \ b) \ X \\<^bsup>M\<^esup> {a \ b . a \ X}" using function_space_rel_char by (auto intro:lam_funtype) ultimately show "(\a\X. a \ b) \ surj\<^bsup>M\<^esup>(X,{a \ b . a \ X})" "M({a \ b . a \ X})" using surj_rel_char function_space_rel_char unfolding surj_def by auto next qed (simp add:\M(X)\) moreover from assms this show ?thesis using Finite_to_one_rel_surj_rel_imp_cardinal_rel_eq by simp qed end \ \\<^locale>\M_cardinal_library_extra\\ context M_library begin subsection\Results on relative cardinal exponentiation\ lemma cexp_rel_eqpoll_rel_cong: assumes "A \\<^bsup>M\<^esup> A'" "B \\<^bsup>M\<^esup> B'" "M(A)" "M(A')" "M(B)" "M(B')" shows "A\<^bsup>\B,M\<^esup> = A'\<^bsup>\B',M\<^esup>" unfolding cexp_rel_def using cardinal_rel_eqpoll_rel_iff function_space_rel_eqpoll_rel_cong assms by simp lemma cexp_rel_cexp_rel_cmult: assumes "M(\)" "M(\1)" "M(\2)" shows "(\\<^bsup>\\1,M\<^esup>)\<^bsup>\\2,M\<^esup> = \\<^bsup>\\2 \\<^bsup>M\<^esup> \1,M\<^esup>" proof - have "(\\<^bsup>\\1,M\<^esup>)\<^bsup>\\2,M\<^esup> = (\1 \\<^bsup>M\<^esup> \)\<^bsup>\\2,M\<^esup>" using cardinal_rel_eqpoll_rel by (intro cexp_rel_eqpoll_rel_cong) (simp_all add:assms cexp_rel_def) also from assms have " \ = \\<^bsup>\\2 \ \1,M\<^esup>" unfolding cexp_rel_def using curry_eqpoll_rel[THEN cardinal_rel_cong] by blast also have " \ = \\<^bsup>\\2 \\<^bsup>M\<^esup> \1,M\<^esup>" using cardinal_rel_eqpoll_rel[THEN eqpoll_rel_sym] unfolding cmult_rel_def by (intro cexp_rel_eqpoll_rel_cong) (auto simp add:assms) finally show ?thesis . qed lemma cardinal_rel_Pow_rel: "M(X) \ |Pow_rel(M,X)|\<^bsup>M\<^esup> = 2\<^bsup>\X,M\<^esup>" \ \Perhaps it's better with |X|\ using cardinal_rel_eqpoll_rel_iff[THEN iffD2, OF _ _ Pow_rel_eqpoll_rel_function_space_rel] unfolding cexp_rel_def by simp lemma cantor_cexp_rel: assumes "Card_rel(M,\)" "M(\)" shows "\ < 2\<^bsup>\\,M\<^esup>" using assms Card_rel_is_Ord Card_rel_cexp_rel proof (intro not_le_iff_lt[THEN iffD1] notI) assume "2\<^bsup>\\,M\<^esup> \ \" with assms have "|Pow_rel(M,\)|\<^bsup>M\<^esup> \ \" using cardinal_rel_Pow_rel[of \] by simp with assms have "Pow_rel(M,\) \\<^bsup>M\<^esup> \" using cardinal_rel_eqpoll_rel_iff Card_rel_le_imp_lepoll_rel Card_rel_cardinal_rel_eq by auto then obtain g where "g \ inj_rel(M,Pow_rel(M,\), \)" by blast moreover note \M(\)\ moreover from calculation have "M(g)" by (auto dest:transM) ultimately show "False" using cantor_inj_rel by simp qed simp lemma countable_iff_lesspoll_rel_Aleph_rel_one: notes iff_trans[trans] assumes "M(C)" shows "countable\<^bsup>M\<^esup>(C) \ C \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" using assms lesspoll_rel_csucc_rel[of \ C] Aleph_rel_succ Aleph_rel_zero unfolding countable_rel_def by simp lemma countable_iff_le_rel_Aleph_rel_one: notes iff_trans[trans] assumes "M(C)" shows "countable\<^bsup>M\<^esup>(C) \ |C|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" proof - from assms have "countable\<^bsup>M\<^esup>(C) \ C \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" using countable_iff_lesspoll_rel_Aleph_rel_one by simp also from assms have "\ \ |C|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" using cardinal_rel_eqpoll_rel[THEN eqpoll_rel_sym, THEN eq_lesspoll_rel_trans] by (auto intro:cardinal_rel_eqpoll_rel[THEN eq_lesspoll_rel_trans]) finally show ?thesis . qed end \ \\<^locale>\M_library\\ (* TODO: This can be generalized. *) lemma (in M_cardinal_library) countable_fun_imp_countable_image: assumes "f:C \\<^bsup>M\<^esup> B" "countable\<^bsup>M\<^esup>(C)" "\c. c\C \ countable\<^bsup>M\<^esup>(f`c)" "M(C)" "M(B)" shows "countable\<^bsup>M\<^esup>(\(f``C))" using assms function_space_rel_char image_fun[of f] cardinal_rel_RepFun_apply_le[of f C B] countable_rel_iff_cardinal_rel_le_nat[THEN iffD1, THEN [2] le_trans, of _ ] countable_rel_iff_cardinal_rel_le_nat by (rule_tac countable_rel_union_countable_rel) (auto dest:transM del:imageE) end \ No newline at end of file diff --git a/thys/Transitive_Models/Cardinal_Relative.thy b/thys/Transitive_Models/Cardinal_Relative.thy --- a/thys/Transitive_Models/Cardinal_Relative.thy +++ b/thys/Transitive_Models/Cardinal_Relative.thy @@ -1,1331 +1,1487 @@ section\Relative, Choice-less Cardinal Numbers\ theory Cardinal_Relative imports Lambda_Replacement Univ_Relative begin txt\The following command avoids that a commonly used one-letter variable be captured by the definition of the constructible universe \<^term>\L\.\ hide_const (open) L txt\We also return to the old notation for \<^term>\sum\ to preserve the old Constructibility code.\ no_notation oadd (infixl \+\ 65) notation sum (infixr \+\ 65) definition Finite_rel :: "[i\o,i]=>o" where "Finite_rel(M,A) \ \om[M]. \n[M]. omega(M,om) \ n\om \ eqpoll_rel(M,A,n)" definition banach_functor :: "[i,i,i,i,i] \ i" where "banach_functor(X,Y,f,g,W) \ X - g``(Y - f``W)" lemma banach_functor_subset: "banach_functor(X,Y,f,g,W) \ X" unfolding banach_functor_def by auto definition is_banach_functor :: "[i\o,i,i,i,i,i,i]\o" where "is_banach_functor(M,X,Y,f,g,W,b) \ \fW[M]. \YfW[M]. \gYfW[M]. image(M,f,W,fW) \ setdiff(M,Y,fW,YfW) \ image(M,g,YfW,gYfW) \ setdiff(M,X,gYfW,b)" lemma (in M_basic) banach_functor_abs : assumes "M(X)" "M(Y)" "M(f)" "M(g)" shows "relation1(M,is_banach_functor(M,X,Y,f,g),banach_functor(X,Y,f,g))" unfolding relation1_def is_banach_functor_def banach_functor_def using assms by simp lemma (in M_basic) banach_functor_closed: assumes "M(X)" "M(Y)" "M(f)" "M(g)" shows "\W[M]. M(banach_functor(X,Y,f,g,W))" unfolding banach_functor_def using assms image_closed by simp context M_trancl begin lemma iterates_banach_functor_closed: assumes "n\\" "M(X)" "M(Y)" "M(f)" "M(g)" shows "M(banach_functor(X,Y,f,g)^n(0))" using assms banach_functor_closed by (induct) simp_all lemma banach_repl_iter': assumes "\A. M(A) \ separation(M, \b. \x\A. x \ \ \ b = banach_functor(X, Y, f, g)^x (0))" "M(X)" "M(Y)" "M(f)" "M(g)" shows "strong_replacement(M, \x y. x\nat \ y = banach_functor(X, Y, f, g)^x (0))" unfolding strong_replacement_def proof clarsimp fix A let ?Z="{b \ Pow\<^bsup>M\<^esup>(X) . (\x\A. x \ \ \ b = banach_functor(X, Y, f, g)^x (0))}" assume "M(A)" moreover note assms moreover from calculation have "M(?Z)" by simp moreover from assms(2-) have "b \ ?Z \ (\x\A. x \ \ \ b = banach_functor(X, Y, f, g)^x (0))" for b by (auto, rename_tac x, induct_tac x, (auto simp:def_Pow_rel)[1], rule_tac def_Pow_rel[THEN iffD2, OF iterates_banach_functor_closed[of _ X Y f g]]) (simp_all add:banach_functor_subset) ultimately show "\Z[M]. \b[M]. b \ Z \ (\x\A. x \ \ \ b = banach_functor(X, Y, f, g)^x (0))" by (intro rexI[of _ ?Z]) simp qed end \ \\<^locale>\M_trancl\\ +context M_Perm +begin + +lemma mem_Pow_rel: "M(r) \ a \ Pow_rel(M,r) \ a \ Pow(r) \ M(a)" + using Pow_rel_char by simp + +lemma mem_bij_abs: "\M(f);M(A);M(B)\ \ f \ bij\<^bsup>M\<^esup>(A,B) \ f\bij(A,B)" + using bij_rel_char by simp + +lemma mem_inj_abs: "\M(f);M(A);M(B)\ \ f \ inj\<^bsup>M\<^esup>(A,B) \ f\inj(A,B)" + using inj_rel_char by simp + +lemma mem_surj_abs: "\M(f);M(A);M(B)\ \ f \ surj\<^bsup>M\<^esup>(A,B) \ f\surj(A,B)" + using surj_rel_char by simp + +end \ \\<^locale>\M_Perm\\ + locale M_cardinals = M_ordertype + M_trancl + M_Perm + M_replacement + assumes radd_separation: "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))" and rmult_separation: "M(b) \ M(d) \ separation(M, \z. \x' y' x y. z = \\x', y'\, x, y\ \ (\x', x\ \ b \ x' = x \ \y', y\ \ d))" - and - banach_repl_iter: "M(X) \ M(Y) \ M(f) \ M(g) \ - strong_replacement(M, \x y. x\nat \ y = banach_functor(X, Y, f, g)^x (0))" + begin lemma rvimage_separation: "M(f) \ M(r) \ separation(M, \z. \x y. z = \x, y\ \ \f ` x, f ` y\ \ r)" - using separation_Pair separation_in - lam_replacement_Pair[THEN[5] lam_replacement_hcomp2] - lam_replacement_constant lam_replacement_apply2[THEN[5] lam_replacement_hcomp2,OF lam_replacement_constant[of f]] - lam_replacement_fst lam_replacement_snd - lam_replacement_identity lam_replacement_hcomp - by(simp_all) + using separation_Pair separation_in lam_replacement_product + lam_replacement_fst lam_replacement_snd lam_replacement_constant + lam_replacement_apply2[THEN [5] lam_replacement_hcomp2] + by simp lemma radd_closed[intro,simp]: "M(a) \ M(b) \ M(c) \ M(d) \ M(radd(a,b,c,d))" using radd_separation by (auto simp add: radd_def) lemma rmult_closed[intro,simp]: "M(a) \ M(b) \ M(c) \ M(d) \ M(rmult(a,b,c,d))" using rmult_separation by (auto simp add: rmult_def) end \ \\<^locale>\M_cardinals\\ lemma (in M_cardinals) is_cardinal_iff_Least: assumes "M(A)" "M(\)" shows "is_cardinal(M,A,\) \ \ = (\ i. M(i) \ i \\<^bsup>M\<^esup> A)" using is_cardinal_iff assms unfolding cardinal_rel_def by simp subsection\The Schroeder-Bernstein Theorem\ text\See Davey and Priestly, page 106\ context M_cardinals begin -(** Lemma: Banach's Decomposition Theorem **) +subsection\Banach's Decomposition Theorem\ lemma bnd_mono_banach_functor: "bnd_mono(X, banach_functor(X,Y,f,g))" unfolding bnd_mono_def banach_functor_def by blast lemma inj_Inter: assumes "g \ inj(Y,X)" "A\0" "\a\A. a \ Y" shows "g``(\A) = (\a\A. g``a)" proof (intro equalityI subsetI) fix x from assms obtain a where "a\A" by blast moreover assume "x \ (\a\A. g `` a)" ultimately have x_in_im: "x \ g``y" if "y\A" for y using that by auto have exists: "\z \ y. x = g`z" if "y\A" for y proof - note that moreover from this and x_in_im have "x \ g``y" by simp moreover from calculation have "x \ g``y" by simp moreover note assms ultimately show ?thesis using image_fun[OF inj_is_fun] by auto qed with \a\A\ obtain z where "z \ a" "x = g`z" by auto moreover have "z \ y" if "y\A" for y proof - from that and exists obtain w where "w \ y" "x = g`w" by auto moreover from this \x = g`z\ assms that \a\A\ \z\a\ have "z = w" unfolding inj_def by blast ultimately show ?thesis by simp qed moreover note assms moreover from calculation have "z \ \A" by auto moreover from calculation have "z \ Y" by blast ultimately show "x \ g `` (\A)" using inj_is_fun[THEN funcI, of g] by fast qed auto lemma contin_banach_functor: assumes "g \ inj(Y,X)" shows "contin(banach_functor(X,Y,f,g))" unfolding contin_def proof (intro allI impI) fix A assume "directed(A)" then have "A \ 0" unfolding directed_def .. have "banach_functor(X, Y, f, g, \A) = X - g``(Y - f``(\A))" unfolding banach_functor_def .. also have " \ = X - g``(Y - (\a\A. f``a))" by auto also from \A\0\ have " \ = X - g``(\a\A. Y-f``a)" by auto also from \A\0\ and assms have " \ = X - (\a\A. g``(Y-f``a))" using inj_Inter[of g Y X "{Y-f``a. a\A}" ] by fastforce also from \A\0\ have " \ = (\a\A. X - g``(Y-f``a))" by simp also have " \ = (\a\A. banach_functor(X, Y, f, g, a))" unfolding banach_functor_def .. finally show "banach_functor(X,Y,f,g,\A) = (\a\A. banach_functor(X,Y,f,g,a))" . qed lemma lfp_banach_functor: assumes "g\inj(Y,X)" shows "lfp(X, banach_functor(X,Y,f,g)) = (\n\nat. banach_functor(X,Y,f,g)^n (0))" using assms lfp_eq_Union bnd_mono_banach_functor contin_banach_functor by simp lemma lfp_banach_functor_closed: assumes "M(g)" "M(X)" "M(Y)" "M(f)" "g\inj(Y,X)" + and banach_repl_iter: "M(X) \ M(Y) \ M(f) \ M(g) \ + strong_replacement(M, \x y. x\nat \ y = banach_functor(X, Y, f, g)^x (0))" shows "M(lfp(X, banach_functor(X,Y,f,g)))" proof - from assms have "M(banach_functor(X,Y,f,g)^n (0))" if "n\nat" for n by(rule_tac nat_induct[OF that],simp_all add:banach_functor_closed) with assms show ?thesis using family_union_closed'[OF banach_repl_iter M_nat] lfp_banach_functor by simp qed lemma banach_decomposition_rel: - "[| M(f); M(g); M(X); M(Y); f \ X->Y; g \ inj(Y,X) |] ==> + assumes + banach_repl_iter: "M(f) \ M(g) \ M(X) \ M(Y) \ + strong_replacement(M, \x y. x\nat \ y = banach_functor(X, Y, f, g)^x (0))" + shows + "[| M(f); M(g); M(X); M(Y); f \ X->Y; g \ inj(Y,X) |] ==> \XA[M]. \XB[M]. \YA[M]. \YB[M]. (XA \ XB = 0) & (XA \ XB = X) & (YA \ YB = 0) & (YA \ YB = Y) & f``XA=YA & g``YB=XB" apply (intro rexI conjI) apply (rule_tac [6] Banach_last_equation) apply (rule_tac [5] refl) apply (assumption | rule inj_is_fun Diff_disjoint Diff_partition fun_is_rel image_subset lfp_subset)+ - using lfp_banach_functor_closed[of g X Y f] + using lfp_banach_functor_closed[of g X Y f] assms unfolding banach_functor_def by simp_all lemma schroeder_bernstein_closed: + assumes + banach_repl_iter: "M(X) \ M(Y) \ M(f) \ M(g) \ + strong_replacement(M, \x y. x\nat \ y = banach_functor(X, Y, f, g)^x (0))" + shows "[| M(f); M(g); M(X); M(Y); f \ inj(X,Y); g \ inj(Y,X) |] ==> \h[M]. h \ bij(X,Y)" - apply (insert banach_decomposition_rel [of f g X Y]) + apply (insert banach_decomposition_rel [of f g X Y, OF banach_repl_iter]) apply (simp add: inj_is_fun) apply (auto) apply (rule_tac x="restrict(f,XA) \ converse(restrict(g,YB))" in rexI) apply (auto intro!: restrict_bij bij_disjoint_Un intro: bij_converse_bij) done -(** Equipollence is an equivalence relation **) - -lemma mem_Pow_rel: "M(r) \ a \ Pow_rel(M,r) \ a \ Pow(r) \ M(a)" - using Pow_rel_char by simp +text\The previous lemmas finish our original, direct relativization +of the material involving the iterative proof (as appearing in \<^theory>\ZF.Cardinal\) +of the Schröder-Bernstein theorem. Next, we formalize +Zermelo's proof that replaces the recursive construction by a fixed point +represented as an intersection \cite[Exr. x4.27]{moschovakis1994notes}. This allows +to avoid at least one replacement assumption.\ -lemma mem_bij_abs[simp]: "\M(f);M(A);M(B)\ \ f \ bij\<^bsup>M\<^esup>(A,B) \ f\bij(A,B)" - using bij_rel_char by simp +lemma dedekind_zermelo: + assumes + "A' \ B" "B \ A" "A \\<^bsup>M\<^esup> A'" + and types: "M(A')" "M(B)" "M(A)" + shows + "A \\<^bsup>M\<^esup> B" +proof - + from \A \\<^bsup>M\<^esup> A'\ and types + obtain f where "f \ bij\<^bsup>M\<^esup>(A,A')" + unfolding eqpoll_rel_def by blast + let ?Q="B - f``A" + from \f \ _\ \A' \ B\ types + have "f``X \ B" for X + using mem_bij_rel[THEN bij_is_fun, THEN Image_sub_codomain, of f A A'] by auto + moreover + note \B \ A\ + moreover from calculation + have "f``X \ A" for X by auto + moreover + define \ where "\ \ { X \ Pow\<^bsup>M\<^esup>(A) . ?Q \ f `` X \ X}" + moreover from calculation + have "A \ \" using def_Pow_rel by (auto simp:types) + ultimately + have "\ \ 0" unfolding \_def by auto + let ?T="\\" + have "?T \ A" + proof + fix x + assume "x \ ?T" + with \\ \ 0\ + obtain X where "x \ X" "X \ \" + by blast + moreover from this + have "X \ Pow\<^bsup>M\<^esup>(A)" + unfolding \_def by simp + moreover from calculation and \M(A)\ + have "M(x)" using Pow_rel_char by (auto dest:transM) + ultimately + show "x \ A" using Pow_rel_char by (auto simp:types) + qed + moreover from \\ \ 0\ + have "?Q \ f``?T \ ?T" + using image_mono unfolding \_def by blast + moreover from \f \ _\ + have "M(\)" + using zermelo_separation unfolding \_def by (auto simp:types dest:transM) + moreover from this + have "M(?T)" by simp + moreover + note \\ \ 0\ + ultimately + have "?T \ \" + unfolding \_def using def_Pow_rel + by (auto simp:types) + have "?T \ ?Q \ f``?T" + proof (intro subsetI, rule ccontr) + fix x + from \_ \ f ``?T \ ?T\ + have "f `` (?T - {x}) \ ?T" "f `` (?T - {x}) \ f`` ?T" by auto + assume "x \ ?T" "x \ ?Q \ f``?T" + then + have "?T - {x} \ \" + proof - + note \f `` (?T - {x}) \ ?T\ + moreover from this and \x \ _ \ f``?T\ + have "x \ f `` ?T" by simp + ultimately + have "f `` (?T - {x}) \ ?T - {x}" by auto + moreover from \x \ ?Q \ _\ \?Q \ _ \ ?T\ + have "?Q \ ?T - {x}" by auto + moreover + note \?T \ _\ \x \ ?T\ \M(?T)\ \?T \ _\ + ultimately + show ?thesis + using def_Pow_rel[of "?T - {x}" A] transM[of _ ?T] + unfolding \_def + by (auto simp:types) + qed + moreover + note \f `` (?T - {x}) \ f`` ?T\ + ultimately + have "?T \ ?T - {x}" by auto + with \x \ ?T\ + show False by auto + qed + with \?Q \ f``?T \ ?T\ + have "?T = ?Q \ f``?T" by auto + from \\X. f``X \ B\ + have "(?Q \ f``?T) \ (f``A - f``?T) \ B" by auto + with \?T = _\ + have "?T \ (f``A - f``?T) \ B" by simp + with \\ \ 0\ \?T = _\ + have "B = ?T \ (f``A - f``?T)" + proof (intro equalityI, intro subsetI) + fix x + assume "x \ B" + with \\ \ 0\ \?T = _\ + show "x \ ?T \ (f``A - f``?T)" + by (subst \?T = _\, clarify) + qed + moreover from \?T \ A\ + have "A = ?T \ (A - ?T)" by auto + moreover from \M(?T)\ \f \ _\ + have "M(id(?T) \ restrict(f,A-?T))" + using bij_rel_closed[THEN [2] transM] id_closed by (auto simp:types) + moreover from \?T = _\ \f \ _\ + have "f``A - f``?T = f``(A - ?T)" + using bij_rel_char bij_is_inj[THEN inj_range_Diff, of f A] + by (auto simp:types) + from \f \ _\ types + have "id(?T) \ restrict(f, A-?T) \ bij(?T \ (A - ?T),?T \ (f``A - f``?T))" + using id_bij mem_bij_rel[of _ A A', THEN bij_is_inj] + by (rule_tac bij_disjoint_Un; clarsimp) + (subst \f``A - f``?T =_ \, auto intro:restrict_bij, subst \?T = _\, auto) + moreover + note types + ultimately + show ?thesis + using bij_rel_char unfolding eqpoll_rel_def by fastforce +qed -lemma mem_inj_abs[simp]: "\M(f);M(A);M(B)\ \ f \ inj\<^bsup>M\<^esup>(A,B) \ f\inj(A,B)" - using inj_rel_char by simp +lemma schroeder_bernstein_closed': + assumes "f \ inj\<^bsup>M\<^esup>(A,C)" " g \ inj\<^bsup>M\<^esup>(C,A)" + and types:"M(A)" "M(C)" + shows "A \\<^bsup>M\<^esup> C" +proof - + from assms + have "f \ inj(A,C)" " g \ inj(C,A)" "M(f)" "M(g)" + using inj_rel_char by auto + moreover from this + have "g \ bij(C,range(g))" "g O f \ bij(A,range(g O f))" + using inj_bij_range comp_inj by auto blast + moreover from calculation + have "range(g O f) \ range(g)" "range(g) \ A" + using inj_is_fun[THEN range_fun_subset_codomain] by auto + moreover from calculation + obtain h where "h \ bij\<^bsup>M\<^esup>(A, range(g))" + using dedekind_zermelo[of "range(g O f)" "range(g)" A] + bij_rel_char def_eqpoll_rel + by (auto simp:types) + ultimately + show ?thesis + using bij_rel_char def_eqpoll_rel comp_bij[of h A "range(g)" "converse(g)" C] + bij_converse_bij[of g C "range(g)"] by (auto simp:types) +qed -lemma mem_surj_abs: "\M(f);M(A);M(B)\ \ f \ surj\<^bsup>M\<^esup>(A,B) \ f\surj(A,B)" - using surj_rel_char by simp +text\Relative equipollence is an equivalence relation\ + +declare mem_bij_abs[simp] mem_inj_abs[simp] lemma bij_imp_eqpoll_rel: assumes "f \ bij(A,B)" "M(f)" "M(A)" "M(B)" shows "A \\<^bsup>M\<^esup> B" using assms by (auto simp add:def_eqpoll_rel) lemma eqpoll_rel_refl: "M(A) \ A \\<^bsup>M\<^esup> A" using bij_imp_eqpoll_rel[OF id_bij, OF id_closed] . lemma eqpoll_rel_sym: "X \\<^bsup>M\<^esup> Y \ M(X) \ M(Y) \ Y \\<^bsup>M\<^esup> X" unfolding def_eqpoll_rel using converse_closed by (auto intro: bij_converse_bij) lemma eqpoll_rel_trans [trans]: "[|X \\<^bsup>M\<^esup> Y; Y \\<^bsup>M\<^esup> Z; M(X); M(Y) ; M(Z) |] ==> X \\<^bsup>M\<^esup> Z" unfolding def_eqpoll_rel by (auto intro: comp_bij) -(** Le-pollence is a partial ordering **) +text\Relative le-pollence is a preorder\ lemma subset_imp_lepoll_rel: "X \ Y \ M(X) \ M(Y) \ X \\<^bsup>M\<^esup> Y" unfolding def_lepoll_rel using id_subset_inj id_closed by simp blast lemmas lepoll_rel_refl = subset_refl [THEN subset_imp_lepoll_rel, simp] lemmas le_imp_lepoll_rel = le_imp_subset [THEN subset_imp_lepoll_rel] lemma eqpoll_rel_imp_lepoll_rel: "X \\<^bsup>M\<^esup> Y ==> M(X) \ M(Y) \ X \\<^bsup>M\<^esup> Y" unfolding def_eqpoll_rel bij_def def_lepoll_rel using bij_is_inj by (auto) lemma lepoll_rel_trans [trans]: assumes "X \\<^bsup>M\<^esup> Y" "Y \\<^bsup>M\<^esup> Z" "M(X)" "M(Y)" "M(Z)" shows "X \\<^bsup>M\<^esup> Z" using assms def_lepoll_rel by (auto intro: comp_inj) lemma eq_lepoll_rel_trans [trans]: assumes "X \\<^bsup>M\<^esup> Y" "Y \\<^bsup>M\<^esup> Z" "M(X)" "M(Y)" "M(Z)" shows "X \\<^bsup>M\<^esup> Z" using assms by (blast intro: eqpoll_rel_imp_lepoll_rel lepoll_rel_trans) lemma lepoll_rel_eq_trans [trans]: assumes "X \\<^bsup>M\<^esup> Y" "Y \\<^bsup>M\<^esup> Z" "M(X)" "M(Y)" "M(Z)" shows "X \\<^bsup>M\<^esup> Z" using assms eqpoll_rel_imp_lepoll_rel[of Y Z] lepoll_rel_trans[of X Y Z] by simp lemma eqpoll_relI: "\ X \\<^bsup>M\<^esup> Y; Y \\<^bsup>M\<^esup> X; M(X) ; M(Y) \ \ X \\<^bsup>M\<^esup> Y" - unfolding def_lepoll_rel def_eqpoll_rel using schroeder_bernstein_closed - by auto + unfolding def_lepoll_rel using schroeder_bernstein_closed' + by (auto simp del:mem_inj_abs) lemma eqpoll_relE: "[| X \\<^bsup>M\<^esup> Y; [| X \\<^bsup>M\<^esup> Y; Y \\<^bsup>M\<^esup> X |] ==> P ; M(X) ; M(Y) |] ==> P" by (blast intro: eqpoll_rel_imp_lepoll_rel eqpoll_rel_sym) lemma eqpoll_rel_iff: "M(X) \ M(Y) \ X \\<^bsup>M\<^esup> Y \ X \\<^bsup>M\<^esup> Y & Y \\<^bsup>M\<^esup> X" by (blast intro: eqpoll_relI elim: eqpoll_relE) lemma lepoll_rel_0_is_0: "A \\<^bsup>M\<^esup> 0 \ M(A) \ A = 0" using def_lepoll_rel by (cases "A=0") (auto simp add: inj_def) -(* \<^term>\M(Y) \ 0 \\<^bsup>M\<^esup> Y\ *) +\ \\<^term>\M(Y) \ 0 \\<^bsup>M\<^esup> Y\\ lemmas empty_lepoll_relI = empty_subsetI [THEN subset_imp_lepoll_rel, OF nonempty] lemma lepoll_rel_0_iff: "M(A) \ A \\<^bsup>M\<^esup> 0 \ A=0" by (blast intro: lepoll_rel_0_is_0 lepoll_rel_refl) lemma Un_lepoll_rel_Un: "[| A \\<^bsup>M\<^esup> B; C \\<^bsup>M\<^esup> D; B \ D = 0; M(A); M(B); M(C); M(D) |] ==> A \ C \\<^bsup>M\<^esup> B \ D" using def_lepoll_rel using inj_disjoint_Un[of _ A B _ C D] if_then_replacement apply (auto) apply (rule, assumption) apply (auto intro!:lam_closed elim:transM)+ done lemma eqpoll_rel_0_is_0: "A \\<^bsup>M\<^esup> 0 \ M(A) \ A = 0" using eqpoll_rel_imp_lepoll_rel lepoll_rel_0_is_0 nonempty by blast lemma eqpoll_rel_0_iff: "M(A) \ A \\<^bsup>M\<^esup> 0 \ A=0" by (blast intro: eqpoll_rel_0_is_0 eqpoll_rel_refl) lemma eqpoll_rel_disjoint_Un: "[| A \\<^bsup>M\<^esup> B; C \\<^bsup>M\<^esup> D; A \ C = 0; B \ D = 0; M(A); M(B); M(C) ; M(D) |] ==> A \ C \\<^bsup>M\<^esup> B \ D" by (auto intro: bij_disjoint_Un simp add:def_eqpoll_rel) -subsection\lesspoll\_rel: contributions by Krzysztof Grabczewski\ +subsection\lesspoll: contributions by Krzysztof Grabczewski\ lemma lesspoll_rel_not_refl: "M(i) \ ~ (i \\<^bsup>M\<^esup> i)" by (simp add: lesspoll_rel_def eqpoll_rel_refl) lemma lesspoll_rel_irrefl: "i \\<^bsup>M\<^esup> i ==> M(i) \ P" by (simp add: lesspoll_rel_def eqpoll_rel_refl) lemma lesspoll_rel_imp_lepoll_rel: "\A \\<^bsup>M\<^esup> B; M(A); M(B)\\ A \\<^bsup>M\<^esup> B" by (unfold lesspoll_rel_def, blast) lemma rvimage_closed [intro,simp]: assumes "M(A)" "M(f)" "M(r)" shows "M(rvimage(A,f,r))" unfolding rvimage_def using assms rvimage_separation by auto lemma lepoll_rel_well_ord: "[| A \\<^bsup>M\<^esup> B; well_ord(B,r); M(A); M(B); M(r) |] ==> \s[M]. well_ord(A,s)" unfolding def_lepoll_rel by (auto intro:well_ord_rvimage) lemma lepoll_rel_iff_leqpoll_rel: "\M(A); M(B)\ \ A \\<^bsup>M\<^esup> B \ A \\<^bsup>M\<^esup> B | A \\<^bsup>M\<^esup> B" apply (unfold lesspoll_rel_def) apply (blast intro: eqpoll_relI elim: eqpoll_relE) done end \ \\<^locale>\M_cardinals\\ context M_cardinals begin lemma inj_rel_is_fun_M: "f \ inj\<^bsup>M\<^esup>(A,B) \ M(f) \ M(A) \ M(B) \ f \ A \\<^bsup>M\<^esup> B" using inj_is_fun function_space_rel_char by simp \ \In porting the following theorem, I tried to follow the Discipline strictly, though finally only an approach maximizing the use of absoluteness results (@{thm function_space_rel_char inj_rel_char}) was the one paying dividends.\ lemma inj_rel_not_surj_rel_succ: notes mem_inj_abs[simp del] assumes fi: "f \ inj\<^bsup>M\<^esup>(A, succ(m))" and fns: "f \ surj\<^bsup>M\<^esup>(A, succ(m))" and types: "M(f)" "M(A)" "M(m)" shows "\f[M]. f \ inj\<^bsup>M\<^esup>(A,m)" proof - from fi [THEN inj_rel_is_fun_M] fns types obtain y where y: "y \ succ(m)" "\x. x\A \ f ` x \ y" "M(y)" by (auto simp add: def_surj_rel) show ?thesis proof from types and \M(y)\ show "M(\z\A. if f ` z = m then y else f ` z)" using transM[OF _ \M(A)\] lam_if_then_apply_replacement2 lam_replacement_iff_lam_closed by (auto) with types y fi have "(\z\A. if f`z = m then y else f`z) \ A\\<^bsup>M\<^esup> m" using function_space_rel_char inj_rel_char inj_is_fun[of f A "succ(m)"] by (auto intro!: if_type [THEN lam_type] dest: apply_funtype) with types y fi show "(\z\A. if f`z = m then y else f`z) \ inj\<^bsup>M\<^esup>(A, m)" by (simp add: def_inj_rel) blast qed qed -(** Variations on transitivity **) +\ \Variations on transitivity\ lemma lesspoll_rel_trans [trans]: "[| X \\<^bsup>M\<^esup> Y; Y \\<^bsup>M\<^esup> Z; M(X); M(Y) ; M(Z) |] ==> X \\<^bsup>M\<^esup> Z" apply (unfold lesspoll_rel_def) apply (blast elim: eqpoll_relE intro: eqpoll_relI lepoll_rel_trans) done lemma lesspoll_rel_trans1 [trans]: "[| X \\<^bsup>M\<^esup> Y; Y \\<^bsup>M\<^esup> Z; M(X); M(Y) ; M(Z) |] ==> X \\<^bsup>M\<^esup> Z" apply (unfold lesspoll_rel_def) apply (blast elim: eqpoll_relE intro: eqpoll_relI lepoll_rel_trans) done lemma lesspoll_rel_trans2 [trans]: "[| X \\<^bsup>M\<^esup> Y; Y \\<^bsup>M\<^esup> Z; M(X); M(Y) ; M(Z)|] ==> X \\<^bsup>M\<^esup> Z" apply (unfold lesspoll_rel_def) apply (blast elim: eqpoll_relE intro: eqpoll_relI lepoll_rel_trans) done lemma eq_lesspoll_rel_trans [trans]: "[| X \\<^bsup>M\<^esup> Y; Y \\<^bsup>M\<^esup> Z; M(X); M(Y) ; M(Z) |] ==> X \\<^bsup>M\<^esup> Z" by (blast intro: eqpoll_rel_imp_lepoll_rel lesspoll_rel_trans1) lemma lesspoll_rel_eq_trans [trans]: "[| X \\<^bsup>M\<^esup> Y; Y \\<^bsup>M\<^esup> Z; M(X); M(Y) ; M(Z) |] ==> X \\<^bsup>M\<^esup> Z" by (blast intro: eqpoll_rel_imp_lepoll_rel lesspoll_rel_trans2) lemma is_cardinal_cong: assumes "X \\<^bsup>M\<^esup> Y" "M(X)" "M(Y)" shows "\\[M]. is_cardinal(M,X,\) \ is_cardinal(M,Y,\)" proof - from assms have "(\ i. M(i) \ i \\<^bsup>M\<^esup> X) = (\ i. M(i) \ i \\<^bsup>M\<^esup> Y)" by (intro Least_cong) (auto intro: comp_bij bij_converse_bij simp add:def_eqpoll_rel) moreover from assms have "M(\ i. M(i) \ i \\<^bsup>M\<^esup> X)" using Least_closed' by fastforce moreover note assms ultimately show ?thesis using is_cardinal_iff_Least by auto qed \ \ported from Cardinal\ lemma cardinal_rel_cong: "X \\<^bsup>M\<^esup> Y \ M(X) \ M(Y) \ |X|\<^bsup>M\<^esup> = |Y|\<^bsup>M\<^esup>" apply (simp add: def_eqpoll_rel cardinal_rel_def) apply (rule Least_cong) apply (auto intro: comp_bij bij_converse_bij) done lemma well_ord_is_cardinal_eqpoll_rel: assumes "well_ord(A,r)" shows "is_cardinal(M,A,\) \ M(A) \ M(\) \ M(r) \ \ \\<^bsup>M\<^esup> A" proof (subst is_cardinal_iff_Least[THEN iffD1, of A \]) assume "M(A)" "M(\)" "M(r)" "is_cardinal(M,A,\)" moreover from assms and calculation obtain f i where "M(f)" "Ord(i)" "M(i)" "f \ bij(A,i)" using ordertype_exists[of A r] ord_iso_is_bij by auto moreover have "M(\ i. M(i) \ i \\<^bsup>M\<^esup> A)" using Least_closed' by fastforce ultimately show "(\ i. M(i) \ i \\<^bsup>M\<^esup> A) \\<^bsup>M\<^esup> A" using assms[THEN well_ord_imp_relativized] LeastI[of "\i. M(i) \ i \\<^bsup>M\<^esup> A" i] Ord_ordertype[OF assms] bij_converse_bij[THEN bij_imp_eqpoll_rel, of f] by simp qed lemmas Ord_is_cardinal_eqpoll_rel = well_ord_Memrel[THEN well_ord_is_cardinal_eqpoll_rel] - -(**********************************************************************) -(****************** Results imported from Cardinal.thy ****************) -(**********************************************************************) - section\Porting from \<^theory>\ZF.Cardinal\\ txt\The following results were ported more or less directly from \<^theory>\ZF.Cardinal\\ \ \This result relies on various closure properties and thus cannot be translated directly\ lemma well_ord_cardinal_rel_eqpoll_rel: assumes r: "well_ord(A,r)" and "M(A)" "M(r)" shows "|A|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> A" using assms well_ord_is_cardinal_eqpoll_rel is_cardinal_iff by blast lemmas Ord_cardinal_rel_eqpoll_rel = well_ord_Memrel[THEN well_ord_cardinal_rel_eqpoll_rel] lemma Ord_cardinal_rel_idem: "Ord(A) \ M(A) \ ||A|\<^bsup>M\<^esup>|\<^bsup>M\<^esup> = |A|\<^bsup>M\<^esup>" by (rule_tac Ord_cardinal_rel_eqpoll_rel [THEN cardinal_rel_cong]) auto lemma well_ord_cardinal_rel_eqE: assumes woX: "well_ord(X,r)" and woY: "well_ord(Y,s)" and eq: "|X|\<^bsup>M\<^esup> = |Y|\<^bsup>M\<^esup>" and types: "M(X)" "M(r)" "M(Y)" "M(s)" shows "X \\<^bsup>M\<^esup> Y" proof - from types have "X \\<^bsup>M\<^esup> |X|\<^bsup>M\<^esup>" by (blast intro: well_ord_cardinal_rel_eqpoll_rel [OF woX] eqpoll_rel_sym) also have "... = |Y|\<^bsup>M\<^esup>" by (rule eq) also from types have "... \\<^bsup>M\<^esup> Y" by (rule_tac well_ord_cardinal_rel_eqpoll_rel [OF woY]) finally show ?thesis by (simp add:types) qed lemma well_ord_cardinal_rel_eqpoll_rel_iff: "[| well_ord(X,r); well_ord(Y,s); M(X); M(r); M(Y); M(s) |] ==> |X|\<^bsup>M\<^esup> = |Y|\<^bsup>M\<^esup> \ X \\<^bsup>M\<^esup> Y" by (blast intro: cardinal_rel_cong well_ord_cardinal_rel_eqE) lemma Ord_cardinal_rel_le: "Ord(i) \ M(i) ==> |i|\<^bsup>M\<^esup> \ i" unfolding cardinal_rel_def using eqpoll_rel_refl Least_le by simp lemma Card_rel_cardinal_rel_eq: "Card\<^bsup>M\<^esup>(K) ==> M(K) \ |K|\<^bsup>M\<^esup> = K" apply (unfold Card_rel_def) apply (erule sym) done lemma Card_relI: "[| Ord(i); !!j. j M(j) ==> ~(j \\<^bsup>M\<^esup> i); M(i) |] ==> Card\<^bsup>M\<^esup>(i)" apply (unfold Card_rel_def cardinal_rel_def) apply (subst Least_equality) apply (blast intro: eqpoll_rel_refl)+ done lemma Card_rel_is_Ord: "Card\<^bsup>M\<^esup>(i) ==> M(i) \ Ord(i)" apply (unfold Card_rel_def cardinal_rel_def) apply (erule ssubst) apply (rule Ord_Least) done lemma Card_rel_cardinal_rel_le: "Card\<^bsup>M\<^esup>(K) ==> M(K) \ K \ |K|\<^bsup>M\<^esup>" apply (simp (no_asm_simp) add: Card_rel_is_Ord Card_rel_cardinal_rel_eq) done lemma Ord_cardinal_rel [simp,intro!]: "M(A) \ Ord(|A|\<^bsup>M\<^esup>)" apply (unfold cardinal_rel_def) apply (rule Ord_Least) done lemma Card_rel_iff_initial: assumes types:"M(K)" shows "Card\<^bsup>M\<^esup>(K) \ Ord(K) & (\j[M]. j ~ (j \\<^bsup>M\<^esup> K))" proof - { fix j assume K: "Card\<^bsup>M\<^esup>(K)" "M(j) \ j \\<^bsup>M\<^esup> K" assume "j < K" also have "... = (\ i. M(i) \ i \\<^bsup>M\<^esup> K)" using K by (simp add: Card_rel_def cardinal_rel_def types) finally have "j < (\ i. M(i) \ i \\<^bsup>M\<^esup> K)" . then have "False" using K by (best intro: less_LeastE[of "\j. M(j) \ j \\<^bsup>M\<^esup> K"]) } with types show ?thesis by (blast intro: Card_relI Card_rel_is_Ord) qed lemma lt_Card_rel_imp_lesspoll_rel: "[| Card\<^bsup>M\<^esup>(a); i i \\<^bsup>M\<^esup> a" apply (unfold lesspoll_rel_def) apply (frule Card_rel_iff_initial [THEN iffD1], assumption) apply (blast intro!: leI [THEN le_imp_lepoll_rel]) done lemma Card_rel_0: "Card\<^bsup>M\<^esup>(0)" apply (rule Ord_0 [THEN Card_relI]) apply (auto elim!: ltE) done lemma Card_rel_Un: "[| Card\<^bsup>M\<^esup>(K); Card\<^bsup>M\<^esup>(L); M(K); M(L) |] ==> Card\<^bsup>M\<^esup>(K \ L)" apply (rule Ord_linear_le [of K L]) apply (simp_all add: subset_Un_iff [THEN iffD1] Card_rel_is_Ord le_imp_subset subset_Un_iff2 [THEN iffD1]) done lemma Card_rel_cardinal_rel [iff]: assumes types:"M(A)" shows "Card\<^bsup>M\<^esup>(|A|\<^bsup>M\<^esup>)" using assms proof (unfold cardinal_rel_def) show "Card\<^bsup>M\<^esup>(\ i. M(i) \ i \\<^bsup>M\<^esup> A)" proof (cases "\i[M]. Ord (i) \ i \\<^bsup>M\<^esup> A") case False thus ?thesis \ \degenerate case\ using Least_0[of "\i. M(i) \ i \\<^bsup>M\<^esup> A"] Card_rel_0 by fastforce next case True \ \real case: \<^term>\A\ is isomorphic to some ordinal\ then obtain i where i: "Ord(i)" "i \\<^bsup>M\<^esup> A" "M(i)" by blast show ?thesis proof (rule Card_relI [OF Ord_Least], rule notI) fix j assume j: "j < (\ i. M(i) \ i \\<^bsup>M\<^esup> A)" and "M(j)" assume "j \\<^bsup>M\<^esup> (\ i. M(i) \ i \\<^bsup>M\<^esup> A)" also have "... \\<^bsup>M\<^esup> A" using i LeastI[of "\i. M(i) \ i \\<^bsup>M\<^esup> A"] by (auto) finally have "j \\<^bsup>M\<^esup> A" using Least_closed'[of "\i. M(i) \ i \\<^bsup>M\<^esup> A"] by (simp add: \M(j)\ types) thus False using \M(j)\ by (blast intro:less_LeastE [OF _ j]) qed (auto intro:Least_closed) qed qed lemma cardinal_rel_eq_lemma: assumes i:"|i|\<^bsup>M\<^esup> \ j" and j: "j \ i" and types: "M(i)" "M(j)" shows "|j|\<^bsup>M\<^esup> = |i|\<^bsup>M\<^esup>" proof (rule eqpoll_relI [THEN cardinal_rel_cong]) show "j \\<^bsup>M\<^esup> i" by (rule le_imp_lepoll_rel [OF j]) (simp_all add:types) next have Oi: "Ord(i)" using j by (rule le_Ord2) with types have "i \\<^bsup>M\<^esup> |i|\<^bsup>M\<^esup>" by (blast intro: Ord_cardinal_rel_eqpoll_rel eqpoll_rel_sym) also from types have "... \\<^bsup>M\<^esup> j" by (blast intro: le_imp_lepoll_rel i) finally show "i \\<^bsup>M\<^esup> j" by (simp_all add:types) qed (simp_all add:types) lemma cardinal_rel_mono: assumes ij: "i \ j" and types:"M(i)" "M(j)" shows "|i|\<^bsup>M\<^esup> \ |j|\<^bsup>M\<^esup>" using Ord_cardinal_rel [OF \M(i)\] Ord_cardinal_rel [OF \M(j)\] proof (cases rule: Ord_linear_le) case le then show ?thesis . next case ge have i: "Ord(i)" using ij by (simp add: lt_Ord) have ci: "|i|\<^bsup>M\<^esup> \ j" by (blast intro: Ord_cardinal_rel_le ij le_trans i types) have "|i|\<^bsup>M\<^esup> = ||i|\<^bsup>M\<^esup>|\<^bsup>M\<^esup>" by (auto simp add: Ord_cardinal_rel_idem i types) also have "... = |j|\<^bsup>M\<^esup>" by (rule cardinal_rel_eq_lemma [OF ge ci]) (simp_all add:types) finally have "|i|\<^bsup>M\<^esup> = |j|\<^bsup>M\<^esup>" . thus ?thesis by (simp add:types) qed lemma cardinal_rel_lt_imp_lt: "[| |i|\<^bsup>M\<^esup> < |j|\<^bsup>M\<^esup>; Ord(i); Ord(j); M(i); M(j) |] ==> i < j" apply (rule Ord_linear2 [of i j], assumption+) apply (erule lt_trans2 [THEN lt_irrefl]) apply (erule cardinal_rel_mono, assumption+) done lemma Card_rel_lt_imp_lt: "[| |i|\<^bsup>M\<^esup> < K; Ord(i); Card\<^bsup>M\<^esup>(K); M(i); M(K)|] ==> i < K" by (simp (no_asm_simp) add: cardinal_rel_lt_imp_lt Card_rel_is_Ord Card_rel_cardinal_rel_eq) lemma Card_rel_lt_iff: "[| Ord(i); Card\<^bsup>M\<^esup>(K); M(i); M(K) |] ==> (|i|\<^bsup>M\<^esup> < K) \ (i < K)" by (blast intro: Card_rel_lt_imp_lt Ord_cardinal_rel_le [THEN lt_trans1]) lemma Card_rel_le_iff: "[| Ord(i); Card\<^bsup>M\<^esup>(K); M(i); M(K) |] ==> (K \ |i|\<^bsup>M\<^esup>) \ (K \ i)" by (simp add: Card_rel_lt_iff Card_rel_is_Ord not_lt_iff_le [THEN iff_sym]) lemma well_ord_lepoll_rel_imp_cardinal_rel_le: assumes wB: "well_ord(B,r)" and AB: "A \\<^bsup>M\<^esup> B" and types: "M(B)" "M(r)" "M(A)" shows "|A|\<^bsup>M\<^esup> \ |B|\<^bsup>M\<^esup>" using Ord_cardinal_rel [OF \M(A)\] Ord_cardinal_rel [OF \M(B)\] proof (cases rule: Ord_linear_le) case le thus ?thesis . next case ge from lepoll_rel_well_ord [OF AB wB] obtain s where s: "well_ord(A, s)" "M(s)" by (blast intro:types) have "B \\<^bsup>M\<^esup> |B|\<^bsup>M\<^esup>" by (blast intro: wB eqpoll_rel_sym well_ord_cardinal_rel_eqpoll_rel types) also have "... \\<^bsup>M\<^esup> |A|\<^bsup>M\<^esup>" by (rule le_imp_lepoll_rel [OF ge]) (simp_all add:types) also have "... \\<^bsup>M\<^esup> A" by (rule well_ord_cardinal_rel_eqpoll_rel [OF s(1) _ s(2)]) (simp_all add:types) finally have "B \\<^bsup>M\<^esup> A" by (simp_all add:types) hence "A \\<^bsup>M\<^esup> B" by (blast intro: eqpoll_relI AB types) hence "|A|\<^bsup>M\<^esup> = |B|\<^bsup>M\<^esup>" by (rule cardinal_rel_cong) (simp_all add:types) thus ?thesis by (simp_all add:types) qed lemma lepoll_rel_cardinal_rel_le: "[| A \\<^bsup>M\<^esup> i; Ord(i); M(A); M(i) |] ==> |A|\<^bsup>M\<^esup> \ i" using Memrel_closed apply (rule_tac le_trans) apply (erule well_ord_Memrel [THEN well_ord_lepoll_rel_imp_cardinal_rel_le], assumption+) apply (erule Ord_cardinal_rel_le, assumption) done lemma lepoll_rel_Ord_imp_eqpoll_rel: "[| A \\<^bsup>M\<^esup> i; Ord(i); M(A); M(i) |] ==> |A|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> A" by (blast intro: lepoll_rel_cardinal_rel_le well_ord_Memrel well_ord_cardinal_rel_eqpoll_rel dest!: lepoll_rel_well_ord) lemma lesspoll_rel_imp_eqpoll_rel: "[| A \\<^bsup>M\<^esup> i; Ord(i); M(A); M(i) |] ==> |A|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> A" using lepoll_rel_Ord_imp_eqpoll_rel[OF lesspoll_rel_imp_lepoll_rel] . lemma lesspoll_cardinal_lt_rel: shows "[| A \\<^bsup>M\<^esup> i; Ord(i); M(i); M(A) |] ==> |A|\<^bsup>M\<^esup> < i" proof - assume assms:"A \\<^bsup>M\<^esup> i" \Ord(i)\ \M(i)\ \M(A)\ then have A:"Ord(|A|\<^bsup>M\<^esup>)" "|A|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> A" "M(|A|\<^bsup>M\<^esup>)" using Ord_cardinal_rel lesspoll_rel_imp_eqpoll_rel by simp_all with assms have "|A|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> i" using eq_lesspoll_rel_trans by auto consider "|A|\<^bsup>M\<^esup>\i" | "|A|\<^bsup>M\<^esup>=i" | "i\|A|\<^bsup>M\<^esup>" using Ord_linear[OF \Ord(i)\ \Ord(|A|\<^bsup>M\<^esup>)\] by auto then have "|A|\<^bsup>M\<^esup> < i" proof(cases) case 1 then show ?thesis using ltI \Ord(i)\ by simp next case 2 with \|A|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> i\ \M(i)\ show ?thesis using lesspoll_rel_irrefl by simp next case 3 with \Ord(|A|\<^bsup>M\<^esup>)\ have "i<|A|\<^bsup>M\<^esup>" using ltI by simp with \M(A)\ A \M(i)\ have "i \\<^bsup>M\<^esup> |A|\<^bsup>M\<^esup>" using lt_Card_rel_imp_lesspoll_rel Card_rel_cardinal_rel by simp with \M(|A|\<^bsup>M\<^esup>)\ \M(i)\ show ?thesis using lesspoll_rel_irrefl lesspoll_rel_trans[OF \|A|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> i\ \i \\<^bsup>M\<^esup> _ \] by simp qed then show ?thesis by simp qed lemma cardinal_rel_subset_Ord: "[|A<=i; Ord(i); M(A); M(i)|] ==> |A|\<^bsup>M\<^esup> \ i" apply (drule subset_imp_lepoll_rel [THEN lepoll_rel_cardinal_rel_le]) apply (auto simp add: lt_def) apply (blast intro: Ord_trans) done \ \The next lemma is the first with several porting issues\ lemma cons_lepoll_rel_consD: "[| cons(u,A) \\<^bsup>M\<^esup> cons(v,B); u\A; v\B; M(u); M(A); M(v); M(B) |] ==> A \\<^bsup>M\<^esup> B" apply (simp add: def_lepoll_rel, unfold inj_def, safe) apply (rule_tac x = "\x\A. if f`x=v then f`u else f`x" in rexI) apply (rule CollectI) - (*Proving it's in the function space A->B*) + \ \Proving it's in the function space \<^term>\A\B\\ apply (rule if_type [THEN lam_type]) apply (blast dest: apply_funtype) apply (blast elim!: mem_irrefl dest: apply_funtype) - (*Proving it's injective*) + \ \Proving it's injective\ apply (auto simp add:transM[of _ A]) using lam_replacement_iff_lam_closed lam_if_then_apply_replacement by simp lemma cons_eqpoll_rel_consD: "[| cons(u,A) \\<^bsup>M\<^esup> cons(v,B); u\A; v\B; M(u); M(A); M(v); M(B) |] ==> A \\<^bsup>M\<^esup> B" apply (simp add: eqpoll_rel_iff) apply (blast intro: cons_lepoll_rel_consD) done lemma succ_lepoll_rel_succD: "succ(m) \\<^bsup>M\<^esup> succ(n) \ M(m) \ M(n) ==> m \\<^bsup>M\<^esup> n" apply (unfold succ_def) apply (erule cons_lepoll_rel_consD) apply (rule mem_not_refl)+ apply assumption+ done lemma nat_lepoll_rel_imp_le: "m \ nat ==> n \ nat \ m \\<^bsup>M\<^esup> n \ M(m) \ M(n) \ m \ n" proof (induct m arbitrary: n rule: nat_induct) case 0 thus ?case by (blast intro!: nat_0_le) next case (succ m) show ?case using \n \ nat\ proof (cases rule: natE) case 0 thus ?thesis using succ by (simp add: def_lepoll_rel inj_def) next case (succ n') thus ?thesis using succ.hyps \ succ(m) \\<^bsup>M\<^esup> n\ by (blast dest!: succ_lepoll_rel_succD) qed qed lemma nat_eqpoll_rel_iff: "[| m \ nat; n \ nat; M(m); M(n) |] ==> m \\<^bsup>M\<^esup> n \ m = n" apply (rule iffI) apply (blast intro: nat_lepoll_rel_imp_le le_anti_sym elim!: eqpoll_relE) apply (simp add: eqpoll_rel_refl) done lemma nat_into_Card_rel: assumes n: "n \ nat" and types: "M(n)" shows "Card\<^bsup>M\<^esup>(n)" using types apply (subst Card_rel_def) proof (unfold cardinal_rel_def, rule sym) have "Ord(n)" using n by auto moreover { fix i assume "i < n" "M(i)" "i \\<^bsup>M\<^esup> n" hence False using n by (auto simp add: lt_nat_in_nat [THEN nat_eqpoll_rel_iff] types) } ultimately show "(\ i. M(i) \ i \\<^bsup>M\<^esup> n) = n" by (auto intro!: Least_equality types eqpoll_rel_refl) qed lemmas cardinal_rel_0 = nat_0I [THEN nat_into_Card_rel, THEN Card_rel_cardinal_rel_eq, simplified, iff] lemmas cardinal_rel_1 = nat_1I [THEN nat_into_Card_rel, THEN Card_rel_cardinal_rel_eq, simplified, iff] lemma succ_lepoll_rel_natE: "[| succ(n) \\<^bsup>M\<^esup> n; n \ nat |] ==> P" by (rule nat_lepoll_rel_imp_le [THEN lt_irrefl], auto) lemma nat_lepoll_rel_imp_ex_eqpoll_rel_n: "[| n \ nat; nat \\<^bsup>M\<^esup> X ; M(n); M(X)|] ==> \Y[M]. Y \ X & n \\<^bsup>M\<^esup> Y" apply (simp add: def_lepoll_rel def_eqpoll_rel) apply (fast del: subsetI subsetCE intro!: subset_SIs dest!: Ord_nat [THEN [2] OrdmemD, THEN [2] restrict_inj] elim!: restrict_bij inj_is_fun [THEN fun_is_rel, THEN image_subset]) done lemma lepoll_rel_succ: "M(i) \ i \\<^bsup>M\<^esup> succ(i)" by (blast intro: subset_imp_lepoll_rel) lemma lepoll_rel_imp_lesspoll_rel_succ: assumes A: "A \\<^bsup>M\<^esup> m" and m: "m \ nat" and types: "M(A)" "M(m)" shows "A \\<^bsup>M\<^esup> succ(m)" proof - { assume "A \\<^bsup>M\<^esup> succ(m)" hence "succ(m) \\<^bsup>M\<^esup> A" by (rule eqpoll_rel_sym) (auto simp add:types) also have "... \\<^bsup>M\<^esup> m" by (rule A) finally have "succ(m) \\<^bsup>M\<^esup> m" by (auto simp add:types) hence False by (rule succ_lepoll_rel_natE) (rule m) } moreover have "A \\<^bsup>M\<^esup> succ(m)" by (blast intro: lepoll_rel_trans A lepoll_rel_succ types) ultimately show ?thesis by (auto simp add: types lesspoll_rel_def) qed lemma lesspoll_rel_succ_imp_lepoll_rel: "[| A \\<^bsup>M\<^esup> succ(m); m \ nat; M(A); M(m) |] ==> A \\<^bsup>M\<^esup> m" proof - { assume "m \ nat" "M(A)" "M(m)" "A \\<^bsup>M\<^esup> succ(m)" "\f\inj\<^bsup>M\<^esup>(A, succ(m)). f \ surj\<^bsup>M\<^esup>(A, succ(m))" moreover from this obtain f where "M(f)" "f\inj\<^bsup>M\<^esup>(A,succ(m))" using def_lepoll_rel by auto moreover from calculation have "f \ surj\<^bsup>M\<^esup>(A, succ(m))" by simp ultimately have "\f[M]. f \ inj\<^bsup>M\<^esup>(A, m)" using inj_rel_not_surj_rel_succ by auto } from this show "\ A \\<^bsup>M\<^esup> succ(m); m \ nat; M(A); M(m) \ \ A \\<^bsup>M\<^esup> m" unfolding lepoll_rel_def eqpoll_rel_def bij_rel_def lesspoll_rel_def by (simp del:mem_inj_abs) qed lemma lesspoll_rel_succ_iff: "m \ nat \ M(A) ==> A \\<^bsup>M\<^esup> succ(m) \ A \\<^bsup>M\<^esup> m" by (blast intro!: lepoll_rel_imp_lesspoll_rel_succ lesspoll_rel_succ_imp_lepoll_rel) lemma lepoll_rel_succ_disj: "[| A \\<^bsup>M\<^esup> succ(m); m \ nat; M(A) ; M(m)|] ==> A \\<^bsup>M\<^esup> m | A \\<^bsup>M\<^esup> succ(m)" apply (rule disjCI) apply (rule lesspoll_rel_succ_imp_lepoll_rel) prefer 2 apply assumption apply (simp (no_asm_simp) add: lesspoll_rel_def, assumption+) done lemma lesspoll_rel_cardinal_rel_lt: "[| A \\<^bsup>M\<^esup> i; Ord(i); M(A); M(i) |] ==> |A|\<^bsup>M\<^esup> < i" apply (unfold lesspoll_rel_def, clarify) apply (frule lepoll_rel_cardinal_rel_le, assumption+) \ \because of types\ apply (blast intro: well_ord_Memrel well_ord_cardinal_rel_eqpoll_rel [THEN eqpoll_rel_sym] dest: lepoll_rel_well_ord elim!: leE) done - lemma lt_not_lepoll_rel: assumes n: "n nat" and types:"M(n)" "M(i)" shows "~ i \\<^bsup>M\<^esup> n" proof - { assume i: "i \\<^bsup>M\<^esup> n" have "succ(n) \\<^bsup>M\<^esup> i" using n by (elim ltE, blast intro: Ord_succ_subsetI [THEN subset_imp_lepoll_rel] types) also have "... \\<^bsup>M\<^esup> n" by (rule i) finally have "succ(n) \\<^bsup>M\<^esup> n" by (simp add:types) hence False by (rule succ_lepoll_rel_natE) (rule n) } thus ?thesis by auto qed text\A slightly weaker version of \nat_eqpoll_rel_iff\\ lemma Ord_nat_eqpoll_rel_iff: assumes i: "Ord(i)" and n: "n \ nat" and types: "M(i)" "M(n)" shows "i \\<^bsup>M\<^esup> n \ i=n" using i nat_into_Ord [OF n] proof (cases rule: Ord_linear_lt) case lt hence "i \ nat" by (rule lt_nat_in_nat) (rule n) thus ?thesis by (simp add: nat_eqpoll_rel_iff n types) next case eq thus ?thesis by (simp add: eqpoll_rel_refl types) next case gt hence "~ i \\<^bsup>M\<^esup> n" using n by (rule lt_not_lepoll_rel) (simp_all add: types) hence "~ i \\<^bsup>M\<^esup> n" using n by (blast intro: eqpoll_rel_imp_lepoll_rel types) moreover have "i \ n" using \n by auto ultimately show ?thesis by blast qed lemma Card_rel_nat: "Card\<^bsup>M\<^esup>(nat)" proof - { fix i assume i: "i < nat" "i \\<^bsup>M\<^esup> nat" "M(i)" hence "~ nat \\<^bsup>M\<^esup> i" by (simp add: lt_def lt_not_lepoll_rel) hence False using i by (simp add: eqpoll_rel_iff) } hence "(\ i. M(i) \ i \\<^bsup>M\<^esup> nat) = nat" by (blast intro: Least_equality eqpoll_rel_refl) thus ?thesis by (auto simp add: Card_rel_def cardinal_rel_def) qed lemma nat_le_cardinal_rel: "nat \ i \ M(i) ==> nat \ |i|\<^bsup>M\<^esup>" apply (rule Card_rel_nat [THEN Card_rel_cardinal_rel_eq, THEN subst], simp_all) apply (erule cardinal_rel_mono, simp_all) done lemma n_lesspoll_rel_nat: "n \ nat ==> n \\<^bsup>M\<^esup> nat" by (blast intro: Card_rel_nat ltI lt_Card_rel_imp_lesspoll_rel) lemma cons_lepoll_rel_cong: "[| A \\<^bsup>M\<^esup> B; b \ B; M(A); M(B); M(b); M(a) |] ==> cons(a,A) \\<^bsup>M\<^esup> cons(b,B)" apply (subst (asm) def_lepoll_rel, simp_all, subst def_lepoll_rel, simp_all, safe) apply (rule_tac x = "\y\cons (a,A) . if y=a then b else f`y" in rexI) apply (rule_tac d = "%z. if z \ B then converse (f) `z else a" in lam_injective) apply (safe elim!: consE') apply simp_all apply (blast intro: inj_is_fun [THEN apply_type])+ apply (auto intro:lam_closed lam_if_then_replacement simp add:transM[of _ A]) done lemma cons_eqpoll_rel_cong: "[| A \\<^bsup>M\<^esup> B; a \ A; b \ B; M(A); M(B); M(a) ; M(b) |] ==> cons(a,A) \\<^bsup>M\<^esup> cons(b,B)" by (simp add: eqpoll_rel_iff cons_lepoll_rel_cong) lemma cons_lepoll_rel_cons_iff: "[| a \ A; b \ B; M(a); M(A); M(b); M(B) |] ==> cons(a,A) \\<^bsup>M\<^esup> cons(b,B) \ A \\<^bsup>M\<^esup> B" by (blast intro: cons_lepoll_rel_cong cons_lepoll_rel_consD) lemma cons_eqpoll_rel_cons_iff: "[| a \ A; b \ B; M(a); M(A); M(b); M(B) |] ==> cons(a,A) \\<^bsup>M\<^esup> cons(b,B) \ A \\<^bsup>M\<^esup> B" by (blast intro: cons_eqpoll_rel_cong cons_eqpoll_rel_consD) lemma singleton_eqpoll_rel_1: "M(a) \ {a} \\<^bsup>M\<^esup> 1" apply (unfold succ_def) apply (blast intro!: eqpoll_rel_refl [THEN cons_eqpoll_rel_cong]) done lemma cardinal_rel_singleton: "M(a) \ |{a}|\<^bsup>M\<^esup> = 1" apply (rule singleton_eqpoll_rel_1 [THEN cardinal_rel_cong, THEN trans]) apply (simp (no_asm) add: nat_into_Card_rel [THEN Card_rel_cardinal_rel_eq]) apply auto done lemma not_0_is_lepoll_rel_1: "A \ 0 ==> M(A) \ 1 \\<^bsup>M\<^esup> A" apply (erule not_emptyE) apply (rule_tac a = "cons (x, A-{x}) " in subst) apply (rule_tac [2] a = "cons(0,0)" and P= "%y. y \\<^bsup>M\<^esup> cons (x, A-{x})" in subst) apply auto proof - fix x assume "M(A)" then show "x \ A \ {0} \\<^bsup>M\<^esup> cons(x, A - {x})" by (auto intro: cons_lepoll_rel_cong transM[OF _ \M(A)\] subset_imp_lepoll_rel) qed - lemma succ_eqpoll_rel_cong: "A \\<^bsup>M\<^esup> B \ M(A) \ M(B) ==> succ(A) \\<^bsup>M\<^esup> succ(B)" apply (unfold succ_def) apply (simp add: cons_eqpoll_rel_cong mem_not_refl) done text\The next result was not straightforward to port, and even a different statement was needed.\ lemma sum_bij_rel: "[| f \ bij\<^bsup>M\<^esup>(A,C); g \ bij\<^bsup>M\<^esup>(B,D); M(f); M(A); M(C); M(g); M(B); M(D)|] ==> (\z\A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) \ bij\<^bsup>M\<^esup>(A+B, C+D)" proof - assume asm:"f \ bij\<^bsup>M\<^esup>(A,C)" "g \ bij\<^bsup>M\<^esup>(B,D)" "M(f)" "M(A)" "M(C)" "M(g)" "M(B)" "M(D)" then have "M(\z\A+B. case(%x. Inl(f`x), %y. Inr(g`y), z))" using transM[OF _ \M(A)\] transM[OF _ \M(B)\] by (auto intro:case_replacement4[THEN lam_closed]) with asm show ?thesis apply simp apply (rule_tac d = "case (%x. Inl (converse(f)`x), %y. Inr(converse(g)`y))" in lam_bijective) apply (typecheck add: bij_is_inj inj_is_fun) apply (auto simp add: left_inverse_bij right_inverse_bij) done qed lemma sum_bij_rel': assumes "f \ bij\<^bsup>M\<^esup>(A,C)" "g \ bij\<^bsup>M\<^esup>(B,D)" "M(f)" "M(A)" "M(C)" "M(g)" "M(B)" "M(D)" shows "(\z\A+B. case(\x. Inl(f`x), \y. Inr(g`y), z)) \ bij(A+B, C+D)" "M(\z\A+B. case(\x. Inl(f`x), \y. Inr(g`y), z))" proof - from assms show "M(\z\A+B. case(\x. Inl(f`x), \y. Inr(g`y), z))" using transM[OF _ \M(A)\] transM[OF _ \M(B)\] by (auto intro:case_replacement4[THEN lam_closed]) with assms show "(\z\A+B. case(\x. Inl(f`x), \y. Inr(g`y), z)) \ bij(A+B, C+D)" apply simp apply (rule_tac d = "case (%x. Inl (converse(f)`x), %y. Inr(converse(g)`y))" in lam_bijective) apply (typecheck add: bij_is_inj inj_is_fun) apply (auto simp add: left_inverse_bij right_inverse_bij) done qed lemma sum_eqpoll_rel_cong: assumes "A \\<^bsup>M\<^esup> C" "B \\<^bsup>M\<^esup> D" "M(A)" "M(C)" "M(B)" "M(D)" shows "A+B \\<^bsup>M\<^esup> C+D" using assms proof (simp add: def_eqpoll_rel, safe, rename_tac g) fix f g assume "M(f)" "f \ bij(A, C)" "M(g)" "g \ bij(B, D)" with assms obtain h where "h\bij(A+B, C+D)" "M(h)" using sum_bij_rel'[of f A C g B D] by simp then show "\f[M]. f \ bij(A + B, C + D)" by auto qed lemma prod_bij_rel': assumes "f \ bij\<^bsup>M\<^esup>(A,C)" "g \ bij\<^bsup>M\<^esup>(B,D)" "M(f)" "M(A)" "M(C)" "M(g)" "M(B)" "M(D)" shows "(\\A*B. ) \ bij(A*B, C*D)" "M(\\A*B. )" proof - from assms show "M((\\A*B. ))" using transM[OF _ \M(A)\] transM[OF _ \M(B)\] transM[OF _ cartprod_closed, of _ A B] by (auto intro:prod_fun_replacement[THEN lam_closed, of f g "A\B"]) with assms show "(\\A*B. ) \ bij(A*B, C*D)" apply simp apply (rule_tac d = "%. " in lam_bijective) apply (typecheck add: bij_is_inj inj_is_fun) apply (auto simp add: left_inverse_bij right_inverse_bij) done qed lemma prod_eqpoll_rel_cong: assumes "A \\<^bsup>M\<^esup> C" "B \\<^bsup>M\<^esup> D" "M(A)" "M(C)" "M(B)" "M(D)" shows "A\B \\<^bsup>M\<^esup> C\D" using assms proof (simp add: def_eqpoll_rel, safe, rename_tac g) fix f g assume "M(f)" "f \ bij(A, C)" "M(g)" "g \ bij(B, D)" with assms obtain h where "h\bij(A\B, C\D)" "M(h)" using prod_bij_rel'[of f A C g B D] by simp then show "\f[M]. f \ bij(A \ B, C \ D)" by auto qed lemma inj_rel_disjoint_eqpoll_rel: "[| f \ inj\<^bsup>M\<^esup>(A,B); A \ B = 0;M(f); M(A);M(B) |] ==> A \ (B - range(f)) \\<^bsup>M\<^esup> B" apply (simp add: def_eqpoll_rel) apply (rule rexI) apply (rule_tac c = "%x. if x \ A then f`x else x" and d = "%y. if y \ range (f) then converse (f) `y else y" in lam_bijective) apply (blast intro!: if_type inj_is_fun [THEN apply_type]) apply (simp (no_asm_simp) add: inj_converse_fun [THEN apply_funtype]) apply (safe elim!: UnE') apply (simp_all add: inj_is_fun [THEN apply_rangeI]) apply (blast intro: inj_converse_fun [THEN apply_type]) proof - assume "f \ inj(A, B)" "A \ B = 0" "M(f)" "M(A)" "M(B)" then show "M(\x\A \ (B - range(f)). if x \ A then f ` x else x)" using transM[OF _ \M(A)\] transM[OF _ \M(B)\] lam_replacement_iff_lam_closed lam_if_then_replacement2 by auto qed lemma Diff_sing_lepoll_rel: "[| a \ A; A \\<^bsup>M\<^esup> succ(n); M(a); M(A); M(n) |] ==> A - {a} \\<^bsup>M\<^esup> n" apply (unfold succ_def) apply (rule cons_lepoll_rel_consD) apply (rule_tac [3] mem_not_refl) apply (erule cons_Diff [THEN ssubst], simp_all) done lemma lepoll_rel_Diff_sing: assumes A: "succ(n) \\<^bsup>M\<^esup> A" and types: "M(n)" "M(A)" "M(a)" shows "n \\<^bsup>M\<^esup> A - {a}" proof - have "cons(n,n) \\<^bsup>M\<^esup> A" using A by (unfold succ_def) also from types have "... \\<^bsup>M\<^esup> cons(a, A-{a})" by (blast intro: subset_imp_lepoll_rel) finally have "cons(n,n) \\<^bsup>M\<^esup> cons(a, A-{a})" by (simp_all add:types) with types show ?thesis by (blast intro: cons_lepoll_rel_consD mem_irrefl) qed lemma Diff_sing_eqpoll_rel: "[| a \ A; A \\<^bsup>M\<^esup> succ(n); M(a); M(A); M(n) |] ==> A - {a} \\<^bsup>M\<^esup> n" by (blast intro!: eqpoll_relI elim!: eqpoll_relE intro: Diff_sing_lepoll_rel lepoll_rel_Diff_sing) lemma lepoll_rel_1_is_sing: "[| A \\<^bsup>M\<^esup> 1; a \ A ;M(a); M(A) |] ==> A = {a}" apply (frule Diff_sing_lepoll_rel, assumption+, simp) apply (drule lepoll_rel_0_is_0, simp) apply (blast elim: equalityE) done lemma Un_lepoll_rel_sum: "M(A) \ M(B) \ A \ B \\<^bsup>M\<^esup> A+B" apply (simp add: def_lepoll_rel) apply (rule_tac x = "\x\A \ B. if x\A then Inl (x) else Inr (x)" in rexI) apply (rule_tac d = "%z. snd (z)" in lam_injective) apply force apply (simp add: Inl_def Inr_def) proof - assume "M(A)" "M(B)" then show "M(\x\A \ B. if x \ A then Inl(x) else Inr(x))" using transM[OF _ \M(A)\] transM[OF _ \M(B)\] if_then_Inj_replacement by (rule_tac lam_closed) auto qed lemma well_ord_Un_M: assumes "well_ord(X,R)" "well_ord(Y,S)" and types: "M(X)" "M(R)" "M(Y)" "M(S)" shows "\T[M]. well_ord(X \ Y, T)" using assms by (erule_tac well_ord_radd [THEN [3] Un_lepoll_rel_sum [THEN lepoll_rel_well_ord]]) (auto simp add: types) lemma disj_Un_eqpoll_rel_sum: "M(A) \ M(B) \ A \ B = 0 \ A \ B \\<^bsup>M\<^esup> A + B" apply (simp add: def_eqpoll_rel) apply (rule_tac x = "\a\A \ B. if a \ A then Inl (a) else Inr (a)" in rexI) apply (rule_tac d = "%z. case (%x. x, %x. x, z)" in lam_bijective) apply auto proof - assume "M(A)" "M(B)" then show "M(\x\A \ B. if x \ A then Inl(x) else Inr(x))" using transM[OF _ \M(A)\] transM[OF _ \M(B)\] if_then_Inj_replacement by (rule_tac lam_closed) auto qed lemma eqpoll_rel_imp_Finite_rel_iff: "A \\<^bsup>M\<^esup> B ==> M(A) \ M(B) \ Finite_rel(M,A) \ Finite_rel(M,B)" apply (unfold Finite_rel_def) apply (blast intro: eqpoll_rel_trans eqpoll_rel_sym) done \ \It seems reasonable to have the absoluteness of \<^term>\Finite\ here, and deduce the rest of the results from this. Perhaps modularize that proof to have absoluteness of injections and bijections of finite sets (cf. @{thm lesspoll_rel_succ_imp_lepoll_rel}.\ -lemma Finite_abs[simp]: assumes "M(A)" shows "Finite_rel(M,A) \ Finite(A)" +lemma Finite_abs[simp]: + assumes "M(A)" + shows "Finite_rel(M,A) \ Finite(A)" unfolding Finite_rel_def Finite_def proof (simp, intro iffI) assume "\n\nat. A \\<^bsup>M\<^esup> n" then obtain n where "A \\<^bsup>M\<^esup> n" "n\nat" by blast with assms show "\n\nat. A \ n" unfolding eqpoll_def using nat_into_M by (auto simp add:def_eqpoll_rel) next fix n assume "\n\nat. A \ n" then obtain n where "A \ n" "n\nat" by blast moreover from this obtain f where "f \ bij(A,n)" unfolding eqpoll_def by auto moreover note assms moreover from calculation have "converse(f) \ n\A" using bij_is_fun by simp moreover from calculation have "M(converse(f))" using transM[of _ "n\A"] by simp moreover from calculation have "M(f)" using bij_is_fun fun_is_rel[of "f" A "\_. n", THEN converse_converse] converse_closed[of "converse(f)"] by simp ultimately show "\n\nat. A \\<^bsup>M\<^esup> n" by (force dest:nat_into_M simp add:def_eqpoll_rel) qed -(* \ \From the next result, the relative versions of @{thm Finite_Fin_lemma} and @{thm Fin_lemma} should follow\ -lemma nat_eqpoll_imp_eqpoll_rel: - assumes "n \ nat" "A \ n" and types:"M(n)" "M(A)" - shows "A \\<^bsup>M\<^esup> n" -*) lemma lepoll_rel_nat_imp_Finite_rel: assumes A: "A \\<^bsup>M\<^esup> n" and n: "n \ nat" and types: "M(A)" "M(n)" shows "Finite_rel(M,A)" proof - have "A \\<^bsup>M\<^esup> n \ Finite_rel(M,A)" using n proof (induct n) case 0 hence "A = 0" by (rule lepoll_rel_0_is_0, simp_all add:types) thus ?case by simp next case (succ n) hence "A \\<^bsup>M\<^esup> n \ A \\<^bsup>M\<^esup> succ(n)" by (blast dest: lepoll_rel_succ_disj intro:types) thus ?case using succ by (auto simp add: Finite_rel_def types) qed thus ?thesis using A . qed lemma lesspoll_rel_nat_is_Finite_rel: "A \\<^bsup>M\<^esup> nat \ M(A) \ Finite_rel(M,A)" apply (unfold Finite_rel_def) apply (auto dest: ltD lesspoll_rel_cardinal_rel_lt lesspoll_rel_imp_eqpoll_rel [THEN eqpoll_rel_sym]) done lemma lepoll_rel_Finite_rel: assumes Y: "Y \\<^bsup>M\<^esup> X" and X: "Finite_rel(M,X)" and types:"M(Y)" "M(X)" shows "Finite_rel(M,Y)" proof - obtain n where n: "n \ nat" "X \\<^bsup>M\<^esup> n" "M(n)" using X by (auto simp add: Finite_rel_def) have "Y \\<^bsup>M\<^esup> X" by (rule Y) also have "... \\<^bsup>M\<^esup> n" by (rule n) finally have "Y \\<^bsup>M\<^esup> n" by (simp_all add:types \M(n)\) thus ?thesis using n by (simp add: lepoll_rel_nat_imp_Finite_rel types \M(n)\ del:Finite_abs) qed lemma succ_lepoll_rel_imp_not_empty: "succ(x) \\<^bsup>M\<^esup> y ==> M(x) \ M(y) \ y \ 0" by (fast dest!: lepoll_rel_0_is_0) lemma eqpoll_rel_succ_imp_not_empty: "x \\<^bsup>M\<^esup> succ(n) ==> M(x) \ M(n) \ x \ 0" by (fast elim!: eqpoll_rel_sym [THEN eqpoll_rel_0_is_0, THEN succ_neq_0]) lemma Finite_subset_closed: assumes "Finite(B)" "B\A" "M(A)" shows "M(B)" proof - from \Finite(B)\ \B\A\ show ?thesis proof(induct,simp) case (cons x D) with assms have "M(D)" "x\A" unfolding cons_def by auto then show ?case using transM[OF _ \M(A)\] by simp qed qed lemma Finite_Pow_abs: assumes "Finite(A)" " M(A)" shows "Pow(A) = Pow_rel(M,A)" using Finite_subset_closed[OF subset_Finite] assms Pow_rel_char by auto lemma Finite_Pow_rel: assumes "Finite(A)" "M(A)" shows "Finite(Pow_rel(M,A))" using Finite_Pow Finite_Pow_abs[symmetric] assms by simp lemma Pow_rel_0 [simp]: "Pow_rel(M,0) = {0}" using Finite_Pow_abs[of 0] by simp lemma eqpoll_rel_imp_Finite: "A \\<^bsup>M\<^esup> B \ Finite(A) \ M(A) \ M(B) \ Finite(B)" proof - assume "A \\<^bsup>M\<^esup> B" "Finite(A)" "M(A)" "M(B)" then obtain f n g where "f\bij(A,B)" "n\nat" "g\bij(A,n)" unfolding Finite_def eqpoll_def eqpoll_rel_def using bij_rel_char by auto then have "g O converse(f) \ bij(B,n)" using bij_converse_bij comp_bij by simp with \n\_\ show"Finite(B)" unfolding Finite_def eqpoll_def by auto qed lemma eqpoll_rel_imp_Finite_iff: "A \\<^bsup>M\<^esup> B \ M(A) \ M(B) \ Finite(A) \ Finite(B)" using eqpoll_rel_imp_Finite eqpoll_rel_sym by force end \ \\<^locale>\M_cardinals\\ -end +end \ No newline at end of file diff --git a/thys/Transitive_Models/Delta_System_Relative.thy b/thys/Transitive_Models/Delta_System_Relative.thy --- a/thys/Transitive_Models/Delta_System_Relative.thy +++ b/thys/Transitive_Models/Delta_System_Relative.thy @@ -1,378 +1,373 @@ section\The Delta System Lemma, Relativized\label{sec:dsl-rel}\ theory Delta_System_Relative imports Cardinal_Library_Relative begin relativize functional "delta_system" "delta_system_rel" external locale M_delta = M_cardinal_library_extra + assumes countable_lepoll_assms: "M(G) \ M(A) \ 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)\)" begin lemma disjoint_separation: "M(c) \ separation(M, \ x. \a. \b. x=\a,b\ \ a \ b = c)" using separation_Pair separation_eq lam_replacement_constant lam_replacement_Int by simp -lemma insnd_ball: "M(G) \ separation(M, \p. \x\G. x \ snd(p) \ fst(p) \ x)" - using separation_ball separation_iff' lam_replacement_fst lam_replacement_snd - separation_in lam_replacement_hcomp - by simp - lemma (in M_trans) mem_F_bound6: fixes F G defines "F \ \_ x. Collect(G, (\)(x))" shows "x\F(G,c) \ c \ (range(f) \ \G)" using apply_0 unfolding F_def by (cases "M(c)", auto simp:F_def) lemma delta_system_Aleph_rel1: assumes "\A\F. Finite(A)" "F \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "M(F)" shows "\D[M]. D \ F \ delta_system(D) \ D \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" proof - have "M(G) \ M(p) \ M({A\G . p \ A})" for G p proof - assume "M(G)" "M(p)" have "{A\G . p \ A} = G \ (Memrel({p} \ G) `` {p})" unfolding Memrel_def by auto with \M(G)\ \M(p)\ show ?thesis by simp qed from \\A\F. Finite(A)\ \M(F)\ have "M(\A\F. |A|\<^bsup>M\<^esup>)" using cardinal_rel_separation Finite_cardinal_rel_in_nat[OF _ transM[of _ F]] separation_imp_lam_closed[of _ "cardinal_rel(M)" \] by simp text\Since all members are finite,\ with \\A\F. Finite(A)\ \M(F)\ have "(\A\F. |A|\<^bsup>M\<^esup>) : F \\<^bsup>M\<^esup> \" (is "?cards : _") by (simp add:mem_function_space_rel_abs, rule_tac lam_type) (force dest:transM) moreover from this have a:"?cards -`` {n} = { A\F . |A|\<^bsup>M\<^esup> = n }" for n using vimage_lam by auto moreover note \F \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ \M(F)\ moreover from calculation text\there are uncountably many have the same cardinal:\ obtain n where "n\\" "|?cards -`` {n}|\<^bsup>M\<^esup> = \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" using eqpoll_rel_Aleph_rel1_cardinal_rel_vimage[of F ?cards] by auto moreover define G where "G \ ?cards -`` {n}" moreover from calculation and \M(?cards)\ have "M(G)" by blast moreover from calculation have "G \ F" by auto ultimately text\Therefore, without loss of generality, we can assume that all elements of the family have cardinality \<^term>\n\\\.\ have "A\G \ |A|\<^bsup>M\<^esup> = n" and "G \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" and "M(G)" for A using cardinal_rel_Card_rel_eqpoll_rel_iff by auto with \n\\\ text\So we prove the result by induction on this \<^term>\n\ and generalizing \<^term>\G\, since the argument requires changing the family in order to apply the inductive hypothesis.\ have "\D[M]. D \ G \ delta_system(D) \ D \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" proof (induct arbitrary:G) case 0 \ \This case is impossible\ then have "G \ {0}" using cardinal_rel_0_iff_0 by (blast dest:transM) with \G \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ \M(G)\ show ?case using nat_lt_Aleph_rel1 subset_imp_le_cardinal_rel[of G "{0}"] lt_trans2 cardinal_rel_Card_rel_eqpoll_rel_iff by auto next case (succ n) have "\a\G. Finite(a)" proof fix a assume "a \ G" moreover note \M(G)\ \n\\\ moreover from calculation have "M(a)" by (auto dest: transM) moreover from succ and \a\G\ have "|a|\<^bsup>M\<^esup> = succ(n)" by simp ultimately show "Finite(a)" using Finite_cardinal_rel_iff' nat_into_Finite[of "succ(n)"] by fastforce qed show "\D[M]. D \ G \ delta_system(D) \ D \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" proof (cases "\p[M]. {A\G . p \ A} \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>") case True \ \the positive case, uncountably many sets with a common element\ then obtain p where "{A\G . p \ A} \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "M(p)" by blast moreover note 1=\M(G)\ \M(G) \ M(p) \ M({A\G . p \ A})\ singleton_closed[OF \M(p)\] moreover from this have "M({x - {p} . x \ {x \ G . p \ x}})" using RepFun_closed[OF lam_replacement_Diff'[THEN lam_replacement_imp_strong_replacement]] Diff_closed[OF transM[OF _ 1(2)]] by auto moreover from 1 have "M(converse(\x\{x \ G . p \ x}. x - {p}))" (is "M(converse(?h))") using converse_closed[of ?h] lam_closed[OF diff_Pair_replacement] Diff_closed[OF transM[OF _ 1(2)]] by auto moreover from calculation have "{A-{p} . A\{X\G. p\X}} \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" (is "?F \\<^bsup>M\<^esup> _") using Diff_bij_rel[of "{A\G . p \ A}" "{p}", THEN comp_bij_rel[OF bij_rel_converse_bij_rel, where C="\\<^bsub>1\<^esub>\<^bsup>M\<^esup>", THEN bij_rel_imp_eqpoll_rel, of _ _ ?F]] unfolding eqpoll_rel_def by (auto simp del:mem_bij_abs) text\Now using the hypothesis of the successor case,\ moreover from \\A. A\G \ |A|\<^bsup>M\<^esup>=succ(n)\ \\a\G. Finite(a)\ and this \M(G)\ have "p\A \ A\G \ |A - {p}|\<^bsup>M\<^esup> = n" for A using Finite_imp_succ_cardinal_rel_Diff[of _ p] by (force dest: transM) moreover have "\a\?F. Finite(a)" proof (clarsimp) fix A assume "p\A" "A\G" with \\A. p \ A \ A \ G \ |A - {p}|\<^bsup>M\<^esup> = n\ and \n\\\ \M(G)\ have "Finite(|A - {p}|\<^bsup>M\<^esup>)" using nat_into_Finite by simp moreover from \p\A\ \A\G\ \M(G)\ have "M(A - {p})" by (auto dest: transM) ultimately show "Finite(A - {p})" using Finite_cardinal_rel_iff' by simp qed moreover text\we may apply the inductive hypothesis to the new family \<^term>\?F\:\ note \(\A. A \ ?F \ |A|\<^bsup>M\<^esup> = n) \ ?F \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ M(?F) \ \D[M]. D \ ?F \ delta_system(D) \ D \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ moreover note 1=\M(G)\ \M(G) \ M(p) \ M({A\G . p \ A})\ singleton_closed[OF \M(p)\] moreover from this have "M({x - {p} . x \ {x \ G . p \ x}})" using RepFun_closed[OF lam_replacement_Diff'[THEN lam_replacement_imp_strong_replacement]] Diff_closed[OF transM[OF _ 1(2)]] by auto ultimately obtain D where "D\{A-{p} . A\{X\G. p\X}}" "delta_system(D)" "D \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "M(D)" by auto moreover from this obtain r where "\A\D. \B\D. A \ B \ A \ B = r" by fastforce then have "\A\D.\B\D. A\{p} \ B\{p}\(A \ {p}) \ (B \ {p}) = r\{p}" by blast ultimately have "delta_system({B \ {p} . B\D})" (is "delta_system(?D)") by fastforce moreover from \M(D)\ \M(p)\ have "M(?D)" using RepFun_closed un_Pair_replacement transM[of _ D] by auto moreover from \D \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ \M(D)\ have "Infinite(D)" "|D|\<^bsup>M\<^esup> = \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" using uncountable_rel_iff_subset_eqpoll_rel_Aleph_rel1[THEN iffD2, THEN uncountable_rel_imp_Infinite, of D] cardinal_rel_eqpoll_rel_iff[of D "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>"] \M(D)\ \D \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ by auto moreover from this \M(?D)\ \M(D)\ \M(p)\ have "?D \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" using cardinal_rel_map_Un[of D "{p}"] naturals_lt_nat cardinal_rel_eqpoll_rel_iff[THEN iffD1] by simp moreover note \D \ {A-{p} . A\{X\G. p\X}}\ have "?D \ G" proof - { fix A assume "A\G" "p\A" moreover from this have "A = A - {p} \ {p}" by blast ultimately have "A -{p} \ {p} \ G" by auto } with \D \ {A-{p} . A\{X\G. p\X}}\ show ?thesis by blast qed moreover note \M(?D)\ ultimately show "\D[M]. D \ G \ delta_system(D) \ D \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" by auto next case False note \\ (\p[M]. {A \ G . p \ A} \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>)\ \ \the other case\ \M(G)\ \\p. M(G) \ M(p) \ M({A\G . p \ A})\ moreover from \G \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ and this have "M(p) \ {A \ G . p \ A} \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" (is "_ \ ?G(p) \\<^bsup>M\<^esup> _") for p by (auto intro!:lepoll_rel_eq_trans[OF subset_imp_lepoll_rel] dest:transM) moreover from calculation have "M(p) \ ?G(p) \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" for p using \M(G) \ M(p) \ M({A\G . p \ A})\ unfolding lesspoll_rel_def by simp moreover from calculation have "M(p) \ ?G(p) \\<^bsup>M\<^esup> \" for p using lesspoll_rel_Aleph_rel_succ[of 0] Aleph_rel_zero by auto moreover have "{A \ G . S \ A \ 0} = (\p\S. ?G(p))" for S by auto moreover from calculation have "M(S) \ i \ S \ M({x \ G . i \ x})" for i S by (auto dest: transM) moreover have "M(S) \ countable_rel(M,S) \ countable_rel(M,{A \ G . S \ A \ 0})" for S proof - from \M(G)\ interpret M_replacement_lepoll M "\_ x. Collect(G, (\)(x))" using countable_lepoll_assms lam_replacement_inj_rel separation_in_rev - lam_replacement_Collect[OF _ _ insnd_ball] mem_F_bound6[of _ G] - lam_replacement_minimum - by unfold_locales + lam_replacement_Collect'[of G "\x y. x\y"] mem_F_bound6[of _ G] + lam_replacement_minimum separation_in lam_replacement_fst lam_replacement_snd + by unfold_locales (auto dest:transM intro:lam_Least_assumption_general[of _ _ _ _ Union]) fix S assume "M(S)" with \M(G)\ \\i. M(S) \ i \ S \ M({x \ G . i \ x})\ interpret M_cardinal_UN_lepoll _ ?G S using lepoll_assumptions by unfold_locales (auto dest:transM) assume "countable_rel(M,S)" with \M(S)\ calculation(6) calculation(7,8)[of S] show "countable_rel(M,{A \ G . S \ A \ 0})" using InfCard_rel_nat Card_rel_nat - le_Card_rel_iff[THEN iffD2, THEN [3] leqpoll_rel_imp_cardinal_rel_UN_le, + le_Card_rel_iff[THEN iffD2, THEN [3] lepoll_rel_imp_cardinal_rel_UN_le, THEN [4] le_Card_rel_iff[THEN iffD1], of \] j.UN_closed unfolding countable_rel_def by (auto dest: transM) qed define Disjoint where "Disjoint = { \ G\G . B \ A = 0}" have "Disjoint = {x \ G\G . \ a b. x= \ a\b=0}" unfolding Disjoint_def by force with \M(G)\ have "M(Disjoint)" using disjoint_separation by simp text\For every countable\_rel subfamily of \<^term>\G\ there is another some element disjoint from all of them:\ have "\A\G. \S\X. \Disjoint" if "|X|\<^bsup>M\<^esup> < \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "X \ G" "M(X)" for X proof - note \n\\\ \M(G)\ moreover from this and \\A. A\G \ |A|\<^bsup>M\<^esup> = succ(n)\ have "|A|\<^bsup>M\<^esup>= succ(n)" "M(A)" if "A\G" for A using that Finite_cardinal_rel_eq_cardinal[of A] Finite_cardinal_rel_iff'[of A] nat_into_Finite transM[of A G] by (auto dest:transM) ultimately have "A\G \ Finite(A)" for A using cardinal_rel_Card_rel_eqpoll_rel_iff[of "succ(n)" A] Finite_cardinal_rel_eq_cardinal[of A] nat_into_Card_rel[of "succ(n)"] nat_into_M[of n] unfolding Finite_def eqpoll_rel_def by (auto) with \X\G\ \M(X)\ have "A\X \ countable_rel(M,A)" for A using Finite_imp_countable_rel by (auto dest: transM) moreover from \M(X)\ have "M(\X)" by simp moreover note \|X|\<^bsup>M\<^esup> < \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ \M(X)\ ultimately have "countable_rel(M,\X)" using Card_rel_nat[THEN cardinal_rel_lt_csucc_rel_iff, of X] countable_rel_union_countable_rel[of X] countable_rel_iff_cardinal_rel_le_nat[of X] Aleph_rel_succ Aleph_rel_zero by simp with \M(\X)\ \M(_) \ countable_rel(M,_) \ countable_rel(M,{A \ G . _ \ A \ 0})\ have "countable_rel(M,{A \ G . (\X) \ A \ 0})" by simp with \G \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ \M(G)\ obtain B where "B\G" "B \ {A \ G . (\X) \ A \ 0}" using nat_lt_Aleph_rel1 cardinal_rel_Card_rel_eqpoll_rel_iff[of "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" G] uncountable_rel_not_subset_countable_rel [of "{A \ G . (\X) \ A \ 0}" G] uncountable_rel_iff_nat_lt_cardinal_rel[of G] by force then have "\A\G. \S\X. A \ S = 0" by auto with \X\G\ show "\A\G. \S\X. \Disjoint" unfolding Disjoint_def using subsetD by simp qed moreover from \G \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>\ \M(G)\ obtain b where "b\G" using uncountable_rel_iff_subset_eqpoll_rel_Aleph_rel1 uncountable_rel_not_empty by blast ultimately text\Hence, the hypotheses to perform a bounded-cardinal selection are satisfied,\ obtain S where "S:\\<^bsub>1\<^esub>\<^bsup>M\<^esup>\\<^bsup>M\<^esup>G" "\\\\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ \\\\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ \<\ \ , S`\> \Disjoint" for \ \ using bounded_cardinal_rel_selection[of "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" G Disjoint] \M(Disjoint)\ by force moreover from this \n\\\ \M(G)\ have inM:"M(S)" "M(n)" "\x. x \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ S ` x \ G" "\x. x \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ M(x)" using function_space_rel_char by (auto dest: transM) ultimately have "\ \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ \ \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ \\\ \ S`\ \ S`\ = 0" for \ \ unfolding Disjoint_def using lt_neq_symmetry[of "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "\\ \. S`\ \ S`\ = 0"] Card_rel_is_Ord by auto (blast) text\and a symmetry argument shows that obtained \<^term>\S\ is an injective \<^term>\\\<^bsub>1\<^esub>\<^bsup>M\<^esup>\-sequence of disjoint elements of \<^term>\G\.\ moreover from this and \\A. A\G \ |A|\<^bsup>M\<^esup> = succ(n)\ inM \S : \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \\<^bsup>M\<^esup> G\ \M(G)\ have "S \ inj_rel(M,\\<^bsub>1\<^esub>\<^bsup>M\<^esup>, G)" using def_inj_rel[OF Aleph_rel_closed \M(G)\, of 1] proof (clarsimp) fix w x from inM have "a \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ b \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup> \ a \ b \ S ` a \ S ` b" for a b using \\A. A\G \ |A|\<^bsup>M\<^esup> = succ(n)\[THEN [4] cardinal_rel_succ_not_0[THEN [4] Int_eq_zero_imp_not_eq[OF calculation, of "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "\x. x"], of "\_.n"], OF _ _ _ _ apply_closed] by auto moreover assume "w \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "x \ \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "S ` w = S ` x" ultimately show "w = x" by blast qed moreover from this \M(G)\ have "range(S) \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" using inj_rel_bij_rel_range eqpoll_rel_sym unfolding eqpoll_rel_def by (blast dest: transM) moreover note \M(G)\ moreover from calculation have "range(S) \ G" using inj_rel_is_fun range_fun_subset_codomain by (fastforce dest: transM) moreover note \M(S)\ ultimately show "\D[M]. D \ G \ delta_system(D) \ D \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" using inj_rel_is_fun ZF_Library.range_eq_image[of S "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" G] image_function[OF fun_is_function, OF inj_rel_is_fun, of S "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" G] domain_of_fun[OF inj_rel_is_fun, of S "\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" G] apply_replacement[of S] by (rule_tac x="S``\\<^bsub>1\<^esub>\<^bsup>M\<^esup>" in rexI) (auto dest:transM intro!:RepFun_closed) text\This finishes the successor case and hence the proof.\ qed qed with \G \ F\ show ?thesis by blast qed lemma delta_system_uncountable_rel: assumes "\A\F. Finite(A)" "uncountable_rel(M,F)" "M(F)" shows "\D[M]. D \ F \ delta_system(D) \ D \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" proof - from assms obtain S where "S \ F" "S \\<^bsup>M\<^esup> \\<^bsub>1\<^esub>\<^bsup>M\<^esup>" "M(S)" using uncountable_rel_iff_subset_eqpoll_rel_Aleph_rel1[of F] by auto moreover from \\A\F. Finite(A)\ and this have "\A\S. Finite(A)" by auto ultimately show ?thesis using delta_system_Aleph_rel1[of S] by (auto dest:transM) qed end \ \\<^locale>\M_delta\\ end \ No newline at end of file diff --git a/thys/Transitive_Models/Discipline_Base.thy b/thys/Transitive_Models/Discipline_Base.thy --- a/thys/Transitive_Models/Discipline_Base.thy +++ b/thys/Transitive_Models/Discipline_Base.thy @@ -1,659 +1,638 @@ theory Discipline_Base imports "ZF-Constructible.Rank" M_Basic_No_Repl Relativization ZF_Miscellanea begin hide_const (open) Order.pred declare [[syntax_ambiguity_warning = false]] subsection\Discipline of relativization of basic concepts\ definition is_singleton :: "[i\o,i,i] \ o" where "is_singleton(A,x,z) \ \c[A]. empty(A,c) \ is_cons(A,x,c,z)" lemma (in M_trivial) singleton_abs[simp] : "\ M(x) ; M(s) \ \ is_singleton(M,x,s) \ s = {x}" unfolding is_singleton_def using nonempty by simp synthesize "singleton" from_definition "is_singleton" notation singleton_fm (\\{_} is _\\) -(* TODO: check if the following lemmas should be here or not? *) -lemma (in M_trivial) singleton_closed [simp]: - "M(x) \ M({x})" - by simp +lemmas (in M_trivial) singleton_closed = singleton_in_MI lemma (in M_trivial) Upair_closed[simp]: "M(a) \ M(b) \ M(Upair(a,b))" using Upair_eq_cons by simp text\The following named theorems gather instances of transitivity that arise from closure theorems\ named_theorems trans_closed definition is_hcomp :: "[i\o,i\i\o,i\i\o,i,i] \ o" where "is_hcomp(M,is_f,is_g,a,w) \ \z[M]. is_g(a,z) \ is_f(z,w)" lemma (in M_trivial) is_hcomp_abs: assumes is_f_abs:"\a z. M(a) \ M(z) \ is_f(a,z) \ z = f(a)" and is_g_abs:"\a z. M(a) \ M(z) \ is_g(a,z) \ z = g(a)" and g_closed:"\a. M(a) \ M(g(a))" "M(a)" "M(w)" shows "is_hcomp(M,is_f,is_g,a,w) \ w = f(g(a))" unfolding is_hcomp_def using assms by simp definition hcomp_fm :: "[i\i\i,i\i\i,i,i] \ i" where "hcomp_fm(pf,pg,a,w) \ Exists(And(pg(succ(a),0),pf(0,succ(w))))" lemma sats_hcomp_fm: assumes f_iff_sats:"\a b z. a\nat \ b\nat \ z\M \ is_f(nth(a,Cons(z,env)),nth(b,Cons(z,env))) \ sats(M,pf(a,b),Cons(z,env))" and g_iff_sats:"\a b z. a\nat \ b\nat \ z\M \ is_g(nth(a,Cons(z,env)),nth(b,Cons(z,env))) \ sats(M,pg(a,b),Cons(z,env))" and "a\nat" "w\nat" "env\list(M)" shows "sats(M,hcomp_fm(pf,pg,a,w),env) \ is_hcomp(##M,is_f,is_g,nth(a,env),nth(w,env))" proof - have "sats(M, pf(0, succ(w)), Cons(x, env)) \ is_f(x,nth(w,env))" if "x\M" "w\nat" for x w using f_iff_sats[of 0 "succ(w)" x] that by simp moreover have "sats(M, pg(succ(a), 0), Cons(x, env)) \ is_g(nth(a,env),x)" if "x\M" "a\nat" for x a using g_iff_sats[of "succ(a)" 0 x] that by simp ultimately show ?thesis unfolding hcomp_fm_def is_hcomp_def using assms by simp qed definition hcomp_r :: "[i\o,[i\o,i,i]\o,[i\o,i,i]\o,i,i] \ o" where "hcomp_r(M,is_f,is_g,a,w) \ \z[M]. is_g(M,a,z) \ is_f(M,z,w)" definition is_hcomp2_2 :: "[i\o,[i\o,i,i,i]\o,[i\o,i,i,i]\o,[i\o,i,i,i]\o,i,i,i] \ o" where "is_hcomp2_2(M,is_f,is_g1,is_g2,a,b,w) \ \g1ab[M]. \g2ab[M]. is_g1(M,a,b,g1ab) \ is_g2(M,a,b,g2ab) \ is_f(M,g1ab,g2ab,w)" lemma (in M_trivial) hcomp_abs: assumes is_f_abs:"\a z. M(a) \ M(z) \ is_f(M,a,z) \ z = f(a)" and is_g_abs:"\a z. M(a) \ M(z) \ is_g(M,a,z) \ z = g(a)" and g_closed:"\a. M(a) \ M(g(a))" "M(a)" "M(w)" shows "hcomp_r(M,is_f,is_g,a,w) \ w = f(g(a))" unfolding hcomp_r_def using assms by simp lemma hcomp_uniqueness: assumes uniq_is_f: "\r d d'. M(r) \ M(d) \ M(d') \ is_f(M, r, d) \ is_f(M, r, d') \ d = d'" and uniq_is_g: "\r d d'. M(r) \ M(d) \ M(d') \ is_g(M, r, d) \ is_g(M, r, d') \ d = d'" and "M(a)" "M(w)" "M(w')" "hcomp_r(M,is_f,is_g,a,w)" "hcomp_r(M,is_f,is_g,a,w')" shows "w=w'" proof - from assms obtain z z' where "is_g(M, a, z)" "is_g(M, a, z')" "is_f(M,z,w)" "is_f(M,z',w')" "M(z)" "M(z')" unfolding hcomp_r_def by blast moreover from this and uniq_is_g and \M(a)\ have "z=z'" by blast moreover note uniq_is_f and \M(w)\ \M(w')\ ultimately show ?thesis by blast qed lemma hcomp_witness: assumes wit_is_f: "\r. M(r) \ \d[M]. is_f(M,r,d)" and wit_is_g: "\r. M(r) \ \d[M]. is_g(M,r,d)" and "M(a)" shows "\w[M]. hcomp_r(M,is_f,is_g,a,w)" proof - from \M(a)\ and wit_is_g obtain z where "is_g(M,a,z)" "M(z)" by blast moreover from this and wit_is_f obtain w where "is_f(M,z,w)" "M(w)" by blast ultimately show ?thesis using assms unfolding hcomp_r_def by auto qed +\ \In the next lemma, the absoluteness hypotheses on \<^term>\g1\ and \<^term>\g2\ can be restricted +to the actual parameters.\ lemma (in M_trivial) hcomp2_2_abs: assumes is_f_abs:"\r1 r2 z. M(r1) \ M(r2) \ M(z) \ is_f(M,r1,r2,z) \ z = f(r1,r2)" and is_g1_abs:"\r1 r2 z. M(r1) \ M(r2) \ M(z) \ is_g1(M,r1,r2,z) \ z = g1(r1,r2)" and is_g2_abs:"\r1 r2 z. M(r1) \ M(r2) \ M(z) \ is_g2(M,r1,r2,z) \ z = g2(r1,r2)" and types: "M(a)" "M(b)" "M(w)" "M(g1(a,b))" "M(g2(a,b))" shows "is_hcomp2_2(M,is_f,is_g1,is_g2,a,b,w) \ w = f(g1(a,b),g2(a,b))" - unfolding is_hcomp2_2_def using assms - \ \We only need some particular cases of the abs assumptions\ - (* is_f_abs types is_g1_abs[of a b] is_g2_abs[of a b] *) + unfolding is_hcomp2_2_def + using assms by simp lemma hcomp2_2_uniqueness: assumes uniq_is_f: "\r1 r2 d d'. M(r1) \ M(r2) \ M(d) \ M(d') \ is_f(M, r1, r2 , d) \ is_f(M, r1, r2, d') \ d = d'" and uniq_is_g1: "\r1 r2 d d'. M(r1) \ M(r2)\ M(d) \ M(d') \ is_g1(M, r1,r2, d) \ is_g1(M, r1,r2, d') \ d = d'" and uniq_is_g2: "\r1 r2 d d'. M(r1) \ M(r2)\ M(d) \ M(d') \ is_g2(M, r1,r2, d) \ is_g2(M, r1,r2, d') \ d = d'" and "M(a)" "M(b)" "M(w)" "M(w')" "is_hcomp2_2(M,is_f,is_g1,is_g2,a,b,w)" "is_hcomp2_2(M,is_f,is_g1,is_g2,a,b,w')" shows "w=w'" proof - from assms obtain z z' y y' where "is_g1(M, a,b, z)" "is_g1(M, a,b, z')" "is_g2(M, a,b, y)" "is_g2(M, a,b, y')" "is_f(M,z,y,w)" "is_f(M,z',y',w')" "M(z)" "M(z')" "M(y)" "M(y')" unfolding is_hcomp2_2_def by force moreover from this and uniq_is_g1 uniq_is_g2 and \M(a)\ \M(b)\ have "z=z'" "y=y'" by blast+ moreover note uniq_is_f and \M(w)\ \M(w')\ ultimately show ?thesis by blast qed lemma hcomp2_2_witness: assumes wit_is_f: "\r1 r2. M(r1) \ M(r2) \ \d[M]. is_f(M,r1,r2,d)" and wit_is_g1: "\r1 r2. M(r1) \ M(r2) \ \d[M]. is_g1(M,r1,r2,d)" and wit_is_g2: "\r1 r2. M(r1) \ M(r2) \ \d[M]. is_g2(M,r1,r2,d)" and "M(a)" "M(b)" shows "\w[M]. is_hcomp2_2(M,is_f,is_g1,is_g2,a,b,w)" proof - from \M(a)\ \M(b)\ and wit_is_g1 obtain g1a where "is_g1(M,a,b,g1a)" "M(g1a)" by blast moreover from \M(a)\ \M(b)\ and wit_is_g2 obtain g2a where "is_g2(M,a,b,g2a)" "M(g2a)" by blast moreover from calculation and wit_is_f obtain w where "is_f(M,g1a,g2a,w)" "M(w)" by blast ultimately show ?thesis using assms unfolding is_hcomp2_2_def by auto qed lemma (in M_trivial) extensionality_trans: assumes "M(d) \ (\x[M]. x\d \ P(x))" "M(d') \ (\x[M]. x\d' \ P(x))" shows "d=d'" proof - from assms have "\x. x\d \ P(x) \ M(x)" using transM[of _ d] by auto moreover from assms have "\x. x\d' \ P(x) \ M(x)" using transM[of _ d'] by auto ultimately show ?thesis by auto qed definition lt_rel :: "[i\o,i,i] \ o" where "lt_rel(M,a,b) \ a\b \ ordinal(M,b)" lemma (in M_trans) lt_abs[absolut]: "M(a) \ M(b) \ lt_rel(M,a,b) \ ao,i,i] \ o" where "le_rel(M,a,b) \ \sb[M]. successor(M,b,sb) \ lt_rel(M,a,sb)" lemma (in M_trivial) le_abs[absolut]: "M(a) \ M(b) \ le_rel(M,a,b) \ a\b" unfolding le_rel_def by (simp add:absolut) subsection\Discipline for \<^term>\Pow\\ definition is_Pow :: "[i\o,i,i] \ o" where "is_Pow(M,A,z) \ M(z) \ (\x[M]. x \ z \ subset(M,x,A))" definition Pow_rel :: "[i\o,i] \ i" (\Pow\<^bsup>_\<^esup>'(_')\) where "Pow_rel(M,r) \ THE d. is_Pow(M,r,d)" abbreviation Pow_r_set :: "[i,i] \ i" (\Pow\<^bsup>_\<^esup>'(_')\) where "Pow_r_set(M) \ Pow_rel(##M)" context M_basic_no_repl begin lemma is_Pow_uniqueness: assumes "M(r)" "is_Pow(M,r,d)" "is_Pow(M,r,d')" shows "d=d'" using assms extensionality_trans unfolding is_Pow_def by simp lemma is_Pow_witness: "M(r) \ \d[M]. is_Pow(M,r,d)" using power_ax unfolding power_ax_def powerset_def is_Pow_def by simp \ \We have to do this by hand, using axioms\ lemma is_Pow_closed : "\ M(r);is_Pow(M,r,d) \ \ M(d)" unfolding is_Pow_def by simp lemma Pow_rel_closed[intro,simp]: "M(r) \ M(Pow_rel(M,r))" unfolding Pow_rel_def using is_Pow_closed theI[OF ex1I[of "\d. is_Pow(M,r,d)"], OF _ is_Pow_uniqueness[of r]] is_Pow_witness by fastforce lemmas trans_Pow_rel_closed[trans_closed] = transM[OF _ Pow_rel_closed] text\The proof of \<^term>\f_rel_iff\ lemma is schematic and it can reused by copy-paste replacing appropriately.\ lemma Pow_rel_iff: assumes "M(r)" "M(d)" shows "is_Pow(M,r,d) \ d = Pow_rel(M,r)" proof (intro iffI) assume "d = Pow_rel(M,r)" with assms show "is_Pow(M, r, d)" using is_Pow_uniqueness[of r] is_Pow_witness theI[OF ex1I[of "\d. is_Pow(M,r,d)"], OF _ is_Pow_uniqueness[of r]] unfolding Pow_rel_def by auto next assume "is_Pow(M, r, d)" with assms show "d = Pow_rel(M,r)" using is_Pow_uniqueness unfolding Pow_rel_def by (auto del:the_equality intro:the_equality[symmetric]) qed text\The next "def\_" result really corresponds to @{thm Pow_iff}\ lemma def_Pow_rel: "M(A) \ M(r) \ A\Pow_rel(M,r) \ A \ r" using Pow_rel_iff[OF _ Pow_rel_closed, of r r] unfolding is_Pow_def by simp lemma Pow_rel_char: "M(r) \ Pow_rel(M,r) = {A\Pow(r). M(A)}" proof - assume "M(r)" moreover from this have "x \ Pow_rel(M,r) \ x\r" "M(x) \ x \ r \ x \ Pow_rel(M,r)" for x using def_Pow_rel by (auto intro!:trans_Pow_rel_closed) ultimately show ?thesis using trans_Pow_rel_closed by blast qed lemma mem_Pow_rel_abs: "M(a) \ M(r) \ a \ Pow_rel(M,r) \ a \ Pow(r)" using Pow_rel_char by simp end \ \\<^locale>\M_basic_no_repl\\ -(****************** end Discipline **********************) - - -(**********************************************************) subsection\Discipline for \<^term>\PiP\\ definition PiP_rel:: "[i\o,i,i]\o" where "PiP_rel(M,A,f) \ \df[M]. is_domain(M,f,df) \ subset(M,A,df) \ is_function(M,f)" context M_basic begin lemma def_PiP_rel: assumes "M(A)" "M(f)" shows "PiP_rel(M,A,f) \ A \ domain(f) \ function(f)" using assms unfolding PiP_rel_def by simp end \ \\<^locale>\M_basic\\ -(****************** end Discipline **********************) - -(* -Sigma(A,B) == \x\A. \y\B(x). {\x,y\} - == \ { (\y\B(x). {\x,y\}) . x\A} - == \ { (\y\B(x). {\x,y\}) . x\A} - == \ { ( \ { {\x,y\} . y\B(x)} ) . x\A} - ---------------------- - Sigfun(x,B) -*) - -definition \ \FIX THIS: not completely relational. Can it be?\ +definition Sigfun :: "[i,i\i]\i" where "Sigfun(x,B) \ \y\B(x). {\x,y\}" +lemma SigFun_char : "Sigfun(x,B) = {x}\B(x)" + unfolding Sigfun_def by auto + lemma Sigma_Sigfun: "Sigma(A,B) = \ {Sigfun(x,B) . x\A}" unfolding Sigma_def Sigfun_def .. -definition \ \FIX THIS: not completely relational. Can it be?\ +definition \ \Note that $B$ below is a higher order argument\ is_Sigfun :: "[i\o,i,i\i,i]\o" where "is_Sigfun(M,x,B,Sd) \ M(Sd) \ (\RB[M]. is_Replace(M,B(x),\y z. z={\x,y\},RB) \ big_union(M,RB,Sd))" - context M_trivial begin lemma is_Sigfun_abs: assumes "strong_replacement(M,\y z. z={\x,y\})" "M(x)" "M(B(x))" "M(Sd)" shows "is_Sigfun(M,x,B,Sd) \ Sd = Sigfun(x,B)" proof - have "\{z . y \ B(x), z = {\x, y\}} = (\y\B(x). {\x, y\})" by auto then show ?thesis using assms transM[OF _ \M(B(x))\] Replace_abs unfolding is_Sigfun_def Sigfun_def by auto qed lemma Sigfun_closed: assumes "strong_replacement(M, \y z. y \ B(x) \ z = {\x, y\})" "M(x)" "M(B(x))" shows "M(Sigfun(x,B))" using assms transM[OF _ \M(B(x))\] RepFun_closed2 unfolding Sigfun_def by simp lemmas trans_Sigfun_closed[trans_closed] = transM[OF _ Sigfun_closed] end \ \\<^locale>\M_trivial\\ definition is_Sigma :: "[i\o,i,i\i,i]\o" where "is_Sigma(M,A,B,S) \ M(S) \ (\RSf[M]. is_Replace(M,A,\x z. z=Sigfun(x,B),RSf) \ big_union(M,RSf,S))" locale M_Pi = M_basic + assumes Pi_separation: "M(A) \ separation(M, PiP_rel(M,A))" and Pi_replacement: "M(x) \ M(y) \ strong_replacement(M, \ya z. ya \ y \ z = {\x, ya\})" "M(y) \ strong_replacement(M, \x z. z = (\xa\y. {\x, xa\}))" locale M_Pi_assumptions = M_Pi + fixes A B assumes Pi_assumptions: "M(A)" "\x. x\A \ M(B(x))" "\x\A. strong_replacement(M, \y z. y \ B(x) \ z = {\x, y\})" "strong_replacement(M,\x z. z=Sigfun(x,B))" begin lemma Sigma_abs[simp]: assumes "M(S)" shows "is_Sigma(M,A,B,S) \ S = Sigma(A,B)" proof - have "\{z . x \ A, z = Sigfun(x, B)} = (\x\A. Sigfun(x, B))" by auto with assms show ?thesis using Replace_abs[of A _ "\x z. z=Sigfun(x,B)"] - Sigfun_closed Sigma_Sigfun[of A B] transM[of _ A] + Sigfun_closed Sigma_Sigfun transM[of _ A] Pi_assumptions is_Sigfun_abs unfolding is_Sigma_def by simp qed lemma Sigma_closed[intro,simp]: "M(Sigma(A,B))" proof - have "(\x\A. Sigfun(x, B)) = \{z . x \ A, z = Sigfun(x, B)}" by auto then show ?thesis - using Sigma_Sigfun[of A B] transM[of _ A] - Sigfun_closed Pi_assumptions + using Sigma_Sigfun[of A B] transM[of _ A] Sigfun_closed Pi_assumptions by simp qed lemmas trans_Sigma_closed[trans_closed] = transM[OF _ Sigma_closed] end \ \\<^locale>\M_Pi_assumptions\\ -(**********************************************************) subsection\Discipline for \<^term>\Pi\\ -definition (* completely relational *) +definition \ \completely relational\ is_Pi :: "[i\o,i,i\i,i]\o" where "is_Pi(M,A,B,I) \ M(I) \ (\S[M]. \PS[M]. is_Sigma(M,A,B,S) \ is_Pow(M,S,PS) \ is_Collect(M,PS,PiP_rel(M,A),I))" definition Pi_rel :: "[i\o,i,i\i] \ i" (\Pi\<^bsup>_\<^esup>'(_,_')\) where "Pi_rel(M,A,B) \ THE d. is_Pi(M,A,B,d)" abbreviation Pi_r_set :: "[i,i,i\i] \ i" (\Pi\<^bsup>_\<^esup>'(_,_')\) where "Pi_r_set(M,A,B) \ Pi_rel(##M,A,B)" - context M_basic begin lemmas Pow_rel_iff = mbnr.Pow_rel_iff lemmas Pow_rel_char = mbnr.Pow_rel_char lemmas mem_Pow_rel_abs = mbnr.mem_Pow_rel_abs lemmas Pow_rel_closed = mbnr.Pow_rel_closed lemmas def_Pow_rel = mbnr.def_Pow_rel lemmas trans_Pow_rel_closed = mbnr.trans_Pow_rel_closed end \ \\<^locale>\M_basic\\ context M_Pi_assumptions begin lemma is_Pi_uniqueness: assumes "is_Pi(M,A,B,d)" "is_Pi(M,A,B,d')" shows "d=d'" using assms Pi_assumptions extensionality_trans Pow_rel_iff unfolding is_Pi_def by simp lemma is_Pi_witness: "\d[M]. is_Pi(M,A,B,d)" using Pow_rel_iff Pi_separation Pi_assumptions unfolding is_Pi_def by simp lemma is_Pi_closed : "is_Pi(M,A,B,d) \ M(d)" unfolding is_Pi_def by simp lemma Pi_rel_closed[intro,simp]: "M(Pi_rel(M,A,B))" proof - have "is_Pi(M, A, B, THE xa. is_Pi(M, A, B, xa))" using Pi_assumptions theI[OF ex1I[of "is_Pi(M,A,B)"], OF _ is_Pi_uniqueness] is_Pi_witness is_Pi_closed by auto then show ?thesis using is_Pi_closed unfolding Pi_rel_def by simp qed -\ \From this point on, the higher order variable \<^term>\y\ must be +\ \From this point on, the higher order variable \<^term>\B\ must be explicitly instantiated, and proof methods are slower\ lemmas trans_Pi_rel_closed[trans_closed] = transM[OF _ Pi_rel_closed] lemma Pi_rel_iff: assumes "M(d)" shows "is_Pi(M,A,B,d) \ d = Pi_rel(M,A,B)" proof (intro iffI) assume "d = Pi_rel(M,A,B)" moreover note assms moreover from this obtain e where "M(e)" "is_Pi(M,A,B,e)" using is_Pi_witness by blast ultimately show "is_Pi(M, A, B, d)" using is_Pi_uniqueness is_Pi_witness is_Pi_closed theI[OF ex1I[of "is_Pi(M,A,B)"], OF _ is_Pi_uniqueness, of e] unfolding Pi_rel_def by simp next assume "is_Pi(M, A, B, d)" with assms show "d = Pi_rel(M,A,B)" using is_Pi_uniqueness is_Pi_closed unfolding Pi_rel_def by (blast del:the_equality intro:the_equality[symmetric]) qed lemma def_Pi_rel: "Pi_rel(M,A,B) = {f\Pow_rel(M,Sigma(A,B)). A\domain(f) \ function(f)}" proof - have "Pi_rel(M,A, B) \ Pow_rel(M,Sigma(A,B))" using Pi_assumptions Pi_rel_iff[of "Pi_rel(M,A,B)"] Pow_rel_iff unfolding is_Pi_def by auto moreover have "f \ Pi_rel(M,A, B) \ A\domain(f) \ function(f)" for f using Pi_assumptions Pi_rel_iff[of "Pi_rel(M,A,B)"] def_PiP_rel[of A f] trans_closed Pow_rel_iff unfolding is_Pi_def by simp moreover have "f \ Pow_rel(M,Sigma(A,B)) \ A\domain(f) \ function(f) \ f \ Pi_rel(M,A, B)" for f using Pi_rel_iff[of "Pi_rel(M,A,B)"] Pi_assumptions def_PiP_rel[of A f] trans_closed Pow_rel_iff unfolding is_Pi_def by simp ultimately show ?thesis by force qed lemma Pi_rel_char: "Pi_rel(M,A,B) = {f\Pi(A,B). M(f)}" using Pi_assumptions def_Pi_rel Pow_rel_char[OF Sigma_closed] unfolding Pi_def by fastforce lemma mem_Pi_rel_abs: assumes "M(f)" shows "f \ Pi_rel(M,A,B) \ f \ Pi(A,B)" using assms Pi_rel_char by simp end \ \\<^locale>\M_Pi_assumptions\\ text\The next locale (and similar ones below) are used to show the relationship between versions of simple (i.e. $\Sigma_1^{\mathit{ZF}}$, $\Pi_1^{\mathit{ZF}}$) concepts in two different transitive models.\ locale M_N_Pi_assumptions = M:M_Pi_assumptions + N:M_Pi_assumptions N for N + assumes M_imp_N:"M(x) \ N(x)" begin lemma Pi_rel_transfer: "Pi\<^bsup>M\<^esup>(A,B) \ Pi\<^bsup>N\<^esup>(A,B)" using M.Pi_rel_char N.Pi_rel_char M_imp_N by auto end \ \\<^locale>\M_N_Pi_assumptions\\ -(****************** end Discipline **********************) - locale M_Pi_assumptions_0 = M_Pi_assumptions _ 0 begin text\This is used in the proof of \<^term>\AC_Pi_rel\\ lemma Pi_rel_empty1[simp]: "Pi\<^bsup>M\<^esup>(0,B) = {0}" using Pi_assumptions Pow_rel_char by (unfold def_Pi_rel function_def) (auto) end \ \\<^locale>\M_Pi_assumptions_0\\ context M_Pi_assumptions begin subsection\Auxiliary ported results on \<^term>\Pi_rel\, now unused\ lemma Pi_rel_iff': assumes types:"M(f)" shows "f \ Pi_rel(M,A,B) \ function(f) \ f \ Sigma(A,B) \ A \ domain(f)" using assms Pow_rel_char by (simp add:def_Pi_rel, blast) lemma lam_type_M: assumes "M(A)" "\x. x\A \ M(B(x))" "\x. x \ A \ b(x)\B(x)" "strong_replacement(M,\x y. y=\x, b(x)\) " shows "(\x\A. b(x)) \ Pi_rel(M,A,B)" proof (auto simp add: lam_def def_Pi_rel function_def) from assms have "M({\x, b(x)\ . x \ A})" using Pi_assumptions transM[OF _ \M(A)\] by (rule_tac RepFun_closed, auto intro!:transM[OF _ \\x. x\A \ M(B(x))\]) with assms show "{\x, b(x)\ . x \ A} \ Pow\<^bsup>M\<^esup>(Sigma(A, B))" using Pow_rel_char by auto qed end \ \\<^locale>\M_Pi_assumptions\\ locale M_Pi_assumptions2 = M_Pi_assumptions + PiC: M_Pi_assumptions _ _ C for C begin lemma Pi_rel_type: assumes "f \ Pi\<^bsup>M\<^esup>(A,C)" "\x. x \ A \ f`x \ B(x)" and types: "M(f)" shows "f \ Pi\<^bsup>M\<^esup>(A,B)" using assms Pi_assumptions by (simp only: Pi_rel_iff' PiC.Pi_rel_iff') (blast dest: function_apply_equality) lemma Pi_rel_weaken_type: assumes "f \ Pi\<^bsup>M\<^esup>(A,B)" "\x. x \ A \ B(x) \ C(x)" and types: "M(f)" shows "f \ Pi\<^bsup>M\<^esup>(A,C)" using assms Pi_assumptions by (simp only: Pi_rel_iff' PiC.Pi_rel_iff') (blast intro: Pi_rel_type dest: apply_type) end \ \\<^locale>\M_Pi_assumptions2\\ - end \ No newline at end of file diff --git a/thys/Transitive_Models/Discipline_Cardinal.thy b/thys/Transitive_Models/Discipline_Cardinal.thy --- a/thys/Transitive_Models/Discipline_Cardinal.thy +++ b/thys/Transitive_Models/Discipline_Cardinal.thy @@ -1,173 +1,179 @@ +section\Basic relativization of cardinality\ + theory Discipline_Cardinal imports Discipline_Function begin relativize functional "cardinal" "cardinal_rel" external relationalize "cardinal_rel" "is_cardinal" synthesize "is_cardinal" from_definition assuming "nonempty" notation is_cardinal_fm (\cardinal'(_') is _\) abbreviation cardinal_r :: "[i,i\o] \ i" (\|_|\<^bsup>_\<^esup>\) where "|x|\<^bsup>M\<^esup> \ cardinal_rel(M,x)" abbreviation cardinal_r_set :: "[i,i]\i" (\|_|\<^bsup>_\<^esup>\) where "|x|\<^bsup>M\<^esup> \ cardinal_rel(##M,x)" -context M_trivial begin +context M_trivial +begin rel_closed for "cardinal" using Least_closed'[of "\i. M(i) \ i \\<^bsup>M\<^esup> A"] unfolding cardinal_rel_def by simp -end +end \ \\<^locale>\M_trivial\\ manual_arity intermediate for "is_Int_fm" unfolding is_Int_fm_def using arity pred_Un_distrib by (simp) arity_theorem for "is_Int_fm" arity_theorem for "is_funspace_fm" arity_theorem for "is_function_space_fm" arity_theorem for "surjP_rel_fm" arity_theorem intermediate for "is_surj_fm" lemma arity_is_surj_fm [arity] : "A \ nat \ B \ nat \ I \ nat \ arity(is_surj_fm(A, B, I)) = succ(A) \ succ(B) \ succ(I)" using arity_is_surj_fm' pred_Un_distrib by auto arity_theorem for "injP_rel_fm" arity_theorem intermediate for "is_inj_fm" lemma arity_is_inj_fm [arity]: "A \ nat \ B \ nat \ I \ nat \ arity(is_inj_fm(A, B, I)) = succ(A) \ succ(B) \ succ(I)" using arity_is_inj_fm' pred_Un_distrib by auto arity_theorem for "is_bij_fm" arity_theorem for "is_eqpoll_fm" arity_theorem for "is_cardinal_fm" -context M_Perm begin - +context M_Perm +begin is_iff_rel for "cardinal" using least_abs'[of "\i. M(i) \ i \\<^bsup>M\<^esup> A"] is_eqpoll_iff unfolding is_cardinal_def cardinal_rel_def by simp -end +end \ \\<^locale>\M_Perm\\ reldb_add functional "Ord" "Ord" reldb_add relational "Ord" "ordinal" reldb_add functional "lt" "lt" reldb_add relational "lt" "lt_rel" synthesize "lt_rel" from_definition notation lt_rel_fm (\\_ < _\\) arity_theorem intermediate for "lt_rel_fm" lemma arity_lt_rel_fm[arity]: "a \ nat \ b \ nat \ arity(lt_rel_fm(a, b)) = succ(a) \ succ(b)" using arity_lt_rel_fm' by auto relativize functional "Card" "Card_rel" external relationalize "Card_rel" "is_Card" synthesize "is_Card" from_definition assuming "nonempty" notation is_Card_fm (\\Card'(_')\\) arity_theorem for "is_Card_fm" notation Card_rel (\Card\<^bsup>_\<^esup>'(_')\) lemma (in M_Perm) is_Card_iff: "M(A) \ is_Card(M, A) \ Card\<^bsup>M\<^esup>(A)" using is_cardinal_iff unfolding is_Card_def Card_rel_def by simp abbreviation Card_r_set :: "[i,i]\o" (\Card\<^bsup>_\<^esup>'(_')\) where "Card\<^bsup>M\<^esup>(i) \ Card_rel(##M,i)" relativize functional "InfCard" "InfCard_rel" external relationalize "InfCard_rel" "is_InfCard" synthesize "is_InfCard" from_definition assuming "nonempty" notation is_InfCard_fm (\\InfCard'(_')\\) arity_theorem for "is_InfCard_fm" notation InfCard_rel (\InfCard\<^bsup>_\<^esup>'(_')\) abbreviation InfCard_r_set :: "[i,i]\o" (\InfCard\<^bsup>_\<^esup>'(_')\) where "InfCard\<^bsup>M\<^esup>(i) \ InfCard_rel(##M,i)" +subsection\Disicpline for \<^term>\cadd\\ relativize functional "cadd" "cadd_rel" external abbreviation cadd_r :: "[i,i\o,i] \ i" (\_ \\<^bsup>_\<^esup> _\ [66,1,66] 65) where "A \\<^bsup>M\<^esup> B \ cadd_rel(M,A,B)" -context M_basic begin +context M_basic +begin rel_closed for "cadd" using cardinal_rel_closed unfolding cadd_rel_def by simp -end - -(* relativization *) +end \ \\<^locale>\M_basic\\ relationalize "cadd_rel" "is_cadd" manual_schematic for "is_cadd" assuming "nonempty" unfolding is_cadd_def by (rule iff_sats sum_iff_sats | simp)+ synthesize "is_cadd" from_schematic arity_theorem for "sum_fm" arity_theorem for "is_cadd_fm" -context M_Perm begin +context M_Perm +begin is_iff_rel for "cadd" using is_cardinal_iff unfolding is_cadd_def cadd_rel_def by simp -end + +end \ \\<^locale>\M_Perm\\ + +subsection\Disicpline for \<^term>\cmult\\ relativize functional "cmult" "cmult_rel" external abbreviation cmult_r :: "[i,i\o,i] \ i" (\_ \\<^bsup>_\<^esup> _\ [66,1,66] 65) where "A \\<^bsup>M\<^esup> B \ cmult_rel(M,A,B)" -(* relativization *) relationalize "cmult_rel" "is_cmult" declare cartprod_iff_sats [iff_sats] synthesize "is_cmult" from_definition assuming "nonempty" arity_theorem for "is_cmult_fm" -context M_Perm begin - +context M_Perm +begin rel_closed for "cmult" using cardinal_rel_closed unfolding cmult_rel_def by simp is_iff_rel for "cmult" using is_cardinal_iff unfolding is_cmult_def cmult_rel_def by simp -end +end \ \\<^locale>\M_Perm\\ end \ No newline at end of file diff --git a/thys/Transitive_Models/Discipline_Function.thy b/thys/Transitive_Models/Discipline_Function.thy --- a/thys/Transitive_Models/Discipline_Function.thy +++ b/thys/Transitive_Models/Discipline_Function.thy @@ -1,963 +1,928 @@ +section\Basic relativization of function spaces\ + theory Discipline_Function imports Arities begin -(**********************************************************) -paragraph\Discipline for \<^term>\fst\\ +subsection\Discipline for \<^term>\fst\ and \<^term>\snd\\ - -(* ftype(p) \ THE a. \b. p = \a, b\ *) arity_theorem for "empty_fm" arity_theorem for "upair_fm" arity_theorem for "pair_fm" definition is_fst :: "(i\o)\i\i\o" where "is_fst(M,x,t) \ (\z[M]. pair(M,t,z,x)) \ (\(\z[M]. \w[M]. pair(M,w,z,x)) \ empty(M,t))" synthesize "fst" from_definition "is_fst" notation fst_fm (\\fst'(_') is _\\) arity_theorem for "fst_fm" definition fst_rel :: "[i\o,i] \ i" where "fst_rel(M,p) \ THE d. M(d) \ is_fst(M,p,d)" reldb_add relational "fst" "is_fst" reldb_add functional "fst" "fst_rel" definition is_snd :: "(i\o)\i\i\o" where "is_snd(M,x,t) \ (\z[M]. pair(M,z,t,x)) \ (\(\z[M]. \w[M]. pair(M,z,w,x)) \ empty(M,t))" synthesize "snd" from_definition "is_snd" notation snd_fm (\\snd'(_') is _\\) arity_theorem for "snd_fm" definition snd_rel :: "[i\o,i] \ i" where "snd_rel(M,p) \ THE d. M(d) \ is_snd(M,p,d)" - reldb_add relational "snd" "is_snd" reldb_add functional "snd" "snd_rel" context M_trans begin lemma fst_snd_closed: assumes "M(p)" shows "M(fst(p)) \ M(snd(p))" unfolding fst_def snd_def using assms by (cases "\a. \b. p = \a, b\";auto) lemma fst_closed[intro,simp]: "M(x) \ M(fst(x))" using fst_snd_closed by auto lemma snd_closed[intro,simp]: "M(x) \ M(snd(x))" using fst_snd_closed by auto lemma fst_abs [absolut]: "\M(p); M(x) \ \ is_fst(M,p,x) \ x = fst(p)" unfolding is_fst_def fst_def by (cases "\a. \b. p = \a, b\";auto) lemma snd_abs [absolut]: "\M(p); M(y) \ \ is_snd(M,p,y) \ y = snd(p)" unfolding is_snd_def snd_def by (cases "\a. \b. p = \a, b\";auto) lemma empty_rel_abs : "M(x) \ M(0) \ x = 0 \ x = (THE d. M(d) \ empty(M, d))" unfolding the_def using transM by auto lemma fst_rel_abs: assumes "M(p)" shows "fst_rel(M,p) = fst(p)" using fst_abs assms unfolding fst_def fst_rel_def by (cases "\a. \b. p = \a, b\";auto;rule_tac the_equality[symmetric],simp_all) lemma snd_rel_abs: assumes "M(p)" shows " snd_rel(M,p) = snd(p)" using snd_abs assms unfolding snd_def snd_rel_def by (cases "\a. \b. p = \a, b\";auto;rule_tac the_equality[symmetric],simp_all) end \ \\<^locale>\M_trans\\ +subsection\Discipline for \<^term>\minimum\\ + relativize functional "first" "first_rel" external relativize functional "minimum" "minimum_rel" external context M_trans begin lemma minimum_closed[simp,intro]: assumes "M(A)" shows "M(minimum(r,A))" using first_is_elem the_equality_if transM[OF _ \M(A)\] by(cases "\x . first(x,A,r)",auto simp:minimum_def) lemma first_abs : assumes "M(B)" shows "first(z,B,r) \ first_rel(M,z,B,r)" unfolding first_def first_rel_def using assms by auto -(* TODO: find a naming convention for absoluteness results like this. -See notes/TODO.txt -*) lemma minimum_abs: assumes "M(B)" shows "minimum(r,B) = minimum_rel(M,r,B)" proof - from assms have "first(b, B, r) \ M(b) \ first_rel(M,b,B,r)" for b using first_abs proof (auto) fix b assume "first_rel(M,b,B,r)" with \M(B)\ have "b\B" using first_abs first_is_elem by simp with \M(B)\ show "M(b)" using transM[OF \b\B\] by simp qed with assms show ?thesis unfolding minimum_rel_def minimum_def by simp qed end \ \\<^locale>\M_trans\\ subsection\Discipline for \<^term>\function_space\\ definition is_function_space :: "[i\o,i,i,i] \ o" where "is_function_space(M,A,B,fs) \ M(fs) \ is_funspace(M,A,B,fs)" definition function_space_rel :: "[i\o,i,i] \ i" where "function_space_rel(M,A,B) \ THE d. is_function_space(M,A,B,d)" reldb_rem absolute "Pi" reldb_add relational "Pi" "is_function_space" reldb_add functional "Pi" "function_space_rel" abbreviation function_space_r :: "[i,i\o,i] \ i" (\_ \\<^bsup>_\<^esup> _\ [61,1,61] 60) where "A \\<^bsup>M\<^esup> B \ function_space_rel(M,A,B)" abbreviation function_space_r_set :: "[i,i,i] \ i" (\_ \\<^bsup>_\<^esup> _\ [61,1,61] 60) where "function_space_r_set(A,M) \ function_space_rel(##M,A)" context M_Pi begin lemma is_function_space_uniqueness: assumes "M(r)" "M(B)" "is_function_space(M,r,B,d)" "is_function_space(M,r,B,d')" shows "d=d'" using assms extensionality_trans unfolding is_function_space_def is_funspace_def by simp lemma is_function_space_witness: assumes "M(A)" "M(B)" shows "\d[M]. is_function_space(M,A,B,d)" proof - from assms interpret M_Pi_assumptions M A "\_. B" using Pi_replacement Pi_separation by unfold_locales (auto dest:transM simp add:Sigfun_def) have "\f[M]. f \ Pi_rel(M,A, \_. B) \ f \ A \ B" using Pi_rel_char by simp with assms show ?thesis unfolding is_funspace_def is_function_space_def by auto qed lemma is_function_space_closed : "is_function_space(M,A,B,d) \ M(d)" unfolding is_function_space_def by simp \ \adding closure to simpset and claset\ lemma function_space_rel_closed[intro,simp]: assumes "M(x)" "M(y)" shows "M(function_space_rel(M,x,y))" proof - have "is_function_space(M, x, y, THE xa. is_function_space(M, x, y, xa))" using assms theI[OF ex1I[of "is_function_space(M,x,y)"], OF _ is_function_space_uniqueness[of x y]] is_function_space_witness by auto then show ?thesis using assms is_function_space_closed unfolding function_space_rel_def by blast qed lemmas trans_function_space_rel_closed[trans_closed] = transM[OF _ function_space_rel_closed] lemma is_function_space_iff: assumes "M(x)" "M(y)" "M(d)" shows "is_function_space(M,x,y,d) \ d = function_space_rel(M,x,y)" proof (intro iffI) assume "d = function_space_rel(M,x,y)" moreover note assms moreover from this obtain e where "M(e)" "is_function_space(M,x,y,e)" using is_function_space_witness by blast ultimately show "is_function_space(M, x, y, d)" using is_function_space_uniqueness[of x y] is_function_space_witness theI[OF ex1I[of "is_function_space(M,x,y)"], OF _ is_function_space_uniqueness[of x y], of e] unfolding function_space_rel_def by auto next assume "is_function_space(M, x, y, d)" with assms show "d = function_space_rel(M,x,y)" using is_function_space_uniqueness unfolding function_space_rel_def by (blast del:the_equality intro:the_equality[symmetric]) qed - lemma def_function_space_rel: assumes "M(A)" "M(y)" shows "function_space_rel(M,A,y) = Pi_rel(M,A,\_. y)" proof - from assms interpret M_Pi_assumptions M A "\_. y" using Pi_replacement Pi_separation by unfold_locales (auto dest:transM simp add:Sigfun_def) from assms have "x\function_space_rel(M,A,y) \ x\Pi_rel(M,A,\_. y)" if "M(x)" for x using that is_function_space_iff[of A y, OF _ _ function_space_rel_closed, of A y] def_Pi_rel Pi_rel_char mbnr.Pow_rel_char unfolding is_function_space_def is_funspace_def by (simp add:Pi_def) with assms show ?thesis \ \At this point, quoting "trans\_rules" doesn't work\ using transM[OF _ function_space_rel_closed, OF _ \M(A)\ \M(y)\] transM[OF _ Pi_rel_closed] by blast qed lemma function_space_rel_char: assumes "M(A)" "M(y)" shows "function_space_rel(M,A,y) = {f \ A \ y. M(f)}" proof - from assms interpret M_Pi_assumptions M A "\_. y" using Pi_replacement Pi_separation by unfold_locales (auto dest:transM simp add:Sigfun_def) show ?thesis using assms def_function_space_rel Pi_rel_char by simp qed lemma mem_function_space_rel_abs: assumes "M(A)" "M(y)" "M(f)" shows "f \ function_space_rel(M,A,y) \ f \ A \ y" using assms function_space_rel_char by simp end \ \\<^locale>\M_Pi\\ locale M_N_Pi = M:M_Pi + N:M_Pi N for N + assumes M_imp_N:"M(x) \ N(x)" begin lemma function_space_rel_transfer: "M(A) \ M(B) \ function_space_rel(M,A,B) \ function_space_rel(N,A,B)" using M.function_space_rel_char N.function_space_rel_char by (auto dest!:M_imp_N) end \ \\<^locale>\M_N_Pi\\ -(***************** end Discipline ***********************) - abbreviation "is_apply \ fun_apply" \ \It is not necessary to perform the Discipline for \<^term>\is_apply\ since it is absolute in this context\ subsection\Discipline for \<^term>\Collect\ terms.\ text\We have to isolate the predicate involved and apply the Discipline to it.\ -(*************** Discipline for injP ******************) - - -definition (* completely relational *) +definition \ \completely relational\ injP_rel:: "[i\o,i,i]\o" where "injP_rel(M,A,f) \ \w[M]. \x[M]. \fw[M]. \fx[M]. w\A \ x\A \ is_apply(M,f,w,fw) \ is_apply(M,f,x,fx) \ fw=fx\ w=x" synthesize "injP_rel" from_definition assuming "nonempty" arity_theorem for "injP_rel_fm" context M_basic begin \ \I'm undecided on keeping the relative quantifiers here. Same with \<^term>\surjP\ below. It might relieve from changing @{thm exI allI} to @{thm rexI rallI} in some proofs. I wonder if this escalates well. Assuming that all terms appearing in the "def\_" theorem are in \<^term>\M\ and using @{thm transM}, it might do.\ lemma def_injP_rel: assumes "M(A)" "M(f)" shows "injP_rel(M,A,f) \ (\w[M]. \x[M]. w\A \ x\A \ f`w=f`x \ w=x)" using assms unfolding injP_rel_def by simp end \ \\<^locale>\M_basic\\ -(****************** end Discipline **********************) - -(**********************************************************) subsection\Discipline for \<^term>\inj\\ -definition (* completely relational *) +definition \ \completely relational\ is_inj :: "[i\o,i,i,i]\o" where "is_inj(M,A,B,I) \ M(I) \ (\F[M]. is_function_space(M,A,B,F) \ is_Collect(M,F,injP_rel(M,A),I))" declare typed_function_iff_sats Collect_iff_sats [iff_sats] synthesize "is_funspace" from_definition assuming "nonempty" arity_theorem for "is_funspace_fm" synthesize "is_function_space" from_definition assuming "nonempty" notation is_function_space_fm (\\_ \ _ is _\\) arity_theorem for "is_function_space_fm" synthesize "is_inj" from_definition assuming "nonempty" notation is_inj_fm (\\inj'(_,_') is _\\) arity_theorem intermediate for "is_inj_fm" lemma arity_is_inj_fm[arity]: "A \ nat \ B \ nat \ I \ nat \ arity(is_inj_fm(A, B, I)) = succ(A) \ succ(B) \ succ(I)" using arity_is_inj_fm' by (auto simp:pred_Un_distrib arity) definition inj_rel :: "[i\o,i,i] \ i" (\inj\<^bsup>_\<^esup>'(_,_')\) where "inj_rel(M,A,B) \ THE d. is_inj(M,A,B,d)" abbreviation inj_r_set :: "[i,i,i] \ i" (\inj\<^bsup>_\<^esup>'(_,_')\) where "inj_r_set(M) \ inj_rel(##M)" locale M_inj = M_Pi + assumes injP_separation: "M(r) \ separation(M,injP_rel(M, r))" begin lemma is_inj_uniqueness: assumes "M(r)" "M(B)" "is_inj(M,r,B,d)" "is_inj(M,r,B,d')" shows "d=d'" using assms is_function_space_iff extensionality_trans unfolding is_inj_def by simp lemma is_inj_witness: "M(r) \ M(B)\ \d[M]. is_inj(M,r,B,d)" using injP_separation is_function_space_iff unfolding is_inj_def by simp lemma is_inj_closed : "is_inj(M,x,y,d) \ M(d)" unfolding is_inj_def by simp lemma inj_rel_closed[intro,simp]: assumes "M(x)" "M(y)" shows "M(inj_rel(M,x,y))" proof - have "is_inj(M, x, y, THE xa. is_inj(M, x, y, xa))" using assms theI[OF ex1I[of "is_inj(M,x,y)"], OF _ is_inj_uniqueness[of x y]] is_inj_witness by auto then show ?thesis using assms is_inj_closed unfolding inj_rel_def by blast qed lemmas trans_inj_rel_closed[trans_closed] = transM[OF _ inj_rel_closed] lemma inj_rel_iff: assumes "M(x)" "M(y)" "M(d)" shows "is_inj(M,x,y,d) \ d = inj_rel(M,x,y)" proof (intro iffI) assume "d = inj_rel(M,x,y)" moreover note assms moreover from this obtain e where "M(e)" "is_inj(M,x,y,e)" using is_inj_witness by blast ultimately show "is_inj(M, x, y, d)" using is_inj_uniqueness[of x y] is_inj_witness theI[OF ex1I[of "is_inj(M,x,y)"], OF _ is_inj_uniqueness[of x y], of e] unfolding inj_rel_def by auto next assume "is_inj(M, x, y, d)" with assms show "d = inj_rel(M,x,y)" using is_inj_uniqueness unfolding inj_rel_def by (blast del:the_equality intro:the_equality[symmetric]) qed lemma def_inj_rel: assumes "M(A)" "M(B)" shows "inj_rel(M,A,B) = {f \ function_space_rel(M,A,B). \w[M]. \x[M]. w\A \ x\A \ f`w = f`x \ w=x}" (is "_ = Collect(_,?P)") proof - from assms have "inj_rel(M,A,B) \ function_space_rel(M,A,B)" using inj_rel_iff[of A B "inj_rel(M,A,B)"] is_function_space_iff unfolding is_inj_def by auto moreover from assms have "f \ inj_rel(M,A,B) \ ?P(f)" for f using inj_rel_iff[of A B "inj_rel(M,A,B)"] is_function_space_iff def_injP_rel transM[OF _ function_space_rel_closed, OF _ \M(A)\ \M(B)\] unfolding is_inj_def by auto moreover from assms have "f \ function_space_rel(M,A,B) \ ?P(f) \ f \ inj_rel(M,A,B)" for f using inj_rel_iff[of A B "inj_rel(M,A,B)"] is_function_space_iff def_injP_rel transM[OF _ function_space_rel_closed, OF _ \M(A)\ \M(B)\] unfolding is_inj_def by auto ultimately show ?thesis by force qed lemma inj_rel_char: assumes "M(A)" "M(B)" shows "inj_rel(M,A,B) = {f \ inj(A,B). M(f)}" proof - from assms interpret M_Pi_assumptions M A "\_. B" using Pi_replacement Pi_separation by unfold_locales (auto dest:transM simp add:Sigfun_def) from assms show ?thesis using def_inj_rel[OF assms] def_function_space_rel[OF assms] transM[OF _ \M(A)\] Pi_rel_char unfolding inj_def by auto qed end \ \\<^locale>\M_inj\\ locale M_N_inj = M:M_inj + N:M_inj N for N + assumes M_imp_N:"M(x) \ N(x)" begin lemma inj_rel_transfer: "M(A) \ M(B) \ inj_rel(M,A,B) \ inj_rel(N,A,B)" using M.inj_rel_char N.inj_rel_char by (auto dest!:M_imp_N) end \ \\<^locale>\M_N_inj\\ -(*************** end Discipline *********************) - -(*************** Discipline for surjP ******************) +subsection\Discipline for \<^term>\surj\\ definition surjP_rel:: "[i\o,i,i,i]\o" where "surjP_rel(M,A,B,f) \ \y[M]. \x[M]. \fx[M]. y\B \ x\A \ is_apply(M,f,x,fx) \ fx=y" synthesize "surjP_rel" from_definition assuming "nonempty" context M_basic begin lemma def_surjP_rel: assumes "M(A)" "M(B)" "M(f)" shows "surjP_rel(M,A,B,f) \ (\y[M]. \x[M]. y\B \ x\A \ f`x=y)" using assms unfolding surjP_rel_def by auto end \ \\<^locale>\M_basic\\ -(****************** end Discipline **********************) - -(**********************************************************) subsection\Discipline for \<^term>\surj\\ -definition (* completely relational *) +definition \ \completely relational\ is_surj :: "[i\o,i,i,i]\o" where "is_surj(M,A,B,I) \ M(I) \ (\F[M]. is_function_space(M,A,B,F) \ is_Collect(M,F,surjP_rel(M,A,B),I))" synthesize "is_surj" from_definition assuming "nonempty" notation is_surj_fm (\\surj'(_,_') is _\\) definition surj_rel :: "[i\o,i,i] \ i" (\surj\<^bsup>_\<^esup>'(_,_')\) where "surj_rel(M,A,B) \ THE d. is_surj(M,A,B,d)" abbreviation surj_r_set :: "[i,i,i] \ i" (\surj\<^bsup>_\<^esup>'(_,_')\) where "surj_r_set(M) \ surj_rel(##M)" locale M_surj = M_Pi + assumes surjP_separation: "M(A)\M(B)\separation(M,\x. surjP_rel(M,A,B,x))" begin lemma is_surj_uniqueness: assumes "M(r)" "M(B)" "is_surj(M,r,B,d)" "is_surj(M,r,B,d')" shows "d=d'" using assms is_function_space_iff extensionality_trans unfolding is_surj_def by simp lemma is_surj_witness: "M(r) \ M(B)\ \d[M]. is_surj(M,r,B,d)" using surjP_separation is_function_space_iff unfolding is_surj_def by simp lemma is_surj_closed : "is_surj(M,x,y,d) \ M(d)" unfolding is_surj_def by simp lemma surj_rel_closed[intro,simp]: assumes "M(x)" "M(y)" shows "M(surj_rel(M,x,y))" proof - have "is_surj(M, x, y, THE xa. is_surj(M, x, y, xa))" using assms theI[OF ex1I[of "is_surj(M,x,y)"], OF _ is_surj_uniqueness[of x y]] is_surj_witness by auto then show ?thesis using assms is_surj_closed unfolding surj_rel_def by blast qed lemmas trans_surj_rel_closed[trans_closed] = transM[OF _ surj_rel_closed] lemma surj_rel_iff: assumes "M(x)" "M(y)" "M(d)" shows "is_surj(M,x,y,d) \ d = surj_rel(M,x,y)" proof (intro iffI) assume "d = surj_rel(M,x,y)" moreover note assms moreover from this obtain e where "M(e)" "is_surj(M,x,y,e)" using is_surj_witness by blast ultimately show "is_surj(M, x, y, d)" using is_surj_uniqueness[of x y] is_surj_witness theI[OF ex1I[of "is_surj(M,x,y)"], OF _ is_surj_uniqueness[of x y], of e] unfolding surj_rel_def by auto next assume "is_surj(M, x, y, d)" with assms show "d = surj_rel(M,x,y)" using is_surj_uniqueness unfolding surj_rel_def by (blast del:the_equality intro:the_equality[symmetric]) qed lemma def_surj_rel: assumes "M(A)" "M(B)" shows "surj_rel(M,A,B) = {f \ function_space_rel(M,A,B). \y[M]. \x[M]. y\B \ x\A \ f`x=y }" (is "_ = Collect(_,?P)") proof - from assms have "surj_rel(M,A,B) \ function_space_rel(M,A,B)" using surj_rel_iff[of A B "surj_rel(M,A,B)"] is_function_space_iff unfolding is_surj_def by auto moreover from assms have "f \ surj_rel(M,A,B) \ ?P(f)" for f using surj_rel_iff[of A B "surj_rel(M,A,B)"] is_function_space_iff def_surjP_rel transM[OF _ function_space_rel_closed, OF _ \M(A)\ \M(B)\] unfolding is_surj_def by auto moreover from assms have "f \ function_space_rel(M,A,B) \ ?P(f) \ f \ surj_rel(M,A,B)" for f using surj_rel_iff[of A B "surj_rel(M,A,B)"] is_function_space_iff def_surjP_rel transM[OF _ function_space_rel_closed, OF _ \M(A)\ \M(B)\] unfolding is_surj_def by auto ultimately show ?thesis by force qed lemma surj_rel_char: assumes "M(A)" "M(B)" shows "surj_rel(M,A,B) = {f \ surj(A,B). M(f)}" proof - from assms interpret M_Pi_assumptions M A "\_. B" using Pi_replacement Pi_separation by unfold_locales (auto dest:transM simp add:Sigfun_def) from assms show ?thesis using def_surj_rel[OF assms] def_function_space_rel[OF assms] transM[OF _ \M(A)\] transM[OF _ \M(B)\] Pi_rel_char unfolding surj_def by auto qed end \ \\<^locale>\M_surj\\ locale M_N_surj = M:M_surj + N:M_surj N for N + assumes M_imp_N:"M(x) \ N(x)" begin lemma surj_rel_transfer: "M(A) \ M(B) \ surj_rel(M,A,B) \ surj_rel(N,A,B)" using M.surj_rel_char N.surj_rel_char by (auto dest!:M_imp_N) end \ \\<^locale>\M_N_surj\\ -(*************** end Discipline *********************) +subsection\Discipline for \<^term>\Inter\\ definition is_Int :: "[i\o,i,i,i]\o" where "is_Int(M,A,B,I) \ M(I) \ (\x[M]. x \ I \ x \ A \ x \ B)" reldb_rem relational "inter" reldb_add absolute relational "ZF_Base.Int" "is_Int" synthesize "is_Int" from_definition assuming "nonempty" notation is_Int_fm (\_ \ _ is _\) context M_basic begin lemma is_Int_closed : "is_Int(M,A,B,I) \ M(I)" unfolding is_Int_def by simp lemma is_Int_abs: assumes "M(A)" "M(B)" "M(I)" shows "is_Int(M,A,B,I) \ I = A \ B" using assms transM[OF _ \M(B)\] transM[OF _ \M(I)\] unfolding is_Int_def by blast lemma is_Int_uniqueness: assumes "M(r)" "M(B)" "is_Int(M,r,B,d)" "is_Int(M,r,B,d')" shows "d=d'" proof - have "M(d)" and "M(d')" using assms is_Int_closed by simp+ then show ?thesis using assms is_Int_abs by simp qed text\Note: @{thm Int_closed} already in \<^theory>\ZF-Constructible.Relative\.\ end \ \\<^locale>\M_basic\\ -(**********************************************************) subsection\Discipline for \<^term>\bij\\ reldb_add functional "inj" "inj_rel" reldb_add functional relational "inj_rel" "is_inj" reldb_add functional "surj" "surj_rel" reldb_add functional relational "surj_rel" "is_surj" relativize functional "bij" "bij_rel" external relationalize "bij_rel" "is_bij" -(* definition (* completely relational *) - is_bij :: "[i\o,i,i,i]\o" where - "is_bij(M,A,B,bj) \ M(bj) \ is_hcomp2_2(M,is_Int,is_inj,is_surj,A,B,bj)" - -definition - bij_rel :: "[i\o,i,i] \ i" (\bij\<^bsup>_\<^esup>'(_,_')\) where - "bij_rel(M,A,B) \ THE d. is_bij(M,A,B,d)" *) - synthesize "is_bij" from_definition assuming "nonempty" notation is_bij_fm (\\bij'(_,_') is _\\) abbreviation bij_r_class :: "[i\o,i,i] \ i" (\bij\<^bsup>_\<^esup>'(_,_')\) where "bij_r_class \ bij_rel" abbreviation bij_r_set :: "[i,i,i] \ i" (\bij\<^bsup>_\<^esup>'(_,_')\) where "bij_r_set(M) \ bij_rel(##M)" locale M_Perm = M_Pi + M_inj + M_surj begin lemma is_bij_closed : "is_bij(M,f,y,d) \ M(d)" unfolding is_bij_def using is_Int_closed is_inj_witness is_surj_witness by auto lemma bij_rel_closed[intro,simp]: assumes "M(x)" "M(y)" shows "M(bij_rel(M,x,y))" unfolding bij_rel_def using assms Int_closed surj_rel_closed inj_rel_closed by auto lemmas trans_bij_rel_closed[trans_closed] = transM[OF _ bij_rel_closed] lemma bij_rel_iff: assumes "M(x)" "M(y)" "M(d)" shows "is_bij(M,x,y,d) \ d = bij_rel(M,x,y)" unfolding is_bij_def bij_rel_def using assms surj_rel_iff inj_rel_iff is_Int_abs by auto lemma def_bij_rel: assumes "M(A)" "M(B)" shows "bij_rel(M,A,B) = inj_rel(M,A,B) \ surj_rel(M,A,B)" using assms bij_rel_iff inj_rel_iff surj_rel_iff is_Int_abs\ \For absolute terms, "\_abs" replaces "\_iff". Also, in this case "\_closed" is in the simpset.\ unfolding is_bij_def by simp lemma bij_rel_char: assumes "M(A)" "M(B)" shows "bij_rel(M,A,B) = {f \ bij(A,B). M(f)}" using assms def_bij_rel inj_rel_char surj_rel_char unfolding bij_def\ \Unfolding this might be a pattern already\ by auto end \ \\<^locale>\M_Perm\\ locale M_N_Perm = M_N_Pi + M_N_inj + M_N_surj + M:M_Perm + N:M_Perm N begin lemma bij_rel_transfer: "M(A) \ M(B) \ bij_rel(M,A,B) \ bij_rel(N,A,B)" using M.bij_rel_char N.bij_rel_char by (auto dest!:M_imp_N) end \ \\<^locale>\M_N_Perm\\ -(*************** end Discipline *********************) - context M_Perm begin lemma mem_bij_rel: "\f \ bij\<^bsup>M\<^esup>(A,B); M(A); M(B)\ \ f\bij(A,B)" using bij_rel_char by simp lemma mem_inj_rel: "\f \ inj\<^bsup>M\<^esup>(A,B); M(A); M(B)\ \ f\inj(A,B)" using inj_rel_char by simp lemma mem_surj_rel: "\f \ surj\<^bsup>M\<^esup>(A,B); M(A); M(B)\ \ f\surj(A,B)" using surj_rel_char by simp end \ \\<^locale>\M_Perm\\ -(******************************************************) subsection\Discipline for \<^term>\eqpoll\\ relativize functional "eqpoll" "eqpoll_rel" external relationalize "eqpoll_rel" "is_eqpoll" synthesize "is_eqpoll" from_definition assuming "nonempty" arity_theorem for "is_eqpoll_fm" notation is_eqpoll_fm (\\_ \ _\\) context M_Perm begin is_iff_rel for "eqpoll" using bij_rel_iff unfolding is_eqpoll_def eqpoll_rel_def by simp end \ \\<^locale>\M_Perm\\ abbreviation eqpoll_r :: "[i,i\o,i] => o" (\_ \\<^bsup>_\<^esup> _\ [51,1,51] 50) where "A \\<^bsup>M\<^esup> B \ eqpoll_rel(M,A,B)" abbreviation eqpoll_r_set :: "[i,i,i] \ o" (\_ \\<^bsup>_\<^esup> _\ [51,1,51] 50) where "eqpoll_r_set(A,M) \ eqpoll_rel(##M,A)" context M_Perm begin lemma def_eqpoll_rel: assumes "M(A)" "M(B)" shows "eqpoll_rel(M,A,B) \ (\f[M]. f \ bij_rel(M,A,B))" using assms bij_rel_iff unfolding eqpoll_rel_def by simp end \ \\<^locale>\M_Perm\\ context M_N_Perm begin -(* the next lemma is not part of the discipline *) +text\The next lemma is not part of the discipline\ lemma eqpoll_rel_transfer: assumes "A \\<^bsup>M\<^esup> B" "M(A)" "M(B)" shows "A \\<^bsup>N\<^esup> B" proof - note assms moreover from this obtain f where "f \ bij\<^bsup>M\<^esup>(A,B)" "N(f)" using M.def_eqpoll_rel by (auto dest!:M_imp_N) moreover from calculation have "f \ bij\<^bsup>N\<^esup>(A,B)" using bij_rel_transfer by (auto) ultimately show ?thesis using N.def_eqpoll_rel by (blast dest!:M_imp_N) qed end \ \\<^locale>\M_N_Perm\\ -(****************** end Discipline ******************) - -(******************************************************) subsection\Discipline for \<^term>\lepoll\\ relativize functional "lepoll" "lepoll_rel" external relationalize "lepoll_rel" "is_lepoll" synthesize "is_lepoll" from_definition assuming "nonempty" notation is_lepoll_fm (\\_ \ _\\) arity_theorem for "is_lepoll_fm" context M_inj begin is_iff_rel for "lepoll" using inj_rel_iff unfolding is_lepoll_def lepoll_rel_def by simp end \ \\<^locale>\M_inj\\ abbreviation lepoll_r :: "[i,i\o,i] => o" (\_ \\<^bsup>_\<^esup> _\ [51,1,51] 50) where "A \\<^bsup>M\<^esup> B \ lepoll_rel(M,A,B)" abbreviation lepoll_r_set :: "[i,i,i] \ o" (\_ \\<^bsup>_\<^esup> _\ [51,1,51] 50) where "lepoll_r_set(A,M) \ lepoll_rel(##M,A)" context M_Perm begin lemma def_lepoll_rel: assumes "M(A)" "M(B)" shows "lepoll_rel(M,A,B) \ (\f[M]. f \ inj_rel(M,A,B))" using assms inj_rel_iff unfolding lepoll_rel_def by simp end \ \\<^locale>\M_Perm\\ context M_N_Perm begin -(* the next lemma is not part of the discipline *) +\ \This lemma is not part of the discipline.\ lemma lepoll_rel_transfer: assumes "A \\<^bsup>M\<^esup> B" "M(A)" "M(B)" shows "A \\<^bsup>N\<^esup> B" proof - note assms moreover from this obtain f where "f \ inj\<^bsup>M\<^esup>(A,B)" "N(f)" using M.def_lepoll_rel by (auto dest!:M_imp_N) moreover from calculation have "f \ inj\<^bsup>N\<^esup>(A,B)" using inj_rel_transfer by (auto) ultimately show ?thesis using N.def_lepoll_rel by (blast dest!:M_imp_N) qed end \ \\<^locale>\M_N_Perm\\ -(****************** end Discipline ******************) - -(******************************************************) subsection\Discipline for \<^term>\lesspoll\\ relativize functional "lesspoll" "lesspoll_rel" external relationalize "lesspoll_rel" "is_lesspoll" synthesize "is_lesspoll" from_definition assuming "nonempty" notation is_lesspoll_fm (\\_ \ _\\) arity_theorem for "is_lesspoll_fm" context M_Perm begin is_iff_rel for "lesspoll" using is_lepoll_iff is_eqpoll_iff unfolding is_lesspoll_def lesspoll_rel_def by simp end \ \\<^locale>\M_Perm\\ abbreviation lesspoll_r :: "[i,i\o,i] => o" (\_ \\<^bsup>_\<^esup> _\ [51,1,51] 50) where "A \\<^bsup>M\<^esup> B \ lesspoll_rel(M,A,B)" abbreviation lesspoll_r_set :: "[i,i,i] \ o" (\_ \\<^bsup>_\<^esup> _\ [51,1,51] 50) where "lesspoll_r_set(A,M) \ lesspoll_rel(##M,A)" text\Since \<^term>\lesspoll_rel\ is defined as a propositional combination of older terms, there is no need for a separate ``def'' theorem for it.\ text\Note that \<^term>\lesspoll_rel\ is neither $\Sigma_1^{\mathit{ZF}}$ nor $\Pi_1^{\mathit{ZF}}$, so there is no ``transfer'' theorem for it.\ definition Powapply :: "[i,i] \ i" where "Powapply(f,y) \ Pow(f`y)" reldb_add functional "Pow" "Pow_rel" reldb_add relational "Pow" "is_Pow" declare Replace_iff_sats[iff_sats] synthesize "is_Pow" from_definition assuming "nonempty" arity_theorem for "is_Pow_fm" relativize functional "Powapply" "Powapply_rel" relationalize "Powapply_rel" "is_Powapply" synthesize "is_Powapply" from_definition assuming "nonempty" arity_theorem for "is_Powapply_fm" notation Powapply_rel (\Powapply\<^bsup>_\<^esup>'(_,_')\) context M_basic begin rel_closed for "Powapply" unfolding Powapply_rel_def by simp is_iff_rel for "Powapply" using Pow_rel_iff unfolding is_Powapply_def Powapply_rel_def by simp end \\\<^locale>\M_basic\\ end \ No newline at end of file diff --git a/thys/Transitive_Models/Higher_Order_Constructs.thy b/thys/Transitive_Models/Higher_Order_Constructs.thy --- a/thys/Transitive_Models/Higher_Order_Constructs.thy +++ b/thys/Transitive_Models/Higher_Order_Constructs.thy @@ -1,152 +1,152 @@ -section\Fully relational versions of higher order construct \ +section\Fully relational versions of higher order constructs\ theory Higher_Order_Constructs imports Recursion_Thms Least begin syntax "_sats" :: "[i, i, i] \ o" ("(_, _ \ _)" [36,36,36] 25) translations "(M,env \ \)" \ "CONST sats(M,\,env)" definition is_If :: "[i\o,o,i,i,i] \ o" where "is_If(M,b,t,f,r) \ (b \ r=t) \ (\b \ r=f)" lemma (in M_trans) If_abs: "is_If(M,b,t,f,r) \ r = If(b,t,f)" by (simp add: is_If_def) definition is_If_fm :: "[i,i,i,i] \ i" where "is_If_fm(\,t,f,r) \ Or(And(\,Equal(t,r)),And(Neg(\),Equal(f,r)))" lemma is_If_fm_type [TC]: "\ \ formula \ t \ nat \ f \ nat \ r \ nat \ is_If_fm(\,t,f,r) \ formula" unfolding is_If_fm_def by auto lemma sats_is_If_fm: assumes Qsats: "Q \ A, env \ \" "env \ list(A)" shows "is_If(##A, Q, nth(t, env), nth(f, env), nth(r, env)) \ A, env \ is_If_fm(\,t,f,r)" using assms unfolding is_If_def is_If_fm_def by auto lemma is_If_fm_iff_sats [iff_sats]: assumes Qsats: "Q \ A, env \ \" and "nth(t, env) = ta" "nth(f, env) = fa" "nth(r, env) = ra" "t \ nat" "f \ nat" "r \ nat" "env \ list(A)" shows "is_If(##A,Q,ta,fa,ra) \ A, env \ is_If_fm(\,t,f,r)" using assms sats_is_If_fm[of Q A \ env t f r] by simp lemma arity_is_If_fm [arity]: "\ \ formula \ t \ nat \ f \ nat \ r \ nat \ arity(is_If_fm(\, t, f, r)) = arity(\) \ succ(t) \ succ(r) \ succ(f)" unfolding is_If_fm_def by auto definition is_The :: "[i\o,i\o,i] \ o" where "is_The(M,Q,i) \ (Q(i) \ (\x[M]. Q(x) \ (\y[M]. Q(y) \ y = x))) \ (\(\x[M]. Q(x) \ (\y[M]. Q(y) \ y = x))) \ empty(M,i) " (* definition is_The_fm :: "[i,i] \ i" where "is_The_fm(q,i) \ Or(And(Exists(And(Equal(succ(i),0),q)), Exists(And(q,Forall(Implies(q,Equal(1,0)))))), And(Neg(Exists(And(q,Forall(Implies(q,Equal(1,0)))))),empty_fm(i)))" (* this doesn't work yet *) lemma sats_The_fm : assumes p_iff_sats: "\a. a \ A \ P(a) \ sats(A, p, Cons(a, env))" shows "\y \ nat; env \ list(A) ; 0\A\ \ sats(A, is_The_fm(p,y), env) \ is_The(##A, P, nth(y,env))" using nth_closed p_iff_sats unfolding is_The_def is_The_fm_def oops lemma The_iff_sats [iff_sats]: assumes is_Q_iff_sats: "\a. a \ A \ is_Q(a) \ sats(A, q, Cons(a,env))" shows "\nth(j,env) = y; j \ nat; env \ list(A); 0\A\ \ is_The(##A, is_Q, y) \ sats(A, is_The_fm(q,j), env)" using sats_The_fm [OF is_Q_iff_sats, of j , symmetric] by simp *) lemma (in M_trans) The_abs: assumes "\x. Q(x) \ M(x)" "M(a)" shows "is_The(M,Q,a) \ a = (THE x. Q(x))" proof (cases "\x[M]. Q(x) \ (\y[M]. Q(y) \ y = x)") case True with assms show ?thesis unfolding is_The_def by (intro iffI the_equality[symmetric]) (auto, blast intro:theI) next case False with \\x. Q(x) \ M(x)\ have " \ (\x. Q(x) \ (\y. Q(y) \ y = x))" by auto then have "The(Q) = 0" by (intro the_0) auto with assms and False show ?thesis unfolding is_The_def by auto qed (* definition recursor :: "[i, [i,i]=>i, i]=>i" where "recursor(a,b,k) \ transrec(k, \n f. nat_case(a, \m. b(m, f`m), n))" *) definition is_recursor :: "[i\o,i,[i,i,i]\o,i,i] \o" where "is_recursor(M,a,is_b,k,r) \ is_transrec(M, \n f ntc. is_nat_case(M,a, \m bmfm. \fm[M]. fun_apply(M,f,m,fm) \ is_b(m,fm,bmfm),n,ntc),k,r)" lemma (in M_eclose) recursor_abs: assumes "Ord(k)" and types: "M(a)" "M(k)" "M(r)" and b_iff: "\m f bmf. M(m) \ M(f) \ M(bmf) \ is_b(m,f,bmf) \ bmf = b(m,f)" and b_closed: "\m f bmf. M(m) \ M(f) \ M(b(m,f))" and repl: "transrec_replacement(M, \n f ntc. is_nat_case(M, a, \m bmfm. \fm[M]. fun_apply(M, f, m, fm) \ is_b( m, fm, bmfm), n, ntc), k)" shows "is_recursor(M,a,is_b,k,r) \ r = recursor(a,b,k)" unfolding is_recursor_def recursor_def using assms apply (rule_tac transrec_abs) apply (auto simp:relation2_def) apply (rule nat_case_abs[THEN iffD1, where is_b1="\m bmfm. \fm[M]. fun_apply(M,_,m,fm) \ is_b(m,fm,bmfm)"]) apply (auto simp:relation1_def) apply (rule nat_case_abs[THEN iffD2, where is_b1="\m bmfm. \fm[M]. fun_apply(M,_,m,fm) \ is_b(m,fm,bmfm)"]) apply (auto simp:relation1_def) done definition is_wfrec_on :: "[i=>o,[i,i,i]=>o,i,i,i, i] => o" where "is_wfrec_on(M,MH,A,r,a,z) == is_wfrec(M,MH,r,a,z)" lemma (in M_trancl) trans_wfrec_on_abs: "[|wf(r); trans(r); relation(r); M(r); M(a); M(z); wfrec_replacement(M,MH,r); relation2(M,MH,H); \x[M]. \g[M]. function(g) \ M(H(x,g)); r-``{a}\A; a \ A|] ==> is_wfrec_on(M,MH,A,r,a,z) \ z=wfrec[A](r,a,H)" using trans_wfrec_abs wfrec_trans_restr unfolding is_wfrec_on_def by simp end \ No newline at end of file diff --git a/thys/Transitive_Models/Lambda_Replacement.thy b/thys/Transitive_Models/Lambda_Replacement.thy --- a/thys/Transitive_Models/Lambda_Replacement.thy +++ b/thys/Transitive_Models/Lambda_Replacement.thy @@ -1,2361 +1,2508 @@ section\Replacements using Lambdas\ theory Lambda_Replacement imports Discipline_Function begin text\In this theory we prove several instances of separation and replacement in @{locale M_basic}. Moreover we introduce a new locale assuming two instances of separation and twelve instances of lambda replacements (ie, replacement of the form $\lambda x y. y=\langle x, f(x) \rangle$) we prove a bunch of other instances.\ definition lam_replacement :: "[i\o,i\i] \ o" where "lam_replacement(M,b) \ strong_replacement(M, \x y. y = \x, b(x)\)" lemma separation_univ : shows "separation(M,M)" unfolding separation_def by auto context M_trivial begin lemma lam_replacement_iff_lam_closed: assumes "\x[M]. M(b(x))" shows "lam_replacement(M, b) \ (\A[M]. M(\x\A. b(x)))" using assms lam_closed lam_funtype[of _ b, THEN Pi_memberD] unfolding lam_replacement_def strong_replacement_def by (auto intro:lamI dest:transM) (rule lam_closed, auto simp add:strong_replacement_def dest:transM) lemma lam_replacement_imp_lam_closed: assumes "lam_replacement(M, b)" "M(A)" "\x\A. M(b(x))" shows "M(\x\A. b(x))" using assms unfolding lam_replacement_def by (rule_tac lam_closed, auto simp add:strong_replacement_def dest:transM) lemma lam_replacement_cong: assumes "lam_replacement(M,f)" "\x[M]. f(x) = g(x)" "\x[M]. M(f(x))" shows "lam_replacement(M,g)" proof - note assms moreover from this have "\A[M]. M(\x\A. f(x))" using lam_replacement_iff_lam_closed by simp moreover from calculation have "(\x\A . f(x)) = (\x\A . g(x))" if "M(A)" for A using lam_cong[OF refl,of A f g] transM[OF _ that] by simp ultimately show ?thesis using lam_replacement_iff_lam_closed by simp qed end \ \\<^locale>\M_trivial\\ context M_basic begin lemma separation_iff': assumes "separation(M,\x . P(x))" "separation(M,\x . Q(x))" shows "separation(M,\x . P(x) \ Q(x))" using assms separation_conj separation_imp iff_def by auto lemma separation_in_constant : assumes "M(a)" shows "separation(M,\x . x\a)" proof - have "{x\A . x\a} = A \ a" for A by auto with \M(a)\ show ?thesis using separation_iff Collect_abs by simp qed lemma separation_equal : shows "separation(M,\x . x=a)" proof - have "{x\A . x=a} = (if a\A then {a} else 0)" for A by auto then have "M({x\A . x=a})" if "M(A)" for A using transM[OF _ \M(A)\] by simp then show ?thesis using separation_iff Collect_abs by simp qed lemma (in M_basic) separation_in_rev: assumes "(M)(a)" shows "separation(M,\x . a\x)" proof - have eq: "{x\A. a\x} = Memrel(A\{a}) `` {a}" for A unfolding ZF_Base.image_def by(intro equalityI,auto simp:mem_not_refl) moreover from assms have "M(Memrel(A\{a}) `` {a})" if "M(A)" for A using that by simp ultimately show ?thesis using separation_iff Collect_abs by simp qed +\ \\ lemma lam_replacement_imp_strong_replacement_aux: assumes "lam_replacement(M, b)" "\x[M]. M(b(x))" shows "strong_replacement(M, \x y. y = b(x))" proof - { fix A note assms moreover assume "M(A)" moreover from calculation have "M(\x\A. b(x))" using lam_replacement_iff_lam_closed by auto ultimately have "M((\x\A. b(x))``A)" "\z[M]. z \ (\x\A. b(x))``A \ (\x\A. z = b(x))" by (auto simp:lam_def) } then show ?thesis unfolding strong_replacement_def by clarsimp (rule_tac x="(\x\A. b(x))``A" in rexI, auto) qed +\ \This lemma could be modularized into the instantiation (fixing \<^term>\X\) +and the projection of the result of \<^term>\f\.\ +lemma strong_lam_replacement_imp_strong_replacement : + assumes "strong_replacement(M,\ x z . P(fst(x),snd(x)) \ z=\x,f(fst(x),snd(x))\)" + "\A . M(A) \ \x\A. P(X,x) \ M(f(X,x))" "M(X)" + shows "strong_replacement(M,\ x z . P(X,x) \ z=f(X,x))" + unfolding strong_replacement_def +proof(clarsimp) + fix A + assume "M(A)" + moreover from this \M(X)\ + have "M({X}\A)" (is "M(?A)") + by simp + moreover + have "fst(x) = X" if "x\?A" for x + using that by auto + moreover from calculation assms + have "M({z . x\?A , P(fst(x),snd(x)) \ z=\x,f(fst(x),snd(x))\})" (is "M(?F)") + using transM[of _ ?A] + by(rule_tac strong_replacement_closed,simp_all) + moreover + have "?F=({\x,f(fst(x),snd(x))\ . x\ {x\?A . P(fst(x),snd(x))}})" (is "_=(?G)") + by auto + moreover + note \M(?A)\ + ultimately + have "M(?G``?A)" + by simp + moreover + have "?G``?A = {y . x\?A , P(fst(x),snd(x)) \ y = f(fst(x),snd(x))}" (is "_=(?H)") + by auto + ultimately + show "\Y[M]. \b[M]. b \ Y \ (\x\A. P(X,x) \ b = f(X,x))" + by(rule_tac rexI[of _ ?H],auto,force) +qed + lemma lam_replacement_imp_RepFun_Lam: assumes "lam_replacement(M, f)" "M(A)" shows "M({y . x\A , M(y) \ y=\x,f(x)\})" proof - from assms obtain Y where 1:"M(Y)" "\b[M]. b \ Y \ (\x[M]. x \ A \ b = \x,f(x)\)" unfolding lam_replacement_def strong_replacement_def by auto moreover from calculation have "Y = {y . x\A , M(y) \ y = \x,f(x)\}" (is "Y=?R") proof(intro equalityI subsetI) fix y assume "y\Y" moreover from this 1 obtain x where "x\A" "y=\x,f(x)\" "M(y)" using transM[OF _ \M(Y)\] by auto ultimately show "y\?R" by auto next fix z assume "z\?R" moreover from this obtain a where "a\A" "z=\a,f(a)\" "M(a)" "M(f(a))" using transM[OF _ \M(A)\] by auto ultimately show "z\Y" using 1 by simp qed ultimately show ?thesis by auto qed lemma lam_closed_imp_closed: assumes "\A[M]. M(\x\A. f(x))" shows "\x[M]. M(f(x))" proof fix x assume "M(x)" moreover from this and assms have "M(\x\{x}. f(x))" by simp ultimately show "M(f(x))" using image_lam[of "{x}" "{x}" f] image_closed[of "{x}" "(\x\{x}. f(x))"] by (auto dest:transM) qed lemma lam_replacement_if: assumes "lam_replacement(M,f)" "lam_replacement(M,g)" "separation(M,b)" "\x[M]. M(f(x))" "\x[M]. M(g(x))" shows "lam_replacement(M, \x. if b(x) then f(x) else g(x))" proof - let ?G="\x. if b(x) then f(x) else g(x)" let ?b="\A . {x\A. b(x)}" and ?b'="\A . {x\A. \b(x)}" have eq:"(\x\A . ?G(x)) = (\x\?b(A) . f(x)) \ (\x\?b'(A).g(x))" for A unfolding lam_def by auto have "?b'(A) = A - ?b(A)" for A by auto moreover have "M(?b(A))" if "M(A)" for A using assms that by simp moreover from calculation have "M(?b'(A))" if "M(A)" for A using that by simp moreover from calculation assms have "M(\x\?b(A). f(x))" "M(\x\?b'(A) . g(x))" if "M(A)" for A using lam_replacement_iff_lam_closed that by simp_all moreover from this have "M((\x\?b(A) . f(x)) \ (\x\?b'(A).g(x)))" if "M(A)" for A using that by simp ultimately have "M(\x\A. if b(x) then f(x) else g(x))" if "M(A)" for A using that eq by simp with assms show ?thesis using lam_replacement_iff_lam_closed by simp qed lemma lam_replacement_constant: "M(b) \ lam_replacement(M,\_. b)" unfolding lam_replacement_def strong_replacement_def by safe (rule_tac x="_\{b}" in rexI; blast) subsection\Replacement instances obtained through Powerset\ txt\The next few lemmas provide bounds for certain constructions.\ lemma not_functional_Replace_0: assumes "\(\y y'. P(y) \ P(y') \ y=y')" shows "{y . x \ A, P(y)} = 0" using assms by (blast elim!: ReplaceE) lemma Replace_in_Pow_rel: assumes "\x b. x \ A \ P(x,b) \ b \ U" "\x\A. \y y'. P(x,y) \ P(x,y') \ y=y'" "separation(M, \y. \x[M]. x \ A \ P(x, y))" "M(U)" "M(A)" shows "{y . x \ A, P(x, y)} \ Pow\<^bsup>M\<^esup>(U)" proof - from assms have "{y . x \ A, P(x, y)} \ U" "z \ {y . x \ A, P(x, y)} \ M(z)" for z by (auto dest:transM) with assms have "{y . x \ A, P(x, y)} = {y\U . \x[M]. x\A \ P(x,y)}" by (intro equalityI) (auto, blast) with assms have "M({y . x \ A, P(x, y)})" by simp with assms show ?thesis using mem_Pow_rel_abs by auto qed lemma Replace_sing_0_in_Pow_rel: assumes "\b. P(b) \ b \ U" "separation(M, \y. P(y))" "M(U)" shows "{y . x \ {0}, P(y)} \ Pow\<^bsup>M\<^esup>(U)" proof (cases "\y y'. P(y) \ P(y') \ y=y'") case True with assms show ?thesis by (rule_tac Replace_in_Pow_rel) auto next case False with assms show ?thesis using nonempty not_functional_Replace_0[of P "{0}"] Pow_rel_char by auto qed lemma The_in_Pow_rel_Union: assumes "\b. P(b) \ b \ U" "separation(M, \y. P(y))" "M(U)" shows "(THE i. P(i)) \ Pow\<^bsup>M\<^esup>(\U)" proof - note assms moreover from this have "(THE i. P(i)) \ Pow(\U)" unfolding the_def by auto moreover from assms have "M(THE i. P(i))" using Replace_sing_0_in_Pow_rel[of P U] unfolding the_def by (auto dest:transM) ultimately show ?thesis using Pow_rel_char by auto qed lemma separation_least: "separation(M, \y. Ord(y) \ P(y) \ (\j. j < y \ \ P(j)))" unfolding separation_def proof fix z assume "M(z)" have "M({x \ z . x \ z \ Ord(x) \ P(x) \ (\j. j < x \ \ P(j))})" (is "M(?y)") proof (cases "\x\z. Ord(x) \ P(x) \ (\j. j < x \ \ P(j))") case True with \M(z)\ have "\x[M]. ?y = {x}" by (safe, rename_tac x, rule_tac x=x in rexI) (auto dest:transM, intro equalityI, auto elim:Ord_linear_lt) then show ?thesis by auto next case False then have "{x \ z . x \ z \ Ord(x) \ P(x) \ (\j. j < x \ \ P(j))} = 0" by auto then show ?thesis by auto qed moreover from this have "\x[M]. x \ ?y \ x \ z \ Ord(x) \ P(x) \ (\j. j < x \ \ P(j))" by simp ultimately show "\y[M]. \x[M]. x \ y \ x \ z \ Ord(x) \ P(x) \ (\j. j < x \ \ P(j))" by blast qed lemma Least_in_Pow_rel_Union: assumes "\b. P(b) \ b \ U" "M(U)" shows "(\ i. P(i)) \ Pow\<^bsup>M\<^esup>(\U)" using assms separation_least unfolding Least_def by (rule_tac The_in_Pow_rel_Union) simp lemma bounded_lam_replacement: fixes U assumes "\X[M]. \x\X. f(x) \ U(X)" and separation_f:"\A[M]. separation(M,\y. \x[M]. x\A \ y = \x, f(x)\)" and U_closed [intro,simp]: "\X. M(X) \ M(U(X))" shows "lam_replacement(M, f)" proof - have "M(\x\A. f(x))" if "M(A)" for A proof - have "(\x\A. f(x)) = {y\ Pow\<^bsup>M\<^esup>(Pow\<^bsup>M\<^esup>(A \ U(A))). \x[M]. x\A \ y = \x, f(x)\}" using \M(A)\ unfolding lam_def proof (intro equalityI, auto) fix x assume "x\A" moreover note \M(A)\ moreover from calculation assms have "f(x) \ U(A)" by simp moreover from calculation have "{x, f(x)} \ Pow\<^bsup>M\<^esup>(A \ U(A))" "{x,x} \ Pow\<^bsup>M\<^esup>(A \ U(A))" using Pow_rel_char[of "A \ U(A)"] by (auto dest:transM) ultimately show "\x, f(x)\ \ Pow\<^bsup>M\<^esup>(Pow\<^bsup>M\<^esup>(A \ U(A)))" using Pow_rel_char[of "Pow\<^bsup>M\<^esup>(A \ U(A))"] unfolding Pair_def by (auto dest:transM) qed moreover from \M(A)\ have "M({y\ Pow\<^bsup>M\<^esup>(Pow\<^bsup>M\<^esup>(A \ U(A))). \x[M]. x\A \ y = \x, f(x)\})" using separation_f by (rule_tac separation_closed) simp_all ultimately show ?thesis by simp qed moreover from this have "\x[M]. M(f(x))" using lam_closed_imp_closed by simp ultimately show ?thesis using assms by (rule_tac lam_replacement_iff_lam_closed[THEN iffD2]) simp_all qed lemma fst_in_double_Union: - assumes "x\X" "M(X)" + assumes "x\X" shows "fst(x) \ {0} \ \\X" proof - have "fst(x) \ {0} \ \x" for x unfolding fst_def using the_0 fst_conv by(cases "\!a.\b. x = \a,b\", auto, simp add:Pair_def) with assms show ?thesis by auto qed lemma snd_in_double_Union: - assumes "x\X" "M(X)" + assumes "x\X" shows "snd(x) \ {0} \ \\X" proof - have "snd(x) \ {0} \ \x" for x unfolding snd_def using the_0 snd_conv by(cases "\!a.\b. x = \a,b\", auto, simp add:Pair_def) with assms show ?thesis by auto qed lemma bounded_lam_replacement_binary: fixes U assumes "\X[M]. \x\X. \y\X. f(x,y) \ U(X)" and separation_f:"\A[M]. separation(M,\y. \x[M]. x\A \ y = \x, f(fst(x),snd(x))\)" and U_closed [intro,simp]: "\X. M(X) \ M(U(X))" shows "lam_replacement(M, \r . f(fst(r),snd(r)))" proof - have "M(\x\A. f(fst(x),snd(x)))" if "M(A)" for A proof - have "(\x\A. f(fst(x),snd(x))) = {y\ Pow\<^bsup>M\<^esup>(Pow\<^bsup>M\<^esup>(A \ U({0} \ \\A))). \x[M]. x\A \ y = \x, f(fst(x),snd(x))\}" using \M(A)\ unfolding lam_def proof (intro equalityI, auto) fix x assume "x\A" moreover note \M(A)\ moreover from calculation assms have "f(fst(x),snd(x)) \ U({0} \ \\A)" using fst_in_double_Union snd_in_double_Union by simp moreover from calculation have "{x, f(fst(x),snd(x))} \ Pow\<^bsup>M\<^esup>(A \ U({0} \ \\A))" "{x,x} \ Pow\<^bsup>M\<^esup>(A \ U({0} \ \\A))" using Pow_rel_char[of "A \ U({0} \ \\A)"] by (auto dest:transM) ultimately show "\x, f(fst(x),snd(x))\ \ Pow\<^bsup>M\<^esup>(Pow\<^bsup>M\<^esup>(A \ U({0} \ \\A)))" using Pow_rel_char[of "Pow\<^bsup>M\<^esup>(A \ U({0} \ \\A))"] unfolding Pair_def by (auto dest:transM) qed moreover from \M(A)\ have "M({y\ Pow\<^bsup>M\<^esup>(Pow\<^bsup>M\<^esup>(A \ U({0} \ \\A))). \x[M]. x\A \ y = \x, f(fst(x),snd(x))\})" using separation_f by (rule_tac separation_closed) simp_all ultimately show ?thesis by simp qed moreover from this have "\x[M]. M(f(fst(x),snd(x)))" using lam_closed_imp_closed[of "\x. f(fst(x),snd(x))"] by simp ultimately show ?thesis using assms by (rule_tac lam_replacement_iff_lam_closed[THEN iffD2]) simp_all qed \ \Below we assume the replacement instance for @{term fst}. Alternatively it follows from the instance of separation assumed in this lemma.\ lemma lam_replacement_fst': assumes "\A[M]. separation(M, \y. \x[M]. x\A \ y = \x, fst(x)\)" shows "lam_replacement(M,fst)" using fst_in_double_Union assms bounded_lam_replacement[of fst "\X. {0} \ \\X"] by simp lemma lam_replacement_snd': assumes "\A[M]. separation(M, \y. \x[M]. x\A \ y = \x, snd(x)\)" shows "lam_replacement(M,snd)" using snd_in_double_Union assms bounded_lam_replacement[of snd "\X. {0} \ \\X"] by simp +\ \We can prove this separation in the locale introduced below.\ lemma lam_replacement_restrict: assumes "\A[M]. separation(M, \y. \x[M]. x\A \ y = \x, restrict(x,B)\)" "M(B)" shows "lam_replacement(M, \r . restrict(r,B))" proof - from assms have "restrict(r,B)\Pow\<^bsup>M\<^esup>(\R)" if "M(R)" "r\R" for r R using Union_upper subset_Pow_Union subset_trans[OF restrict_subset] Pow_rel_char that transM[of _ R] by auto with assms show ?thesis using bounded_lam_replacement[of "\r . restrict(r,B)" "\X. Pow\<^bsup>M\<^esup>(\X)"] by simp qed -lemma domain_sub_Union3: - assumes "r\R" shows "domain(r) \ (\\\R)" - using assms - unfolding domain_def Pair_def - by auto - -lemma range_sub_Union3: - assumes "r\R" shows "range(r) \ (\\\R)" - unfolding range_def converse_def domain_def - by (clarsimp, auto simp add:Pair_def,rule_tac x="r" in rev_bexI,auto simp add:assms) - -lemma lam_replacement_converse': - assumes "\A[M]. separation(M, \y. \x[M]. x\A \ y = \x, converse(x)\)" - shows "lam_replacement(M, \r . converse(r))" -proof - - have "converse(r) \ range(r) \ domain(r)" for r - unfolding range_def domain_def converse_def - by auto - moreover from this - have "converse(r) \ Pow(\\\R \ \\\R)" if "r\R" for r R - using that domain_sub_Union3 range_sub_Union3 - subset_trans[OF \converse(_)\_\] Sigma_mono - by auto - ultimately - have "converse(r) \ Pow\<^bsup>M\<^esup>((\\\R \ \\\R))" if "M(R)" "r\R" for r R - using that Pow_rel_char transM[of _ R] by auto - with assms - show ?thesis - using bounded_lam_replacement[of "\r . converse(r)" "\X. Pow\<^bsup>M\<^esup>((\\\X \ \\\X))"] - by simp -qed - -lemma lam_replacement_domain' : - assumes "\A[M]. separation(M, \y. \x[M]. x \ A \ y = \x, domain(x)\)" - shows "lam_replacement(M,domain)" -proof - - have "domain(x) \ Pow\<^bsup>M\<^esup>(\\\A)" if "M(A)" "x\A" for x A - using that domain_sub_Union3 transM[of x A] Pow_rel_char - by auto - with assms - show ?thesis - using bounded_lam_replacement[where U="\A . Pow\<^bsup>M\<^esup>(\\\A)"] - by auto -qed - lemma lam_replacement_Union' : assumes "\A[M]. separation(M, \y. \x[M]. x \ A \ y = \x, \x\)" shows "lam_replacement(M,Union)" proof - have "\x \ Pow(\\A)" if "x\A" for x A using that by auto then have "\x \ Pow\<^bsup>M\<^esup>(\\A)" if "M(A)" "x\A" for x A using that transM[of x A] Pow_rel_char by auto with assms show ?thesis using bounded_lam_replacement[where U="\A . Pow\<^bsup>M\<^esup>(\\A)"] by auto qed -lemma Upair_in_Pow_rel: - assumes "x\X" "y\X" "M(X)" - shows "Upair(x,y) \ Pow\<^bsup>M\<^esup>(X)" - using assms transM[of _ X] mem_Pow_rel_abs - by auto - -lemma lam_replacement_Upair': - assumes - "\A[M]. separation(M,\y. \x[M]. x\A \ y = \x, Upair(fst(x),snd(x))\)" - shows - "lam_replacement(M, \r . Upair(fst(r),snd(r)))" - using assms Upair_in_Pow_rel - by (rule_tac bounded_lam_replacement_binary[of _ "\X. Pow\<^bsup>M\<^esup>(X)"]) auto - -lemma Diff_in_Pow_rel_Union: - assumes "x\X" "y\X" "M(X)" - shows "Diff(x,y) \ Pow\<^bsup>M\<^esup>(\X)" - using assms transM[of _ X] mem_Pow_rel_abs - by auto - -lemma lam_replacement_Diff'': - assumes - "\A[M]. separation(M,\y. \x[M]. x\A \ y = \x, Diff(fst(x),snd(x))\)" - shows - "lam_replacement(M, \r . Diff(fst(r),snd(r)))" - using assms Diff_in_Pow_rel_Union - by (rule_tac bounded_lam_replacement_binary[of _ "\X. Pow\<^bsup>M\<^esup>(\X)"]) auto - lemma Image_in_Pow_rel_Union3: assumes "x\X" "y\X" "M(X)" shows "Image(x,y) \ Pow\<^bsup>M\<^esup>(\\\X)" proof - from assms have "\a, b\ \ x \ b \ \\x" for a b unfolding Pair_def by (auto dest:transM) moreover from this and assms have "\a, b\ \ x \ M(b)" for a b using transM[of _ X] transM[of _ x] by auto ultimately show ?thesis using assms transM[of _ X] transM[of _ x] mem_Pow_rel_abs by auto fast qed lemma lam_replacement_Image': assumes "\A[M]. separation(M,\y. \x[M]. x\A \ y = \x, Image(fst(x),snd(x))\)" shows "lam_replacement(M, \r . Image(fst(r),snd(r)))" using assms Image_in_Pow_rel_Union3 by (rule_tac bounded_lam_replacement_binary[of _ "\X. Pow\<^bsup>M\<^esup>(\\\X)"]) auto -lemma comp_in_Pow_rel: - assumes "x\X" "y\X" "M(X)" - shows "comp(x,y) \ Pow\<^bsup>M\<^esup>((\\\X \ \\\X))" -proof - - have "comp(x,y) \ domain(y) \ range(x)" - using M_comp_iff by auto - moreover - note assms - moreover from this - have "comp(x,y) \ Pow(\\\X \ \\\X)" - using domain_sub_Union3 range_sub_Union3 - subset_trans[OF \comp(_,_)\_\] Sigma_mono - by auto - ultimately - show "comp(x,y) \ Pow\<^bsup>M\<^esup>((\\\X \ \\\X))" - using Pow_rel_char transM[of _ X] by auto -qed - -lemma lam_replacement_comp'': - assumes - "\A[M]. separation(M,\y. \x[M]. x\A \ y = \x, comp(fst(x),snd(x))\)" - shows - "lam_replacement(M, \r . comp(fst(r),snd(r)))" - using assms comp_in_Pow_rel - by (rule_tac bounded_lam_replacement_binary[of _ "\X. Pow\<^bsup>M\<^esup>((\\\X \ \\\X))"]) auto - lemma minimum_in_Union: - assumes "x\X" "y\X" "M(X)" + assumes "x\X" "y\X" shows "minimum(x,y) \ {0} \ \X" using assms minimum_in' by auto lemma lam_replacement_minimum': assumes "\A[M]. separation(M,\y. \x[M]. x\A \ y = \x, minimum(fst(x),snd(x))\)" shows "lam_replacement(M, \r . minimum(fst(r),snd(r)))" using assms minimum_in_Union by (rule_tac bounded_lam_replacement_binary[of _ "\X. {0} \ \X"]) auto lemma lam_replacement_Pow_rel: assumes "\A[M]. separation(M, \y. \x[M]. x \ A \ y = \x, Pow_rel(M,x)\)" shows "lam_replacement(M,Pow_rel(M))" proof - have "Pow_rel(M,x) \ Pow(Pow(\A))" if "x\A" "M(A)" for x A using that transM[of x A] def_Pow_rel[of _ x] by (auto dest:transM) then have "Pow_rel(M,x) \ Pow(Pow\<^bsup>M\<^esup>(\A))" if "M(A)" "x\A" for x A using that transM[of x A] Pow_rel_char by auto then have "Pow_rel(M,x) \ Pow\<^bsup>M\<^esup>(Pow\<^bsup>M\<^esup>(\A))" if "M(A)" "x\A" for x A using that transM[of x A] Pow_rel_char[of "Pow\<^bsup>M\<^esup>(\A)"] by auto with assms show ?thesis using bounded_lam_replacement[where U="\A. Pow\<^bsup>M\<^esup>(Pow\<^bsup>M\<^esup>(\A))"] by auto qed end \ \\<^locale>\M_basic\\ definition middle_del :: "i\i\i" where "middle_del(x,y) \ \fst(x),snd(y)\" relativize functional "middle_del" "middle_del_rel" relationalize "middle_del_rel" "is_middle_del" synthesize "is_middle_del" from_definition arity_theorem for "is_middle_del_fm" context M_basic begin lemma middle_del_in_cartprod: assumes "x\X" "y\X" "M(X)" shows "middle_del(x,y) \ ({0} \ \\X) \ ({0} \ \\X)" using assms Pow_rel_char transM[of _ X] fst_in_double_Union snd_in_double_Union unfolding middle_del_def by auto lemma lam_replacement_middle_del': assumes "\A[M]. separation(M,\y. \x[M]. x\A \ y = \x, middle_del(fst(x),snd(x))\)" shows "lam_replacement(M, \r . middle_del(fst(r),snd(r)))" using assms middle_del_in_cartprod by (rule_tac bounded_lam_replacement_binary[of _ "\X. ({0} \ \\X) \ ({0} \ \\X)"]) auto end \ \\<^locale>\M_basic\\ definition prodRepl :: "i\i\i" where "prodRepl(x,y) \ \snd(x),\fst(x),snd(y)\\" relativize functional "prodRepl" "prodRepl_rel" relationalize "prodRepl_rel" "is_prodRepl" synthesize "is_prodRepl" from_definition arity_theorem for "is_prodRepl_fm" context M_basic begin lemma prodRepl_in_cartprod: assumes "x\X" "y\X" "M(X)" shows "prodRepl(x,y) \ ({0} \ \\X) \ ({0} \ \\X) \ ({0} \ \\X)" using assms Pow_rel_char transM[of _ X] fst_in_double_Union snd_in_double_Union unfolding prodRepl_def by auto lemma lam_replacement_prodRepl': assumes "\A[M]. separation(M,\y. \x[M]. x\A \ y = \x, prodRepl(fst(x),snd(x))\)" shows "lam_replacement(M, \r . prodRepl(fst(r),snd(r)))" using assms prodRepl_in_cartprod by (rule_tac bounded_lam_replacement_binary[of _ "\X. ({0} \ \\X) \ ({0} \ \\X) \ ({0} \ \\X)"]) auto end \ \\<^locale>\M_basic\\ context M_Perm begin lemma inj_rel_in_Pow_rel_Pow_rel: assumes "x\X" "y\X" "M(X)" shows "inj\<^bsup>M\<^esup>(x,y) \ Pow\<^bsup>M\<^esup>(Pow\<^bsup>M\<^esup>(\X \ \X))" using assms transM[of _ X] mem_Pow_rel_abs inj_def Pi_def by clarsimp (use inj_rel_char in auto) lemma lam_replacement_inj_rel': assumes "\A[M]. separation(M,\y. \x[M]. x\A \ y = \x, inj\<^bsup>M\<^esup>(fst(x),snd(x))\)" shows "lam_replacement(M, \r . inj\<^bsup>M\<^esup>(fst(r),snd(r)))" using assms inj_rel_in_Pow_rel_Pow_rel by (rule_tac bounded_lam_replacement_binary[of _ "\X. Pow\<^bsup>M\<^esup>(Pow\<^bsup>M\<^esup>(\X \ \X))"]) auto end \ \\<^locale>\M_Perm\\ locale M_replacement = M_basic + assumes - lam_replacement_domain: "lam_replacement(M,domain)" - and lam_replacement_fst: "lam_replacement(M,fst)" and lam_replacement_snd: "lam_replacement(M,snd)" and lam_replacement_Union: "lam_replacement(M,Union)" and lam_replacement_middle_del: "lam_replacement(M, \r. middle_del(fst(r),snd(r)))" and lam_replacement_prodRepl: "lam_replacement(M, \r. prodRepl(fst(r),snd(r)))" and - lam_replacement_Upair:"lam_replacement(M, \p. Upair(fst(p),snd(p)))" - and - lam_replacement_Diff:"lam_replacement(M, \p. fst(p) - snd(p))" - and lam_replacement_Image:"lam_replacement(M, \p. fst(p) `` snd(p))" and middle_separation: "separation(M, \x. snd(fst(x))=fst(snd(x)))" and separation_fst_in_snd: "separation(M, \y. fst(snd(y)) \ snd(snd(y)))" - and - lam_replacement_converse : "lam_replacement(M,converse)" - and - lam_replacement_comp: "lam_replacement(M, \x. fst(x) O snd(x))" begin +\ \This lemma is similar to @{thm strong_lam_replacement_imp_strong_replacement} +and @{thm lam_replacement_imp_strong_replacement_aux} but does not require +\<^term>\g\ to be closed under \<^term>\M\.\ lemma lam_replacement_imp_strong_replacement: assumes "lam_replacement(M, f)" shows "strong_replacement(M, \x y. y = f(x))" proof - { fix A assume "M(A)" moreover from calculation assms obtain Y where 1:"M(Y)" "\b[M]. b \ Y \ (\x[M]. x \ A \ b = \x,f(x)\)" unfolding lam_replacement_def strong_replacement_def by auto moreover from this have "M({snd(b) . b \ Y})" using transM[OF _ \M(Y)\] lam_replacement_snd lam_replacement_imp_strong_replacement_aux RepFun_closed by simp moreover have "{snd(b) . b \ Y} = {y . x\A , M(f(x)) \ y=f(x)}" (is "?L=?R") proof(intro equalityI subsetI) fix x assume "x\?L" moreover from this obtain b where "b\Y" "x=snd(b)" "M(b)" using transM[OF _ \M(Y)\] by auto moreover from this 1 obtain a where "a\A" "b=\a,f(a)\" by auto moreover from calculation have "x=f(a)" by simp ultimately show "x\?R" by auto next fix z assume "z\?R" moreover from this obtain a where "a\A" "z=f(a)" "M(a)" "M(f(a))" using transM[OF _ \M(A)\] by auto moreover from calculation this 1 have "z=snd(\a,f(a)\)" "\a,f(a)\ \ Y" by auto ultimately show "z\?L" by force qed ultimately have "\Z[M]. \z[M]. z\Z \ (\a[M]. a\A \ z=f(a))" by (rule_tac rexI[where x="{snd(b) . b \ Y}"],auto) } then show ?thesis unfolding strong_replacement_def by simp qed lemma Collect_middle: "{p \ (\x\A. f(x)) \ (\x\{f(x) . x\A}. g(x)) . snd(fst(p))=fst(snd(p))} = { \\x,f(x)\,\f(x),g(f(x))\\ . x\A }" by (intro equalityI; auto simp:lam_def) lemma RepFun_middle_del: "{ \fst(fst(p)),snd(snd(p))\ . p \ { \\x,f(x)\,\f(x),g(f(x))\\ . x\A }} = { \x,g(f(x))\ . x\A }" by auto lemma lam_replacement_imp_RepFun: assumes "lam_replacement(M, f)" "M(A)" shows "M({y . x\A , M(y) \ y=f(x)})" proof - from assms obtain Y where 1:"M(Y)" "\b[M]. b \ Y \ (\x[M]. x \ A \ b = \x,f(x)\)" unfolding lam_replacement_def strong_replacement_def by auto moreover from this have "M({snd(b) . b \ Y})" using transM[OF _ \M(Y)\] lam_replacement_snd lam_replacement_imp_strong_replacement_aux RepFun_closed by simp moreover have "{snd(b) . b \ Y} = {y . x\A , M(y) \ y=f(x)}" (is "?L=?R") proof(intro equalityI subsetI) fix x assume "x\?L" moreover from this obtain b where "b\Y" "x=snd(b)" "M(b)" using transM[OF _ \M(Y)\] by auto moreover from this 1 obtain a where "a\A" "b=\a,f(a)\" by auto moreover from calculation have "x=f(a)" by simp ultimately show "x\?R" by auto next fix z assume "z\?R" moreover from this obtain a where "a\A" "z=f(a)" "M(a)" "M(f(a))" using transM[OF _ \M(A)\] by auto moreover from calculation this 1 have "z=snd(\a,f(a)\)" "\a,f(a)\ \ Y" by auto ultimately show "z\?L" by force qed ultimately show ?thesis by simp qed lemma lam_replacement_product: assumes "lam_replacement(M,f)" "lam_replacement(M,g)" shows "lam_replacement(M, \x. \f(x),g(x)\)" proof - { fix A let ?Y="{y . x\A , M(y) \ y=f(x)}" let ?Y'="{y . x\A ,M(y) \ y=\x,f(x)\}" let ?Z="{y . x\A , M(y) \ y=g(x)}" let ?Z'="{y . x\A ,M(y) \ y=\x,g(x)\}" have "x\C \ y\C \ fst(x) = fst(y) \ M(fst(y)) \ M(snd(x)) \ M(snd(y))" if "M(C)" for C y x using transM[OF _ that] by auto moreover note assms moreover assume "M(A)" moreover from \M(A)\ assms(1) have "M(converse(?Y'))" "M(?Y)" using lam_replacement_imp_RepFun_Lam lam_replacement_imp_RepFun by auto moreover from calculation have "M(?Z)" "M(?Z')" using lam_replacement_imp_RepFun_Lam lam_replacement_imp_RepFun by auto moreover from calculation have "M(converse(?Y')\?Z')" by simp moreover from this have "M({p \ converse(?Y')\?Z' . snd(fst(p))=fst(snd(p))})" (is "M(?P)") using middle_separation by simp moreover from calculation have "M({ \snd(fst(p)),\fst(fst(p)),snd(snd(p))\\ . p\?P })" (is "M(?R)") using RepFun_closed[OF lam_replacement_prodRepl[THEN lam_replacement_imp_strong_replacement] \M(?P)\ ] unfolding prodRepl_def by simp ultimately have "b \ ?R \ (\x[M]. x \ A \ b = \x,\f(x),g(x)\\)" if "M(b)" for b using that apply(intro iffI) apply(auto)[1] proof - assume " \x[M]. x \ A \ b = \x, f(x), g(x)\" moreover from this obtain x where "M(x)" "x\A" "b= \x, \f(x), g(x)\\" by auto moreover from calculation that have "M(\x,f(x)\)" "M(\x,g(x)\)" by auto moreover from calculation have "\f(x),x\ \ converse(?Y')" "\x,g(x)\ \ ?Z'" by auto moreover from calculation have "\\f(x),x\,\x,g(x)\\\converse(?Y')\?Z'" by auto moreover from calculation have "\\f(x),x\,\x,g(x)\\ \ ?P" (is "?p\?P") by auto moreover from calculation have "b = \snd(fst(?p)),\fst(fst(?p)),snd(snd(?p))\\" by auto moreover from calculation have "\snd(fst(?p)),\fst(fst(?p)),snd(snd(?p))\\\?R" by(rule_tac RepFunI[of ?p ?P], simp) ultimately show "b\?R" by simp qed with \M(?R)\ have "\Y[M]. \b[M]. b \ Y \ (\x[M]. x \ A \ b = \x,\f(x),g(x)\\)" by (rule_tac rexI[where x="?R"],simp_all) } with assms show ?thesis using lam_replacement_def strong_replacement_def by simp qed lemma lam_replacement_hcomp: assumes "lam_replacement(M,f)" "lam_replacement(M,g)" "\x[M]. M(f(x))" shows "lam_replacement(M, \x. g(f(x)))" proof - { fix A let ?Y="{y . x\A , y=f(x)}" let ?Y'="{y . x\A , y=\x,f(x)\}" have "\x\C. M(\fst(fst(x)),snd(snd(x))\)" if "M(C)" for C using transM[OF _ that] by auto moreover note assms moreover assume "M(A)" moreover from assms have eq:"?Y = {y . x\A ,M(y) \ y=f(x)}" "?Y' = {y . x\A ,M(y) \ y=\x,f(x)\}" using transM[OF _ \M(A)\] by auto moreover from \M(A)\ assms(1) have "M(?Y')" "M(?Y)" using lam_replacement_imp_RepFun_Lam lam_replacement_imp_RepFun eq by auto moreover from calculation have "M({z . y\?Y , M(z) \ z=\y,g(y)\})" (is "M(?Z)") using lam_replacement_imp_RepFun_Lam by auto moreover from calculation have "M(?Y'\?Z)" by simp moreover from this have "M({p \ ?Y'\?Z . snd(fst(p))=fst(snd(p))})" (is "M(?P)") using middle_separation by simp moreover from calculation have "M({ \fst(fst(p)),snd(snd(p))\ . p\?P })" (is "M(?R)") using RepFun_closed[OF lam_replacement_middle_del[THEN lam_replacement_imp_strong_replacement] \M(?P)\] unfolding middle_del_def by simp ultimately have "b \ ?R \ (\x[M]. x \ A \ b = \x,g(f(x))\)" if "M(b)" for b using that assms(3) apply(intro iffI) apply(auto)[1] proof - assume "\x[M]. x \ A \ b = \x, g(f(x))\" moreover from this obtain x where "M(x)" "x\A" "b= \x, g(f(x))\" by auto moreover from calculation that assms(3) have "M(f(x))" "M(g(f(x)))" by auto moreover from calculation have "\x,f(x)\ \ ?Y'" by auto moreover from calculation have "\f(x),g(f(x))\\?Z" by auto moreover from calculation have "\\x,f(x)\,\f(x),g(f(x))\\ \ ?P" (is "?p\?P") by auto moreover from calculation have "b = \fst(fst(?p)),snd(snd(?p))\" by auto moreover from calculation have "\fst(fst(?p)),snd(snd(?p))\\?R" by(rule_tac RepFunI[of ?p ?P], simp) ultimately show "b\?R" by simp qed with \M(?R)\ have "\Y[M]. \b[M]. b \ Y \ (\x[M]. x \ A \ b = \x,g(f(x))\)" by (rule_tac rexI[where x="?R"],simp_all) } with assms show ?thesis using lam_replacement_def strong_replacement_def by simp qed -lemma lam_replacement_Collect : - assumes "M(A)" "\x[M]. separation(M,F(x))" - "separation(M,\p . \x\A. x\snd(p) \ F(fst(p),x))" - shows "lam_replacement(M,\x. {y\A . F(x,y)})" -proof - - { - fix Z - let ?Y="\z.{x\A . F(z,x)}" - assume "M(Z)" - moreover from this - have "M(?Y(z))" if "z\Z" for z - using assms that transM[of _ Z] by simp - moreover from this - have "?Y(z)\Pow\<^bsup>M\<^esup>(A)" if "z\Z" for z - using Pow_rel_char that assms by auto - moreover from calculation \M(A)\ - have "M(Z\Pow\<^bsup>M\<^esup>(A))" by simp - moreover from this - have "M({p \ Z\Pow\<^bsup>M\<^esup>(A) . \x\A. x\snd(p) \ F(fst(p),x)})" (is "M(?P)") - using assms by simp - ultimately - have "b \ ?P \ (\z[M]. z\Z \ b=\z,?Y(z)\)" if "M(b)" for b - using assms(1) Pow_rel_char[OF \M(A)\] that - by(intro iffI,auto,intro equalityI,auto) - with \M(?P)\ - have "\Y[M]. \b[M]. b \ Y \ (\z[M]. z \ Z \ b = \z,?Y(z)\)" - by (rule_tac rexI[where x="?P"],simp_all) - } - then - show ?thesis - unfolding lam_replacement_def strong_replacement_def - by simp -qed - lemma lam_replacement_hcomp2: assumes "lam_replacement(M,f)" "lam_replacement(M,g)" "\x[M]. M(f(x))" "\x[M]. M(g(x))" "lam_replacement(M, \p. h(fst(p),snd(p)))" - "\x[M]. \y[M]. M(h(x,y))" shows "lam_replacement(M, \x. h(f(x),g(x)))" using assms lam_replacement_product[of f g] lam_replacement_hcomp[of "\x. \f(x), g(x)\" "\\x,y\. h(x,y)"] unfolding split_def by simp lemma lam_replacement_identity: "lam_replacement(M,\x. x)" proof - { fix A assume "M(A)" moreover from this have "id(A) = {\snd(fst(z)),fst(snd(z))\ . z\ {z\ (A\A)\(A\A). snd(fst(z)) = fst(snd(z))}}" unfolding id_def lam_def by(intro equalityI subsetI,simp_all,auto) moreover from calculation have "M({z\ (A\A)\(A\A). snd(fst(z)) = fst(snd(z))})" (is "M(?A')") using middle_separation by simp moreover from calculation have "M({\snd(fst(z)),fst(snd(z))\ . z\ ?A'})" using transM[of _ A] lam_replacement_product lam_replacement_hcomp lam_replacement_fst lam_replacement_snd lam_replacement_imp_strong_replacement[THEN RepFun_closed] by simp_all ultimately have "M(id(A))" by simp } then show ?thesis using lam_replacement_iff_lam_closed unfolding id_def by simp qed lemma strong_replacement_separation_aux : assumes "strong_replacement(M,\ x y . y=f(x))" "separation(M,P)" shows "strong_replacement(M, \x y . P(x) \ y=f(x))" proof - { fix A let ?Q="\X. \b[M]. b \ X \ (\x[M]. x \ A \ P(x) \ b = f(x))" assume "M(A)" moreover from this have "M({x\A . P(x)})" (is "M(?B)") using assms by simp moreover from calculation assms obtain Y where "M(Y)" "\b[M]. b \ Y \ (\x[M]. x \ ?B \ b = f(x))" unfolding strong_replacement_def by auto then have "\Y[M]. \b[M]. b \ Y \ (\x[M]. x \ A \ P(x) \ b = f(x))" using rexI[of ?Q _ M] by simp } then show ?thesis unfolding strong_replacement_def by simp qed lemma separation_in: assumes "\x[M]. M(f(x))" "lam_replacement(M,f)" "\x[M]. M(g(x))" "lam_replacement(M,g)" shows "separation(M,\x . f(x)\g(x))" proof - let ?Z="\A. {\x,\f(x),g(x)\\. x\A}" have "M(?Z(A))" if "M(A)" for A using assms lam_replacement_iff_lam_closed that lam_replacement_product[of f g] unfolding lam_def by auto then have "M({u\?Z(A) . fst(snd(u)) \snd(snd(u))})" (is "M(?W(A))") if "M(A)" for A using that separation_fst_in_snd assms by auto then have "M({fst(u) . u \ ?W(A)})" if "M(A)" for A using that lam_replacement_imp_strong_replacement[OF lam_replacement_fst,THEN RepFun_closed] fst_closed[OF transM] by auto moreover have "{x\A. f(x)\g(x)} = {fst(u) . u\?W(A)}" for A by auto ultimately show ?thesis using separation_iff by auto qed lemma lam_replacement_swap: "lam_replacement(M, \x. \snd(x),fst(x)\)" using lam_replacement_fst lam_replacement_snd lam_replacement_product[of "snd" "fst"] by simp -lemma lam_replacement_range : "lam_replacement(M,range)" - unfolding range_def - using lam_replacement_hcomp[OF lam_replacement_converse lam_replacement_domain] - by auto - -lemma separation_in_range : "M(a) \ separation(M, \x. a\range(x))" - using lam_replacement_range lam_replacement_constant separation_in - by auto - -lemma separation_in_domain : "M(a) \ separation(M, \x. a\domain(x))" - using lam_replacement_domain lam_replacement_constant separation_in - by auto - -lemma lam_replacement_separation : - assumes "lam_replacement(M,f)" "separation(M,P)" - shows "strong_replacement(M, \x y . P(x) \ y=\x,f(x)\)" - using strong_replacement_separation_aux assms - unfolding lam_replacement_def - by simp - -lemmas strong_replacement_separation = - strong_replacement_separation_aux[OF lam_replacement_imp_strong_replacement] - -lemma id_closed: "M(A) \ M(id(A))" - using lam_replacement_identity lam_replacement_iff_lam_closed - unfolding id_def by simp lemma relation_separation: "separation(M, \z. \x y. z = \x, y\)" unfolding separation_def proof (clarify) fix A assume "M(A)" moreover from this have "{z\A. \x y. z = \x, y\} = {z\A. \x\domain(A). \y\range(A). pair(M, x, y, z)}" (is "?rel = _") by (intro equalityI, auto dest:transM) (intro bexI, auto dest:transM simp:Pair_def) moreover from calculation have "M(?rel)" using cartprod_separation[THEN separation_closed, of "domain(A)" "range(A)" A] by simp ultimately show "\y[M]. \x[M]. x \ y \ x \ A \ (\w y. x = \w, y\)" by (rule_tac x="{z\A. \x y. z = \x, y\}" in rexI) auto qed lemma separation_Pair: assumes "separation(M, \y . P(fst(y), snd(y)))" shows "separation(M, \y. \ u v . y=\u,v\ \ P(u,v))" unfolding separation_def proof(clarify) fix A assume "M(A)" moreover from this have "M({z\A. \x y. z = \x, y\})" (is "M(?P)") using relation_separation by simp moreover from this assms have "M({z\?P . P(fst(z),snd(z))})" by(rule_tac separation_closed,simp_all) moreover have "{y\A . \ u v . y=\u,v\ \ P(u,v) } = {z\?P . P(fst(z),snd(z))}" by(rule equalityI subsetI,auto) moreover from calculation have "M({y\A . \ u v . y=\u,v\ \ P(u,v) })" by simp ultimately show "\y[M]. \x[M]. x \ y \ x \ A \ (\w y. x = \w, y\ \ P(w,y))" by (rule_tac x="{z\A. \x y. z = \x, y\ \ P(x,y)}" in rexI) auto qed +lemma lam_replacement_separation : + assumes "lam_replacement(M,f)" "separation(M,P)" + shows "strong_replacement(M, \x y . P(x) \ y=\x,f(x)\)" + using strong_replacement_separation_aux assms + unfolding lam_replacement_def + by simp + +lemmas strong_replacement_separation = + strong_replacement_separation_aux[OF lam_replacement_imp_strong_replacement] + +lemma id_closed: "M(A) \ M(id(A))" + using lam_replacement_identity lam_replacement_iff_lam_closed + unfolding id_def by simp + lemma lam_replacement_Pair: shows "lam_replacement(M, \x. \fst(x), snd(x)\)" - unfolding lam_replacement_def strong_replacement_def -proof (clarsimp) - fix A - assume "M(A)" + using lam_replacement_product lam_replacement_fst lam_replacement_snd + by simp + +lemma lam_replacement_Upair: "lam_replacement(M, \p. Upair(fst(p), snd(p)))" +proof - + have "(\(\a b . x = )) \ fst(x) = 0 \ snd(x) = 0" for x + unfolding fst_def snd_def + by (simp add:the_0) then - show "\Y[M]. \b[M]. b \ Y \ (\x\A. b = \x, fst(x), snd(x)\)" - unfolding lam_replacement_def strong_replacement_def - proof (cases "relation(A)") - case True - with \M(A)\ - show ?thesis - using id_closed unfolding relation_def - by (rule_tac x="id(A)" in rexI) auto - next - case False - moreover - note \M(A)\ - moreover from this - have "M({z\A. \x y. z = \x, y\})" (is "M(?rel)") - using relation_separation by auto - moreover - have "z = \fst(z), snd(z)\" if "fst(z) \ 0 \ snd(z) \ 0" for z - using that - by (cases "\a b. z=\a,b\") (auto simp add: the_0 fst_def snd_def) - ultimately - show ?thesis - using id_closed unfolding relation_def - by (rule_tac x="id(?rel) \ (A-?rel)\{0}\{0}" in rexI) - (force simp:fst_def snd_def)+ - qed + have "Upair(fst(x),snd(x)) = (if (\w . \ z . x=) then (\x) else ({0}))" for x + using fst_conv snd_conv Upair_eq_cons + by (auto simp add:Pair_def) + moreover + have "lam_replacement(M, \x . (if (\w . \ z . x=) then (\x) else ({0})))" + using lam_replacement_Union lam_replacement_constant relation_separation lam_replacement_if + by simp + ultimately + show ?thesis + using lam_replacement_cong + by simp qed lemma lam_replacement_Un: "lam_replacement(M, \p. fst(p) \ snd(p))" using lam_replacement_Upair lam_replacement_Union lam_replacement_hcomp[where g=Union and f="\p. Upair(fst(p),snd(p))"] unfolding Un_def by simp lemma lam_replacement_cons: "lam_replacement(M, \p. cons(fst(p),snd(p)))" using lam_replacement_Upair lam_replacement_hcomp2[of _ _ "(\)"] lam_replacement_hcomp2[of fst fst "Upair"] lam_replacement_Un lam_replacement_fst lam_replacement_snd unfolding cons_def by auto lemma lam_replacement_sing_fun: assumes "lam_replacement(M, f)" "\x[M]. M(f(x))" shows "lam_replacement(M, \x. {f(x)})" using lam_replacement_constant lam_replacement_cons lam_replacement_hcomp2[of "f" "\_. 0" cons] assms by auto lemma lam_replacement_sing: "lam_replacement(M, \x. {x})" using lam_replacement_sing_fun lam_replacement_identity by simp - lemma lam_replacement_CartProd: assumes "lam_replacement(M,f)" "lam_replacement(M,g)" "\x[M]. M(f(x))" "\x[M]. M(g(x))" shows "lam_replacement(M, \x. f(x) \ g(x))" proof - note rep_closed = lam_replacement_imp_strong_replacement[THEN RepFun_closed] { fix A assume "M(A)" moreover note transM[OF _ \M(A)\] moreover from calculation assms have "M({\x,\f(x),g(x)\\ . x\A})" (is "M(?A')") using lam_replacement_product[THEN lam_replacement_imp_lam_closed[unfolded lam_def]] by simp moreover from calculation have "M(\{f(x) . x\A})" (is "M(?F)") using rep_closed[OF assms(1)] assms(3) by simp moreover from calculation have "M(\{g(x) . x\A})" (is "M(?G)") using rep_closed[OF assms(2)] assms(4) by simp moreover from calculation have "M(?A' \ (?F \ ?G))" (is "M(?T)") by simp moreover from this have "M({t \ ?T . fst(snd(t)) \ fst(snd(fst(t))) \ snd(snd(t)) \ snd(snd(fst(t)))})" (is "M(?Q)") using lam_replacement_hcomp[OF lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_snd] _ ] - lam_replacement_hcomp lam_replacement_identity lam_replacement_fst lam_replacement_snd + lam_replacement_hcomp lam_replacement_identity lam_replacement_fst lam_replacement_snd separation_in separation_conj by simp moreover from this have "M({\fst(fst(t)),snd(t)\ . t\?Q})" (is "M(?R)") - using rep_closed lam_replacement_Pair[THEN [5] lam_replacement_hcomp2] + using rep_closed lam_replacement_product lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_fst] lam_replacement_snd transM[of _ ?Q] by simp moreover from calculation have "M({\x,?R``{x}\ . x\A})" using lam_replacement_imp_lam_closed[unfolded lam_def] lam_replacement_sing lam_replacement_Image[THEN [5] lam_replacement_hcomp2] lam_replacement_constant[of ?R] by simp moreover have "?R``{x} = f(x)\g(x)" if "x\A" for x by(rule equalityI subsetI,force,rule subsetI,rule_tac a="x" in imageI) (auto simp:that,(rule_tac rev_bexI[of x],simp_all add:that)+) ultimately have "M({\x,f(x) \ g(x)\ . x\A})" by auto } with assms show ?thesis using lam_replacement_iff_lam_closed[THEN iffD2,unfolded lam_def] by simp qed lemma lam_replacement_RepFun : assumes "lam_replacement(M,\x . f(fst(x),snd(x)))" "lam_replacement(M,g)" - "\x[M].\y[M].M(f(x,y))" "\x[M].M(g(x))" + "\x[M].\y\g(x).M(f(x,y))" "\x[M].M(g(x))" shows "lam_replacement(M,\x . {f(x,z) . z\g(x)})" proof - + note rep_closed = lam_replacement_imp_strong_replacement[THEN RepFun_closed] from assms have "lam_replacement(M,\x.{x}\g(x))" using lam_replacement_CartProd lam_replacement_sing by auto moreover from assms have "M({f(x,z) . z \ g(x)})" if "M(x)" for x - using that transM[of _ "g(_)"] - lam_replacement_imp_strong_replacement + using that transM[of _ "g(_)"] rep_closed lam_replacement_product[OF lam_replacement_fst] lam_replacement_snd lam_replacement_identity assms(1)[THEN [5] lam_replacement_hcomp2] lam_replacement_constant[of x] - by(rule_tac RepFun_closed,simp_all) + by simp moreover have "M(\z\A.{f(z,u) . u \ g(z)})" if "M(A)" for A proof - from that assms calculation have "M(\{{x}\g(x) . x\A})" (is "M(?C)") - using lam_replacement_imp_strong_replacement transM[of _ A] RepFun_closed + using rep_closed transM[of _ A] by simp with assms \M(A)\ have "M({\fst(x),f(fst(x),snd(x))\ . x \?C})" (is "M(?B)") - using singleton_closed transM[of _ A] lam_replacement_imp_strong_replacement + using singleton_closed transM[of _ A] transM[of _ "g(_)"] rep_closed lam_replacement_product[OF lam_replacement_fst] lam_replacement_hcomp[OF lam_replacement_snd] - transM[of _ "g(_)"] RepFun_closed by simp with \M(A)\ have "M({\x,?B``{x}\ . x\A})" - using transM[of _ A] + using transM[of _ A] rep_closed lam_replacement_product[OF lam_replacement_identity] lam_replacement_Image[THEN [5] lam_replacement_hcomp2] - lam_replacement_constant lam_replacement_sing RepFun_closed[OF lam_replacement_imp_strong_replacement] + lam_replacement_constant lam_replacement_sing by simp moreover have "?B``{z} = {f(z,u) . u \ g(z)}" if "z\A" for z using that by(intro equalityI subsetI,auto simp:imageI) moreover from this have "{\x,?B``{x}\ . x\A} = {\z,{f(z,u) . u \ g(z)}\ . z\A}" by auto ultimately show ?thesis unfolding lam_def by auto qed ultimately show ?thesis using lam_replacement_iff_lam_closed[THEN iffD2] by simp qed +lemma lam_replacement_Collect : + assumes "lam_replacement(M,f)" "\x[M].M(f(x))" + "separation(M,\x . F(fst(x),snd(x)))" + shows "lam_replacement(M,\x. {y\f(x) . F(x,y)})" +proof - + note rep_closed = lam_replacement_imp_strong_replacement[THEN RepFun_closed] + from assms + have 1:"lam_replacement(M,\x.{x}\f(x))" + using lam_replacement_CartProd lam_replacement_sing + by auto + have "M({u\f(x) . F(x,u)})" if "M(x)" for x + proof - + from assms that + have "M({u \ {x}\f(x) . F(fst(u),snd(u))})" (is "M(?F)") + by simp + with assms that + have "M({snd(u) . u \ ?F})" + using rep_closed transM[of _ "f(x)"] lam_replacement_snd + by simp + moreover + have "{snd(u) . u \ ?F} = {u\f(x) . F(x,u)}" + by auto + ultimately show ?thesis by simp + qed + moreover + have "M(\z\A.{y \ f(z) . F(z,y)})" if "M(A)" for A + proof - + from 1 that assms + have "M(\{{x}\f(x) . x\A})" (is "M(?C)") + using rep_closed transM[of _ A] + by simp + moreover from this assms + have "M({p \ ?C . F(fst(p),snd(p))})" (is "M(?B)") + by(rule_tac separation_closed,simp_all) + with \M(A)\ + have "M({\x,?B``{x}\ . x\A})" + using transM[of _ A] rep_closed + lam_replacement_product[OF lam_replacement_identity] + lam_replacement_Image[THEN [5] lam_replacement_hcomp2] + lam_replacement_constant lam_replacement_sing + by simp + moreover + have "?B``{z} = {u\f(z) . F(z,u)}" if "z\A" for z + using that + by(intro equalityI subsetI,auto simp:imageI) + moreover from this + have "{\x,?B``{x}\ . x\A} = {\z,{u\f(z) . F(z,u)}\ . z\A}" + by auto + ultimately + show ?thesis + unfolding lam_def by auto + qed + ultimately + show ?thesis + using lam_replacement_iff_lam_closed[THEN iffD2] + by simp +qed + + +lemma separation_eq: + assumes "\x[M]. M(f(x))" "lam_replacement(M,f)" + "\x[M]. M(g(x))" "lam_replacement(M,g)" + shows "separation(M,\x . f(x) = g(x))" +proof - + let ?Z="\A. {\x,\f(x),\g(x),x\\\. x\A}" + let ?Y="\A. {\\x,f(x)\,\g(x),x\\. x\A}" + note sndsnd = lam_replacement_hcomp[OF lam_replacement_snd lam_replacement_snd] + note fstsnd = lam_replacement_hcomp[OF lam_replacement_snd lam_replacement_fst] + note sndfst = lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_snd] + have "M(?Z(A))" if "M(A)" for A + using assms lam_replacement_iff_lam_closed that + lam_replacement_product[OF assms(2) + lam_replacement_product[OF assms(4) lam_replacement_identity]] + unfolding lam_def + by auto + moreover + have "?Y(A) = {\\fst(x), fst(snd(x))\, fst(snd(snd(x))), snd(snd(snd(x)))\ . x \ ?Z(A)}" for A + by auto + moreover from calculation + have "M(?Y(A))" if "M(A)" for A + using + lam_replacement_imp_strong_replacement[OF + lam_replacement_product[OF + lam_replacement_product[OF lam_replacement_fst fstsnd] + lam_replacement_product[OF + lam_replacement_hcomp[OF sndsnd lam_replacement_fst] + lam_replacement_hcomp[OF lam_replacement_snd sndsnd] + ] + ], THEN RepFun_closed,simplified,of "?Z(A)"] + fst_closed[OF transM] snd_closed[OF transM] that + by auto + then + have "M({u\?Y(A) . snd(fst(u)) = fst(snd(u))})" (is "M(?W(A))") if "M(A)" for A + using that middle_separation assms + by auto + then + have "M({fst(fst(u)) . u \ ?W(A)})" if "M(A)" for A + using that lam_replacement_imp_strong_replacement[OF + lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_fst], THEN RepFun_closed] + fst_closed[OF transM] + by auto + moreover + have "{x\A. f(x) = g(x)} = {fst(fst(u)) . u\?W(A)}" for A + by auto + ultimately + show ?thesis + using separation_iff by auto +qed +lemma separation_comp : + assumes "separation(M,P)" "lam_replacement(M,f)" "\x[M]. M(f(x))" + shows "separation(M,\x. P(f(x)))" + unfolding separation_def +proof(clarify) + fix A + assume "M(A)" + let ?B="{f(a) . a \ A}" + let ?C="A\{b\?B . P(b)}" + note \M(A)\ + moreover from calculation + have "M(?C)" + using lam_replacement_imp_strong_replacement assms RepFun_closed transM[of _ A] + by simp + moreover from calculation + have "M({p\?C . f(fst(p)) = snd(p)})" (is "M(?Prod)") + using assms separation_eq lam_replacement_fst lam_replacement_snd + lam_replacement_hcomp + by simp + moreover from calculation + have "M({fst(p) . p\?Prod})" (is "M(?L)") + using lam_replacement_imp_strong_replacement lam_replacement_fst RepFun_closed + transM[of _ ?Prod] + by simp + moreover + have "?L = {z\A . P(f(z))}" + by(intro equalityI subsetI,auto) + ultimately + show " \y[M]. \z[M]. z \ y \ z \ A \ P(f(z))" + by (rule_tac x="?L" in rexI,simp_all) +qed + +lemma lam_replacement_domain: "lam_replacement(M,domain)" +proof - + have 1:"{fst(z) . z\ {u\x . (\ a b. u = )}} =domain(x)" for x + unfolding domain_def + by (intro equalityI subsetI,auto,rule_tac x="" in bexI,auto) + moreover + have "lam_replacement(M,\ x .{fst(z) . z\ {u\x . (\ a b. u = )}})" + using lam_replacement_hcomp lam_replacement_snd lam_replacement_fst +snd_closed[OF transM] fst_closed[OF transM] lam_replacement_identity +relation_separation relation_separation[THEN separation_comp,where f=snd] +lam_replacement_Collect + by(rule_tac lam_replacement_RepFun,simp_all) + ultimately + show ?thesis + using lam_replacement_cong + by auto +qed + +lemma separation_in_domain : "M(a) \ separation(M, \x. a\domain(x))" + using lam_replacement_domain lam_replacement_constant separation_in + by auto + + +lemma separation_all: + assumes "separation(M, \x . P(fst(fst(x)),snd(x)))" + "lam_replacement(M,f)" "lam_replacement(M,g)" + "\x[M]. M(f(x))" "\x[M]. M(g(x))" + shows "separation(M, \z. \x\f(z). P(g(z),x))" + unfolding separation_def +proof(clarify) + note rep_closed = lam_replacement_imp_strong_replacement[THEN RepFun_closed] + note lr_fs = lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_snd] + fix A + assume "M(A)" + with assms + have "M({f(x) . x\A})" (is "M(?F)") "M({ . x\A})" (is "M(?G)") + using rep_closed transM[of _ A] + lam_replacement_product lam_replacement_imp_lam_closed + unfolding lam_def + by auto + let ?B="\?F" + let ?C="?G\?B" + note \M(A)\ \M(?F)\ \M(?G)\ + moreover from calculation + have "M(?C)" + by simp + moreover from calculation + have "M({p\?C . P(fst(fst(p)),snd(p)) \ snd(p)\snd(fst(p))})" (is "M(?Prod)") + using assms separation_conj separation_in lam_replacement_fst lam_replacement_snd + lam_replacement_hcomp + by simp + moreover from calculation + have "M({z\?G . snd(z)=?Prod``{z}})" (is "M(?L)") + using separation_eq + lam_replacement_constant[of ?Prod] lam_replacement_constant[of 0] + lam_replacement_hcomp[OF _ lam_replacement_sing] + lam_replacement_Image[THEN [5] lam_replacement_hcomp2] + lam_replacement_identity lam_replacement_snd + by simp + moreover from calculation assms + have "M({x\A . \ ?L})" (is "M(?N)") + using lam_replacement_product lam_replacement_constant + by (rule_tac separation_closed,rule_tac separation_in,auto) + moreover from calculation + have "?N = {z\A . \x\f(z). P(g(z),x)}" + proof - + have "P(g(z),x)" if "z\A" "x\f(z)" "f(z) = ?Prod``{}" for z x + using that + by(rule_tac imageE[of x ?Prod "{}"],simp_all) + moreover + have "f(z) = ?Prod `` {}" if "z\A" "\x\f(z). P(g(z), x)" for z + using that + by force + ultimately + show ?thesis + by auto + qed + ultimately + show " \y[M]. \z[M]. z \ y \ z \ A \ (\x\f(z) . P(g(z),x))" + by (rule_tac x="?N" in rexI,simp_all) +qed + +lemma separation_ex: + assumes "separation(M, \x . P(fst(fst(x)),snd(x)))" + "lam_replacement(M,f)" "lam_replacement(M,g)" + "\x[M]. M(f(x))" "\x[M]. M(g(x))" + shows "separation(M, \z. \x\f(z). P(g(z),x))" + unfolding separation_def +proof(clarify) + note rep_closed = lam_replacement_imp_strong_replacement[THEN RepFun_closed] + note lr_fs = lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_snd] + fix A + assume "M(A)" + with assms + have "M({f(x) . x\A})" (is "M(?F)") "M({ . x\A})" (is "M(?G)") + using rep_closed transM[of _ A] + lam_replacement_product lam_replacement_imp_lam_closed + unfolding lam_def + by auto + let ?B="\?F" + let ?C="?G\?B" + note \M(A)\ \M(?F)\ \M(?G)\ + moreover from calculation + have "M(?C)" + by simp + moreover from calculation + have "M({p\?C . P(fst(fst(p)),snd(p)) \ snd(p)\snd(fst(p))})" (is "M(?Prod)") + using assms separation_conj separation_in lam_replacement_fst lam_replacement_snd + lam_replacement_hcomp + by simp + moreover from calculation + have "M({z\?G . 0\?Prod``{z}})" (is "M(?L)") + using separation_eq separation_neg + lam_replacement_constant[of ?Prod] lam_replacement_constant[of 0] + lam_replacement_hcomp[OF _ lam_replacement_sing] + lam_replacement_Image[THEN [5] lam_replacement_hcomp2] + lam_replacement_identity + by simp + moreover from calculation assms + have "M({x\A . \ ?L})" (is "M(?N)") + using lam_replacement_product lam_replacement_constant + by (rule_tac separation_closed,rule_tac separation_in,auto) + moreover from calculation + have "?N = {z\A . \x\f(z). P(g(z),x)}" + by(intro equalityI subsetI,simp_all,force) (rule ccontr,force) + ultimately + show " \y[M]. \z[M]. z \ y \ z \ A \ (\x\f(z) . P(g(z),x))" + by (rule_tac x="?N" in rexI,simp_all) +qed + +lemma lam_replacement_converse: "lam_replacement(M,converse)" +proof- + note lr_Un = lam_replacement_Union[THEN [2] lam_replacement_hcomp] + note lr_fs = lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_snd] + note lr_ff = lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_fst] + note lr_ss = lam_replacement_hcomp[OF lam_replacement_snd lam_replacement_snd] + note lr_sf = lam_replacement_hcomp[OF lr_ff lam_replacement_snd] + note lr_fff = lam_replacement_hcomp[OF lam_replacement_fst lr_ff] + note lr_f4 = lam_replacement_hcomp[OF lr_ff lr_ff] + have eq:"converse(x) = {\snd(y),fst(y)\ . y \ {z \ x. \ u \\\x . \v\\\x . z=}}" for x + unfolding converse_def + by(intro equalityI subsetI,auto,rename_tac a b,rule_tac x="" in bexI,simp_all) + (auto simp : Pair_def) + have "M(x) \ M({z \ x. \ u \\\x . \v\\\x . z=})" for x + using lam_replacement_identity lr_Un[OF lr_Un] lr_Un lr_Un[OF lam_replacement_Union] + lam_replacement_snd lr_ff lam_replacement_fst lam_replacement_hcomp lr_fff + lam_replacement_product[OF lr_sf lam_replacement_snd] lr_f4 + lam_replacement_hcomp[OF lr_f4 lam_replacement_snd] separation_eq + lam_replacement_constant + by(rule_tac separation_closed,rule_tac separation_ex,rule_tac separation_ex,simp_all) + moreover + have sep:"separation(M, \x. \u\\\fst(x). \v\\\fst(x). snd(x) = \u, v\)" + using lam_replacement_identity lr_Un[OF lr_Un] lr_Un + lam_replacement_snd lr_ff lam_replacement_fst lam_replacement_hcomp lr_fff + lam_replacement_product[OF lr_sf lam_replacement_snd] + lam_replacement_hcomp[OF lr_f4 lam_replacement_snd] separation_eq + by(rule_tac separation_ex,rule_tac separation_ex,simp_all) + moreover from calculation + have "lam_replacement(M, \x. {z \ x . \u\\\x. \v\\\x. z = \u, v\})" + using lam_replacement_identity + by(rule_tac lam_replacement_Collect,simp_all) + ultimately + have 1:"lam_replacement(M, \x . {\snd(y),fst(y)\ . y \ {z \ x. \u \\\x . \v\\\x. z=}})" + using lam_replacement_product lam_replacement_fst lam_replacement_hcomp + lam_replacement_identity lr_ff lr_ss + lam_replacement_snd lam_replacement_identity sep + by(rule_tac lam_replacement_RepFun,simp_all,force dest:transM) + with eq[symmetric] + show ?thesis + by(rule_tac lam_replacement_cong[OF 1],simp_all) +qed + +lemma lam_replacement_Diff: "lam_replacement(M, \x . fst(x) - snd(x))" +proof - + have eq:"X - Y = {x\X . x\Y}" for X Y + by auto + moreover + have "lam_replacement(M, \x . {y \ fst(x) . y\snd(x)})" + using lam_replacement_fst lam_replacement_hcomp + lam_replacement_snd lam_replacement_identity separation_in separation_neg + by(rule_tac lam_replacement_Collect,simp_all) + then + show ?thesis + by(rule_tac lam_replacement_cong,auto,subst eq[symmetric],simp) +qed + +lemma lam_replacement_range : "lam_replacement(M,range)" + unfolding range_def + using lam_replacement_hcomp[OF lam_replacement_converse lam_replacement_domain] + by auto + +lemma separation_in_range : "M(a) \ separation(M, \x. a\range(x))" + using lam_replacement_range lam_replacement_constant separation_in + by auto + lemmas tag_replacement = lam_replacement_constant[unfolded lam_replacement_def] +lemma tag_lam_replacement : "M(X) \ lam_replacement(M,\x. )" + using lam_replacement_product[OF lam_replacement_constant] lam_replacement_identity + by simp + +lemma strong_lam_replacement_imp_lam_replacement_RepFun : + assumes "strong_replacement(M,\ x z . P(fst(x),snd(x)) \ z=\x,f(fst(x),snd(x))\)" + "lam_replacement(M,g)" + "\A y . M(y) \ M(A) \ \x\A. P(y,x) \ M(f(y,x))" + "\x[M]. M(g(x))" + "separation(M, \x. P(fst(x),snd(x)))" + shows "lam_replacement(M, \x. {y . r\ g(x) , P(x,r) \ y=f(x,r)}) " +proof - + note rep_closed = lam_replacement_imp_strong_replacement[THEN RepFun_closed] + moreover + have "{f(x, xa) . xa \ {xa \ g(x) . P(x, xa)}} = {y . z \ g(x), P(x, z) \ y = f(x, z)}" for x + by(intro equalityI subsetI,auto) + moreover from assms + have 0:"M({xa \ g(x) . P(x, xa)})" if "M(x)" for x + using that separation_closed assms(5)[THEN separation_comp,OF tag_lam_replacement] + by simp + moreover from assms + have 1:"lam_replacement(M,\x.{x}\{u\g(x) . P(x,u)})" (is "lam_replacement(M,\x.?R(x))") + using separation_closed assms(5)[THEN separation_comp,OF tag_lam_replacement] + by(rule_tac lam_replacement_CartProd[OF lam_replacement_sing lam_replacement_Collect],simp_all) + moreover from assms + have "M({y . z\g(x) , P(x,z) \ y=f(x,z)})" (is "M(?Q(x))") if "M(x)" for x + using that transM[of _ "g(_)"] + separation_closed assms(5)[THEN separation_comp,OF tag_lam_replacement] + assms(3)[of "x" "g(x)"] strong_lam_replacement_imp_strong_replacement + by simp + moreover + have "M(\z\A.{f(z,r) . r \ {u\ g(z) . P(z,u)}})" if "M(A)" for A + proof - + from that assms calculation + have "M(\{?R(x) . x\A})" (is "M(?C)") + using transM[of _ A] rep_closed + by simp + moreover from assms \M(A)\ + have "x \ {y} \ {x \ g(y) . P(y, x)} \ M(x) \ M(f(fst(x),snd(x)))" if "y\A" for y x + using assms(3)[of "y" "g(y)"] transM[of _ A] transM[of _ "g(y)"] that + by force + moreover from this + have "\y\A . x \ {y} \ {x \ g(y) . P(y, x)} \ M(x) \ M(f(fst(x),snd(x)))" for x + by auto + moreover note assms \M(A)\ + ultimately + have "M({z . x\?C , P(fst(x),snd(x)) \ z = \x,f(fst(x),snd(x))\})" (is "M(?B)") + using singleton_closed transM[of _ A] transM[of _ "g(_)"] rep_closed + lam_replacement_product[OF lam_replacement_fst] + lam_replacement_hcomp[OF lam_replacement_snd] transM[OF _ 0] + by(rule_tac strong_replacement_closed,simp_all) + then + have "M({\fst(fst(x)),snd(x)\ . x\?B})" (is "M(?D)") + using rep_closed transM[of _ ?B] + lam_replacement_product[OF + lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_fst] + lam_replacement_snd] + by simp + with \M(A)\ + have "M({\x,?D``{x}\ . x\A})" + using transM[of _ A] rep_closed + lam_replacement_product[OF lam_replacement_identity] + lam_replacement_Image[THEN [5] lam_replacement_hcomp2] + lam_replacement_constant lam_replacement_sing + by simp + moreover from calculation + have "?D``{z} = {f(z,r) . r \ {u\ g(z) . P(z,u)}}" if "z\A" for z + using that + by (intro equalityI subsetI,auto,intro imageI,force,auto) + moreover from this + have "{\x,?D``{x}\ . x\A} = {\z,{f(z,r) . r \ {u\ g(z) . P(z,u)}}\ . z\A}" + by auto + ultimately + show ?thesis + unfolding lam_def by auto + qed + ultimately + show ?thesis + using lam_replacement_iff_lam_closed[THEN iffD2] + by simp +qed + +lemma lam_replacement_Collect' : + assumes "M(A)" "separation(M,\p . F(fst(p),snd(p)))" + shows "lam_replacement(M,\x. {y\A . F(x,y)})" + using assms lam_replacement_constant + by(rule_tac lam_replacement_Collect, simp_all) + lemma lam_replacement_id2: "lam_replacement(M, \x. \x, x\)" using lam_replacement_identity lam_replacement_product[of "\x. x" "\x. x"] by simp lemmas id_replacement = lam_replacement_id2[unfolded lam_replacement_def] lemma lam_replacement_apply2:"lam_replacement(M, \p. fst(p) ` snd(p))" using lam_replacement_fst lam_replacement_sing_fun[OF lam_replacement_snd] lam_replacement_Image lam_replacement_Union lam_replacement_hcomp[of _ Union] lam_replacement_hcomp2[of _ _ "(``)"] unfolding apply_def by simp lemma lam_replacement_apply:"M(S) \ lam_replacement(M, \x. S ` x)" using lam_replacement_constant[of S] lam_replacement_identity lam_replacement_apply2[THEN [5] lam_replacement_hcomp2] by simp lemma apply_replacement:"M(S) \ strong_replacement(M, \x y. y = S ` x)" using lam_replacement_apply lam_replacement_imp_strong_replacement by simp lemma lam_replacement_id_const: "M(b) \ lam_replacement(M, \x. \x, b\)" using lam_replacement_identity lam_replacement_constant lam_replacement_product[of "\x. x" "\x. b"] by simp lemmas pospend_replacement = lam_replacement_id_const[unfolded lam_replacement_def] lemma lam_replacement_const_id: "M(b) \ lam_replacement(M, \z. \b, z\)" using lam_replacement_identity lam_replacement_constant lam_replacement_product[of "\x. b" "\x. x"] by simp lemmas prepend_replacement = lam_replacement_const_id[unfolded lam_replacement_def] lemma lam_replacement_apply_const_id: "M(f) \ M(z) \ lam_replacement(M, \x. f ` \z, x\)" using lam_replacement_const_id[of z] lam_replacement_apply lam_replacement_hcomp[of "\x. \z, x\"] by simp lemmas apply_replacement2 = lam_replacement_apply_const_id[unfolded lam_replacement_def] lemma lam_replacement_Inl: "lam_replacement(M, Inl)" using lam_replacement_identity lam_replacement_constant lam_replacement_product[of "\x. 0" "\x. x"] unfolding Inl_def by simp lemma lam_replacement_Inr: "lam_replacement(M, Inr)" using lam_replacement_identity lam_replacement_constant lam_replacement_product[of "\x. 1" "\x. x"] unfolding Inr_def by simp lemmas Inl_replacement1 = lam_replacement_Inl[unfolded lam_replacement_def] lemma lam_replacement_Diff': "M(X) \ lam_replacement(M, \x. x - X)" using lam_replacement_Diff[THEN [5] lam_replacement_hcomp2] lam_replacement_constant lam_replacement_identity by simp lemmas Pair_diff_replacement = lam_replacement_Diff'[unfolded lam_replacement_def] lemma diff_Pair_replacement: "M(p) \ strong_replacement(M, \x y . y=\x,x-{p}\)" using Pair_diff_replacement by simp lemma swap_replacement:"strong_replacement(M, \x y. y = \x, (\\x,y\. \y, x\)(x)\)" using lam_replacement_swap unfolding lam_replacement_def split_def by simp lemma lam_replacement_Un_const:"M(b) \ lam_replacement(M, \x. x \ b)" using lam_replacement_Un lam_replacement_hcomp2[of _ _ "(\)"] lam_replacement_constant[of b] lam_replacement_identity by simp lemmas tag_union_replacement = lam_replacement_Un_const[unfolded lam_replacement_def] lemma lam_replacement_csquare: "lam_replacement(M,\p. \fst(p) \ snd(p), fst(p), snd(p)\)" using lam_replacement_Un lam_replacement_fst lam_replacement_snd by (fast intro: lam_replacement_product lam_replacement_hcomp2) lemma csquare_lam_replacement:"strong_replacement(M, \x y. y = \x, (\\x,y\. \x \ y, x, y\)(x)\)" using lam_replacement_csquare unfolding split_def lam_replacement_def . lemma lam_replacement_assoc:"lam_replacement(M,\x. \fst(fst(x)), snd(fst(x)), snd(x)\)" using lam_replacement_fst lam_replacement_snd by (force intro: lam_replacement_product lam_replacement_hcomp) lemma assoc_replacement:"strong_replacement(M, \x y. y = \x, (\\\x,y\,z\. \x, y, z\)(x)\)" using lam_replacement_assoc unfolding split_def lam_replacement_def . lemma lam_replacement_prod_fun: "M(f) \ M(g) \ lam_replacement(M,\x. \f ` fst(x), g ` snd(x)\)" using lam_replacement_fst lam_replacement_snd by (force intro: lam_replacement_product lam_replacement_hcomp lam_replacement_apply) lemma prod_fun_replacement:"M(f) \ M(g) \ strong_replacement(M, \x y. y = \x, (\\w,y\. \f ` w, g ` y\)(x)\)" using lam_replacement_prod_fun unfolding split_def lam_replacement_def . -lemma separation_eq: - assumes "\x[M]. M(f(x))" "lam_replacement(M,f)" - "\x[M]. M(g(x))" "lam_replacement(M,g)" - shows "separation(M,\x . f(x) = g(x))" -proof - - let ?Z="\A. {\x,\f(x),\g(x),x\\\. x\A}" - let ?Y="\A. {\\x,f(x)\,\g(x),x\\. x\A}" - note sndsnd = lam_replacement_hcomp[OF lam_replacement_snd lam_replacement_snd] - note fstsnd = lam_replacement_hcomp[OF lam_replacement_snd lam_replacement_fst] - note sndfst = lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_snd] - have "M(?Z(A))" if "M(A)" for A - using assms lam_replacement_iff_lam_closed that - lam_replacement_product[OF assms(2) - lam_replacement_product[OF assms(4) lam_replacement_identity]] - unfolding lam_def - by auto - moreover - have "?Y(A) = {\\fst(x), fst(snd(x))\, fst(snd(snd(x))), snd(snd(snd(x)))\ . x \ ?Z(A)}" for A - by auto - moreover from calculation - have "M(?Y(A))" if "M(A)" for A - using - lam_replacement_imp_strong_replacement[OF - lam_replacement_product[OF - lam_replacement_product[OF lam_replacement_fst fstsnd] - lam_replacement_product[OF - lam_replacement_hcomp[OF sndsnd lam_replacement_fst] - lam_replacement_hcomp[OF lam_replacement_snd sndsnd] - ] - ], THEN RepFun_closed,simplified,of "?Z(A)"] - fst_closed[OF transM] snd_closed[OF transM] that - by auto - then - have "M({u\?Y(A) . snd(fst(u)) = fst(snd(u))})" (is "M(?W(A))") if "M(A)" for A - using that middle_separation assms - by auto - then - have "M({fst(fst(u)) . u \ ?W(A)})" if "M(A)" for A - using that lam_replacement_imp_strong_replacement[OF - lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_fst], THEN RepFun_closed] - fst_closed[OF transM] - by auto - moreover - have "{x\A. f(x) = g(x)} = {fst(fst(u)) . u\?W(A)}" for A - by auto - ultimately - show ?thesis - using separation_iff by auto -qed - lemma separation_subset: assumes "\x[M]. M(f(x))" "lam_replacement(M,f)" "\x[M]. M(g(x))" "lam_replacement(M,g)" shows "separation(M,\x . f(x) \ g(x))" proof - have "f(x) \ g(x) \ f(x)\g(x) = g(x)" for x using subset_Un_iff by simp moreover from assms have "separation(M,\x . f(x)\g(x) = g(x))" using separation_eq lam_replacement_Un lam_replacement_hcomp2 by simp ultimately show ?thesis using separation_cong[THEN iffD1] by auto qed lemma lam_replacement_twist: "lam_replacement(M,\\\x,y\,z\. \x,y,z\)" using lam_replacement_fst lam_replacement_snd lam_replacement_Pair[THEN [5] lam_replacement_hcomp2, of "\x. snd(fst(x))" "\x. snd(x)", THEN [2] lam_replacement_Pair[ THEN [5] lam_replacement_hcomp2, of "\x. fst(fst(x))"]] lam_replacement_hcomp unfolding split_def by simp lemma twist_closed[intro,simp]: "M(x) \ M((\\\x,y\,z\. \x,y,z\)(x))" unfolding split_def by simp lemma lam_replacement_vimage : shows "lam_replacement(M, \x. fst(x)-``snd(x))" unfolding vimage_def using lam_replacement_Image[THEN [5] lam_replacement_hcomp2] lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_converse] lam_replacement_snd by simp lemma lam_replacement_vimage_sing: "lam_replacement(M, \p. fst(p) -`` {snd(p)})" using lam_replacement_snd lam_replacement_sing_fun lam_replacement_vimage[THEN [5] lam_replacement_hcomp2] lam_replacement_fst by simp lemma lam_replacement_vimage_sing_fun: "M(f) \ lam_replacement(M, \x. f -`` {x})" using lam_replacement_vimage_sing[THEN [5] lam_replacement_hcomp2] lam_replacement_constant[of f] lam_replacement_identity by simp lemma lam_replacement_image_sing_fun: "M(f) \ lam_replacement(M, \x. f `` {x})" using lam_replacement_Image[THEN [5] lam_replacement_hcomp2] lam_replacement_constant[of f] lam_replacement_sing by simp lemma converse_apply_projs: "\x[M]. \ (fst(x) -`` {snd(x)}) = converse(fst(x)) ` (snd(x))" using converse_apply_eq by auto lemma lam_replacement_converse_app: "lam_replacement(M, \p. converse(fst(p)) ` snd(p))" using lam_replacement_cong[OF _ converse_apply_projs] lam_replacement_hcomp[OF lam_replacement_vimage_sing lam_replacement_Union] by simp lemmas cardinal_lib_assms4 = lam_replacement_vimage_sing_fun[unfolded lam_replacement_def] lemma lam_replacement_sing_const_id: "M(x) \ lam_replacement(M, \y. {\x, y\})" using lam_replacement_const_id[of x] lam_replacement_sing_fun by simp lemma tag_singleton_closed: "M(x) \ M(z) \ M({{\z, y\} . y \ x})" using RepFun_closed lam_replacement_imp_strong_replacement lam_replacement_sing_const_id transM[of _ x] by simp lemma separation_ball: assumes "separation(M, \y. f(fst(y),snd(y)))" "M(X)" shows "separation(M, \y. \u\X. f(y,u))" unfolding separation_def proof(clarify) fix A assume "M(A)" moreover note \M(X)\ moreover from calculation have "M(A\X)" by simp then have "M({p \ A\X . f(fst(p),snd(p))})" (is "M(?P)") using assms(1) by auto moreover from calculation have "M({a\A . ?P``{a} = X})" (is "M(?A')") using separation_eq lam_replacement_image_sing_fun[of "?P"] lam_replacement_constant by simp moreover have "f(a,x)" if "a\?A'" and "x\X" for a x proof - from that have "a\A" "?P``{a}=X" by auto then have "x\?P``{a}" using that by simp then show ?thesis using image_singleton_iff by simp qed moreover from this have "\a[M]. a \ ?A' \ a \ A \ (\x\X. f(a, x))" using image_singleton_iff by auto with \M(?A')\ show "\y[M]. \a[M]. a \ y \ a \ A \ (\x\X. f(a, x))" by (rule_tac x="?A'" in rexI,simp_all) qed lemma separation_bex: assumes "separation(M, \y. f(fst(y),snd(y)))" "M(X)" shows "separation(M, \y. \u\X. f(y,u))" unfolding separation_def proof(clarify) fix A assume "M(A)" moreover note \M(X)\ moreover from calculation have "M(A\X)" by simp then have "M({p \ A\X . f(fst(p),snd(p))})" (is "M(?P)") using assms(1) by auto moreover from calculation have "M({a\A . ?P``{a} \ 0})" (is "M(?A')") using separation_eq lam_replacement_image_sing_fun[of "?P"] lam_replacement_constant separation_neg by simp moreover from this have "\a[M]. a \ ?A' \ a \ A \ (\x\X. f(a, x))" using image_singleton_iff by auto with \M(?A')\ show "\y[M]. \a[M]. a \ y \ a \ A \ (\x\X. f(a, x))" by (rule_tac x="?A'" in rexI,simp_all) qed lemma lam_replacement_Lambda: assumes "lam_replacement(M, \y. b(fst(y), snd(y)))" "\w[M]. \y[M]. M(b(w, y))" "M(W)" shows "lam_replacement(M, \x. \w\W. b(x, w))" using assms lam_replacement_constant lam_replacement_product lam_replacement_snd - lam_replacement_RepFun + lam_replacement_RepFun transM[of _ W] unfolding lam_def by simp lemma lam_replacement_apply_Pair: assumes "M(y)" shows "lam_replacement(M, \x. y ` \fst(x), snd(x)\)" using assms lam_replacement_constant lam_replacement_Pair lam_replacement_apply2[THEN [5] lam_replacement_hcomp2] by auto lemma lam_replacement_apply_fst_snd: shows "lam_replacement(M, \w. fst(w) ` fst(snd(w)) ` snd(snd(w)))" using lam_replacement_fst lam_replacement_snd lam_replacement_hcomp lam_replacement_apply2[THEN [5] lam_replacement_hcomp2] by auto lemma lam_replacement_RepFun_apply : assumes "M(f)" shows "lam_replacement(M, \x . {f`y . y \ x})" using assms lam_replacement_identity lam_replacement_RepFun lam_replacement_hcomp lam_replacement_snd lam_replacement_apply - by simp + by(rule_tac lam_replacement_RepFun,auto dest:transM) lemma separation_snd_in_fst: "separation(M, \x. snd(x) \ fst(x))" using separation_in lam_replacement_fst lam_replacement_snd by auto lemma lam_replacement_if_mem: "lam_replacement(M, \x. if snd(x) \ fst(x) then 1 else 0)" using separation_snd_in_fst lam_replacement_constant lam_replacement_if by auto lemma lam_replacement_Lambda_apply_fst_snd: assumes "M(X)" shows "lam_replacement(M, \x. \w\X. x ` fst(w) ` snd(w))" using assms lam_replacement_apply_fst_snd lam_replacement_Lambda by simp lemma lam_replacement_Lambda_apply_Pair: assumes "M(X)" "M(y)" shows "lam_replacement(M, \x. \w\X. y ` \x, w\)" using assms lam_replacement_apply_Pair lam_replacement_Lambda by simp lemma lam_replacement_Lambda_if_mem: assumes "M(X)" shows "lam_replacement(M, \x. \xa\X. if xa \ x then 1 else 0)" using assms lam_replacement_if_mem lam_replacement_Lambda by simp -lemma lam_replacement_comp': - "M(f) \ M(g) \ lam_replacement(M, \x . f O x O g)" - using lam_replacement_comp[THEN [5] lam_replacement_hcomp2, - OF lam_replacement_constant lam_replacement_comp, - THEN [5] lam_replacement_hcomp2] lam_replacement_constant - lam_replacement_identity by simp - lemma case_closed : assumes "\x[M]. M(f(x))" "\x[M]. M(g(x))" shows "\x[M]. M(case(f,g,x))" unfolding case_def split_def cond_def using assms by simp lemma separation_fst_equal : "M(a) \ separation(M,\x . fst(x)=a)" using separation_eq lam_replacement_fst lam_replacement_constant by auto lemma lam_replacement_case : assumes "lam_replacement(M,f)" "lam_replacement(M,g)" "\x[M]. M(f(x))" "\x[M]. M(g(x))" shows "lam_replacement(M, \x . case(f,g,x))" unfolding case_def split_def cond_def using lam_replacement_if separation_fst_equal lam_replacement_hcomp[of "snd" g] lam_replacement_hcomp[of "snd" f] lam_replacement_snd assms by simp lemma Pi_replacement1: "M(x) \ M(y) \ strong_replacement(M, \ya z. ya \ y \ z = {\x, ya\})" using lam_replacement_imp_strong_replacement strong_replacement_separation[OF lam_replacement_sing_const_id[of x],where P="\x . x \y"] separation_in_constant by simp lemma surj_imp_inj_replacement1: "M(f) \ M(x) \ strong_replacement(M, \y z. y \ f -`` {x} \ z = {\x, y\})" using Pi_replacement1 vimage_closed singleton_closed by simp lemmas domain_replacement = lam_replacement_domain[unfolded lam_replacement_def] lemma domain_replacement_simp: "strong_replacement(M, \x y. y=domain(x))" using lam_replacement_domain lam_replacement_imp_strong_replacement by simp lemma un_Pair_replacement: "M(p) \ strong_replacement(M, \x y . y = x\{p})" using lam_replacement_Un_const[THEN lam_replacement_imp_strong_replacement] by simp lemma diff_replacement: "M(X) \ strong_replacement(M, \x y. y = x - X)" using lam_replacement_Diff'[THEN lam_replacement_imp_strong_replacement] by simp lemma lam_replacement_succ: "lam_replacement(M,\z . succ(z))" unfolding succ_def using lam_replacement_hcomp2[of "\x. x" "\x. x" cons] lam_replacement_cons lam_replacement_identity by simp lemma lam_replacement_hcomp_Least: assumes "lam_replacement(M, g)" "lam_replacement(M,\x. \ i. x\F(i,x))" "\x[M]. M(g(x))" "\x i. M(x) \ i \ F(i, x) \ M(i)" shows "lam_replacement(M,\x. \ i. g(x)\F(i,g(x)))" using assms by (rule_tac lam_replacement_hcomp[of _ "\x. \ i. x\F(i,x)"]) (auto intro:Least_closed') lemma domain_mem_separation: "M(A) \ separation(M, \x . domain(x)\A)" using separation_in lam_replacement_constant lam_replacement_domain by auto lemma domain_eq_separation: "M(p) \ separation(M, \x . domain(x) = p)" using separation_eq lam_replacement_domain lam_replacement_constant by auto lemma lam_replacement_Int: shows "lam_replacement(M, \x. fst(x) \ snd(x))" proof - have "A\B = (A\B) - ((A- B) \ (B-A))" (is "_=?f(A,B)")for A B by auto then show ?thesis using lam_replacement_cong lam_replacement_Diff[THEN[5] lam_replacement_hcomp2] lam_replacement_Un[THEN[5] lam_replacement_hcomp2] lam_replacement_fst lam_replacement_snd by simp qed lemma restrict_eq_separation': "M(B) \ \A[M]. separation(M, \y. \x[M]. x\A \ y = \x, restrict(x, B)\)" proof(clarify) fix A have "restrict(r,B) = r \ (B \ range(r))" for r unfolding restrict_def by(rule equalityI subsetI,auto) moreover assume "M(A)" "M(B)" moreover from this have "separation(M, \y. \x\A. y = \x, x \ (B \ range(x))\)" using lam_replacement_Int[THEN[5] lam_replacement_hcomp2] lam_replacement_Pair[THEN[5] lam_replacement_hcomp2] using lam_replacement_fst lam_replacement_snd lam_replacement_constant lam_replacement_hcomp lam_replacement_range lam_replacement_identity lam_replacement_CartProd separation_bex separation_eq by simp_all ultimately show "separation(M, \y. \x[M]. x\A \ y = \x, restrict(x, B)\)" by simp qed lemmas lam_replacement_restrict' = lam_replacement_restrict[OF restrict_eq_separation'] lemma restrict_strong_replacement: "M(A) \ strong_replacement(M, \x y. y=restrict(x,A))" using lam_replacement_restrict restrict_eq_separation' lam_replacement_imp_strong_replacement by simp lemma restrict_eq_separation: "M(r) \ M(p) \ separation(M, \x . restrict(x,r) = p)" using separation_eq lam_replacement_restrict' lam_replacement_constant by auto lemma separation_equal_fst2 : "M(a) \ separation(M,\x . fst(fst(x))=a)" using separation_eq lam_replacement_hcomp lam_replacement_fst lam_replacement_constant by auto lemma separation_equal_apply: "M(f) \ M(a) \ separation(M,\x. f`x=a)" using separation_eq lam_replacement_apply[of f] lam_replacement_constant by auto lemma lam_apply_replacement: "M(A) \ M(f) \ lam_replacement(M, \x . \n\A. f ` \x, n\)" using lam_replacement_Lambda lam_replacement_hcomp[OF _ lam_replacement_apply[of f]] lam_replacement_Pair by simp -lemma separation_all: - assumes "separation(M, \x . P(fst(fst(x)),snd(x)))" - "lam_replacement(M,f)" "lam_replacement(M,g)" - "\x[M]. M(f(x))" "\x[M]. M(g(x))" - shows "separation(M, \z. \x\f(z). P(g(z),x))" - unfolding separation_def -proof(clarify) - note rep_closed = lam_replacement_imp_strong_replacement[THEN RepFun_closed] +\ \We need a tactic to deal with composition of replacements!\ +lemma lam_replacement_comp : + "lam_replacement(M, \x . fst(x) O snd(x))" +proof - note lr_fs = lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_snd] - fix A - assume "M(A)" - with assms - have "M({f(x) . x\A})" (is "M(?F)") "M({ . x\A})" (is "M(?G)") - using rep_closed transM[of _ A] - lam_replacement_product lam_replacement_imp_lam_closed - unfolding lam_def - by auto - let ?B="\?F" - let ?C="?G\?B" - note \M(A)\ \M(?F)\ \M(?G)\ - moreover from calculation - have "M(?C)" - by simp - moreover from calculation - have "M({p\?C . P(fst(fst(p)),snd(p)) \ snd(p)\snd(fst(p))})" (is "M(?Prod)") - using assms separation_conj separation_in lam_replacement_fst lam_replacement_snd - lam_replacement_hcomp - by simp - moreover from calculation - have "M({z\?G . snd(z)=?Prod``{z}})" (is "M(?L)") - using separation_eq - lam_replacement_constant[of ?Prod] lam_replacement_constant[of 0] - lam_replacement_hcomp[OF _ lam_replacement_sing] - lam_replacement_Image[THEN [5] lam_replacement_hcomp2] - lam_replacement_identity lam_replacement_snd - by simp - moreover from calculation assms - have "M({x\A . \ ?L})" (is "M(?N)") - using lam_replacement_product lam_replacement_constant - by (rule_tac separation_closed,rule_tac separation_in,auto) - moreover from calculation - have "?N = {z\A . \x\f(z). P(g(z),x)}" - proof - - have "P(g(z),x)" if "z\A" "x\f(z)" "f(z) = ?Prod``{}" for z x - using that - by(rule_tac imageE[of x ?Prod "{}"],simp_all) - moreover - have "f(z) = ?Prod `` {}" if "z\A" "\x\f(z). P(g(z), x)" for z - using that - by force - ultimately - show ?thesis - by auto - qed + note lr_ff = lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_fst] + note lr_ss = lam_replacement_hcomp[OF lam_replacement_snd lam_replacement_snd] + note lr_sf = lam_replacement_hcomp[OF lam_replacement_snd lam_replacement_fst] + note lr_fff = lam_replacement_hcomp[OF lam_replacement_fst lr_ff] + have eq:"R O S = + {xz \ domain(S) \ range(R) . + \y\range(S)\domain(R) . \fst(xz),y\\S \ \y,snd(xz)\\R}" for R S + unfolding comp_def + by(intro equalityI subsetI,auto) + moreover + have "lam_replacement(M, \x . {xz \ domain(snd(x)) \ range(fst(x)) . + \y\range(snd(x))\domain(fst(x)) . \fst(xz),y\\snd(x) \ \y,snd(xz)\\fst(x)})" + using lam_replacement_CartProd lam_replacement_hcomp + lam_replacement_range lam_replacement_domain lam_replacement_fst lam_replacement_snd + lam_replacement_identity lr_fs lr_ff lr_ss lam_replacement_product + lam_replacement_Int[THEN [5] lam_replacement_hcomp2] + lam_replacement_hcomp[OF lr_ff lam_replacement_domain] + lam_replacement_hcomp[OF lr_fs lam_replacement_range] + lam_replacement_hcomp[OF lr_ff lr_ff] + lam_replacement_hcomp[OF lr_ff lr_ss] + lam_replacement_hcomp[OF lr_ff lr_sf] + lr_fff + lam_replacement_hcomp[OF lr_fff lam_replacement_snd ] + separation_ex separation_conj separation_in + by(rule_tac lam_replacement_Collect,simp_all,rule_tac separation_ex,rule_tac separation_conj,auto) ultimately - show " \y[M]. \z[M]. z \ y \ z \ A \ (\x\f(z) . P(g(z),x))" - by (rule_tac x="?N" in rexI,simp_all) + show ?thesis + by(rule_tac lam_replacement_cong,auto,subst eq[symmetric],rule_tac comp_closed,auto) qed -lemma separation_ex: - assumes "separation(M, \x . P(fst(fst(x)),snd(x)))" - "lam_replacement(M,f)" "lam_replacement(M,g)" - "\x[M]. M(f(x))" "\x[M]. M(g(x))" - shows "separation(M, \z. \x\f(z). P(g(z),x))" - unfolding separation_def -proof(clarify) - note rep_closed = lam_replacement_imp_strong_replacement[THEN RepFun_closed] - note lr_fs = lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_snd] - fix A - assume "M(A)" - with assms - have "M({f(x) . x\A})" (is "M(?F)") "M({ . x\A})" (is "M(?G)") - using rep_closed transM[of _ A] - lam_replacement_product lam_replacement_imp_lam_closed - unfolding lam_def - by auto - let ?B="\?F" - let ?C="?G\?B" - note \M(A)\ \M(?F)\ \M(?G)\ - moreover from calculation - have "M(?C)" - by simp - moreover from calculation - have "M({p\?C . P(fst(fst(p)),snd(p)) \ snd(p)\snd(fst(p))})" (is "M(?Prod)") - using assms separation_conj separation_in lam_replacement_fst lam_replacement_snd - lam_replacement_hcomp - by simp - moreover from calculation - have "M({z\?G . 0\?Prod``{z}})" (is "M(?L)") - using separation_eq separation_neg - lam_replacement_constant[of ?Prod] lam_replacement_constant[of 0] - lam_replacement_hcomp[OF _ lam_replacement_sing] - lam_replacement_Image[THEN [5] lam_replacement_hcomp2] - lam_replacement_identity - by simp - moreover from calculation assms - have "M({x\A . \ ?L})" (is "M(?N)") - using lam_replacement_product lam_replacement_constant - by (rule_tac separation_closed,rule_tac separation_in,auto) - moreover from calculation - have "?N = {z\A . \x\f(z). P(g(z),x)}" - by(intro equalityI subsetI,simp_all,force) (rule ccontr,force) - ultimately - show " \y[M]. \z[M]. z \ y \ z \ A \ (\x\f(z) . P(g(z),x))" - by (rule_tac x="?N" in rexI,simp_all) -qed +lemma lam_replacement_comp': + "M(f) \ M(g) \ lam_replacement(M, \x . f O x O g)" + using lam_replacement_comp[THEN [5] lam_replacement_hcomp2, + OF lam_replacement_constant lam_replacement_comp, + THEN [5] lam_replacement_hcomp2] lam_replacement_constant + lam_replacement_identity by simp lemma RepFun_SigFun_closed: "M(x)\ M(z) \ M({{\z, u\} . u \ x})" using lam_replacement_sing_const_id lam_replacement_imp_strong_replacement RepFun_closed transM[of _ x] singleton_in_M_iff pair_in_M_iff by simp -lemma separation_comp : - assumes "separation(M,P)" "lam_replacement(M,f)" "\x[M]. M(f(x))" - shows "separation(M,\x. P(f(x)))" - unfolding separation_def -proof(clarify) - fix A - assume "M(A)" - let ?B="{f(a) . a \ A}" - let ?C="A\{b\?B . P(b)}" - note \M(A)\ - moreover from calculation - have "M(?C)" - using lam_replacement_imp_strong_replacement assms RepFun_closed transM[of _ A] - by simp - moreover from calculation - have "M({p\?C . f(fst(p)) = snd(p)})" (is "M(?Prod)") - using assms separation_eq lam_replacement_fst lam_replacement_snd - lam_replacement_hcomp - by simp - moreover from calculation - have "M({fst(p) . p\?Prod})" (is "M(?L)") - using lam_replacement_imp_strong_replacement lam_replacement_fst RepFun_closed - transM[of _ ?Prod] - by simp - moreover - have "?L = {z\A . P(f(z))}" - by(intro equalityI subsetI,auto) - ultimately - show " \y[M]. \z[M]. z \ y \ z \ A \ P(f(z))" - by (rule_tac x="?L" in rexI,simp_all) -qed - -(* Gives the orthogonal of x with respect Q *) +\ \Gives the orthogonal of \<^term>\x\ with respect to the relation \<^term>\Q\.\ lemma separation_orthogonal: "M(x) \ M(Q) \ separation(M, \a . \s\x. \s, a\ \ Q)" using separation_in[OF _ lam_replacement_hcomp2[OF _ _ _ _ lam_replacement_Pair] _ lam_replacement_constant] separation_ball lam_replacement_hcomp lam_replacement_fst lam_replacement_snd by simp_all lemma separation_Transset: "separation(M,Transset)" unfolding Transset_def using separation_subset lam_replacement_fst lam_replacement_snd lam_replacement_hcomp lam_replacement_identity by(rule_tac separation_all,auto) lemma separation_Ord: "separation(M,Ord)" unfolding Ord_def using separation_Transset separation_comp separation_Transset lam_replacement_snd lam_replacement_identity by(rule_tac separation_conj,auto,rule_tac separation_all,auto) lemma ord_iso_separation: "M(A) \ M(r) \ M(s) \ separation(M, \f. \x\A. \y\A. \x, y\ \ r \ \f ` x, f ` y\ \ s)" using lam_replacement_Pair[THEN[5] lam_replacement_hcomp2] lam_replacement_apply2[THEN[5] lam_replacement_hcomp2] lam_replacement_hcomp lam_replacement_fst lam_replacement_snd lam_replacement_identity lam_replacement_constant separation_in separation_ball[of _ A] separation_iff' by simp lemma separation_dist: "separation(M, \ x . \a. \b . x=\a,b\ \ a\b)" using separation_Pair separation_neg separation_eq lam_replacement_fst lam_replacement_snd by simp lemma separation_imp_lam_closed: assumes "\x\A. F(x) \ B" "separation(M, \\x,y\. F(x) = y)" "M(A)" "M(B)" shows "M(\x\A. F(x))" proof - from \\x\A. F(x) \ B\ have "(\x\A. F(x)) \ A\B" using times_subset_iff fun_is_rel[OF lam_funtype, of A F] by auto moreover from this have "(\x\A. F(x)) = {\x,y\ \ A \ B. F(x) = y}" unfolding lam_def by force moreover note \M(A)\ \M(B)\ \separation(M, \\x,y\. F(x) = y)\ moreover from this have "M({\x,y\ \ A \ B. F(x) = y})" by simp ultimately show ?thesis unfolding lam_def by auto qed lemma curry_closed : assumes "M(f)" "M(A)" "M(B)" shows "M(curry(A,B,f))" using assms lam_replacement_apply_const_id lam_apply_replacement lam_replacement_iff_lam_closed unfolding curry_def by auto end \ \\<^locale>\M_replacement\\ definition RepFun_cons :: "i\i\i" where "RepFun_cons(x,y) \ RepFun(x, \z. {\y,z\})" context M_replacement begin lemma lam_replacement_RepFun_cons'': shows "lam_replacement(M, \x . RepFun_cons(fst(x),snd(x)))" unfolding RepFun_cons_def - using lam_replacement_fst fst_closed snd_closed -lam_replacement_Pair[THEN [5] lam_replacement_hcomp2] + using lam_replacement_fst fst_closed snd_closed lam_replacement_product lam_replacement_snd lam_replacement_hcomp lam_replacement_sing_fun - lam_replacement_RepFun[of "\ x z . {}" "fst"] - by(simp_all) + transM[OF _ fst_closed] +lam_replacement_RepFun[of "\ x z . {}" "fst"] + by simp lemma RepFun_cons_replacement: "lam_replacement(M, \p. RepFun(fst(p), \x. {\snd(p),x\}))" using lam_replacement_RepFun_cons'' unfolding RepFun_cons_def by simp lemma lam_replacement_Sigfun: assumes "lam_replacement(M,f)" "\y[M]. M(f(y))" shows "lam_replacement(M, \x. Sigfun(x,f))" using assms lam_replacement_Union lam_replacement_sing_fun lam_replacement_Pair - lam_replacement_hcomp[of _ Union] tag_singleton_closed lam_replacement_RepFun + tag_singleton_closed transM[of _ "f(_)"] lam_replacement_hcomp[of _ Union] + lam_replacement_RepFun unfolding Sigfun_def by simp lemma phrank_repl: assumes "M(f)" shows "strong_replacement(M, \x y. y = succ(f`x))" using assms lam_replacement_succ lam_replacement_apply lam_replacement_imp_strong_replacement lam_replacement_hcomp by auto lemma lam_replacement_Powapply_rel: assumes "\A[M]. separation(M, \y. \x[M]. x \ A \ y = \x, Pow_rel(M,x)\)" "M(f)" shows "lam_replacement(M, Powapply_rel(M,f))" using assms lam_replacement_Pow_rel lam_replacement_apply[THEN lam_replacement_hcomp] unfolding Powapply_rel_def by auto lemmas Powapply_rel_replacement = lam_replacement_Powapply_rel[THEN lam_replacement_imp_strong_replacement] lemma surj_imp_inj_replacement2: "M(f) \ strong_replacement(M, \x z. z = Sigfun(x, \y. f -`` {y}))" using lam_replacement_imp_strong_replacement lam_replacement_Sigfun lam_replacement_vimage_sing_fun by simp lemma lam_replacement_Pi: "M(y) \ lam_replacement(M, \x. \xa\y. {\x, xa\})" using lam_replacement_constant lam_replacement_Sigfun[unfolded Sigfun_def,of "\x. y"] by (simp) lemma Pi_replacement2: "M(y) \ strong_replacement(M, \x z. z = (\xa\y. {\x, xa\}))" using lam_replacement_Pi[THEN lam_replacement_imp_strong_replacement, of y] by simp lemma if_then_Inj_replacement: shows "M(A) \ strong_replacement(M, \x y. y = \x, if x \ A then Inl(x) else Inr(x)\)" using lam_replacement_if lam_replacement_Inl lam_replacement_Inr separation_in_constant unfolding lam_replacement_def by simp lemma lam_if_then_replacement: "M(b) \ M(a) \ M(f) \ strong_replacement(M, \y ya. ya = \y, if y = a then b else f ` y\)" using lam_replacement_if lam_replacement_apply lam_replacement_constant separation_equal unfolding lam_replacement_def by simp lemma if_then_replacement: "M(A) \ M(f) \ M(g) \ strong_replacement(M, \x y. y = \x, if x \ A then f ` x else g ` x\)" using lam_replacement_if lam_replacement_apply[of f] lam_replacement_apply[of g] separation_in_constant unfolding lam_replacement_def by simp lemma ifx_replacement: "M(f) \ M(b) \ strong_replacement(M, \x y. y = \x, if x \ range(f) then converse(f) ` x else b\)" using lam_replacement_if lam_replacement_apply lam_replacement_constant separation_in_constant unfolding lam_replacement_def by simp lemma if_then_range_replacement2: "M(A) \ M(C) \ strong_replacement(M, \x y. y = \x, if x = Inl(A) then C else x\)" using lam_replacement_if lam_replacement_constant lam_replacement_identity separation_equal unfolding lam_replacement_def by simp lemma if_then_range_replacement: "M(u) \ M(f) \ strong_replacement (M, \z y. y = \z, if z = u then f ` 0 else if z \ range(f) then f ` succ(converse(f) ` z) else z\)" using lam_replacement_if separation_equal separation_in_constant lam_replacement_constant lam_replacement_identity lam_replacement_succ lam_replacement_apply lam_replacement_hcomp[of "\x. converse(f)`x" "succ"] lam_replacement_hcomp[of "\x. succ(converse(f)`x)" "\x . f`x"] unfolding lam_replacement_def by simp lemma Inl_replacement2: "M(A) \ strong_replacement(M, \x y. y = \x, if fst(x) = A then Inl(snd(x)) else Inr(x)\)" using lam_replacement_if separation_fst_equal lam_replacement_hcomp[of "snd" "Inl"] lam_replacement_Inl lam_replacement_Inr lam_replacement_snd unfolding lam_replacement_def by simp lemma case_replacement1: "strong_replacement(M, \z y. y = \z, case(Inr, Inl, z)\)" using lam_replacement_case lam_replacement_Inl lam_replacement_Inr unfolding lam_replacement_def by simp lemma case_replacement2: "strong_replacement(M, \z y. y = \z, case(case(Inl, \y. Inr(Inl(y))), \y. Inr(Inr(y)), z)\)" using lam_replacement_case lam_replacement_hcomp case_closed[of Inl "\x. Inr(Inl(x))"] lam_replacement_Inl lam_replacement_Inr unfolding lam_replacement_def by simp lemma case_replacement4: "M(f) \ M(g) \ strong_replacement(M, \z y. y = \z, case(\w. Inl(f ` w), \y. Inr(g ` y), z)\)" using lam_replacement_case lam_replacement_hcomp lam_replacement_Inl lam_replacement_Inr lam_replacement_apply unfolding lam_replacement_def by simp lemma case_replacement5: "strong_replacement(M, \x y. y = \x, (\\x,z\. case(\y. Inl(\y, z\), \y. Inr(\y, z\), x))(x)\)" unfolding split_def case_def cond_def using lam_replacement_if separation_equal_fst2 lam_replacement_snd lam_replacement_Inl lam_replacement_Inr lam_replacement_hcomp[OF lam_replacement_product[OF lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_snd]]] unfolding lam_replacement_def by simp end \ \\<^locale>\M_replacement\\ locale M_Pi_replacement = M_Pi + M_replacement begin lemma curry_rel_exp : assumes "M(f)" "M(A)" "M(B)" "M(C)" "f \ A \ B \ C" shows "curry(A,B,f) \ A \\<^bsup>M\<^esup> (B \\<^bsup>M\<^esup> C)" using assms transM[of _ A] lam_closed[OF apply_replacement2] lam_replacement_apply_const_id lam_apply_replacement lam_replacement_iff_lam_closed unfolding curry_def by (intro lam_type mem_function_space_rel_abs[THEN iffD2],auto) end \ \\<^locale>\M_Pi_replacement\\ \ \To be used in the relativized treatment of Cohen posets\ definition \ \"domain collect F"\ dC_F :: "i \ i \ i" where "dC_F(A,d) \ {p \ A. domain(p) = d }" definition \ \"domain restrict SepReplace Y"\ drSR_Y :: "i \ i \ i \ i \ i" where "drSR_Y(B,D,A,x) \ {y . r\A , restrict(r,B) = x \ y = domain(r) \ domain(r) \ D}" lemma drSR_Y_equality: "drSR_Y(B,D,A,x) = { dr\D . (\r\A . restrict(r,B) = x \ dr=domain(r)) }" unfolding drSR_Y_def by auto context M_replacement begin lemma separation_restrict_eq_dom_eq:"\x[M].separation(M, \dr. \r\A . restrict(r,B) = x \ dr=domain(r))" if "M(A)" and "M(B)" for A B using that separation_eq[OF _ lam_replacement_fst _ lam_replacement_hcomp[OF lam_replacement_snd lam_replacement_domain ]] separation_eq[OF _ lam_replacement_hcomp[OF lam_replacement_snd lam_replacement_restrict'] _ lam_replacement_constant] by(clarify,rule_tac separation_bex[OF _ \M(A)\],rule_tac separation_conj,simp_all) -lemma separation_is_insnd_restrict_eq_dom : "separation(M, \p. \x\D. x \ snd(p) \ (\r\A. restrict(r, B) = fst(p) \ x = domain(r)))" +lemma separation_is_insnd_restrict_eq_dom : "separation(M, \p. (\r\A. restrict(r, B) = fst(p) \ snd(p) = domain(r)))" if "M(B)" "M(D)" "M(A)" for A B D - using that lam_replacement_fst lam_replacement_hcomp lam_replacement_snd separation_in - separation_eq[OF _ - lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_snd] _ - lam_replacement_hcomp[OF lam_replacement_snd lam_replacement_domain]] - separation_eq separation_restrict_eq_dom_eq - lam_replacement_hcomp[OF lam_replacement_snd lam_replacement_restrict'] - lam_replacement_hcomp[OF lam_replacement_fst - lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_fst]] - by(rule_tac separation_ball,rule_tac separation_iff',simp_all, - rule_tac separation_bex[OF _ \M(A)\],rule_tac separation_conj,simp_all) + using that lam_replacement_fst lam_replacement_snd + separation_eq lam_replacement_hcomp lam_replacement_domain + lam_replacement_hcomp[OF _ lam_replacement_restrict'] + by( rule_tac separation_bex[OF _ \M(A)\],rule_tac separation_conj,simp_all) lemma lam_replacement_drSR_Y: - assumes - "M(B)" "M(D)" "M(A)" + assumes "M(B)" "M(D)" "M(A)" shows "lam_replacement(M, drSR_Y(B,D,A))" - using lam_replacement_cong lam_replacement_Collect[OF \M(D)\ separation_restrict_eq_dom_eq[of A B]] - assms drSR_Y_equality separation_is_insnd_restrict_eq_dom separation_restrict_eq_dom_eq + using lam_replacement_cong[OF lam_replacement_Collect'[OF _ separation_is_insnd_restrict_eq_dom]] + assms drSR_Y_equality separation_restrict_eq_dom_eq by simp lemma drSR_Y_closed: - assumes - "M(B)" "M(D)" "M(A)" "M(f)" + assumes "M(B)" "M(D)" "M(A)" "M(f)" shows "M(drSR_Y(B,D,A,f))" - using assms drSR_Y_equality lam_replacement_Collect[OF \M(D)\ separation_restrict_eq_dom_eq[of A B]] + using assms drSR_Y_equality lam_replacement_Collect'[OF \M(D)\] assms drSR_Y_equality separation_is_insnd_restrict_eq_dom separation_restrict_eq_dom_eq by simp lemma lam_if_then_apply_replacement: "M(f) \ M(v) \ M(u) \ lam_replacement(M, \x. if f ` x = v then f ` u else f ` x)" using lam_replacement_if separation_equal_apply lam_replacement_constant lam_replacement_apply by simp lemma lam_if_then_apply_replacement2: "M(f) \ M(m) \ M(y) \ lam_replacement(M, \z . if f ` z = m then y else f ` z)" using lam_replacement_if separation_equal_apply lam_replacement_constant lam_replacement_apply by simp lemma lam_if_then_replacement2: "M(A) \ M(f) \ lam_replacement(M, \x . if x \ A then f ` x else x)" using lam_replacement_if separation_in_constant lam_replacement_identity lam_replacement_apply by simp lemma lam_if_then_replacement_apply: "M(G) \ lam_replacement(M, \x. if M(x) then G ` x else 0)" using lam_replacement_if separation_in_constant lam_replacement_identity lam_replacement_apply lam_replacement_constant[of 0] separation_univ by simp lemma lam_replacement_dC_F: assumes "M(A)" shows "lam_replacement(M, dC_F(A))" proof - - have "separation(M, \p. \x\A. x \ snd(p) \ domain(x) = fst(p))" if "M(A)" for A - using separation_ball separation_iff' + have "separation(M, \p. domain(snd(p)) = fst(p))" if "M(A)" for A + using separation_eq that lam_replacement_hcomp lam_replacement_fst lam_replacement_snd lam_replacement_domain - separation_in separation_eq that by simp_all then show ?thesis unfolding dC_F_def - using assms lam_replacement_Collect[of A "\ d x . domain(x) = d"] + using assms lam_replacement_Collect'[of A "\ d x . domain(x) = d"] separation_eq[OF _ lam_replacement_domain _ lam_replacement_constant] by simp -qed + qed lemma dCF_closed: assumes "M(A)" "M(f)" shows "M(dC_F(A,f))" unfolding dC_F_def - using assms lam_replacement_Collect[of A "\ d x . domain(x) = d"] + using assms lam_replacement_Collect'[of A "\ d x . domain(x) = d"] separation_eq[OF _ lam_replacement_domain _ lam_replacement_constant] by simp lemma lam_replacement_Collect_ball_Pair: - assumes "separation(M, \p. \x\G. x \ snd(p) \ (\s\fst(p). \s, x\ \ Q))" "M(G)" "M(Q)" + assumes "M(G)" "M(Q)" shows "lam_replacement(M, \x . {a \ G . \s\x. \s, a\ \ Q})" proof - - have 1:"\x[M]. separation(M, \a . \s\x. \s, a\ \ Q)" if "M(Q)" for Q - using separation_in lam_replacement_hcomp2[OF _ _ _ _ lam_replacement_Pair] - lam_replacement_constant separation_ball - lam_replacement_hcomp lam_replacement_fst lam_replacement_snd that - by simp - then - show ?thesis - using assms lam_replacement_Collect + have "separation(M, \p. (\s\fst(p). \s, snd(p)\ \ Q))" + using separation_in assms lam_replacement_identity + lam_replacement_constant lam_replacement_hcomp[OF lam_replacement_fst + lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_snd]] + lam_replacement_snd lam_replacement_fst + lam_replacement_hcomp[OF lam_replacement_fst lam_replacement_fst] + by(rule_tac separation_all,simp_all,rule_tac separation_in,simp_all,rule_tac lam_replacement_product[of "snd" "\x . snd(fst(fst(x)))"],simp_all) + then + show ?thesis + using assms lam_replacement_Collect' by simp_all qed lemma surj_imp_inj_replacement3: - "(\x. M(x) \ separation(M, \y. \s\x. \s, y\ \ Q)) \ M(G) \ M(Q) \ M(x) \ - strong_replacement(M, \y z. y \ {a \ G . \s\x. \s, a\ \ Q} \ z = {\x, y\})" - using lam_replacement_imp_strong_replacement - using lam_replacement_sing_const_id[THEN lam_replacement_imp_strong_replacement, of x] - unfolding strong_replacement_def - by (simp, safe, drule_tac x="A \ {a \ G . \s\x. \s, a\ \ Q}" in rspec, - simp, erule_tac rexE, rule_tac x=Y in rexI) auto + assumes "M(G)" "M(Q)" "M(x)" + shows "strong_replacement(M, \y z. y \ {a \ G . \s\x. \s, a\ \ Q} \ z = {\x, y\})" +proof - + from assms + have "separation(M, \z. \s\x. \s, z\ \ Q)" + using lam_replacement_swap lam_replacement_constant separation_in separation_ball + by simp + with assms + show ?thesis + using separation_in_constant lam_replacement_sing_const_id[of x] separation_conj + by(rule_tac strong_replacement_separation,simp_all) +qed lemmas replacements = Pair_diff_replacement id_replacement tag_replacement pospend_replacement prepend_replacement Inl_replacement1 diff_Pair_replacement swap_replacement tag_union_replacement csquare_lam_replacement assoc_replacement prod_fun_replacement cardinal_lib_assms4 domain_replacement apply_replacement un_Pair_replacement restrict_strong_replacement diff_replacement if_then_Inj_replacement lam_if_then_replacement if_then_replacement ifx_replacement if_then_range_replacement2 if_then_range_replacement Inl_replacement2 case_replacement1 case_replacement2 case_replacement4 case_replacement5 +lemma zermelo_separation: "M(Q) \ M(f) \ separation(M, \X. Q \ f `` X \ X)" + using lam_replacement_identity lam_replacement_Un[THEN [5] lam_replacement_hcomp2] + lam_replacement_constant lam_replacement_Image[THEN [5] lam_replacement_hcomp2] + separation_subset + by simp + end \ \\<^locale>\M_replacement\\ subsection\Some basic replacement lemmas\ lemma (in M_trans) strong_replacement_conj: assumes "\A. M(A) \ univalent(M,A,P)" "strong_replacement(M,P)" "separation(M, \x. \b[M]. Q(x,b) \ P(x,b))" shows "strong_replacement(M, \x z. Q(x,z) \ P(x,z))" proof - { fix A assume "M(A)" with \separation(M, \x. \b[M]. Q(x,b) \ P(x,b))\ have "M({x\A. \b[M]. Q(x,b) \ P(x,b)})" by simp with \M(_) \ univalent(M,{x\A. \b[M]. Q(x,b) \ P(x,b)}, P)\ \strong_replacement(M, P)\ have "\Y[M]. \b[M]. b \ Y \ (\x[M]. x \ {x\A. \b[M]. Q(x,b) \ P(x,b)} \ P(x, b))" unfolding strong_replacement_def by blast then obtain Y where "\b[M]. b \ Y \ (\x[M]. x \ {x\A. \b[M]. Q(x,b) \ P(x,b)} \ P(x, b))" "M(Y)" by blast with \M(A)\ \M(A) \ univalent(M,A,P)\ have "\b[M]. b \ Y \ (\x[M]. x \ A \ Q(x, b) \ P(x, b))" unfolding univalent_def by auto with \M(Y)\ have "\Y[M]. \b[M]. b \ Y \ (\x[M]. x \ A \ Q(x, b) \ P(x, b))" by auto } then show ?thesis unfolding strong_replacement_def by simp qed lemma strong_replacement_iff_bounded_M: "strong_replacement(M,P) \ strong_replacement(M,\ x z . M(z) \ M(x) \ P(x,z))" unfolding strong_replacement_def by auto -end \ No newline at end of file +end diff --git a/thys/Transitive_Models/Least.thy b/thys/Transitive_Models/Least.thy --- a/thys/Transitive_Models/Least.thy +++ b/thys/Transitive_Models/Least.thy @@ -1,164 +1,163 @@ section\The binder \<^term>\Least\\ theory Least imports "Internalizations" begin text\We have some basic results on the least ordinal satisfying a predicate.\ lemma Least_Ord: "(\ \. R(\)) = (\ \. Ord(\) \ R(\))" unfolding Least_def by (simp add:lt_Ord) lemma Ord_Least_cong: assumes "\y. Ord(y) \ R(y) \ Q(y)" shows "(\ \. R(\)) = (\ \. Q(\))" proof - from assms have "(\ \. Ord(\) \ R(\)) = (\ \. Ord(\) \ Q(\))" by simp then show ?thesis using Least_Ord by simp qed definition least :: "[i\o,i\o,i] \ o" where "least(M,Q,i) \ ordinal(M,i) \ ( (empty(M,i) \ (\b[M]. ordinal(M,b) \ \Q(b))) \ (Q(i) \ (\b[M]. ordinal(M,b) \ b\i\ \Q(b))))" definition least_fm :: "[i,i] \ i" where "least_fm(q,i) \ And(ordinal_fm(i), Or(And(empty_fm(i),Forall(Implies(ordinal_fm(0),Neg(q)))), And(Exists(And(q,Equal(0,succ(i)))), Forall(Implies(And(ordinal_fm(0),Member(0,succ(i))),Neg(q))))))" lemma least_fm_type[TC] :"i \ nat \ q\formula \ least_fm(q,i) \ formula" unfolding least_fm_def by simp -(* Refactorize Formula and Relative to include the following three lemmas *) lemmas basic_fm_simps = sats_subset_fm' sats_transset_fm' sats_ordinal_fm' lemma sats_least_fm : assumes p_iff_sats: "\a. a \ A \ P(a) \ sats(A, p, Cons(a, env))" shows "\y \ nat; env \ list(A) ; 0\A\ \ sats(A, least_fm(p,y), env) \ least(##A, P, nth(y,env))" using nth_closed p_iff_sats unfolding least_def least_fm_def by (simp add:basic_fm_simps) lemma least_iff_sats [iff_sats]: assumes is_Q_iff_sats: "\a. a \ A \ is_Q(a) \ sats(A, q, Cons(a,env))" shows "\nth(j,env) = y; j \ nat; env \ list(A); 0\A\ \ least(##A, is_Q, y) \ sats(A, least_fm(q,j), env)" using sats_least_fm [OF is_Q_iff_sats, of j , symmetric] by simp lemma least_conj: "a\M \ least(##M, \x. x\M \ Q(x),a) \ least(##M,Q,a)" unfolding least_def by simp context M_trivial begin subsection\Uniqueness, absoluteness and closure under \<^term>\Least\\ lemma unique_least: assumes "M(a)" "M(b)" "least(M,Q,a)" "least(M,Q,b)" shows "a=b" proof - from assms have "Ord(a)" "Ord(b)" unfolding least_def by simp_all then consider (le) "a\b" | "a=b" | (ge) "b\a" using Ord_linear[OF \Ord(a)\ \Ord(b)\] by auto then show ?thesis proof(cases) case le then show ?thesis using assms unfolding least_def by auto next case ge then show ?thesis using assms unfolding least_def by auto qed qed lemma least_abs: assumes "\x. Q(x) \ Ord(x) \ \y[M]. Q(y) \ Ord(y)" "M(a)" shows "least(M,Q,a) \ a = (\ x. Q(x))" unfolding least_def proof (cases "\b[M]. Ord(b) \ \ Q(b)"; intro iffI; simp add:assms) case True with assms have "\ (\i. Ord(i) \ Q(i)) " by blast then show "0 =(\ x. Q(x))" using Least_0 by simp then show "ordinal(M, \ x. Q(x)) \ (empty(M, Least(Q)) \ Q(Least(Q)))" by simp next assume "\b[M]. Ord(b) \ Q(b)" then obtain i where "M(i)" "Ord(i)" "Q(i)" by blast assume "a = (\ x. Q(x))" moreover note \M(a)\ moreover from \Q(i)\ \Ord(i)\ have "Q(\ x. Q(x))" (is ?G) by (blast intro:LeastI) moreover have "(\b[M]. Ord(b) \ b \ (\ x. Q(x)) \ \ Q(b))" (is "?H") using less_LeastE[of Q _ False] by (auto, drule_tac ltI, simp, blast) ultimately show "ordinal(M, \ x. Q(x)) \ (empty(M, \ x. Q(x)) \ (\b[M]. Ord(b) \ \ Q(b)) \ ?G \ ?H)" by simp next assume 1:"\b[M]. Ord(b) \ Q(b)" then obtain i where "M(i)" "Ord(i)" "Q(i)" by blast assume "Ord(a) \ (a = 0 \ (\b[M]. Ord(b) \ \ Q(b)) \ Q(a) \ (\b[M]. Ord(b) \ b \ a \ \ Q(b)))" with 1 have "Ord(a)" "Q(a)" "\b[M]. Ord(b) \ b \ a \ \ Q(b)" by blast+ moreover from this and assms have "Ord(b) \ b \ a \ \ Q(b)" for b by (auto dest:transM) moreover from this and \Ord(a)\ have "b < a \ \ Q(b)" for b unfolding lt_def using Ord_in_Ord by blast ultimately show "a = (\ x. Q(x))" using Least_equality by simp qed lemma Least_closed: assumes "\x. Q(x) \ Ord(x) \ \y[M]. Q(y) \ Ord(y)" shows "M(\ x. Q(x))" using assms Least_le[of Q] Least_0[of Q] by (cases "(\i[M]. Ord(i) \ Q(i))") (force dest:transM ltD)+ text\Older, easier to apply versions (with a simpler assumption on \<^term>\Q\).\ lemma least_abs': assumes "\x. Q(x) \ M(x)" "M(a)" shows "least(M,Q,a) \ a = (\ x. Q(x))" using assms least_abs[of Q] by auto lemma Least_closed': assumes "\x. Q(x) \ M(x)" shows "M(\ x. Q(x))" using assms Least_closed[of Q] by auto end \ \\<^locale>\M_trivial\\ end \ No newline at end of file diff --git a/thys/Transitive_Models/M_Basic_No_Repl.thy b/thys/Transitive_Models/M_Basic_No_Repl.thy --- a/thys/Transitive_Models/M_Basic_No_Repl.thy +++ b/thys/Transitive_Models/M_Basic_No_Repl.thy @@ -1,320 +1,304 @@ theory M_Basic_No_Repl imports "ZF-Constructible.Relative" begin txt\This locale is exactly \<^locale>\M_basic\ without its only replacement instance.\ locale M_basic_no_repl = M_trivial + assumes Inter_separation: "M(A) ==> separation(M, \x. \y[M]. y\A \ x\y)" and Diff_separation: "M(B) ==> separation(M, \x. x \ B)" and cartprod_separation: "[| M(A); M(B) |] ==> separation(M, \z. \x[M]. x\A & (\y[M]. y\B & pair(M,x,y,z)))" and image_separation: "[| M(A); M(r) |] ==> separation(M, \y. \p[M]. p\r & (\x[M]. x\A & pair(M,x,y,p)))" and converse_separation: "M(r) ==> separation(M, \z. \p[M]. p\r & (\x[M]. \y[M]. pair(M,x,y,p) & pair(M,y,x,z)))" and restrict_separation: "M(A) ==> separation(M, \z. \x[M]. x\A & (\y[M]. pair(M,x,y,z)))" and comp_separation: "[| M(r); M(s) |] ==> 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)" and pred_separation: "[| M(r); M(x) |] ==> separation(M, \y. \p[M]. p\r & pair(M,y,x,p))" and Memrel_separation: "separation(M, \z. \x[M]. \y[M]. pair(M,x,y,z) & x \ y)" and is_recfun_separation: \ \for well-founded recursion: used to prove \is_recfun_equal\\ "[| M(r); M(f); M(g); M(a); M(b) |] ==> 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))" and power_ax: "power_ax(M)" lemma (in M_basic_no_repl) cartprod_iff: "[| M(A); M(B); M(C) |] ==> cartprod(M,A,B,C) \ (\p1[M]. \p2[M]. powerset(M,A \ B,p1) & powerset(M,p1,p2) & C = {z \ p2. \x\A. \y\B. z = })" apply (simp add: Pair_def cartprod_def, safe) defer 1 apply (simp add: powerset_def) apply blast txt\Final, difficult case: the left-to-right direction of the theorem.\ apply (insert power_ax, simp add: power_ax_def) apply (frule_tac x="A \ B" and P="\x. rex(M,Q(x))" for Q in rspec) apply (blast, clarify) apply (drule_tac x=z and P="\x. rex(M,Q(x))" for Q in rspec) apply assumption apply (blast intro: cartprod_iff_lemma) done lemma (in M_basic_no_repl) cartprod_closed_lemma: "[| M(A); M(B) |] ==> \C[M]. cartprod(M,A,B,C)" apply (simp del: cartprod_abs add: cartprod_iff) apply (insert power_ax, simp add: power_ax_def) apply (frule_tac x="A \ B" and P="\x. rex(M,Q(x))" for Q in rspec) apply (blast, clarify) apply (drule_tac x=z and P="\x. rex(M,Q(x))" for Q in rspec, auto) apply (intro rexI conjI, simp+) apply (insert cartprod_separation [of A B], simp) done text\All the lemmas above are necessary because Powerset is not absolute. I should have used Replacement instead!\ lemma (in M_basic_no_repl) cartprod_closed [intro,simp]: "[| M(A); M(B) |] ==> M(A*B)" by (frule cartprod_closed_lemma, assumption, force) lemma (in M_basic_no_repl) sum_closed [intro,simp]: "[| M(A); M(B) |] ==> M(A+B)" by (simp add: sum_def) lemma (in M_basic_no_repl) sum_abs [simp]: "[| M(A); M(B); M(Z) |] ==> is_sum(M,A,B,Z) \ (Z = A+B)" by (simp add: is_sum_def sum_def singleton_0 nat_into_M) lemma (in M_basic_no_repl) M_converse_iff: "M(r) ==> converse(r) = {z \ \(\(r)) * \(\(r)). \p\r. \x[M]. \y[M]. p = \x,y\ & z = \y,x\}" apply (rule equalityI) prefer 2 apply (blast dest: transM, clarify, simp) apply (simp add: Pair_def) apply (blast dest: transM) done lemma (in M_basic_no_repl) converse_closed [intro,simp]: "M(r) ==> M(converse(r))" apply (simp add: M_converse_iff) apply (insert converse_separation [of r], simp) done lemma (in M_basic_no_repl) converse_abs [simp]: "[| M(r); M(z) |] ==> is_converse(M,r,z) \ z = converse(r)" apply (simp add: is_converse_def) apply (rule iffI) prefer 2 apply blast apply (rule M_equalityI) apply simp apply (blast dest: transM)+ done subsubsection \image, preimage, domain, range\ lemma (in M_basic_no_repl) image_closed [intro,simp]: "[| M(A); M(r) |] ==> M(r``A)" apply (simp add: image_iff_Collect) apply (insert image_separation [of A r], simp) done lemma (in M_basic_no_repl) vimage_abs [simp]: "[| M(r); M(A); M(z) |] ==> pre_image(M,r,A,z) \ z = r-``A" apply (simp add: pre_image_def) apply (rule iffI) apply (blast intro!: equalityI dest: transM, blast) done lemma (in M_basic_no_repl) vimage_closed [intro,simp]: "[| M(A); M(r) |] ==> M(r-``A)" by (simp add: vimage_def) subsubsection\Domain, range and field\ lemma (in M_basic_no_repl) domain_closed [intro,simp]: "M(r) ==> M(domain(r))" apply (simp add: domain_eq_vimage) done lemma (in M_basic_no_repl) range_closed [intro,simp]: "M(r) ==> M(range(r))" apply (simp add: range_eq_image) done lemma (in M_basic_no_repl) field_abs [simp]: "[| M(r); M(z) |] ==> is_field(M,r,z) \ z = field(r)" by (simp add: is_field_def field_def) lemma (in M_basic_no_repl) field_closed [intro,simp]: "M(r) ==> M(field(r))" by (simp add: field_def) subsubsection\Relations, functions and application\ lemma (in M_basic_no_repl) apply_closed [intro,simp]: "[|M(f); M(a)|] ==> M(f`a)" by (simp add: apply_def) lemma (in M_basic_no_repl) apply_abs [simp]: "[| M(f); M(x); M(y) |] ==> fun_apply(M,f,x,y) \ f`x = y" apply (simp add: fun_apply_def apply_def, blast) done lemma (in M_basic_no_repl) injection_abs [simp]: "[| M(A); M(f) |] ==> injection(M,A,B,f) \ f \ inj(A,B)" apply (simp add: injection_def apply_iff inj_def) apply (blast dest: transM [of _ A]) done lemma (in M_basic_no_repl) surjection_abs [simp]: "[| M(A); M(B); M(f) |] ==> surjection(M,A,B,f) \ f \ surj(A,B)" by (simp add: surjection_def surj_def) lemma (in M_basic_no_repl) bijection_abs [simp]: "[| M(A); M(B); M(f) |] ==> bijection(M,A,B,f) \ f \ bij(A,B)" by (simp add: bijection_def bij_def) subsubsection\Composition of relations\ lemma (in M_basic_no_repl) M_comp_iff: "[| M(r); M(s) |] ==> r O s = {xz \ domain(s) * range(r). \x[M]. \y[M]. \z[M]. xz = \x,z\ & \x,y\ \ s & \y,z\ \ r}" apply (simp add: comp_def) apply (rule equalityI) apply clarify apply simp apply (blast dest: transM)+ done lemma (in M_basic_no_repl) comp_closed [intro,simp]: "[| M(r); M(s) |] ==> M(r O s)" apply (simp add: M_comp_iff) apply (insert comp_separation [of r s], simp) done lemma (in M_basic_no_repl) composition_abs [simp]: "[| M(r); M(s); M(t) |] ==> composition(M,r,s,t) \ t = r O s" apply safe txt\Proving \<^term>\composition(M, r, s, r O s)\\ prefer 2 apply (simp add: composition_def comp_def) apply (blast dest: transM) txt\Opposite implication\ apply (rule M_equalityI) apply (simp add: composition_def comp_def) apply (blast del: allE dest: transM)+ done text\no longer needed\ lemma (in M_basic_no_repl) restriction_is_function: "[| restriction(M,f,A,z); function(f); M(f); M(A); M(z) |] ==> function(z)" apply (simp add: restriction_def ball_iff_equiv) apply (unfold function_def, blast) done lemma (in M_basic_no_repl) restrict_closed [intro,simp]: "[| M(A); M(r) |] ==> M(restrict(r,A))" apply (simp add: M_restrict_iff) apply (insert restrict_separation [of A], simp) done lemma (in M_basic_no_repl) Inter_closed [intro,simp]: "M(A) ==> M(\(A))" by (insert Inter_separation, simp add: Inter_def) lemma (in M_basic_no_repl) Int_closed [intro,simp]: "[| M(A); M(B) |] ==> M(A \ B)" apply (subgoal_tac "M({A,B})") apply (frule Inter_closed, force+) done lemma (in M_basic_no_repl) Diff_closed [intro,simp]: "[|M(A); M(B)|] ==> M(A-B)" by (insert Diff_separation, simp add: Diff_def) subsubsection\Some Facts About Separation Axioms\ lemma (in M_basic_no_repl) separation_conj: "[|separation(M,P); separation(M,Q)|] ==> separation(M, \z. P(z) & Q(z))" by (simp del: separation_closed add: separation_iff Collect_Int_Collect_eq [symmetric]) lemma (in M_basic_no_repl) separation_disj: "[|separation(M,P); separation(M,Q)|] ==> separation(M, \z. P(z) | Q(z))" by (simp del: separation_closed add: separation_iff Collect_Un_Collect_eq [symmetric]) lemma (in M_basic_no_repl) separation_neg: "separation(M,P) ==> separation(M, \z. ~P(z))" by (simp del: separation_closed add: separation_iff Diff_Collect_eq [symmetric]) lemma (in M_basic_no_repl) separation_imp: "[|separation(M,P); separation(M,Q)|] ==> separation(M, \z. P(z) \ Q(z))" by (simp add: separation_neg separation_disj not_disj_iff_imp [symmetric]) text\This result is a hint of how little can be done without the Reflection Theorem. The quantifier has to be bounded by a set. We also need another instance of Separation!\ lemma (in M_basic_no_repl) separation_rall: "[|M(Y); \y[M]. separation(M, \x. P(x,y)); \z[M]. strong_replacement(M, \x y. y = {u \ z . P(u,x)})|] ==> separation(M, \x. \y[M]. y\Y \ P(x,y))" apply (simp del: separation_closed rall_abs add: separation_iff Collect_rall_eq) apply (blast intro!: RepFun_closed dest: transM) done subsubsection\Functions and function space\ lemma (in M_basic_no_repl) succ_fun_eq2: "[|M(B); M(n->B)|] ==> succ(n) -> B = \{z. p \ (n->B)*B, \f[M]. \b[M]. p = & z = {cons(, f)}}" apply (simp add: succ_fun_eq) apply (blast dest: transM) done -(* lemma (in M_basic_no_repl) funspace_succ: - "[|M(n); M(B); M(n->B) |] ==> M(succ(n) -> B)" -apply (insert funspace_succ_replacement [of n], simp) -apply (force simp add: succ_fun_eq2 univalent_def) -done - -text\\<^term>\M\ contains all finite function spaces. Needed to prove the -absoluteness of transitive closure. See the definition of -\rtrancl_alt\ in in \WF_absolute.thy\.\ -lemma (in M_basic_no_repl) finite_funspace_closed [intro,simp]: - "[|n\nat; M(B)|] ==> M(n->B)" -apply (induct_tac n, simp) -apply (simp add: funspace_succ nat_into_M) -done - *) - lemma (in M_basic_no_repl) list_case'_closed [intro,simp]: "[|M(k); M(a); \x[M]. \y[M]. M(b(x,y))|] ==> M(list_case'(a,b,k))" apply (case_tac "quasilist(k)") apply (simp add: quasilist_def, force) apply (simp add: non_list_case) done lemma (in M_basic_no_repl) tl'_closed: "M(x) ==> M(tl'(x))" apply (simp add: tl'_def) apply (force simp add: quasilist_def) done sublocale M_basic \ mbnr:M_basic_no_repl using Inter_separation Diff_separation cartprod_separation image_separation converse_separation restrict_separation comp_separation pred_separation Memrel_separation is_recfun_separation power_ax by unfold_locales end diff --git a/thys/Transitive_Models/Partial_Functions_Relative.thy b/thys/Transitive_Models/Partial_Functions_Relative.thy --- a/thys/Transitive_Models/Partial_Functions_Relative.thy +++ b/thys/Transitive_Models/Partial_Functions_Relative.thy @@ -1,947 +1,947 @@ section\Cohen forcing notions\ theory Partial_Functions_Relative imports Cardinal_Library_Relative begin text\In this theory we introduce bounded partial functions and its relative version; for historical reasons the relative version is based on a proper definition of partial functions. We note that finite partial functions are easier and are used to prove some lemmas about finite sets in the theory \<^theory>\Transitive_Models.ZF_Library_Relative\.\ definition Fn :: "[i,i,i] \ i" where "Fn(\,I,J) \ \{y . d \ Pow(I), y=(d\J) \ d\\}" lemma domain_function_lepoll : assumes "function(r)" shows "domain(r) \ r" proof - let ?f="\x\domain(r) . \x,r`x\" have "\x, r ` x\ \ r" if "\x, y\ \ r" for x y proof - have "x\domain(r)" using that by auto with assms show ?thesis using function_apply_Pair by auto qed then have "?f \ inj(domain(r), r)" by(rule_tac lam_injective,auto) then show ?thesis unfolding lepoll_def by force qed lemma function_lepoll: assumes "r:d\J" shows "r \ d" proof - let ?f="\x\r . fst(x)" note assms Pi_iff[THEN iffD1,OF assms] moreover from calculation have "r`fst(x) = snd(x)" if "x\r" for x using that apply_equality by auto ultimately have "?f\inj(r,d)" by(rule_tac d= "\u . \u, r`u\" in lam_injective,auto) then show ?thesis unfolding lepoll_def by auto qed lemma function_eqpoll : assumes "r:d\J" shows "r \ d" using assms domain_of_fun domain_function_lepoll Pi_iff[THEN iffD1,OF assms] eqpollI[OF function_lepoll[OF assms]] subset_imp_lepoll by force lemma Fn_char : "Fn(\,I,J) = {f \ Pow(I\J) . function(f) \ f \ \}" (is "?L=?R") proof (intro equalityI subsetI) fix x assume "x \ ?R" moreover from this have "domain(x) \ Pow(I)" "domain(x) \ x" "x\ \" using domain_function_lepoll by auto ultimately show "x \ ?L" unfolding Fn_def using lesspoll_trans1 Pi_iff by (auto,rule_tac rev_bexI[of "domain(x) \ J"],auto) next fix x assume "x \ ?L" then obtain d where "x:d\J" "d \ Pow(I)" "d\\" unfolding Fn_def by auto moreover from this have "x\\" using function_lepoll[THEN lesspoll_trans1] by auto moreover from calculation have "x \ Pow(I\J)" "function(x)" using Pi_iff by auto ultimately show "x \ ?R" by simp qed lemma zero_in_Fn: assumes "0 < \" shows "0 \ Fn(\, I, J)" using lt_Card_imp_lesspoll assms zero_lesspoll unfolding Fn_def by (simp,rule_tac x="0\J" in bexI,simp) (rule ReplaceI[of _ 0],simp_all) lemma Fn_nat_eq_FiniteFun: "Fn(nat,I,J) = I -||> J" proof (intro equalityI subsetI) fix x assume "x \ I -||> J" then show "x \ Fn(nat,I,J)" proof (induct) case emptyI then show ?case using zero_in_Fn ltI by simp next case (consI a b h) then obtain d where "h:d\J" "d\nat" "d\I" unfolding Fn_def by auto moreover from this have "Finite(d)" using lesspoll_nat_is_Finite by simp ultimately have "h : d -||> J" using fun_FiniteFunI Finite_into_Fin by blast note \h:d\J\ moreover from this have "domain(cons(\a, b\, h)) = cons(a,d)" (is "domain(?h) = ?d") and "domain(h) = d" using domain_of_fun by simp_all moreover note consI moreover from calculation have "cons(\a, b\, h) \ cons(a,d) \ J" using fun_extend3 by simp moreover from \Finite(d)\ have "Finite(cons(a,d))" by simp moreover from this have "cons(a,d) \ nat" using Finite_imp_lesspoll_nat by simp ultimately show ?case unfolding Fn_def by (simp,rule_tac x="?d\J" in bexI) (force dest:app_fun)+ qed next fix x assume "x \ Fn(nat,I,J)" then obtain d where "x:d\J" "d \ Pow(I)" "d\nat" unfolding Fn_def by auto moreover from this have "Finite(d)" using lesspoll_nat_is_Finite by simp moreover from calculation have "d \ Fin(I)" using Finite_into_Fin[of d] Fin_mono by auto ultimately show "x \ I -||> J" using fun_FiniteFunI FiniteFun_mono by blast qed lemma Fn_nat_subset_Pow: "Fn(\,I,J) \ Pow(I\J)" using Fn_char by auto lemma FnI: assumes "p : d \ J" "d \ I" "d \ \" shows "p \ Fn(\,I,J)" using assms unfolding Fn_def by auto lemma FnD[dest]: assumes "p \ Fn(\,I,J)" shows "\d. p : d \ J \ d \ I \ d \ \" using assms unfolding Fn_def by auto lemma Fn_is_function: "p \ Fn(\,I,J) \ function(p)" unfolding Fn_def using fun_is_function by auto lemma Fn_csucc: assumes "Ord(\)" shows "Fn(csucc(\),I,J) = \{y . d \ Pow(I), y=(d\J) \ d\\}" using assms unfolding Fn_def using lesspoll_csucc by (simp) definition FnleR :: "i \ i \ o" (infixl \\\ 50) where "f \ g \ g \ f" lemma FnleR_iff_subset [iff]: "f \ g \ g \ f" unfolding FnleR_def .. definition Fnlerel :: "i \ i" where "Fnlerel(A) \ Rrel(\x y. x \ y,A)" definition Fnle :: "[i,i,i] \ i" where "Fnle(\,I,J) \ Fnlerel(Fn(\,I,J))" lemma FnleI[intro]: assumes "p \ Fn(\,I,J)" "q \ Fn(\,I,J)" "p \ q" shows "\p,q\ \ Fnle(\,I,J)" using assms unfolding Fnlerel_def Fnle_def FnleR_def Rrel_def by auto lemma FnleD[dest]: assumes "\p,q\ \ Fnle(\,I,J)" shows "p \ Fn(\,I,J)" "q \ Fn(\,I,J)" "p \ q" using assms unfolding Fnlerel_def Fnle_def FnleR_def Rrel_def by auto definition PFun_Space_Rel :: "[i,i\o, i] \ i" ("_\\<^bsup>_\<^esup>_") where "A \\<^bsup>M\<^esup> B \ {f \ Pow(A\B) . M(f) \ function(f)}" lemma (in M_library) PFun_Space_subset_Powrel : assumes "M(A)" "M(B)" shows "A \\<^bsup>M\<^esup> B = {f \ Pow\<^bsup>M\<^esup>(A\B) . function(f)}" using Pow_rel_char assms unfolding PFun_Space_Rel_def by auto lemma (in M_library) PFun_Space_closed : assumes "M(A)" "M(B)" shows "M(A \\<^bsup>M\<^esup> B)" using assms PFun_Space_subset_Powrel separation_is_function by auto lemma pfun_is_function : "f \ A\\<^bsup>M\<^esup> B \ function(f)" unfolding PFun_Space_Rel_def by simp lemma pfun_range : "f \ A \\<^bsup>M\<^esup> B \ range(f) \ B" unfolding PFun_Space_Rel_def by auto lemma pfun_domain : "f \ A \\<^bsup>M\<^esup> B \ domain(f) \ A" unfolding PFun_Space_Rel_def by auto lemma Un_filter_fun_space_closed: assumes "G \ I \ J" "\ f g . f\G \ g\G \ \d\I\ J . d \ f \ d \ g" shows "\G \ Pow(I\J)" "function(\G)" proof - from assms show "\G \ Pow(I\J)" using Union_Pow_iff unfolding Pi_def by auto next show "function(\G)" unfolding function_def proof(auto) fix B B' x y y' assume "B \ G" "\x, y\ \ B" "B' \ G" "\x, y'\ \ B'" moreover from assms this have "B \ I \ J" "B' \ I \ J" by auto moreover from calculation assms(2)[of B B'] obtain d where "d \ B" "d \ B'" "d\I \ J" "\x, y\ \ d" "\x, y'\ \ d" using subsetD[OF \G\_\] by auto then show "y=y'" using fun_is_function[OF \d\_\] unfolding function_def by force qed qed lemma Un_filter_is_fun : assumes "G \ I \ J" "\ f g . f\G \ g\G \ \d\I\ J . d\f \ d\g" "G\0" shows "\G \ I \ J" using assms Un_filter_fun_space_closed Pi_iff proof(simp_all) show "I\domain(\G)" proof - from \G\0\ obtain f where "f\\G" "f\G" by auto with \G\_\ have "f\I\J" by auto then show ?thesis using subset_trans[OF _ domain_mono[OF \f\\G\],of I] unfolding Pi_def by auto qed qed context M_Pi begin lemma mem_function_space_relD: assumes "f \ function_space_rel(M,A,y)" "M(A)" "M(y)" shows "f \ A \ y" and "M(f)" using assms function_space_rel_char by simp_all lemma pfunI : assumes "C\A" "f \ C \\<^bsup>M\<^esup> B" "M(C)" "M(B)" shows "f\ A \\<^bsup>M\<^esup> B" proof - from assms have "f \ C\B" "M(f)" using mem_function_space_relD by simp_all with assms show ?thesis using Pi_iff unfolding PFun_Space_Rel_def by auto qed lemma zero_in_PFun_rel: assumes "M(I)" "M(J)" shows "0 \ I \\<^bsup>M\<^esup> J" using pfunI[of 0] nonempty mem_function_space_rel_abs assms by simp lemma pfun_subsetI : assumes "f \ A \\<^bsup>M\<^esup> B" "g\f" "M(g)" shows "g\ A \\<^bsup>M\<^esup> B" using assms function_subset unfolding PFun_Space_Rel_def by auto lemma pfun_Un_filter_closed: assumes "G \I\\<^bsup>M\<^esup> J" "\ f g . f\G \ g\G \ \d\I\\<^bsup>M\<^esup> J . d\f \ d\g" shows "\G \ Pow(I\J)" "function(\G)" proof - from assms show "\G \ Pow(I\J)" using Union_Pow_iff unfolding PFun_Space_Rel_def by auto next show "function(\G)" unfolding function_def proof(auto) fix B B' x y y' assume "B \ G" "\x, y\ \ B" "B' \ G" "\x, y'\ \ B'" moreover from calculation assms obtain d where "d \ I \\<^bsup>M\<^esup> J" "function(d)" "\x, y\ \ d" "\x, y'\ \ d" using pfun_is_function by force ultimately show "y=y'" unfolding function_def by auto qed qed lemma pfun_Un_filter_closed'': assumes "G \I\\<^bsup>M\<^esup> J" "\ f g . f\G \ g\G \ \d\G . d\f \ d\g" shows "\G \ Pow(I\J)" "function(\G)" proof - from assms have "\ f g . f\G \ g\G \ \d\I\\<^bsup>M\<^esup> J . d\f \ d\g" using subsetD[OF assms(1),THEN [2] bexI] by force then show "\G \ Pow(I\J)" "function(\G)" using assms pfun_Un_filter_closed unfolding PFun_Space_Rel_def by auto qed lemma pfun_Un_filter_closed': assumes "G \I\\<^bsup>M\<^esup> J" "\ f g . f\G \ g\G \ \d\G . d\f \ d\g" "M(G)" shows "\G \ I\\<^bsup>M\<^esup> J" using assms pfun_Un_filter_closed'' unfolding PFun_Space_Rel_def by auto lemma pfunD : assumes "f \ A\\<^bsup>M\<^esup> B" shows "\C[M]. C\A \ f \ C\B" proof - note assms moreover from this have "f\Pow(A\B)" "function(f)" "M(f)" unfolding PFun_Space_Rel_def by simp_all moreover from this have "domain(f) \ A" "f \ domain(f) \ B" "M(domain(f))" using assms Pow_iff[of f "A\B"] domain_subset Pi_iff by auto ultimately show ?thesis by auto qed lemma pfunD_closed : assumes "f \ A\\<^bsup>M\<^esup> B" shows "M(f)" using assms unfolding PFun_Space_Rel_def by simp lemma pfun_singletonI : assumes "x \ A" "b \ B" "M(A)" "M(B)" shows "{\x,b\} \ A\\<^bsup>M\<^esup> B" using assms transM[of x A] transM[of b B] unfolding PFun_Space_Rel_def function_def by auto lemma pfun_unionI : assumes "f \ A\\<^bsup>M\<^esup> B" "g \ A\\<^bsup>M\<^esup> B" "domain(f)\domain(g)=0" shows "f\g \ A\\<^bsup>M\<^esup> B" using assms unfolding PFun_Space_Rel_def function_def by blast lemma (in M_library) pfun_restrict_eq_imp_compat: assumes "f \ I\\<^bsup>M\<^esup> J" "g \ I\\<^bsup>M\<^esup> J" "M(J)" "restrict(f, domain(f) \ domain(g)) = restrict(g, domain(f) \ domain(g))" shows "f \ g \ I\\<^bsup>M\<^esup> J" proof - note assms moreover from this obtain C D where "f : C \ J" "C \ I" "D \ I" "M(C)" "M(D)" "g : D \ J" using pfunD[of f] pfunD[of g] by force moreover from calculation have "f\g \ C\D \ J" using restrict_eq_imp_Un_into_Pi'[OF \f\C\_\ \g\D\_\] by auto ultimately show ?thesis using pfunI[of "C\D" _ "f\g"] Un_subset_iff pfunD_closed function_space_rel_char by auto qed lemma FiniteFun_pfunI : assumes "f \ A-||> B" "M(A)" "M(B)" shows "f \ A \\<^bsup>M\<^esup> B" using assms(1) proof(induct) case emptyI then show ?case using assms nonempty mem_function_space_rel_abs pfunI[OF empty_subsetI, of 0] by simp next case (consI a b h) note consI moreover from this have "M(a)" "M(b)" "M(h)" "domain(h) \ A" using transM[OF _ \M(A)\] transM[OF _ \M(B)\] FinD FiniteFun_domain_Fin pfunD_closed by simp_all moreover from calculation have "{a}\domain(h)\A" "M({a}\domain(h))" "M(cons(,h))" "domain(cons(,h)) = {a}\domain(h)" by auto moreover from calculation have "cons(,h) \ {a}\domain(h) \ B" using FiniteFun_is_fun[OF FiniteFun.consI, of a A b B h] by auto ultimately show "cons(,h) \ A \\<^bsup>M\<^esup> B" using assms mem_function_space_rel_abs pfunI by simp_all qed lemma PFun_FiniteFunI : assumes "f \ A \\<^bsup>M\<^esup> B" "Finite(f)" shows "f \ A-||> B" proof - from assms have "f\Fin(A\B)" "function(f)" using Finite_Fin Pow_iff unfolding PFun_Space_Rel_def by auto then show ?thesis using FiniteFunI by simp qed end \ \\<^locale>\M_Pi\\ -(* Fn_rel should be the relativization *) definition Fn_rel :: "[i\o,i,i,i] \ i" (\Fn\<^bsup>_\<^esup>'(_,_,_')\) where "Fn_rel(M,\,I,J) \ {f \ I\\<^bsup>M\<^esup> J . f \\<^bsup>M\<^esup> \}" +\ \Eventually we can define \<^term>\Fn_rel\ as the relativization of \<^term>\Fn\\ context M_library begin lemma Fn_rel_subset_PFun_rel : "Fn\<^bsup>M\<^esup>(\,I,J) \ I\\<^bsup>M\<^esup> J" unfolding Fn_rel_def by auto lemma Fn_relI[intro]: assumes "f : d \ J" "d \ I" "f \\<^bsup>M\<^esup> \" "M(d)" "M(J)" "M(f)" shows "f \ Fn_rel(M,\,I,J)" using assms pfunI mem_function_space_rel_abs unfolding Fn_rel_def by auto lemma Fn_relD[dest]: assumes "p \ Fn_rel(M,\,I,J)" shows "\C[M]. C\I \ p : C \ J \ p \\<^bsup>M\<^esup> \" using assms pfunD unfolding Fn_rel_def by simp lemma Fn_rel_is_function: assumes "p \ Fn_rel(M,\,I,J)" shows "function(p)" "M(p)" "p \\<^bsup>M\<^esup> \" "p\ I\\<^bsup>M\<^esup> J" using assms unfolding Fn_rel_def PFun_Space_Rel_def by simp_all lemma Fn_rel_mono: assumes "p \ Fn_rel(M,\,I,J)" "\ \\<^bsup>M\<^esup> \'" "M(\)" "M(\')" shows "p \ Fn_rel(M,\',I,J)" using assms lesspoll_rel_trans[OF _ assms(2)] cardinal_rel_closed Fn_rel_is_function(2)[OF assms(1)] unfolding Fn_rel_def by simp lemma Fn_rel_mono': assumes "p \ Fn_rel(M,\,I,J)" "\ \\<^bsup>M\<^esup> \'" "M(\)" "M(\')" shows "p \ Fn_rel(M,\',I,J)" proof - note assms then consider "\ \\<^bsup>M\<^esup> \'" | "\ \\<^bsup>M\<^esup> \'" using lepoll_rel_iff_leqpoll_rel by auto then show ?thesis proof(cases) case 1 with assms show ?thesis using Fn_rel_mono by simp next case 2 then show ?thesis using assms cardinal_rel_closed Fn_rel_is_function[OF assms(1)] lesspoll_rel_eq_trans unfolding Fn_rel_def by simp qed qed lemma Fn_csucc: assumes "Ord(\)" "M(\)" shows "Fn_rel(M,(\\<^sup>+)\<^bsup>M\<^esup>,I,J) = {p\ I\\<^bsup>M\<^esup> J . p \\<^bsup>M\<^esup> \}" (is "?L=?R") using assms proof(intro equalityI) show "?L \ ?R" proof(intro subsetI) fix p assume "p\?L" then have "p \\<^bsup>M\<^esup> csucc_rel(M,\)" "M(p)" "p\ I\\<^bsup>M\<^esup> J" using Fn_rel_is_function by simp_all then show "p\?R" using assms lesspoll_rel_csucc_rel by simp qed next show "?R\?L" proof(intro subsetI) fix p assume "p\?R" then have "p\ I\\<^bsup>M\<^esup> J" "p \\<^bsup>M\<^esup> \" using assms lesspoll_rel_csucc_rel by simp_all then show "p\?L" using assms lesspoll_rel_csucc_rel pfunD_closed unfolding Fn_rel_def by simp qed qed lemma Finite_imp_lesspoll_nat: assumes "Finite(A)" shows "A \ nat" using assms subset_imp_lepoll[OF naturals_subset_nat] eq_lepoll_trans n_lesspoll_nat eq_lesspoll_trans unfolding Finite_def lesspoll_def by auto lemma FinD_Finite : assumes "a\Fin(A)" shows "Finite(a)" using assms by(induct,simp_all) lemma Fn_rel_nat_eq_FiniteFun: assumes "M(I)" "M(J)" shows "I -||> J = Fn_rel(M,\,I,J)" proof(intro equalityI subsetI) fix p assume "p\I -||> J" with assms have "p\ I \\<^bsup>M\<^esup> J" "Finite(p)" using FiniteFun_pfunI FinD_Finite[OF subsetD[OF FiniteFun.dom_subset,OF \p\_\]] by auto moreover from this have "p \\<^bsup>M\<^esup> \" using Finite_lesspoll_rel_nat pfunD_closed[of p] n_lesspoll_rel_nat by simp ultimately show "p\ Fn_rel(M,\,I,J)" unfolding Fn_rel_def by simp next fix p assume "p\ Fn_rel(M,\,I,J)" then have "p\ I \\<^bsup>M\<^esup> J" "p \\<^bsup>M\<^esup> \" unfolding Fn_rel_def by simp_all moreover from this have "Finite(p)" using Finite_cardinal_rel_Finite lesspoll_rel_nat_is_Finite_rel pfunD_closed cardinal_rel_closed[of p] Finite_cardinal_rel_iff'[THEN iffD1] by simp ultimately show "p\I -||> J" using PFun_FiniteFunI by simp qed lemma Fn_nat_abs: assumes "M(I)" "M(J)" shows "Fn(nat,I,J) = Fn_rel(M,\,I,J)" using assms Fn_rel_nat_eq_FiniteFun Fn_nat_eq_FiniteFun by simp lemma Fn_rel_singletonI: assumes "x \ I" "j \ J" "1 \\<^bsup>M\<^esup> \" "M(\)" "M(I)" "M(J)" shows "{\x,j\} \ Fn\<^bsup>M\<^esup>(\,I,J)" using pfun_singletonI transM[of x] transM[of j] assms eq_lesspoll_rel_trans[OF singleton_eqpoll_rel_1] unfolding Fn_rel_def by auto end \ \\<^locale>\M_library\\ definition Fnle_rel :: "[i\o,i,i,i] \ i" (\Fnle\<^bsup>_\<^esup>'(_,_,_')\) where "Fnle_rel(M,\,I,J) \ Fnlerel(Fn\<^bsup>M\<^esup>(\,I,J))" abbreviation Fn_r_set :: "[i,i,i,i] \ i" (\Fn\<^bsup>_\<^esup>'(_,_,_')\) where "Fn_r_set(M) \ Fn_rel(##M)" abbreviation Fnle_r_set :: "[i,i,i,i] \ i" (\Fnle\<^bsup>_\<^esup>'(_,_,_')\) where "Fnle_r_set(M) \ Fnle_rel(##M)" context M_library begin lemma Fnle_relI[intro]: assumes "p \ Fn_rel(M,\,I,J)" "q \ Fn_rel(M,\,I,J)" "p \ q" shows "\p, q\ \ Fnle_rel(M,\,I,J)" using assms unfolding Fnlerel_def Fnle_rel_def FnleR_def Rrel_def by auto lemma Fnle_relD[dest]: assumes "\p, q\ \ Fnle_rel(M,\,I,J)" shows "p \ Fn_rel(M,\,I,J)" "q \ Fn_rel(M,\,I,J)" "p \ q" using assms unfolding Fnlerel_def Fnle_rel_def FnleR_def Rrel_def by auto lemma Fn_rel_closed[intro,simp]: assumes "M(\)" "M(I)" "M(J)" shows "M(Fn\<^bsup>M\<^esup>(\,I,J))" using assms separation_cardinal_rel_lesspoll_rel PFun_Space_closed unfolding Fn_rel_def by auto lemma Fn_rel_subset_Pow: assumes "M(\)" "M(I)" "M(J)" shows "Fn\<^bsup>M\<^esup>(\,I,J) \ Pow(I\J)" unfolding Fn_rel_def PFun_Space_Rel_def by auto lemma Fnle_rel_closed[intro,simp]: assumes "M(\)" "M(I)" "M(J)" shows "M(Fnle\<^bsup>M\<^esup>(\,I,J))" unfolding Fnle_rel_def Fnlerel_def Rrel_def FnleR_def using assms supset_separation Fn_rel_closed by auto lemma zero_in_Fn_rel: assumes "0<\" "M(\)" "M(I)" "M(J)" shows "0 \ Fn\<^bsup>M\<^esup>(\, I, J)" unfolding Fn_rel_def using zero_in_PFun_rel zero_lesspoll_rel assms by simp lemma zero_top_Fn_rel: assumes "p\Fn\<^bsup>M\<^esup>(\, I, J)" "0<\" "M(\)" "M(I)" "M(J)" shows "\p, 0\ \ Fnle\<^bsup>M\<^esup>(\, I, J)" using assms zero_in_Fn_rel unfolding preorder_on_def refl_def trans_on_def by auto lemma preorder_on_Fnle_rel: assumes "M(\)" "M(I)" "M(J)" shows "preorder_on(Fn\<^bsup>M\<^esup>(\, I, J), Fnle\<^bsup>M\<^esup>(\, I, J))" unfolding preorder_on_def refl_def trans_on_def by blast end \ \\<^locale>\M_library\\ 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 end \ \\<^locale>\M_cardinal_library\\ context M_cardinal_library_extra begin 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 lemma 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 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 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 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 dense_dom_dense: assumes "x \ I" "InfCard\<^bsup>M\<^esup>(\)" "M(I)" "M(J)" "z\J" "M(\)" "p\Fn\<^bsup>M\<^esup>(\, I, J)" shows "\d\{ p \ Fn\<^bsup>M\<^esup>(\, I, J) . x \ domain(p) }. \d,p\ \ Fnle\<^bsup>M\<^esup>(\, I, J)" proof (cases "x \ domain(p)") case True with \x \ I\ \p \ Fn\<^bsup>M\<^esup>(\, I, J)\ show ?thesis by auto next case False note \p \ Fn\<^bsup>M\<^esup>(\, I, J)\ moreover from this and False and assms have "cons(\x,z\, p) \ Fn\<^bsup>M\<^esup>(\, I, J)" "M(x)" using cons_in_Fn_rel by (auto dest:transM) ultimately show ?thesis using Fn_relD by blast qed lemma (in M_cardinal_library) domain_lepoll_rel : assumes "function(r)" "M(r)" shows "domain(r) \\<^bsup>M\<^esup> r" proof - let ?f="\x\domain(r) . \x,r`x\" have "\x, r ` x\ \ r" if "\x, y\ \ r" for x y proof - have "x\domain(r)" using that by auto with assms show ?thesis using function_apply_Pair by auto qed then have "?f \ inj(domain(r), r)" by(rule_tac lam_injective,auto) moreover note assms moreover from calculation have "M(?f)" using lam_replacement_constant[of r] lam_replacement_identity assms lam_replacement_apply lam_replacement_Pair[THEN [5] lam_replacement_hcomp2] by(rule_tac lam_replacement_imp_lam_closed,auto dest:transM[of _ r]) ultimately have "?f \ inj\<^bsup>M\<^esup>(domain(r),r)" using inj_rel_char by auto with \M(?f)\ show ?thesis using lepoll_relI by simp qed lemma dense_surj_dense: assumes "x \ J" "InfCard\<^bsup>M\<^esup>(\)" "M(I)" "M(J)" "M(\)" "p\Fn\<^bsup>M\<^esup>(\, I, J)" "\ \\<^bsup>M\<^esup> I" shows "\d\{ p \ Fn\<^bsup>M\<^esup>(\, I, J) . x \ range(p) }. \d,p\ \ Fnle\<^bsup>M\<^esup>(\, I, J)" proof - from \p \ _\ \M(J)\ \x\_\ lesspoll_rel_def have "domain(p) \ I" "M(p)" "p \\<^bsup>M\<^esup> \" "M(x)" "function(p)" using pfun_domain[OF Fn_rel_is_function(4)] Fn_rel_is_function transM[of x J] by simp_all moreover from calculation assms have 1:"domain(p) \\<^bsup>M\<^esup> \" "M(domain(p))" "domain(p) \\<^bsup>M\<^esup> I" using domain_lepoll_rel lesspoll_rel_trans1[of "domain(p)" p \] lesspoll_rel_trans2 by auto with calculation \\ \\<^bsup>M\<^esup> I\ \M(I)\ \M(\)\ have "domain(p) \ I" proof - {assume "domain(p) = I" with 1 have "domain(p) \\<^bsup>M\<^esup> domain(p)" by auto with \M(domain(p))\ have False using lesspoll_rel_irrefl[of "domain(p)"] by simp } then show ?thesis by auto qed ultimately obtain \ where "\ \ domain(p)" "\\I" by force moreover note assms moreover from calculation have "cons(\\,x\, p) \ Fn\<^bsup>M\<^esup>(\, I, J)" using InfCard_rel_Aleph_rel cons_in_Fn_rel[of \ p \ I J x] by auto ultimately show ?thesis using Fnle_relI by blast qed end \ \\<^locale>\M_cardinal_library_extra\\ end diff --git a/thys/Transitive_Models/Pointed_DC_Relative.thy b/thys/Transitive_Models/Pointed_DC_Relative.thy --- a/thys/Transitive_Models/Pointed_DC_Relative.thy +++ b/thys/Transitive_Models/Pointed_DC_Relative.thy @@ -1,469 +1,459 @@ section\Relative DC\ theory Pointed_DC_Relative imports Cardinal_Library_Relative begin consts dc_witness :: "i \ i \ i \ i \ i \ i" primrec wit0 : "dc_witness(0,A,a,s,R) = a" witrec : "dc_witness(succ(n),A,a,s,R) = s`{x\A. \dc_witness(n,A,a,s,R),x\\R}" lemmas dc_witness_def = dc_witness_nat_def relativize functional "dc_witness" "dc_witness_rel" relationalize "dc_witness_rel" "is_dc_witness" (* definition is_dc_witness_fm where "is_dc_witness_fm(na, A, a, s, R, e) \ is_transrec_fm (is_nat_case_fm (a +\<^sub>\ 8, (\\\\4`2 is 0\ \ (\\\\s +\<^sub>\ 12`0 is 2\ \ Collect_fm(A +\<^sub>\ 12, \(\\\0 = 0\\) \ (\\\\0 \ R +\<^sub>\ 14\ \ pair_fm(3, 1, 0) \\)\, 0) \\)\\), 2, 0), na, e)" *) schematic_goal sats_is_dc_witness_fm_auto: assumes "na < length(env)" "e < length(env)" shows " na \ \ \ A \ \ \ a \ \ \ s \ \ \ R \ \ \ e \ \ \ env \ list(Aa) \ 0 \ Aa \ is_dc_witness(##Aa, nth(na, env), nth(A, env), nth(a, env), nth(s, env), nth(R, env), nth(e, env)) \ Aa, env \ ?fm(nat, A, a, s, R, e)" unfolding is_dc_witness_def is_recursor_def by (rule is_transrec_iff_sats | simp_all) (rule iff_sats is_nat_case_iff_sats is_eclose_iff_sats sep_rules | simp add:assms)+ synthesize "is_dc_witness" from_schematic manual_arity for "is_dc_witness_fm" unfolding is_dc_witness_fm_def apply (subst arity_transrec_fm) apply (simp add:arity) defer 3 apply (simp add:arity) defer apply (subst arity_is_nat_case_fm) apply (simp add:arity del:arity_transrec_fm) prefer 5 by (simp add:arity del:arity_transrec_fm)+ definition dcwit_body :: "[i,i,i,i,i] \ o" where "dcwit_body(A,a,g,R) \ \p. snd(p) = dc_witness(fst(p), A, a, g, R)" relativize functional "dcwit_body" "dcwit_body_rel" relationalize "dcwit_body_rel" "is_dcwit_body" synthesize "is_dcwit_body" from_definition assuming "nonempty" arity_theorem for "is_dcwit_body_fm" context M_replacement begin lemma dc_witness_closed[intro,simp]: assumes "M(n)" "M(A)" "M(a)" "M(s)" "M(R)" "n\nat" shows "M(dc_witness(n,A,a,s,R))" using \n\nat\ proof(induct) case 0 with \M(a)\ show ?case unfolding dc_witness_def by simp next case (succ x) with assms have "M(dc_witness(x,A,a,s,R))" (is "M(?b)") by simp moreover from this assms have "M(({?b}\A)\R)" (is "M(?X)") by auto moreover have "{x\A. \?b,x\\R} = {snd(y) . y\?X}" (is "_ = ?Y") by(intro equalityI subsetI,force,auto) moreover from calculation have "M(?Y)" using lam_replacement_snd lam_replacement_imp_strong_replacement RepFun_closed snd_closed[OF transM] by auto ultimately show ?case using \M(s)\ apply_closed unfolding dc_witness_def by simp qed lemma dc_witness_rel_char: assumes "M(A)" shows "dc_witness_rel(M,n,A,a,s,R) = dc_witness(n,A,a,s,R)" proof - from assms have "{x \ A . \rec, x\ \ R} = {x \ A . M(x) \ \rec, x\ \ R}" for rec by (auto dest:transM) then show ?thesis unfolding dc_witness_def dc_witness_rel_def by simp qed lemma (in M_basic) first_section_closed: assumes "M(A)" "M(a)" "M(R)" shows "M({x \ A . \a, x\ \ R})" proof - have "{x \ A . \a, x\ \ R} = range({a} \ range(R) \ R) \ A" by (intro equalityI) auto with assms show ?thesis by simp qed lemma witness_into_A [TC]: assumes "a\A" "\X[M]. X\0 \ X\A \ s`X\A" "\y\A. {x\A. \y,x\\R } \ 0" "n\nat" "M(A)" "M(a)" "M(s)" "M(R)" shows "dc_witness(n, A, a, s, R)\A" using \n\nat\ proof(induct n) case 0 then show ?case using \a\A\ by simp next case (succ x) with succ assms(1,3-) show ?case using nat_into_M first_section_closed by (simp, rule_tac rev_subsetD, rule_tac assms(2)[rule_format]) auto qed end \ \\<^locale>\M_replacement\\ locale M_DC = M_trancl + M_replacement + M_eclose + assumes separation_is_dcwit_body: "M(A) \ M(a) \ M(g) \ M(R) \ separation(M,is_dcwit_body(M, A, a, g, R))" and dcwit_replacement:"Ord(na) \ M(na) \ M(A) \ M(a) \ M(s) \ M(R) \ transrec_replacement (M, \n f ntc. 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, ntc),na)" begin lemma is_dc_witness_iff: assumes "Ord(na)" "M(na)" "M(A)" "M(a)" "M(s)" "M(R)" "M(res)" shows "is_dc_witness(M, na, A, a, s, R, res) \ res = dc_witness_rel(M, na, A, a, s, R)" proof - note assms moreover from this have "{x \ A . M(x) \ \f, x\ \ R} = {x \ A . \f, x\ \ R}" for f by (auto dest:transM) moreover from calculation have "M(fm) \ M({x \ A . M(x) \ \fm, x\ \ R})" for fm using first_section_closed by (auto dest:transM) moreover from calculation have "M(x) \ M(z) \ M(mesa) \ (\ya[M]. pair(M, x, ya, z) \ is_wfrec(M, \n f. is_nat_case(M, a, \m bmfm. \fm[M]. is_apply(M, f, m, fm) \ is_apply(M, s, {x \ A . \fm, x\ \ R}, bmfm), n), mesa, x, ya)) \ (\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. M(x) \ \fm, x\ \ R, cp) \ is_apply(M, s, cp, bmfm),n), mesa, x, y))" for x z mesa by simp moreover from calculation show ?thesis using assms dcwit_replacement[of na A a s R] unfolding is_dc_witness_def dc_witness_rel_def by (rule_tac recursor_abs) (auto dest:transM) qed lemma dcwit_body_abs: "fst(x) \ \ \ M(A) \ M(a) \ M(g) \ M(R) \ M(x) \ is_dcwit_body(M,A,a,g,R,x) \ dcwit_body(A,a,g,R,x)" using pair_in_M_iff apply_closed transM[of _ A] is_dc_witness_iff[of "fst(x)" "A" "a" "g" "R" "snd(x)"] fst_snd_closed dc_witness_closed unfolding dcwit_body_def is_dcwit_body_def by (auto dest:transM simp:absolut dc_witness_rel_char del:bexI intro!:bexI) lemma Lambda_dc_witness_closed: assumes "g \ Pow\<^bsup>M\<^esup>(A)-{0} \ A" "a\A" "\y\A. {x \ A . \y, x\ \ R} \ 0" "M(g)" "M(A)" "M(a)" "M(R)" shows "M(\n\nat. dc_witness(n,A,a,g,R))" proof - from assms have "(\n\nat. dc_witness(n,A,a,g,R)) = {p \ \ \ A . is_dcwit_body(M,A,a,g,R,p)}" using witness_into_A[of a A g R] Pow_rel_char apply_type[of g "{x \ Pow(A) . M(x)}-{0}" "\_.A"] unfolding lam_def split_def apply (intro equalityI subsetI) apply (auto) (* slow *) by (subst dcwit_body_abs, simp_all add:transM[of _ \] dcwit_body_def, subst (asm) dcwit_body_abs, auto dest:transM simp:dcwit_body_def) (* by (intro equalityI subsetI, auto) (* Extremely slow *) (subst dcwit_body_abs, simp_all add:transM[of _ \] dcwit_body_def, subst (asm) dcwit_body_abs, auto dest:transM simp:dcwit_body_def) *) with assms show ?thesis using separation_is_dcwit_body dc_witness_rel_char unfolding split_def by simp qed lemma witness_related: assumes "a\A" "\X[M]. X\0 \ X\A \ s`X\X" "\y\A. {x\A. \y,x\\R } \ 0" "n\nat" "M(a)" "M(A)" "M(s)" "M(R)" "M(n)" shows "\dc_witness(n, A, a, s, R),dc_witness(succ(n), A, a, s, R)\\R" proof - note assms moreover from this have "(\X[M]. X\0 \ X\A \ s`X\A)" by auto moreover from calculation have "dc_witness(n, A, a, s, R)\A" (is "?x \ A") using witness_into_A[of _ _ s R n] by simp moreover from assms have "M({x \ A . \dc_witness(n, A, a, s, R), x\ \ R})" using first_section_closed[of A "dc_witness(n, A, a, s, R)"] by simp ultimately show ?thesis by auto qed lemma witness_funtype: assumes "a\A" "\X[M]. X\0 \ X\A \ s`X \ A" "\y\A. {x\A. \y,x\\R} \ 0" "M(A)" "M(a)" "M(s)" "M(R)" shows "(\n\nat. dc_witness(n, A, a, s, R)) \ nat \ A" (is "?f \ _ \ _") proof - have "?f \ nat \ {dc_witness(n, A, a, s, R). n\nat}" (is "_ \ _ \ ?B") using lam_funtype assms by simp then have "?B \ A" using witness_into_A assms by auto with \?f \ _\ show ?thesis using fun_weaken_type by simp qed lemma witness_to_fun: assumes "a\A" "\X[M]. X\0 \ X\A \ s`X\A" "\y\A. {x\A. \y,x\\R } \ 0" "M(A)" "M(a)" "M(s)" "M(R)" shows "\f \ nat\A. \n\nat. f`n =dc_witness(n,A,a,s,R)" using assms bexI[of _ "\n\nat. dc_witness(n,A,a,s,R)"] witness_funtype by simp end \ \\<^locale>\M_DC\\ locale M_library_DC = M_library + M_DC begin -(* Should port the whole AC theory, including the absolute version - of the following theorem *) lemma AC_M_func: assumes "\x. x \ A \ (\y. y \ x)" "M(A)" shows "\f \ A \\<^bsup>M\<^esup> \(A). \x \ A. f`x \ x" proof - from \M(A)\ - interpret mpiA:M_Pi_assumptions _ A "\x. x" - using Pi_replacement Pi_separation lam_replacement_identity - lam_replacement_Sigfun[THEN lam_replacement_imp_strong_replacement] - by unfold_locales (simp_all add:transM[of _ A]) + interpret mpiA:M_Pi_assumption_repl _ A "\x . x" + by unfold_locales (simp_all add:lam_replacement_identity) + from \M(A)\ + interpret mpi2:M_Pi_assumption_repl _ "\A" "\x . x" + by unfold_locales (simp_all add:lam_replacement_identity) from \M(A)\ interpret mpic_A:M_Pi_assumptions_choice _ A "\x. x" apply unfold_locales using lam_replacement_constant lam_replacement_identity lam_replacement_identity[THEN lam_replacement_imp_strong_replacement] lam_replacement_minimum[THEN [5] lam_replacement_hcomp2] unfolding lam_replacement_def[symmetric] by auto - from \M(A)\ - interpret mpi2:M_Pi_assumptions2 _ A "\_. \A" "\x. x" - using Pi_replacement Pi_separation lam_replacement_constant - lam_replacement_Sigfun[THEN lam_replacement_imp_strong_replacement, - of "\_. \A"] Pi_replacement1[of _ "\A"] transM[of _ "A"] - by unfold_locales auto from assms show ?thesis - using mpi2.Pi_rel_type apply_type mpiA.mem_Pi_rel_abs mpi2.mem_Pi_rel_abs - function_space_rel_char - by (rule_tac mpic_A.AC_Pi_rel[THEN rexE], simp, rule_tac x=x in bexI) - (auto, rule_tac C="\x. x" in Pi_type, auto) + using apply_type mpiA.mem_Pi_rel_abs mpi2.mem_Pi_rel_abs function_space_rel_char + by (rule_tac mpic_A.AC_Pi_rel[THEN rexE], simp,rename_tac f,rule_tac x=f in bexI) + (auto, rule_tac C="\x. x" in Pi_type,auto) qed lemma non_empty_family: "[| 0 \ A; x \ A |] ==> \y. y \ x" by (subgoal_tac "x \ 0", blast+) lemma AC_M_func0: "0 \ A \ M(A) \ \f \ A \\<^bsup>M\<^esup> \(A). \x \ A. f`x \ x" by (rule AC_M_func) (simp_all add: non_empty_family) lemma AC_M_func_Pow_rel: assumes "M(C)" shows "\f \ (Pow\<^bsup>M\<^esup>(C)-{0}) \\<^bsup>M\<^esup> C. \x \ Pow\<^bsup>M\<^esup>(C)-{0}. f`x \ x" proof - have "0\Pow\<^bsup>M\<^esup>(C)-{0}" by simp with assms obtain f where "f \ (Pow\<^bsup>M\<^esup>(C)-{0}) \\<^bsup>M\<^esup> \(Pow\<^bsup>M\<^esup>(C)-{0})" "\x \ Pow\<^bsup>M\<^esup>(C)-{0}. f`x \ x" using AC_M_func0[OF \0\_\] by auto moreover have "x\C" if "x \ Pow\<^bsup>M\<^esup>(C) - {0}" for x using that Pow_rel_char assms by auto moreover have "\(Pow\<^bsup>M\<^esup>(C) - {0}) \ C" using assms Pow_rel_char by auto ultimately show ?thesis using assms function_space_rel_char by (rule_tac bexI,auto,rule_tac Pi_weaken_type,simp_all) qed theorem pointed_DC: assumes "\x\A. \y\A. \x,y\\ R" "M(A)" "M(R)" shows "\a\A. \f \ nat\\<^bsup>M\<^esup> A. f`0 = a \ (\n \ nat. \f`n,f`succ(n)\\R)" proof - have 0:"\y\A. {x \ A . \y, x\ \ R} \ 0" using assms by auto from \M(A)\ obtain g where 1: "g \ Pow\<^bsup>M\<^esup>(A)-{0} \ A" "\X[M]. X \ 0 \ X \ A \ g ` X \ X" "M(g)" using AC_M_func_Pow_rel[of A] mem_Pow_rel_abs function_space_rel_char[of "Pow\<^bsup>M\<^esup>(A)-{0}" A] by auto then have 2:"\X[M]. X \ 0 \ X \ A \ g ` X \ A" by auto { fix a assume "a\A" let ?f ="\n\nat. dc_witness(n,A,a,g,R)" note \a\A\ moreover from this have f0: "?f`0 = a" by simp moreover note \a\A\ \M(g)\ \M(A)\ \M(R)\ moreover from calculation have "\?f ` n, ?f ` succ(n)\ \ R" if "n\nat" for n using witness_related[OF \a\A\ _ 0, of g n] 1 that by (auto dest:transM) ultimately have "\f[M]. f\nat \ A \ f ` 0 = a \ (\n\nat. \f ` n, f ` succ(n)\ \ R)" using 0 1 \a\_\ by (rule_tac x="(\n\\. dc_witness(n, A, a, g, R))" in rexI) (simp_all, rule_tac witness_funtype, auto intro!: Lambda_dc_witness_closed dest:transM) } with \M(A)\ show ?thesis using function_space_rel_char by auto qed lemma aux_DC_on_AxNat2 : "\x\A\nat. \y\A. \x,\y,succ(snd(x))\\ \ R \ \x\A\nat. \y\A\nat. \x,y\ \ {\a,b\\R. snd(b) = succ(snd(a))}" by (rule ballI, erule_tac x="x" in ballE, simp_all) lemma infer_snd : "c\ A\B \ snd(c) = k \ c=\fst(c),k\" by auto corollary DC_on_A_x_nat : assumes "(\x\A\nat. \y\A. \x,\y,succ(snd(x))\\ \ R)" "a\A" "M(A)" "M(R)" shows "\f \ nat\\<^bsup>M\<^esup>A. f`0 = a \ (\n \ nat. \\f`n,n\,\f`succ(n),succ(n)\\\R)" (is "\x\_.?P(x)") proof - let ?R'="{\a,b\\R. snd(b) = succ(snd(a))}" from assms(1) have "\x\A\nat. \y\A\nat. \x,y\ \ ?R'" using aux_DC_on_AxNat2 by simp moreover note \a\_\ \M(A)\ moreover from this have "M(A \ \)" by simp have "lam_replacement(M, \x. succ(snd(fst(x))))" using lam_replacement_fst lam_replacement_snd lam_replacement_hcomp lam_replacement_hcomp[of _ "\x. succ(snd(x))"] lam_replacement_succ by simp with \M(R)\ have "M(?R')" using separation_eq lam_replacement_fst lam_replacement_snd lam_replacement_succ lam_replacement_hcomp lam_replacement_identity unfolding split_def by simp ultimately obtain f where F:"f\nat\\<^bsup>M\<^esup> A\\" "f ` 0 = \a,0\" "\n\nat. \f ` n, f ` succ(n)\ \ ?R'" using pointed_DC[of "A\nat" ?R'] by blast with \M(A)\ have "M(f)" using function_space_rel_char by simp then have "M(\x\nat. fst(f`x))" (is "M(?f)") using lam_replacement_fst lam_replacement_hcomp lam_replacement_constant lam_replacement_identity lam_replacement_apply - by (rule_tac lam_replacement_iff_lam_closed[THEN iffD1, rule_format]) - auto + by (rule_tac lam_replacement_imp_lam_closed, auto) with F \M(A)\ have "?f\nat\\<^bsup>M\<^esup> A" "?f ` 0 = a" using function_space_rel_char by auto have 1:"n\ nat \ f`n= \?f`n, n\" for n proof(induct n set:nat) case 0 then show ?case using F by simp next case (succ x) with \M(A)\ have "\f`x, f`succ(x)\ \ ?R'" "f`x \ A\nat" "f`succ(x)\A\nat" using F function_space_rel_char[of \ "A\\"] by auto then have "snd(f`succ(x)) = succ(snd(f`x))" by simp with succ \f`x\_\ show ?case using infer_snd[OF \f`succ(_)\_\] by auto qed have "\\?f`n,n\,\?f`succ(n),succ(n)\\ \ R" if "n\nat" for n using that 1[of "succ(n)"] 1[OF \n\_\] F(3) by simp with \f`0=\a,0\\ show ?thesis using rev_bexI[OF \?f\_\] by simp qed lemma aux_sequence_DC : assumes "\x\A. \n\nat. \y\A. \x,y\ \ S`n" "R={\\x,n\,\y,m\\ \ (A\nat)\(A\nat). \x,y\\S`m }" shows "\ x\A\nat . \y\A. \x,\y,succ(snd(x))\\ \ R" using assms Pair_fst_snd_eq by auto lemma aux_sequence_DC2 : "\x\A. \n\nat. \y\A. \x,y\ \ S`n \ \x\A\nat. \y\A. \x,\y,succ(snd(x))\\ \ {\\x,n\,\y,m\\\(A\nat)\(A\nat). \x,y\\S`m }" by auto lemma sequence_DC: assumes "\x\A. \n\nat. \y\A. \x,y\ \ S`n" "M(A)" "M(S)" shows "\a\A. (\f \ nat\\<^bsup>M\<^esup> A. f`0 = a \ (\n \ nat. \f`n,f`succ(n)\\S`succ(n)))" proof - from \M(S)\ have "lam_replacement(M, \x. S ` snd(snd(x)))" using lam_replacement_snd lam_replacement_hcomp lam_replacement_hcomp[of _ "\x. S`snd(x)"] lam_replacement_apply by simp with assms have "M({x \ (A \ \) \ A \ \ . (\\\x,n\,y,m\. \x, y\ \ S ` m)(x)})" - using lam_replacement_fst lam_replacement_snd - lam_replacement_Pair[THEN [5] lam_replacement_hcomp2, - of "\x. fst(fst(x))" "\x. fst(snd(x))", THEN [2] separation_in, - of "\x. S ` snd(snd(x))"] lam_replacement_apply[of S] + using lam_replacement_fst lam_replacement_snd lam_replacement_apply[of S] + lam_replacement_product[of "\x. fst(fst(x))" "\x. fst(snd(x))", + THEN [2] separation_in, of "\x. S ` snd(snd(x))"] lam_replacement_hcomp unfolding split_def by simp with assms show ?thesis by (rule_tac ballI) (drule aux_sequence_DC2, drule DC_on_A_x_nat, auto) qed end \ \\<^locale>\M_library_DC\\ end \ No newline at end of file diff --git a/thys/Transitive_Models/Recursion_Thms.thy b/thys/Transitive_Models/Recursion_Thms.thy --- a/thys/Transitive_Models/Recursion_Thms.thy +++ b/thys/Transitive_Models/Recursion_Thms.thy @@ -1,364 +1,363 @@ section\Some enhanced theorems on recursion\ theory Recursion_Thms imports "ZF-Constructible.Datatype_absolute" begin hide_const (open) Order.pred \ \Removing arities from inherited simpset\ declare arity_And [simp del] arity_Or[simp del] arity_Implies[simp del] arity_Exists[simp del] arity_Iff[simp del] arity_subset_fm [simp del] arity_ordinal_fm[simp del] arity_transset_fm[simp del] text\We prove results concerning definitions by well-founded recursion on some relation \<^term>\R\ and its transitive closure \<^term>\R^*\\ - (* Restrict the relation r to the field A*A *) lemma fld_restrict_eq : "a \ A \ (r \ A\A)-``{a} = (r-``{a} \ A)" by(force) lemma fld_restrict_mono : "relation(r) \ A \ B \ r \ A\A \ r \ B\B" by(auto) lemma fld_restrict_dom : assumes "relation(r)" "domain(r) \ A" "range(r)\ A" shows "r\ A\A = r" proof (rule equalityI,blast,rule subsetI) { fix x assume xr: "x \ r" from xr assms have "\ a b . x = \a,b\" by (simp add: relation_def) then obtain a b where "\a,b\ \ r" "\a,b\ \ r\A\A" "x \ r\A\A" using assms xr by force then have "x\ r \ A\A" by simp } then show "x \ r \ x\ r\A\A" for x . qed definition tr_down :: "[i,i] \ i" where "tr_down(r,a) = (r^+)-``{a}" lemma tr_downD : "x \ tr_down(r,a) \ \x,a\ \ r^+" by (simp add: tr_down_def vimage_singleton_iff) lemma pred_down : "relation(r) \ r-``{a} \ tr_down(r,a)" by(simp add: tr_down_def vimage_mono r_subset_trancl) lemma tr_down_mono : "relation(r) \ x \ r-``{a} \ tr_down(r,x) \ tr_down(r,a)" by(rule subsetI,simp add:tr_down_def,auto dest: underD,force simp add: underI r_into_trancl trancl_trans) lemma rest_eq : assumes "relation(r)" and "r-``{a} \ B" and "a \ B" shows "r-``{a} = (r\B\B)-``{a}" proof (intro equalityI subsetI) fix x assume "x \ r-``{a}" then have "x \ B" using assms by (simp add: subsetD) from \x\ r-``{a}\ have "\x,a\ \ r" using underD by simp then show "x \ (r\B\B)-``{a}" using \x\B\ \a\B\ underI by simp next from assms show "x \ r -`` {a}" if "x \ (r \ B\B) -`` {a}" for x using vimage_mono that by auto qed lemma wfrec_restr_eq : "r' = r \ A\A \ wfrec[A](r,a,H) = wfrec(r',a,H)" by(simp add:wfrec_on_def) lemma wfrec_restr : assumes rr: "relation(r)" and wfr:"wf(r)" shows "a \ A \ tr_down(r,a) \ A \ wfrec(r,a,H) = wfrec[A](r,a,H)" proof (induct a arbitrary:A rule:wf_induct_raw[OF wfr] ) case (1 a) have wfRa : "wf[A](r)" using wf_subset wfr wf_on_def Int_lower1 by simp from pred_down rr have "r -`` {a} \ tr_down(r, a)" . with 1 have "r-``{a} \ A" by (force simp add: subset_trans) { fix x assume x_a : "x \ r-``{a}" with \r-``{a} \ A\ have "x \ A" .. from pred_down rr have b : "r -``{x} \ tr_down(r,x)" . then have "tr_down(r,x) \ tr_down(r,a)" using tr_down_mono x_a rr by simp with 1 have "tr_down(r,x) \ A" using subset_trans by force have "\x,a\ \ r" using x_a underD by simp with 1 \tr_down(r,x) \ A\ \x \ A\ have "wfrec(r,x,H) = wfrec[A](r,x,H)" by simp } then have "x\ r-``{a} \ wfrec(r,x,H) = wfrec[A](r,x,H)" for x . then have Eq1 :"(\ x \ r-``{a} . wfrec(r,x,H)) = (\ x \ r-``{a} . wfrec[A](r,x,H))" using lam_cong by simp from assms have "wfrec(r,a,H) = H(a,\ x \ r-``{a} . wfrec(r,x,H))" by (simp add:wfrec) also have "... = H(a,\ x \ r-``{a} . wfrec[A](r,x,H))" using assms Eq1 by simp also from 1 \r-``{a} \ A\ have "... = H(a,\ x \ (r\A\A)-``{a} . wfrec[A](r,x,H))" using assms rest_eq by simp also from \a\A\ have "... = H(a,\ x \ (r-``{a})\A . wfrec[A](r,x,H))" using fld_restrict_eq by simp also from \a\A\ \wf[A](r)\ have "... = wfrec[A](r,a,H)" using wfrec_on by simp finally show ?case . qed lemmas wfrec_tr_down = wfrec_restr[OF _ _ _ subset_refl] lemma wfrec_trans_restr : "relation(r) \ wf(r) \ trans(r) \ r-``{a}\A \ a \ A \ wfrec(r, a, H) = wfrec[A](r, a, H)" by(subgoal_tac "tr_down(r,a) \ A",auto simp add : wfrec_restr tr_down_def trancl_eq_r) lemma field_trancl : "field(r^+) = field(r)" by (blast intro: r_into_trancl dest!: trancl_type [THEN subsetD]) definition Rrel :: "[i\i\o,i] \ i" where "Rrel(R,A) \ {z\A\A. \x y. z = \x, y\ \ R(x,y)}" lemma RrelI : "x \ A \ y \ A \ R(x,y) \ \x,y\ \ Rrel(R,A)" unfolding Rrel_def by simp lemma Rrel_mem: "Rrel(mem,x) = Memrel(x)" unfolding Rrel_def Memrel_def .. lemma relation_Rrel: "relation(Rrel(R,d))" unfolding Rrel_def relation_def by simp lemma field_Rrel: "field(Rrel(R,d)) \ d" unfolding Rrel_def by auto lemma Rrel_mono : "A \ B \ Rrel(R,A) \ Rrel(R,B)" unfolding Rrel_def by blast lemma Rrel_restr_eq : "Rrel(R,A) \ B\B = Rrel(R,A\B)" unfolding Rrel_def by blast -(* now a consequence of the previous lemmas *) +\ \We obtain this lemmas as a consequence of the previous one; +alternatively it can be obtained using @{thm Ordinal.Memrel_type}\ lemma field_Memrel : "field(Memrel(A)) \ A" - (* unfolding field_def using Ordinal.Memrel_type by blast *) using Rrel_mem field_Rrel by blast lemma restrict_trancl_Rrel: assumes "R(w,y)" shows "restrict(f,Rrel(R,d)-``{y})`w = restrict(f,(Rrel(R,d)^+)-``{y})`w" proof (cases "y\d") let ?r="Rrel(R,d)" and ?s="(Rrel(R,d))^+" case True show ?thesis proof (cases "w\d") case True with \y\d\ assms have "\w,y\\?r" unfolding Rrel_def by blast then have "\w,y\\?s" using r_subset_trancl[of ?r] relation_Rrel[of R d] by blast with \\w,y\\?r\ have "w\?r-``{y}" "w\?s-``{y}" using vimage_singleton_iff by simp_all then show ?thesis by simp next case False then have "w\domain(restrict(f,?r-``{y}))" using subsetD[OF field_Rrel[of R d]] by auto moreover from \w\d\ have "w\domain(restrict(f,?s-``{y}))" using subsetD[OF field_Rrel[of R d], of w] field_trancl[of ?r] fieldI1[of w y ?s] by auto ultimately have "restrict(f,?r-``{y})`w = 0" "restrict(f,?s-``{y})`w = 0" unfolding apply_def by auto then show ?thesis by simp qed next let ?r="Rrel(R,d)" let ?s="?r^+" case False then have "?r-``{y}=0" unfolding Rrel_def by blast then have "w\?r-``{y}" by simp with \y\d\ assms have "y\field(?s)" using field_trancl subsetD[OF field_Rrel[of R d]] by force then have "w\?s-``{y}" using vimage_singleton_iff by blast with \w\?r-``{y}\ show ?thesis by simp qed lemma restrict_trans_eq: assumes "w \ y" shows "restrict(f,Memrel(eclose({x}))-``{y})`w = restrict(f,(Memrel(eclose({x}))^+)-``{y})`w" using assms restrict_trancl_Rrel[of mem ] Rrel_mem by (simp) lemma wf_eq_trancl: assumes "\ f y . H(y,restrict(f,R-``{y})) = H(y,restrict(f,R^+-``{y}))" shows "wfrec(R, x, H) = wfrec(R^+, x, H)" (is "wfrec(?r,_,_) = wfrec(?r',_,_)") proof - have "wfrec(R, x, H) = wftrec(?r^+, x, \y f. H(y, restrict(f,?r-``{y})))" unfolding wfrec_def .. also have " ... = wftrec(?r^+, x, \y f. H(y, restrict(f,(?r^+)-``{y})))" using assms by simp also have " ... = wfrec(?r^+, x, H)" unfolding wfrec_def using trancl_eq_r[OF relation_trancl trans_trancl] by simp finally show ?thesis . qed lemma transrec_equal_on_Ord: assumes "\x f . Ord(x) \ foo(x,f) = bar(x,f)" "Ord(\)" shows "transrec(\, foo) = transrec(\, bar)" proof - have "transrec(\,foo) = transrec(\,bar)" if "Ord(\)" for \ using that proof (induct rule:trans_induct) case (step \) have "transrec(\, foo) = foo(\, \x\\. transrec(x, foo))" using def_transrec[of "\x. transrec(x, foo)" foo] by blast also from assms and step have " \ = bar(\, \x\\. transrec(x, foo))" by simp also from step have " \ = bar(\, \x\\. transrec(x, bar))" by (auto) also have " \ = transrec(\, bar)" using def_transrec[of "\x. transrec(x, bar)" bar, symmetric] by blast finally show "transrec(\, foo) = transrec(\, bar)" . qed with assms show ?thesis by simp qed \ \Next theorem is very similar to @{thm transrec_equal_on_Ord}\ lemma (in M_eclose) transrec_equal_on_M: assumes "\x f . M(x) \ M(f) \ foo(x,f) = bar(x,f)" "\\. M(\) \ transrec_replacement(M,is_foo,\)" "relation2(M,is_foo,foo)" "strong_replacement(M, \x y. y = \x, transrec(x, foo)\)" "\x[M]. \g[M]. function(g) \ M(foo(x,g))" "M(\)" "Ord(\)" shows "transrec(\, foo) = transrec(\, bar)" proof - have "M(transrec(x, foo))" if "Ord(x)" and "M(x)" for x using that assms transrec_closed[of is_foo] by simp have "transrec(\,foo) = transrec(\,bar)" "M(transrec(\,foo))" if "Ord(\)" "M(\)" for \ using that proof (induct rule:trans_induct) case (step \) moreover assume "M(\)" moreover note \Ord(\)\ M(\) \ M(transrec(\, foo))\ ultimately show "M(transrec(\, foo))" by blast with step \M(\)\ \\x. Ord(x)\ M(x) \ M(transrec(x, foo))\ \strong_replacement(M, \x y. y = \x, transrec(x, foo)\)\ have "M(\x\\. transrec(x, foo))" using Ord_in_Ord transM[of _ \] by (rule_tac lam_closed) auto have "transrec(\, foo) = foo(\, \x\\. transrec(x, foo))" using def_transrec[of "\x. transrec(x, foo)" foo] by blast also from assms and \M(\x\\. transrec(x, foo))\ \M(\)\ have " \ = bar(\, \x\\. transrec(x, foo))" by simp also from step and \M(\)\ have " \ = bar(\, \x\\. transrec(x, bar))" using transM[of _ \] by (auto) also have " \ = transrec(\, bar)" using def_transrec[of "\x. transrec(x, bar)" bar, symmetric] by blast finally show "transrec(\, foo) = transrec(\, bar)" . qed with assms show ?thesis by simp qed lemma ordermap_restr_eq: assumes "well_ord(X,r)" shows "ordermap(X, r) = ordermap(X, r \ X\X)" proof - let ?A="\x . Order.pred(X, x, r)" let ?B="\x . Order.pred(X, x, r \ X \ X)" let ?F="\x f. f `` ?A(x)" let ?G="\x f. f `` ?B(x)" let ?P="\ z. z\X \ wfrec(r \ X \ X,z,\x f. f `` ?A(x)) = wfrec(r \ X \ X,z,\x f. f `` ?B(x))" have pred_eq: "Order.pred(X, x, r \ X \ X) = Order.pred(X, x, r)" if "x\X" for x unfolding Order.pred_def using that by auto from assms have wf_onX:"wf(r \ X \ X)" unfolding well_ord_def wf_on_def by simp { have "?P(z)" for z proof(induct rule:wf_induct[where P="?P",OF wf_onX]) case (1 x) { assume "x\X" from 1 have lam_eq: "(\w\(r \ X \ X) -`` {x}. wfrec(r \ X \ X, w, ?F)) = (\w\(r \ X \ X) -`` {x}. wfrec(r \ X \ X, w, ?G))" (is "?L=?R") proof - have "wfrec(r \ X \ X, w, ?F) = wfrec(r \ X \ X, w, ?G)" if "w\(r\X\X)-``{x}" for w using 1 that by auto then show ?thesis using lam_cong[OF refl] by simp qed then have "wfrec(r \ X \ X, x, ?F) = ?L `` ?A(x)" using wfrec[OF wf_onX,of x ?F] by simp also have "... = ?R `` ?B(x)" using lam_eq pred_eq[OF \x\_\] by simp also have "... = wfrec(r \ X \ X, x, ?G)" using wfrec[OF wf_onX,of x ?G] by simp finally have "wfrec(r \ X \ X, x, ?F) = wfrec(r \ X \ X, x, ?G)" by simp } then show ?case by simp qed } then show ?thesis unfolding ordermap_def wfrec_on_def using Int_ac by simp qed end diff --git a/thys/Transitive_Models/Relativization.ml b/thys/Transitive_Models/Relativization.ml new file mode 100644 --- /dev/null +++ b/thys/Transitive_Models/Relativization.ml @@ -0,0 +1,714 @@ +signature Relativization = + sig + structure Data: GENERIC_DATA + val Rel_add: attribute + val Rel_del: attribute + val add_rel_const : Database.mode -> term -> term -> Data.T -> Data.T + val add_constant : Database.mode -> string -> string -> Proof.context -> Proof.context + val rem_constant : (term -> Data.T -> Data.T) -> string -> Proof.context -> Proof.context + val db: Data.T + val init_db : Data.T -> theory -> theory + val get_db : Proof.context -> Data.T + val relativ_fm: bool -> bool -> term -> Data.T -> (term * (term * term)) list * Proof.context * term list * bool -> term -> term * ((term * (term * term)) list * term list * term list * Proof.context) + val relativ_tm: bool -> bool -> term option -> term -> Data.T -> (term * (term * term)) list * Proof.context -> term -> term * (term * (term * term)) list * Proof.context + val read_new_const : Proof.context -> string -> term + val relativ_tm_frm': bool -> bool -> term -> Data.T -> Proof.context -> term -> term option * term + val relativize_def: bool -> bool -> bool -> bstring -> string -> Position.T -> Proof.context -> Proof.context + val relativize_tm: bool -> bstring -> string -> Position.T -> Proof.context -> Proof.context + val rel_closed_goal : string -> Position.T -> Proof.context -> Proof.state + val iff_goal : string -> Position.T -> Proof.context -> Proof.state + val univalent_goal : string -> Position.T -> Proof.context -> Proof.state + end + +structure Relativization : Relativization = struct + +infix 6 &&& +val op &&& = Utils.&&& + +infix 6 *** +val op *** = Utils.*** + +infix 6 @@ +val op @@ = Utils.@@ + +infix 6 --- +val op --- = Utils.--- + +fun insert_abs2rel ((t, u), db) = ((t, u), Database.insert Database.abs2rel (t, t) db) + +fun insert_rel2is ((t, u), db) = Database.insert Database.rel2is (t, u) db + +(* relativization db of relation constructors *) +val db = [ (@{const relation}, @{const Relative.is_relation}) + , (@{const function}, @{const Relative.is_function}) + , (@{const mem}, @{const mem}) + , (@{const True}, @{const True}) + , (@{const False}, @{const False}) + , (@{const Memrel}, @{const membership}) + , (@{const trancl}, @{const tran_closure}) + , (@{const IFOL.eq(i)}, @{const IFOL.eq(i)}) + , (@{const Subset}, @{const Relative.subset}) + , (@{const quasinat}, @{const Relative.is_quasinat}) + , (@{const apply}, @{const Relative.fun_apply}) + , (@{const Upair}, @{const Relative.upair}) + ] + |> List.foldr (insert_rel2is o insert_abs2rel) Database.empty + |> Database.insert Database.abs2is (@{const Pi}, @{const is_funspace}) + +fun var_i v = Free (v, @{typ i}) +fun var_io v = Free (v, @{typ "i \ o"}) +val const_name = #1 o dest_Const + +val lookup_tm = AList.lookup (op aconv) +val update_tm = AList.update (op aconv) +val join_tm = AList.join (op aconv) (K #1) + +val conj_ = Utils.binop @{const "IFOL.conj"} + +(* generic data *) +structure Data = Generic_Data +( + type T = Database.db + val empty = Database.empty (* Should we initialize this outside this file? *) + val merge = Database.merge +); + +fun init_db db = Context.theory_map (Data.put db) + +fun get_db thy = Data.get (Context.Proof thy) + +val read_const = Proof_Context.read_const {proper = true, strict = true} +val read_new_const = Proof_Context.read_term_pattern + +fun add_rel_const mode c t = Database.insert mode (c, t) + +fun get_consts thm = + let val (c_rel, rhs) = Thm.concl_of thm |> Utils.dest_trueprop |> + Utils.dest_iff_tms |>> head_of + in case try Utils.dest_eq_tms rhs of + SOME tm => (c_rel, tm |> #2 |> head_of) + | NONE => (c_rel, rhs |> Utils.dest_mem_tms |> #2 |> head_of) + end + +fun add_rule thm rs = + let val (c_rel,c_abs) = get_consts thm + (* in (add_rel_const Database.rel2is c_abs c_rel o add_rel_const Database.abs2rel c_abs c_abs) rs *) + in (add_rel_const Database.abs2rel c_abs c_abs o add_rel_const Database.abs2is c_abs c_rel) rs +end + +fun get_mode is_functional relationalising = if relationalising then Database.rel2is else if is_functional then Database.abs2rel else Database.abs2is + +fun add_constant mode abs rel thy = + let + val c_abs = read_new_const thy abs + val c_rel = read_new_const thy rel + val db_map = Data.map (Database.insert mode (c_abs, c_rel)) + fun add_to_context ctxt' = Context.proof_map db_map ctxt' + fun add_to_theory ctxt' = Local_Theory.raw_theory (Context.theory_map db_map) ctxt' + in + Local_Theory.target (add_to_theory o add_to_context) thy + end + +fun rem_constant rem_op c thy = + let + val c = read_new_const thy c + val db_map = Data.map (rem_op c) + fun add_to_context ctxt' = Context.proof_map db_map ctxt' + fun add_to_theory ctxt' = Local_Theory.raw_theory (Context.theory_map db_map) ctxt' + in + Local_Theory.target (add_to_theory o add_to_context) thy + end + +val del_rel_const = Database.remove_abs + +fun del_rule thm = del_rel_const (thm |> get_consts |> #2) + +val Rel_add = + Thm.declaration_attribute (fn thm => fn context => + Data.map (add_rule (Thm.trim_context thm)) context); + +val Rel_del = + Thm.declaration_attribute (fn thm => fn context => + Data.map (del_rule (Thm.trim_context thm)) context); + +(* Conjunction of a list of terms *) +fun conjs [] = @{term IFOL.True} + | conjs (fs as _ :: _) = foldr1 (uncurry conj_) fs + +(* Produces a relativized existential quantification of the term t *) +fun rex p t (Free v) = @{const rex} $ p $ lambda (Free v) t + | rex _ t (Bound _) = t + | rex _ t tm = raise TERM ("rex shouldn't handle this.",[tm,t]) + +(* Constants that do not take the class predicate *) +val absolute_rels = [ @{const ZF_Base.mem} + , @{const IFOL.eq(i)} + , @{const Memrel} + , @{const True} + , @{const False} + ] + +(* Creates the relational term corresponding to a term of type i. If the last + argument is (SOME v) then that variable is not bound by an existential + quantifier. +*) +fun close_rel_tm pred tm tm_var rs = + let val news = filter (not o (fn x => is_Free x orelse is_Bound x) o #1) rs + val (vars, tms) = split_list (map #2 news) ||> (curry op @) (the_list tm) + val vars = case tm_var of + SOME w => filter (fn v => not (v = w)) vars + | NONE => vars + in fold (fn v => fn t => rex pred (incr_boundvars 1 t) v) vars (conjs tms) + end + +fun relativ_tms __ _ _ rs ctxt [] = ([], rs, ctxt) + | relativ_tms is_functional relationalising pred rel_db rs ctxt (u :: us) = + let val (w_u, rs_u, ctxt_u) = relativ_tm is_functional relationalising NONE pred rel_db (rs, ctxt) u + val (w_us, rs_us, ctxt_us) = relativ_tms is_functional relationalising pred rel_db rs_u ctxt_u us + in (w_u :: w_us, join_tm (rs_u , rs_us), ctxt_us) + end +and + (* The result of the relativization of a term is a triple consisting of + a. the relativized term (it can be a free or a bound variable but also a Collect) + b. a list of (term * (term, term)), taken as a map, which is used + to reuse relativization of different occurrences of the same term. The + first element is the original term, the second its relativized version, + and the last one is the predicate corresponding to it. + c. the resulting context of created variables. + *) + relativ_tm is_functional relationalising mv pred rel_db (rs,ctxt) tm = + let + (* relativization of a fully applied constant *) + fun mk_rel_const mv c (args, after) abs_args ctxt = + case Database.lookup (get_mode is_functional relationalising) c rel_db of + SOME p => + let + val args' = List.filter (not o member (op =) (Utils.frees p)) args + val (v, ctxt1) = + the_default + (Variable.variant_fixes [""] ctxt |>> var_i o hd) + (Utils.map_option (I &&& K ctxt) mv) + val args' = + (* FIXME: This special case for functional relativization of sigma should not be needed *) + if c = @{const Sigma} andalso is_functional + then + let + val t = hd args' + val t' = Abs ("uu_", @{typ "i"}, (hd o tl) args' |> incr_boundvars 1) + in + [t, t'] + end + else + args' + val arg_list = if after then abs_args @ args' else args' @ abs_args + val r_tm = + if is_functional + then list_comb (p, if p = c then arg_list else pred :: arg_list) + else list_comb (p, if (not o null) args' andalso hd args' = pred then arg_list @ [v] else pred :: arg_list @ [v]) + in + if is_functional + then (r_tm, r_tm, ctxt) + else (v, r_tm, ctxt1) + end + | NONE => raise TERM ("Constant " ^ const_name c ^ " is not present in the db." , nil) + (* relativization of a partially applied constant *) + fun relativ_app mv mctxt tm abs_args (Const c) (args, after) rs = + let + val (w_ts, rs_ts, ctxt_ts) = relativ_tms is_functional relationalising pred rel_db rs (the_default ctxt mctxt) args + val (w_tm, r_tm, ctxt_tm) = mk_rel_const mv (Const c) (w_ts, after) abs_args ctxt_ts + val rs_ts' = if is_functional then rs_ts else update_tm (tm, (w_tm, r_tm)) rs_ts + in + (w_tm, rs_ts', ctxt_tm) + end + | relativ_app _ _ _ _ t _ _ = + raise TERM ("Tried to relativize an application with a non-constant in head position",[t]) + + (* relativization of non dependent product and sum *) + fun relativ_app_no_dep mv tm c t t' rs = + if loose_bvar1 (t', 0) + then + raise TERM("A dependency was found when trying to relativize", [tm]) + else + relativ_app mv NONE tm [] c ([t, incr_boundvars ~1 t'], false) rs + + fun relativ_replace mv t body after ctxt' = + let + val (v, b) = Utils.dest_abs body |>> var_i ||> after + val (b', (rs', ctxt'')) = + relativ_fm is_functional relationalising pred rel_db (rs, ctxt', single v, false) b |>> incr_boundvars 1 ||> #1 &&& #4 + in + relativ_app mv (SOME ctxt'') tm [lambda v b'] @{const Replace} ([t], false) rs' + end + + fun get_abs_body (Abs body) = body + | get_abs_body t = raise TERM ("Term is not Abs", [t]) + + fun go _ (Var _) = raise TERM ("Var: Is this possible?",[]) + | go mv (@{const Replace} $ t $ Abs body) = relativ_replace mv t body I ctxt + (* It is easier to rewrite RepFun as Replace before relativizing, + since { f(x) . x \ t } = { y . x \ t, y = f(x) } *) + | go mv (@{const RepFun} $ t $ Abs body) = + let + val (y, ctxt') = Variable.variant_fixes [""] ctxt |>> var_i o hd + in + relativ_replace mv t body (lambda y o Utils.eq_ y o incr_boundvars 1) ctxt' + end + | go mv (@{const Collect} $ t $ pc) = + let + val (pc', (rs', ctxt')) = relativ_fm is_functional relationalising pred rel_db (rs,ctxt, [], false) pc ||> #1 &&& #4 + in + relativ_app mv (SOME ctxt') tm [pc'] @{const Collect} ([t], false) rs' + end + | go mv (@{const Least} $ pc) = + let + val (pc', (rs', ctxt')) = relativ_fm is_functional relationalising pred rel_db (rs,ctxt, [], false) pc ||> #1 &&& #4 + in + relativ_app mv (SOME ctxt') tm [pc'] @{const Least} ([], false) rs' + end + | go mv (@{const transrec} $ t $ Abs body) = + let + val (res, ctxt') = Variable.variant_fixes [if is_functional then "_aux" else ""] ctxt |>> var_i o hd + val (x, b') = Utils.dest_abs body |>> var_i + val (y, b) = get_abs_body b' |> Utils.dest_abs |>> var_i + val p = Utils.eq_ res b |> lambda res + val (p', (rs', ctxt'')) = relativ_fm is_functional relationalising pred rel_db (rs, ctxt', [x, y], true) p |>> incr_boundvars 3 ||> #1 &&& #4 + val p' = if is_functional then p' |> #2 o Utils.dest_eq_tms o #2 o Utils.dest_abs o get_abs_body else p' + in + relativ_app mv (SOME ctxt'') tm [p' |> lambda x o lambda y] @{const transrec} ([t], not is_functional) rs' + end + | go mv (tm as @{const Sigma} $ t $ Abs (_, _, t')) = + relativ_app_no_dep mv tm @{const Sigma} t t' rs + | go mv (tm as @{const Pi} $ t $ Abs (_, _, t')) = + relativ_app_no_dep mv tm @{const Pi} t t' rs + | go mv (tm as @{const bool_of_o} $ t) = + let + val (t', (rs', ctxt')) = relativ_fm is_functional relationalising pred rel_db (rs, ctxt, [], false) t ||> #1 &&& #4 + in + relativ_app mv (SOME ctxt') tm [t'] @{const bool_of_o} ([], false) rs' + end + | go mv (tm as @{const If} $ b $ t $ t') = + let + val (br, (rs', ctxt')) = relativ_fm is_functional relationalising pred rel_db (rs, ctxt, [], false) b ||> #1 &&& #4 + in + relativ_app mv (SOME ctxt') tm [br] @{const If} ([t,t'], true) rs' + end + | go mv (@{const The} $ pc) = + let + val (pc', (rs', ctxt')) = relativ_fm is_functional relationalising pred rel_db (rs,ctxt, [], false) pc ||> #1 &&& #4 + in + relativ_app mv (SOME ctxt') tm [pc'] @{const The} ([], false) rs' + end + | go mv (@{const recursor} $ t $ Abs body $ t') = + let + val (res, ctxt') = Variable.variant_fixes [if is_functional then "_aux" else ""] ctxt |>> var_i o hd + val (x, b') = Utils.dest_abs body |>> var_i + val (y, b) = get_abs_body b' |> Utils.dest_abs |>> var_i + val p = Utils.eq_ res b |> lambda res + val (p', (rs', ctxt'')) = relativ_fm is_functional relationalising pred rel_db (rs, ctxt', [x, y], true) p |>> incr_boundvars 3 ||> #1 &&& #4 + val p' = if is_functional then p' |> #2 o Utils.dest_eq_tms o #2 o Utils.dest_abs o get_abs_body else p' + val (tr, rs'', ctxt''') = relativ_tm is_functional relationalising NONE pred rel_db (rs', ctxt'') t + in + relativ_app mv (SOME ctxt''') tm [tr, p' |> lambda x o lambda y] @{const recursor} ([t'], true) rs'' + end + | go mv (@{const wfrec} $ t1 $ t2 $ Abs body) = + let + val (res, ctxt') = Variable.variant_fixes [if is_functional then "_aux" else ""] ctxt |>> var_i o hd + val (x, b') = Utils.dest_abs body |>> var_i + val (y, b) = get_abs_body b' |> Utils.dest_abs |>> var_i + val p = Utils.eq_ res b |> lambda res + val (p', (rs', ctxt'')) = relativ_fm is_functional relationalising pred rel_db (rs, ctxt', [x, y], true) p |>> incr_boundvars 3 ||> #1 &&& #4 + val p' = if is_functional then p' |> #2 o Utils.dest_eq_tms o #2 o Utils.dest_abs o get_abs_body else p' + in + relativ_app mv (SOME ctxt'') tm [p' |> lambda x o lambda y] @{const wfrec} ([t1,t2], not is_functional) rs' + end + | go mv (@{const wfrec_on} $ t1 $ t2 $ t3 $ Abs body) = + let + val (res, ctxt') = Variable.variant_fixes [if is_functional then "_aux" else ""] ctxt |>> var_i o hd + val (x, b') = Utils.dest_abs body |>> var_i + val (y, b) = get_abs_body b' |> Utils.dest_abs |>> var_i + val p = Utils.eq_ res b |> lambda res + val (p', (rs', ctxt'')) = relativ_fm is_functional relationalising pred rel_db (rs, ctxt', [x, y], true) p |>> incr_boundvars 3 ||> #1 &&& #4 + val p' = if is_functional then p' |> #2 o Utils.dest_eq_tms o #2 o Utils.dest_abs o get_abs_body else p' + in + relativ_app mv (SOME ctxt'') tm [p' |> lambda x o lambda y] @{const wfrec_on} ([t1,t2,t3], not is_functional) rs' + end + | go mv (@{const Lambda} $ t $ Abs body) = + let + val (res, ctxt') = Variable.variant_fixes [if is_functional then "_aux" else ""] ctxt |>> var_i o hd + val (x, b) = Utils.dest_abs body |>> var_i + val p = Utils.eq_ res b |> lambda res + val (p', (rs', ctxt'')) = relativ_fm is_functional relationalising pred rel_db (rs, ctxt', [x], true) p |>> incr_boundvars 2 ||> #1 &&& #4 + val p' = if is_functional then p' |> #2 o Utils.dest_eq_tms o #2 o Utils.dest_abs o get_abs_body else p' + val (tr, rs'', ctxt''') = relativ_tm is_functional relationalising NONE pred rel_db (rs', ctxt'') t + in + relativ_app mv (SOME ctxt''') tm [tr, p' |> lambda x] @{const Lambda} ([], true) rs'' + end + (* The following are the generic cases *) + | go mv (tm as Const _) = relativ_app mv NONE tm [] tm ([], false) rs + | go mv (tm as _ $ _) = (strip_comb tm ||> I &&& K false |> uncurry (relativ_app mv NONE tm [])) rs + | go _ tm = if is_functional then (tm, rs, ctxt) else (tm, update_tm (tm,(tm,tm)) rs, ctxt) + + (* we first check if the term has been already relativized as a variable *) + in case lookup_tm rs tm of + NONE => go mv tm + | SOME (w, _) => (w, rs, ctxt) + end +and + relativ_fm is_functional relationalising pred rel_db (rs, ctxt, vs, is_term) fm = + let + + (* relativization of a fully applied constant *) + fun relativ_app (ctxt, rs) c args = case Database.lookup (get_mode is_functional relationalising) c rel_db of + SOME p => + let (* flag indicates whether the relativized constant is absolute or not. *) + val flag = not (exists (curry op aconv c) absolute_rels orelse c = p) + val (args, rs_ts, ctxt') = relativ_tms is_functional relationalising pred rel_db rs ctxt args + (* TODO: Verify if next line takes care of locales' definitions *) + val args' = List.filter (not o member (op =) (Utils.frees p)) args + val args'' = if not (null args') andalso hd args' = pred then args' else pred :: args' + val tm = list_comb (p, if flag then args'' else args') + (* TODO: Verify if next line is necessary *) + val news = filter (not o (fn x => is_Free x orelse is_Bound x) o #1) rs_ts + val (vars, tms) = split_list (map #2 news) + (* val vars = filter (fn v => not (v = tm)) vars *) (* Verify if this line is necessary *) + in (tm, (rs_ts, vars, tms, ctxt')) + end + | NONE => raise TERM ("Constant " ^ const_name c ^ " is not present in the db." , nil) + + fun close_fm quantifier (f, (rs, vars, tms, ctxt)) = + let + fun contains_b0 t = loose_bvar1 (t, 0) + + fun contains_extra_var t = fold (fn v => fn acc => acc orelse fold_aterms (fn t => fn acc => t = v orelse acc) t false) vs false + + fun contains_b0_extra t = contains_b0 t orelse contains_extra_var t + + (* t1 $ v \ t2 iff v \ FV(t2) *) + fun chained_frees (_ $ v) t2 = member (op =) (Utils.frees t2) v + | chained_frees t _ = raise TERM ("Malformed term", [t]) + + val tms_to_close = filter contains_b0_extra tms |> Utils.reachable chained_frees tms + val tms_to_keep = map (incr_boundvars ~1) (tms --- tms_to_close) + val vars_to_close = inter (op =) (map (List.last o #2 o strip_comb) tms_to_close) vars + val vars_to_keep = vars --- vars_to_close + val new_rs = + rs + |> filter (fn (k, (v, rel)) => not (contains_b0_extra k orelse contains_b0_extra v orelse contains_b0_extra rel)) + |> map (fn (k, (v, rel)) => (incr_boundvars ~1 k, (incr_boundvars ~1 v, incr_boundvars ~1 rel))) + + val f' = + if not is_term andalso not quantifier andalso is_functional + then pred $ Bound 0 :: (map (curry (op $) pred) vs) @ [f] + else [f] + in + (fold (fn v => fn t => rex pred (incr_boundvars 1 t) v) vars_to_close (conjs (f' @ tms_to_close)), + (new_rs, vars_to_keep, tms_to_keep, ctxt)) + end + + (* Handling of bounded quantifiers. *) + fun bquant (ctxt, rs) quant conn dom pred = + let val (v,pred') = Utils.dest_abs pred |>> var_i + in + go (ctxt, rs, false) (quant $ (lambda v o incr_boundvars 1) (conn $ (@{const mem} $ v $ dom) $ pred')) + end + and + bind_go (ctxt, rs) const f f' = + let + val (r , (rs1, vars1, tms1, ctxt1)) = go (ctxt, rs, false) f + val (r', (rs2, vars2, tms2, ctxt2)) = go (ctxt1, rs1, false) f' + in + (const $ r $ r', (rs2, vars1 @@ vars2, tms1 @@ tms2, ctxt2)) + end + and + relativ_eq_var (ctxt, rs) v t = + let + val (_, rs', ctxt') = relativ_tm is_functional relationalising (SOME v) pred rel_db (rs, ctxt) t + val f = lookup_tm rs' t |> #2 o the + val rs'' = filter (not o (curry (op =) t) o #1) rs' + val news = filter (not o (fn x => is_Free x orelse is_Bound x) o #1) rs'' + val (vars, tms) = split_list (map #2 news) + in + (f, (rs'', vars, tms, ctxt')) + end + and + relativ_eq (ctxt, rs) t1 t2 = + if is_functional orelse ((is_Free t1 orelse is_Bound t1) andalso (is_Free t2 orelse is_Bound t2)) then + relativ_app (ctxt, rs) @{const IFOL.eq(i)} [t1, t2] + else if is_Free t1 orelse is_Bound t1 then + relativ_eq_var (ctxt, rs) t1 t2 + else if is_Free t2 orelse is_Bound t2 then + relativ_eq_var (ctxt, rs) t2 t1 + else + relativ_app (ctxt, rs) @{const IFOL.eq(i)} [t1, t2] + and + go (ctxt, rs, _ ) (@{const IFOL.conj} $ f $ f') = bind_go (ctxt, rs) @{const IFOL.conj} f f' + | go (ctxt, rs, _ ) (@{const IFOL.disj} $ f $ f') = bind_go (ctxt, rs) @{const IFOL.disj} f f' + | go (ctxt, rs, _ ) (@{const IFOL.Not} $ f) = go (ctxt, rs, false) f |>> ((curry op $) @{const IFOL.Not}) + | go (ctxt, rs, _ ) (@{const IFOL.iff} $ f $ f') = bind_go (ctxt, rs) @{const IFOL.iff} f f' + | go (ctxt, rs, _ ) (@{const IFOL.imp} $ f $ f') = bind_go (ctxt, rs) @{const IFOL.imp} f f' + | go (ctxt, rs, _ ) (@{const IFOL.All(i)} $ f) = go (ctxt, rs, true) f |>> ((curry op $) (@{const OrdQuant.rall} $ pred)) + | go (ctxt, rs, _ ) (@{const IFOL.Ex(i)} $ f) = go (ctxt, rs, true) f |>> ((curry op $) (@{const OrdQuant.rex} $ pred)) + | go (ctxt, rs, _ ) (@{const Bex} $ f $ Abs p) = bquant (ctxt, rs) @{const Ex(i)} @{const IFOL.conj} f p + | go (ctxt, rs, _ ) (@{const Ball} $ f $ Abs p) = bquant (ctxt, rs) @{const All(i)} @{const IFOL.imp} f p + | go (ctxt, rs, _ ) (@{const rall} $ _ $ p) = go (ctxt, rs, true) p |>> (curry op $) (@{const rall} $ pred) + | go (ctxt, rs, _ ) (@{const rex} $ _ $ p) = go (ctxt, rs, true) p |>> (curry op $) (@{const rex} $ pred) + | go (ctxt, rs, _ ) (@{const IFOL.eq(i)} $ t1 $ t2) = relativ_eq (ctxt, rs) t1 t2 + | go (ctxt, rs, _ ) (Const c) = relativ_app (ctxt, rs) (Const c) [] + | go (ctxt, rs, _ ) (tm as _ $ _) = strip_comb tm |> uncurry (relativ_app (ctxt, rs)) + | go (ctxt, rs, quantifier) (Abs (v, _, t)) = + let + val new_rs = map (fn (k, (v, rel)) => (incr_boundvars 1 k, (incr_boundvars 1 v, incr_boundvars 1 rel))) rs + in + go (ctxt, new_rs, false) t |> close_fm quantifier |>> lambda (var_i v) + end + | go _ t = raise TERM ("Relativization of formulas cannot handle this case.",[t]) + in + go (ctxt, rs, false) fm + end + + +fun relativ_tm_frm' is_functional relationalising cls_pred db ctxt tm = + let + fun get_bounds (l as Abs _) = op @@ (strip_abs l |>> map (op #1) ||> get_bounds) + | get_bounds (t as _$_) = strip_comb t |> op :: |> map get_bounds |> flat + | get_bounds _ = [] + + val ty = fastype_of tm + val initial_ctxt = fold Utils.add_to_context (get_bounds tm) ctxt + in + case ty of + @{typ i} => + let + val (w, rs, _) = relativ_tm is_functional relationalising NONE cls_pred db ([], initial_ctxt) tm + in + if is_functional + then (NONE, w) + else (SOME w, close_rel_tm cls_pred NONE (SOME w) rs) + end + | @{typ o} => + let + fun close_fm (f, (_, vars, tms, _)) = + fold (fn v => fn t => rex cls_pred (incr_boundvars 1 t) v) vars (conjs (f :: tms)) + in + (NONE, relativ_fm is_functional relationalising cls_pred db ([], initial_ctxt, [], false) tm |> close_fm) + end + | ty' => raise TYPE ("We can relativize only terms of types i and o", [ty'], [tm]) + end + +fun lname ctxt = Local_Theory.full_name ctxt o Binding.name + +fun destroy_first_lambdas (Abs (body as (_, ty, _))) = + Utils.dest_abs body ||> destroy_first_lambdas |> (#1 o #2) &&& ((fn v => Free (v, ty)) *** #2) ||> op :: + | destroy_first_lambdas t = (t, []) + +fun freeType (Free (_, ty)) = ty + | freeType t = raise TERM ("freeType", [t]) + +fun relativize_def is_external is_functional relationalising def_name thm_ref pos lthy = + let + val ctxt = lthy + val (vars,tm,ctxt1) = Utils.thm_concl_tm ctxt (thm_ref ^ "_def") + val db' = Data.get (Context.Proof lthy) + val (tm, lambdavars) = tm |> destroy_first_lambdas o #2 o Utils.dest_eq_tms' o Utils.dest_trueprop + val ctxt1 = fold Utils.add_to_context (map Utils.freeName lambdavars) ctxt1 + val (cls_pred, ctxt1, vars, lambdavars) = + if (not o null) vars andalso (#2 o #1 o hd) vars = @{typ "i \ o"} then + ((Thm.term_of o #2 o hd) vars, ctxt1, tl vars, lambdavars) + else if null vars andalso (not o null) lambdavars andalso (freeType o hd) lambdavars = @{typ "i \ o"} then + (hd lambdavars, ctxt1, vars, tl lambdavars) + else Variable.variant_fixes ["N"] ctxt1 |>> var_io o hd |> (fn (cls, ctxt) => (cls, ctxt, vars, lambdavars)) + val db' = db' |> Database.insert Database.abs2rel (cls_pred, cls_pred) + o Database.insert Database.rel2is (cls_pred, cls_pred) + val (v,t) = relativ_tm_frm' is_functional relationalising cls_pred db' ctxt1 tm + val t_vars = sort_strings (Term.add_free_names tm []) + val vs' = List.filter (#1 #> #1 #> #1 #> Ord_List.member String.compare t_vars) vars + val vs = cls_pred :: map (Thm.term_of o #2) vs' @ lambdavars @ the_list v + val at = List.foldr (uncurry lambda) t vs + val abs_const = read_const lthy (if is_external then thm_ref else lname lthy thm_ref) + fun new_const ctxt' = read_new_const ctxt' def_name + fun db_map ctxt' = + Data.map (add_rel_const (get_mode is_functional relationalising) abs_const (new_const ctxt')) + fun add_to_context ctxt' = Context.proof_map (db_map ctxt') ctxt' + fun add_to_theory ctxt' = Local_Theory.raw_theory (Context.theory_map (db_map ctxt')) ctxt' + in + lthy + |> Local_Theory.define ((Binding.name def_name, NoSyn), ((Binding.name (def_name ^ "_def"), []), at)) + |>> (#2 #> (fn (s,t) => (s,[t]))) + |> Utils.display "theorem" pos + |> Local_Theory.target (add_to_theory o add_to_context) + end + +fun relativize_tm is_functional def_name term pos lthy = + let + val ctxt = lthy + val (cls_pred, ctxt1) = Variable.variant_fixes ["N"] ctxt |>> var_io o hd + val tm = Syntax.read_term ctxt1 term + val db' = Data.get (Context.Proof lthy) + val db' = db' |> Database.insert Database.abs2rel (cls_pred, cls_pred) + o Database.insert Database.rel2is (cls_pred, cls_pred) + val vs' = Variable.add_frees ctxt1 tm [] + val ctxt2 = fold Utils.add_to_context (map #1 vs') ctxt1 + val (v,t) = relativ_tm_frm' is_functional false cls_pred db' ctxt2 tm + val vs = cls_pred :: map Free vs' @ the_list v + val at = List.foldr (uncurry lambda) t vs + in + lthy + |> Local_Theory.define ((Binding.name def_name, NoSyn), ((Binding.name (def_name ^ "_def"), []), at)) + |>> (#2 #> (fn (s,t) => (s,[t]))) + |> Utils.display "theorem" pos + end + +val op $` = curry ((op $) o swap) +infix $` + +fun is_free_i (Free (_, @{typ "i"})) = true + | is_free_i _ = false + +fun rel_closed_goal target pos lthy = + let + val (_, tm, _) = Utils.thm_concl_tm lthy (target ^ "_rel_def") + val (def, tm) = tm |> Utils.dest_eq_tms' + fun first_lambdas (Abs (body as (_, ty, _))) = + if ty = @{typ "i"} + then (op ::) (Utils.dest_abs body |>> Utils.var_i ||> first_lambdas) + else Utils.dest_abs body |> first_lambdas o #2 + | first_lambdas _ = [] + val (def, vars) = Term.strip_comb def ||> filter is_free_i + val vs = vars @ first_lambdas tm + val class = Free ("M", @{typ "i \ o"}) + val def = fold (op $`) (class :: vs) def + val hyps = map (fn v => class $ v |> Utils.tp) vs + val concl = class $ def + val goal = Logic.list_implies (hyps, Utils.tp concl) + val attribs = @{attributes [intro, simp]} + in + Proof.theorem NONE (fn thmss => Utils.display "theorem" pos + o Local_Theory.note ((Binding.name (target ^ "_rel_closed"), attribs), hd thmss)) + [[(goal, [])]] lthy + end + +fun iff_goal target pos lthy = + let + val (_, tm, ctxt') = Utils.thm_concl_tm lthy (target ^ "_rel_def") + val (_, is_def, ctxt) = Utils.thm_concl_tm ctxt' ("is_" ^ target ^ "_def") + val is_def = is_def |> Utils.dest_eq_tms' |> #1 |> Term.strip_comb |> #1 + val (def, tm) = tm |> Utils.dest_eq_tms' + fun first_lambdas (Abs (body as (_, ty, _))) = + if ty = @{typ "i"} + then (op ::) (Utils.dest_abs body |>> Utils.var_i ||> first_lambdas) + else Utils.dest_abs body |> first_lambdas o #2 + | first_lambdas _ = [] + val (def, vars) = Term.strip_comb def ||> filter is_free_i + val vs = vars @ first_lambdas tm + val class = Free ("M", @{typ "i \ o"}) + val def = fold (op $`) (class :: vs) def + val ty = fastype_of def + val res = if ty = @{typ "i"} + then Variable.variant_fixes ["res"] ctxt |> SOME o Utils.var_i o hd o #1 + else NONE + val is_def = fold (op $`) (class :: vs @ the_list res) is_def + val hyps = map (fn v => class $ v |> Utils.tp) (vs @ the_list res) + val concl = @{const "IFOL.iff"} $ is_def + $ (if ty = @{typ "i"} then (@{const IFOL.eq(i)} $ the res $ def) else def) + val goal = Logic.list_implies (hyps, Utils.tp concl) + in + Proof.theorem NONE (fn thmss => Utils.display "theorem" pos + o Local_Theory.note ((Binding.name ("is_" ^ target ^ "_iff"), []), hd thmss)) + [[(goal, [])]] lthy + end + +fun univalent_goal target pos lthy = + let + val (_, tm, ctxt) = Utils.thm_concl_tm lthy ("is_" ^ target ^ "_def") + val (def, tm) = tm |> Utils.dest_eq_tms' + fun first_lambdas (Abs (body as (_, ty, _))) = + if ty = @{typ "i"} + then (op ::) (Utils.dest_abs body |>> Utils.var_i ||> first_lambdas) + else Utils.dest_abs body |> first_lambdas o #2 + | first_lambdas _ = [] + val (def, vars) = Term.strip_comb def ||> filter is_free_i + val vs = vars @ first_lambdas tm + val n = length vs + val vs = List.take (vs, n - 2) + val class = Free ("M", @{typ "i \ o"}) + val def = fold (op $`) (class :: vs) def + val v = Variable.variant_fixes ["A"] ctxt |> Utils.var_i o hd o #1 + val hyps = map (fn v => class $ v |> Utils.tp) (v :: vs) + val concl = @{const "Relative.univalent"} $ class $ v $ def + val goal = Logic.list_implies (hyps, Utils.tp concl) + in + Proof.theorem NONE (fn thmss => Utils.display "theorem" pos + o Local_Theory.note ((Binding.name ("univalent_is_" ^ target), []), hd thmss)) + [[(goal, [])]] lthy + end + +end + +local + val full_mode_parser = + Scan.option (((Parse.$$$ "functional" |-- Parse.$$$ "relational") >> K Database.rel2is) + || (((Scan.option (Parse.$$$ "absolute")) |-- Parse.$$$ "functional") >> K Database.abs2rel) + || (((Scan.option (Parse.$$$ "absolute")) |-- Parse.$$$ "relational") >> K Database.abs2is)) + >> (fn mode => the_default Database.abs2is mode) + + val reldb_parser = + Parse.position (full_mode_parser -- (Parse.string -- Parse.string)); + + val singlemode_parser = (Parse.$$$ "absolute" >> K Database.remove_abs) + || (Parse.$$$ "functional" >> K Database.remove_rel) + || (Parse.$$$ "relational" >> K Database.remove_is) + + val reldb_rem_parser = Parse.position (singlemode_parser -- Parse.string) + + val mode_parser = + Scan.option ((Parse.$$$ "relational" >> K false) || (Parse.$$$ "functional" >> K true)) + >> (fn mode => if is_none mode then false else the mode) + + val relativize_parser = + Parse.position (mode_parser -- (Parse.string -- Parse.string) -- (Scan.optional (Parse.$$$ "external" >> K true) false)); + + val _ = + Outer_Syntax.local_theory \<^command_keyword>\reldb_add\ "ML setup for adding relativized/absolute pairs" + (reldb_parser >> (fn ((mode, (abs_term,rel_term)),_) => + Relativization.add_constant mode abs_term rel_term)) + + val _ = + Outer_Syntax.local_theory \<^command_keyword>\reldb_rem\ "ML setup for adding relativized/absolute pairs" + (reldb_rem_parser >> (uncurry Relativization.rem_constant o #1)) + + val _ = + Outer_Syntax.local_theory \<^command_keyword>\relativize\ "ML setup for relativizing definitions" + (relativize_parser >> (fn (((is_functional, (bndg,thm)), is_external),pos) => + Relativization.relativize_def is_external is_functional false thm bndg pos)) + + val _ = + Outer_Syntax.local_theory \<^command_keyword>\relativize_tm\ "ML setup for relativizing definitions" + (relativize_parser >> (fn (((is_functional, (bndg,term)), _),pos) => + Relativization.relativize_tm is_functional term bndg pos)) + + val _ = + Outer_Syntax.local_theory \<^command_keyword>\relationalize\ "ML setup for relativizing definitions" + (relativize_parser >> (fn (((is_functional, (bndg,thm)), is_external),pos) => + Relativization.relativize_def is_external is_functional true thm bndg pos)) + + val _ = + Outer_Syntax.local_theory_to_proof \<^command_keyword>\rel_closed\ "ML setup for rel_closed theorem" + (Parse.position (Parse.$$$ "for" |-- Parse.string) >> (fn (target,pos) => + Relativization.rel_closed_goal target pos)) + + val _ = + Outer_Syntax.local_theory_to_proof \<^command_keyword>\is_iff_rel\ "ML setup for rel_closed theorem" + (Parse.position (Parse.$$$ "for" |-- Parse.string) >> (fn (target,pos) => + Relativization.iff_goal target pos)) + + val _ = + Outer_Syntax.local_theory_to_proof \<^command_keyword>\univalent\ "ML setup for rel_closed theorem" + (Parse.position (Parse.$$$ "for" |-- Parse.string) >> (fn (target,pos) => + Relativization.univalent_goal target pos)) + +val _ = + Theory.setup + (Attrib.setup \<^binding>\Rel\ (Attrib.add_del Relativization.Rel_add Relativization.Rel_del) + "declaration of relativization rule") ; +in +end diff --git a/thys/Transitive_Models/Relativization.thy b/thys/Transitive_Models/Relativization.thy --- a/thys/Transitive_Models/Relativization.thy +++ b/thys/Transitive_Models/Relativization.thy @@ -1,856 +1,140 @@ section\Automatic relativization of terms and formulas\ text\Relativization of terms and formulas. Relativization of formulas shares relativized terms as far as possible; assuming that the witnesses for the relativized terms are always unique.\ theory Relativization imports "ZF-Constructible.Datatype_absolute" Higher_Order_Constructs keywords "relativize" :: thy_decl % "ML" and "relativize_tm" :: thy_decl % "ML" and "reldb_add" :: thy_decl % "ML" and "reldb_rem" :: thy_decl % "ML" and "relationalize" :: thy_decl % "ML" and "rel_closed" :: thy_goal_stmt % "ML" and "is_iff_rel" :: thy_goal_stmt % "ML" and "univalent" :: thy_goal_stmt % "ML" and "absolute" and "functional" and "relational" and "external" and "for" begin ML_file\Relativization_Database.ml\ ML\ structure Absoluteness = Named_Thms (val name = @{binding "absolut"} val description = "Theorems of absoulte terms and predicates.") \ setup\Absoluteness.setup\ lemmas relative_abs = M_trans.empty_abs M_trans.pair_abs M_trivial.cartprod_abs M_trans.union_abs M_trans.inter_abs M_trans.setdiff_abs M_trans.Union_abs M_trivial.cons_abs (*M_trans.upair_abs*) M_trivial.successor_abs M_trans.Collect_abs M_trans.Replace_abs M_trivial.lambda_abs2 M_trans.image_abs (*M_trans.powerset_abs*) M_trivial.nat_case_abs (* M_trans.transitive_set_abs M_trans.ordinal_abs M_trivial.limit_ordinal_abs M_trivial.successor_ordinal_abs M_trivial.finite_ordinal_abs *) M_trivial.omega_abs M_basic.sum_abs M_trivial.Inl_abs M_trivial.Inr_abs M_basic.converse_abs M_basic.vimage_abs M_trans.domain_abs M_trans.range_abs M_basic.field_abs (* M_basic.apply_abs *) (* M_trivial.typed_function_abs M_basic.injection_abs M_basic.surjection_abs M_basic.bijection_abs *) M_basic.composition_abs M_trans.restriction_abs M_trans.Inter_abs M_trivial.bool_of_o_abs M_trivial.not_abs M_trivial.and_abs M_trivial.or_abs M_trivial.Nil_abs M_trivial.Cons_abs (*M_trivial.quasilist_abs*) M_trivial.list_case_abs M_trivial.hd_abs M_trivial.tl_abs M_trivial.least_abs' M_eclose.transrec_abs M_trans.If_abs M_trans.The_abs M_eclose.recursor_abs M_trancl.trans_wfrec_abs M_trancl.trans_wfrec_on_abs lemmas datatype_abs = M_datatypes.list_N_abs M_datatypes.list_abs M_datatypes.formula_N_abs M_datatypes.formula_abs M_eclose.is_eclose_n_abs M_eclose.eclose_abs M_datatypes.length_abs M_datatypes.nth_abs M_trivial.Member_abs M_trivial.Equal_abs M_trivial.Nand_abs M_trivial.Forall_abs M_datatypes.depth_abs M_datatypes.formula_case_abs declare relative_abs[absolut] declare datatype_abs[absolut] -ML\ -signature Relativization = - sig - structure Data: GENERIC_DATA - val Rel_add: attribute - val Rel_del: attribute - val add_rel_const : Database.mode -> term -> term -> Data.T -> Data.T - val add_constant : Database.mode -> string -> string -> Proof.context -> Proof.context - val rem_constant : (term -> Data.T -> Data.T) -> string -> Proof.context -> Proof.context - val db: Data.T - val init_db : Data.T -> theory -> theory - val get_db : Proof.context -> Data.T - val relativ_fm: bool -> bool -> term -> Data.T -> (term * (term * term)) list * Proof.context * term list * bool -> term -> term * ((term * (term * term)) list * term list * term list * Proof.context) - val relativ_tm: bool -> bool -> term option -> term -> Data.T -> (term * (term * term)) list * Proof.context -> term -> term * (term * (term * term)) list * Proof.context - val read_new_const : Proof.context -> string -> term - val relativ_tm_frm': bool -> bool -> term -> Data.T -> Proof.context -> term -> term option * term - val relativize_def: bool -> bool -> bool -> bstring -> string -> Position.T -> Proof.context -> Proof.context - val relativize_tm: bool -> bstring -> string -> Position.T -> Proof.context -> Proof.context - val rel_closed_goal : string -> Position.T -> Proof.context -> Proof.state - val iff_goal : string -> Position.T -> Proof.context -> Proof.state - val univalent_goal : string -> Position.T -> Proof.context -> Proof.state - end - -structure Relativization : Relativization = struct - -infix 6 &&& -val op &&& = Utils.&&& - -infix 6 *** -val op *** = Utils.*** - -infix 6 @@ -val op @@ = Utils.@@ - -infix 6 --- -val op --- = Utils.--- - -fun insert_abs2rel ((t, u), db) = ((t, u), Database.insert Database.abs2rel (t, t) db) - -fun insert_rel2is ((t, u), db) = Database.insert Database.rel2is (t, u) db - -(* relativization db of relation constructors *) -val db = [ (@{const relation}, @{const Relative.is_relation}) - , (@{const function}, @{const Relative.is_function}) - , (@{const mem}, @{const mem}) - , (@{const True}, @{const True}) - , (@{const False}, @{const False}) - , (@{const Memrel}, @{const membership}) - , (@{const trancl}, @{const tran_closure}) - , (@{const IFOL.eq(i)}, @{const IFOL.eq(i)}) - , (@{const Subset}, @{const Relative.subset}) - , (@{const quasinat}, @{const Relative.is_quasinat}) - , (@{const apply}, @{const Relative.fun_apply}) - , (@{const Upair}, @{const Relative.upair}) - ] - |> List.foldr (insert_rel2is o insert_abs2rel) Database.empty - |> Database.insert Database.abs2is (@{const Pi}, @{const is_funspace}) - -fun var_i v = Free (v, @{typ i}) -fun var_io v = Free (v, @{typ "i \ o"}) -val const_name = #1 o dest_Const - -val lookup_tm = AList.lookup (op aconv) -val update_tm = AList.update (op aconv) -val join_tm = AList.join (op aconv) (K #1) - -val conj_ = Utils.binop @{const "IFOL.conj"} - -(* generic data *) -structure Data = Generic_Data -( - type T = Database.db - val empty = Database.empty (* Should we initialize this outside this file? *) - val merge = Database.merge -); - -fun init_db db = Context.theory_map (Data.put db) - -fun get_db thy = Data.get (Context.Proof thy) - -val read_const = Proof_Context.read_const {proper = true, strict = true} -val read_new_const = Proof_Context.read_term_pattern - -fun add_rel_const mode c t = Database.insert mode (c, t) - -fun get_consts thm = - let val (c_rel, rhs) = Thm.concl_of thm |> Utils.dest_trueprop |> - Utils.dest_iff_tms |>> head_of - in case try Utils.dest_eq_tms rhs of - SOME tm => (c_rel, tm |> #2 |> head_of) - | NONE => (c_rel, rhs |> Utils.dest_mem_tms |> #2 |> head_of) - end - -fun add_rule thm rs = - let val (c_rel,c_abs) = get_consts thm - (* in (add_rel_const Database.rel2is c_abs c_rel o add_rel_const Database.abs2rel c_abs c_abs) rs *) - in (add_rel_const Database.abs2rel c_abs c_abs o add_rel_const Database.abs2is c_abs c_rel) rs -end - -fun get_mode is_functional relationalising = if relationalising then Database.rel2is else if is_functional then Database.abs2rel else Database.abs2is - -fun add_constant mode abs rel thy = - let - val c_abs = read_new_const thy abs - val c_rel = read_new_const thy rel - val db_map = Data.map (Database.insert mode (c_abs, c_rel)) - fun add_to_context ctxt' = Context.proof_map db_map ctxt' - fun add_to_theory ctxt' = Local_Theory.raw_theory (Context.theory_map db_map) ctxt' - in - Local_Theory.target (add_to_theory o add_to_context) thy - end - -fun rem_constant rem_op c thy = - let - val c = read_new_const thy c - val db_map = Data.map (rem_op c) - fun add_to_context ctxt' = Context.proof_map db_map ctxt' - fun add_to_theory ctxt' = Local_Theory.raw_theory (Context.theory_map db_map) ctxt' - in - Local_Theory.target (add_to_theory o add_to_context) thy - end - -val del_rel_const = Database.remove_abs - -fun del_rule thm = del_rel_const (thm |> get_consts |> #2) - -val Rel_add = - Thm.declaration_attribute (fn thm => fn context => - Data.map (add_rule (Thm.trim_context thm)) context); - -val Rel_del = - Thm.declaration_attribute (fn thm => fn context => - Data.map (del_rule (Thm.trim_context thm)) context); - -(* Conjunction of a list of terms *) -fun conjs [] = @{term IFOL.True} - | conjs (fs as _ :: _) = foldr1 (uncurry conj_) fs - -(* Produces a relativized existential quantification of the term t *) -fun rex p t (Free v) = @{const rex} $ p $ lambda (Free v) t - | rex _ t (Bound _) = t - | rex _ t tm = raise TERM ("rex shouldn't handle this.",[tm,t]) - -(* Constants that do not take the class predicate *) -val absolute_rels = [ @{const ZF_Base.mem} - , @{const IFOL.eq(i)} - , @{const Memrel} - , @{const True} - , @{const False} - ] - -(* Creates the relational term corresponding to a term of type i. If the last - argument is (SOME v) then that variable is not bound by an existential - quantifier. -*) -fun close_rel_tm pred tm tm_var rs = - let val news = filter (not o (fn x => is_Free x orelse is_Bound x) o #1) rs - val (vars, tms) = split_list (map #2 news) ||> (curry op @) (the_list tm) - val vars = case tm_var of - SOME w => filter (fn v => not (v = w)) vars - | NONE => vars - in fold (fn v => fn t => rex pred (incr_boundvars 1 t) v) vars (conjs tms) - end - -fun relativ_tms __ _ _ rs ctxt [] = ([], rs, ctxt) - | relativ_tms is_functional relationalising pred rel_db rs ctxt (u :: us) = - let val (w_u, rs_u, ctxt_u) = relativ_tm is_functional relationalising NONE pred rel_db (rs, ctxt) u - val (w_us, rs_us, ctxt_us) = relativ_tms is_functional relationalising pred rel_db rs_u ctxt_u us - in (w_u :: w_us, join_tm (rs_u , rs_us), ctxt_us) - end -and - (* The result of the relativization of a term is a triple consisting of - a. the relativized term (it can be a free or a bound variable but also a Collect) - b. a list of (term * (term, term)), taken as a map, which is used - to reuse relativization of different occurrences of the same term. The - first element is the original term, the second its relativized version, - and the last one is the predicate corresponding to it. - c. the resulting context of created variables. - *) - relativ_tm is_functional relationalising mv pred rel_db (rs,ctxt) tm = - let - (* relativization of a fully applied constant *) - fun mk_rel_const mv c (args, after) abs_args ctxt = - case Database.lookup (get_mode is_functional relationalising) c rel_db of - SOME p => - let - val args' = List.filter (not o member (op =) (Utils.frees p)) args - val (v, ctxt1) = - the_default - (Variable.variant_fixes [""] ctxt |>> var_i o hd) - (Utils.map_option (I &&& K ctxt) mv) - val args' = - (* FIXME: This special case for functional relativization of sigma should not be needed *) - if c = @{const Sigma} andalso is_functional - then - let - val t = hd args' - val t' = Abs ("uu_", @{typ "i"}, (hd o tl) args' |> incr_boundvars 1) - in - [t, t'] - end - else - args' - val arg_list = if after then abs_args @ args' else args' @ abs_args - val r_tm = - if is_functional - then list_comb (p, if p = c then arg_list else pred :: arg_list) - else list_comb (p, if (not o null) args' andalso hd args' = pred then arg_list @ [v] else pred :: arg_list @ [v]) - in - if is_functional - then (r_tm, r_tm, ctxt) - else (v, r_tm, ctxt1) - end - | NONE => raise TERM ("Constant " ^ const_name c ^ " is not present in the db." , nil) - (* relativization of a partially applied constant *) - fun relativ_app mv mctxt tm abs_args (Const c) (args, after) rs = - let - val (w_ts, rs_ts, ctxt_ts) = relativ_tms is_functional relationalising pred rel_db rs (the_default ctxt mctxt) args - val (w_tm, r_tm, ctxt_tm) = mk_rel_const mv (Const c) (w_ts, after) abs_args ctxt_ts - val rs_ts' = if is_functional then rs_ts else update_tm (tm, (w_tm, r_tm)) rs_ts - in - (w_tm, rs_ts', ctxt_tm) - end - | relativ_app _ _ _ _ t _ _ = - raise TERM ("Tried to relativize an application with a non-constant in head position",[t]) - - (* relativization of non dependent product and sum *) - fun relativ_app_no_dep mv tm c t t' rs = - if loose_bvar1 (t', 0) - then - raise TERM("A dependency was found when trying to relativize", [tm]) - else - relativ_app mv NONE tm [] c ([t, incr_boundvars ~1 t'], false) rs - - fun relativ_replace mv t body after ctxt' = - let - val (v, b) = Utils.dest_abs body |>> var_i ||> after - val (b', (rs', ctxt'')) = - relativ_fm is_functional relationalising pred rel_db (rs, ctxt', single v, false) b |>> incr_boundvars 1 ||> #1 &&& #4 - in - relativ_app mv (SOME ctxt'') tm [lambda v b'] @{const Replace} ([t], false) rs' - end - - fun get_abs_body (Abs body) = body - | get_abs_body t = raise TERM ("Term is not Abs", [t]) - - fun go _ (Var _) = raise TERM ("Var: Is this possible?",[]) - | go mv (@{const Replace} $ t $ Abs body) = relativ_replace mv t body I ctxt - (* It is easier to rewrite RepFun as Replace before relativizing, - since { f(x) . x \ t } = { y . x \ t, y = f(x) } *) - | go mv (@{const RepFun} $ t $ Abs body) = - let - val (y, ctxt') = Variable.variant_fixes [""] ctxt |>> var_i o hd - in - relativ_replace mv t body (lambda y o Utils.eq_ y o incr_boundvars 1) ctxt' - end - | go mv (@{const Collect} $ t $ pc) = - let - val (pc', (rs', ctxt')) = relativ_fm is_functional relationalising pred rel_db (rs,ctxt, [], false) pc ||> #1 &&& #4 - in - relativ_app mv (SOME ctxt') tm [pc'] @{const Collect} ([t], false) rs' - end - | go mv (@{const Least} $ pc) = - let - val (pc', (rs', ctxt')) = relativ_fm is_functional relationalising pred rel_db (rs,ctxt, [], false) pc ||> #1 &&& #4 - in - relativ_app mv (SOME ctxt') tm [pc'] @{const Least} ([], false) rs' - end - | go mv (@{const transrec} $ t $ Abs body) = - let - val (res, ctxt') = Variable.variant_fixes [if is_functional then "_aux" else ""] ctxt |>> var_i o hd - val (x, b') = Utils.dest_abs body |>> var_i - val (y, b) = get_abs_body b' |> Utils.dest_abs |>> var_i - val p = Utils.eq_ res b |> lambda res - val (p', (rs', ctxt'')) = relativ_fm is_functional relationalising pred rel_db (rs, ctxt', [x, y], true) p |>> incr_boundvars 3 ||> #1 &&& #4 - val p' = if is_functional then p' |> #2 o Utils.dest_eq_tms o #2 o Utils.dest_abs o get_abs_body else p' - in - relativ_app mv (SOME ctxt'') tm [p' |> lambda x o lambda y] @{const transrec} ([t], not is_functional) rs' - end - | go mv (tm as @{const Sigma} $ t $ Abs (_, _, t')) = - relativ_app_no_dep mv tm @{const Sigma} t t' rs - | go mv (tm as @{const Pi} $ t $ Abs (_, _, t')) = - relativ_app_no_dep mv tm @{const Pi} t t' rs - | go mv (tm as @{const bool_of_o} $ t) = - let - val (t', (rs', ctxt')) = relativ_fm is_functional relationalising pred rel_db (rs, ctxt, [], false) t ||> #1 &&& #4 - in - relativ_app mv (SOME ctxt') tm [t'] @{const bool_of_o} ([], false) rs' - end - | go mv (tm as @{const If} $ b $ t $ t') = - let - val (br, (rs', ctxt')) = relativ_fm is_functional relationalising pred rel_db (rs, ctxt, [], false) b ||> #1 &&& #4 - in - relativ_app mv (SOME ctxt') tm [br] @{const If} ([t,t'], true) rs' - end - | go mv (@{const The} $ pc) = - let - val (pc', (rs', ctxt')) = relativ_fm is_functional relationalising pred rel_db (rs,ctxt, [], false) pc ||> #1 &&& #4 - in - relativ_app mv (SOME ctxt') tm [pc'] @{const The} ([], false) rs' - end - | go mv (@{const recursor} $ t $ Abs body $ t') = - let - val (res, ctxt') = Variable.variant_fixes [if is_functional then "_aux" else ""] ctxt |>> var_i o hd - val (x, b') = Utils.dest_abs body |>> var_i - val (y, b) = get_abs_body b' |> Utils.dest_abs |>> var_i - val p = Utils.eq_ res b |> lambda res - val (p', (rs', ctxt'')) = relativ_fm is_functional relationalising pred rel_db (rs, ctxt', [x, y], true) p |>> incr_boundvars 3 ||> #1 &&& #4 - val p' = if is_functional then p' |> #2 o Utils.dest_eq_tms o #2 o Utils.dest_abs o get_abs_body else p' - val (tr, rs'', ctxt''') = relativ_tm is_functional relationalising NONE pred rel_db (rs', ctxt'') t - in - relativ_app mv (SOME ctxt''') tm [tr, p' |> lambda x o lambda y] @{const recursor} ([t'], true) rs'' - end - | go mv (@{const wfrec} $ t1 $ t2 $ Abs body) = - let - val (res, ctxt') = Variable.variant_fixes [if is_functional then "_aux" else ""] ctxt |>> var_i o hd - val (x, b') = Utils.dest_abs body |>> var_i - val (y, b) = get_abs_body b' |> Utils.dest_abs |>> var_i - val p = Utils.eq_ res b |> lambda res - val (p', (rs', ctxt'')) = relativ_fm is_functional relationalising pred rel_db (rs, ctxt', [x, y], true) p |>> incr_boundvars 3 ||> #1 &&& #4 - val p' = if is_functional then p' |> #2 o Utils.dest_eq_tms o #2 o Utils.dest_abs o get_abs_body else p' - in - relativ_app mv (SOME ctxt'') tm [p' |> lambda x o lambda y] @{const wfrec} ([t1,t2], not is_functional) rs' - end - | go mv (@{const wfrec_on} $ t1 $ t2 $ t3 $ Abs body) = - let - val (res, ctxt') = Variable.variant_fixes [if is_functional then "_aux" else ""] ctxt |>> var_i o hd - val (x, b') = Utils.dest_abs body |>> var_i - val (y, b) = get_abs_body b' |> Utils.dest_abs |>> var_i - val p = Utils.eq_ res b |> lambda res - val (p', (rs', ctxt'')) = relativ_fm is_functional relationalising pred rel_db (rs, ctxt', [x, y], true) p |>> incr_boundvars 3 ||> #1 &&& #4 - val p' = if is_functional then p' |> #2 o Utils.dest_eq_tms o #2 o Utils.dest_abs o get_abs_body else p' - in - relativ_app mv (SOME ctxt'') tm [p' |> lambda x o lambda y] @{const wfrec_on} ([t1,t2,t3], not is_functional) rs' - end - | go mv (@{const Lambda} $ t $ Abs body) = - let - val (res, ctxt') = Variable.variant_fixes [if is_functional then "_aux" else ""] ctxt |>> var_i o hd - val (x, b) = Utils.dest_abs body |>> var_i - val p = Utils.eq_ res b |> lambda res - val (p', (rs', ctxt'')) = relativ_fm is_functional relationalising pred rel_db (rs, ctxt', [x], true) p |>> incr_boundvars 2 ||> #1 &&& #4 - val p' = if is_functional then p' |> #2 o Utils.dest_eq_tms o #2 o Utils.dest_abs o get_abs_body else p' - val (tr, rs'', ctxt''') = relativ_tm is_functional relationalising NONE pred rel_db (rs', ctxt'') t - in - relativ_app mv (SOME ctxt''') tm [tr, p' |> lambda x] @{const Lambda} ([], true) rs'' - end - (* The following are the generic cases *) - | go mv (tm as Const _) = relativ_app mv NONE tm [] tm ([], false) rs - | go mv (tm as _ $ _) = (strip_comb tm ||> I &&& K false |> uncurry (relativ_app mv NONE tm [])) rs - | go _ tm = if is_functional then (tm, rs, ctxt) else (tm, update_tm (tm,(tm,tm)) rs, ctxt) +ML_file\Relativization.ml\ - (* we first check if the term has been already relativized as a variable *) - in case lookup_tm rs tm of - NONE => go mv tm - | SOME (w, _) => (w, rs, ctxt) - end -and - relativ_fm is_functional relationalising pred rel_db (rs, ctxt, vs, is_term) fm = - let - - (* relativization of a fully applied constant *) - fun relativ_app (ctxt, rs) c args = case Database.lookup (get_mode is_functional relationalising) c rel_db of - SOME p => - let (* flag indicates whether the relativized constant is absolute or not. *) - val flag = not (exists (curry op aconv c) absolute_rels orelse c = p) - val (args, rs_ts, ctxt') = relativ_tms is_functional relationalising pred rel_db rs ctxt args - (* TODO: Verify if next line takes care of locales' definitions *) - val args' = List.filter (not o member (op =) (Utils.frees p)) args - val args'' = if not (null args') andalso hd args' = pred then args' else pred :: args' - val tm = list_comb (p, if flag then args'' else args') - (* TODO: Verify if next line is necessary *) - val news = filter (not o (fn x => is_Free x orelse is_Bound x) o #1) rs_ts - val (vars, tms) = split_list (map #2 news) - (* val vars = filter (fn v => not (v = tm)) vars *) (* Verify if this line is necessary *) - in (tm, (rs_ts, vars, tms, ctxt')) - end - | NONE => raise TERM ("Constant " ^ const_name c ^ " is not present in the db." , nil) - - fun close_fm quantifier (f, (rs, vars, tms, ctxt)) = - let - fun contains_b0 t = loose_bvar1 (t, 0) - - fun contains_extra_var t = fold (fn v => fn acc => acc orelse fold_aterms (fn t => fn acc => t = v orelse acc) t false) vs false - - fun contains_b0_extra t = contains_b0 t orelse contains_extra_var t - - (* t1 $ v \ t2 iff v \ FV(t2) *) - fun chained_frees (_ $ v) t2 = member (op =) (Utils.frees t2) v - | chained_frees t _ = raise TERM ("Malformed term", [t]) - - val tms_to_close = filter contains_b0_extra tms |> Utils.reachable chained_frees tms - val tms_to_keep = map (incr_boundvars ~1) (tms --- tms_to_close) - val vars_to_close = inter (op =) (map (List.last o #2 o strip_comb) tms_to_close) vars - val vars_to_keep = vars --- vars_to_close - val new_rs = - rs - |> filter (fn (k, (v, rel)) => not (contains_b0_extra k orelse contains_b0_extra v orelse contains_b0_extra rel)) - |> map (fn (k, (v, rel)) => (incr_boundvars ~1 k, (incr_boundvars ~1 v, incr_boundvars ~1 rel))) - - val f' = - if not is_term andalso not quantifier andalso is_functional - then pred $ Bound 0 :: (map (curry (op $) pred) vs) @ [f] - else [f] - in - (fold (fn v => fn t => rex pred (incr_boundvars 1 t) v) vars_to_close (conjs (f' @ tms_to_close)), - (new_rs, vars_to_keep, tms_to_keep, ctxt)) - end - - (* Handling of bounded quantifiers. *) - fun bquant (ctxt, rs) quant conn dom pred = - let val (v,pred') = Utils.dest_abs pred |>> var_i - in - go (ctxt, rs, false) (quant $ (lambda v o incr_boundvars 1) (conn $ (@{const mem} $ v $ dom) $ pred')) - end - and - bind_go (ctxt, rs) const f f' = - let - val (r , (rs1, vars1, tms1, ctxt1)) = go (ctxt, rs, false) f - val (r', (rs2, vars2, tms2, ctxt2)) = go (ctxt1, rs1, false) f' - in - (const $ r $ r', (rs2, vars1 @@ vars2, tms1 @@ tms2, ctxt2)) - end - and - relativ_eq_var (ctxt, rs) v t = - let - val (_, rs', ctxt') = relativ_tm is_functional relationalising (SOME v) pred rel_db (rs, ctxt) t - val f = lookup_tm rs' t |> #2 o the - val rs'' = filter (not o (curry (op =) t) o #1) rs' - val news = filter (not o (fn x => is_Free x orelse is_Bound x) o #1) rs'' - val (vars, tms) = split_list (map #2 news) - in - (f, (rs'', vars, tms, ctxt')) - end - and - relativ_eq (ctxt, rs) t1 t2 = - if is_functional orelse ((is_Free t1 orelse is_Bound t1) andalso (is_Free t2 orelse is_Bound t2)) then - relativ_app (ctxt, rs) @{const IFOL.eq(i)} [t1, t2] - else if is_Free t1 orelse is_Bound t1 then - relativ_eq_var (ctxt, rs) t1 t2 - else if is_Free t2 orelse is_Bound t2 then - relativ_eq_var (ctxt, rs) t2 t1 - else - relativ_app (ctxt, rs) @{const IFOL.eq(i)} [t1, t2] - and - go (ctxt, rs, _ ) (@{const IFOL.conj} $ f $ f') = bind_go (ctxt, rs) @{const IFOL.conj} f f' - | go (ctxt, rs, _ ) (@{const IFOL.disj} $ f $ f') = bind_go (ctxt, rs) @{const IFOL.disj} f f' - | go (ctxt, rs, _ ) (@{const IFOL.Not} $ f) = go (ctxt, rs, false) f |>> ((curry op $) @{const IFOL.Not}) - | go (ctxt, rs, _ ) (@{const IFOL.iff} $ f $ f') = bind_go (ctxt, rs) @{const IFOL.iff} f f' - | go (ctxt, rs, _ ) (@{const IFOL.imp} $ f $ f') = bind_go (ctxt, rs) @{const IFOL.imp} f f' - | go (ctxt, rs, _ ) (@{const IFOL.All(i)} $ f) = go (ctxt, rs, true) f |>> ((curry op $) (@{const OrdQuant.rall} $ pred)) - | go (ctxt, rs, _ ) (@{const IFOL.Ex(i)} $ f) = go (ctxt, rs, true) f |>> ((curry op $) (@{const OrdQuant.rex} $ pred)) - | go (ctxt, rs, _ ) (@{const Bex} $ f $ Abs p) = bquant (ctxt, rs) @{const Ex(i)} @{const IFOL.conj} f p - | go (ctxt, rs, _ ) (@{const Ball} $ f $ Abs p) = bquant (ctxt, rs) @{const All(i)} @{const IFOL.imp} f p - | go (ctxt, rs, _ ) (@{const rall} $ _ $ p) = go (ctxt, rs, true) p |>> (curry op $) (@{const rall} $ pred) - | go (ctxt, rs, _ ) (@{const rex} $ _ $ p) = go (ctxt, rs, true) p |>> (curry op $) (@{const rex} $ pred) - | go (ctxt, rs, _ ) (@{const IFOL.eq(i)} $ t1 $ t2) = relativ_eq (ctxt, rs) t1 t2 - | go (ctxt, rs, _ ) (Const c) = relativ_app (ctxt, rs) (Const c) [] - | go (ctxt, rs, _ ) (tm as _ $ _) = strip_comb tm |> uncurry (relativ_app (ctxt, rs)) - | go (ctxt, rs, quantifier) (Abs (v, _, t)) = - let - val new_rs = map (fn (k, (v, rel)) => (incr_boundvars 1 k, (incr_boundvars 1 v, incr_boundvars 1 rel))) rs - in - go (ctxt, new_rs, false) t |> close_fm quantifier |>> lambda (var_i v) - end - | go _ t = raise TERM ("Relativization of formulas cannot handle this case.",[t]) - in - go (ctxt, rs, false) fm - end - - -fun relativ_tm_frm' is_functional relationalising cls_pred db ctxt tm = - let - fun get_bounds (l as Abs _) = op @@ (strip_abs l |>> map (op #1) ||> get_bounds) - | get_bounds (t as _$_) = strip_comb t |> op :: |> map get_bounds |> flat - | get_bounds _ = [] - - val ty = fastype_of tm - val initial_ctxt = fold Utils.add_to_context (get_bounds tm) ctxt - in - case ty of - @{typ i} => - let - val (w, rs, _) = relativ_tm is_functional relationalising NONE cls_pred db ([], initial_ctxt) tm - in - if is_functional - then (NONE, w) - else (SOME w, close_rel_tm cls_pred NONE (SOME w) rs) - end - | @{typ o} => - let - fun close_fm (f, (_, vars, tms, _)) = - fold (fn v => fn t => rex cls_pred (incr_boundvars 1 t) v) vars (conjs (f :: tms)) - in - (NONE, relativ_fm is_functional relationalising cls_pred db ([], initial_ctxt, [], false) tm |> close_fm) - end - | ty' => raise TYPE ("We can relativize only terms of types i and o", [ty'], [tm]) - end - -fun lname ctxt = Local_Theory.full_name ctxt o Binding.name - -fun destroy_first_lambdas (Abs (body as (_, ty, _))) = - Utils.dest_abs body ||> destroy_first_lambdas |> (#1 o #2) &&& ((fn v => Free (v, ty)) *** #2) ||> op :: - | destroy_first_lambdas t = (t, []) - -fun freeType (Free (_, ty)) = ty - | freeType t = raise TERM ("freeType", [t]) - -fun relativize_def is_external is_functional relationalising def_name thm_ref pos lthy = - let - val ctxt = lthy - val (vars,tm,ctxt1) = Utils.thm_concl_tm ctxt (thm_ref ^ "_def") - val db' = Data.get (Context.Proof lthy) - val (tm, lambdavars) = tm |> destroy_first_lambdas o #2 o Utils.dest_eq_tms' o Utils.dest_trueprop - val ctxt1 = fold Utils.add_to_context (map Utils.freeName lambdavars) ctxt1 - val (cls_pred, ctxt1, vars, lambdavars) = - if (not o null) vars andalso (#2 o #1 o hd) vars = @{typ "i \ o"} then - ((Thm.term_of o #2 o hd) vars, ctxt1, tl vars, lambdavars) - else if null vars andalso (not o null) lambdavars andalso (freeType o hd) lambdavars = @{typ "i \ o"} then - (hd lambdavars, ctxt1, vars, tl lambdavars) - else Variable.variant_fixes ["N"] ctxt1 |>> var_io o hd |> (fn (cls, ctxt) => (cls, ctxt, vars, lambdavars)) - val db' = db' |> Database.insert Database.abs2rel (cls_pred, cls_pred) - o Database.insert Database.rel2is (cls_pred, cls_pred) - val (v,t) = relativ_tm_frm' is_functional relationalising cls_pred db' ctxt1 tm - val t_vars = sort_strings (Term.add_free_names tm []) - val vs' = List.filter (#1 #> #1 #> #1 #> Ord_List.member String.compare t_vars) vars - val vs = cls_pred :: map (Thm.term_of o #2) vs' @ lambdavars @ the_list v - val at = List.foldr (uncurry lambda) t vs - val abs_const = read_const lthy (if is_external then thm_ref else lname lthy thm_ref) - fun new_const ctxt' = read_new_const ctxt' def_name - fun db_map ctxt' = - Data.map (add_rel_const (get_mode is_functional relationalising) abs_const (new_const ctxt')) - fun add_to_context ctxt' = Context.proof_map (db_map ctxt') ctxt' - fun add_to_theory ctxt' = Local_Theory.raw_theory (Context.theory_map (db_map ctxt')) ctxt' - in - lthy - |> Local_Theory.define ((Binding.name def_name, NoSyn), ((Binding.name (def_name ^ "_def"), []), at)) - |>> (#2 #> (fn (s,t) => (s,[t]))) - |> Utils.display "theorem" pos - |> Local_Theory.target (add_to_theory o add_to_context) - end - -fun relativize_tm is_functional def_name term pos lthy = - let - val ctxt = lthy - val (cls_pred, ctxt1) = Variable.variant_fixes ["N"] ctxt |>> var_io o hd - val tm = Syntax.read_term ctxt1 term - val db' = Data.get (Context.Proof lthy) - val db' = db' |> Database.insert Database.abs2rel (cls_pred, cls_pred) - o Database.insert Database.rel2is (cls_pred, cls_pred) - val vs' = Variable.add_frees ctxt1 tm [] - val ctxt2 = fold Utils.add_to_context (map #1 vs') ctxt1 - val (v,t) = relativ_tm_frm' is_functional false cls_pred db' ctxt2 tm - val vs = cls_pred :: map Free vs' @ the_list v - val at = List.foldr (uncurry lambda) t vs - in - lthy - |> Local_Theory.define ((Binding.name def_name, NoSyn), ((Binding.name (def_name ^ "_def"), []), at)) - |>> (#2 #> (fn (s,t) => (s,[t]))) - |> Utils.display "theorem" pos - end - -val op $` = curry ((op $) o swap) -infix $` - -fun is_free_i (Free (_, @{typ "i"})) = true - | is_free_i _ = false - -fun rel_closed_goal target pos lthy = - let - val (_, tm, _) = Utils.thm_concl_tm lthy (target ^ "_rel_def") - val (def, tm) = tm |> Utils.dest_eq_tms' - fun first_lambdas (Abs (body as (_, ty, _))) = - if ty = @{typ "i"} - then (op ::) (Utils.dest_abs body |>> Utils.var_i ||> first_lambdas) - else Utils.dest_abs body |> first_lambdas o #2 - | first_lambdas _ = [] - val (def, vars) = Term.strip_comb def ||> filter is_free_i - val vs = vars @ first_lambdas tm - val class = Free ("M", @{typ "i \ o"}) - val def = fold (op $`) (class :: vs) def - val hyps = map (fn v => class $ v |> Utils.tp) vs - val concl = class $ def - val goal = Logic.list_implies (hyps, Utils.tp concl) - val attribs = @{attributes [intro, simp]} - in - Proof.theorem NONE (fn thmss => Utils.display "theorem" pos - o Local_Theory.note ((Binding.name (target ^ "_rel_closed"), attribs), hd thmss)) - [[(goal, [])]] lthy - end - -fun iff_goal target pos lthy = - let - val (_, tm, ctxt') = Utils.thm_concl_tm lthy (target ^ "_rel_def") - val (_, is_def, ctxt) = Utils.thm_concl_tm ctxt' ("is_" ^ target ^ "_def") - val is_def = is_def |> Utils.dest_eq_tms' |> #1 |> Term.strip_comb |> #1 - val (def, tm) = tm |> Utils.dest_eq_tms' - fun first_lambdas (Abs (body as (_, ty, _))) = - if ty = @{typ "i"} - then (op ::) (Utils.dest_abs body |>> Utils.var_i ||> first_lambdas) - else Utils.dest_abs body |> first_lambdas o #2 - | first_lambdas _ = [] - val (def, vars) = Term.strip_comb def ||> filter is_free_i - val vs = vars @ first_lambdas tm - val class = Free ("M", @{typ "i \ o"}) - val def = fold (op $`) (class :: vs) def - val ty = fastype_of def - val res = if ty = @{typ "i"} - then Variable.variant_fixes ["res"] ctxt |> SOME o Utils.var_i o hd o #1 - else NONE - val is_def = fold (op $`) (class :: vs @ the_list res) is_def - val hyps = map (fn v => class $ v |> Utils.tp) (vs @ the_list res) - val concl = @{const "IFOL.iff"} $ is_def - $ (if ty = @{typ "i"} then (@{const IFOL.eq(i)} $ the res $ def) else def) - val goal = Logic.list_implies (hyps, Utils.tp concl) - in - Proof.theorem NONE (fn thmss => Utils.display "theorem" pos - o Local_Theory.note ((Binding.name ("is_" ^ target ^ "_iff"), []), hd thmss)) - [[(goal, [])]] lthy - end - -fun univalent_goal target pos lthy = - let - val (_, tm, ctxt) = Utils.thm_concl_tm lthy ("is_" ^ target ^ "_def") - val (def, tm) = tm |> Utils.dest_eq_tms' - fun first_lambdas (Abs (body as (_, ty, _))) = - if ty = @{typ "i"} - then (op ::) (Utils.dest_abs body |>> Utils.var_i ||> first_lambdas) - else Utils.dest_abs body |> first_lambdas o #2 - | first_lambdas _ = [] - val (def, vars) = Term.strip_comb def ||> filter is_free_i - val vs = vars @ first_lambdas tm - val n = length vs - val vs = List.take (vs, n - 2) - val class = Free ("M", @{typ "i \ o"}) - val def = fold (op $`) (class :: vs) def - val v = Variable.variant_fixes ["A"] ctxt |> Utils.var_i o hd o #1 - val hyps = map (fn v => class $ v |> Utils.tp) (v :: vs) - val concl = @{const "Relative.univalent"} $ class $ v $ def - val goal = Logic.list_implies (hyps, Utils.tp concl) - in - Proof.theorem NONE (fn thmss => Utils.display "theorem" pos - o Local_Theory.note ((Binding.name ("univalent_is_" ^ target), []), hd thmss)) - [[(goal, [])]] lthy - end - -end -\ - -ML\ -local - val full_mode_parser = - Scan.option (((Parse.$$$ "functional" |-- Parse.$$$ "relational") >> K Database.rel2is) - || (((Scan.option (Parse.$$$ "absolute")) |-- Parse.$$$ "functional") >> K Database.abs2rel) - || (((Scan.option (Parse.$$$ "absolute")) |-- Parse.$$$ "relational") >> K Database.abs2is)) - >> (fn mode => the_default Database.abs2is mode) - - val reldb_parser = - Parse.position (full_mode_parser -- (Parse.string -- Parse.string)); - - val singlemode_parser = (Parse.$$$ "absolute" >> K Database.remove_abs) - || (Parse.$$$ "functional" >> K Database.remove_rel) - || (Parse.$$$ "relational" >> K Database.remove_is) - - val reldb_rem_parser = Parse.position (singlemode_parser -- Parse.string) - - val mode_parser = - Scan.option ((Parse.$$$ "relational" >> K false) || (Parse.$$$ "functional" >> K true)) - >> (fn mode => if is_none mode then false else the mode) - - val relativize_parser = - Parse.position (mode_parser -- (Parse.string -- Parse.string) -- (Scan.optional (Parse.$$$ "external" >> K true) false)); - - val _ = - Outer_Syntax.local_theory \<^command_keyword>\reldb_add\ "ML setup for adding relativized/absolute pairs" - (reldb_parser >> (fn ((mode, (abs_term,rel_term)),_) => - Relativization.add_constant mode abs_term rel_term)) - - val _ = - Outer_Syntax.local_theory \<^command_keyword>\reldb_rem\ "ML setup for adding relativized/absolute pairs" - (reldb_rem_parser >> (uncurry Relativization.rem_constant o #1)) - - val _ = - Outer_Syntax.local_theory \<^command_keyword>\relativize\ "ML setup for relativizing definitions" - (relativize_parser >> (fn (((is_functional, (bndg,thm)), is_external),pos) => - Relativization.relativize_def is_external is_functional false thm bndg pos)) - - val _ = - Outer_Syntax.local_theory \<^command_keyword>\relativize_tm\ "ML setup for relativizing definitions" - (relativize_parser >> (fn (((is_functional, (bndg,term)), _),pos) => - Relativization.relativize_tm is_functional term bndg pos)) - - val _ = - Outer_Syntax.local_theory \<^command_keyword>\relationalize\ "ML setup for relativizing definitions" - (relativize_parser >> (fn (((is_functional, (bndg,thm)), is_external),pos) => - Relativization.relativize_def is_external is_functional true thm bndg pos)) - - val _ = - Outer_Syntax.local_theory_to_proof \<^command_keyword>\rel_closed\ "ML setup for rel_closed theorem" - (Parse.position (Parse.$$$ "for" |-- Parse.string) >> (fn (target,pos) => - Relativization.rel_closed_goal target pos)) - - val _ = - Outer_Syntax.local_theory_to_proof \<^command_keyword>\is_iff_rel\ "ML setup for rel_closed theorem" - (Parse.position (Parse.$$$ "for" |-- Parse.string) >> (fn (target,pos) => - Relativization.iff_goal target pos)) - - val _ = - Outer_Syntax.local_theory_to_proof \<^command_keyword>\univalent\ "ML setup for rel_closed theorem" - (Parse.position (Parse.$$$ "for" |-- Parse.string) >> (fn (target,pos) => - Relativization.univalent_goal target pos)) - -val _ = - Theory.setup - (Attrib.setup \<^binding>\Rel\ (Attrib.add_del Relativization.Rel_add Relativization.Rel_del) - "declaration of relativization rule") ; -in -end -\ setup\Relativization.init_db Relativization.db \ declare relative_abs[Rel] - (*todo: check all the duplicate cases here.*) +(*TODO: check all the duplicate cases here.*) declare datatype_abs[Rel] ML\ val db = Relativization.get_db @{context} \ end diff --git a/thys/Transitive_Models/Renaming_Auto.thy b/thys/Transitive_Models/Renaming_Auto.thy --- a/thys/Transitive_Models/Renaming_Auto.thy +++ b/thys/Transitive_Models/Renaming_Auto.thy @@ -1,56 +1,23 @@ theory Renaming_Auto imports Utils Renaming keywords "rename" :: thy_decl % "ML" and "simple_rename" :: thy_decl % "ML" and "src" and "tgt" abbrevs "simple_rename" = "" begin hide_const (open) Order.pred lemmas nat_succI = nat_succ_iff[THEN iffD2] ML_file\Renaming_ML.ml\ -ML\ - open Renaming_ML - - fun renaming_def mk_ren name from to ctxt = - let val to = to |> Syntax.read_term ctxt - val from = from |> Syntax.read_term ctxt - val (tc_lemma,action_lemma,fvs,r) = mk_ren from to ctxt - val (tc_lemma,action_lemma) = (fix_vars tc_lemma fvs ctxt , fix_vars action_lemma fvs ctxt) - val ren_fun_name = Binding.name (name ^ "_fn") - val ren_fun_def = Binding.name (name ^ "_fn_def") - val ren_thm = Binding.name (name ^ "_thm") - in - Local_Theory.note ((ren_thm, []), [tc_lemma,action_lemma]) ctxt |> snd |> - Local_Theory.define ((ren_fun_name, NoSyn), ((ren_fun_def, []), r)) |> snd - end; -\ -ML\ -local - - val ren_parser = Parse.position (Parse.string -- - (Parse.$$$ "src" |-- Parse.string --| Parse.$$$ "tgt" -- Parse.string)); - - val _ = - Outer_Syntax.local_theory \<^command_keyword>\rename\ "ML setup for synthetic definitions" - (ren_parser >> (fn ((name,(from,to)),_) => renaming_def sum_rename name from to )) - - val _ = - Outer_Syntax.local_theory \<^command_keyword>\simple_rename\ "ML setup for synthetic definitions" - (ren_parser >> (fn ((name,(from,to)),_) => renaming_def ren_thm name from to )) - -in -end -\ end \ No newline at end of file diff --git a/thys/Transitive_Models/Renaming_ML.ml b/thys/Transitive_Models/Renaming_ML.ml --- a/thys/Transitive_Models/Renaming_ML.ml +++ b/thys/Transitive_Models/Renaming_ML.ml @@ -1,179 +1,205 @@ structure Renaming_ML = struct open Utils fun sum_ f g m n p = @{const Renaming.rsum} $ f $ g $ m $ n $ p (*Builds a finite mapping from rho to rho'.*) fun mk_ren rho rho' ctxt = let val rs = to_ML_list rho val rs' = to_ML_list rho' val ixs = 0 upto (length rs-1) fun err t = "The element " ^ Syntax.string_of_term ctxt t ^ " is missing in the target environment" fun mkp i = case find_index (fn x => x = nth rs i) rs' of ~1 => nth rs i |> err |> error | j => mk_Pair (mk_ZFnat i) (mk_ZFnat j) in map mkp ixs |> mk_FinSet end fun mk_dom_lemma ren rho = let val n = rho |> to_ML_list |> length |> mk_ZFnat in eq_ n (@{const domain} $ ren) |> tp end fun ren_tc_goal fin ren rho rho' = let val n = rho |> to_ML_list |> length |> mk_ZFnat val m = rho' |> to_ML_list |> length |> mk_ZFnat val fun_ty = if fin then @{const_name "FiniteFun"} else @{const_abbrev "function_space"} val ty = Const (fun_ty,@{typ "i \ i \ i"}) $ n $ m in mem_ ren ty |> tp end fun ren_action_goal ren rho rho' ctxt = let val setV = Variable.variant_frees ctxt [] [("A",@{typ i})] |> hd |> Free val j = Variable.variant_frees ctxt [] [("j",@{typ i})] |> hd |> Free val vs = rho |> to_ML_list val ws = rho' |> to_ML_list |> filter Term.is_Free val h1 = subset_ (mk_FinSet vs) setV val h2 = lt_ j (length vs |> mk_ZFnat) val fvs = [j,setV ] @ ws |> filter Term.is_Free |> map freeName val lhs = nth_ j rho val rhs = nth_ (app_ ren j) rho' val concl = eq_ lhs rhs in (Logic.list_implies([tp h1,tp h2],tp concl),fvs) end fun sum_tc_goal f m n p = let val m_length = m |> to_ML_list |> length |> mk_ZFnat val n_length = n |> to_ML_list |> length |> mk_ZFnat val p_length = p |> length_ val id_fun = @{const id} $ p_length val sum_fun = sum_ f id_fun m_length n_length p_length val dom = add_ m_length p_length val codom = add_ n_length p_length val fun_ty = @{const_abbrev "function_space"} val ty = Const (fun_ty,@{typ "i \ i \ i"}) $ dom $ codom in (sum_fun, mem_ sum_fun ty |> tp) end fun sum_action_goal ren rho rho' ctxt = let val setV = Variable.variant_frees ctxt [] [("A",@{typ i})] |> hd |> Free val envV = Variable.variant_frees ctxt [] [("env",@{typ i})] |> hd |> Free val j = Variable.variant_frees ctxt [] [("j",@{typ i})] |> hd |> Free val vs = rho |> to_ML_list val ws = rho' |> to_ML_list |> filter Term.is_Free val envL = envV |> length_ val rhoL = vs |> length |> mk_ZFnat val h1 = subset_ (append vs ws |> mk_FinSet) setV val h2 = lt_ j (add_ rhoL envL) val h3 = mem_ envV (list_ setV) val fvs = ([j,setV,envV] @ ws |> filter Term.is_Free) |> map freeName val lhs = nth_ j (concat_ rho envV) val rhs = nth_ (app_ ren j) (concat_ rho' envV) val concl = eq_ lhs rhs in (Logic.list_implies([tp h1,tp h2,tp h3],tp concl),fvs) end (* Tactics *) fun fin ctxt = REPEAT (resolve_tac ctxt [@{thm nat_succI}] 1) THEN resolve_tac ctxt [@{thm nat_0I}] 1 fun step ctxt thm = asm_full_simp_tac ctxt 1 THEN asm_full_simp_tac ctxt 1 THEN EqSubst.eqsubst_tac ctxt [1] [@{thm app_fun} OF [thm]] 1 THEN simp_tac ctxt 1 THEN simp_tac ctxt 1 fun fin_fun_tac ctxt = REPEAT ( resolve_tac ctxt [@{thm consI}] 1 THEN resolve_tac ctxt [@{thm ltD}] 1 THEN simp_tac ctxt 1 THEN resolve_tac ctxt [@{thm ltD}] 1 THEN simp_tac ctxt 1) THEN resolve_tac ctxt [@{thm emptyI}] 1 THEN REPEAT (simp_tac ctxt 1) fun ren_thm e e' ctxt = let val r = mk_ren e e' ctxt val fin_tc_goal = ren_tc_goal true r e e' val dom_goal = mk_dom_lemma r e val tc_goal = ren_tc_goal false r e e' val (action_goal,fvs) = ren_action_goal r e e' ctxt val fin_tc_lemma = Goal.prove ctxt [] [] fin_tc_goal (fn _ => fin_fun_tac ctxt) val dom_lemma = Goal.prove ctxt [] [] dom_goal (fn _ => blast_tac ctxt 1) val tc_lemma = Goal.prove ctxt [] [] tc_goal (fn _ => EqSubst.eqsubst_tac ctxt [1] [dom_lemma] 1 THEN resolve_tac ctxt [@{thm FiniteFun_is_fun}] 1 THEN resolve_tac ctxt [fin_tc_lemma] 1) val action_lemma = Goal.prove ctxt [] [] action_goal (fn _ => forward_tac ctxt [@{thm le_natI}] 1 THEN fin ctxt THEN REPEAT (resolve_tac ctxt [@{thm natE}] 1 THEN step ctxt tc_lemma) THEN (step ctxt tc_lemma) ) in (action_lemma, tc_lemma, fvs, r) end (* Returns the sum renaming, the goal for type_checking, and the actual lemmas for the left part of the sum. *) fun sum_ren_aux e e' ctxt = let val env = Variable.variant_frees ctxt [] [("env",@{typ i})] |> hd |> Free val (left_action_lemma,left_tc_lemma,_,r) = ren_thm e e' ctxt val (sum_ren,sum_goal_tc) = sum_tc_goal r e e' env val setV = Variable.variant_frees ctxt [] [("A",@{typ i})] |> hd |> Free fun hyp en = mem_ en (list_ setV) in (sum_ren, freeName env, Logic.list_implies (map (fn e => e |> hyp |> tp) [env], sum_goal_tc), left_tc_lemma, left_action_lemma) end fun sum_tc_lemma rho rho' ctxt = let val (sum_ren, envVar, tc_goal, left_tc_lemma, left_action_lemma) = sum_ren_aux rho rho' ctxt val (goal,fvs) = sum_action_goal sum_ren rho rho' ctxt val r = mk_ren rho rho' ctxt in (sum_ren, goal,envVar, r,left_tc_lemma, left_action_lemma ,fvs, Goal.prove ctxt [] [] tc_goal (fn _ => resolve_tac ctxt [@{thm sum_type_id_aux2}] 1 THEN asm_simp_tac ctxt 4 THEN simp_tac ctxt 1 THEN resolve_tac ctxt [left_tc_lemma] 1 THEN (fin ctxt) THEN (fin ctxt) )) end fun sum_rename rho rho' ctxt = let val (_, goal, _, left_rename, left_tc_lemma, left_action_lemma, fvs, sum_tc_lemma) = sum_tc_lemma rho rho' ctxt val action_lemma = fix_vars left_action_lemma fvs ctxt in (sum_tc_lemma, Goal.prove ctxt [] [] goal (fn _ => resolve_tac ctxt [@{thm sum_action_id_aux}] 1 THEN (simp_tac ctxt 4) THEN (simp_tac ctxt 1) THEN (resolve_tac ctxt [left_tc_lemma] 1) THEN (asm_full_simp_tac ctxt 1) THEN (asm_full_simp_tac ctxt 1) THEN (simp_tac ctxt 1) THEN (simp_tac ctxt 1) THEN (simp_tac ctxt 1) THEN (full_simp_tac ctxt 1) THEN (resolve_tac ctxt [action_lemma] 1) THEN (blast_tac ctxt 1) THEN (full_simp_tac ctxt 1) THEN (full_simp_tac ctxt 1) ), fvs, left_rename ) -end ; +end + +fun renaming_def mk_ren name from to ctxt = + let val to = to |> Syntax.read_term ctxt + val from = from |> Syntax.read_term ctxt + val (tc_lemma,action_lemma,fvs,r) = mk_ren from to ctxt + val (tc_lemma,action_lemma) = (fix_vars tc_lemma fvs ctxt , fix_vars action_lemma fvs ctxt) + val ren_fun_name = Binding.name (name ^ "_fn") + val ren_fun_def = Binding.name (name ^ "_fn_def") + val ren_thm = Binding.name (name ^ "_thm") + in + Local_Theory.note ((ren_thm, []), [tc_lemma,action_lemma]) ctxt |> snd |> + Local_Theory.define ((ren_fun_name, NoSyn), ((ren_fun_def, []), r)) |> snd +end; + +val ren_parser = Parse.position (Parse.string -- + (Parse.$$$ "src" |-- Parse.string --| Parse.$$$ "tgt" -- Parse.string)); + +val _ = + Outer_Syntax.local_theory \<^command_keyword>\rename\ "ML setup for synthetic definitions" + (ren_parser >> (fn ((name,(from,to)),_) => renaming_def sum_rename name from to )) + +val _ = + Outer_Syntax.local_theory \<^command_keyword>\simple_rename\ "ML setup for synthetic definitions" + (ren_parser >> (fn ((name,(from,to)),_) => renaming_def ren_thm name from to )) + + end \ No newline at end of file diff --git a/thys/Transitive_Models/Synthetic_Definition.ml b/thys/Transitive_Models/Synthetic_Definition.ml new file mode 100644 --- /dev/null +++ b/thys/Transitive_Models/Synthetic_Definition.ml @@ -0,0 +1,328 @@ +val $` = curry ((op $) o swap) +infix $` + +infix 6 &&& +val op &&& = Utils.&&& + +infix 6 *** +val op *** = Utils.*** + +fun arity_goal intermediate def_name lthy = + let + val thm = Proof_Context.get_thm lthy (def_name ^ "_def") + val (_, tm, _) = Utils.thm_concl_tm lthy (def_name ^ "_def") + val (def, tm) = tm |> Utils.dest_eq_tms' + fun first_lambdas (Abs (body as (_, ty, _))) = + if ty = @{typ "i"} + then (op ::) (Utils.dest_abs body |>> Utils.var_i ||> first_lambdas) + else Utils.dest_abs body |> first_lambdas o #2 + | first_lambdas _ = [] + val (def, vars) = Term.strip_comb def + val vs = vars @ first_lambdas tm + val def = fold (op $`) vs def + val hyps = map (fn v => Utils.mem_ v Utils.nat_ |> Utils.tp) vs + val concl = @{const IFOL.eq(i)} $ (@{const arity} $ def) $ Var (("ar", 0), @{typ "i"}) + val g_iff = Logic.list_implies (hyps, Utils.tp concl) + val attribs = if intermediate then [] else @{attributes [arity]} + in + (g_iff, "arity_" ^ def_name ^ (if intermediate then "'" else ""), attribs, thm, vs) + end + +fun manual_arity intermediate def_name pos lthy = + let + val (goal, thm_name, attribs, _, _) = arity_goal intermediate def_name lthy + in + Proof.theorem NONE (fn thmss => Utils.display "theorem" pos + o Local_Theory.note ((Binding.name thm_name, attribs), hd thmss)) + [[(goal, [])]] lthy + end + +fun prove_arity thms goal ctxt = + let + val rules = (Named_Theorems.get ctxt \<^named_theorems>\arity\) @ + (Named_Theorems.get ctxt \<^named_theorems>\arity_aux\) + in + Goal.prove ctxt [] [] goal + (K (rewrite_goal_tac ctxt thms 1 THEN Method.insert_tac ctxt rules 1 THEN asm_simp_tac ctxt 1)) + end + +fun auto_arity intermediate def_name pos lthy = + let + val (goal, thm_name, attribs, def_thm, vs) = arity_goal intermediate def_name lthy + val intermediate_text = if intermediate then "intermediate" else "" + val help = "You can manually prove the arity_theorem by typing:\n" + ^ "manual_arity " ^ intermediate_text ^ " for \"" ^ def_name ^ "\"\n" + val thm = prove_arity [def_thm] goal lthy + handle ERROR s => help ^ "\n\n" ^ s |> Exn.reraise o ERROR + val thm = Utils.fix_vars thm (map Utils.freeName vs) lthy + in + Local_Theory.note ((Binding.name thm_name, attribs), [thm]) lthy |> Utils.display "theorem" pos + end + +fun prove_tc_form goal thms ctxt = + Goal.prove ctxt [] [] goal (K (rewrite_goal_tac ctxt thms 1 THEN auto_tac ctxt)) + +fun prove_sats_tm thm_auto thms goal ctxt = + let + val ctxt' = ctxt |> Simplifier.add_simp (hd thm_auto) + in + Goal.prove ctxt [] [] goal + (K (rewrite_goal_tac ctxt thms 1 THEN PARALLEL_ALLGOALS (asm_simp_tac ctxt'))) + end + +fun prove_sats_iff goal ctxt = Goal.prove ctxt [] [] goal (K (asm_simp_tac ctxt 1)) + +fun is_mem (@{const mem} $ _ $ _) = true + | is_mem _ = false + +fun pre_synth_thm_sats term set env vars vs lthy = + let + val (_, tm, ctxt1) = Utils.thm_concl_tm lthy term + val (thm_refs, ctxt2) = Variable.import true [Proof_Context.get_thm lthy term] ctxt1 |>> #2 + val vs' = map (Thm.term_of o #2) vs + val vars' = map (Thm.term_of o #2) vars + val r_tm = tm |> Utils.dest_lhs_def |> fold (op $`) vs' + val sats = @{const apply} $ (@{const satisfies} $ set $ r_tm) $ env + val sats' = @{const IFOL.eq(i)} $ sats $ (@{const succ} $ @{const zero}) + in + { vars = vars' + , vs = vs' + , sats = sats' + , thm_refs = thm_refs + , lthy = ctxt2 + , env = env + , set = set + } + end + +fun synth_thm_sats_gen name lhs hyps pos attribs aux_funs environment lthy = + let + val ctxt = (#prepare_ctxt aux_funs) lthy + val ctxt = Utils.add_to_context (Utils.freeName (#set environment)) ctxt + val (new_vs, ctxt') = (#create_variables aux_funs) (#vs environment, ctxt) + val new_hyps = (#create_hyps aux_funs) (#vs environment, new_vs) + val concl = (#make_concl aux_funs) (lhs, #sats environment, new_vs) + val g_iff = Logic.list_implies (new_hyps @ hyps, Utils.tp concl) + val thm = (#prover aux_funs) g_iff ctxt' + val thm = Utils.fix_vars thm (map Utils.freeName ((#vars environment) @ new_vs)) lthy + in + Local_Theory.note ((name, attribs), [thm]) lthy |> Utils.display "theorem" pos + end + +fun synth_thm_sats_iff def_name lhs hyps pos environment = + let + val subst = Utils.zip_with (I *** I) (#vs environment) + fun subst_nth (@{const "nth"} $ v $ _) new_vs = AList.lookup (op =) (subst new_vs) v |> the + | subst_nth (t1 $ t2) new_vs = (subst_nth t1 new_vs) $ (subst_nth t2 new_vs) + | subst_nth (Abs (v, ty, t)) new_vs = Abs (v, ty, subst_nth t new_vs) + | subst_nth t _ = t + val name = Binding.name (def_name ^ "_iff_sats") + val iff_sats_attrib = @{attributes [iff_sats]} + val aux_funs = { prepare_ctxt = fold Utils.add_to_context (map Utils.freeName (#vs environment)) + , create_variables = fn (vs, ctxt) => Variable.variant_fixes (map Utils.freeName vs) ctxt |>> map Utils.var_i + , create_hyps = fn (vs, new_vs) => Utils.zip_with (fn (v, nv) => Utils.eq_ (Utils.nth_ v (#env environment)) nv) vs new_vs |> map Utils.tp + , make_concl = fn (lhs, rhs, new_vs) => @{const IFOL.iff} $ (subst_nth lhs new_vs) $ rhs + , prover = prove_sats_iff + } + in + synth_thm_sats_gen name lhs hyps pos iff_sats_attrib aux_funs environment + end + +fun synth_thm_sats_fm def_name lhs hyps pos thm_auto environment = + let + val name = Binding.name ("sats_" ^ def_name ^ "_fm") + val simp_attrib = @{attributes [simp]} + val aux_funs = { prepare_ctxt = I + , create_variables = K [] *** I + , create_hyps = K [] + , make_concl = fn (rhs, lhs, _) => @{const IFOL.iff} $ lhs $ rhs + , prover = prove_sats_tm thm_auto (#thm_refs environment) + } + in + synth_thm_sats_gen name lhs hyps pos simp_attrib aux_funs environment + end + +fun synth_thm_tc def_name term hyps vars pos lthy = + let + val (_,tm,ctxt1) = Utils.thm_concl_tm lthy term + val (thm_refs,ctxt2) = Variable.import true [Proof_Context.get_thm lthy term] ctxt1 |>> #2 + val vars' = map (Thm.term_of o #2) vars + val tc_attrib = @{attributes [TC]} + val r_tm = tm |> Utils.dest_lhs_def |> fold (op $`) vars' + val concl = @{const mem} $ r_tm $ @{const formula} + val g = Logic.list_implies(hyps, Utils.tp concl) + val thm = prove_tc_form g thm_refs ctxt2 + val name = Binding.name (def_name ^ "_fm_type") + val thm = Utils.fix_vars thm (map Utils.freeName vars') ctxt2 + in + Local_Theory.note ((name, tc_attrib), [thm]) lthy |> Utils.display "theorem" pos + end + +fun synthetic_def def_name thm_ref pos tc auto thy = + let + val thm = Proof_Context.get_thm thy thm_ref + val thm_vars = rev (Term.add_vars (Thm.full_prop_of thm) []) + val (((_,inst),thm_tms),_) = Variable.import true [thm] thy + val vars = map (fn v => (v, the (Vars.lookup inst v))) thm_vars + val (tm,hyps) = thm_tms |> hd |> Thm.concl_of &&& Thm.prems_of + val (lhs,rhs) = tm |> Utils.dest_iff_tms o Utils.dest_trueprop + val ((set,t),env) = rhs |> Utils.dest_sats_frm + fun relevant ts (@{const mem} $ t $ _) = + (not (t = @{term "0"})) andalso + (not (Term.is_Free t) orelse member (op =) ts (t |> Term.dest_Free |> #1)) + | relevant _ _ = false + val t_vars = sort_strings (Term.add_free_names t []) + val vs = filter (Ord_List.member String.compare t_vars o #1 o #1 o #1) vars + val at = fold_rev (lambda o Thm.term_of o #2) vs t + val hyps' = filter (relevant t_vars o Utils.dest_trueprop) hyps + val def_attrs = @{attributes [fm_definitions]} + in + Local_Theory.define ((Binding.name (def_name ^ "_fm"), NoSyn), + ((Binding.name (def_name ^ "_fm_def"), def_attrs), at)) thy + |>> (#2 #> I *** single) |> Utils.display "theorem" pos |> + (if tc then synth_thm_tc def_name (def_name ^ "_fm_def") hyps' vs pos else I) |> + (if auto then + pre_synth_thm_sats (def_name ^ "_fm_def") set env vars vs + #> I &&& #lthy + #> #1 &&& uncurry (synth_thm_sats_fm def_name lhs hyps pos thm_tms) + #> uncurry (synth_thm_sats_iff def_name lhs hyps pos) + else I) + end + +fun prove_schematic thms goal ctxt = + let + val rules = Named_Theorems.get ctxt \<^named_theorems>\iff_sats\ + in + Goal.prove ctxt [] [] goal + (K (rewrite_goal_tac ctxt thms 1 THEN REPEAT1 (CHANGED (resolve_tac ctxt rules 1 ORELSE asm_simp_tac ctxt 1)))) + end + +val valid_assumptions = [ ("nonempty", Utils.mem_ @{term "0"}) + ] + +fun pre_schematic_def target assuming lthy = +let + val thm = Proof_Context.get_thm lthy (target ^ "_def") + val (vars, tm, ctxt1) = Utils.thm_concl_tm lthy (target ^ "_def") + val (const, tm) = tm |> Utils.dest_eq_tms' o Utils.dest_trueprop |>> #1 o strip_comb + val t_vars = sort_strings (Term.add_free_names tm []) + val vs = List.filter (#1 #> #1 #> #1 #> Ord_List.member String.compare t_vars) vars + |> List.filter ((curry op = @{typ "i"}) o #2 o #1) + |> List.map (Utils.var_i o #1 o #1 o #1) + fun first_lambdas (Abs (body as (_, ty, _))) = + if ty = @{typ "i"} + then (op ::) (Utils.dest_abs body |>> Utils.var_i ||> first_lambdas) + else Utils.dest_abs body |> first_lambdas o #2 + | first_lambdas _ = [] + val vs = vs @ (first_lambdas tm) + val ctxt1' = fold Utils.add_to_context (map Utils.freeName vs) ctxt1 + val (set, ctxt2) = Variable.variant_fixes ["A"] ctxt1' |>> Utils.var_i o hd + val class = @{const "setclass"} $ set + val (env, ctxt3) = Variable.variant_fixes ["env"] ctxt2 |>> Utils.var_i o hd + val assumptions = filter (member (op =) assuming o #1) valid_assumptions |> map #2 + val hyps = (List.map (fn v => Utils.tp (Utils.mem_ v Utils.nat_)) vs) + @ [Utils.tp (Utils.mem_ env (Utils.list_ set))] + @ Utils.zip_with (fn (f,x) => Utils.tp (f x)) assumptions (replicate (length assumptions) set) + val args = class :: map (fn v => Utils.nth_ v env) vs + val (fm_name, ctxt4) = Variable.variant_fixes ["fm"] ctxt3 |>> hd + val fm_type = fold (K (fn acc => Type ("fun", [@{typ "i"}, acc]))) vs @{typ "i"} + val fm = Var ((fm_name, 0), fm_type) + val lhs = fold (op $`) args const + val fm_app = fold (op $`) vs fm + val sats = @{const apply} $ (@{const satisfies} $ set $ fm_app) $ env + val rhs = @{const IFOL.eq(i)} $ sats $ (@{const succ} $ @{const zero}) + val concl = @{const "IFOL.iff"} $ lhs $ rhs + val schematic = Logic.list_implies (hyps, Utils.tp concl) + in + (schematic, ctxt4, thm, set, env, vs) + end + +fun str_join _ [] = "" + | str_join _ [s] = s + | str_join c (s :: ss) = s ^ c ^ (str_join c ss) + +fun schematic_def def_name target assuming pos lthy = + let + val (schematic, ctxt, thm, set, env, vs) = pre_schematic_def target assuming lthy + val assuming_text = if null assuming then "" else "assuming " ^ (map (fn s => "\"" ^ s ^ "\"") assuming |> str_join " ") + val help = "You can manually prove the schematic_goal by typing:\n" + ^ "manual_schematic [sch_name] for \"" ^ target ^ "\"" ^ assuming_text ^"\n" + ^ "And then complete the synthesis with:\n" + ^ "synthesize \"" ^ target ^ "\" from_schematic [sch_name]\n" + ^ "In both commands, \sch_name\ must be the same and, if ommited, will be defaulted to sats_" ^ target ^ "_fm_auto.\n" + ^ "You can also try adding new assumptions and/or synthetizing definitions of sub-terms." + val thm = prove_schematic [thm] schematic ctxt + handle ERROR s => help ^ "\n\n" ^ s |> Exn.reraise o ERROR + val thm = Utils.fix_vars thm (map Utils.freeName (set :: env :: vs)) lthy + in + Local_Theory.note ((Binding.name def_name, []), [thm]) lthy |> Utils.display "theorem" pos + end + +fun schematic_synthetic_def def_name target assuming pos tc auto = + (synthetic_def def_name ("sats_" ^ def_name ^ "_fm_auto") pos tc auto) + o (schematic_def ("sats_" ^ def_name ^ "_fm_auto") target assuming pos) + +fun manual_schematic def_name target assuming pos lthy = + let + val (schematic, _, _, _, _, _) = pre_schematic_def target assuming lthy + in + Proof.theorem NONE (fn thmss => Utils.display "theorem" pos + o Local_Theory.note ((Binding.name def_name, []), hd thmss)) + [[(schematic, [])]] lthy + end + +local + val simple_from_schematic_synth_constdecl = + Parse.string --| (Parse.$$$ "from_schematic") + >> (fn bndg => synthetic_def bndg ("sats_" ^ bndg ^ "_fm_auto")) + + val full_from_schematic_synth_constdecl = + Parse.string -- ((Parse.$$$ "from_schematic" |-- Parse.thm)) + >> (fn (bndg,thm) => synthetic_def bndg (#1 (thm |>> Facts.ref_name))) + + val full_from_definition_synth_constdecl = + Parse.string -- ((Parse.$$$ "from_definition" |-- Parse.string)) -- (Scan.optional (Parse.$$$ "assuming" |-- Scan.repeat Parse.string) []) + >> (fn ((bndg,target), assuming) => schematic_synthetic_def bndg target assuming) + + val simple_from_definition_synth_constdecl = + Parse.string -- (Parse.$$$ "from_definition" |-- (Scan.optional (Parse.$$$ "assuming" |-- Scan.repeat Parse.string)) []) + >> (fn (bndg, assuming) => schematic_synthetic_def bndg bndg assuming) + + val synth_constdecl = + Parse.position (full_from_schematic_synth_constdecl || simple_from_schematic_synth_constdecl + || full_from_definition_synth_constdecl + || simple_from_definition_synth_constdecl) + + val full_schematic_decl = + Parse.string -- ((Parse.$$$ "for" |-- Parse.string)) -- (Scan.optional (Parse.$$$ "assuming" |-- Scan.repeat Parse.string) []) + + val simple_schematic_decl = + (Parse.$$$ "for" |-- Parse.string >> (fn name => "sats_" ^ name ^ "_fm_auto") &&& I) -- (Scan.optional (Parse.$$$ "assuming" |-- Scan.repeat Parse.string) []) + + val schematic_decl = Parse.position (full_schematic_decl || simple_schematic_decl) + + val _ = + Outer_Syntax.local_theory \<^command_keyword>\synthesize\ "ML setup for synthetic definitions" + (synth_constdecl >> (fn (f,p) => f p true true)) + + val _ = + Outer_Syntax.local_theory \<^command_keyword>\synthesize_notc\ "ML setup for synthetic definitions" + (synth_constdecl >> (fn (f,p) => f p false false)) + + val _ = Outer_Syntax.local_theory \<^command_keyword>\generate_schematic\ "ML setup for schematic goals" + (schematic_decl >> (fn (((bndg,target), assuming),p) => schematic_def bndg target assuming p)) + + val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\manual_schematic\ "ML setup for schematic goals" + (schematic_decl >> (fn (((bndg,target), assuming),p) => manual_schematic bndg target assuming p)) + + val arity_parser = Parse.position ((Scan.option (Parse.$$$ "intermediate") >> is_some) -- (Parse.$$$ "for" |-- Parse.string)) + + val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\manual_arity\ "ML setup for arities" + (arity_parser >> (fn ((intermediate, target), pos) => manual_arity intermediate target pos)) + + val _ = Outer_Syntax.local_theory \<^command_keyword>\arity_theorem\ "ML setup for arities" + (arity_parser >> (fn ((intermediate, target), pos) => auto_arity intermediate target pos)) + +in + +end \ No newline at end of file diff --git a/thys/Transitive_Models/Synthetic_Definition.thy b/thys/Transitive_Models/Synthetic_Definition.thy --- a/thys/Transitive_Models/Synthetic_Definition.thy +++ b/thys/Transitive_Models/Synthetic_Definition.thy @@ -1,384 +1,53 @@ section\Automatic synthesis of formulas\ theory Synthetic_Definition imports Utils keywords "synthesize" :: thy_decl % "ML" and "synthesize_notc" :: thy_decl % "ML" and "generate_schematic" :: thy_decl % "ML" and "arity_theorem" :: thy_decl % "ML" and "manual_schematic" :: thy_goal_stmt % "ML" and "manual_arity" :: thy_goal_stmt % "ML" and "from_schematic" and "for" and "from_definition" and "assuming" and "intermediate" begin named_theorems fm_definitions "Definitions of synthetized formulas." named_theorems iff_sats "Theorems for synthetising formulas." named_theorems arity "Theorems for arity of formulas." named_theorems arity_aux "Auxiliary theorems for calculating arities." -ML\ -val $` = curry ((op $) o swap) -infix $` - -infix 6 &&& -val op &&& = Utils.&&& - -infix 6 *** -val op *** = Utils.*** - -fun arity_goal intermediate def_name lthy = - let - val thm = Proof_Context.get_thm lthy (def_name ^ "_def") - val (_, tm, _) = Utils.thm_concl_tm lthy (def_name ^ "_def") - val (def, tm) = tm |> Utils.dest_eq_tms' - fun first_lambdas (Abs (body as (_, ty, _))) = - if ty = @{typ "i"} - then (op ::) (Utils.dest_abs body |>> Utils.var_i ||> first_lambdas) - else Utils.dest_abs body |> first_lambdas o #2 - | first_lambdas _ = [] - val (def, vars) = Term.strip_comb def - val vs = vars @ first_lambdas tm - val def = fold (op $`) vs def - val hyps = map (fn v => Utils.mem_ v Utils.nat_ |> Utils.tp) vs - val concl = @{const IFOL.eq(i)} $ (@{const arity} $ def) $ Var (("ar", 0), @{typ "i"}) - val g_iff = Logic.list_implies (hyps, Utils.tp concl) - val attribs = if intermediate then [] else @{attributes [arity]} - in - (g_iff, "arity_" ^ def_name ^ (if intermediate then "'" else ""), attribs, thm, vs) - end - -fun manual_arity intermediate def_name pos lthy = - let - val (goal, thm_name, attribs, _, _) = arity_goal intermediate def_name lthy - in - Proof.theorem NONE (fn thmss => Utils.display "theorem" pos - o Local_Theory.note ((Binding.name thm_name, attribs), hd thmss)) - [[(goal, [])]] lthy - end - -fun prove_arity thms goal ctxt = - let - val rules = (Named_Theorems.get ctxt \<^named_theorems>\arity\) @ - (Named_Theorems.get ctxt \<^named_theorems>\arity_aux\) - in - Goal.prove ctxt [] [] goal - (K (rewrite_goal_tac ctxt thms 1 THEN Method.insert_tac ctxt rules 1 THEN asm_simp_tac ctxt 1)) - end - -fun auto_arity intermediate def_name pos lthy = - let - val (goal, thm_name, attribs, def_thm, vs) = arity_goal intermediate def_name lthy - val intermediate_text = if intermediate then "intermediate" else "" - val help = "You can manually prove the arity_theorem by typing:\n" - ^ "manual_arity " ^ intermediate_text ^ " for \"" ^ def_name ^ "\"\n" - val thm = prove_arity [def_thm] goal lthy - handle ERROR s => help ^ "\n\n" ^ s |> Exn.reraise o ERROR - val thm = Utils.fix_vars thm (map Utils.freeName vs) lthy - in - Local_Theory.note ((Binding.name thm_name, attribs), [thm]) lthy |> Utils.display "theorem" pos - end - -fun prove_tc_form goal thms ctxt = - Goal.prove ctxt [] [] goal (K (rewrite_goal_tac ctxt thms 1 THEN auto_tac ctxt)) - -fun prove_sats_tm thm_auto thms goal ctxt = - let - val ctxt' = ctxt |> Simplifier.add_simp (hd thm_auto) - in - Goal.prove ctxt [] [] goal - (K (rewrite_goal_tac ctxt thms 1 THEN PARALLEL_ALLGOALS (asm_simp_tac ctxt'))) - end - -fun prove_sats_iff goal ctxt = Goal.prove ctxt [] [] goal (K (asm_simp_tac ctxt 1)) - -fun is_mem (@{const mem} $ _ $ _) = true - | is_mem _ = false - -fun pre_synth_thm_sats term set env vars vs lthy = - let - val (_, tm, ctxt1) = Utils.thm_concl_tm lthy term - val (thm_refs, ctxt2) = Variable.import true [Proof_Context.get_thm lthy term] ctxt1 |>> #2 - val vs' = map (Thm.term_of o #2) vs - val vars' = map (Thm.term_of o #2) vars - val r_tm = tm |> Utils.dest_lhs_def |> fold (op $`) vs' - val sats = @{const apply} $ (@{const satisfies} $ set $ r_tm) $ env - val sats' = @{const IFOL.eq(i)} $ sats $ (@{const succ} $ @{const zero}) - in - { vars = vars' - , vs = vs' - , sats = sats' - , thm_refs = thm_refs - , lthy = ctxt2 - , env = env - , set = set - } - end - -fun synth_thm_sats_gen name lhs hyps pos attribs aux_funs environment lthy = - let - val ctxt = (#prepare_ctxt aux_funs) lthy - val ctxt = Utils.add_to_context (Utils.freeName (#set environment)) ctxt - val (new_vs, ctxt') = (#create_variables aux_funs) (#vs environment, ctxt) - val new_hyps = (#create_hyps aux_funs) (#vs environment, new_vs) - val concl = (#make_concl aux_funs) (lhs, #sats environment, new_vs) - val g_iff = Logic.list_implies (new_hyps @ hyps, Utils.tp concl) - val thm = (#prover aux_funs) g_iff ctxt' - val thm = Utils.fix_vars thm (map Utils.freeName ((#vars environment) @ new_vs)) lthy - in - Local_Theory.note ((name, attribs), [thm]) lthy |> Utils.display "theorem" pos - end - -fun synth_thm_sats_iff def_name lhs hyps pos environment = - let - val subst = Utils.zip_with (I *** I) (#vs environment) - fun subst_nth (@{const "nth"} $ v $ _) new_vs = AList.lookup (op =) (subst new_vs) v |> the - | subst_nth (t1 $ t2) new_vs = (subst_nth t1 new_vs) $ (subst_nth t2 new_vs) - | subst_nth (Abs (v, ty, t)) new_vs = Abs (v, ty, subst_nth t new_vs) - | subst_nth t _ = t - val name = Binding.name (def_name ^ "_iff_sats") - val iff_sats_attrib = @{attributes [iff_sats]} - val aux_funs = { prepare_ctxt = fold Utils.add_to_context (map Utils.freeName (#vs environment)) - , create_variables = fn (vs, ctxt) => Variable.variant_fixes (map Utils.freeName vs) ctxt |>> map Utils.var_i - , create_hyps = fn (vs, new_vs) => Utils.zip_with (fn (v, nv) => Utils.eq_ (Utils.nth_ v (#env environment)) nv) vs new_vs |> map Utils.tp - , make_concl = fn (lhs, rhs, new_vs) => @{const IFOL.iff} $ (subst_nth lhs new_vs) $ rhs - , prover = prove_sats_iff - } - in - synth_thm_sats_gen name lhs hyps pos iff_sats_attrib aux_funs environment - end - -fun synth_thm_sats_fm def_name lhs hyps pos thm_auto environment = - let - val name = Binding.name ("sats_" ^ def_name ^ "_fm") - val simp_attrib = @{attributes [simp]} - val aux_funs = { prepare_ctxt = I - , create_variables = K [] *** I - , create_hyps = K [] - , make_concl = fn (rhs, lhs, _) => @{const IFOL.iff} $ lhs $ rhs - , prover = prove_sats_tm thm_auto (#thm_refs environment) - } - in - synth_thm_sats_gen name lhs hyps pos simp_attrib aux_funs environment - end - -fun synth_thm_tc def_name term hyps vars pos lthy = - let - val (_,tm,ctxt1) = Utils.thm_concl_tm lthy term - val (thm_refs,ctxt2) = Variable.import true [Proof_Context.get_thm lthy term] ctxt1 |>> #2 - val vars' = map (Thm.term_of o #2) vars - val tc_attrib = @{attributes [TC]} - val r_tm = tm |> Utils.dest_lhs_def |> fold (op $`) vars' - val concl = @{const mem} $ r_tm $ @{const formula} - val g = Logic.list_implies(hyps, Utils.tp concl) - val thm = prove_tc_form g thm_refs ctxt2 - val name = Binding.name (def_name ^ "_fm_type") - val thm = Utils.fix_vars thm (map Utils.freeName vars') ctxt2 - in - Local_Theory.note ((name, tc_attrib), [thm]) lthy |> Utils.display "theorem" pos - end - -fun synthetic_def def_name thm_ref pos tc auto thy = - let - val thm = Proof_Context.get_thm thy thm_ref - val thm_vars = rev (Term.add_vars (Thm.full_prop_of thm) []) - val (((_,inst),thm_tms),_) = Variable.import true [thm] thy - val vars = map (fn v => (v, the (Vars.lookup inst v))) thm_vars - val (tm,hyps) = thm_tms |> hd |> Thm.concl_of &&& Thm.prems_of - val (lhs,rhs) = tm |> Utils.dest_iff_tms o Utils.dest_trueprop - val ((set,t),env) = rhs |> Utils.dest_sats_frm - fun relevant ts (@{const mem} $ t $ _) = - (not (t = @{term "0"})) andalso - (not (Term.is_Free t) orelse member (op =) ts (t |> Term.dest_Free |> #1)) - | relevant _ _ = false - val t_vars = sort_strings (Term.add_free_names t []) - val vs = filter (Ord_List.member String.compare t_vars o #1 o #1 o #1) vars - val at = fold_rev (lambda o Thm.term_of o #2) vs t - val hyps' = filter (relevant t_vars o Utils.dest_trueprop) hyps - val def_attrs = @{attributes [fm_definitions]} - in - Local_Theory.define ((Binding.name (def_name ^ "_fm"), NoSyn), - ((Binding.name (def_name ^ "_fm_def"), def_attrs), at)) thy - |>> (#2 #> I *** single) |> Utils.display "theorem" pos |> - (if tc then synth_thm_tc def_name (def_name ^ "_fm_def") hyps' vs pos else I) |> - (if auto then - pre_synth_thm_sats (def_name ^ "_fm_def") set env vars vs - #> I &&& #lthy - #> #1 &&& uncurry (synth_thm_sats_fm def_name lhs hyps pos thm_tms) - #> uncurry (synth_thm_sats_iff def_name lhs hyps pos) - else I) - end - -fun prove_schematic thms goal ctxt = - let - val rules = Named_Theorems.get ctxt \<^named_theorems>\iff_sats\ - in - Goal.prove ctxt [] [] goal - (K (rewrite_goal_tac ctxt thms 1 THEN REPEAT1 (CHANGED (resolve_tac ctxt rules 1 ORELSE asm_simp_tac ctxt 1)))) - end - -val valid_assumptions = [ ("nonempty", Utils.mem_ @{term "0"}) - ] - -fun pre_schematic_def target assuming lthy = -let - val thm = Proof_Context.get_thm lthy (target ^ "_def") - val (vars, tm, ctxt1) = Utils.thm_concl_tm lthy (target ^ "_def") - val (const, tm) = tm |> Utils.dest_eq_tms' o Utils.dest_trueprop |>> #1 o strip_comb - val t_vars = sort_strings (Term.add_free_names tm []) - val vs = List.filter (#1 #> #1 #> #1 #> Ord_List.member String.compare t_vars) vars - |> List.filter ((curry op = @{typ "i"}) o #2 o #1) - |> List.map (Utils.var_i o #1 o #1 o #1) - fun first_lambdas (Abs (body as (_, ty, _))) = - if ty = @{typ "i"} - then (op ::) (Utils.dest_abs body |>> Utils.var_i ||> first_lambdas) - else Utils.dest_abs body |> first_lambdas o #2 - | first_lambdas _ = [] - val vs = vs @ (first_lambdas tm) - val ctxt1' = fold Utils.add_to_context (map Utils.freeName vs) ctxt1 - val (set, ctxt2) = Variable.variant_fixes ["A"] ctxt1' |>> Utils.var_i o hd - val class = @{const "setclass"} $ set - val (env, ctxt3) = Variable.variant_fixes ["env"] ctxt2 |>> Utils.var_i o hd - val assumptions = filter (member (op =) assuming o #1) valid_assumptions |> map #2 - val hyps = (List.map (fn v => Utils.tp (Utils.mem_ v Utils.nat_)) vs) - @ [Utils.tp (Utils.mem_ env (Utils.list_ set))] - @ Utils.zip_with (fn (f,x) => Utils.tp (f x)) assumptions (replicate (length assumptions) set) - val args = class :: map (fn v => Utils.nth_ v env) vs - val (fm_name, ctxt4) = Variable.variant_fixes ["fm"] ctxt3 |>> hd - val fm_type = fold (K (fn acc => Type ("fun", [@{typ "i"}, acc]))) vs @{typ "i"} - val fm = Var ((fm_name, 0), fm_type) - val lhs = fold (op $`) args const - val fm_app = fold (op $`) vs fm - val sats = @{const apply} $ (@{const satisfies} $ set $ fm_app) $ env - val rhs = @{const IFOL.eq(i)} $ sats $ (@{const succ} $ @{const zero}) - val concl = @{const "IFOL.iff"} $ lhs $ rhs - val schematic = Logic.list_implies (hyps, Utils.tp concl) - in - (schematic, ctxt4, thm, set, env, vs) - end - -fun str_join _ [] = "" - | str_join _ [s] = s - | str_join c (s :: ss) = s ^ c ^ (str_join c ss) - -fun schematic_def def_name target assuming pos lthy = - let - val (schematic, ctxt, thm, set, env, vs) = pre_schematic_def target assuming lthy - val assuming_text = if null assuming then "" else "assuming " ^ (map (fn s => "\"" ^ s ^ "\"") assuming |> str_join " ") - val help = "You can manually prove the schematic_goal by typing:\n" - ^ "manual_schematic [sch_name] for \"" ^ target ^ "\"" ^ assuming_text ^"\n" - ^ "And then complete the synthesis with:\n" - ^ "synthesize \"" ^ target ^ "\" from_schematic [sch_name]\n" - ^ "In both commands, \sch_name\ must be the same and, if ommited, will be defaulted to sats_" ^ target ^ "_fm_auto.\n" - ^ "You can also try adding new assumptions and/or synthetizing definitions of sub-terms." - val thm = prove_schematic [thm] schematic ctxt - handle ERROR s => help ^ "\n\n" ^ s |> Exn.reraise o ERROR - val thm = Utils.fix_vars thm (map Utils.freeName (set :: env :: vs)) lthy - in - Local_Theory.note ((Binding.name def_name, []), [thm]) lthy |> Utils.display "theorem" pos - end - -fun schematic_synthetic_def def_name target assuming pos tc auto = - (synthetic_def def_name ("sats_" ^ def_name ^ "_fm_auto") pos tc auto) - o (schematic_def ("sats_" ^ def_name ^ "_fm_auto") target assuming pos) - -fun manual_schematic def_name target assuming pos lthy = - let - val (schematic, _, _, _, _, _) = pre_schematic_def target assuming lthy - in - Proof.theorem NONE (fn thmss => Utils.display "theorem" pos - o Local_Theory.note ((Binding.name def_name, []), hd thmss)) - [[(schematic, [])]] lthy - end -\ - -ML\ -local - val simple_from_schematic_synth_constdecl = - Parse.string --| (Parse.$$$ "from_schematic") - >> (fn bndg => synthetic_def bndg ("sats_" ^ bndg ^ "_fm_auto")) - - val full_from_schematic_synth_constdecl = - Parse.string -- ((Parse.$$$ "from_schematic" |-- Parse.thm)) - >> (fn (bndg,thm) => synthetic_def bndg (#1 (thm |>> Facts.ref_name))) - - val full_from_definition_synth_constdecl = - Parse.string -- ((Parse.$$$ "from_definition" |-- Parse.string)) -- (Scan.optional (Parse.$$$ "assuming" |-- Scan.repeat Parse.string) []) - >> (fn ((bndg,target), assuming) => schematic_synthetic_def bndg target assuming) - - val simple_from_definition_synth_constdecl = - Parse.string -- (Parse.$$$ "from_definition" |-- (Scan.optional (Parse.$$$ "assuming" |-- Scan.repeat Parse.string)) []) - >> (fn (bndg, assuming) => schematic_synthetic_def bndg bndg assuming) - - val synth_constdecl = - Parse.position (full_from_schematic_synth_constdecl || simple_from_schematic_synth_constdecl - || full_from_definition_synth_constdecl - || simple_from_definition_synth_constdecl) - - val full_schematic_decl = - Parse.string -- ((Parse.$$$ "for" |-- Parse.string)) -- (Scan.optional (Parse.$$$ "assuming" |-- Scan.repeat Parse.string) []) - - val simple_schematic_decl = - (Parse.$$$ "for" |-- Parse.string >> (fn name => "sats_" ^ name ^ "_fm_auto") &&& I) -- (Scan.optional (Parse.$$$ "assuming" |-- Scan.repeat Parse.string) []) - - val schematic_decl = Parse.position (full_schematic_decl || simple_schematic_decl) - - val _ = - Outer_Syntax.local_theory \<^command_keyword>\synthesize\ "ML setup for synthetic definitions" - (synth_constdecl >> (fn (f,p) => f p true true)) - - val _ = - Outer_Syntax.local_theory \<^command_keyword>\synthesize_notc\ "ML setup for synthetic definitions" - (synth_constdecl >> (fn (f,p) => f p false false)) - - val _ = Outer_Syntax.local_theory \<^command_keyword>\generate_schematic\ "ML setup for schematic goals" - (schematic_decl >> (fn (((bndg,target), assuming),p) => schematic_def bndg target assuming p)) - - val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\manual_schematic\ "ML setup for schematic goals" - (schematic_decl >> (fn (((bndg,target), assuming),p) => manual_schematic bndg target assuming p)) - - val arity_parser = Parse.position ((Scan.option (Parse.$$$ "intermediate") >> is_some) -- (Parse.$$$ "for" |-- Parse.string)) - - val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\manual_arity\ "ML setup for arities" - (arity_parser >> (fn ((intermediate, target), pos) => manual_arity intermediate target pos)) - - val _ = Outer_Syntax.local_theory \<^command_keyword>\arity_theorem\ "ML setup for arities" - (arity_parser >> (fn ((intermediate, target), pos) => auto_arity intermediate target pos)) - -in - -end -\ +ML_file\Synthetic_Definition.ml\ text\The \<^ML>\synthetic_def\ function extracts definitions from schematic goals. A new definition is added to the context. \ (* example of use *) (* schematic_goal mem_formula_ex : assumes "m\nat" "n\ nat" "env \ list(M)" shows "nth(m,env) \ nth(n,env) \ sats(M,?frm,env)" by (insert assms ; (rule sep_rules empty_iff_sats cartprod_iff_sats | simp del:sats_cartprod_fm)+) synthesize "\" from_schematic mem_formula_ex *) end diff --git a/thys/Transitive_Models/ZF_Library_Relative.thy b/thys/Transitive_Models/ZF_Library_Relative.thy --- a/thys/Transitive_Models/ZF_Library_Relative.thy +++ b/thys/Transitive_Models/ZF_Library_Relative.thy @@ -1,1072 +1,1077 @@ section\Library of basic $\mathit{ZF}$ results\label{sec:zf-lib}\ theory ZF_Library_Relative imports - Aleph_Relative\ \must be before Cardinal\_AC\_Relative!\ + Aleph_Relative \ \must be before \<^theory>\Transitive_Models.Cardinal_AC_Relative\!\ Cardinal_AC_Relative FiniteFun_Relative begin +locale M_Pi_assumption_repl = M_Pi_replacement + + fixes A f + assumes A_in_M: "M(A)" and + f_repl : "lam_replacement(M,f)" and + f_closed : "\x[M]. M(f(x))" +begin + +sublocale M_Pi_assumptions M A f + using Pi_replacement Pi_separation A_in_M f_repl f_closed + lam_replacement_Sigfun[THEN lam_replacement_imp_strong_replacement] + by unfold_locales (simp_all add:transM[of _ A]) + +end \ \\<^locale>\M_Pi_assumption_repl\\ + no_notation sum (infixr \+\ 65) notation oadd (infixl \+\ 65) lemma (in M_cardinal_arith_jump) csucc_rel_cardinal_rel: assumes "Ord(\)" "M(\)" shows "(|\|\<^bsup>M\<^esup>\<^sup>+)\<^bsup>M\<^esup> = (\\<^sup>+)\<^bsup>M\<^esup>" -proof (intro le_anti_sym)\ \show both inequalities\ +proof (intro le_anti_sym) \ \we show both inequalities\ from assms have hips:"M((\\<^sup>+)\<^bsup>M\<^esup>)" "Ord((\\<^sup>+)\<^bsup>M\<^esup>)" "\ < (\\<^sup>+)\<^bsup>M\<^esup>" using Card_rel_csucc_rel[THEN Card_rel_is_Ord] csucc_rel_basic by simp_all then show "(|\|\<^bsup>M\<^esup>\<^sup>+)\<^bsup>M\<^esup> \ (\\<^sup>+)\<^bsup>M\<^esup>" using Ord_cardinal_rel_le[THEN lt_trans1] Card_rel_csucc_rel unfolding csucc_rel_def by (rule_tac Least_antitone) (assumption, simp_all add:assms) from assms have "\ < L" if "Card\<^bsup>M\<^esup>(L)" "|\|\<^bsup>M\<^esup> < L" "M(L)" for L - using (* Card_rel_le_iff[THEN iffD1, THEN le_trans, of \ _ L] *) that - Card_rel_is_Ord leI Card_rel_le_iff[of \ L] + using that Card_rel_is_Ord leI Card_rel_le_iff[of \ L] by (rule_tac ccontr, auto dest:not_lt_imp_le) (fast dest: le_imp_not_lt) with hips show "(\\<^sup>+)\<^bsup>M\<^esup> \ (|\|\<^bsup>M\<^esup>\<^sup>+)\<^bsup>M\<^esup>" using Ord_cardinal_rel_le[THEN lt_trans1] Card_rel_csucc_rel unfolding csucc_rel_def by (rule_tac Least_antitone) (assumption, auto simp add:assms) qed lemma (in M_cardinal_arith_jump) csucc_rel_le_mono: assumes "\ \ \" "M(\)" "M(\)" shows "(\\<^sup>+)\<^bsup>M\<^esup> \ (\\<^sup>+)\<^bsup>M\<^esup>" proof (cases "\ = \") case True with assms show ?thesis using Card_rel_csucc_rel [THEN Card_rel_is_Ord] by simp next case False with assms have "\ < \" using le_neq_imp_lt by simp - show ?thesis\ \by way of contradiction\ + show ?thesis \ \by way of contradiction\ proof (rule ccontr) assume "\ (\\<^sup>+)\<^bsup>M\<^esup> \ (\\<^sup>+)\<^bsup>M\<^esup>" with assms have "(\\<^sup>+)\<^bsup>M\<^esup> < (\\<^sup>+)\<^bsup>M\<^esup>" using Card_rel_csucc_rel[THEN Card_rel_is_Ord] le_Ord2 lt_Ord by (intro not_le_iff_lt[THEN iffD1]) auto with assms have "(\\<^sup>+)\<^bsup>M\<^esup> \ |\|\<^bsup>M\<^esup>" using le_Ord2[THEN Card_rel_csucc_rel, of \ \] Card_rel_lt_csucc_rel_iff[of "(\\<^sup>+)\<^bsup>M\<^esup>" "|\|\<^bsup>M\<^esup>", THEN iffD1] csucc_rel_cardinal_rel[OF lt_Ord] leI[of "(\\<^sup>+)\<^bsup>M\<^esup>" "(\\<^sup>+)\<^bsup>M\<^esup>"] by simp with assms have "(\\<^sup>+)\<^bsup>M\<^esup> \ \" using Ord_cardinal_rel_le[OF lt_Ord] le_trans by auto with assms have "\ < \" using csucc_rel_basic le_Ord2[THEN Card_rel_csucc_rel, of \ \] Card_rel_is_Ord le_Ord2 by (rule_tac j="(\\<^sup>+)\<^bsup>M\<^esup>" in lt_trans2) simp_all with \\ < \\ show "False" using le_imp_not_lt leI by blast qed qed lemma (in M_cardinal_AC) cardinal_rel_succ_not_0: "|A|\<^bsup>M\<^esup> = succ(n) \ M(A) \ M(n) \ A \ 0" by auto (* "Finite_to_one(X,Y) \ {f:X\Y. \y\Y. Finite({x\X . f`x = y})}" *) reldb_add functional "Finite" "Finite" \ \wrongly done? Finite is absolute\ relativize functional "Finite_to_one" "Finite_to_one_rel" external - (* reldb_add relational "Finite" "is_Finite" \ \don't have is_Finite yet\ -relationalize "Finite_to_one_rel" "is_Finite_to_one" *) +(* reldb_add relational "Finite" "is_Finite" *) \ \don't have \<^term>\is_Finite\ yet\ +(* relationalize "Finite_to_one_rel" "is_Finite_to_one" *) notation Finite_to_one_rel (\Finite'_to'_one\<^bsup>_\<^esup>'(_,_')\) abbreviation Finite_to_one_r_set :: "[i,i,i] \ i" (\Finite'_to'_one\<^bsup>_\<^esup>'(_,_')\) where "Finite_to_one\<^bsup>M\<^esup>(X,Y) \ Finite_to_one_rel(##M,X,Y)" locale M_ZF_library = M_aleph + M_FiniteFun begin lemma Finite_Collect_imp: "Finite({x\X . Q(x)}) \ Finite({x\X . M(x) \ Q(x)})" (is "Finite(?A) \ Finite(?B)") using subset_Finite[of ?B ?A] by auto lemma Finite_to_one_relI[intro]: assumes "f:X\\<^bsup>M\<^esup>Y" "\y. y\Y \ Finite({x\X . f`x = y})" and types:"M(f)" "M(X)" "M(Y)" shows "f \ Finite_to_one\<^bsup>M\<^esup>(X,Y)" using assms Finite_Collect_imp unfolding Finite_to_one_rel_def by (simp) lemma Finite_to_one_relI'[intro]: assumes "f:X\\<^bsup>M\<^esup>Y" "\y. y\Y \ Finite({x\X . M(x) \ f`x = y})" and types:"M(f)" "M(X)" "M(Y)" shows "f \ Finite_to_one\<^bsup>M\<^esup>(X,Y)" using assms unfolding Finite_to_one_rel_def by (simp) lemma Finite_to_one_relD[dest]: "f \ Finite_to_one\<^bsup>M\<^esup>(X,Y) \f:X\\<^bsup>M\<^esup>Y" "f \ Finite_to_one\<^bsup>M\<^esup>(X,Y) \ y\Y \ M(Y) \ Finite({x\X . M(x) \ f`x = y})" unfolding Finite_to_one_rel_def by simp_all lemma Diff_bij_rel: assumes "\A\F. X \ A" and types: "M(F)" "M(X)" shows "(\A\F. A-X) \ bij\<^bsup>M\<^esup>(F, {A-X. A\F})" - using assms def_inj_rel def_surj_rel unfolding bij_rel_def - apply (auto) - apply (subgoal_tac "M(\A\F. A - X)" "M({A - X . A \ F})") - apply (auto simp add:mem_function_space_rel_abs) - apply (rule_tac lam_type, auto) - prefer 4 - apply (subgoal_tac "M(\A\F. A - X)" "M({A - X . A \ F})") - apply(tactic \distinct_subgoals_tac\) - apply (auto simp add:mem_function_space_rel_abs) - apply (rule_tac lam_type, auto) prefer 3 - apply (subst subset_Diff_Un[of X]) - apply auto proof - - from types - show "M({A - X . A \ F})" - using diff_replacement - by (rule_tac RepFun_closed) (auto dest:transM[of _ F]) - from types - show "M(\A\F. A - X)" - using Pair_diff_replacement - by (rule_tac lam_closed, auto dest:transM) + from assms + have "A - X = C - X \ A = C" if "A\F" "C\F" for A C + using that subset_Diff_Un[of X A] subset_Diff_Un[of X C] + by simp + moreover + note assms + moreover from this + have "M({A-X . A\F})" "(\A\F. A-X) \ F \ {A-X. A\F}" (is "?H \ _") "M(\A\F. A-X)" + using lam_type lam_replacement_Diff[THEN [5] lam_replacement_hcomp2] lam_replacement_constant + lam_replacement_identity transM[of _ F] lam_replacement_imp_strong_replacement RepFun_closed + lam_replacement_imp_lam_closed + by (simp,rule_tac lam_type,auto) + ultimately + show ?thesis + using assms def_inj_rel def_surj_rel function_space_rel_char + unfolding bij_rel_def + by auto qed lemma function_space_rel_nonempty: assumes "b\B" and types: "M(B)" "M(A)" shows "(\x\A. b) : A \\<^bsup>M\<^esup> B" proof - note assms moreover from this have "M(\x\A. b)" using tag_replacement by (rule_tac lam_closed, auto dest:transM) ultimately show ?thesis by (simp add:mem_function_space_rel_abs) qed lemma mem_function_space_rel: assumes "f \ A \\<^bsup>M\<^esup> y" "M(A)" "M(y)" shows "f \ A \ y" using assms function_space_rel_char by simp lemmas range_fun_rel_subset_codomain = range_fun_subset_codomain[OF mem_function_space_rel] end \ \\<^locale>\M_ZF_library\\ context M_Pi_assumptions begin lemma mem_Pi_rel: "f \ Pi\<^bsup>M\<^esup>(A,B) \ f \ Pi(A, B)" using trans_closed mem_Pi_rel_abs by force lemmas Pi_rel_rangeD = Pi_rangeD[OF mem_Pi_rel] lemmas rel_apply_Pair = apply_Pair[OF mem_Pi_rel] lemmas rel_apply_rangeI = apply_rangeI[OF mem_Pi_rel] lemmas Pi_rel_range_eq = Pi_range_eq[OF mem_Pi_rel] lemmas Pi_rel_vimage_subset = Pi_vimage_subset[OF mem_Pi_rel] end \ \\<^locale>\M_Pi_assumptions\\ context M_ZF_library begin lemmas rel_apply_in_range = apply_in_codomain_Ord[OF _ _ mem_function_space_rel] lemmas rel_range_eq_image = ZF_Library.range_eq_image[OF mem_function_space_rel] lemmas rel_Image_sub_codomain = Image_sub_codomain[OF mem_function_space_rel] lemma rel_inj_to_Image: "\f:A\\<^bsup>M\<^esup>B; f \ inj\<^bsup>M\<^esup>(A,B); M(A); M(B)\ \ f \ inj\<^bsup>M\<^esup>(A,f``A)" using inj_to_Image[OF mem_function_space_rel mem_inj_rel] transM[OF _ function_space_rel_closed] by simp lemma inj_rel_imp_surj_rel: fixes f b defines [simp]: "ifx(x) \ if x\range(f) then converse(f)`x else b" assumes "f \ inj\<^bsup>M\<^esup>(B,A)" "b\B" and types: "M(f)" "M(B)" "M(A)" shows "(\x\A. ifx(x)) \ surj\<^bsup>M\<^esup>(A,B)" proof - from types and \b\B\ have "M(\x\A. ifx(x))" using ifx_replacement by (rule_tac lam_closed) (auto dest:transM) with assms(2-) show ?thesis using inj_imp_surj mem_surj_abs by simp qed lemma function_space_rel_disjoint_Un: assumes "f \ A\\<^bsup>M\<^esup>B" "g \ C\\<^bsup>M\<^esup>D" "A \ C = 0" and types:"M(A)" "M(B)" "M(C)" "M(D)" shows "f \ g \ (A \ C)\\<^bsup>M\<^esup> (B \ D)" using assms fun_Pi_disjoint_Un[OF mem_function_space_rel mem_function_space_rel, OF assms(1) _ _ assms(2)] function_space_rel_char by auto lemma restrict_eq_imp_Un_into_function_space_rel: assumes "f \ A\\<^bsup>M\<^esup>B" "g \ C\\<^bsup>M\<^esup>D" "restrict(f, A \ C) = restrict(g, A \ C)" and types:"M(A)" "M(B)" "M(C)" "M(D)" shows "f \ g \ (A \ C)\\<^bsup>M\<^esup> (B \ D)" using assms restrict_eq_imp_Un_into_Pi[OF mem_function_space_rel mem_function_space_rel, OF assms(1) _ _ assms(2)] function_space_rel_char by auto lemma lepoll_relD[dest]: "A \\<^bsup>M\<^esup> B \ \f[M]. f \ inj\<^bsup>M\<^esup>(A, B)" unfolding lepoll_rel_def . \ \Should the assumptions be on \<^term>\f\ or on \<^term>\A\ and \<^term>\B\? Should BOTH be intro rules?\ lemma lepoll_relI[intro]: "f \ inj\<^bsup>M\<^esup>(A, B) \ M(f) \ A \\<^bsup>M\<^esup> B" unfolding lepoll_rel_def by blast lemma eqpollD[dest]: "A \\<^bsup>M\<^esup> B \ \f[M]. f \ bij\<^bsup>M\<^esup>(A, B)" unfolding eqpoll_rel_def . \ \Same as @{thm lepoll_relI}\ lemma bij_rel_imp_eqpoll_rel[intro]: "f \ bij\<^bsup>M\<^esup>(A,B) \ M(f) \ A \\<^bsup>M\<^esup> B" unfolding eqpoll_rel_def by blast lemma restrict_bij_rel:\ \Unused\ assumes "f \ inj\<^bsup>M\<^esup>(A,B)" "C\A" and types:"M(A)" "M(B)" "M(C)" shows "restrict(f,C)\ bij\<^bsup>M\<^esup>(C, f``C)" using assms restrict_bij inj_rel_char bij_rel_char by auto lemma range_of_subset_eqpoll_rel: assumes "f \ inj\<^bsup>M\<^esup>(X,Y)" "S \ X" and types:"M(X)" "M(Y)" "M(S)" shows "S \\<^bsup>M\<^esup> f `` S" using assms restrict_bij bij_rel_char trans_inj_rel_closed[OF \f \ inj\<^bsup>M\<^esup>(X,Y)\] unfolding eqpoll_rel_def by (rule_tac x="restrict(f,S)" in rexI) auto lemmas inj_rel_is_fun = inj_is_fun[OF mem_inj_rel] lemma inj_rel_bij_rel_range: "f \ inj\<^bsup>M\<^esup>(A,B) \ M(A) \ M(B) \ f \ bij\<^bsup>M\<^esup>(A,range(f))" using bij_rel_char inj_rel_char inj_bij_range by force lemma bij_rel_is_inj_rel: "f \ bij\<^bsup>M\<^esup>(A,B) \ M(A) \ M(B) \ f \ inj\<^bsup>M\<^esup>(A,B)" unfolding bij_rel_def by simp lemma inj_rel_weaken_type: "[| f \ inj\<^bsup>M\<^esup>(A,B); B\D; M(A); M(B); M(D) |] ==> f \ inj\<^bsup>M\<^esup>(A,D)" using inj_rel_char inj_rel_is_fun inj_weaken_type by auto lemma bij_rel_converse_bij_rel [TC]: "f \ bij\<^bsup>M\<^esup>(A,B) \ M(A) \ M(B) ==> converse(f): bij\<^bsup>M\<^esup>(B,A)" using bij_rel_char by force lemma bij_rel_is_fun_rel: "f \ bij\<^bsup>M\<^esup>(A,B) \ M(A) \ M(B) \ f \ A\\<^bsup>M\<^esup>B" using bij_rel_char mem_function_space_rel_abs bij_is_fun by simp lemmas bij_rel_is_fun = bij_rel_is_fun_rel[THEN mem_function_space_rel] lemma comp_bij_rel: "g \ bij\<^bsup>M\<^esup>(A,B) \ f \ bij\<^bsup>M\<^esup>(B,C) \ M(A) \ M(B) \ M(C) \ (f O g) \ bij\<^bsup>M\<^esup>(A,C)" using bij_rel_char comp_bij by force lemma inj_rel_converse_fun: "f \ inj\<^bsup>M\<^esup>(A,B) \ M(A) \ M(B) \ converse(f) \ range(f)\\<^bsup>M\<^esup>A" proof - assume "f \ inj\<^bsup>M\<^esup>(A,B)" "M(A)" "M(B)" then have "M(f)" "M(converse(f))" "M(range(f))" "f\inj(A,B)" using inj_rel_char converse_closed range_closed by auto then show ?thesis using inj_converse_inj function_space_rel_char inj_is_fun \M(A)\ by auto qed lemma fg_imp_bijective_rel: assumes "f \ A \\<^bsup>M\<^esup>B" "g \ B\\<^bsup>M\<^esup>A" "f O g = id(B)" "g O f = id(A)" "M(A)" "M(B)" shows "f \ bij\<^bsup>M\<^esup>(A,B)" using assms mem_bij_abs fg_imp_bijective mem_function_space_rel_abs[THEN iffD2] function_space_rel_char by auto end \ \\<^locale>\M_ZF_library\\ -(************* Discipline for cexp ****************) +subsection\Discipline for \<^term>\cexp\\ + relativize functional "cexp" "cexp_rel" external relationalize "cexp_rel" "is_cexp" context M_ZF_library begin is_iff_rel for "cexp" using is_cardinal_iff is_function_space_iff unfolding cexp_rel_def is_cexp_def by (simp) rel_closed for "cexp" unfolding cexp_rel_def by simp end \ \\<^locale>\M_ZF_library\\ synthesize "is_cexp" from_definition assuming "nonempty" notation is_cexp_fm (\\_\<^bsup>\_\<^esup> is _\\) arity_theorem for "is_cexp_fm" abbreviation cexp_r :: "[i,i,i\o] \ i" (\_\<^bsup>\_,_\<^esup>\) where "cexp_r(x,y,M) \ cexp_rel(M,x,y)" abbreviation cexp_r_set :: "[i,i,i] \ i" (\_\<^bsup>\_,_\<^esup>\) where "cexp_r_set(x,y,M) \ cexp_rel(##M,x,y)" context M_ZF_library begin lemma Card_rel_cexp_rel: "M(\) \ M(\) \ Card\<^bsup>M\<^esup>(\\<^bsup>\\,M\<^esup>)" unfolding cexp_rel_def by simp -\ \Restoring congruence rule, but NOTE: beware\ +\ \Restoring congruence rule\ declare conj_cong[cong] lemma eq_csucc_rel_ord: "Ord(i) \ M(i) \ (i\<^sup>+)\<^bsup>M\<^esup> = (|i|\<^bsup>M\<^esup>\<^sup>+)\<^bsup>M\<^esup>" using Card_rel_lt_iff Least_cong unfolding csucc_rel_def by auto lemma lesspoll_succ_rel: assumes "Ord(\)" "M(\)" shows "\ \\<^bsup>M\<^esup> (\\<^sup>+)\<^bsup>M\<^esup>" using csucc_rel_basic assms lt_Card_rel_imp_lesspoll_rel Card_rel_csucc_rel lepoll_rel_iff_leqpoll_rel by auto lemma lesspoll_rel_csucc_rel: assumes "Ord(\)" and types:"M(\)" "M(d)" shows "d \\<^bsup>M\<^esup> (\\<^sup>+)\<^bsup>M\<^esup> \ d \\<^bsup>M\<^esup> \" proof assume "d \\<^bsup>M\<^esup> (\\<^sup>+)\<^bsup>M\<^esup>" moreover note Card_rel_csucc_rel assms Card_rel_is_Ord moreover from calculation have "Card\<^bsup>M\<^esup>((\\<^sup>+)\<^bsup>M\<^esup>)" "M((\\<^sup>+)\<^bsup>M\<^esup>)" "Ord((\\<^sup>+)\<^bsup>M\<^esup>)" using Card_rel_is_Ord by simp_all moreover from calculation have "d \\<^bsup>M\<^esup> (|\|\<^bsup>M\<^esup>\<^sup>+)\<^bsup>M\<^esup>" "d \\<^bsup>M\<^esup> |d|\<^bsup>M\<^esup>" using eq_csucc_rel_ord[OF _ \M(\)\] lesspoll_rel_imp_eqpoll_rel eqpoll_rel_sym by simp_all moreover from calculation have "|d|\<^bsup>M\<^esup> < (|\|\<^bsup>M\<^esup>\<^sup>+)\<^bsup>M\<^esup>" using lesspoll_cardinal_lt_rel by simp moreover from calculation have "|d|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> |\|\<^bsup>M\<^esup>" using Card_rel_lt_csucc_rel_iff le_imp_lepoll_rel by simp moreover from calculation have "|d|\<^bsup>M\<^esup> \\<^bsup>M\<^esup> \" using Ord_cardinal_rel_eqpoll_rel lepoll_rel_eq_trans by simp ultimately show "d \\<^bsup>M\<^esup> \" using eq_lepoll_rel_trans by simp next from \Ord(\)\ have "\ < (\\<^sup>+)\<^bsup>M\<^esup>" "Card\<^bsup>M\<^esup>((\\<^sup>+)\<^bsup>M\<^esup>)" "M((\\<^sup>+)\<^bsup>M\<^esup>)" using Card_rel_csucc_rel lt_csucc_rel_iff types eq_csucc_rel_ord[OF _ \M(\)\] by simp_all then have "\ \\<^bsup>M\<^esup> (\\<^sup>+)\<^bsup>M\<^esup>" using lt_Card_rel_imp_lesspoll_rel[OF _ \\ <_\] types by simp moreover assume "d \\<^bsup>M\<^esup> \" ultimately have "d \\<^bsup>M\<^esup> (\\<^sup>+)\<^bsup>M\<^esup>" using Card_rel_csucc_rel types lesspoll_succ_rel lepoll_rel_trans \Ord(\)\ by simp moreover from \d \\<^bsup>M\<^esup> \\ \Ord(\)\ have "(\\<^sup>+)\<^bsup>M\<^esup> \\<^bsup>M\<^esup> \" if "d \\<^bsup>M\<^esup> (\\<^sup>+)\<^bsup>M\<^esup>" using eqpoll_rel_sym[OF that] types eq_lepoll_rel_trans[OF _ \d\\<^bsup>M\<^esup>\\] by simp moreover from calculation \\ \\<^bsup>M\<^esup> (\\<^sup>+)\<^bsup>M\<^esup>\ have False if "d \\<^bsup>M\<^esup> (\\<^sup>+)\<^bsup>M\<^esup>" using lesspoll_rel_irrefl[OF _ \M((\\<^sup>+)\<^bsup>M\<^esup>)\] lesspoll_rel_trans1 types that by auto ultimately show "d \\<^bsup>M\<^esup> (\\<^sup>+)\<^bsup>M\<^esup>" unfolding lesspoll_rel_def by auto qed lemma Infinite_imp_nats_lepoll: assumes "Infinite(X)" "n \ \" shows "n \ X" using \n \ \\ proof (induct) case 0 then show ?case using empty_lepollI by simp next case (succ x) show ?case proof - from \Infinite(X)\ and \x \ \\ have "\ (x \ X)" using eqpoll_sym unfolding Finite_def by auto with \x \ X\ obtain f where "f \ inj(x,X)" "f \ surj(x,X)" unfolding bij_def eqpoll_def by auto moreover from this obtain b where "b \ X" "\a\x. f`a \ b" using inj_is_fun unfolding surj_def by auto ultimately have "f \ inj(x,X-{b})" unfolding inj_def by (auto intro:Pi_type) then have "cons(\x, b\, f) \ inj(succ(x), cons(b, X - {b}))" using inj_extend[of f x "X-{b}" x b] unfolding succ_def by (auto dest:mem_irrefl) moreover from \b\X\ have "cons(b, X - {b}) = X" by auto ultimately show "succ(x) \ X" by auto qed qed lemma nepoll_imp_nepoll_rel : assumes "\ x \ X" "M(x)" "M(X)" shows "\ (x \\<^bsup>M\<^esup> X)" using assms unfolding eqpoll_def eqpoll_rel_def by simp lemma Infinite_imp_nats_lepoll_rel: assumes "Infinite(X)" "n \ \" and types: "M(X)" shows "n \\<^bsup>M\<^esup> X" using \n \ \\ proof (induct) case 0 then show ?case using empty_lepoll_relI types by simp next case (succ x) show ?case proof - from \Infinite(X)\ and \x \ \\ have "\ (x \ X)" "M(x)" "M(succ(x))" using eqpoll_sym unfolding Finite_def by auto then have "\ (x \\<^bsup>M\<^esup> X)" using nepoll_imp_nepoll_rel types by simp with \x \\<^bsup>M\<^esup> X\ obtain f where "f \ inj\<^bsup>M\<^esup>(x,X)" "f \ surj\<^bsup>M\<^esup>(x,X)" "M(f)" unfolding bij_rel_def eqpoll_rel_def by auto with \M(X)\ \M(x)\ have "f\surj(x,X)" "f\inj(x,X)" using surj_rel_char by simp_all moreover from this obtain b where "b \ X" "\a\x. f`a \ b" using inj_is_fun unfolding surj_def by auto moreover from this calculation \M(x)\ have "f \ inj(x,X-{b})" "M()" unfolding inj_def using transM[OF _ \M(X)\] by (auto intro:Pi_type) moreover from this have "cons(\x, b\, f) \ inj(succ(x), cons(b, X - {b}))" (is "?g\_") using inj_extend[of f x "X-{b}" x b] unfolding succ_def by (auto dest:mem_irrefl) moreover note \M()\ \M(f)\ \b\X\ \M(X)\ \M(succ(x))\ moreover from this have "M(?g)" "cons(b, X - {b}) = X" by auto moreover from calculation have "?g\inj_rel(M,succ(x),X)" using mem_inj_abs by simp with \M(?g)\ show "succ(x) \\<^bsup>M\<^esup> X" using lepoll_relI by simp qed qed lemma lepoll_rel_imp_lepoll: "A \\<^bsup>M\<^esup> B \ M(A) \ M(B) \ A \ B" unfolding lepoll_rel_def by auto lemma zero_lesspoll_rel: assumes "0<\" "M(\)" shows "0 \\<^bsup>M\<^esup> \" using assms eqpoll_rel_0_iff[THEN iffD1, of \] eqpoll_rel_sym unfolding lesspoll_rel_def lepoll_rel_def by (auto simp add:inj_def) lemma lepoll_rel_nat_imp_Infinite: "\ \\<^bsup>M\<^esup> X \ M(X) \ Infinite(X)" using lepoll_nat_imp_Infinite lepoll_rel_imp_lepoll by simp lemma InfCard_rel_imp_Infinite: "InfCard\<^bsup>M\<^esup>(\) \ M(\) \ Infinite(\)" using le_imp_lepoll_rel[THEN lepoll_rel_nat_imp_Infinite, of \] unfolding InfCard_rel_def by simp lemma lt_surj_rel_empty_imp_Card_rel: assumes "Ord(\)" "\\. \ < \ \ surj\<^bsup>M\<^esup>(\,\) = 0" and types:"M(\)" shows "Card\<^bsup>M\<^esup>(\)" proof - { define min where "min\\ x. \f[M]. f \ bij\<^bsup>M\<^esup>(x,\)" moreover note \Ord(\)\ \M(\)\ moreover assume "|\|\<^bsup>M\<^esup> < \" moreover from calculation have "\f. f \ bij\<^bsup>M\<^esup>(min,\)" using LeastI[of "\i. i \\<^bsup>M\<^esup> \" \, OF eqpoll_rel_refl] unfolding Card_rel_def cardinal_rel_def eqpoll_rel_def by (auto) moreover from calculation have "min < \" using lt_trans1[of min "\ i. M(i) \ (\f[M]. f \ bij\<^bsup>M\<^esup>(i, \))" \] Least_le[of "\i. i \\<^bsup>M\<^esup> \" "|\|\<^bsup>M\<^esup>", OF Ord_cardinal_rel_eqpoll_rel] unfolding Card_rel_def cardinal_rel_def eqpoll_rel_def by (simp) moreover note \min < \ \ surj\<^bsup>M\<^esup>(min,\) = 0\ ultimately have "False" unfolding bij_rel_def by simp } with assms show ?thesis using Ord_cardinal_rel_le[of \] not_lt_imp_le[of "|\|\<^bsup>M\<^esup>" \] le_anti_sym unfolding Card_rel_def by auto qed end \ \\<^locale>\M_ZF_library\\ relativize functional "mono_map" "mono_map_rel" external relationalize "mono_map_rel" "is_mono_map" synthesize "is_mono_map" from_definition assuming "nonempty" notation mono_map_rel (\mono'_map\<^bsup>_\<^esup>'(_,_,_,_')\) abbreviation mono_map_r_set :: "[i,i,i,i,i]\i" (\mono'_map\<^bsup>_\<^esup>'(_,_,_,_')\) where "mono_map\<^bsup>M\<^esup>(a,r,b,s) \ mono_map_rel(##M,a,r,b,s)" context M_ZF_library begin lemma mono_map_rel_char: assumes "M(a)" "M(b)" shows "mono_map\<^bsup>M\<^esup>(a,r,b,s) = {f\mono_map(a,r,b,s) . M(f)}" using assms function_space_rel_char unfolding mono_map_rel_def mono_map_def by auto text\Just a sample of porting results on \<^term>\mono_map\\ lemma mono_map_rel_mono: assumes "f \ mono_map\<^bsup>M\<^esup>(A,r,B,s)" "B \ C" and types:"M(A)" "M(B)" "M(C)" shows "f \ mono_map\<^bsup>M\<^esup>(A,r,C,s)" using assms mono_map_mono mono_map_rel_char by auto lemma nats_le_InfCard_rel: assumes "n \ \" "InfCard\<^bsup>M\<^esup>(\)" shows "n \ \" using assms Ord_is_Transset le_trans[of n \ \, OF le_subset_iff[THEN iffD2]] unfolding InfCard_rel_def Transset_def by simp lemma nat_into_InfCard_rel: assumes "n \ \" "InfCard\<^bsup>M\<^esup>(\)" shows "n \ \" using assms le_imp_subset[of \ \] unfolding InfCard_rel_def by auto lemma Finite_lesspoll_rel_nat: assumes "Finite(x)" "M(x)" shows "x \\<^bsup>M\<^esup> nat" proof - note assms moreover from this obtain n where "n \ \" "M(n)" "x \ n" unfolding Finite_def by auto moreover from calculation obtain f where "f \ bij(x,n)" "f: x-||>n" using Finite_Fin[THEN fun_FiniteFunI, OF _ subset_refl] bij_is_fun unfolding eqpoll_def by auto ultimately have "x\\<^bsup>M\<^esup> n" unfolding eqpoll_rel_def by (auto dest:transM) with assms and \M(n)\ have "n \\<^bsup>M\<^esup> x" using eqpoll_rel_sym by simp moreover note \n\\\ \M(n)\ ultimately show ?thesis using assms eq_lesspoll_rel_trans[OF \x\\<^bsup>M\<^esup> n\ n_lesspoll_rel_nat] by simp qed lemma Finite_cardinal_rel_in_nat [simp]: assumes "Finite(A)" "M(A)" shows "|A|\<^bsup>M\<^esup> \ \" proof - note assms moreover from this obtain n where "n \ \" "M(n)" "A \ n" unfolding Finite_def by auto moreover from calculation obtain f where "f \ bij(A,n)" "f: A-||>n" using Finite_Fin[THEN fun_FiniteFunI, OF _ subset_refl] bij_is_fun unfolding eqpoll_def by auto ultimately have "A \\<^bsup>M\<^esup> n" unfolding eqpoll_rel_def by (auto dest:transM) with assms and \M(n)\ have "n \\<^bsup>M\<^esup> A" using eqpoll_rel_sym by simp moreover note \n\\\ \M(n)\ ultimately show ?thesis using assms Least_le[of "\i. M(i) \ i \\<^bsup>M\<^esup> A" n] lt_trans1[of _ n \, THEN ltD] unfolding cardinal_rel_def Finite_def by (auto dest!:naturals_lt_nat) qed lemma Finite_cardinal_rel_eq_cardinal: assumes "Finite(A)" "M(A)" shows "|A|\<^bsup>M\<^esup> = |A|" proof - \ \Copy-paste from @{thm Finite_cardinal_rel_in_nat}\ note assms moreover from this obtain n where "n \ \" "M(n)" "A \ n" unfolding Finite_def by auto moreover from this have "|A| = n" using cardinal_cong[of A n] nat_into_Card[THEN Card_cardinal_eq, of n] by simp moreover from calculation obtain f where "f \ bij(A,n)" "f: A-||>n" using Finite_Fin[THEN fun_FiniteFunI, OF _ subset_refl] bij_is_fun unfolding eqpoll_def by auto ultimately have "A \\<^bsup>M\<^esup> n" unfolding eqpoll_rel_def by (auto dest:transM) with assms and \M(n)\ \n\\\ have "|A|\<^bsup>M\<^esup> = n" using cardinal_rel_cong[of A n] nat_into_Card_rel[THEN Card_rel_cardinal_rel_eq, of n] by simp with \|A| = n\ show ?thesis by simp qed lemma Finite_imp_cardinal_rel_cons: assumes FA: "Finite(A)" and a: "a\A" and types:"M(A)" "M(a)" shows "|cons(a,A)|\<^bsup>M\<^esup> = succ(|A|\<^bsup>M\<^esup>)" using assms Finite_imp_cardinal_cons Finite_cardinal_rel_eq_cardinal by simp lemma Finite_imp_succ_cardinal_rel_Diff: assumes "Finite(A)" "a \ A" "M(A)" shows "succ(|A-{a}|\<^bsup>M\<^esup>) = |A|\<^bsup>M\<^esup>" proof - from assms have inM: "M(A-{a})" "M(a)" "M(A)" by (auto dest:transM) with \Finite(A)\ have "succ(|A-{a}|\<^bsup>M\<^esup>) = succ(|A-{a}|)" using Diff_subset[THEN subset_Finite, THEN Finite_cardinal_rel_eq_cardinal, of A "{a}"] by simp also from assms have "\ = |A|" using Finite_imp_succ_cardinal_Diff by simp also from assms have "\ = |A|\<^bsup>M\<^esup>" using Finite_cardinal_rel_eq_cardinal by simp finally show ?thesis . qed lemma InfCard_rel_Aleph_rel: notes Aleph_rel_zero[simp] assumes "Ord(\)" and types: "M(\)" shows "InfCard\<^bsup>M\<^esup>(\\<^bsub>\\<^esub>\<^bsup>M\<^esup>)" proof - have "\ (\\<^bsub>\\<^esub>\<^bsup>M\<^esup> \ \)" proof (cases "\=0") case True then show ?thesis using mem_irrefl by auto next case False with assms have "\ \ \\<^bsub>\\<^esub>\<^bsup>M\<^esup>" using Ord_0_lt[of \] ltD by (auto dest:Aleph_rel_increasing) then show ?thesis using foundation by blast qed with assms have "\ (|\\<^bsub>\\<^esub>\<^bsup>M\<^esup>|\<^bsup>M\<^esup> \ \)" using Card_rel_cardinal_rel_eq by auto with assms have "Infinite(\\<^bsub>\\<^esub>\<^bsup>M\<^esup>)" using Ord_Aleph_rel by clarsimp with assms show ?thesis using Inf_Card_rel_is_InfCard_rel by simp qed lemmas Limit_Aleph_rel = InfCard_rel_Aleph_rel[THEN InfCard_rel_is_Limit] bundle Ord_dests = Limit_is_Ord[dest] Card_rel_is_Ord[dest] bundle Aleph_rel_dests = Aleph_rel_cont[dest] bundle Aleph_rel_intros = Aleph_rel_increasing[intro!] bundle Aleph_rel_mem_dests = Aleph_rel_increasing[OF ltI, THEN ltD, dest] lemma f_imp_injective_rel: assumes "f \ A \\<^bsup>M\<^esup> B" "\x\A. d(f ` x) = x" "M(A)" "M(B)" shows "f \ inj\<^bsup>M\<^esup>(A, B)" using assms - apply (simp (no_asm_simp) add: def_inj_rel) - apply (auto intro: subst_context [THEN box_equals]) - done + by (simp (no_asm_simp) add: def_inj_rel,auto intro: subst_context [THEN box_equals]) lemma lam_injective_rel: assumes "\x. x \ A \ c(x) \ B" "\x. x \ A \ d(c(x)) = x" "\x[M]. M(c(x))" "lam_replacement(M,c)" "M(A)" "M(B)" shows "(\x\A. c(x)) \ inj\<^bsup>M\<^esup>(A, B)" using assms function_space_rel_char lam_replacement_iff_lam_closed by (rule_tac d = d in f_imp_injective_rel) (auto simp add: lam_type) lemma f_imp_surjective_rel: assumes "f \ A \\<^bsup>M\<^esup> B" "\y. y \ B \ d(y) \ A" "\y. y \ B \ f ` d(y) = y" "M(A)" "M(B)" shows "f \ surj\<^bsup>M\<^esup>(A, B)" using assms by (simp add: def_surj_rel, blast) lemma lam_surjective_rel: assumes "\x. x \ A \ c(x) \ B" "\y. y \ B \ d(y) \ A" "\y. y \ B \ c(d(y)) = y" "\x[M]. M(c(x))" "lam_replacement(M,c)" "M(A)" "M(B)" shows "(\x\A. c(x)) \ surj\<^bsup>M\<^esup>(A, B)" using assms function_space_rel_char lam_replacement_iff_lam_closed by (rule_tac d = d in f_imp_surjective_rel) (auto simp add: lam_type) lemma lam_bijective_rel: assumes "\x. x \ A \ c(x) \ B" "\y. y \ B \ d(y) \ A" "\x. x \ A \ d(c(x)) = x" "\y. y \ B \ c(d(y)) = y" "\x[M]. M(c(x))" "lam_replacement(M,c)" "M(A)" "M(B)" shows "(\x\A. c(x)) \ bij\<^bsup>M\<^esup>(A, B)" using assms - apply (unfold bij_rel_def) - apply (blast intro!: lam_injective_rel lam_surjective_rel) - done + unfolding bij_rel_def + by (blast intro!: lam_injective_rel lam_surjective_rel) lemma function_space_rel_eqpoll_rel_cong: assumes "A \\<^bsup>M\<^esup> A'" "B \\<^bsup>M\<^esup> B'" "M(A)" "M(A')" "M(B)" "M(B')" shows "A \\<^bsup>M\<^esup> B \\<^bsup>M\<^esup> A' \\<^bsup>M\<^esup> B'" proof - from assms(1)[THEN eqpoll_rel_sym] assms(2) assms lam_type obtain f g where "f \ bij\<^bsup>M\<^esup>(A',A)" "g \ bij\<^bsup>M\<^esup>(B,B')" by blast with assms have "converse(g) : bij\<^bsup>M\<^esup>(B', B)" "converse(f): bij\<^bsup>M\<^esup>(A, A')" using bij_converse_bij by auto let ?H="\ h \ A \\<^bsup>M\<^esup> B . g O h O f" let ?I="\ h \ A' \\<^bsup>M\<^esup> B' . converse(g) O h O converse(f)" have go:"g O F O f : A' \\<^bsup>M\<^esup> B'" if "F: A \\<^bsup>M\<^esup> B" for F proof - note assms \f\_\ \g\_\ that moreover from this have "g O F O f : A' \ B'" using bij_rel_is_fun[OF \g\_\] bij_rel_is_fun[OF \f\_\] comp_fun mem_function_space_rel[OF \F\_\] by blast ultimately show "g O F O f : A' \\<^bsup>M\<^esup> B'" using comp_closed function_space_rel_char bij_rel_char by auto qed have og:"converse(g) O F O converse(f) : A \\<^bsup>M\<^esup> B" if "F: A' \\<^bsup>M\<^esup> B'" for F proof - note assms that \converse(f) \ _\ \converse(g) \ _\ moreover from this have "converse(g) O F O converse(f) : A \ B" using bij_rel_is_fun[OF \converse(g)\_\] bij_rel_is_fun[OF \converse(f)\_\] comp_fun mem_function_space_rel[OF \F\_\] by blast ultimately show "converse(g) O F O converse(f) : A \\<^bsup>M\<^esup> B" (is "?G\_") using comp_closed function_space_rel_char bij_rel_char by auto qed with go have tc:"?H \ (A \\<^bsup>M\<^esup> B) \ (A'\\<^bsup>M\<^esup> B')" "?I \ (A' \\<^bsup>M\<^esup> B') \ (A\\<^bsup>M\<^esup> B)" using lam_type by auto with assms \f\_\ \g\_\ have "M(g O x O f)" and "M(converse(g) O x O converse(f))" if "M(x)" for x using bij_rel_char comp_closed that by auto with assms \f\_\ \g\_\ have "M(?H)" "M(?I)" using lam_replacement_iff_lam_closed[THEN iffD1,OF _ lam_replacement_comp'] bij_rel_char by auto show ?thesis unfolding eqpoll_rel_def proof (intro rexI[of _ ?H] fg_imp_bijective_rel) from og go have "(\x. x \ A' \\<^bsup>M\<^esup> B' \ converse(g) O x O converse(f) \ A \\<^bsup>M\<^esup> B)" by simp next show "M(A \\<^bsup>M\<^esup> B)" using assms by simp next show "M(A' \\<^bsup>M\<^esup> B')" using assms by simp next from og assms have "?H O ?I = (\x\A' \\<^bsup>M\<^esup> B' . (g O converse(g)) O x O (converse(f) O f))" using lam_cong[OF refl[of "A' \\<^bsup>M\<^esup> B'"]] comp_assoc comp_lam by auto also have "... = (\x\A' \\<^bsup>M\<^esup> B' . id(B') O x O (id(A')))" using left_comp_inverse[OF mem_inj_rel[OF bij_rel_is_inj_rel]] \f\_\ right_comp_inverse[OF bij_is_surj[OF mem_bij_rel]] \g\_\ assms by auto also have "... = (\x\A' \\<^bsup>M\<^esup> B' . x)" using left_comp_id[OF fun_is_rel[OF mem_function_space_rel]] right_comp_id[OF fun_is_rel[OF mem_function_space_rel]] assms by auto also have "... = id(A'\\<^bsup>M\<^esup>B')" unfolding id_def by simp finally show "?H O ?I = id(A' \\<^bsup>M\<^esup> B')" . next from go assms have "?I O ?H = (\x\A \\<^bsup>M\<^esup> B . (converse(g) O g) O x O (f O converse(f)))" using lam_cong[OF refl[of "A \\<^bsup>M\<^esup> B"]] comp_assoc comp_lam by auto also have "... = (\x\A \\<^bsup>M\<^esup> B . id(B) O x O (id(A)))" using left_comp_inverse[OF mem_inj_rel[OF bij_rel_is_inj_rel[OF \g\_\]]] right_comp_inverse[OF bij_is_surj[OF mem_bij_rel[OF \f\_\]]] assms by auto also have "... = (\x\A \\<^bsup>M\<^esup> B . x)" using left_comp_id[OF fun_is_rel[OF mem_function_space_rel]] right_comp_id[OF fun_is_rel[OF mem_function_space_rel]] assms by auto also have "... = id(A\\<^bsup>M\<^esup>B)" unfolding id_def by simp finally show "?I O ?H = id(A \\<^bsup>M\<^esup> B)" . next from assms tc \M(?H)\ \M(?I)\ show "?H \ (A\\<^bsup>M\<^esup> B) \\<^bsup>M\<^esup> (A'\\<^bsup>M\<^esup> B')" "M(?H)" "?I \ (A'\\<^bsup>M\<^esup> B') \\<^bsup>M\<^esup> (A\\<^bsup>M\<^esup> B)" using mem_function_space_rel_abs by auto qed qed lemma curry_eqpoll_rel: fixes \1 \2 \ assumes "M(\1)" "M(\2)" "M(\)" shows "\1 \\<^bsup>M\<^esup> (\2 \\<^bsup>M\<^esup> \) \\<^bsup>M\<^esup> \1 \ \2 \\<^bsup>M\<^esup> \" unfolding eqpoll_rel_def proof (intro rexI, rule lam_bijective_rel, rule_tac [1-2] mem_function_space_rel_abs[THEN iffD2], rule_tac [4] lam_type, rule_tac [8] lam_type, rule_tac [8] mem_function_space_rel_abs[THEN iffD2], rule_tac [11] lam_type, simp_all add:assms) let ?cur="\x. \w\\1 \ \2. x ` fst(w) ` snd(w)" fix f z assume "f : \1 \\<^bsup>M\<^esup> (\2 \\<^bsup>M\<^esup> \)" moreover note assms moreover from calculation have "M(\2 \\<^bsup>M\<^esup> \)" using function_space_rel_closed by simp moreover from calculation have "M(f)" "f : \1 \ (\2 \\<^bsup>M\<^esup> \)" using function_space_rel_char by (auto dest:transM) moreover from calculation have "x \ \1 \ f`x : \2 \ \" for x by (auto dest:transM intro!:mem_function_space_rel_abs[THEN iffD1]) moreover from this show "(\a\\1. \b\\2. ?cur(f) ` \a, b\) = f" using Pi_type[OF \f \ \1 \ \2 \\<^bsup>M\<^esup> \\, of "\_.\2 \ \"] by simp moreover assume "z \ \1 \ \2" moreover from calculation have "f`fst(z): \2 \\<^bsup>M\<^esup> \" by simp ultimately show "f`fst(z)`snd(z) \ \" using mem_function_space_rel_abs by (auto dest:transM) next \ \one composition is the identity:\ let ?cur="\x. \w\\1 \ \2. x ` fst(w) ` snd(w)" fix f assume "f : \1 \ \2 \\<^bsup>M\<^esup> \" with assms show "?cur(\x\\1. \xa\\2. f ` \x, xa\) = f" using function_space_rel_char mem_function_space_rel_abs by (auto dest:transM intro:fun_extension) fix x y assume "x\\1" "y\\2" with assms \f : \1 \ \2 \\<^bsup>M\<^esup> \\ show "f`\x,y\ \ \" using function_space_rel_char mem_function_space_rel_abs by (auto dest:transM[of _ "\1 \ \2 \\<^bsup>M\<^esup> \"]) next let ?cur="\x. \w\\1 \ \2. x ` fst(w) ` snd(w)" note assms moreover from this show "\x[M]. M(?cur(x))" using lam_replacement_fst lam_replacement_snd lam_replacement_apply2[THEN [5] lam_replacement_hcomp2, THEN [1] lam_replacement_hcomp2, where h="(`)", OF lam_replacement_constant] lam_replacement_apply2 by (auto intro: lam_replacement_iff_lam_closed[THEN iffD1, rule_format]) moreover from calculation show "x \ \1 \\<^bsup>M\<^esup> (\2 \\<^bsup>M\<^esup> \) \ M(?cur(x))" for x by (auto dest:transM) moreover from assms show "lam_replacement(M, ?cur)" using lam_replacement_Lambda_apply_fst_snd by simp ultimately show "M(\x\\1 \\<^bsup>M\<^esup> (\2 \\<^bsup>M\<^esup> \). ?cur(x))" using lam_replacement_iff_lam_closed by (auto dest:transM) from assms show "y \ \1 \ \2 \\<^bsup>M\<^esup> \ \ x \ \1 \ M(\xa\\2. y ` \x, xa\)" for x y using lam_replacement_apply_const_id by (rule_tac lam_replacement_iff_lam_closed[THEN iffD1, rule_format]) (auto dest:transM) from assms show "y \ \1 \ \2 \\<^bsup>M\<^esup> \ \ M(\x\\1. \xa\\2. y ` \x, xa\)" for y using lam_replacement_apply2[THEN [5] lam_replacement_hcomp2, OF lam_replacement_constant lam_replacement_const_id] lam_replacement_Lambda_apply_Pair[of \2] by (auto dest:transM intro!: lam_replacement_iff_lam_closed[THEN iffD1, rule_format]) qed lemma Pow_rel_eqpoll_rel_function_space_rel: fixes d X notes bool_of_o_def [simp] defines [simp]:"d(A) \ (\x\X. bool_of_o(x\A))" \ \the witnessing map for the thesis:\ assumes "M(X)" shows "Pow\<^bsup>M\<^esup>(X) \\<^bsup>M\<^esup> X \\<^bsup>M\<^esup> 2" proof - from assms - interpret M_Pi_assumptions M X "\_. 2" - using Pi_replacement Pi_separation lam_replacement_identity - lam_replacement_Sigfun[THEN lam_replacement_imp_strong_replacement] - Pi_replacement1[of _ 2] transM[of _ X] lam_replacement_constant - by unfold_locales auto + interpret M_Pi_assumption_repl M X "\x . 2" + by unfold_locales (simp_all add:lam_replacement_constant) have "lam_replacement(M, \x. bool_of_o(x\A))" if "M(A)" for A using that lam_replacement_if lam_replacement_constant separation_in_constant by simp with assms have "lam_replacement(M, \x. d(x))" using separation_in_constant[THEN [3] lam_replacement_if, of "\_.1" "\_.0"] lam_replacement_identity lam_replacement_constant lam_replacement_Lambda_if_mem by simp show ?thesis unfolding eqpoll_rel_def proof (intro rexI, rule lam_bijective_rel) \ \We give explicit mutual inverses\ fix A assume "A\Pow\<^bsup>M\<^esup>(X)" moreover note \M(X)\ moreover from calculation have "M(A)" by (auto dest:transM) moreover note \_ \ lam_replacement(M, \x. bool_of_o(x\A))\ ultimately show "d(A) : X \\<^bsup>M\<^esup> 2" using function_space_rel_char lam_replacement_iff_lam_closed[THEN iffD1] by (simp, rule_tac lam_type[of X "\x. bool_of_o(x\A)" "\_. 2", simplified]) auto from \A\Pow\<^bsup>M\<^esup>(X)\ \M(X)\ show "{y\X. d(A)`y = 1} = A" using Pow_rel_char by auto next fix f assume "f: X\\<^bsup>M\<^esup> 2" with assms have "f: X\ 2" "M(f)" using function_space_rel_char by simp_all then show "d({y \ X . f ` y = 1}) = f" using apply_type[OF \f: X\2\] by (force intro:fun_extension) from \M(X)\ \M(f)\ show "{ya \ X . f ` ya = 1} \ Pow\<^bsup>M\<^esup>(X)" using Pow_rel_char separation_equal_apply by auto next from assms \lam_replacement(M, \x. d(x))\ \\A. _ \ lam_replacement(M, \x. bool_of_o(x\A))\ show "M(\x\Pow\<^bsup>M\<^esup>(X). d(x))" "lam_replacement(M, \x. d(x))" "\x[M]. M(d(x))" using lam_replacement_iff_lam_closed[THEN iffD1] by auto qed (auto simp:\M(X)\) qed lemma Pow_rel_bottom: "M(B) \ 0 \ Pow\<^bsup>M\<^esup>(B)" using Pow_rel_char by simp lemma cantor_surj_rel: assumes "M(f)" "M(A)" shows "f \ surj\<^bsup>M\<^esup>(A,Pow\<^bsup>M\<^esup>(A))" proof assume "f \ surj\<^bsup>M\<^esup>(A,Pow\<^bsup>M\<^esup>(A))" with assms have "f \ surj(A,Pow\<^bsup>M\<^esup>(A))" using surj_rel_char by simp moreover note assms moreover from this have "M({x \ A . x \ f ` x})" "{x \ A . x \ f ` x} = A - {x \ A . x \ f ` x}" using lam_replacement_apply[THEN [4] separation_in, of "\x. x"] lam_replacement_identity lam_replacement_constant by auto with \M(A)\ have "{x\A . x \ f`x} \ Pow\<^bsup>M\<^esup>(A)" by (intro mem_Pow_rel_abs[THEN iffD2]) auto ultimately obtain d where "d\A" "f`d = {x\A . x \ f`x}" unfolding surj_def by blast show False proof (cases "d \ f`d") case True note \d \ f`d\ also note \f`d = {x\A . x \ f`x}\ finally have "d \ f`d" using \d\A\ by simp then show False using \d \ f ` d\ by simp next case False with \d\A\ have "d \ {x\A . x \ f`x}" by simp also from \f`d = \\ have "\ = f`d" by simp finally show False using \d \ f`d\ by simp qed qed lemma cantor_inj_rel: "M(f) \ M(A) \ f \ inj\<^bsup>M\<^esup>(Pow\<^bsup>M\<^esup>(A),A)" using inj_rel_imp_surj_rel[OF _ Pow_rel_bottom, of f A A] cantor_surj_rel[of "\x\A. if x \ range(f) then converse(f) ` x else 0" A] lam_replacement_if separation_in_constant[of "range(f)"] lam_replacement_converse_app[THEN [5] lam_replacement_hcomp2] lam_replacement_identity lam_replacement_constant lam_replacement_iff_lam_closed by auto end \ \\<^locale>\M_ZF_library\\ end \ No newline at end of file diff --git a/thys/Transitive_Models/ZF_Miscellanea.thy b/thys/Transitive_Models/ZF_Miscellanea.thy --- a/thys/Transitive_Models/ZF_Miscellanea.thy +++ b/thys/Transitive_Models/ZF_Miscellanea.thy @@ -1,204 +1,209 @@ section\Various results missing from ZF.\ theory ZF_Miscellanea imports ZF Nat_Miscellanea begin lemma rex_mono : assumes "\ d \ A . P(d)" "A\B" shows "\ d \ B. P(d)" using assms by auto lemma function_subset: "function(f) \ g\f \ function(g)" unfolding function_def subset_def by auto lemma converse_refl : "refl(A,r) \ refl(A,converse(r))" unfolding refl_def by simp lemma Ord_lt_subset : "Ord(b) \ a a\b" by(intro subsetI,frule ltD,rule_tac Ord_trans,simp_all) lemma funcI : "f \ A \ B \ a \ A \ b= f ` a \ \a, b\ \ f" by(simp_all add: apply_Pair) lemma vimage_fun_sing: assumes "f\A\B" "b\B" shows "{a\A . f`a=b} = f-``{b}" using assms vimage_singleton_iff function_apply_equality Pi_iff funcI by auto lemma image_fun_subset: "S\A\B \ C\A\ {S ` x . x\ C} = S``C" using image_function[symmetric,of S C] domain_of_fun Pi_iff by auto +lemma inj_range_Diff: + assumes "f \ inj(A,A')" + shows "f``A - f``T = f``(A - T)" + using inj_equality[OF _ _ assms] by auto + lemma subset_Diff_Un: "X \ A \ A = (A - X) \ X " by auto lemma Diff_bij: assumes "\A\F. X \ A" shows "(\A\F. A-X) \ bij(F, {A-X. A\F})" using assms unfolding bij_def inj_def surj_def by (auto intro:lam_type, subst subset_Diff_Un[of X]) auto lemma function_space_nonempty: assumes "b\B" shows "(\x\A. b) : A \ B" using assms lam_type by force lemma lam_constant_eq_cartprod: "(\_\A. y) = A \ {y}" unfolding lam_def by auto lemma vimage_lam: "(\x\A. f(x)) -`` B = { x\A . f(x) \ B }" using lam_funtype[of A f, THEN [2] domain_type] lam_funtype[of A f, THEN [2] apply_equality] lamI[of _ A f] by auto blast lemma range_fun_subset_codomain: assumes "h:B \ C" shows "range(h) \ C" - unfolding range_def domain_def converse_def using range_type[OF _ assms] by auto + unfolding range_def domain_def converse_def using range_type[OF _ assms] by auto lemma Pi_rangeD: assumes "f\Pi(A,B)" "b \ range(f)" shows "\a\A. f`a = b" using assms apply_equality[OF _ assms(1), of _ b] domain_type[OF _ assms(1)] by auto lemma Pi_range_eq: "f \ Pi(A,B) \ range(f) = {f ` x . x \ A}" using Pi_rangeD[of f A B] apply_rangeI[of f A B] by blast lemma Pi_vimage_subset : "f \ Pi(A,B) \ f-``C \ A" unfolding Pi_def by auto definition minimum :: "i \ i \ i" where "minimum(r,B) \ THE b. first(b,B,r)" lemma minimum_in': "minimum(r,B) \ B \ {0}" using the_0 first_is_elem unfolding minimum_def by (cases "\!b. first(b, B, r)") (auto dest!:theI[of "\b. first(b, B, r)"]) lemma minimum_in: "\ well_ord(A,r); B\A; B\0 \ \ minimum(r,B) \ B" using the_first_in unfolding minimum_def by simp lemma well_ord_surj_imp_inj_inverse: assumes "well_ord(A,r)" "h \ surj(A,B)" shows "(\b\B. minimum(r, {a\A. h`a=b})) \ inj(B,A)" proof - let ?f="\b\B. minimum(r, {a\A. h`a=b})" have "minimum(r, {a \ A . h ` a = b}) \ {a\A. h`a=b}" if "b\B" for b proof - from \h \ surj(A,B)\ that have "{a\A. h`a=b} \ 0" unfolding surj_def by blast with \well_ord(A,r)\ show "minimum(r,{a\A. h`a=b}) \ {a\A. h`a=b}" using minimum_in by blast qed moreover from this have "?f : B \ A" using lam_type[of B _ "\_.A"] by simp moreover have "?f ` w = ?f ` x \ w = x" if "w\B" "x\B" for w x proof - from calculation that have "w = h ` minimum(r,{a\A. h`a=w})" "x = h ` minimum(r,{a\A. h`a=x})" by simp_all moreover assume "?f ` w = ?f ` x" moreover from this and that have "minimum(r, {a \ A . h ` a = w}) = minimum(r, {a \ A . h ` a = x})" unfolding minimum_def by simp_all moreover from calculation(1,2,4) show "w=x" by simp qed ultimately show ?thesis unfolding inj_def by blast qed lemma well_ord_surj_imp_lepoll: assumes "well_ord(A,r)" "h \ surj(A,B)" shows "B\A" unfolding lepoll_def using well_ord_surj_imp_inj_inverse[OF assms] by blast \ \New result\ lemma surj_imp_well_ord: assumes "well_ord(A,r)" "h \ surj(A,B)" shows "\s. well_ord(B,s)" using assms lepoll_well_ord[OF well_ord_surj_imp_lepoll] by force lemma Pow_sing : "Pow({a}) = {0,{a}}" proof(intro equalityI,simp_all) have "z \ {0,{a}}" if "z \ {a}" for z using that by auto then show " Pow({a}) \ {0, {a}}" by auto qed lemma Pow_cons: shows "Pow(cons(a,A)) = Pow(A) \ {{a} \ X . X: Pow(A)}" using Un_Pow_subset Pow_sing proof(intro equalityI,auto simp add:Un_Pow_subset) { fix C D assume "\ B . B\Pow(A) \ C \ {a} \ B" "C \ {a} \ A" "D \ C" moreover from this have "\x\C . x=a \ x\A" by auto moreover from calculation consider (a) "D=a" | (b) "D\A" by auto from this have "D\A" proof(cases) case a with calculation show ?thesis by auto next case b then show ?thesis by simp qed } then show "\x xa. (\xa\Pow(A). x \ {a} \ xa) \ x \ cons(a, A) \ xa \ x \ xa \ A" by auto qed lemma app_nm : assumes "n\nat" "m\nat" "f\n\m" "x \ nat" shows "f`x \ nat" proof(cases "x\n") case True then show ?thesis using assms in_n_in_nat apply_type by simp next case False then show ?thesis using assms apply_0 domain_of_fun by simp qed lemma Upair_eq_cons: "Upair(a,b) = {a,b}" unfolding cons_def by auto lemma converse_apply_eq : "converse(f) ` x = \(f -`` {x})" unfolding apply_def vimage_def by simp lemmas app_fun = apply_iff[THEN iffD1] lemma Finite_imp_lesspoll_nat: assumes "Finite(A)" shows "A \ nat" using assms subset_imp_lepoll[OF naturals_subset_nat] eq_lepoll_trans n_lesspoll_nat eq_lesspoll_trans unfolding Finite_def lesspoll_def by auto definition curry :: "[i,i,i] \ i" where "curry(A,B,f) \ \x\A . \y\B . f`\x,y\" lemma curry_type : assumes "f \ A\B \ C" shows "curry(A,B,f) \ A \ (B \ C)" using assms lam_funtype unfolding curry_def by auto end \ No newline at end of file diff --git a/thys/Transitive_Models/document/root.bib b/thys/Transitive_Models/document/root.bib --- a/thys/Transitive_Models/document/root.bib +++ b/thys/Transitive_Models/document/root.bib @@ -1,92 +1,102 @@ @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} } @article {MR2051585, AUTHOR = {Paulson, Lawrence C.}, TITLE = {The relative consistency of the axiom of choice mechanized using {I}sabelle/{ZF}}, NOTE = {Appendix A available electronically at \url{http://www.lms.ac.uk/jcm/6/lms2003-001/appendix-a/}}, JOURNAL = {LMS J. Comput. Math.}, FJOURNAL = {LMS Journal of Computation and Mathematics}, VOLUME = {6}, YEAR = {2003}, PAGES = {198--248}, ISSN = {1461-1570}, MRCLASS = {03B35 (03E25 03E35 68T15)}, MRNUMBER = {2051585}, DOI = {10.1007/978-3-540-69407-6_52}, URL = {http://dx.doi.org/10.1007/978-3-540-69407-6_52}, } @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} } + +@book{moschovakis1994notes, + title={Notes on Set Theory}, + author={Moschovakis, Yiannis N.}, + isbn={9783540941804}, + lccn={93035825}, + series={Springer Texts in Electrical Engineering}, + year={1994}, + publisher={Springer-Verlag} +} diff --git a/thys/Transitive_Models/document/root.tex b/thys/Transitive_Models/document/root.tex --- a/thys/Transitive_Models/document/root.tex +++ b/thys/Transitive_Models/document/root.tex @@ -1,184 +1,191 @@ \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}} \begin{document} \title{Transitive Models of Fragments of 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 extend the ZF-Constructibility library by relativizing theories of the Isabelle/ZF and Delta System Lemma sessions to a transitive class. We also relativize Paulson's work on Aleph and our former treatment of the Axiom of Dependent Choices. This work is a prerequisite to our formalization of the independence of the Continuum Hypothesis. \end{abstract} \tableofcontents % sane default for proof documents \parindent 0pt\parskip 0.5ex \section{Introduction} Relativization of concepts is a key tool to obtain results in forcing, as it is explained in \cite[Sect.~3]{2020arXiv200109715G} and elsewhere. In this session, we cast some theories in relative form, in a way that they now refer to a fixed class $M$ as the universe of discourse. Whenever it was possible, we tried to minimize the changes to the structure and proof scripts of the absolute concepts. For this reason, some comments of the original text as well as outdated \textbf{apply} commands appear profusely in the following theories. A repeated pattern that appears is that the relativized result can be proved \emph{mutatis mutandis}, with remaining proof obligations that the objects constructed actually belong to the model $M$. Another aspect was that the management of higher order constructs always posed some extra problems, already noted by Paulson \cite[Sect.~7.3]{MR2051585}. In the theory \theory{Lambda\_Replacement} we introduce a new locale assuming -two instances of separation and twelve instances of ``lambda replacements'' +two instances of separation and six instances of ``lambda replacements'' (i.e., replacements using definable functions of the form $\lambda x y. y=\langle x, f(x) \rangle$) -that allow for having some form of compositionality of further instances +that allow some form of compositionality of further instances of separations and replacements. We proceed to enumerate the theories that were ``ported'' to relative form, briefly commenting on each of them. Below, we refer to the original theories as the \emph{source} and, correspondingly, call \emph{target} the relativized version. We omit the \theory{.thy} suffixes. \begin{enumerate} \item From \session{ZF}: \begin{enumerate} \item \theory{Univ}. Here we decided to relativize only the term \isatt{Vfrom} that constructs the cumulative hierarchy up to some ordinal length and starting from an arbitrary set. \item \theory{Cardinal}. There are two targets for this source, \theory{Least} and \theory{Cardinal\_Relative}. Both require some fair amount of preparation, trying to take advantage of absolute concepts. It is not straightforward to compare source and targets in a line-by-line fashion at this point. \item \theory{CardinalArith}. The hardest part was to formalize the - cardinal successor function. We also disregarded the part treating + cardinal successor function and ordertypes. + We also disregarded the part treating finite cardinals since it is an absolute concept. Apart from that, the relative version closely parallels the source. \item \theory{Cardinal\_AC}. After some boilerplate, porting was rather straightforward, excepting cardinal arithmetic involving the higher-order union operator. \end{enumerate} \item From \session{ZF-Constructible}: \begin{enumerate} \item \theory{Normal}. The target here is \theory{Aleph\_Relative} since that is the only concept that we ported. Instead of porting all the machinery of normal functions (since it involved higher-order variables), we particularized the results for the Aleph function. We also used an alternative definition of the latter that worked better with our relativization discipline. \end{enumerate} \item From \session{Delta\_System\_Lemma}: \begin{enumerate} \item \theory{ZF\_Library}. The target includes a big section of auxiliary lemmas and commands that aid the relativization. We needed to make explicit the witnesses (mainly functions) in some of the existential results proved in the source, since only in that way we would be able to show that they belonged to the model. \item \theory{Cardinal\_Library}. Porting was relatively straightforward; most of the extra work laid in adjusting locale assumptions to obtain an appropriate context to state and prove the theorems. \item \theory{Delta\_System}. Same comments as in the case of \theory{Cardinal\_Library} apply here. \end{enumerate} \item From \session{Forcing}: \begin{enumerate} \item \theory{Pointed\_DC}. This case was similar to \theory{Cardinal\_AC} above, although a bit of care was needed to handle the recursive construction. Also, a fraction of the theory \theory{AC} from \session{ZF} was ported here as it was a prerequisite. A complete relativization of \theory{AC} would be desirable but still missing. \end{enumerate} \end{enumerate} + +Finally, we had to repeat some material from \theory{Relative} in +\session{ZF-Constructible} (in our theory \theory{M\_Basic\_No\_Repl}) +in order to eliminate one replacement instance appearing in the locale +\texttt{M\_Basic}. + % generated text of all theories \input{session} \bibliographystyle{root} \bibliography{root} \end{document} %%% Local Variables: %%% mode: latex %%% TeX-master: t %%% ispell-local-dictionary: "american" %%% End: