# HG changeset patch # User blanchet # Date 1384794285 -3600 # Node ID b4d644be252c4278b7304d20e4ee30f90c48a541 # Parent 6d5941722fae1afe401a4d528a6715c735c33ee6 moved theorems out of LFP diff -r 6d5941722fae -r b4d644be252c src/HOL/Cardinals/Cardinal_Arithmetic.thy --- a/src/HOL/Cardinals/Cardinal_Arithmetic.thy Mon Nov 18 18:04:44 2013 +0100 +++ b/src/HOL/Cardinals/Cardinal_Arithmetic.thy Mon Nov 18 18:04:45 2013 +0100 @@ -8,7 +8,7 @@ header {* Cardinal Arithmetic *} theory Cardinal_Arithmetic -imports Cardinal_Arithmetic_GFP +imports Cardinal_Arithmetic_GFP Cardinal_Order_Relation begin @@ -99,7 +99,7 @@ lemma Cinfinite_ctwo_cexp: "Cinfinite r \ Cinfinite (ctwo ^c r)" unfolding ctwo_def cexp_def cinfinite_def Field_card_of -by (rule conjI, rule infinite_Func, auto, rule card_of_card_order_on) +by (rule conjI, rule infinite_Func, auto) lemma cone_ordLeq_iff_Field: assumes "cone \o r" @@ -202,6 +202,6 @@ unfolding clists_def by (rule card_of_Card_order) lemma Cnotzero_clists: "Cnotzero (clists r)" -by (simp add: clists_def card_of_ordIso_czero_iff_empty lists_not_empty) (rule card_of_Card_order) +by (simp add: clists_def card_of_ordIso_czero_iff_empty lists_not_empty) end diff -r 6d5941722fae -r b4d644be252c src/HOL/Cardinals/Cardinal_Order_Relation.thy --- a/src/HOL/Cardinals/Cardinal_Order_Relation.thy Mon Nov 18 18:04:44 2013 +0100 +++ b/src/HOL/Cardinals/Cardinal_Order_Relation.thy Mon Nov 18 18:04:45 2013 +0100 @@ -34,7 +34,6 @@ Card_order_singl_ordLeq[simp] card_of_Pow[simp] Card_order_Pow[simp] - card_of_set_type[simp] card_of_Plus1[simp] Card_order_Plus1[simp] card_of_Plus2[simp] @@ -44,25 +43,19 @@ card_of_Plus_mono[simp] card_of_Plus_cong2[simp] card_of_Plus_cong[simp] - card_of_Un1[simp] - card_of_diff[simp] card_of_Un_Plus_ordLeq[simp] card_of_Times1[simp] card_of_Times2[simp] card_of_Times3[simp] card_of_Times_mono1[simp] card_of_Times_mono2[simp] - card_of_Times_cong1[simp] - card_of_Times_cong2[simp] card_of_ordIso_finite[simp] - finite_ordLess_infinite2[simp] card_of_Times_same_infinite[simp] card_of_Times_infinite_simps[simp] card_of_Plus_infinite1[simp] card_of_Plus_infinite2[simp] card_of_Plus_ordLess_infinite[simp] card_of_Plus_ordLess_infinite_Field[simp] - card_of_lists_infinite[simp] infinite_cartesian_product[simp] cardSuc_Card_order[simp] cardSuc_greater[simp] @@ -143,6 +136,25 @@ subsection {* Cardinals versus set operations on arbitrary sets *} +lemma infinite_Pow: +assumes "infinite A" +shows "infinite (Pow A)" +proof- + have "|A| \o |Pow A|" by (metis card_of_Pow ordLess_imp_ordLeq) + thus ?thesis by (metis assms finite_Pow_iff) +qed + +corollary card_of_set_type[simp]: "|UNIV::'a set| o |A \ B| " +using inj_on_id[of A] card_of_ordLeq[of A _] by fastforce + +lemma card_of_diff[simp]: +shows "|A - B| \o |A|" +using inj_on_id[of "A - B"] card_of_ordLeq[of "A - B" _] by fastforce + lemma subset_ordLeq_strict: assumes "A \ B" and "|A| C| =o |B \ C|" +using assms by (simp add: ordIso_iff_ordLeq card_of_Times_mono1) + +lemma card_of_Times_cong2[simp]: +assumes "|A| =o |B|" +shows "|C \ A| =o |C \ B|" +using assms by (simp add: ordIso_iff_ordLeq card_of_Times_mono2) + lemma card_of_Times_mono[simp]: assumes "|A| \o |B|" and "|C| \o |D|" shows "|A \ C| \o |B \ D|" @@ -323,6 +345,11 @@ shows "|(Field r) \ C| =o |(Field r') \ C|" using assms card_of_cong card_of_Times_cong1 by blast +corollary ordIso_Times_cong2: +assumes "r =o r'" +shows "|A \ (Field r)| =o |A \ (Field r')|" +using assms card_of_cong card_of_Times_cong2 by blast + lemma card_of_Times_cong[simp]: assumes "|A| =o |B|" and "|C| =o |D|" shows "|A \ C| =o |B \ D|" @@ -501,11 +528,55 @@ using assms Plus_infinite_bij_betw[of "UNIV::'a set" g "UNIV::'b set"] by auto +lemma card_of_Un_infinite: +assumes INF: "infinite A" and LEQ: "|B| \o |A|" +shows "|A \ B| =o |A| \ |B \ A| =o |A|" +proof- + have "|A \ B| \o |A <+> B|" by (rule card_of_Un_Plus_ordLeq) + moreover have "|A <+> B| =o |A|" + using assms by (metis card_of_Plus_infinite) + ultimately have "|A \ B| \o |A|" using ordLeq_ordIso_trans by blast + hence "|A \ B| =o |A|" using card_of_Un1 ordIso_iff_ordLeq by blast + thus ?thesis using Un_commute[of B A] by auto +qed + lemma card_of_Un_infinite_simps[simp]: "\infinite A; |B| \o |A| \ \ |A \ B| =o |A|" "\infinite A; |B| \o |A| \ \ |B \ A| =o |A|" using card_of_Un_infinite by auto +lemma card_of_Un_diff_infinite: +assumes INF: "infinite A" and LESS: "|B| B| =o |A|" + using assms ordLeq_iff_ordLess_or_ordIso card_of_Un_infinite by blast + moreover have "C \ B = A \ B" unfolding C_def by auto + ultimately have 1: "|C \ B| =o |A|" by auto + (* *) + {assume *: "|C| \o |B|" + moreover + {assume **: "finite B" + hence "finite C" + using card_of_ordLeq_finite * by blast + hence False using ** INF card_of_ordIso_finite 1 by blast + } + hence "infinite B" by auto + ultimately have False + using card_of_Un_infinite 1 ordIso_equivalence(1,3) LESS not_ordLess_ordIso by metis + } + hence 2: "|B| \o |C|" using card_of_Well_order ordLeq_total by blast + {assume *: "finite C" + hence "finite B" using card_of_ordLeq_finite 2 by blast + hence False using * INF card_of_ordIso_finite 1 by blast + } + hence "infinite C" by auto + hence "|C| =o |A|" + using card_of_Un_infinite 1 2 ordIso_equivalence(1,3) by metis + thus ?thesis unfolding C_def . +qed + corollary Card_order_Un_infinite: assumes INF: "infinite(Field r)" and CARD: "Card_order r" and LEQ: "p \o r" @@ -597,6 +668,33 @@ thus ?thesis using 1 ordLess_ordIso_trans by blast qed + +subsection {* Cardinals versus set operations involving infinite sets *} + +lemma finite_iff_cardOf_nat: +"finite A = ( |A| {a}" by simp + show ?thesis + using infinite_imp_bij_betw2[OF assms] unfolding iA + by (metis bij_betw_inv card_of_ordIso) +qed + lemma card_of_Un_singl_ordLess_infinite1: assumes "infinite B" and "|A| nat \ 'a list set" +where "nlists A n \ {l. set l \ A \ length l = n}" + +lemma lists_def2: "lists A = {l. set l \ A}" +using in_listsI by blast + +lemma lists_UNION_nlists: "lists A = (\ n. nlists A n)" +unfolding lists_def2 nlists_def by blast + +lemma card_of_lists: "|A| \o |lists A|" +proof- + let ?h = "\ a. [a]" + have "inj_on ?h A \ ?h ` A \ lists A" + unfolding inj_on_def lists_def2 by auto + thus ?thesis by (metis card_of_ordLeq) +qed + +lemma nlists_0: "nlists A 0 = {[]}" +unfolding nlists_def by auto + +lemma nlists_not_empty: +assumes "A \ {}" +shows "nlists A n \ {}" +proof(induct n, simp add: nlists_0) + fix n assume "nlists A n \ {}" + then obtain a and l where "a \ A \ l \ nlists A n" using assms by auto + hence "a # l \ nlists A (Suc n)" unfolding nlists_def by auto + thus "nlists A (Suc n) \ {}" by auto +qed + +lemma Nil_in_lists: "[] \ lists A" +unfolding lists_def2 by auto + +lemma lists_not_empty: "lists A \ {}" +using Nil_in_lists by blast + +lemma card_of_nlists_Succ: "|nlists A (Suc n)| =o |A \ (nlists A n)|" +proof- + let ?B = "A \ (nlists A n)" let ?h = "\(a,l). a # l" + have "inj_on ?h ?B \ ?h ` ?B \ nlists A (Suc n)" + unfolding inj_on_def nlists_def by auto + moreover have "nlists A (Suc n) \ ?h ` ?B" + proof(auto) + fix l assume "l \ nlists A (Suc n)" + hence 1: "length l = Suc n \ set l \ A" unfolding nlists_def by auto + then obtain a and l' where 2: "l = a # l'" by (auto simp: length_Suc_conv) + hence "a \ A \ set l' \ A \ length l' = n" using 1 by auto + thus "l \ ?h ` ?B" using 2 unfolding nlists_def by auto + qed + ultimately have "bij_betw ?h ?B (nlists A (Suc n))" + unfolding bij_betw_def by auto + thus ?thesis using card_of_ordIso ordIso_symmetric by blast +qed + +lemma card_of_nlists_infinite: +assumes "infinite A" +shows "|nlists A n| \o |A|" +proof(induct n) + have "A \ {}" using assms by auto + thus "|nlists A 0| \o |A|" by (simp add: nlists_0 card_of_singl_ordLeq) +next + fix n assume IH: "|nlists A n| \o |A|" + have "|nlists A (Suc n)| =o |A \ (nlists A n)|" + using card_of_nlists_Succ by blast + moreover + {have "nlists A n \ {}" using assms nlists_not_empty[of A] by blast + hence "|A \ (nlists A n)| =o |A|" + using assms IH by (auto simp add: card_of_Times_infinite) + } + ultimately show "|nlists A (Suc n)| \o |A|" + using ordIso_transitive ordIso_iff_ordLeq by blast +qed lemma Card_order_lists: "Card_order r \ r \o |lists(Field r) |" using card_of_lists card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast @@ -690,6 +864,22 @@ thus ?thesis using card_of_ordIso[of "lists A"] by auto qed +lemma card_of_lists_infinite[simp]: +assumes "infinite A" +shows "|lists A| =o |A|" +proof- + have "|lists A| \o |A|" + using assms + by (auto simp add: lists_UNION_nlists card_of_UNION_ordLeq_infinite + infinite_iff_card_of_nat card_of_nlists_infinite) + thus ?thesis using card_of_lists ordIso_iff_ordLeq by blast +qed + +lemma Card_order_lists_infinite: +assumes "Card_order r" and "infinite(Field r)" +shows "|lists(Field r)| =o r" +using assms card_of_lists_infinite card_of_Field_ordIso ordIso_transitive by blast + lemma ordIso_lists_cong: assumes "r =o r'" shows "|lists(Field r)| =o |lists(Field r')|" @@ -827,13 +1017,22 @@ lemma Field_natLess: "Field natLess = (UNIV::nat set)" by(unfold Field_def, auto) +lemma natLeq_well_order_on: "well_order_on UNIV natLeq" +using natLeq_Well_order Field_natLeq by auto + +lemma natLeq_wo_rel: "wo_rel natLeq" +unfolding wo_rel_def using natLeq_Well_order . + lemma natLeq_ofilter_less: "ofilter natLeq {0 ..< n}" by(auto simp add: natLeq_wo_rel wo_rel.ofilter_def, - simp add: Field_natLeq, unfold rel.under_def, auto) + simp add: Field_natLeq, unfold rel.under_def, auto) lemma natLeq_ofilter_leq: "ofilter natLeq {0 .. n}" by(auto simp add: natLeq_wo_rel wo_rel.ofilter_def, - simp add: Field_natLeq, unfold rel.under_def, auto) + simp add: Field_natLeq, unfold rel.under_def, auto) + +lemma natLeq_UNIV_ofilter: "wo_rel.ofilter natLeq UNIV" +using natLeq_wo_rel Field_natLeq wo_rel.Field_ofilter[of natLeq] by auto lemma natLeq_ofilter_iff: "ofilter natLeq A = (A = UNIV \ (\n. A = {0 ..< n}))" @@ -900,7 +1099,7 @@ qed -subsubsection {* "Backwards compatibility" with the numeric cardinal operator for finite sets *} +subsubsection {* "Backward compatibility" with the numeric cardinal operator for finite sets *} lemma finite_card_of_iff_card: assumes FIN: "finite A" and FIN': "finite B" @@ -993,6 +1192,11 @@ shows "relChain r (\ i. under r i)" using assms unfolding relChain_def by auto +lemma card_of_infinite_diff_finite: +assumes "infinite A" and "finite B" +shows "|A - B| =o |A|" +by (metis assms card_of_Un_diff_infinite finite_ordLess_infinite2) + lemma infinite_card_of_diff_singl: assumes "infinite A" shows "|A - {a}| =o |A|" @@ -1110,6 +1314,30 @@ thus "f \ Pfunc A B" unfolding Func_option_def Pfunc_def by auto qed +lemma card_of_Func_mono: +fixes A1 A2 :: "'a set" and B :: "'b set" +assumes A12: "A1 \ A2" and B: "B \ {}" +shows "|Func A1 B| \o |Func A2 B|" +proof- + obtain bb where bb: "bb \ B" using B by auto + def F \ "\ (f1::'a \ 'b) a. if a \ A2 then (if a \ A1 then f1 a else bb) + else undefined" + show ?thesis unfolding card_of_ordLeq[symmetric] proof(intro exI[of _ F] conjI) + show "inj_on F (Func A1 B)" unfolding inj_on_def proof safe + fix f g assume f: "f \ Func A1 B" and g: "g \ Func A1 B" and eq: "F f = F g" + show "f = g" + proof(rule ext) + fix a show "f a = g a" + proof(cases "a \ A1") + case True + thus ?thesis using eq A12 unfolding F_def fun_eq_iff + by (elim allE[of _ a]) auto + qed(insert f g, unfold Func_def, fastforce) + qed + qed + qed(insert bb, unfold Func_def F_def, force) +qed + lemma card_of_Func_option_mono: fixes A1 A2 :: "'a set" and B :: "'b set" assumes A12: "A1 \ A2" and B: "B \ {}" @@ -1178,4 +1406,18 @@ "|Func (UNIV::'a set) (UNIV::'b set)| =o |UNIV::('a \ 'b) set|" using card_of_Func_UNIV[of "UNIV::'b set"] by auto +lemma ordLeq_Func: +assumes "{b1,b2} \ B" "b1 \ b2" +shows "|A| \o |Func A B|" +unfolding card_of_ordLeq[symmetric] proof(intro exI conjI) + let ?F = "\ aa a. if a \ A then (if a = aa then b1 else b2) else undefined" + show "inj_on ?F A" using assms unfolding inj_on_def fun_eq_iff by auto + show "?F ` A \ Func A B" using assms unfolding Func_def by auto +qed + +lemma infinite_Func: +assumes A: "infinite A" and B: "{b1,b2} \ B" "b1 \ b2" +shows "infinite (Func A B)" +using ordLeq_Func[OF B] by (metis A card_of_ordLeq_finite) + end diff -r 6d5941722fae -r b4d644be252c src/HOL/Cardinals/Cardinal_Order_Relation_LFP.thy --- a/src/HOL/Cardinals/Cardinal_Order_Relation_LFP.thy Mon Nov 18 18:04:44 2013 +0100 +++ b/src/HOL/Cardinals/Cardinal_Order_Relation_LFP.thy Mon Nov 18 18:04:45 2013 +0100 @@ -156,7 +156,7 @@ lemma card_of_Well_order: "Well_order |A|" -using card_of_Card_order unfolding card_order_on_def by auto +using card_of_Card_order unfolding card_order_on_def by auto lemma card_of_refl: "|A| =o |A|" @@ -481,24 +481,11 @@ Pow_not_empty[of A] by auto -lemma infinite_Pow: -assumes "infinite A" -shows "infinite (Pow A)" -proof- - have "|A| \o |Pow A|" by (metis card_of_Pow ordLess_imp_ordLeq) - thus ?thesis by (metis assms finite_Pow_iff) -qed - - corollary Card_order_Pow: "Card_order r \ r o |A <+> B|" proof- have "Inl ` A \ A <+> B" by auto @@ -681,16 +668,6 @@ using assms card_of_cong[of r r'] card_of_cong[of p p'] card_of_Plus_cong by blast -lemma card_of_Un1: -shows "|A| \o |A \ B| " -using inj_on_id[of A] card_of_ordLeq[of A _] by fastforce - - -lemma card_of_diff: -shows "|A - B| \o |A|" -using inj_on_id[of "A - B"] card_of_ordLeq[of "A - B" _] by fastforce - - lemma card_of_Un_Plus_ordLeq: "|A \ B| \o |A <+> B|" proof- @@ -712,12 +689,6 @@ qed -corollary Card_order_Times1: -"\Card_order r; B \ {}\ \ r \o |(Field r) \ B|" -using card_of_Times1[of B] card_of_Field_ordIso - ordIso_ordLeq_trans ordIso_symmetric by blast - - lemma card_of_Times_commute: "|A \ B| =o |B \ A|" proof- let ?f = "\(a::'a,b::'b). (b,a)" @@ -733,6 +704,12 @@ ordLeq_ordIso_trans by blast +corollary Card_order_Times1: +"\Card_order r; B \ {}\ \ r \o |(Field r) \ B|" +using card_of_Times1[of B] card_of_Field_ordIso + ordIso_ordLeq_trans ordIso_symmetric by blast + + corollary Card_order_Times2: "\Card_order r; A \ {}\ \ r \o |A \ (Field r)|" using card_of_Times2[of A] card_of_Field_ordIso @@ -811,24 +788,6 @@ using assms card_of_mono2 card_of_Times_mono2 by blast -lemma card_of_Times_cong1: -assumes "|A| =o |B|" -shows "|A \ C| =o |B \ C|" -using assms by (simp add: ordIso_iff_ordLeq card_of_Times_mono1) - - -lemma card_of_Times_cong2: -assumes "|A| =o |B|" -shows "|C \ A| =o |C \ B|" -using assms by (simp add: ordIso_iff_ordLeq card_of_Times_mono2) - - -corollary ordIso_Times_cong2: -assumes "r =o r'" -shows "|A \ (Field r)| =o |A \ (Field r')|" -using assms card_of_cong card_of_Times_cong2 by blast - - lemma card_of_Sigma_mono1: assumes "\i \ I. |A i| \o |B i|" shows "|SIGMA i : I. A i| \o |SIGMA i : I. B i|" @@ -975,21 +934,6 @@ by (auto simp add: infinite_iff_countable_subset card_of_ordLeq) -lemma finite_iff_cardOf_nat: -"finite A = ( |A| infinite A; B \ {}; |B| \o |A|\ \ |A \ B| =o |A|" -"\infinite A; B \ {}; |B| \o |A|\ \ |A| =o |A \ B|" -"\infinite A; B \ {}; |B| \o |A|\ \ |B \ A| =o |A|" -"\infinite A; B \ {}; |B| \o |A|\ \ |A| =o |B \ A|" -by (auto simp add: card_of_Times_infinite ordIso_symmetric) - - corollary Card_order_Times_infinite: assumes INF: "infinite(Field r)" and CARD: "Card_order r" and NE: "Field p \ {}" and LEQ: "p \o r" @@ -1216,6 +1152,14 @@ by(simp add: card_of_Sigma_ordLeq_infinite_Field) +lemma card_of_Times_infinite_simps: +"\infinite A; B \ {}; |B| \o |A|\ \ |A \ B| =o |A|" +"\infinite A; B \ {}; |B| \o |A|\ \ |A| =o |A \ B|" +"\infinite A; B \ {}; |B| \o |A|\ \ |B \ A| =o |A|" +"\infinite A; B \ {}; |B| \o |A|\ \ |A| =o |B \ A|" +by (auto simp add: card_of_Times_infinite ordIso_symmetric) + + lemma card_of_UNION_ordLeq_infinite: assumes INF: "infinite B" and LEQ_I: "|I| \o |B|" and LEQ: "\i \ I. |A i| \o |B|" @@ -1313,226 +1257,6 @@ qed -lemma card_of_Un_infinite: -assumes INF: "infinite A" and LEQ: "|B| \o |A|" -shows "|A \ B| =o |A| \ |B \ A| =o |A|" -proof- - have "|A \ B| \o |A <+> B|" by (rule card_of_Un_Plus_ordLeq) - moreover have "|A <+> B| =o |A|" - using assms by (metis card_of_Plus_infinite) - ultimately have "|A \ B| \o |A|" using ordLeq_ordIso_trans by blast - hence "|A \ B| =o |A|" using card_of_Un1 ordIso_iff_ordLeq by blast - thus ?thesis using Un_commute[of B A] by auto -qed - - -lemma card_of_Un_diff_infinite: -assumes INF: "infinite A" and LESS: "|B| B| =o |A|" - using assms ordLeq_iff_ordLess_or_ordIso card_of_Un_infinite by blast - moreover have "C \ B = A \ B" unfolding C_def by auto - ultimately have 1: "|C \ B| =o |A|" by auto - (* *) - {assume *: "|C| \o |B|" - moreover - {assume **: "finite B" - hence "finite C" - using card_of_ordLeq_finite * by blast - hence False using ** INF card_of_ordIso_finite 1 by blast - } - hence "infinite B" by auto - ultimately have False - using card_of_Un_infinite 1 ordIso_equivalence(1,3) LESS not_ordLess_ordIso by metis - } - hence 2: "|B| \o |C|" using card_of_Well_order ordLeq_total by blast - {assume *: "finite C" - hence "finite B" using card_of_ordLeq_finite 2 by blast - hence False using * INF card_of_ordIso_finite 1 by blast - } - hence "infinite C" by auto - hence "|C| =o |A|" - using card_of_Un_infinite 1 2 ordIso_equivalence(1,3) by metis - thus ?thesis unfolding C_def . -qed - - -lemma card_of_Plus_ordLess_infinite: -assumes INF: "infinite C" and - LESS1: "|A| B| B = {}") - assume Case1: "A = {} \ B = {}" - hence "|A| =o |A <+> B| \ |B| =o |A <+> B|" - using card_of_Plus_empty1 card_of_Plus_empty2 by blast - hence "|A <+> B| =o |A| \ |A <+> B| =o |B|" - using ordIso_symmetric[of "|A|"] ordIso_symmetric[of "|B|"] by blast - thus ?thesis using LESS1 LESS2 - ordIso_ordLess_trans[of "|A <+> B|" "|A|"] - ordIso_ordLess_trans[of "|A <+> B|" "|B|"] by blast -next - assume Case2: "\(A = {} \ B = {})" - {assume *: "|C| \o |A <+> B|" - hence "infinite (A <+> B)" using INF card_of_ordLeq_finite by blast - hence 1: "infinite A \ infinite B" using finite_Plus by blast - {assume Case21: "|A| \o |B|" - hence "infinite B" using 1 card_of_ordLeq_finite by blast - hence "|A <+> B| =o |B|" using Case2 Case21 - by (auto simp add: card_of_Plus_infinite) - hence False using LESS2 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast - } - moreover - {assume Case22: "|B| \o |A|" - hence "infinite A" using 1 card_of_ordLeq_finite by blast - hence "|A <+> B| =o |A|" using Case2 Case22 - by (auto simp add: card_of_Plus_infinite) - hence False using LESS1 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast - } - ultimately have False using ordLeq_total card_of_Well_order[of A] - card_of_Well_order[of B] by blast - } - thus ?thesis using ordLess_or_ordLeq[of "|A <+> B|" "|C|"] - card_of_Well_order[of "A <+> B"] card_of_Well_order[of "C"] by auto -qed - - -lemma card_of_Plus_ordLess_infinite_Field: -assumes INF: "infinite (Field r)" and r: "Card_order r" and - LESS1: "|A| B| |?C| =o r" using r card_of_Field_ordIso - ordIso_symmetric by blast - hence "|A| B| {a}" by simp - show ?thesis - using infinite_imp_bij_betw2[OF assms] unfolding iA - by (metis bij_betw_inv card_of_ordIso) -qed - - -subsection {* Cardinals versus lists *} - - -text{* The next is an auxiliary operator, which shall be used for inductive -proofs of facts concerning the cardinality of @{text "List"} : *} - -definition nlists :: "'a set \ nat \ 'a list set" -where "nlists A n \ {l. set l \ A \ length l = n}" - - -lemma lists_def2: "lists A = {l. set l \ A}" -using in_listsI by blast - - -lemma lists_UNION_nlists: "lists A = (\ n. nlists A n)" -unfolding lists_def2 nlists_def by blast - - -lemma card_of_lists: "|A| \o |lists A|" -proof- - let ?h = "\ a. [a]" - have "inj_on ?h A \ ?h ` A \ lists A" - unfolding inj_on_def lists_def2 by auto - thus ?thesis by (metis card_of_ordLeq) -qed - - -lemma nlists_0: "nlists A 0 = {[]}" -unfolding nlists_def by auto - - -lemma nlists_not_empty: -assumes "A \ {}" -shows "nlists A n \ {}" -proof(induct n, simp add: nlists_0) - fix n assume "nlists A n \ {}" - then obtain a and l where "a \ A \ l \ nlists A n" using assms by auto - hence "a # l \ nlists A (Suc n)" unfolding nlists_def by auto - thus "nlists A (Suc n) \ {}" by auto -qed - - -lemma Nil_in_lists: "[] \ lists A" -unfolding lists_def2 by auto - - -lemma lists_not_empty: "lists A \ {}" -using Nil_in_lists by blast - - -lemma card_of_nlists_Succ: "|nlists A (Suc n)| =o |A \ (nlists A n)|" -proof- - let ?B = "A \ (nlists A n)" let ?h = "\(a,l). a # l" - have "inj_on ?h ?B \ ?h ` ?B \ nlists A (Suc n)" - unfolding inj_on_def nlists_def by auto - moreover have "nlists A (Suc n) \ ?h ` ?B" - proof(auto) - fix l assume "l \ nlists A (Suc n)" - hence 1: "length l = Suc n \ set l \ A" unfolding nlists_def by auto - then obtain a and l' where 2: "l = a # l'" by (auto simp: length_Suc_conv) - hence "a \ A \ set l' \ A \ length l' = n" using 1 by auto - thus "l \ ?h ` ?B" using 2 unfolding nlists_def by auto - qed - ultimately have "bij_betw ?h ?B (nlists A (Suc n))" - unfolding bij_betw_def by auto - thus ?thesis using card_of_ordIso ordIso_symmetric by blast -qed - - -lemma card_of_nlists_infinite: -assumes "infinite A" -shows "|nlists A n| \o |A|" -proof(induct n) - have "A \ {}" using assms by auto - thus "|nlists A 0| \o |A|" by (simp add: nlists_0 card_of_singl_ordLeq) -next - fix n assume IH: "|nlists A n| \o |A|" - have "|nlists A (Suc n)| =o |A \ (nlists A n)|" - using card_of_nlists_Succ by blast - moreover - {have "nlists A n \ {}" using assms nlists_not_empty[of A] by blast - hence "|A \ (nlists A n)| =o |A|" - using assms IH by (auto simp add: card_of_Times_infinite) - } - ultimately show "|nlists A (Suc n)| \o |A|" - using ordIso_transitive ordIso_iff_ordLeq by blast -qed - - -lemma card_of_lists_infinite: -assumes "infinite A" -shows "|lists A| =o |A|" -proof- - have "|lists A| \o |A|" - using assms - by (auto simp add: lists_UNION_nlists card_of_UNION_ordLeq_infinite - infinite_iff_card_of_nat card_of_nlists_infinite) - thus ?thesis using card_of_lists ordIso_iff_ordLeq by blast -qed - - -lemma Card_order_lists_infinite: -assumes "Card_order r" and "infinite(Field r)" -shows "|lists(Field r)| =o r" -using assms card_of_lists_infinite card_of_Field_ordIso ordIso_transitive by blast - - - subsection {* The cardinal $\omega$ and the finite cardinals *} @@ -1559,7 +1283,6 @@ qed - subsubsection {* First as well-orders *} @@ -1607,18 +1330,6 @@ using natLeq_Linear_order wf_less natLeq_natLess_Id by auto -corollary natLeq_well_order_on: "well_order_on UNIV natLeq" -using natLeq_Well_order Field_natLeq by auto - - -lemma natLeq_wo_rel: "wo_rel natLeq" -unfolding wo_rel_def using natLeq_Well_order . - - -lemma natLeq_UNIV_ofilter: "wo_rel.ofilter natLeq UNIV" -using natLeq_wo_rel Field_natLeq wo_rel.Field_ofilter[of natLeq] by auto - - lemma closed_nat_set_iff: assumes "\(m::nat) n. n \ A \ m \ n \ m \ A" shows "A = UNIV \ (\n. A = {0 ..< n})" @@ -1786,7 +1497,7 @@ -subsubsection {* "Backwards compatibility" with the numeric cardinal operator for finite sets *} +subsubsection {* "Backward compatibility" with the numeric cardinal operator for finite sets *} lemma finite_card_of_iff_card2: @@ -2048,6 +1759,61 @@ qed +lemma card_of_Plus_ordLess_infinite: +assumes INF: "infinite C" and + LESS1: "|A| B| B = {}") + assume Case1: "A = {} \ B = {}" + hence "|A| =o |A <+> B| \ |B| =o |A <+> B|" + using card_of_Plus_empty1 card_of_Plus_empty2 by blast + hence "|A <+> B| =o |A| \ |A <+> B| =o |B|" + using ordIso_symmetric[of "|A|"] ordIso_symmetric[of "|B|"] by blast + thus ?thesis using LESS1 LESS2 + ordIso_ordLess_trans[of "|A <+> B|" "|A|"] + ordIso_ordLess_trans[of "|A <+> B|" "|B|"] by blast +next + assume Case2: "\(A = {} \ B = {})" + {assume *: "|C| \o |A <+> B|" + hence "infinite (A <+> B)" using INF card_of_ordLeq_finite by blast + hence 1: "infinite A \ infinite B" using finite_Plus by blast + {assume Case21: "|A| \o |B|" + hence "infinite B" using 1 card_of_ordLeq_finite by blast + hence "|A <+> B| =o |B|" using Case2 Case21 + by (auto simp add: card_of_Plus_infinite) + hence False using LESS2 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast + } + moreover + {assume Case22: "|B| \o |A|" + hence "infinite A" using 1 card_of_ordLeq_finite by blast + hence "|A <+> B| =o |A|" using Case2 Case22 + by (auto simp add: card_of_Plus_infinite) + hence False using LESS1 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast + } + ultimately have False using ordLeq_total card_of_Well_order[of A] + card_of_Well_order[of B] by blast + } + thus ?thesis using ordLess_or_ordLeq[of "|A <+> B|" "|C|"] + card_of_Well_order[of "A <+> B"] card_of_Well_order[of "C"] by auto +qed + + +lemma card_of_Plus_ordLess_infinite_Field: +assumes INF: "infinite (Field r)" and r: "Card_order r" and + LESS1: "|A| B| |?C| =o r" using r card_of_Field_ordIso + ordIso_symmetric by blast + hence "|A| B| o r" and B: "|B| \o r" and c: "Card_order r" @@ -2199,11 +1965,6 @@ subsection {* Others *} -lemma card_of_infinite_diff_finite: -assumes "infinite A" and "finite B" -shows "|A - B| =o |A|" -by (metis assms card_of_Un_diff_infinite finite_ordLess_infinite2) - (* function space *) definition Func where "Func A B = {f . (\ a \ A. f a \ B) \ (\ a. a \ A \ f a = undefined)}" @@ -2383,44 +2144,6 @@ thus ?thesis unfolding card_of_ordIso[symmetric] by blast qed -lemma card_of_Func_mono: -fixes A1 A2 :: "'a set" and B :: "'b set" -assumes A12: "A1 \ A2" and B: "B \ {}" -shows "|Func A1 B| \o |Func A2 B|" -proof- - obtain bb where bb: "bb \ B" using B by auto - def F \ "\ (f1::'a \ 'b) a. if a \ A2 then (if a \ A1 then f1 a else bb) - else undefined" - show ?thesis unfolding card_of_ordLeq[symmetric] proof(intro exI[of _ F] conjI) - show "inj_on F (Func A1 B)" unfolding inj_on_def proof safe - fix f g assume f: "f \ Func A1 B" and g: "g \ Func A1 B" and eq: "F f = F g" - show "f = g" - proof(rule ext) - fix a show "f a = g a" - proof(cases "a \ A1") - case True - thus ?thesis using eq A12 unfolding F_def fun_eq_iff - by (elim allE[of _ a]) auto - qed(insert f g, unfold Func_def, fastforce) - qed - qed - qed(insert bb, unfold Func_def F_def, force) -qed - -lemma ordLeq_Func: -assumes "{b1,b2} \ B" "b1 \ b2" -shows "|A| \o |Func A B|" -unfolding card_of_ordLeq[symmetric] proof(intro exI conjI) - let ?F = "\ aa a. if a \ A then (if a = aa then b1 else b2) else undefined" - show "inj_on ?F A" using assms unfolding inj_on_def fun_eq_iff by auto - show "?F ` A \ Func A B" using assms unfolding Func_def by auto -qed - -lemma infinite_Func: -assumes A: "infinite A" and B: "{b1,b2} \ B" "b1 \ b2" -shows "infinite (Func A B)" -using ordLeq_Func[OF B] by (metis A card_of_ordLeq_finite) - lemma card_of_Func_UNIV: "|Func (UNIV::'a set) (B::'b set)| =o |{f::'a \ 'b. range f \ B}|" apply(rule ordIso_symmetric) proof(intro card_of_ordIsoI)