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,680 +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"}\ + \ \Compare @{thm [source] "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}.\ + @{thm [source] 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\ 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/Cofinality.thy b/thys/Delta_System_Lemma/Cofinality.thy --- a/thys/Delta_System_Lemma/Cofinality.thy +++ b/thys/Delta_System_Lemma/Cofinality.thy @@ -1,1205 +1,1206 @@ section\Cofinality\label{sec:cofinality}\ theory Cofinality imports ZF_Library begin subsection\Basic results and definitions\ text\A set \<^term>\X\ is \<^emph>\cofinal\ in \<^term>\A\ (with respect to the relation \<^term>\r\) if every element of \<^term>\A\ is ``bounded above'' by some element of \<^term>\X\. Note that \<^term>\X\ does not need to be a subset of \<^term>\A\.\ definition cofinal :: "[i,i,i] \ o" where "cofinal(X,A,r) \ \a\A. \x\X. \a,x\\r \ a = x" (* (* Alternative definitions *) definition cofinal_predic :: "[i,i,[i,i]\o] \ o" where "cofinal_predic(X,A,r) \ \a\A. \x\X. r(a,x) \ a = x" definition f_cofinal :: "[i\i,i,i,i] \ o" where "f_cofinal(f,C,A,r) \ \a\A. \x\C. \a,f(x)\\r \ a = f(x)" (* The next definition doesn't work if 0 is the top element of A. But it works for limit ordinals. *) definition cofinal_fun' :: "[i,i,i] \ o" where "cofinal_fun'(f,A,r) == f_cofinal(\x. f`x,domain(f),A, r)" *) text\A function is cofinal if it range is.\ definition cofinal_fun :: "[i,i,i] \ o" where "cofinal_fun(f,A,r) \ \a\A. \x\domain(f). \a,f`x\\r \ a = f`x" lemma cofinal_funI: assumes "\a. a\A \ \x\domain(f). \a,f`x\\r \ a = f`x" shows "cofinal_fun(f,A,r)" using assms unfolding cofinal_fun_def by simp lemma cofinal_funD: assumes "cofinal_fun(f,A,r)" "a\A" shows "\x\domain(f). \a,f`x\\r \ a = f`x" using assms unfolding cofinal_fun_def by simp lemma cofinal_in_cofinal: assumes "trans(r)" "cofinal(Y,X,r)" "cofinal(X,A,r)" shows "cofinal(Y,A,r)" unfolding cofinal_def proof fix a assume "a\A" moreover from \cofinal(X,A,r)\ have "b\A\\x\X. \b,x\\r \ b=x" for b unfolding cofinal_def by simp ultimately obtain y where "y\X" "\a,y\\r \ a=y" by auto moreover from \cofinal(Y,X,r)\ have "c\X\\y\Y. \c,y\\r \ c=y" for c unfolding cofinal_def by simp ultimately obtain x where "x\Y" "\y,x\\r \ y=x" by auto with \a\A\ \y\X\ \\a,y\\r \ a=y\ \trans(r)\ show "\x\Y. \a,x\\r \ a=x" unfolding trans_def by auto qed lemma codomain_is_cofinal: assumes "cofinal_fun(f,A,r)" "f:C \ D" shows "cofinal(D,A,r)" unfolding cofinal_def proof fix b assume "b \ A" moreover from assms have "a\A \ \x\domain(f). \a, f ` x\ \ r \ a = f`x" for a unfolding cofinal_fun_def by simp ultimately obtain x where "x\domain(f)" "\b, f ` x\ \ r \ b = f`x" by blast moreover from \f:C \ D\ \x\domain(f)\ have "f`x\D" using domain_of_fun apply_rangeI by simp ultimately show "\y\D. \b, y\ \ r \ b = y" by auto qed lemma cofinal_range_iff_cofinal_fun: assumes "function(f)" shows "cofinal(range(f),A,r) \ cofinal_fun(f,A,r)" unfolding cofinal_fun_def proof (intro iffI ballI) fix a assume "a\A" \cofinal(range(f),A,r)\ then obtain y where "y\range(f)" "\a,y\ \ r \ a = y" unfolding cofinal_def by blast moreover from this obtain x where "\x,y\\f" unfolding range_def domain_def converse_def by blast moreover note \function(f)\ ultimately have "\a, f ` x\ \ r \ a = f ` x" using function_apply_equality by blast with \\x,y\\f\ show "\x\domain(f). \a, f ` x\ \ r \ a = f ` x" by blast next assume "\a\A. \x\domain(f). \a, f ` x\ \ r \ a = f ` x" with assms show "cofinal(range(f), A, r)" using function_apply_Pair[of f] unfolding cofinal_def by fast qed lemma cofinal_comp: assumes "f\ mono_map(C,s,D,r)" "cofinal_fun(f,D,r)" "h:B \ C" "cofinal_fun(h,C,s)" "trans(r)" shows "cofinal_fun(f O h,D,r)" unfolding cofinal_fun_def proof fix a from \f\ mono_map(C,s,D,r)\ have "f:C \ D" using mono_map_is_fun by simp with \h:B \ C\ have "domain(f) = C" "domain(h) = B" using domain_of_fun by simp_all moreover assume "a \ D" moreover note \cofinal_fun(f,D,r)\ ultimately obtain c where "c\C" "\a, f ` c\ \ r \ a = f ` c" unfolding cofinal_fun_def by blast with \cofinal_fun(h,C,s)\ \domain(h) = B\ obtain b where "b \ B" "\c, h ` b\ \ s \ c = h ` b" unfolding cofinal_fun_def by blast moreover from this and \h:B \ C\ have "h`b \ C" by simp moreover note \f \ mono_map(C,s,D,r)\ \c\C\ ultimately have "\f`c, f` (h ` b)\ \ r \ f`c = f` (h ` b)" unfolding mono_map_def by blast with \\a, f ` c\ \ r \ a = f ` c\ \trans(r)\ \h:B \ C\ \b\B\ have "\a, (f O h) ` b\ \ r \ a = (f O h) ` b" using transD by auto moreover from \h:B \ C\ \domain(f) = C\ \domain(h) = B\ have "domain(f O h) = B" using range_fun_subset_codomain by blast moreover note \b\B\ ultimately show "\x\domain(f O h). \a, (f O h) ` x\ \ r \ a = (f O h) ` x" by blast qed definition cf_fun :: "[i,i] \ o" where "cf_fun(f,\) \ cofinal_fun(f,\,Memrel(\))" lemma cf_funI[intro!]: "cofinal_fun(f,\,Memrel(\)) \ cf_fun(f,\)" unfolding cf_fun_def by simp lemma cf_funD[dest!]: "cf_fun(f,\) \ cofinal_fun(f,\,Memrel(\))" unfolding cf_fun_def by simp lemma cf_fun_comp: assumes "Ord(\)" "f\ mono_map(C,s,\,Memrel(\))" "cf_fun(f,\)" "h:B \ C" "cofinal_fun(h,C,s)" shows "cf_fun(f O h,\)" using assms cofinal_comp[OF _ _ _ _ trans_Memrel] by auto definition cf :: "i\i" where "cf(\) \ \ \. \A. A\\ \ cofinal(A,\,Memrel(\)) \ \ = ordertype(A,Memrel(\))" lemma Ord_cf [TC]: "Ord(cf(\))" unfolding cf_def using Ord_Least by simp lemma gamma_cofinal_gamma: assumes "Ord(\)" shows "cofinal(\,\,Memrel(\))" unfolding cofinal_def by auto lemma cf_is_ordertype: assumes "Ord(\)" shows "\A. A\\ \ cofinal(A,\,Memrel(\)) \ cf(\) = ordertype(A,Memrel(\))" (is "?P(cf(\))") using gamma_cofinal_gamma LeastI[of ?P \] ordertype_Memrel[symmetric] assms unfolding cf_def by blast lemma cf_fun_succ': assumes "Ord(\)" "Ord(\)" "f:\\succ(\)" shows "(\x\\. f`x=\) \ cf_fun(f,succ(\))" proof (intro iffI) assume "(\x\\. f`x=\)" with assms show "cf_fun(f,succ(\))" using domain_of_fun[OF \f:\\succ(\)\] unfolding cf_fun_def cofinal_fun_def by auto next assume "cf_fun(f,succ(\))" with assms obtain x where "x\\" "\\,f`x\ \ Memrel(succ(\)) \ \ = f ` x" using domain_of_fun[OF \f:\\succ(\)\] unfolding cf_fun_def cofinal_fun_def by auto moreover from \Ord(\)\ have "\\,y\ \ Memrel(succ(\))" for y using foundation unfolding Memrel_def by blast ultimately show "\x\\. f ` x = \" by blast qed lemma cf_fun_succ: "Ord(\) \ f:1\succ(\) \ f`0=\ \ cf_fun(f,succ(\))" using cf_fun_succ' by blast lemma ordertype_0_not_cofinal_succ: assumes "ordertype(A,Memrel(succ(i))) = 0" "A\succ(i)" "Ord(i)" shows "\cofinal(A,succ(i),Memrel(succ(i)))" proof have 1:"ordertype(A,Memrel(succ(i))) = ordertype(0,Memrel(0))" using \ordertype(A,Memrel(succ(i))) = 0\ ordertype_0 by simp from \A\succ(i)\ \Ord(i)\ have "\f. f \ \A, Memrel(succ(i))\ \ \0, Memrel(0)\" using well_ord_Memrel well_ord_subset ordertype_eq_imp_ord_iso[OF 1] Ord_0 by blast then have "A=0" using ord_iso_is_bij bij_imp_eqpoll eqpoll_0_is_0 by blast moreover assume "cofinal(A, succ(i), Memrel(succ(i)))" moreover note \Ord(i)\ ultimately show "False" using not_mem_empty unfolding cofinal_def by auto qed text\I thank Edwin Pacheco Rodríguez for the following lemma.\ lemma cf_succ: assumes "Ord(\)" shows "cf(succ(\)) = 1" proof - define f where "f \ {\0,\\}" then have "f : 1\succ(\)" "f`0 = \" using fun_extend3[of 0 0 "succ(\)" 0 \] singleton_0 by auto with assms have "cf_fun(f,succ(\))" using cf_fun_succ unfolding cofinal_fun_def by simp from \f:1\succ(\)\ have "0\domain(f)" using domain_of_fun by simp define A where "A={f`0}" with \cf_fun(f,succ(\))\ \0\domain(f)\ \f`0=\\ have "cofinal(A,succ(\),Memrel(succ(\)))" unfolding cofinal_def cofinal_fun_def by simp moreover from \f`0=\\ \A={f`0}\ have "A \ succ(\)" unfolding succ_def by auto moreover from \Ord(\)\ \A\ succ(\)\ have "well_ord(A,Memrel(succ(\)))" using Ord_succ well_ord_Memrel well_ord_subset relation_Memrel by blast moreover from \Ord(\)\ have "\(\A. A \ succ(\) \ cofinal(A, succ(\), Memrel(succ(\))) \ 0 = ordertype(A, Memrel(succ(\))))" (is "\?P(0)") using ordertype_0_not_cofinal_succ unfolding cf_def by auto moreover have "1 = ordertype(A,Memrel(succ(\)))" proof - from \A={f`0}\ have "A\1" using singleton_eqpoll_1 by simp with \well_ord(A,Memrel(succ(\)))\ show ?thesis using nat_1I ordertype_eq_n by simp qed ultimately show "cf(succ(\)) = 1" using Ord_1 Least_equality[of ?P 1] unfolding cf_def by blast qed lemma cf_zero [simp]: "cf(0) = 0" unfolding cf_def cofinal_def using ordertype_0 subset_empty_iff Least_le[of _ 0] by auto lemma surj_is_cofinal: "f \ surj(\,\) \ cf_fun(f,\)" unfolding surj_def cofinal_fun_def cf_fun_def using domain_of_fun by force lemma cf_zero_iff: "Ord(\) \ cf(\) = 0 \ \ = 0" proof (intro iffI) assume "\ = 0" "Ord(\)" then show "cf(\) = 0" using cf_zero by simp next assume "cf(\) = 0" "Ord(\)" moreover from this obtain A where "A\\" "cf(\) = ordertype(A,Memrel(\))" "cofinal(A,\,Memrel(\))" using cf_is_ordertype by blast ultimately have "cofinal(0,\,Memrel(\))" using ordertype_zero_imp_zero[of A "Memrel(\)"] by simp then show "\=0" unfolding cofinal_def by blast qed \ \TODO: define Succ (predicate for successor ordinals)\ lemma cf_eq_one_iff: assumes "Ord(\)" shows "cf(\) = 1 \ (\\. Ord(\) \ \ = succ(\))" proof (intro iffI) assume "\\. Ord(\) \ \ = succ(\)" then show "cf(\) = 1" using cf_succ by auto next assume "cf(\) = 1" moreover from assms obtain A where "A \ \" "cf(\) = ordertype(A,Memrel(\))" "cofinal(A,\,Memrel(\))" using cf_is_ordertype by blast ultimately have "ordertype(A,Memrel(\)) = 1" by simp moreover define f where "f\converse(ordermap(A,Memrel(\)))" moreover from this \ordertype(A,Memrel(\)) = 1\ \A \ \\ assms have "f \ surj(1,A)" using well_ord_subset[OF well_ord_Memrel, THEN ordermap_bij, THEN bij_converse_bij, of \ A] bij_is_surj by simp with \cofinal(A,\,Memrel(\))\ have "\a\\. \a, f`0\ \ Memrel(\) \ a = f`0" unfolding cofinal_def surj_def by auto with assms \A \ \\ \f \ surj(1,A)\ show "\\. Ord(\) \ \ = succ(\)" using Ord_has_max_imp_succ[of \ "f`0"] surj_is_fun[of f 1 A] apply_type[of f 1 "\_.A" 0] unfolding lt_def by (auto intro:Ord_in_Ord) qed lemma ordertype_in_cf_imp_not_cofinal: assumes "ordertype(A,Memrel(\)) \ cf(\)" "A \ \" shows "\ cofinal(A,\,Memrel(\))" proof note \A \ \\ moreover assume "cofinal(A,\,Memrel(\))" ultimately have "\B. B \ \ \ cofinal(B, \, Memrel(\)) \ ordertype(A,Memrel(\)) = ordertype(B, Memrel(\))" (is "?P(ordertype(A,_))") by blast moreover from assms have "ordertype(A,Memrel(\)) < cf(\)" using Ord_cf ltI by blast ultimately show "False" unfolding cf_def using less_LeastE[of ?P "ordertype(A,Memrel(\))"] by auto qed lemma cofinal_mono_map_cf: assumes "Ord(\)" shows "\j \ mono_map(cf(\), Memrel(cf(\)), \, Memrel(\)) . cf_fun(j,\)" proof - note assms moreover from this obtain A where "A \ \" "cf(\) = ordertype(A,Memrel(\))" "cofinal(A,\,Memrel(\))" using cf_is_ordertype by blast moreover define j where "j \ converse(ordermap(A,Memrel(\)))" moreover from calculation have "j :cf(\) \\<^sub>< \" using ordertype_ord_iso[THEN ord_iso_sym, THEN ord_iso_is_mono_map, THEN mono_map_mono, of A "Memrel(\)" \] well_ord_Memrel[THEN well_ord_subset] by simp moreover from calculation have "j \ surj(cf(\),A)" using well_ord_Memrel[THEN well_ord_subset, THEN ordertype_ord_iso, THEN ord_iso_sym, of \ A, THEN ord_iso_is_bij, THEN bij_is_surj] by simp with \cofinal(A,\,Memrel(\))\ have "cf_fun(j,\)" using cofinal_range_iff_cofinal_fun[of j \ "Memrel(\)"] surj_range[of j "cf(\)" A] surj_is_fun fun_is_function by fastforce with \j \ mono_map(_,_,_,_)\ show ?thesis by auto qed subsection\The factorization lemma\ text\In this subsection we prove a factorization lemma for cofinal functions into ordinals, which shows that any cofinal function between ordinals can be ``decomposed'' in such a way that a commutative triangle of strictly increasing maps arises. The factorization lemma has a kind of fundamental character, in that the rest of the basic results on cofinality (for, instance, idempotence) follow easily from it, in a more algebraic way. This is a consequence that the proof encapsulates uses of transfinite recursion in the basic theory of cofinality; indeed, only one use is needed. In the setting of Isabelle/ZF, this is convenient since the machinery of recursion is pretty clumsy. On the downside, this way of presenting things results in a longer proof of the factorization lemma. This approach was taken by the author in the notes \cite{apunte_st} for an introductory course in Set Theory. To organize the use of the hypotheses of the factorization lemma, we set up a locale containing all the relevant ingredients. \ locale cofinal_factor = fixes j \ \ \ f assumes j_mono: "j :\ \\<^sub>< \" and ords: "Ord(\)" "Ord(\)" "Limit(\)" and f_type: "f: \ \ \" begin text\Here, \<^term>\f\ is cofinal function from \<^term>\\\ to \<^term>\\\, and the ordinal \<^term>\\\ is meant to be the cofinality of \<^term>\\\. Hence, there exists an increasing map \<^term>\j\ from \<^term>\\\ to \<^term>\\\ by the last lemma. The main goal is to construct an increasing function \<^term>\g:\\\\ such that the composition \<^term>\f O g\ is still cofinal but also increasing.\ definition factor_body :: "[i,i,i] \ o" where "factor_body(\,h,x) \ (x\\ \ j`\ \ f`x \ (\\<\ . f`(h`\) < f`x)) \ x=\" definition factor_rec :: "[i,i] \ i" where "factor_rec(\,h) \ \ x. factor_body(\,h,x)" text\\<^term>\factor_rec\ is the inductive step for the definition by transfinite recursion of the \<^emph>\factor\ function (called \<^term>\g\ above), which in turn is obtained by minimizing the predicate \<^term>\factor_body\. Next we show that this predicate is monotonous.\ lemma factor_body_mono: assumes "\\\" "\<\" "factor_body(\,\x\\. G(x),x)" shows "factor_body(\,\x\\. G(x),x)" proof - from \\<\\ have "\\\" using ltD by simp moreover note \\\\\ moreover from calculation have "\\\" using ords ltD Ord_cf Ord_trans by blast ultimately have "j`\ \ j`\" using j_mono mono_map_increasing by blast moreover from \\\\\ have "j`\\\" using j_mono domain_of_fun apply_rangeI mono_map_is_fun by force moreover from this have "Ord(j`\)" using Ord_in_Ord ords Limit_is_Ord by auto ultimately have "j`\ \ j`\" unfolding lt_def by blast then have "j`\ \ f`\ \ j`\ \ f`\" for \ using le_trans by blast moreover have "f`((\w\\. G(w))`y) < f`z" if "z\\" "\x<\. f`((\w\\. G(w))`x) < f`z" "y<\" for y z proof - note \y<\\ also note \\<\\ finally have "y<\" by simp with \\x<\. f`((\w\\. G(w))`x) < f`z\ have "f ` ((\w\\. G(w)) ` y) < f ` z" by simp moreover from \y<\\ \y<\\ have "(\w\\. G(w)) ` y = (\w\\. G(w)) ` y" using beta_if by (auto dest:ltD) ultimately show ?thesis by simp qed moreover note \factor_body(\,\x\\. G(x),x)\ ultimately show ?thesis unfolding factor_body_def by blast qed lemma factor_body_simp[simp]: "factor_body(\,g,\)" unfolding factor_body_def by simp lemma factor_rec_mono: assumes "\\\" "\<\" shows "factor_rec(\,\x\\. G(x)) \ factor_rec(\,\x\\. G(x))" unfolding factor_rec_def using assms ords factor_body_mono Least_antitone by simp text\We now define the factor as higher-order function. Later it will be restricted to a set to obtain a bona fide function of type @{typ i}.\ definition factor :: "i \ i" where "factor(\) \ transrec(\,factor_rec)" lemma factor_unfold: "factor(\) = factor_rec(\,\x\\. factor(x))" using def_transrec[OF factor_def] . lemma factor_mono: assumes "\\\" "\<\" "factor(\)\\" "factor(\)\\" shows "factor(\) \ factor(\)" proof - have "factor(\) = factor_rec(\, \x\\. factor(x))" using factor_unfold . also from assms and factor_rec_mono have "... \ factor_rec(\, \x\\. factor(x))" by simp also have "factor_rec(\, \x\\. factor(x)) = factor(\)" using def_transrec[OF factor_def, symmetric] . finally show ?thesis . qed text\The factor satisfies the predicate body of the minimization.\ lemma factor_body_factor: "factor_body(\,\x\\. factor(x),factor(\))" using ords factor_unfold[of \] LeastI[of "factor_body(_,_)" \] unfolding factor_rec_def by simp lemma factor_type [TC]: "Ord(factor(\))" using ords factor_unfold[of \] unfolding factor_rec_def by simp text\The value \<^term>\\\ in \<^term>\factor_body\ (and therefore, in \<^term>\factor\) is meant to be a ``default value''. Whenever it is not attained, the factor function behaves as expected: It is increasing and its composition with \<^term>\f\ also is.\ lemma f_factor_increasing: assumes "\\\" "\<\" "factor(\)\\" shows "f`factor(\) < f`factor(\)" proof - from assms have "f ` ((\x\\. factor(x)) ` \) < f ` factor(\)" using factor_unfold[of \] ords LeastI[of "factor_body(\,\x\\. factor(x))"] unfolding factor_rec_def factor_body_def by (auto simp del:beta_if) with \\<\\ show ?thesis using ltD by auto qed lemma factor_increasing: assumes "\\\" "\<\" "factor(\)\\" "factor(\)\\" shows "factor(\))" using assms f_factor_increasing factor_mono by (force intro:le_neq_imp_lt) lemma factor_in_delta: assumes "factor(\) \ \" shows "factor(\) \ \" using assms factor_body_factor ords unfolding factor_body_def by auto text\Finally, we define the (set) factor function as the restriction of factor to the ordinal \<^term>\\\.\ definition fun_factor :: "i" where "fun_factor \ \\\\. factor(\)" lemma fun_factor_is_mono_map: assumes "\\. \ \ \ \ factor(\) \ \" shows "fun_factor \ mono_map(\, Memrel(\), \, Memrel(\))" unfolding mono_map_def proof (intro CollectI ballI impI) - (* Proof that \<^term>\fun_factor\ respects membership *) + text\Proof that \<^term>\fun_factor\ respects membership:\ fix \ \ assume "\\\" "\\\" moreover note assms moreover from calculation have "factor(\)\\" "factor(\)\\" "Ord(\)" using factor_in_delta Ord_in_Ord ords by auto moreover assume "\\, \\ \ Memrel(\)" ultimately show "\fun_factor ` \, fun_factor ` \\ \ Memrel(\)" unfolding fun_factor_def using ltI factor_increasing[THEN ltD] factor_in_delta by simp next - (* Proof of type *) + text\Proof that it has the appropriate type:\ from assms show "fun_factor : \ \ \" unfolding fun_factor_def using ltI lam_type factor_in_delta by simp qed lemma f_fun_factor_is_mono_map: assumes "\\. \ \ \ \ factor(\) \ \" shows "f O fun_factor \ mono_map(\, Memrel(\), \, Memrel(\))" unfolding mono_map_def using f_type proof (intro CollectI ballI impI comp_fun[of _ _ \]) from assms show "fun_factor : \ \ \" using fun_factor_is_mono_map mono_map_is_fun by simp - (* Proof that f O ?g respects membership *) + text\Proof that \<^term>\f O fun_factor\ respects membership\ fix \ \ assume "\\, \\ \ Memrel(\)" then have "\<\" using Ord_in_Ord[of "\"] ltI ords by blast assume "\\\" "\\\" moreover from this and assms have "factor(\)\\" "factor(\)\\" by auto moreover have "Ord(\)" "\\0" using ords Limit_is_Ord by auto moreover note \\<\\ \fun_factor : \ \ \\ ultimately show "\(f O fun_factor) ` \, (f O fun_factor) ` \\ \ Memrel(\)" using ltD[of "f ` factor(\)" "f ` factor(\)"] f_factor_increasing apply_in_codomain_Ord f_type unfolding fun_factor_def by auto qed -end (* cofinal_factor *) +end \ \\<^locale>\cofinal_factor\\ text\We state next the factorization lemma.\ lemma cofinal_fun_factorization: notes le_imp_subset [dest] lt_trans2 [trans] assumes "Ord(\)" "Limit(\)" "f: \ \ \" "cf_fun(f,\)" shows "\g \ cf(\) \\<^sub>< \. f O g : cf(\) \\<^sub>< \ \ cofinal_fun(f O g,\,Memrel(\))" proof - from \Limit(\)\ have "Ord(\)" using Limit_is_Ord by simp then obtain j where "j :cf(\) \\<^sub>< \" "cf_fun(j,\)" using cofinal_mono_map_cf by blast then have "domain(j) = cf(\)" using domain_of_fun mono_map_is_fun by force from \j \ _\ assms interpret cofinal_factor j \ "cf(\)" by (unfold_locales) (simp_all) text\The core of the argument is to show that the factor function indeed maps into \<^term>\\\, therefore its values satisfy the first disjunct of \<^term>\factor_body\. This holds in turn because no restriction of the factor composed with \<^term>\f\ to a proper initial segment of \<^term>\cf(\)\ can be cofinal in \<^term>\\\ by definition of cofinality. Hence there must be a witness that satisfies the first disjunct.\ have factor_not_delta: "factor(\)\\" if "\ \ cf(\)" for \ text\For this, we induct on \<^term>\\\ ranging over \<^term>\cf(\)\.\ proof (induct \ rule:Ord_induct[OF _ Ord_cf[of \]]) case 1 with that show ?case . next case (2 \) then have IH: "z\\ \ factor(z)\\" for z by simp define h where "h \ \x\\. f`factor(x)" from IH have "z\\ \ factor(z) \ \" for z using factor_in_delta by blast with \f:\\\\ have "h : \ \ \" unfolding h_def using apply_funtype lam_type by auto then have "h : \ \\<^sub>< \" unfolding mono_map_def proof (intro CollectI ballI impI) fix x y assume "x\\" "y\\" moreover from this and IH have "factor(y) \ \" by simp moreover from calculation and \h \ \ \ \\ have "h`x \ \" "h`y \ \" by simp_all moreover from \\\cf(\)\ and \y\\\ have "y \ cf(\)" using Ord_trans Ord_cf by blast moreover from this have "Ord(y)" using Ord_cf Ord_in_Ord by blast moreover assume "\x,y\ \ Memrel(\)" moreover from calculation have "xh ` x, h ` y\ \ Memrel(\)" unfolding h_def using f_factor_increasing ltD by (auto) qed with \\\cf(\)\ \Ord(\)\ have "ordertype(h``\,Memrel(\)) = \" (* Maybe should use range(h) *) using mono_map_ordertype_image[of \] Ord_cf Ord_in_Ord by blast also note \\ \cf(\)\ finally have "ordertype(h``\,Memrel(\)) \ cf(\)" by simp moreover from \h \ \ \ \\ have "h``\ \ \" using mono_map_is_fun Image_sub_codomain by blast ultimately have "\ cofinal(h``\,\,Memrel(\))" using ordertype_in_cf_imp_not_cofinal by simp then obtain \_0 where "\_0\\" "\x\h `` \. \ \\_0, x\ \ Memrel(\) \ \_0 \ x" unfolding cofinal_def by auto with \Ord(\)\ \h``\ \ \\ have "\x\h `` \. x \ \_0" using well_ord_Memrel[of \] well_ord_is_linear[of \ "Memrel(\)"] unfolding linear_def by blast from \\_0 \ \\ \j \ mono_map(_,_,\,_)\ \Ord(\)\ have "j`\ \ \" using mono_map_is_fun apply_in_codomain_Ord by force with \\_0 \ \\ \Ord(\)\ have "\_0 \ j`\ \ \" using Un_least_mem_iff Ord_in_Ord by auto with \cf_fun(f,\)\ obtain \ where "\\domain(f)" "\\_0 \ j`\, f ` \\ \ Memrel(\) \ \_0 \ j`\ = f ` \" by (auto simp add:cofinal_fun_def) blast moreover from this and \f:\\\\ have "\ \ \" using domain_of_fun by auto moreover note \Ord(\)\ moreover from this and \f:\\\\ \\_0 \ \\ have "Ord(f`\)" using apply_in_codomain_Ord Ord_in_Ord by blast moreover from calculation and \\_0 \ \\ and \Ord(\)\ and \j`\ \ \\ have "Ord(\_0)" "Ord(j`\)" "Ord(\)" using Ord_in_Ord by auto moreover from \\x\h `` \. x \ \_0\ \Ord(\_0)\ \h:\\\\ have "x\\ \ h`x < \_0" for x using fun_is_function[of h \ "\_. \"] Image_subset_Ord_imp_lt domain_of_fun[of h \ "\_. \"] by blast moreover have "x\\ \ h`x < f`\" for x proof - fix x assume "x\\" with \\x\h `` \. x \ \_0\ \Ord(\_0)\ \h:\\\\ have "h`x < \_0" using fun_is_function[of h \ "\_. \"] Image_subset_Ord_imp_lt domain_of_fun[of h \ "\_. \"] by blast also from \\\_0 \ _, f ` \\ \ Memrel(\) \ \_0 \ _= f ` \\ \Ord(f`\)\ \Ord(\_0)\ \Ord(j`\)\ have "\_0 \ f`\" using Un_leD1[OF leI [OF ltI]] Un_leD1[OF le_eqI] by blast finally show "h`x < f`\" . qed ultimately have "factor_body(\,\x\\. factor(x),\)" unfolding h_def factor_body_def using ltD by (auto dest:Un_memD2 Un_leD2[OF le_eqI]) with \Ord(\)\ have "factor(\) \ \" using factor_unfold[of \] Least_le unfolding factor_rec_def by auto with \\\\\ \Ord(\)\ have "factor(\) \ \" using leI[of \] ltI[of \] by (auto dest:ltD) then show ?case by (auto elim:mem_irrefl) qed moreover have "cofinal_fun(f O fun_factor, \, Memrel(\))" proof (intro cofinal_funI) fix a assume "a \ \" with \cf_fun(j,\)\ \domain(j) = cf(\)\ obtain x where "x\cf(\)" "a \ j`x \ a = j`x" by (auto simp add:cofinal_fun_def) blast with factor_not_delta have "x \ domain(f O fun_factor)" using f_fun_factor_is_mono_map mono_map_is_fun domain_of_fun by force moreover have "a \ (f O fun_factor) `x \ a = (f O fun_factor) `x" proof - from \x\cf(\)\ factor_not_delta have "j ` x \ f ` factor(x)" using mem_not_refl factor_body_factor factor_in_delta unfolding factor_body_def by auto with \a \ j`x \ a = j`x\ have "a \ f ` factor(x) \ a = f ` factor(x)" using ltD by blast with \x\cf(\)\ show ?thesis using lam_funtype[of "cf(\)" factor] unfolding fun_factor_def by auto qed moreover note \a \ \\ moreover from calculation and \Ord(\)\ and factor_not_delta have "(f O fun_factor) `x \ \" using Limit_nonzero apply_in_codomain_Ord mono_map_is_fun[of "f O fun_factor"] f_fun_factor_is_mono_map by blast ultimately show "\x \ domain(f O fun_factor). \a, (f O fun_factor) ` x\ \ Memrel(\) \ a = (f O fun_factor) `x" by blast qed ultimately show ?thesis using fun_factor_is_mono_map f_fun_factor_is_mono_map by blast qed text\As a final observation in this part, we note that if the original cofinal map was increasing, then the factor function is also cofinal.\ lemma factor_is_cofinal: assumes "Ord(\)" "Ord(\)" "f :\ \\<^sub>< \" "f O g \ mono_map(\,r,\,Memrel(\))" "cofinal_fun(f O g,\,Memrel(\))" "g: \ \ \" shows "cf_fun(g,\)" unfolding cf_fun_def cofinal_fun_def proof fix a assume "a \ \" with \f \ mono_map(\,_,\,_)\ have "f`a \ \" using mono_map_is_fun by force with \cofinal_fun(f O g,\,_)\ obtain y where "y\\" "\f`a, (f O g) ` y\ \ Memrel(\) \ f`a = (f O g) ` y" unfolding cofinal_fun_def using domain_of_fun[OF \g:\ \ \\] by blast with \g:\ \ \\ have "\f`a, f ` (g ` y)\ \ Memrel(\) \ f`a = f ` (g ` y)" "g`y \ \" using comp_fun_apply[of g \ \ y f] by auto with assms(1-3) and \a\\\ have "\a, g ` y\ \ Memrel(\) \ a = g ` y" using Memrel_mono_map_reflects Memrel_mono_map_is_inj[of \ f \ \] inj_apply_equality[of f \ \] by blast with \y\\\ show "\x\domain(g). \a, g ` x\ \ Memrel(\) \ a = g ` x" using domain_of_fun[OF \g:\ \ \\] by blast qed subsection\Classical results on cofinalities\ text\Now the rest of the results follow in a more algebraic way. The next proof one invokes a case analysis on whether the argument is zero, a successor ordinal or a limit one; the last case being the most relevant one and is immediate from the factorization lemma.\ lemma cf_le_domain_cofinal_fun: assumes "Ord(\)" "Ord(\)" "f:\ \ \" "cf_fun(f,\)" shows "cf(\)\\" using assms proof (cases rule:Ord_cases) case 0 with \Ord(\)\ show ?thesis using Ord_0_le by simp next case (succ \) with assms obtain x where "x\\" "f`x=\" using cf_fun_succ' by blast then have "\\0" by blast let ?f="{\0,f`x\}" from \f`x=\\ have "?f:1\succ(\)" using singleton_0 singleton_fun[of 0 \] singleton_subsetI fun_weaken_type by simp with \Ord(\)\ \f`x=\\ have "cf(succ(\)) = 1" using cf_succ by simp with \\\0\ succ show ?thesis using Ord_0_lt_iff succ_leI \Ord(\)\ by simp next case (limit) with assms obtain g where "g :cf(\) \\<^sub>< \" using cofinal_fun_factorization by blast with assms show ?thesis using mono_map_imp_le by simp qed lemma cf_ordertype_cofinal: assumes "Limit(\)" "A\\" "cofinal(A,\,Memrel(\))" shows "cf(\) = cf(ordertype(A,Memrel(\)))" proof (intro le_anti_sym) + text\We show the result by proving the two inequalities.\ from \Limit(\)\ have "Ord(\)" using Limit_is_Ord by simp with \A \ \\ have "well_ord(A,Memrel(\))" using well_ord_Memrel well_ord_subset by blast then obtain f \ where "f:\\, Memrel(\)\ \ \A,Memrel(\)\" "Ord(\)" "\ = ordertype(A,Memrel(\))" using ordertype_ord_iso Ord_ordertype ord_iso_sym by blast moreover from this have "f: \ \ A" using ord_iso_is_mono_map mono_map_is_fun[of f _ "Memrel(\)"] by blast moreover from this have "function(f)" using fun_is_function by simp moreover from \f:\\, Memrel(\)\ \ \A,Memrel(\)\\ have "range(f) = A" using ord_iso_is_bij bij_is_surj surj_range by blast moreover note \cofinal(A,\,_)\ ultimately have "cf_fun(f,\)" using cofinal_range_iff_cofinal_fun by blast moreover from \Ord(\)\ obtain h where "h :cf(\) \\<^sub>< \" "cf_fun(h,\)" using cofinal_mono_map_cf by blast moreover from \Ord(\)\ have "trans(Memrel(\))" using trans_Memrel by simp moreover note \A\\\ ultimately have "cofinal_fun(f O h,\,Memrel(\))" using cofinal_comp ord_iso_is_mono_map[OF \f:\\,_\ \ \A,_\\] mono_map_is_fun mono_map_mono by blast moreover from \f:\\A\ \A\\\ \h\mono_map(cf(\),_,\,_)\ have "f O h : cf(\) \ \" using Pi_mono[of A \] comp_fun mono_map_is_fun by blast moreover note \Ord(\)\ \Ord(\)\ \\ = ordertype(A,Memrel(\))\ ultimately show "cf(\) \ cf(ordertype(A,Memrel(\)))" using cf_le_domain_cofinal_fun[of _ _ "f O h"] by (auto simp add:cf_fun_def) - (********************************************************) + text\That finishes the first inequality. Now we go the other side.\ from \f:\\, _\ \ \A,_\\ \A\\\ have "f :\ \\<^sub>< \" using mono_map_mono[OF ord_iso_is_mono_map] by simp then have "f: \ \ \" using mono_map_is_fun by simp with \cf_fun(f,\)\ \Limit(\)\ \Ord(\)\ obtain g where "g :cf(\) \\<^sub>< \" "f O g :cf(\) \\<^sub>< \" "cofinal_fun(f O g,\,Memrel(\))" using cofinal_fun_factorization by blast moreover from this have "g:cf(\)\\" using mono_map_is_fun by simp moreover note \Ord(\)\ moreover from calculation and \f :\ \\<^sub>< \\ \Ord(\)\ have "cf_fun(g,\)" using factor_is_cofinal by blast moreover note \\ = ordertype(A,Memrel(\))\ ultimately show "cf(ordertype(A,Memrel(\))) \ cf(\)" using cf_le_domain_cofinal_fun[OF _ Ord_cf mono_map_is_fun] by simp qed lemma cf_idemp: assumes "Limit(\)" shows "cf(\) = cf(cf(\))" proof - from assms obtain A where "A\\" "cofinal(A,\,Memrel(\))" "cf(\) = ordertype(A,Memrel(\))" using Limit_is_Ord cf_is_ordertype by blast with assms have "cf(\) = cf(ordertype(A,Memrel(\)))" using cf_ordertype_cofinal by simp also have "... = cf(cf(\))" using \cf(\) = ordertype(A,Memrel(\))\ by simp finally show "cf(\) = cf(cf(\))" . qed lemma cf_le_cardinal: assumes "Limit(\)" shows "cf(\) \ |\|" proof - from assms have \Ord(\)\ using Limit_is_Ord by simp then obtain f where "f \ surj(|\|,\)" using Ord_cardinal_eqpoll unfolding eqpoll_def bij_def by blast with \Ord(\)\ show ?thesis using Card_is_Ord[OF Card_cardinal] surj_is_cofinal cf_le_domain_cofinal_fun[of \] surj_is_fun by blast qed lemma regular_is_Card: notes le_imp_subset [dest] assumes "Limit(\)" "\ = cf(\)" shows "Card(\)" proof - from assms have "|\| \ \" using Limit_is_Ord Ord_cardinal_le by blast also from \\ = cf(\)\ have "\ \ cf(\)" by simp finally have "|\| \ cf(\)" . with assms show ?thesis unfolding Card_def using cf_le_cardinal by force qed lemma Limit_cf: assumes "Limit(\)" shows "Limit(cf(\))" using Ord_cf[of \, THEN Ord_cases] - \ \\<^term>\cf(\)\ being 0 or successor leads to contradiction\ + \ \\<^term>\cf(\)\ being $0$ or successor leads to contradiction\ proof (cases) case 1 with \Limit(\)\ show ?thesis using cf_zero_iff Limit_is_Ord by simp next case (2 \) moreover note \Limit(\)\ moreover from calculation have "cf(\) = 1" using cf_idemp cf_succ by fastforce ultimately show ?thesis using succ_LimitE cf_eq_one_iff Limit_is_Ord by auto qed lemma InfCard_cf: "Limit(\) \ InfCard(cf(\))" using regular_is_Card cf_idemp Limit_cf nat_le_Limit Limit_cf unfolding InfCard_def by simp lemma cf_le_cf_fun: notes [dest] = Limit_is_Ord assumes "cf(\) \ \" "Limit(\)" shows "\f. f:\ \ \ \ cf_fun(f, \)" proof - note assms moreover from this obtain h where h_cofinal_mono: "cf_fun(h,\)" "h :cf(\) \\<^sub>< \" "h : cf(\) \ \" using cofinal_mono_map_cf mono_map_is_fun by force moreover from calculation obtain g where "g \ inj(cf(\), \)" using le_imp_lepoll by blast from this and calculation(2,3,5) obtain f where "f \ surj(\, cf(\))" "f: \ \ cf(\)" using inj_imp_surj[OF _ Limit_has_0[THEN ltD]] surj_is_fun Limit_cf by blast moreover from this have "cf_fun(f,cf(\))" using surj_is_cofinal by simp moreover note h_cofinal_mono \Limit(\)\ moreover from calculation have "cf_fun(h O f,\)" using cf_fun_comp by blast moreover from calculation have "h O f \ \ -> \" using comp_fun by simp ultimately show ?thesis by blast qed lemma Limit_cofinal_fun_lt: notes [dest] = Limit_is_Ord assumes "Limit(\)" "f: \ \ \" "cf_fun(f,\)" "n\\" shows "\\\\. n < f`\" proof - from \Limit(\)\ \n\\\ have "succ(n) \ \" using Limit_has_succ[OF _ ltI, THEN ltD] by auto moreover note \f: \ \ _\ moreover from this have "domain(f) = \" using domain_of_fun by simp moreover note \cf_fun(f,\)\ ultimately obtain \ where "\ \ \" "succ(n) \ f`\ \ succ(n) = f `\" using cf_funD[THEN cofinal_funD] by blast moreover from this consider (1) "succ(n) \ f`\" | (2) "succ(n) = f `\" by blast then have "n < f`\" proof (cases) case 1 moreover have "n \ succ(n)" by simp moreover note \Limit(\)\ \f: \ \ _\ \\ \ \\ moreover from this have "Ord(f ` \)" using apply_type[of f \ "\_. \", THEN [2] Ord_in_Ord] by blast ultimately show ?thesis using Ord_trans[of n "succ(n)" "f ` \"] ltI by blast next case 2 have "n \ f ` \" by (simp add:2[symmetric]) with \Limit(\)\ \f: \ \ _\ \\ \ \\ show ?thesis using ltI apply_type[of f \ "\_. \", THEN [2] Ord_in_Ord] by blast qed ultimately show ?thesis by blast qed context includes Ord_dests Aleph_dests Aleph_intros Aleph_mem_dests mono_map_rules begin text\We end this section by calculating the cofinality of Alephs, for the zero and limit case. The successor case depends on $\AC$.\ lemma cf_nat: "cf(\) = \" using Limit_nat[THEN InfCard_cf] cf_le_cardinal[of \] Card_nat[THEN Card_cardinal_eq] le_anti_sym unfolding InfCard_def by auto lemma cf_Aleph_zero: "cf(\\<^bsub>0\<^esub>) = \\<^bsub>0\<^esub>" using cf_nat unfolding Aleph_def by simp lemma cf_Aleph_Limit: assumes "Limit(\)" shows "cf(\\<^bsub>\\<^esub>) = cf(\)" proof - note \Limit(\)\ moreover from this have "(\x\\. \\<^bsub>x\<^esub>) : \ \ \\<^bsub>\\<^esub>" (is "?f : _ \ _") using lam_funtype[of _ Aleph] fun_weaken_type[of _ _ _ "\\<^bsub>\\<^esub>"] by blast moreover from \Limit(\)\ have "x \ y \ \\<^bsub>x\<^esub> \ \\<^bsub>y\<^esub>" if "x \ \" "y \ \" for x y using that Ord_in_Ord[of \] Ord_trans[of _ _ \] by blast ultimately have "?f \ mono_map(\,Memrel(\),\\<^bsub>\\<^esub>, Memrel(\\<^bsub>\\<^esub>))" by auto with \Limit(\)\ have "?f \ \\, Memrel(\)\ \ \?f``\, Memrel(\\<^bsub>\\<^esub>)\" using mono_map_imp_ord_iso_Memrel[of \ "\\<^bsub>\\<^esub>" ?f] Card_Aleph (* Already an intro rule, but need it explicitly *) by blast then have "converse(?f) \ \?f``\, Memrel(\\<^bsub>\\<^esub>)\ \ \\, Memrel(\)\" using ord_iso_sym by simp with \Limit(\)\ have "ordertype(?f``\, Memrel(\\<^bsub>\\<^esub>)) = \" using ordertype_eq[OF _ well_ord_Memrel] ordertype_Memrel by auto moreover from \Limit(\)\ have "cofinal(?f``\, \\<^bsub>\\<^esub>, Memrel(\\<^bsub>\\<^esub>))" unfolding cofinal_def proof (standard, intro ballI) fix a assume "a\\\<^bsub>\\<^esub>" "\\<^bsub>\\<^esub> = (\i<\. \\<^bsub>i\<^esub>)" moreover from this obtain i where "i<\" "a\\\<^bsub>i\<^esub>" by auto moreover from this and \Limit(\)\ have "Ord(i)" using ltD Ord_in_Ord by blast moreover from \Limit(\)\ and calculation have "succ(i) \ \" using ltD by auto moreover from this and \Ord(i)\ have "\\<^bsub>i\<^esub> < \\<^bsub>succ(i)\<^esub>" by (auto) ultimately have "\a,\\<^bsub>i\<^esub>\ \ Memrel(\\<^bsub>\\<^esub>)" using ltD by (auto dest:Aleph_increasing) moreover from \i<\\ have "\\<^bsub>i\<^esub> \ ?f``\" using ltD apply_in_image[OF \?f : _ \ _\] by auto ultimately show "\x\?f `` \. \a, x\ \ Memrel(\\<^bsub>\\<^esub>) \ a = x" by blast qed moreover note \?f: \ \ \\<^bsub>\\<^esub>\ \Limit(\)\ ultimately show "cf(\\<^bsub>\\<^esub>) = cf(\)" using cf_ordertype_cofinal[OF Limit_Aleph Image_sub_codomain, of \ ?f \ \ ] Limit_is_Ord by simp qed -end (* includes *) +end \ \includes\ 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,1079 +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}\ + \ \almost the same proof as @{thm [source] 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}\ +\ \Same proof as Paulson's @{thm [source] 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/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,488 +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 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.\ + @{thm [source] 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_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(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_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 with assms show ?thesis 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(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,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 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' 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/Demonstrations.thy b/thys/Independence_CH/Demonstrations.thy --- a/thys/Independence_CH/Demonstrations.thy +++ b/thys/Independence_CH/Demonstrations.thy @@ -1,168 +1,168 @@ section\Some demonstrations\ theory Demonstrations imports Definitions_Main begin text\The following theory is only intended to explore some details of the formalization and to show the appearance of relevant internalized formulas. It is \<^bold>\not\ intended as the entry point of the session. For that purpose, consult \<^theory>\Independence_CH.Definitions_Main\\ locale Demo = M_trivial + M_AC + fixes t\<^sub>1 t\<^sub>2 assumes ts_in_nat[simp]: "t\<^sub>1\\" "t\<^sub>2\\" and power_infty: "power_ax(M)" "M(\)" begin text\The next fake lemma is intended to explore the instances of the axiom schemes that are needed to build our forcing models. They are categorized as plain replacements (using \<^term>\strong_replacement\), “lambda-replacements” using a higher order function, replacements to perform transfinite and general well-founded recursion (using \<^term>\transrec_replacement\ and \<^term>\wfrec_replacement\ respectively) and for the construction of fixpoints (using \<^term>\iterates_replacement\). Lastly, separations instances.\ lemma assumes sorried_replacements: "\P. strong_replacement(M,P)" "\F. lam_replacement(M,F)" "\Q S. iterates_replacement(M,Q,S)" "\Q S. wfrec_replacement(M,Q,S)" "\Q S. transrec_replacement(M,Q,S)" and sorried_separations: "\Q. separation(M,Q)" shows "M_master(M)" apply unfold_locales apply (simp_all add: sorried_replacements(1-2) sorried_separations power_infty) oops \ \NOTE: Only for pretty-printing purposes, overrides previous fundamental notations\ no_notation mem (infixl \\\ 50) no_notation conj (infixr \\\ 35) no_notation disj (infixr \\\ 30) no_notation iff (infixr \\\ 25) no_notation imp (infixr \\\ 25) no_notation not (\\ _\ [40] 40) no_notation All (\'(\_')\) no_notation Ex (\'(\_')\) no_notation Member (\\_ \/ _\\) no_notation Equal (\\_ =/ _\\) no_notation Nand (\\\'(_ \/ _')\\) no_notation And (\\_ \/ _\\) no_notation Or (\\_ \/ _\\) no_notation Iff (\\_ \/ _\\) no_notation Implies (\\_ \/ _\\) no_notation Neg (\\\_\\) no_notation Forall (\'(\\(/_)\')\) no_notation Exists (\'(\\(/_)\')\) notation Member (infixl \\\ 50) notation Equal (infixl \\\ 50) notation Nand (\\'(_ \/ _')\) notation And (infixr \\\ 35) notation Or (infixr \\\ 30) notation Iff (infixr \\\ 25) notation Implies (infixr \\\ 25) notation Neg (\\ _\ [40] 40) notation Forall (\'(\_')\) notation Exists (\'(\_')\) lemma "forces(t\<^sub>1\t\<^sub>2) = (0 \ 1 \ forces_mem_fm(1, 2, 0, t\<^sub>1+\<^sub>\4, t\<^sub>2+\<^sub>\4))" unfolding forces_def by simp (* \ \Prefix abbreviated notation\ notation Member (\M\) notation Equal (\Eq\) notation Nand (\Na\) notation And (\A\) notation Or (\O\) notation Iff (\If\) notation Implies (\Im\) notation Neg (\Ne\) notation Forall (\Fo\) notation Exists (\Ex\) *) (* forces_mem_fm(1, 2, 0, t\<^sub>1+\<^sub>\4, t\<^sub>1+\<^sub>\4) = forces_mem_fm(1, 2, 0, succ(succ(succ(succ(t\<^sub>1)))), succ(succ(succ(succ(t\<^sub>2))))) = \ *) definition forces_0_mem_1 where "forces_0_mem_1\forces_mem_fm(1,2,0,t\<^sub>1+\<^sub>\4,t\<^sub>2+\<^sub>\4)" thm forces_0_mem_1_def[ unfolded frc_at_fm_def ftype_fm_def name1_fm_def name2_fm_def snd_snd_fm_def hcomp_fm_def ecloseN_fm_def eclose_n1_fm_def eclose_n2_fm_def is_eclose_fm_def mem_eclose_fm_def eclose_n_fm_def is_If_fm_def least_fm_def Replace_fm_def Collect_fm_def fm_definitions,simplified] - (* NOTE: in view of the above, @{thm fm_definitions} might be incomplete *) + (* NOTE: in view of the above, @{thm [source] fm_definitions} might be incomplete *) named_theorems incr_bv_new_simps schematic_goal incr_bv_Neg(* [incr_bv_new_simps] *): "mem(n,\) \ mem(\,formula) \ incr_bv(Neg(\))`n = ?x" unfolding Neg_def by simp schematic_goal incr_bv_Exists [incr_bv_new_simps]: "mem(n,\) \ mem(\,formula) \ incr_bv(Exists(\))`n = ?x" unfolding Exists_def by (simp add: incr_bv_Neg) (* schematic_goal incr_bv_And [incr_bv_new_simps]: "mem(n,\) \ mem(\,formula) \mem(\,formula)\ incr_bv(And(\,\))`n = ?x" unfolding And_def by (simp add: incr_bv_Neg) schematic_goal incr_bv_Or [incr_bv_new_simps]: "mem(n,\) \ mem(\,formula) \mem(\,formula)\ incr_bv(Or(\,\))`n = ?x" unfolding Or_def by (simp add: incr_bv_Neg) schematic_goal incr_bv_Implies [incr_bv_new_simps]: "mem(n,\) \ mem(\,formula) \mem(\,formula)\ incr_bv(Implies(\,\))`n = ?x" unfolding Implies_def by (simp add: incr_bv_Neg) *) \ \The two renamings involved in the definition of \<^term>\forces\ depend on the recursive function \<^term>\incr_bv\. Here we have an apparently exponential bottleneck, since all the propositional connectives (even \<^term>\Neg\) duplicate the appearances of \<^term>\incr_bv\. Not even the double negation of an atomic formula can be managed by the system (in version 2021-1).\ (* schematic_goal "forces(\\0\1) = ?x" unfolding forces_def Neg_def by (simp add:ren_forces_nand_def ren_forces_forall_def frc_at_fm_def ftype_fm_def name1_fm_def name2_fm_def snd_snd_fm_def hcomp_fm_def ecloseN_fm_def eclose_n1_fm_def eclose_n2_fm_def is_eclose_fm_def mem_eclose_fm_def eclose_n_fm_def is_If_fm_def least_fm_def Collect_fm_def fm_definitions incr_bv_Neg incr_bv_Exists) (* exception Size raised (line 183 of "./basis/LibrarySupport.sml") *) *) (* declare is_ContHyp_fm_def[fm_definitions del] thm is_ContHyp_fm_def[unfolded is_eclose_fm_def mem_eclose_fm_def eclose_n_fm_def is_If_fm_def least_fm_def Replace_fm_def Collect_fm_def fm_definitions, simplified] *) end \ \\<^locale>\Demo\\ end \ No newline at end of file diff --git a/thys/Independence_CH/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,1011 +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 +\ \same as @{thm [source] 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" 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(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" 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_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 \ No newline at end of file diff --git a/thys/Transitive_Models/Aleph_Relative.thy b/thys/Transitive_Models/Aleph_Relative.thy --- a/thys/Transitive_Models/Aleph_Relative.thy +++ b/thys/Transitive_Models/Aleph_Relative.thy @@ -1,464 +1,464 @@ theory Aleph_Relative imports CardinalArith_Relative begin definition HAleph :: "[i,i] \ i" where "HAleph(i,r) \ if(\(Ord(i)),i,if(i=0, nat, if(\Limit(i) \ i\0, csucc(r`( \ i )), \j\i. r`j)))" reldb_add functional "Limit" "Limit" relationalize "Limit" "is_Limit" external synthesize "is_Limit" from_definition arity_theorem for "is_Limit_fm" relativize functional "HAleph" "HAleph_rel" relationalize "HAleph_rel" "is_HAleph" synthesize "is_HAleph" from_definition assuming "nonempty" arity_theorem intermediate for "is_HAleph_fm" lemma arity_is_HAleph_fm_aux: assumes "i \ nat" "r \ nat" \ \NOTE: assumptions are \<^bold>\not\ used, but if omitted, next lemma fails!\ shows "arity(Replace_fm(8 +\<^sub>\ i, \10 +\<^sub>\ r`0 is 1\, 3)) = 9 +\<^sub>\ i \ pred(pred(11 +\<^sub>\ r))" using arity_Replace_fm[of "\ (10+\<^sub>\r)`0 is 1\" "8+\<^sub>\i" 3 "(11+\<^sub>\r) \ 1 \ 2"] ord_simp_union by (auto simp:arity) lemma arity_is_HAleph_fm[arity]: assumes "i \ nat" "r \ nat" "l \ nat" shows "arity(is_HAleph_fm(i, r, l)) = succ(i) \ succ(l) \ succ(r)" using assms pred_Un arity_is_HAleph_fm_aux arity_is_HAleph_fm' by auto definition Aleph' :: "i => i" where "Aleph'(a) == transrec(a,\i r. HAleph(i,r))" relativize functional "Aleph'" "Aleph_rel" relationalize "Aleph_rel" "is_Aleph" txt\The extra assumptions \<^term>\a < length(env)\ and \<^term>\c < length(env)\ in this schematic goal (and the following results on synthesis that - depend on it) are imposed by @{thm is_transrec_iff_sats}.\ + depend on it) are imposed by @{thm [source] is_transrec_iff_sats}.\ schematic_goal sats_is_Aleph_fm_auto: "a \ nat \ c \ nat \ env \ list(A) \ a < length(env) \ c < length(env) \ 0 \ A \ is_Aleph(##A, nth(a, env), nth(c, env)) \ A, env \ ?fm(a, c)" unfolding is_Aleph_def proof (rule is_transrec_iff_sats, rule_tac [1] is_HAleph_iff_sats) fix a0 a1 a2 a3 a4 a5 a6 a7 let ?env' = "Cons(a0, Cons(a1, Cons(a2, Cons(a3, Cons(a4, Cons(a5, Cons(a6, Cons(a7, env))))))))" show "nth(2, ?env') = a2" "nth(1, ?env') = a1" "nth(0, ?env') = a0" "nth(c, env) = nth(c, env)" by simp_all qed simp_all synthesize_notc "is_Aleph" from_schematic notation is_Aleph_fm (\\\'(_') is _\\) lemma is_Aleph_fm_type [TC]: "a \ nat \ c \ nat \ is_Aleph_fm(a, c) \ formula" unfolding is_Aleph_fm_def by simp lemma sats_is_Aleph_fm: assumes "f\nat" "r\nat" "env \ list(A)" "0\A" "f < length(env)" "r< length(env)" shows "is_Aleph(##A, nth(f, env), nth(r, env)) \ A, env \ is_Aleph_fm(f,r)" using assms sats_is_Aleph_fm_auto unfolding is_Aleph_def is_Aleph_fm_def by simp lemma is_Aleph_iff_sats [iff_sats]: assumes "nth(f, env) = fa" "nth(r, env) = ra" "f < length(env)" "r< length(env)" "f \ nat" "r \ nat" "env \ list(A)" "0\A" shows "is_Aleph(##A,fa,ra) \ A, env \ is_Aleph_fm(f,r)" using assms sats_is_Aleph_fm[of f r env A] by simp arity_theorem for "is_Aleph_fm" lemma (in M_cardinal_arith_jump) is_Limit_iff: assumes "M(a)" shows "is_Limit(M,a) \ Limit(a)" unfolding is_Limit_def Limit_def using lt_abs transM[OF ltD \M(a)\] assms by auto lemma HAleph_eq_Aleph_recursive: "Ord(i) \ HAleph(i,r) = (if i = 0 then nat else if \j. i = succ(j) then csucc(r ` (THE j. i = succ(j))) else \j (\succ(j)) = j" for j using Ord_Union_succ_eq by simp moreover from \Ord(i)\ have "(\j. i = succ(j)) \ \Limit(i) \ i \ 0" using Ord_cases_disj by auto ultimately show ?thesis unfolding HAleph_def OUnion_def by auto qed lemma Aleph'_eq_Aleph: "Ord(a) \ Aleph'(a) = Aleph(a)" unfolding Aleph'_def Aleph_def transrec2_def using HAleph_eq_Aleph_recursive by (intro transrec_equal_on_Ord) auto reldb_rem functional "Aleph'" reldb_rem relational "is_Aleph" reldb_add functional "Aleph" "Aleph_rel" reldb_add relational "Aleph" "is_Aleph" abbreviation Aleph_r :: "[i,i\o] \ i" (\\\<^bsub>_\<^esub>\<^bsup>_\<^esup>\) where "Aleph_r(a,M) \ Aleph_rel(M,a)" abbreviation Aleph_r_set :: "[i,i] \ i" (\\\<^bsub>_\<^esub>\<^bsup>_\<^esup>\) where "Aleph_r_set(a,M) \ Aleph_rel(##M,a)" lemma Aleph_rel_def': "Aleph_rel(M,a) \ transrec(a, \i r. HAleph_rel(M, i, r))" unfolding Aleph_rel_def . lemma succ_mem_Limit: "Limit(j) \ i \ j \ succ(i) \ j" using Limit_has_succ[THEN ltD] ltI Limit_is_Ord by auto locale M_pre_aleph = M_eclose + M_cardinal_arith_jump + assumes haleph_transrec_replacement: "M(a) \ transrec_replacement(M,is_HAleph(M),a)" begin lemma aux_ex_Replace_funapply: assumes "M(a)" "M(f)" shows "\x[M]. is_Replace(M, a, \j y. f ` j = y, x)" proof - have "{f`j . j\a} = {y . j\a , f ` j=y}" "{y . j\a , f ` j=y} = {y . j\a , y =f ` j}" by auto moreover note assms moreover from calculation have "x \ a \ y = f `x \ M(y)" for x y using transM[OF _ \M(a)\] by auto moreover from assms have "M({f`j . j\a})" using transM[OF _ \M(a)\] RepFun_closed[OF apply_replacement] by simp ultimately have 2:"is_Replace(M, a, \j y. y = f ` j, {f`j . j\a})" using Replace_abs[of _ _ "\j y. y = f ` j",OF \M(a)\,THEN iffD2] by auto with \M({f`j . j\a})\ show ?thesis using is_Replace_cong[of _ _ M "\j y. y = f ` j" "\j y. f ` j = y", THEN iffD1,OF _ _ _ 2] by auto qed lemma is_HAleph_zero: assumes "M(f)" shows "is_HAleph(M,0,f,res) \ res = nat" unfolding is_HAleph_def using Ord_0 If_abs is_Limit_iff is_csucc_iff assms aux_ex_Replace_funapply by auto lemma is_HAleph_succ: assumes "M(f)" "M(x)" "Ord(x)" "M(res)" shows "is_HAleph(M,succ(x),f,res) \ res = csucc_rel(M,f`x)" unfolding is_HAleph_def using assms is_Limit_iff is_csucc_iff aux_ex_Replace_funapply If_abs Ord_Union_succ_eq by simp lemma is_HAleph_limit: assumes "M(f)" "M(x)" "Limit(x)" "M(res)" shows "is_HAleph(M,x,f,res) \ res = (\{y . i\x ,M(i) \ M(y) \ y = f`i})" proof - from assms have "univalent(M, x, \j y. y = f ` j )" "(\xa y. xa \ x \ f ` xa = y \ M(y))" "{y . x \ x, f ` x = y} = {y . i\x ,M(i) \ M(y) \ y = f`i}" using univalent_triv[of M x "\j .f ` j"] transM[OF _ \M(x)\] by auto moreover from this have "univalent(M, x, \j y. f ` j = y )" by (rule_tac univalent_cong[of x x M " \j y. y = f ` j" " \j y. f ` j=y",THEN iffD1], auto) moreover from this have "univalent(M, x, \j y. M(j) \ M(y) \ f ` j = y )" by auto ultimately show ?thesis unfolding is_HAleph_def using assms is_Limit_iff Limit_is_Ord zero_not_Limit If_abs is_csucc_iff Replace_abs apply_replacement by auto qed lemma is_HAleph_iff: assumes "M(a)" "M(f)" "M(res)" shows "is_HAleph(M, a, f, res) \ res = HAleph_rel(M, a, f)" proof(cases "Ord(a)") case True note Ord_cases[OF \Ord(a)\] then show ?thesis proof(cases ) case 1 with True assms show ?thesis using is_HAleph_zero unfolding HAleph_rel_def by simp next case (2 j) with True assms show ?thesis using is_HAleph_succ Ord_Union_succ_eq unfolding HAleph_rel_def by simp next case 3 with assms show ?thesis using is_HAleph_limit zero_not_Limit Limit_is_Ord unfolding HAleph_rel_def by auto qed next case False then have "\Limit(a)" "a\0" "\ x . Ord(x) \ a\succ(x)" using Limit_is_Ord by auto with False show ?thesis unfolding is_HAleph_def HAleph_rel_def using assms is_Limit_iff If_abs is_csucc_iff aux_ex_Replace_funapply by auto qed lemma HAleph_rel_closed [intro,simp]: assumes "function(f)" "M(a)" "M(f)" shows "M(HAleph_rel(M,a,f))" unfolding HAleph_rel_def using assms apply_replacement by simp lemma Aleph_rel_closed[intro, simp]: assumes "Ord(a)" "M(a)" shows "M(Aleph_rel(M,a))" proof - have "relation2(M, is_HAleph(M), HAleph_rel(M))" unfolding relation2_def using is_HAleph_iff assms by simp moreover have "\x[M]. \g[M]. function(g) \ M(HAleph_rel(M, x, g))" using HAleph_rel_closed by simp moreover note assms ultimately show ?thesis unfolding Aleph_rel_def using transrec_closed[of "is_HAleph(M)" a "HAleph_rel(M)"] haleph_transrec_replacement by simp qed lemma Aleph_rel_zero: "\\<^bsub>0\<^esub>\<^bsup>M\<^esup> = nat" using def_transrec [OF Aleph_rel_def',of _ 0] unfolding HAleph_rel_def by simp lemma Aleph_rel_succ: "Ord(\) \ M(\) \ \\<^bsub>succ(\)\<^esub>\<^bsup>M\<^esup> = (\\<^bsub>\\<^esub>\<^bsup>M\<^esup>\<^sup>+)\<^bsup>M\<^esup>" using Ord_Union_succ_eq by (subst def_transrec [OF Aleph_rel_def']) (simp add:HAleph_rel_def) lemma Aleph_rel_limit: assumes "Limit(\)" "M(\)" shows "\\<^bsub>\\<^esub>\<^bsup>M\<^esup> = \{\\<^bsub>j\<^esub>\<^bsup>M\<^esup> . j \ \}" proof - note trans=transM[OF _ \M(\)\] from \M(\)\ have "\\<^bsub>\\<^esub>\<^bsup>M\<^esup> = HAleph_rel(M, \, \x\\. \\<^bsub>x\<^esub>\<^bsup>M\<^esup>)" using def_transrec [OF Aleph_rel_def',of M \] by simp also have "... = \{a . j \ \, M(a) \ a = \\<^bsub>j\<^esub>\<^bsup>M\<^esup>}" unfolding HAleph_rel_def using assms zero_not_Limit Limit_is_Ord trans by auto also have "... = \{\\<^bsub>j\<^esub>\<^bsup>M\<^esup> . j \ \}" using Aleph_rel_closed[OF _ trans] Ord_in_Ord Limit_is_Ord[OF \Limit(\)\] by auto finally show ?thesis . qed lemma is_Aleph_iff: assumes "Ord(a)" "M(a)" "M(res)" shows "is_Aleph(M, a, res) \ res = \\<^bsub>a\<^esub>\<^bsup>M\<^esup>" proof - have "relation2(M, is_HAleph(M), HAleph_rel(M))" unfolding relation2_def using is_HAleph_iff assms by simp moreover have "\x[M]. \g[M]. function(g) \ M(HAleph_rel(M, x, g))" using HAleph_rel_closed by simp ultimately show ?thesis using assms transrec_abs haleph_transrec_replacement unfolding is_Aleph_def Aleph_rel_def by simp qed end \ \\<^locale>\M_pre_aleph\\ locale M_aleph = M_pre_aleph + assumes aleph_rel_separation: "Ord(x) \ M(x) \ separation(M, \y. \z\x. y = \\<^bsub>z\<^esub>\<^bsup>M\<^esup>)" begin lemma Aleph_rel_cont: "Limit(l) \ M(l) \ \\<^bsub>l\<^esub>\<^bsup>M\<^esup> = (\i\<^bsub>i\<^esub>\<^bsup>M\<^esup>)" using Limit_is_Ord Aleph_rel_limit by (simp add:OUnion_def) lemma Ord_Aleph_rel: assumes "Ord(a)" shows "M(a) \ Ord(\\<^bsub>a\<^esub>\<^bsup>M\<^esup>)" using \Ord(a)\ proof(induct a rule:trans_induct3) case 0 show ?case using Aleph_rel_zero by simp next case (succ x) with \Ord(x)\ have "M(x)" "Ord(\\<^bsub>x\<^esub>\<^bsup>M\<^esup>)" by simp_all with \Ord(x)\ have "Ord(csucc_rel(M,\\<^bsub>x\<^esub>\<^bsup>M\<^esup>))" using Card_rel_is_Ord Card_rel_csucc_rel by simp with \Ord(x)\ \M(x)\ show ?case using Aleph_rel_succ by simp next case (limit x) note trans=transM[OF _ \M(x)\] from limit have "\\<^bsub>x\<^esub>\<^bsup>M\<^esup> = (\i\x. \\<^bsub>i\<^esub>\<^bsup>M\<^esup>)" using Aleph_rel_cont OUnion_def Limit_is_Ord by auto with limit show ?case using Ord_UN trans by auto qed lemma Aleph_rel_increasing: assumes "a < b" and types: "M(a)" "M(b)" shows "\\<^bsub>a\<^esub>\<^bsup>M\<^esup> < \\<^bsub>b\<^esub>\<^bsup>M\<^esup>" proof - { fix x from assms have "Ord(b)" by (blast intro: lt_Ord2) moreover assume "M(x)" moreover note \M(b)\ ultimately have "x < b \ \\<^bsub>x\<^esub>\<^bsup>M\<^esup> < \\<^bsub>b\<^esub>\<^bsup>M\<^esup>" proof (induct b arbitrary: x rule: trans_induct3) case 0 thus ?case by simp next case (succ b) then show ?case using Card_rel_csucc_rel Ord_Aleph_rel Ord_Union_succ_eq lt_csucc_rel lt_trans[of _ "\\<^bsub>b\<^esub>\<^bsup>M\<^esup>" "csucc\<^bsup>M\<^esup>(\\<^bsub>b\<^esub>\<^bsup>M\<^esup>)"] by (subst (2) def_transrec[OF Aleph_rel_def']) (auto simp add: le_iff HAleph_rel_def) next case (limit l) then have sc: "succ(x) < l" by (blast intro: Limit_has_succ) then have "\\<^bsub>x\<^esub>\<^bsup>M\<^esup> < (\j\<^bsub>j\<^esub>\<^bsup>M\<^esup>)" using limit Ord_Aleph_rel Ord_OUN proof(rule_tac OUN_upper_lt,blast intro: Card_rel_is_Ord ltD lt_Ord) from \x \Limit(l)\ have "Ord(x)" using Limit_is_Ord Ord_in_Ord by (auto dest!:ltD) with \M(x)\ show "\\<^bsub>x\<^esub>\<^bsup>M\<^esup> < \\<^bsub>succ(x)\<^esub>\<^bsup>M\<^esup>" using Card_rel_csucc_rel Ord_Aleph_rel lt_csucc_rel ltD[THEN [2] Ord_in_Ord] succ_in_MI[OF \M(x)\] Aleph_rel_succ[of x] by (simp) next from \M(l)\ \Limit(l)\ show "Ord(\j\<^bsub>j\<^esub>\<^bsup>M\<^esup>)" using Ord_Aleph_rel lt_Ord Limit_is_Ord Ord_in_Ord by (rule_tac Ord_OUN) (auto dest:transM ltD intro!:Ord_Aleph_rel) qed then show ?case using limit Aleph_rel_cont by simp qed } with types assms show ?thesis by simp qed lemma Card_rel_Aleph_rel [simp, intro]: assumes "Ord(a)" and types: "M(a)" shows "Card\<^bsup>M\<^esup>(\\<^bsub>a\<^esub>\<^bsup>M\<^esup>)" using assms proof (induct rule:trans_induct3) case 0 then show ?case using Aleph_rel_zero Card_rel_nat by simp next case (succ x) then show ?case using Card_rel_csucc_rel Ord_Aleph_rel Aleph_rel_succ by simp next case (limit x) moreover from this have "Ord(x)" using Limit_is_Ord by simp from this have "{y . z \ x, M(y) \ M(z) \ y = \\<^bsub>z\<^esub>\<^bsup>M\<^esup>} = {y . z \ x, M(y) \ M(z) \ Ord(z) \ y = \\<^bsub>z\<^esub>\<^bsup>M\<^esup>}" using Ord_in_Ord by simp moreover from \Ord(x)\ have "{y . z \ x, M(y) \ M(z) \ Ord(z) \ y = \\<^bsub>z\<^esub>\<^bsup>M\<^esup>} = {y . z \ x, M(z) \ y = \\<^bsub>z\<^esub>\<^bsup>M\<^esup>}" using Ord_in_Ord by blast moreover from \Ord(x)\ \M(x)\ have "{y . z \ x, M(z) \ y = \\<^bsub>z\<^esub>\<^bsup>M\<^esup>} \ \\<^bsub>x\<^esub>\<^bsup>M\<^esup>" using Aleph_rel_increasing by (auto dest:ltD transM intro:ltI) with calculation have "{y . z \ x, M(z) \ y = \\<^bsub>z\<^esub>\<^bsup>M\<^esup>} = {y \ \\<^bsub>x\<^esub>\<^bsup>M\<^esup> . (\z \ x. y = \\<^bsub>z\<^esub>\<^bsup>M\<^esup>)}" by (blast dest:transM) with calculation have "{y . z \ x, M(y) \ M(z) \ y = \\<^bsub>z\<^esub>\<^bsup>M\<^esup>} = {y \ \\<^bsub>x\<^esub>\<^bsup>M\<^esup> . (\z \ x. y = \\<^bsub>z\<^esub>\<^bsup>M\<^esup>)}" by simp moreover from \Ord(x)\ \M(x)\ have "M({y \ \\<^bsub>x\<^esub>\<^bsup>M\<^esup> . (\z \ x. y = \\<^bsub>z\<^esub>\<^bsup>M\<^esup>)})" using aleph_rel_separation by simp ultimately show ?case using Ord_Aleph_rel Card_nat Limit_is_Ord Card_relI by (subst def_transrec [OF Aleph_rel_def']) (auto simp add:HAleph_rel_def) qed lemmas nat_subset_Aleph_rel_1 = Ord_lt_subset[OF Ord_Aleph_rel[of 1] Aleph_rel_increasing[of 0 1,simplified],simplified] end \ \\<^locale>\M_aleph\\ 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,1611 +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 "\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) 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 by (simp add:absolut csquare_lam_eq_lam[unfolded csquare_lam_def]) end \ \\<^locale>\M_pre_cardinal_arith\\ 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>'(_')\) 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) text\We 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 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 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 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 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 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\ 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\\ 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 \<^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\ (* reldb_add functional relational "Ord" "ordinal" *) relativize functional "ord_iso" "ord_iso_rel" external \ \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:" (\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,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,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 . \x,y\\r \ x\A" 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 \r\Pow(A\A)\ show ?thesis 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) 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,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)" "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 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(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" 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 is_ordertype_replacement : "strong_replacement(M,\ x z . \y[M]. is_order_body(M,x,y) \ z=\x,y\)" and 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 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 - 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 "M(jump_cardinal_body(Pow\<^bsup>M\<^esup>(X \ X), X))" if "M(X)" for X using that by simp moreover from calculation have 1:"lam_replacement(M, \x . {snd(p) . p \ lam_jump_cardinal_body(x)})" using lam_replacement_cong by auto ultimately 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 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)})" 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, THEN bij_converse_bij]) 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}\ + \ \Note the weakened hypothesis @{thm [source] 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) 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) 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) 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 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) 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) 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) 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_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"] 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)"],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) 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) 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) using def_Pow_rel 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) using Memrel_closed 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)})" 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_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) 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] 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 (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)+ 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)+ 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 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) 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,436 +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 lam_replacement_identity lam_replacement_hcomp2[of _ _ minimum] by simp lemmas surj_imp_inj_replacement4 = lam_replacement_minimum_vimage[unfolded lam_replacement_def] 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}\ + \ \The next assumption is required by @{thm [source] 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_Relative.thy b/thys/Transitive_Models/Cardinal_Relative.thy --- a/thys/Transitive_Models/Cardinal_Relative.thy +++ b/thys/Transitive_Models/Cardinal_Relative.thy @@ -1,1487 +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))" 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_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 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: 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] 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, 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 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 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 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 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) 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 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\\ 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: 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 +absoluteness results (@{thm [source] 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\ 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] 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 \<^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\ 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}.\ +bijections of finite sets (cf. @{thm [source] lesspoll_rel_succ_imp_lepoll_rel}.\ 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\ +@{thm [source] Finite_Fin_lemma} and @{thm [source] Fin_lemma} should follow\ 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 \ 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,638 +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 _\\) 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 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}\ +text\The next "def\_" result really corresponds to @{thm [source] 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\\ 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\\ 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 \ \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 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 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\ 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>\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\\ 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_Function.thy b/thys/Transitive_Models/Discipline_Function.thy --- a/thys/Transitive_Models/Discipline_Function.thy +++ b/thys/Transitive_Models/Discipline_Function.thy @@ -1,928 +1,928 @@ section\Basic relativization of function spaces\ theory Discipline_Function imports Arities begin subsection\Discipline for \<^term>\fst\ and \<^term>\snd\\ 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 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\\ 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.\ 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. + @{thm [source] exI allI} to @{thm [source] 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.\ + @{thm [source] 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\\ subsection\Discipline for \<^term>\inj\\ 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\\ 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\\ subsection\Discipline for \<^term>\surj\\ 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\\ 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\.\ +text\Note: @{thm [source] 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" 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\\ 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 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\\ 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 \ \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\\ 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/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,2508 +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 +of separation and six instances of lambda replacements (ie, replacement of +the form $\lambda x y.\ y=\langle x, f(x) \rangle$) and 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" 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" 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 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 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 minimum_in_Union: 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_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_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)))" 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 +\ \This lemma is similar to @{thm [source] strong_lam_replacement_imp_strong_replacement} +and @{thm [source] 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_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)))" 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 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)\)" 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 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 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_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\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(_)"] 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 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 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] transM[of _ "g(_)"] rep_closed lam_replacement_product[OF lam_replacement_fst] lam_replacement_hcomp[OF lam_replacement_snd] by simp 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} = {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_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 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(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 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 \ \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] 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 ?thesis by(rule_tac lam_replacement_cong,auto,subst eq[symmetric],rule_tac comp_closed,auto) 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 \ \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_product lam_replacement_snd lam_replacement_hcomp lam_replacement_sing_fun 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 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. (\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_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)" shows "lam_replacement(M, drSR_Y(B,D,A))" 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)" shows "M(drSR_Y(B,D,A,f))" 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. 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 by simp_all then show ?thesis unfolding dC_F_def using assms lam_replacement_Collect'[of A "\ d x . domain(x) = d"] separation_eq[OF _ lam_replacement_domain _ lam_replacement_constant] by simp 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"] separation_eq[OF _ lam_replacement_domain _ lam_replacement_constant] by simp lemma lam_replacement_Collect_ball_Pair: assumes "M(G)" "M(Q)" shows "lam_replacement(M, \x . {a \ G . \s\x. \s, a\ \ Q})" proof - 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: 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 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,363 +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^*\\ 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 \ \We obtain this lemmas as a consequence of the previous one; -alternatively it can be obtained using @{thm Ordinal.Memrel_type}\ +alternatively it can be obtained using @{thm [source] Ordinal.Memrel_type}\ lemma field_Memrel : "field(Memrel(A)) \ A" 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}\ +\ \Next theorem is very similar to @{thm [source] 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/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,1077 +1,1077 @@ section\Library of basic $\mathit{ZF}$ results\label{sec:zf-lib}\ theory ZF_Library_Relative imports 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) \ \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 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\ 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 \<^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})" proof - 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}\ +\ \Same as @{thm [source] 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\\ 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\ 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}\ + \ \Copy-paste from @{thm [source] 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 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 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_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