# HG changeset patch # User paulson # Date 1573561985 0 # Node ID ec7cc76e88e51773b81ee19e6dca85f0265c3b87 # Parent 038727567817a16541e03c343372e2f398be03c7 New library material from the AFP entry ZFC_in_HOL diff -r 038727567817 -r ec7cc76e88e5 src/HOL/Conditionally_Complete_Lattices.thy --- a/src/HOL/Conditionally_Complete_Lattices.thy Mon Nov 11 07:16:17 2019 +0000 +++ b/src/HOL/Conditionally_Complete_Lattices.thy Tue Nov 12 12:33:05 2019 +0000 @@ -337,7 +337,8 @@ lemma cINF_superset_mono: "A \ {} \ bdd_below (g ` B) \ A \ B \ (\x. x \ B \ g x \ f x) \ \(g ` B) \ \(f ` A)" by (rule cINF_mono) auto -lemma cSUP_subset_mono: "A \ {} \ bdd_above (g ` B) \ A \ B \ (\x. x \ B \ f x \ g x) \ \(f ` A) \ \(g ` B)" +lemma cSUP_subset_mono: + "\A \ {}; bdd_above (g ` B); A \ B; \x. x \ A \ f x \ g x\ \ \ (f ` A) \ \ (g ` B)" by (rule cSUP_mono) auto lemma less_eq_cInf_inter: "bdd_below A \ bdd_below B \ A \ B \ {} \ inf (Inf A) (Inf B) \ Inf (A \ B)" diff -r 038727567817 -r ec7cc76e88e5 src/HOL/Library/Equipollence.thy --- a/src/HOL/Library/Equipollence.thy Mon Nov 11 07:16:17 2019 +0000 +++ b/src/HOL/Library/Equipollence.thy Tue Nov 12 12:33:05 2019 +0000 @@ -1,7 +1,7 @@ section \Equipollence and Other Relations Connected with Cardinality\ theory "Equipollence" - imports FuncSet + imports FuncSet begin subsection\Eqpoll\ @@ -21,6 +21,9 @@ lemma eqpoll_iff_card_of_ordIso: "A \ B \ ordIso2 (card_of A) (card_of B)" by (simp add: card_of_ordIso eqpoll_def) +lemma eqpoll_refl [iff]: "A \ A" + by (simp add: card_of_refl eqpoll_iff_card_of_ordIso) + lemma eqpoll_finite_iff: "A \ B \ finite A \ finite B" by (meson bij_betw_finite eqpoll_def) @@ -63,6 +66,9 @@ lemma subset_imp_lepoll: "A \ B \ A \ B" by (force simp: lepoll_def) +lemma lepoll_refl [iff]: "A \ A" + by (simp add: subset_imp_lepoll) + lemma lepoll_iff: "A \ B \ (\g. A \ g ` B)" unfolding lepoll_def proof safe @@ -71,6 +77,9 @@ by (rule_tac x="inv_into B g" in exI) (auto simp: inv_into_into inj_on_inv_into) qed (metis image_mono the_inv_into_onto) +lemma empty_lepoll [iff]: "{} \ A" + by (simp add: lepoll_iff) + lemma subset_image_lepoll: "B \ f ` A \ B \ A" by (auto simp: lepoll_iff) @@ -82,6 +91,13 @@ apply (simp add: infinite_countable_subset) using infinite_iff_countable_subset by blast +lemma lepoll_Pow_self: "A \ Pow A" + unfolding lepoll_def inj_def + proof (intro exI conjI) + show "inj_on (\x. {x}) A" + by (auto simp: inj_on_def) +qed auto + lemma bij_betw_iff_bijections: "bij_betw f A B \ (\g. (\x \ A. f x \ B \ g(f x) = x) \ (\y \ B. g y \ A \ f(g y) = y))" (is "?lhs = ?rhs") @@ -127,6 +143,35 @@ by (metis (no_types, hide_lams) image_empty image_insert lepoll_iff that) qed +lemma infinite_insert_lepoll: + assumes "infinite A" shows "insert a A \ A" +proof - + obtain f :: "nat \ 'a" where "inj f" and f: "range f \ A" + using assms infinite_countable_subset by blast + let ?g = "(\z. if z=a then f 0 else if z \ range f then f (Suc (inv f z)) else z)" + show ?thesis + unfolding lepoll_def + proof (intro exI conjI) + show "inj_on ?g (insert a A)" + using inj_on_eq_iff [OF \inj f\] + by (auto simp: inj_on_def) + show "?g ` insert a A \ A" + using f by auto + qed +qed + +lemma infinite_insert_eqpoll: "infinite A \ insert a A \ A" + by (simp add: lepoll_antisym infinite_insert_lepoll subset_imp_lepoll subset_insertI) + +lemma finite_lepoll_infinite: + assumes "infinite A" "finite B" shows "B \ A" +proof - + have "B \ (UNIV::nat set)" + unfolding lepoll_def + using finite_imp_inj_to_nat_seg [OF \finite B\] by blast + then show ?thesis + using \infinite A\ infinite_le_lepoll lepoll_trans by auto +qed subsection\The strict relation\ @@ -154,7 +199,300 @@ lemma lesspoll_eq_trans [trans]: "\X \ Y; Y \ Z\ \ X \ Z" using eqpoll_imp_lepoll lesspoll_trans2 by blast -subsection\Cartesian products\ +lemma lesspoll_Pow_self: "A \ Pow A" + unfolding lesspoll_def bij_betw_def eqpoll_def + by (meson lepoll_Pow_self Cantors_paradox) + +lemma finite_lesspoll_infinite: + assumes "infinite A" "finite B" shows "B \ A" + by (meson assms eqpoll_finite_iff finite_lepoll_infinite lesspoll_def) + +subsection\Mapping by an injection\ + +lemma inj_on_image_eqpoll_self: "inj_on f A \ f ` A \ A" + by (meson bij_betw_def eqpoll_def eqpoll_sym) + +lemma inj_on_image_lepoll_1 [simp]: + assumes "inj_on f A" shows "f ` A \ B \ A \ B" + by (meson assms image_lepoll lepoll_def lepoll_trans order_refl) + +lemma inj_on_image_lepoll_2 [simp]: + assumes "inj_on f B" shows "A \ f ` B \ A \ B" + by (meson assms eq_iff image_lepoll lepoll_def lepoll_trans) + +lemma inj_on_image_lesspoll_1 [simp]: + assumes "inj_on f A" shows "f ` A \ B \ A \ B" + by (meson assms image_lepoll le_less lepoll_def lesspoll_trans1) + +lemma inj_on_image_lesspoll_2 [simp]: + assumes "inj_on f B" shows "A \ f ` B \ A \ B" + by (meson assms eqpoll_sym inj_on_image_eqpoll_self lesspoll_eq_trans) + +lemma inj_on_image_eqpoll_1 [simp]: + assumes "inj_on f A" shows "f ` A \ B \ A \ B" + by (metis assms eqpoll_trans inj_on_image_eqpoll_self eqpoll_sym) + +lemma inj_on_image_eqpoll_2 [simp]: + assumes "inj_on f B" shows "A \ f ` B \ A \ B" + by (metis assms inj_on_image_eqpoll_1 eqpoll_sym) + +subsection \Inserting elements into sets\ + +lemma insert_lepoll_insertD: + assumes "insert u A \ insert v B" "u \ A" "v \ B" shows "A \ B" +proof - + obtain f where inj: "inj_on f (insert u A)" and fim: "f ` (insert u A) \ insert v B" + by (meson assms lepoll_def) + show ?thesis + unfolding lepoll_def + proof (intro exI conjI) + let ?g = "\x\A. if f x = v then f u else f x" + show "inj_on ?g A" + using inj \u \ A\ by (auto simp: inj_on_def) + show "?g ` A \ B" + using fim \u \ A\ image_subset_iff inj inj_on_image_mem_iff by fastforce + qed +qed + +lemma insert_eqpoll_insertD: "\insert u A \ insert v B; u \ A; v \ B\ \ A \ B" + by (meson insert_lepoll_insertD eqpoll_imp_lepoll eqpoll_sym lepoll_antisym) + +lemma insert_lepoll_cong: + assumes "A \ B" "b \ B" shows "insert a A \ insert b B" +proof - + obtain f where f: "inj_on f A" "f ` A \ B" + by (meson assms lepoll_def) + let ?f = "\u \ insert a A. if u=a then b else f u" + show ?thesis + unfolding lepoll_def + proof (intro exI conjI) + show "inj_on ?f (insert a A)" + using f \b \ B\ by (auto simp: inj_on_def) + show "?f ` insert a A \ insert b B" + using f \b \ B\ by auto + qed +qed + +lemma insert_eqpoll_cong: + "\A \ B; a \ A; b \ B\ \ insert a A \ insert b B" + apply (rule lepoll_antisym) + apply (simp add: eqpoll_imp_lepoll insert_lepoll_cong)+ + by (meson eqpoll_imp_lepoll eqpoll_sym insert_lepoll_cong) + +lemma insert_eqpoll_insert_iff: + "\a \ A; b \ B\ \ insert a A \ insert b B \ A \ B" + by (meson insert_eqpoll_insertD insert_eqpoll_cong) + +lemma insert_lepoll_insert_iff: + " \a \ A; b \ B\ \ (insert a A \ insert b B) \ (A \ B)" + by (meson insert_lepoll_insertD insert_lepoll_cong) + +lemma less_imp_insert_lepoll: + assumes "A \ B" shows "insert a A \ B" +proof - + obtain f where "inj_on f A" "f ` A \ B" + using assms by (metis bij_betw_def eqpoll_def lepoll_def lesspoll_def psubset_eq) + then obtain b where b: "b \ B" "b \ f ` A" + by auto + show ?thesis + unfolding lepoll_def + proof (intro exI conjI) + show "inj_on (f(a:=b)) (insert a A)" + using b \inj_on f A\ by (auto simp: inj_on_def) + show "(f(a:=b)) ` insert a A \ B" + using \f ` A \ B\ by (auto simp: b) + qed +qed + +lemma finite_insert_lepoll: "finite A \ (insert a A \ A) \ (a \ A)" +proof (induction A rule: finite_induct) + case (insert x A) + then show ?case + apply (auto simp: insert_absorb) + by (metis insert_commute insert_iff insert_lepoll_insertD) +qed auto + + +subsection\Binary sums and unions\ + +lemma Un_lepoll_mono: + assumes "A \ C" "B \ D" "disjnt C D" shows "A \ B \ C \ D" +proof - + obtain f g where inj: "inj_on f A" "inj_on g B" and fg: "f ` A \ C" "g ` B \ D" + by (meson assms lepoll_def) + have "inj_on (\x. if x \ A then f x else g x) (A \ B)" + using inj \disjnt C D\ fg unfolding disjnt_iff + by (fastforce intro: inj_onI dest: inj_on_contraD split: if_split_asm) + with fg show ?thesis + unfolding lepoll_def + by (rule_tac x="\x. if x \ A then f x else g x" in exI) auto +qed + +lemma Un_eqpoll_cong: "\A \ C; B \ D; disjnt A B; disjnt C D\ \ A \ B \ C \ D" + by (meson Un_lepoll_mono eqpoll_imp_lepoll eqpoll_sym lepoll_antisym) + +lemma sum_lepoll_mono: + assumes "A \ C" "B \ D" shows "A <+> B \ C <+> D" +proof - + obtain f g where "inj_on f A" "f ` A \ C" "inj_on g B" "g ` B \ D" + by (meson assms lepoll_def) + then show ?thesis + unfolding lepoll_def + by (rule_tac x="case_sum (Inl \ f) (Inr \ g)" in exI) (force simp: inj_on_def) +qed + +lemma sum_eqpoll_cong: "\A \ C; B \ D\ \ A <+> B \ C <+> D" + by (meson eqpoll_imp_lepoll eqpoll_sym lepoll_antisym sum_lepoll_mono) + +subsection\Binary Cartesian products\ + +lemma times_square_lepoll: "A \ A \ A" + unfolding lepoll_def inj_def +proof (intro exI conjI) + show "inj_on (\x. (x,x)) A" + by (auto simp: inj_on_def) +qed auto + +lemma times_commute_eqpoll: "A \ B \ B \ A" + unfolding eqpoll_def + by (force intro: bij_betw_byWitness [where f = "\(x,y). (y,x)" and f' = "\(x,y). (y,x)"]) + +lemma times_assoc_eqpoll: "(A \ B) \ C \ A \ (B \ C)" + unfolding eqpoll_def + by (force intro: bij_betw_byWitness [where f = "\((x,y),z). (x,(y,z))" and f' = "\(x,(y,z)). ((x,y),z)"]) + +lemma times_singleton_eqpoll: "{a} \ A \ A" +proof - + have "{a} \ A = (\x. (a,x)) ` A" + by auto + also have "\ \ A" + proof (rule inj_on_image_eqpoll_self) + show "inj_on (Pair a) A" + by (auto simp: inj_on_def) + qed + finally show ?thesis . +qed + +lemma times_lepoll_mono: + assumes "A \ C" "B \ D" shows "A \ B \ C \ D" +proof - + obtain f g where "inj_on f A" "f ` A \ C" "inj_on g B" "g ` B \ D" + by (meson assms lepoll_def) + then show ?thesis + unfolding lepoll_def + by (rule_tac x="\(x,y). (f x, g y)" in exI) (auto simp: inj_on_def) +qed + +lemma times_eqpoll_cong: "\A \ C; B \ D\ \ A \ B \ C \ D" + by (metis eqpoll_imp_lepoll eqpoll_sym lepoll_antisym times_lepoll_mono) + +lemma + assumes "B \ {}" shows lepoll_times1: "A \ A \ B" and lepoll_times2: "A \ B \ A" + using assms lepoll_iff by fastforce+ + +lemma times_0_eqpoll: "{} \ A \ {}" + by (simp add: eqpoll_iff_bijections) + +lemma Sigma_lepoll_mono: + assumes "A \ C" "\x. x \ A \ B x \ D x" shows "Sigma A B \ Sigma C D" +proof - + have "\x. x \ A \ \f. inj_on f (B x) \ f ` (B x) \ D x" + by (meson assms lepoll_def) + then obtain f where "\x. x \ A \ inj_on (f x) (B x) \ f x ` B x \ D x" + by metis + with \A \ C\ show ?thesis + unfolding lepoll_def inj_on_def + by (rule_tac x="\(x,y). (x, f x y)" in exI) force +qed + +lemma sum_times_distrib_eqpoll: "(A <+> B) \ C \ (A \ C) <+> (B \ C)" + unfolding eqpoll_def +proof + show "bij_betw (\(x,z). case_sum(\y. Inl(y,z)) (\y. Inr(y,z)) x) ((A <+> B) \ C) (A \ C <+> B \ C)" + by (rule bij_betw_byWitness [where f' = "case_sum (\(x,z). (Inl x, z)) (\(y,z). (Inr y, z))"]) auto +qed + +lemma prod_insert_eqpoll: + assumes "a \ A" shows "insert a A \ B \ B <+> A \ B" + unfolding eqpoll_def + proof + show "bij_betw (\(x,y). if x=a then Inl y else Inr (x,y)) (insert a A \ B) (B <+> A \ B)" + by (rule bij_betw_byWitness [where f' = "case_sum (\y. (a,y)) id"]) (auto simp: assms) +qed + +subsection\General Unions\ + +lemma Union_eqpoll_Times: + assumes B: "\x. x \ A \ F x \ B" and disj: "pairwise (\x y. disjnt (F x) (F y)) A" + shows "(\x\A. F x) \ A \ B" +proof (rule lepoll_antisym) + obtain b where b: "\x. x \ A \ bij_betw (b x) (F x) B" + using B unfolding eqpoll_def by metis + show "\(F ` A) \ A \ B" + unfolding lepoll_def + proof (intro exI conjI) + define \ where "\ \ \z. THE x. x \ A \ z \ F x" + have \: "\ z = x" if "x \ A" "z \ F x" for x z + unfolding \_def + apply (rule the_equality) + apply (simp add: that) + by (metis disj disjnt_iff pairwiseD that) + let ?f = "\z. (\ z, b (\ z) z)" + show "inj_on ?f (\(F ` A))" + unfolding inj_on_def + by clarify (metis \ b bij_betw_inv_into_left) + show "?f ` \(F ` A) \ A \ B" + using \ b bij_betwE by blast + qed + show "A \ B \ \(F ` A)" + unfolding lepoll_def + proof (intro exI conjI) + let ?f = "\(x,y). inv_into (F x) (b x) y" + have *: "inv_into (F x) (b x) y \ F x" if "x \ A" "y \ B" for x y + by (metis b bij_betw_imp_surj_on inv_into_into that) + then show "inj_on ?f (A \ B)" + unfolding inj_on_def + by clarsimp (metis (mono_tags, lifting) b bij_betw_inv_into_right disj disjnt_iff pairwiseD) + show "?f ` (A \ B) \ \ (F ` A)" + by clarsimp (metis b bij_betw_imp_surj_on inv_into_into) + qed +qed + +lemma UN_lepoll_UN: + assumes A: "\x. x \ A \ B x \ C x" + and disj: "pairwise (\x y. disjnt (C x) (C y)) A" + shows "\ (B`A) \ \ (C`A)" +proof - + obtain f where f: "\x. x \ A \ inj_on (f x) (B x) \ f x ` (B x) \ (C x)" + using A unfolding lepoll_def by metis + show ?thesis + unfolding lepoll_def + proof (intro exI conjI) + define \ where "\ \ \z. @x. x \ A \ z \ B x" + have \: "\ z \ A \ z \ B (\ z)" if "x \ A" "z \ B x" for x z + unfolding \_def by (metis (mono_tags, lifting) someI_ex that) + let ?f = "\z. (f (\ z) z)" + show "inj_on ?f (\(B ` A))" + using disj f unfolding inj_on_def disjnt_iff pairwise_def image_subset_iff + by (metis UN_iff \) + show "?f ` \ (B ` A) \ \ (C ` A)" + using \ f unfolding image_subset_iff by blast + qed +qed + +lemma UN_eqpoll_UN: + assumes A: "\x. x \ A \ B x \ C x" + and B: "pairwise (\x y. disjnt (B x) (B y)) A" + and C: "pairwise (\x y. disjnt (C x) (C y)) A" + shows "(\x\A. B x) \ (\x\A. C x)" +proof (rule lepoll_antisym) + show "\ (B ` A) \ \ (C ` A)" + by (meson A C UN_lepoll_UN eqpoll_imp_lepoll) + show "\ (C ` A) \ \ (B ` A)" + by (simp add: A B UN_lepoll_UN eqpoll_imp_lepoll eqpoll_sym) +qed + +subsection\General Cartesian products (Pi)\ lemma PiE_sing_eqpoll_self: "({a} \\<^sub>E B) \ B" proof -