# HG changeset patch # User wenzelm # Date 1683492683 -7200 # Node ID e30a56be9c7bdfb154ad74760aa2a2617c5ed07f # Parent ffdad62bc2354c0df68169b5fb502a766068c4ba# Parent c1b8fdd6588ac232110e61559a4e7b09236fe309 merged diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Abstract_Limits.thy --- a/src/HOL/Analysis/Abstract_Limits.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Analysis/Abstract_Limits.thy Sun May 07 22:51:23 2023 +0200 @@ -41,12 +41,38 @@ finally show ?thesis by (simp add: True) qed auto +lemma nontrivial_limit_atin: + "atin X a \ bot \ a \ X derived_set_of topspace X" +proof + assume L: "atin X a \ bot" + then have "a \ topspace X" + by (meson atin_degenerate) + moreover have "\ openin X {a}" + using L by (auto simp: eventually_atin trivial_limit_eq) + ultimately + show "a \ X derived_set_of topspace X" + by (auto simp: derived_set_of_topspace) +next + assume a: "a \ X derived_set_of topspace X" + show "atin X a \ bot" + proof + assume "atin X a = bot" + then have "eventually (\_. False) (atin X a)" + by simp + then show False + by (smt (verit, best) a eventually_atin in_derived_set_of insertE insert_Diff) + qed +qed + subsection\Limits in a topological space\ definition limitin :: "'a topology \ ('b \ 'a) \ 'a \ 'b filter \ bool" where "limitin X f l F \ l \ topspace X \ (\U. openin X U \ l \ U \ eventually (\x. f x \ U) F)" +lemma limitinD: "\limitin X f l F; openin X U; l \ U\ \ eventually (\x. f x \ U) F" + by (simp add: limitin_def) + lemma limitin_canonical_iff [simp]: "limitin euclidean f l F \ (f \ l) F" by (auto simp: limitin_def tendsto_def) diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Abstract_Topological_Spaces.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Abstract_Topological_Spaces.thy Sun May 07 22:51:23 2023 +0200 @@ -0,0 +1,3531 @@ +(* Author: L C Paulson, University of Cambridge [ported from HOL Light] *) + +section \Various Forms of Topological Spaces\ + +theory Abstract_Topological_Spaces + imports Lindelof_Spaces Locally Sum_Topology FSigma +begin + + +subsection\Connected topological spaces\ + +lemma connected_space_eq_frontier_eq_empty: + "connected_space X \ (\S. S \ topspace X \ X frontier_of S = {} \ S = {} \ S = topspace X)" + by (meson clopenin_eq_frontier_of connected_space_clopen_in) + +lemma connected_space_frontier_eq_empty: + "connected_space X \ S \ topspace X + \ (X frontier_of S = {} \ S = {} \ S = topspace X)" + by (meson connected_space_eq_frontier_eq_empty frontier_of_empty frontier_of_topspace) + +lemma connectedin_eq_subset_separated_union: + "connectedin X C \ + C \ topspace X \ (\S T. separatedin X S T \ C \ S \ T \ C \ S \ C \ T)" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + using connectedin_subset_topspace connectedin_subset_separated_union by blast +next + assume ?rhs + then show ?lhs + by (metis closure_of_subset connectedin_separation dual_order.eq_iff inf.orderE separatedin_def sup.boundedE) +qed + + +lemma connectedin_clopen_cases: + "\connectedin X C; closedin X T; openin X T\ \ C \ T \ disjnt C T" + by (metis Diff_eq_empty_iff Int_empty_right clopenin_eq_frontier_of connectedin_Int_frontier_of disjnt_def) + +lemma connected_space_quotient_map_image: + "\quotient_map X X' q; connected_space X\ \ connected_space X'" + by (metis connectedin_continuous_map_image connectedin_topspace quotient_imp_continuous_map quotient_imp_surjective_map) + +lemma connected_space_retraction_map_image: + "\retraction_map X X' r; connected_space X\ \ connected_space X'" + using connected_space_quotient_map_image retraction_imp_quotient_map by blast + +lemma connectedin_imp_perfect_gen: + assumes X: "t1_space X" and S: "connectedin X S" and nontriv: "\a. S = {a}" + shows "S \ X derived_set_of S" +unfolding derived_set_of_def +proof (intro subsetI CollectI conjI strip) + show XS: "x \ topspace X" if "x \ S" for x + using that S connectedin by fastforce + show "\y. y \ x \ y \ S \ y \ T" + if "x \ S" and "x \ T \ openin X T" for x T + proof - + have opeXx: "openin X (topspace X - {x})" + by (meson X openin_topspace t1_space_openin_delete_alt) + moreover + have "S \ T \ (topspace X - {x})" + using XS that(2) by auto + moreover have "(topspace X - {x}) \ S \ {}" + by (metis Diff_triv S connectedin double_diff empty_subsetI inf_commute insert_subsetI nontriv that(1)) + ultimately show ?thesis + using that connectedinD [OF S, of T "topspace X - {x}"] + by blast + qed +qed + +lemma connectedin_imp_perfect: + "\Hausdorff_space X; connectedin X S; \a. S = {a}\ \ S \ X derived_set_of S" + by (simp add: Hausdorff_imp_t1_space connectedin_imp_perfect_gen) + + +proposition connected_space_prod_topology: + "connected_space(prod_topology X Y) \ + topspace(prod_topology X Y) = {} \ connected_space X \ connected_space Y" (is "?lhs=?rhs") +proof (cases "topspace(prod_topology X Y) = {}") + case True + then show ?thesis + using connected_space_topspace_empty by blast +next + case False + then have nonempty: "topspace X \ {}" "topspace Y \ {}" + by force+ + show ?thesis + proof + assume ?lhs + then show ?rhs + by (meson connected_space_quotient_map_image nonempty quotient_map_fst quotient_map_snd) + next + assume ?rhs + then have conX: "connected_space X" and conY: "connected_space Y" + using False by blast+ + have False + if "openin (prod_topology X Y) U" and "openin (prod_topology X Y) V" + and UV: "topspace X \ topspace Y \ U \ V" "U \ V = {}" + and "U \ {}" and "V \ {}" + for U V + proof - + have Usub: "U \ topspace X \ topspace Y" and Vsub: "V \ topspace X \ topspace Y" + using that by (metis openin_subset topspace_prod_topology)+ + obtain a b where ab: "(a,b) \ U" and a: "a \ topspace X" and b: "b \ topspace Y" + using \U \ {}\ Usub by auto + have "\ topspace X \ topspace Y \ U" + using Usub Vsub \U \ V = {}\ \V \ {}\ by auto + then obtain x y where x: "x \ topspace X" and y: "y \ topspace Y" and "(x,y) \ U" + by blast + have oX: "openin X {x \ topspace X. (x,y) \ U}" "openin X {x \ topspace X. (x,y) \ V}" + and oY: "openin Y {y \ topspace Y. (a,y) \ U}" "openin Y {y \ topspace Y. (a,y) \ V}" + by (force intro: openin_continuous_map_preimage [where Y = "prod_topology X Y"] + simp: that continuous_map_pairwise o_def x y a)+ + have 1: "topspace Y \ {y \ topspace Y. (a,y) \ U} \ {y \ topspace Y. (a,y) \ V}" + using a that(3) by auto + have 2: "{y \ topspace Y. (a,y) \ U} \ {y \ topspace Y. (a,y) \ V} = {}" + using that(4) by auto + have 3: "{y \ topspace Y. (a,y) \ U} \ {}" + using ab b by auto + have 4: "{y \ topspace Y. (a,y) \ V} \ {}" + proof - + show ?thesis + using connected_spaceD [OF conX oX] UV \(x,y) \ U\ a x y + disjoint_iff_not_equal by blast + qed + show ?thesis + using connected_spaceD [OF conY oY 1 2 3 4] by auto + qed + then show ?lhs + unfolding connected_space_def topspace_prod_topology by blast + qed +qed + +lemma connectedin_Times: + "connectedin (prod_topology X Y) (S \ T) \ + S = {} \ T = {} \ connectedin X S \ connectedin Y T" + by (force simp: connectedin_def subtopology_Times connected_space_prod_topology) + + +subsection\The notion of "separated between" (complement of "connected between)"\ + +definition separated_between + where "separated_between X S T \ + \U V. openin X U \ openin X V \ U \ V = topspace X \ disjnt U V \ S \ U \ T \ V" + +lemma separated_between_alt: + "separated_between X S T \ + (\U V. closedin X U \ closedin X V \ U \ V = topspace X \ disjnt U V \ S \ U \ T \ V)" + unfolding separated_between_def + by (metis separatedin_open_sets separation_closedin_Un_gen subtopology_topspace + separatedin_closed_sets separation_openin_Un_gen) + +lemma separated_between: + "separated_between X S T \ + (\U. closedin X U \ openin X U \ S \ U \ T \ topspace X - U)" + unfolding separated_between_def closedin_def disjnt_def + by (smt (verit, del_insts) Diff_cancel Diff_disjoint Diff_partition Un_Diff Un_Diff_Int openin_subset) + +lemma separated_between_mono: + "\separated_between X S T; S' \ S; T' \ T\ \ separated_between X S' T'" + by (meson order.trans separated_between) + +lemma separated_between_refl: + "separated_between X S S \ S = {}" + unfolding separated_between_def + by (metis Un_empty_right disjnt_def disjnt_empty2 disjnt_subset2 disjnt_sym le_iff_inf openin_empty openin_topspace) + +lemma separated_between_sym: + "separated_between X S T \ separated_between X T S" + by (metis disjnt_sym separated_between_alt sup_commute) + +lemma separated_between_imp_subset: + "separated_between X S T \ S \ topspace X \ T \ topspace X" + by (metis le_supI1 le_supI2 separated_between_def) + +lemma separated_between_empty: + "(separated_between X {} S \ S \ topspace X) \ (separated_between X S {} \ S \ topspace X)" + by (metis Diff_empty bot.extremum closedin_empty openin_empty separated_between separated_between_imp_subset separated_between_sym) + + +lemma separated_between_Un: + "separated_between X S (T \ U) \ separated_between X S T \ separated_between X S U" + by (auto simp: separated_between) + +lemma separated_between_Un': + "separated_between X (S \ T) U \ separated_between X S U \ separated_between X T U" + by (simp add: separated_between_Un separated_between_sym) + +lemma separated_between_imp_disjoint: + "separated_between X S T \ disjnt S T" + by (meson disjnt_iff separated_between_def subsetD) + +lemma separated_between_imp_separatedin: + "separated_between X S T \ separatedin X S T" + by (meson separated_between_def separatedin_mono separatedin_open_sets) + +lemma separated_between_full: + assumes "S \ T = topspace X" + shows "separated_between X S T \ disjnt S T \ closedin X S \ openin X S \ closedin X T \ openin X T" +proof - + have "separated_between X S T \ separatedin X S T" + by (simp add: separated_between_imp_separatedin) + then show ?thesis + unfolding separated_between_def + by (metis assms separation_closedin_Un_gen separation_openin_Un_gen subset_refl subtopology_topspace) +qed + +lemma separated_between_eq_separatedin: + "S \ T = topspace X \ (separated_between X S T \ separatedin X S T)" + by (simp add: separated_between_full separatedin_full) + +lemma separated_between_pointwise_left: + assumes "compactin X S" + shows "separated_between X S T \ + (S = {} \ T \ topspace X) \ (\x \ S. separated_between X {x} T)" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + using separated_between_imp_subset separated_between_mono by fastforce +next + assume R: ?rhs + then have "T \ topspace X" + by (meson equals0I separated_between_imp_subset) + show ?lhs + proof - + obtain U where U: "\x \ S. openin X (U x)" + "\x \ S. \V. openin X V \ U x \ V = topspace X \ disjnt (U x) V \ {x} \ U x \ T \ V" + using R unfolding separated_between_def by metis + then have "S \ \(U ` S)" + by blast + then obtain K where "finite K" "K \ S" and K: "S \ (\i\K. U i)" + using assms U unfolding compactin_def by (smt (verit) finite_subset_image imageE) + show ?thesis + unfolding separated_between + proof (intro conjI exI) + have "\x. x \ K \ closedin X (U x)" + by (smt (verit) \K \ S\ Diff_cancel U(2) Un_Diff Un_Diff_Int disjnt_def openin_closedin_eq subsetD) + then show "closedin X (\ (U ` K))" + by (metis (mono_tags, lifting) \finite K\ closedin_Union finite_imageI image_iff) + show "openin X (\ (U ` K))" + using U(1) \K \ S\ by blast + show "S \ \ (U ` K)" + by (simp add: K) + have "\x i. \x \ T; i \ K; x \ U i\ \ False" + by (meson U(2) \K \ S\ disjnt_iff subsetD) + then show "T \ topspace X - \ (U ` K)" + using \T \ topspace X\ by auto + qed + qed +qed + +lemma separated_between_pointwise_right: + "compactin X T + \ separated_between X S T \ (T = {} \ S \ topspace X) \ (\y \ T. separated_between X S {y})" + by (meson separated_between_pointwise_left separated_between_sym) + +lemma separated_between_closure_of: + "S \ topspace X \ separated_between X (X closure_of S) T \ separated_between X S T" + by (meson closure_of_minimal_eq separated_between_alt) + + +lemma separated_between_closure_of': + "T \ topspace X \ separated_between X S (X closure_of T) \ separated_between X S T" + by (meson separated_between_closure_of separated_between_sym) + +lemma separated_between_closure_of_eq: + "separated_between X S T \ S \ topspace X \ separated_between X (X closure_of S) T" + by (metis separated_between_closure_of separated_between_imp_subset) + +lemma separated_between_closure_of_eq': + "separated_between X S T \ T \ topspace X \ separated_between X S (X closure_of T)" + by (metis separated_between_closure_of' separated_between_imp_subset) + +lemma separated_between_frontier_of_eq': + "separated_between X S T \ + T \ topspace X \ disjnt S T \ separated_between X S (X frontier_of T)" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + by (metis interior_of_union_frontier_of separated_between_Un separated_between_closure_of_eq' + separated_between_imp_disjoint) +next + assume R: ?rhs + then obtain U where U: "closedin X U" "openin X U" "S \ U" "X closure_of T - X interior_of T \ topspace X - U" + by (metis frontier_of_def separated_between) + show ?lhs + proof (rule separated_between_mono [of _ S "X closure_of T"]) + have "separated_between X S T" + unfolding separated_between + proof (intro conjI exI) + show "S \ U - T" "T \ topspace X - (U - T)" + using R U(3) by (force simp: disjnt_iff)+ + have "T \ X closure_of T" + by (simp add: R closure_of_subset) + then have *: "U - T = U - X interior_of T" + using U(4) interior_of_subset by fastforce + then show "closedin X (U - T)" + by (simp add: U(1) closedin_diff) + have "U \ X frontier_of T = {}" + using U(4) frontier_of_def by fastforce + then show "openin X (U - T)" + by (metis * Diff_Un U(2) Un_Diff_Int Un_Int_eq(1) closedin_closure_of interior_of_union_frontier_of openin_diff sup_bot_right) + qed + then show "separated_between X S (X closure_of T)" + by (simp add: R separated_between_closure_of') + qed (auto simp: R closure_of_subset) +qed + +lemma separated_between_frontier_of_eq: + "separated_between X S T \ S \ topspace X \ disjnt S T \ separated_between X (X frontier_of S) T" + by (metis disjnt_sym separated_between_frontier_of_eq' separated_between_sym) + +lemma separated_between_frontier_of: + "\S \ topspace X; disjnt S T\ + \ (separated_between X (X frontier_of S) T \ separated_between X S T)" + using separated_between_frontier_of_eq by blast + +lemma separated_between_frontier_of': + "\T \ topspace X; disjnt S T\ + \ (separated_between X S (X frontier_of T) \ separated_between X S T)" + using separated_between_frontier_of_eq' by auto + +lemma connected_space_separated_between: + "connected_space X \ (\S T. separated_between X S T \ S = {} \ T = {})" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + by (metis Diff_cancel connected_space_clopen_in separated_between subset_empty) +next + assume ?rhs then show ?lhs + by (meson connected_space_eq_not_separated separated_between_eq_separatedin) +qed + +lemma connected_space_imp_separated_between_trivial: + "connected_space X + \ (separated_between X S T \ S = {} \ T \ topspace X \ S \ topspace X \ T = {})" + by (metis connected_space_separated_between separated_between_empty) + + +subsection\Connected components\ + +lemma connected_component_of_subtopology_eq: + "connected_component_of (subtopology X U) a = connected_component_of X a \ + connected_component_of_set X a \ U" + by (force simp: connected_component_of_set connectedin_subtopology connected_component_of_def fun_eq_iff subset_iff) + +lemma connected_components_of_subtopology: + assumes "C \ connected_components_of X" "C \ U" + shows "C \ connected_components_of (subtopology X U)" +proof - + obtain a where a: "connected_component_of_set X a \ U" and "a \ topspace X" + and Ceq: "C = connected_component_of_set X a" + using assms by (force simp: connected_components_of_def) + then have "a \ U" + by (simp add: connected_component_of_refl in_mono) + then have "connected_component_of_set X a = connected_component_of_set (subtopology X U) a" + by (metis a connected_component_of_subtopology_eq) + then show ?thesis + by (simp add: Ceq \a \ U\ \a \ topspace X\ connected_component_in_connected_components_of) +qed + +thm connected_space_iff_components_eq + +lemma open_in_finite_connected_components: + assumes "finite(connected_components_of X)" "C \ connected_components_of X" + shows "openin X C" +proof - + have "closedin X (topspace X - C)" + by (metis DiffD1 assms closedin_Union closedin_connected_components_of complement_connected_components_of_Union finite_Diff) + then show ?thesis + by (simp add: assms connected_components_of_subset openin_closedin) +qed +thm connected_component_of_eq_overlap + +lemma connected_components_of_disjoint: + assumes "C \ connected_components_of X" "C' \ connected_components_of X" + shows "(disjnt C C' \ (C \ C'))" +proof - + have "C \ {}" + using nonempty_connected_components_of assms by blast + with assms show ?thesis + by (metis disjnt_self_iff_empty pairwiseD pairwise_disjoint_connected_components_of) +qed + +lemma connected_components_of_overlap: + "\C \ connected_components_of X; C' \ connected_components_of X\ \ C \ C' \ {} \ C = C'" + by (meson connected_components_of_disjoint disjnt_def) + +lemma pairwise_separated_connected_components_of: + "pairwise (separatedin X) (connected_components_of X)" + by (simp add: closedin_connected_components_of connected_components_of_disjoint pairwiseI separatedin_closed_sets) + +lemma finite_connected_components_of_finite: + "finite(topspace X) \ finite(connected_components_of X)" + by (simp add: Union_connected_components_of finite_UnionD) + +lemma connected_component_of_unique: + "\x \ C; connectedin X C; \C'. x \ C' \ connectedin X C' \ C' \ C\ + \ connected_component_of_set X x = C" + by (meson connected_component_of_maximal connectedin_connected_component_of subsetD subset_antisym) + +lemma closedin_connected_component_of_subtopology: + "\C \ connected_components_of (subtopology X s); X closure_of C \ s\ \ closedin X C" + by (metis closedin_Int_closure_of closedin_connected_components_of closure_of_eq inf.absorb_iff2) + +lemma connected_component_of_discrete_topology: + "connected_component_of_set (discrete_topology U) x = (if x \ U then {x} else {})" + by (simp add: locally_path_connected_space_discrete_topology flip: path_component_eq_connected_component_of) + +lemma connected_components_of_discrete_topology: + "connected_components_of (discrete_topology U) = (\x. {x}) ` U" + by (simp add: connected_component_of_discrete_topology connected_components_of_def) + +lemma connected_component_of_continuous_image: + "\continuous_map X Y f; connected_component_of X x y\ + \ connected_component_of Y (f x) (f y)" + by (meson connected_component_of_def connectedin_continuous_map_image image_eqI) + +lemma homeomorphic_map_connected_component_of: + assumes "homeomorphic_map X Y f" and x: "x \ topspace X" + shows "connected_component_of_set Y (f x) = f ` (connected_component_of_set X x)" +proof - + obtain g where g: "continuous_map X Y f" + "continuous_map Y X g " "\x. x \ topspace X \ g (f x) = x" + "\y. y \ topspace Y \ f (g y) = y" + using assms(1) homeomorphic_map_maps homeomorphic_maps_def by fastforce + show ?thesis + using connected_component_in_topspace [of Y] x g + connected_component_of_continuous_image [of X Y f] + connected_component_of_continuous_image [of Y X g] + by force +qed + +lemma homeomorphic_map_connected_components_of: + assumes "homeomorphic_map X Y f" + shows "connected_components_of Y = (image f) ` (connected_components_of X)" +proof - + have "topspace Y = f ` topspace X" + by (metis assms homeomorphic_imp_surjective_map) + with homeomorphic_map_connected_component_of [OF assms] show ?thesis + by (auto simp: connected_components_of_def image_iff) +qed + +lemma connected_component_of_pair: + "connected_component_of_set (prod_topology X Y) (x,y) = + connected_component_of_set X x \ connected_component_of_set Y y" +proof (cases "x \ topspace X \ y \ topspace Y") + case True + show ?thesis + proof (rule connected_component_of_unique) + show "(x, y) \ connected_component_of_set X x \ connected_component_of_set Y y" + using True by (simp add: connected_component_of_refl) + show "connectedin (prod_topology X Y) (connected_component_of_set X x \ connected_component_of_set Y y)" + by (metis connectedin_Times connectedin_connected_component_of) + show "C \ connected_component_of_set X x \ connected_component_of_set Y y" + if "(x, y) \ C \ connectedin (prod_topology X Y) C" for C + using that unfolding connected_component_of_def + apply clarsimp + by (metis (no_types) connectedin_continuous_map_image continuous_map_fst continuous_map_snd fst_conv imageI snd_conv) + qed +next + case False then show ?thesis + by (metis Sigma_empty1 Sigma_empty2 connected_component_of_eq_empty mem_Sigma_iff topspace_prod_topology) +qed + +lemma connected_components_of_prod_topology: + "connected_components_of (prod_topology X Y) = + {C \ D |C D. C \ connected_components_of X \ D \ connected_components_of Y}" (is "?lhs=?rhs") +proof + show "?lhs \ ?rhs" + apply (clarsimp simp: connected_components_of_def) + by (metis (no_types) connected_component_of_pair imageI) +next + show "?rhs \ ?lhs" + using connected_component_of_pair + by (fastforce simp: connected_components_of_def) +qed + + +lemma connected_component_of_product_topology: + "connected_component_of_set (product_topology X I) x = + (if x \ extensional I then PiE I (\i. connected_component_of_set (X i) (x i)) else {})" + (is "?lhs = If _ ?R _") +proof (cases "x \ topspace(product_topology X I)") + case True + have "?lhs = (\\<^sub>E i\I. connected_component_of_set (X i) (x i))" + if xX: "\i. i\I \ x i \ topspace (X i)" and ext: "x \ extensional I" + proof (rule connected_component_of_unique) + show "x \ ?R" + by (simp add: PiE_iff connected_component_of_refl local.ext xX) + show "connectedin (product_topology X I) ?R" + by (simp add: connectedin_PiE connectedin_connected_component_of) + show "C \ ?R" + if "x \ C \ connectedin (product_topology X I) C" for C + proof - + have "C \ extensional I" + using PiE_def connectedin_subset_topspace that by fastforce + have "\y. y \ C \ y \ (\ i\I. connected_component_of_set (X i) (x i))" + apply (simp add: connected_component_of_def Pi_def) + by (metis connectedin_continuous_map_image continuous_map_product_projection imageI that) + then show ?thesis + using PiE_def \C \ extensional I\ by fastforce + qed + qed + with True show ?thesis + by (simp add: PiE_iff) +next + case False + then show ?thesis + apply (simp add: PiE_iff) + by (smt (verit) Collect_empty_eq False PiE_eq_empty_iff PiE_iff connected_component_of_eq_empty) +qed + + +lemma connected_components_of_product_topology: + "connected_components_of (product_topology X I) = + {PiE I B |B. \i \ I. B i \ connected_components_of(X i)}" (is "?lhs=?rhs") +proof + show "?lhs \ ?rhs" + by (auto simp: connected_components_of_def connected_component_of_product_topology PiE_iff) + show "?rhs \ ?lhs" + proof + fix F + assume "F \ ?rhs" + then obtain B where Feq: "F = Pi\<^sub>E I B" and + "\i\I. \x\topspace (X i). B i = connected_component_of_set (X i) x" + by (force simp: connected_components_of_def connected_component_of_product_topology image_iff) + then obtain f where + f: "\i. i \ I \ f i \ topspace (X i) \ B i = connected_component_of_set (X i) (f i)" + by metis + then have "(\i\I. f i) \ ((\\<^sub>E i\I. topspace (X i)) \ extensional I)" + by simp + with f show "F \ ?lhs" + unfolding Feq connected_components_of_def connected_component_of_product_topology image_iff + by (smt (verit, del_insts) PiE_cong restrict_PiE_iff restrict_apply' restrict_extensional topspace_product_topology) + qed +qed + + +subsection \Monotone maps (in the general topological sense)\ + + +definition monotone_map + where "monotone_map X Y f == + f ` (topspace X) \ topspace Y \ + (\y \ topspace Y. connectedin X {x \ topspace X. f x = y})" + +lemma monotone_map: + "monotone_map X Y f \ + f ` (topspace X) \ topspace Y \ (\y. connectedin X {x \ topspace X. f x = y})" + apply (simp add: monotone_map_def) + by (metis (mono_tags, lifting) connectedin_empty [of X] Collect_empty_eq image_subset_iff) + + +lemma monotone_map_in_subtopology: + "monotone_map X (subtopology Y S) f \ monotone_map X Y f \ f ` (topspace X) \ S" + by (smt (verit, del_insts) le_inf_iff monotone_map topspace_subtopology) + +lemma monotone_map_from_subtopology: + assumes "monotone_map X Y f" + "\x y. x \ topspace X \ y \ topspace X \ x \ S \ f x = f y \ y \ S" + shows "monotone_map (subtopology X S) Y f" + using assms + unfolding monotone_map_def connectedin_subtopology + by (smt (verit, del_insts) Collect_cong Collect_empty_eq IntE IntI connectedin_empty image_subset_iff mem_Collect_eq subsetI topspace_subtopology) + +lemma monotone_map_restriction: + "monotone_map X Y f \ {x \ topspace X. f x \ v} = u + \ monotone_map (subtopology X u) (subtopology Y v) f" + by (smt (verit, best) IntI Int_Collect image_subset_iff mem_Collect_eq monotone_map monotone_map_from_subtopology topspace_subtopology) + +lemma injective_imp_monotone_map: + assumes "f ` topspace X \ topspace Y" "inj_on f (topspace X)" + shows "monotone_map X Y f" + unfolding monotone_map_def +proof (intro conjI assms strip) + fix y + assume "y \ topspace Y" + then have "{x \ topspace X. f x = y} = {} \ (\a \ topspace X. {x \ topspace X. f x = y} = {a})" + using assms(2) unfolding inj_on_def by blast + then show "connectedin X {x \ topspace X. f x = y}" + by (metis (no_types, lifting) connectedin_empty connectedin_sing) +qed + +lemma embedding_imp_monotone_map: + "embedding_map X Y f \ monotone_map X Y f" + by (metis (no_types) embedding_map_def homeomorphic_eq_everything_map inf.absorb_iff2 injective_imp_monotone_map topspace_subtopology) + +lemma section_imp_monotone_map: + "section_map X Y f \ monotone_map X Y f" + by (simp add: embedding_imp_monotone_map section_imp_embedding_map) + +lemma homeomorphic_imp_monotone_map: + "homeomorphic_map X Y f \ monotone_map X Y f" + by (meson section_and_retraction_eq_homeomorphic_map section_imp_monotone_map) + +proposition connected_space_monotone_quotient_map_preimage: + assumes f: "monotone_map X Y f" "quotient_map X Y f" and "connected_space Y" + shows "connected_space X" +proof (rule ccontr) + assume "\ connected_space X" + then obtain U V where "openin X U" "openin X V" "U \ V = {}" + "U \ {}" "V \ {}" and topUV: "topspace X \ U \ V" + by (auto simp: connected_space_def) + then have UVsub: "U \ topspace X" "V \ topspace X" + by (auto simp: openin_subset) + have "\ connected_space Y" + unfolding connected_space_def not_not + proof (intro exI conjI) + show "topspace Y \ f`U \ f`V" + by (metis f(2) image_Un quotient_imp_surjective_map subset_Un_eq topUV) + show "f`U \ {}" + by (simp add: \U \ {}\) + show "(f`V) \ {}" + by (simp add: \V \ {}\) + have *: "y \ f ` V" if "y \ f ` U" for y + proof - + have \
: "connectedin X {x \ topspace X. f x = y}" + using f(1) monotone_map by fastforce + show ?thesis + using connectedinD [OF \
\openin X U\ \openin X V\] UVsub topUV \U \ V = {}\ that + by (force simp: disjoint_iff) + qed + then show "f`U \ f`V = {}" + by blast + show "openin Y (f`U)" + using f \openin X U\ topUV * unfolding quotient_map_saturated_open by force + show "openin Y (f`V)" + using f \openin X V\ topUV * unfolding quotient_map_saturated_open by force + qed + then show False + by (simp add: assms) +qed + +lemma connectedin_monotone_quotient_map_preimage: + assumes "monotone_map X Y f" "quotient_map X Y f" "connectedin Y C" "openin Y C \ closedin Y C" + shows "connectedin X {x \ topspace X. f x \ C}" +proof - + have "connected_space (subtopology X {x \ topspace X. f x \ C})" + proof - + have "connected_space (subtopology Y C)" + using \connectedin Y C\ connectedin_def by blast + moreover have "quotient_map (subtopology X {a \ topspace X. f a \ C}) (subtopology Y C) f" + by (simp add: assms quotient_map_restriction) + ultimately show ?thesis + using \monotone_map X Y f\ connected_space_monotone_quotient_map_preimage monotone_map_restriction by blast + qed + then show ?thesis + by (simp add: connectedin_def) +qed + +lemma monotone_open_map: + assumes "continuous_map X Y f" "open_map X Y f" and fim: "f ` (topspace X) = topspace Y" + shows "monotone_map X Y f \ (\C. connectedin Y C \ connectedin X {x \ topspace X. f x \ C})" + (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + unfolding connectedin_def + proof (intro strip conjI) + fix C + assume C: "C \ topspace Y \ connected_space (subtopology Y C)" + show "connected_space (subtopology X {x \ topspace X. f x \ C})" + proof (rule connected_space_monotone_quotient_map_preimage) + show "monotone_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" + by (simp add: L monotone_map_restriction) + show "quotient_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" + proof (rule continuous_open_imp_quotient_map) + show "continuous_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" + using assms continuous_map_from_subtopology continuous_map_in_subtopology by fastforce + qed (use open_map_restriction assms in fastforce)+ + qed (simp add: C) + qed auto +next + assume ?rhs + then have "\y. connectedin Y {y} \ connectedin X {x \ topspace X. f x = y}" + by (smt (verit) Collect_cong singletonD singletonI) + then show ?lhs + by (simp add: fim monotone_map_def) +qed + +lemma monotone_closed_map: + assumes "continuous_map X Y f" "closed_map X Y f" and fim: "f ` (topspace X) = topspace Y" + shows "monotone_map X Y f \ (\C. connectedin Y C \ connectedin X {x \ topspace X. f x \ C})" + (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + unfolding connectedin_def + proof (intro strip conjI) + fix C + assume C: "C \ topspace Y \ connected_space (subtopology Y C)" + show "connected_space (subtopology X {x \ topspace X. f x \ C})" + proof (rule connected_space_monotone_quotient_map_preimage) + show "monotone_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" + by (simp add: L monotone_map_restriction) + show "quotient_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" + proof (rule continuous_closed_imp_quotient_map) + show "continuous_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" + using assms continuous_map_from_subtopology continuous_map_in_subtopology by fastforce + qed (use closed_map_restriction assms in fastforce)+ + qed (simp add: C) + qed auto +next + assume ?rhs + then have "\y. connectedin Y {y} \ connectedin X {x \ topspace X. f x = y}" + by (smt (verit) Collect_cong singletonD singletonI) + then show ?lhs + by (simp add: fim monotone_map_def) +qed + +subsection\Other countability properties\ + +definition second_countable + where "second_countable X \ + \\. countable \ \ (\V \ \. openin X V) \ + (\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U))" + +definition first_countable + where "first_countable X \ + \x \ topspace X. + \\. countable \ \ (\V \ \. openin X V) \ + (\U. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U))" + +definition separable_space + where "separable_space X \ + \C. countable C \ C \ topspace X \ X closure_of C = topspace X" + +lemma second_countable: + "second_countable X \ + (\\. countable \ \ openin X = arbitrary union_of (\x. x \ \))" + by (smt (verit) openin_topology_base_unique second_countable_def) + +lemma second_countable_subtopology: + assumes "second_countable X" + shows "second_countable (subtopology X S)" +proof - + obtain \ where \: "countable \" "\V. V \ \ \ openin X V" + "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" + using assms by (auto simp: second_countable_def) + show ?thesis + unfolding second_countable_def + proof (intro exI conjI) + show "\V\((\)S) ` \. openin (subtopology X S) V" + using openin_subtopology_Int2 \ by blast + show "\U x. openin (subtopology X S) U \ x \ U \ (\V\((\)S) ` \. x \ V \ V \ U)" + using \ unfolding openin_subtopology + by (smt (verit, del_insts) IntI image_iff inf_commute inf_le1 subset_iff) + qed (use \ in auto) +qed + + +lemma second_countable_discrete_topology: + "second_countable(discrete_topology U) \ countable U" (is "?lhs=?rhs") +proof + assume L: ?lhs + then + obtain \ where \: "countable \" "\V. V \ \ \ V \ U" + "\W x. W \ U \ x \ W \ (\V \ \. x \ V \ V \ W)" + by (auto simp: second_countable_def) + then have "{x} \ \" if "x \ U" for x + by (metis empty_subsetI insertCI insert_subset subset_antisym that) + then show ?rhs + by (smt (verit) countable_subset image_subsetI \countable \\ countable_image_inj_on [OF _ inj_singleton]) +next + assume ?rhs + then show ?lhs + unfolding second_countable_def + by (rule_tac x="(\x. {x}) ` U" in exI) auto +qed + +lemma second_countable_open_map_image: + assumes "continuous_map X Y f" "open_map X Y f" + and fim: "f ` (topspace X) = topspace Y" and "second_countable X" + shows "second_countable Y" +proof - + have openXYf: "\U. openin X U \ openin Y (f ` U)" + using assms by (auto simp: open_map_def) + obtain \ where \: "countable \" "\V. V \ \ \ openin X V" + and *: "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" + using assms by (auto simp: second_countable_def) + show ?thesis + unfolding second_countable_def + proof (intro exI conjI strip) + fix V y + assume V: "openin Y V \ y \ V" + then obtain x where "x \ topspace X" and x: "f x = y" + by (metis fim image_iff openin_subset subsetD) + + then obtain W where "W\\" "x \ W" "W \ {x \ topspace X. f x \ V}" + using * [of "{x \ topspace X. f x \ V}" x] V assms openin_continuous_map_preimage + by force + then show "\W \ (image f) ` \. y \ W \ W \ V" + using x by auto + qed (use \ openXYf in auto) +qed + +lemma homeomorphic_space_second_countability: + "X homeomorphic_space Y \ (second_countable X \ second_countable Y)" + by (meson homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym second_countable_open_map_image) + +lemma second_countable_retraction_map_image: + "\retraction_map X Y r; second_countable X\ \ second_countable Y" + using hereditary_imp_retractive_property homeomorphic_space_second_countability second_countable_subtopology by blast + +lemma second_countable_imp_first_countable: + "second_countable X \ first_countable X" + by (metis first_countable_def second_countable_def) + +lemma first_countable_subtopology: + assumes "first_countable X" + shows "first_countable (subtopology X S)" + unfolding first_countable_def +proof + fix x + assume "x \ topspace (subtopology X S)" + then obtain \ where "countable \" and \: "\V. V \ \ \ openin X V" + "\U. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" + using assms first_countable_def by force + show "\\. countable \ \ (\V\\. openin (subtopology X S) V) \ (\U. openin (subtopology X S) U \ x \ U \ (\V\\. x \ V \ V \ U))" + proof (intro exI conjI strip) + show "countable (((\)S) ` \)" + using \countable \\ by blast + show "openin (subtopology X S) V" if "V \ ((\)S) ` \" for V + using \ openin_subtopology_Int2 that by fastforce + show "\V\((\)S) ` \. x \ V \ V \ U" + if "openin (subtopology X S) U \ x \ U" for U + using that \(2) by (clarsimp simp: openin_subtopology) (meson le_infI2) + qed +qed + +lemma first_countable_discrete_topology: + "first_countable (discrete_topology U)" + unfolding first_countable_def topspace_discrete_topology openin_discrete_topology +proof + fix x assume "x \ U" + show "\\. countable \ \ (\V\\. V \ U) \ (\Ua. Ua \ U \ x \ Ua \ (\V\\. x \ V \ V \ Ua))" + using \x \ U\ by (rule_tac x="{{x}}" in exI) auto +qed + +lemma first_countable_open_map_image: + assumes "continuous_map X Y f" "open_map X Y f" + and fim: "f ` (topspace X) = topspace Y" and "first_countable X" + shows "first_countable Y" + unfolding first_countable_def +proof + fix y + assume "y \ topspace Y" + have openXYf: "\U. openin X U \ openin Y (f ` U)" + using assms by (auto simp: open_map_def) + then obtain x where x: "x \ topspace X" "f x = y" + by (metis \y \ topspace Y\ fim imageE) + obtain \ where \: "countable \" "\V. V \ \ \ openin X V" + and *: "\U. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" + using assms x first_countable_def by force + show "\\. countable \ \ + (\V\\. openin Y V) \ (\U. openin Y U \ y \ U \ (\V\\. y \ V \ V \ U))" + proof (intro exI conjI strip) + fix V assume "openin Y V \ y \ V" + then have "\W\\. x \ W \ W \ {x \ topspace X. f x \ V}" + using * [of "{x \ topspace X. f x \ V}"] assms openin_continuous_map_preimage x + by fastforce + then show "\V' \ (image f) ` \. y \ V' \ V' \ V" + using image_mono x by auto + qed (use \ openXYf in force)+ +qed + +lemma homeomorphic_space_first_countability: + "X homeomorphic_space Y \ first_countable X \ first_countable Y" + by (meson first_countable_open_map_image homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym) + +lemma first_countable_retraction_map_image: + "\retraction_map X Y r; first_countable X\ \ first_countable Y" + using first_countable_subtopology hereditary_imp_retractive_property homeomorphic_space_first_countability by blast + +lemma separable_space_open_subset: + assumes "separable_space X" "openin X S" + shows "separable_space (subtopology X S)" +proof - + obtain C where C: "countable C" "C \ topspace X" "X closure_of C = topspace X" + by (meson assms separable_space_def) + then have "\x T. \x \ topspace X; x \ T; openin (subtopology X S) T\ + \ \y. y \ S \ y \ C \ y \ T" + by (smt (verit) \openin X S\ in_closure_of openin_open_subtopology subsetD) + with C \openin X S\ show ?thesis + unfolding separable_space_def + by (rule_tac x="S \ C" in exI) (force simp: in_closure_of) +qed + +lemma separable_space_continuous_map_image: + assumes "separable_space X" "continuous_map X Y f" + and fim: "f ` (topspace X) = topspace Y" + shows "separable_space Y" +proof - + have cont: "\S. f ` (X closure_of S) \ Y closure_of f ` S" + by (simp add: assms continuous_map_image_closure_subset) + obtain C where C: "countable C" "C \ topspace X" "X closure_of C = topspace X" + by (meson assms separable_space_def) + then show ?thesis + unfolding separable_space_def + by (metis cont fim closure_of_subset_topspace countable_image image_mono subset_antisym) +qed + +lemma separable_space_quotient_map_image: + "\quotient_map X Y q; separable_space X\ \ separable_space Y" + by (meson quotient_imp_continuous_map quotient_imp_surjective_map separable_space_continuous_map_image) + +lemma separable_space_retraction_map_image: + "\retraction_map X Y r; separable_space X\ \ separable_space Y" + using retraction_imp_quotient_map separable_space_quotient_map_image by blast + +lemma homeomorphic_separable_space: + "X homeomorphic_space Y \ (separable_space X \ separable_space Y)" + by (meson homeomorphic_eq_everything_map homeomorphic_maps_map homeomorphic_space_def separable_space_continuous_map_image) + +lemma separable_space_discrete_topology: + "separable_space(discrete_topology U) \ countable U" + by (metis countable_Int2 discrete_topology_closure_of dual_order.refl inf.orderE separable_space_def topspace_discrete_topology) + +lemma second_countable_imp_separable_space: + assumes "second_countable X" + shows "separable_space X" +proof - + obtain \ where \: "countable \" "\V. V \ \ \ openin X V" + and *: "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" + using assms by (auto simp: second_countable_def) + obtain c where c: "\V. \V \ \; V \ {}\ \ c V \ V" + by (metis all_not_in_conv) + then have **: "\x. x \ topspace X \ x \ X closure_of c ` (\ - {{}})" + using * by (force simp: closure_of_def) + show ?thesis + unfolding separable_space_def + proof (intro exI conjI) + show "countable (c ` (\-{{}}))" + using \(1) by blast + show "(c ` (\-{{}})) \ topspace X" + using \(2) c openin_subset by fastforce + show "X closure_of (c ` (\-{{}})) = topspace X" + by (meson ** closure_of_subset_topspace subsetI subset_antisym) + qed +qed + +lemma second_countable_imp_Lindelof_space: + assumes "second_countable X" + shows "Lindelof_space X" +unfolding Lindelof_space_def +proof clarify + fix \ + assume "\U \ \. openin X U" and UU: "\\ = topspace X" + obtain \ where \: "countable \" "\V. V \ \ \ openin X V" + and *: "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" + using assms by (auto simp: second_countable_def) + define \' where "\' = {B \ \. \U. U \ \ \ B \ U}" + have \': "countable \'" "\\' = \\" + using \ using "*" \\U\\. openin X U\ by (fastforce simp: \'_def)+ + have "\b. \U. b \ \' \ U \ \ \ b \ U" + by (simp add: \'_def) + then obtain G where G: "\b. b \ \' \ G b \ \ \ b \ G b" + by metis + with \' UU show "\\. countable \ \ \ \ \ \ \\ = topspace X" + by (rule_tac x="G ` \'" in exI) fastforce +qed + +subsection \Neigbourhood bases EXTRAS\ +(* Neigbourhood bases (useful for "local" properties of various kind). *) + +lemma openin_topology_neighbourhood_base_unique: + "openin X = arbitrary union_of P \ + (\u. P u \ openin X u) \ neighbourhood_base_of P X" + by (smt (verit, best) open_neighbourhood_base_of openin_topology_base_unique) + +lemma neighbourhood_base_at_topology_base: + " openin X = arbitrary union_of b + \ (neighbourhood_base_at x P X \ + (\w. b w \ x \ w \ (\u v. openin X u \ P v \ x \ u \ u \ v \ v \ w)))" + apply (simp add: neighbourhood_base_at_def) + by (smt (verit, del_insts) openin_topology_base_unique subset_trans) + +lemma neighbourhood_base_of_unlocalized: + assumes "\S t. P S \ openin X t \ (t \ {}) \ t \ S \ P t" + shows "neighbourhood_base_of P X \ + (\x \ topspace X. \u v. openin X u \ P v \ x \ u \ u \ v \ v \ topspace X)" + apply (simp add: neighbourhood_base_of_def) + by (smt (verit, ccfv_SIG) assms empty_iff neighbourhood_base_at_unlocalized) + +lemma neighbourhood_base_at_discrete_topology: + "neighbourhood_base_at x P (discrete_topology u) \ x \ u \ P {x}" + apply (simp add: neighbourhood_base_at_def) + by (smt (verit) empty_iff empty_subsetI insert_subset singletonI subsetD subset_singletonD) + +lemma neighbourhood_base_of_discrete_topology: + "neighbourhood_base_of P (discrete_topology u) \ (\x \ u. P {x})" + apply (simp add: neighbourhood_base_of_def) + using neighbourhood_base_at_discrete_topology[of _ P u] + by (metis empty_subsetI insert_subset neighbourhood_base_at_def openin_discrete_topology singletonI) + +lemma second_countable_neighbourhood_base_alt: + "second_countable X \ + (\\. countable \ \ (\V \ \. openin X V) \ neighbourhood_base_of (\A. A\\) X)" + by (metis (full_types) openin_topology_neighbourhood_base_unique second_countable) + +lemma first_countable_neighbourhood_base_alt: + "first_countable X \ + (\x \ topspace X. \\. countable \ \ (\V \ \. openin X V) \ neighbourhood_base_at x (\V. V \ \) X)" + unfolding first_countable_def + apply (intro ball_cong refl ex_cong conj_cong) + by (metis (mono_tags, lifting) open_neighbourhood_base_at) + +lemma second_countable_neighbourhood_base: + "second_countable X \ + (\\. countable \ \ neighbourhood_base_of (\V. V \ \) X)" (is "?lhs=?rhs") +proof + assume ?lhs + then show ?rhs + using second_countable_neighbourhood_base_alt by blast +next + assume ?rhs + then obtain \ where "countable \" + and \: "\W x. openin X W \ x \ W \ (\U. openin X U \ (\V. V \ \ \ x \ U \ U \ V \ V \ W))" + by (metis neighbourhood_base_of) + then show ?lhs + unfolding second_countable_neighbourhood_base_alt neighbourhood_base_of + apply (rule_tac x="(\u. X interior_of u) ` \" in exI) + by (smt (verit, best) interior_of_eq interior_of_mono countable_image image_iff openin_interior_of) +qed + +lemma first_countable_neighbourhood_base: + "first_countable X \ + (\x \ topspace X. \\. countable \ \ neighbourhood_base_at x (\V. V \ \) X)" (is "?lhs=?rhs") +proof + assume ?lhs + then show ?rhs + by (metis first_countable_neighbourhood_base_alt) +next + assume R: ?rhs + show ?lhs + unfolding first_countable_neighbourhood_base_alt + proof + fix x + assume "x \ topspace X" + with R obtain \ where "countable \" and \: "neighbourhood_base_at x (\V. V \ \) X" + by blast + then + show "\\. countable \ \ Ball \ (openin X) \ neighbourhood_base_at x (\V. V \ \) X" + unfolding neighbourhood_base_at_def + apply (rule_tac x="(\u. X interior_of u) ` \" in exI) + by (smt (verit, best) countable_image image_iff interior_of_eq interior_of_mono openin_interior_of) + qed +qed + + +subsection\$T_0$ spaces and the Kolmogorov quotient\ + +definition t0_space where + "t0_space X \ + \x \ topspace X. \y \ topspace X. x \ y \ (\U. openin X U \ (x \ U \ y \ U))" + +lemma t0_space_expansive: + "\topspace Y = topspace X; \U. openin X U \ openin Y U\ \ t0_space X \ t0_space Y" + by (metis t0_space_def) + +lemma t1_imp_t0_space: "t1_space X \ t0_space X" + by (metis t0_space_def t1_space_def) + +lemma t1_eq_symmetric_t0_space_alt: + "t1_space X \ + t0_space X \ + (\x \ topspace X. \y \ topspace X. x \ X closure_of {y} \ y \ X closure_of {x})" + apply (simp add: t0_space_def t1_space_def closure_of_def) + by (smt (verit, best) openin_topspace) + +lemma t1_eq_symmetric_t0_space: + "t1_space X \ t0_space X \ (\x y. x \ X closure_of {y} \ y \ X closure_of {x})" + by (auto simp: t1_eq_symmetric_t0_space_alt in_closure_of) + +lemma Hausdorff_imp_t0_space: + "Hausdorff_space X \ t0_space X" + by (simp add: Hausdorff_imp_t1_space t1_imp_t0_space) + +lemma t0_space: + "t0_space X \ + (\x \ topspace X. \y \ topspace X. x \ y \ (\C. closedin X C \ (x \ C \ y \ C)))" + unfolding t0_space_def by (metis Diff_iff closedin_def openin_closedin_eq) + +lemma homeomorphic_t0_space: + assumes "X homeomorphic_space Y" + shows "t0_space X \ t0_space Y" +proof - + obtain f where f: "homeomorphic_map X Y f" and F: "inj_on f (topspace X)" and "topspace Y = f ` topspace X" + by (metis assms homeomorphic_imp_injective_map homeomorphic_imp_surjective_map homeomorphic_space) + with inj_on_image_mem_iff [OF F] + show ?thesis + apply (simp add: t0_space_def homeomorphic_eq_everything_map continuous_map_def open_map_def inj_on_def) + by (smt (verit) mem_Collect_eq openin_subset) +qed + +lemma t0_space_closure_of_sing: + "t0_space X \ + (\x \ topspace X. \y \ topspace X. X closure_of {x} = X closure_of {y} \ x = y)" + by (simp add: t0_space_def closure_of_def set_eq_iff) (smt (verit)) + +lemma t0_space_discrete_topology: "t0_space (discrete_topology S)" + by (simp add: Hausdorff_imp_t0_space) + +lemma t0_space_subtopology: "t0_space X \ t0_space (subtopology X U)" + by (simp add: t0_space_def openin_subtopology) (metis Int_iff) + +lemma t0_space_retraction_map_image: + "\retraction_map X Y r; t0_space X\ \ t0_space Y" + using hereditary_imp_retractive_property homeomorphic_t0_space t0_space_subtopology by blast + +lemma XY: "{x}\{y} = {(x,y)}" + by simp + +lemma t0_space_prod_topologyI: "\t0_space X; t0_space Y\ \ t0_space (prod_topology X Y)" + by (simp add: t0_space_closure_of_sing closure_of_Times closure_of_eq_empty_gen times_eq_iff flip: XY insert_Times_insert) + + +lemma t0_space_prod_topology_iff: + "t0_space (prod_topology X Y) \ topspace (prod_topology X Y) = {} \ t0_space X \ t0_space Y" (is "?lhs=?rhs") +proof + assume ?lhs + then show ?rhs + by (metis Sigma_empty1 Sigma_empty2 retraction_map_fst retraction_map_snd t0_space_retraction_map_image topspace_prod_topology) +qed (metis empty_iff t0_space_def t0_space_prod_topologyI) + +proposition t0_space_product_topology: + "t0_space (product_topology X I) \ + topspace(product_topology X I) = {} \ (\i \ I. t0_space (X i))" (is "?lhs=?rhs") +proof + assume ?lhs + then show ?rhs + by (meson retraction_map_product_projection t0_space_retraction_map_image) +next + assume R: ?rhs + show ?lhs + proof (cases "topspace(product_topology X I) = {}") + case True + then show ?thesis + by (simp add: t0_space_def) + next + case False + show ?thesis + unfolding t0_space + proof (intro strip) + fix x y + assume x: "x \ topspace (product_topology X I)" + and y: "y \ topspace (product_topology X I)" + and "x \ y" + then obtain i where "i \ I" "x i \ y i" + by (metis PiE_ext topspace_product_topology) + then have "t0_space (X i)" + using False R by blast + then obtain U where "closedin (X i) U" "(x i \ U \ y i \ U)" + by (metis t0_space PiE_mem \i \ I\ \x i \ y i\ topspace_product_topology x y) + with \i \ I\ x y show "\U. closedin (product_topology X I) U \ (x \ U) = (y \ U)" + by (rule_tac x="PiE I (\j. if j = i then U else topspace(X j))" in exI) + (simp add: closedin_product_topology PiE_iff) + qed + qed +qed + + +subsection \Kolmogorov quotients\ + +definition Kolmogorov_quotient + where "Kolmogorov_quotient X \ \x. @y. \U. openin X U \ (y \ U \ x \ U)" + +lemma Kolmogorov_quotient_in_open: + "openin X U \ (Kolmogorov_quotient X x \ U \ x \ U)" + by (smt (verit, ccfv_SIG) Kolmogorov_quotient_def someI_ex) + +lemma Kolmogorov_quotient_in_topspace: + "Kolmogorov_quotient X x \ topspace X \ x \ topspace X" + by (simp add: Kolmogorov_quotient_in_open) + +lemma Kolmogorov_quotient_in_closed: + "closedin X C \ (Kolmogorov_quotient X x \ C \ x \ C)" + unfolding closedin_def + by (meson DiffD2 DiffI Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace in_mono) + +lemma continuous_map_Kolmogorov_quotient: + "continuous_map X X (Kolmogorov_quotient X)" + using Kolmogorov_quotient_in_open openin_subopen openin_subset + by (fastforce simp: continuous_map_def Kolmogorov_quotient_in_topspace) + +lemma open_map_Kolmogorov_quotient_explicit: + "openin X U \ Kolmogorov_quotient X ` U = Kolmogorov_quotient X ` topspace X \ U" + using Kolmogorov_quotient_in_open openin_subset by fastforce + + +lemma open_map_Kolmogorov_quotient_gen: + "open_map (subtopology X S) (subtopology X (image (Kolmogorov_quotient X) S)) (Kolmogorov_quotient X)" +proof (clarsimp simp: open_map_def openin_subtopology_alt image_iff) + fix U + assume "openin X U" + then have "Kolmogorov_quotient X ` (S \ U) = Kolmogorov_quotient X ` S \ U" + using Kolmogorov_quotient_in_open [of X U] by auto + then show "\V. openin X V \ Kolmogorov_quotient X ` (S \ U) = Kolmogorov_quotient X ` S \ V" + using \openin X U\ by blast +qed + +lemma open_map_Kolmogorov_quotient: + "open_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) + (Kolmogorov_quotient X)" + by (metis open_map_Kolmogorov_quotient_gen subtopology_topspace) + +lemma closed_map_Kolmogorov_quotient_explicit: + "closedin X U \ Kolmogorov_quotient X ` U = Kolmogorov_quotient X ` topspace X \ U" + using closedin_subset by (fastforce simp: Kolmogorov_quotient_in_closed) + +lemma closed_map_Kolmogorov_quotient_gen: + "closed_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) + (Kolmogorov_quotient X)" + using Kolmogorov_quotient_in_closed by (force simp: closed_map_def closedin_subtopology_alt image_iff) + +lemma closed_map_Kolmogorov_quotient: + "closed_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) + (Kolmogorov_quotient X)" + by (metis closed_map_Kolmogorov_quotient_gen subtopology_topspace) + +lemma quotient_map_Kolmogorov_quotient_gen: + "quotient_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" +proof (intro continuous_open_imp_quotient_map) + show "continuous_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" + by (simp add: continuous_map_Kolmogorov_quotient continuous_map_from_subtopology continuous_map_in_subtopology image_mono) + show "open_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" + using open_map_Kolmogorov_quotient_gen by blast + show "Kolmogorov_quotient X ` topspace (subtopology X S) = topspace (subtopology X (Kolmogorov_quotient X ` S))" + by (force simp: Kolmogorov_quotient_in_open) +qed + +lemma quotient_map_Kolmogorov_quotient: + "quotient_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X)" + by (metis quotient_map_Kolmogorov_quotient_gen subtopology_topspace) + +lemma Kolmogorov_quotient_eq: + "Kolmogorov_quotient X x = Kolmogorov_quotient X y \ + (\U. openin X U \ (x \ U \ y \ U))" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + by (metis Kolmogorov_quotient_in_open) +next + assume ?rhs then show ?lhs + by (simp add: Kolmogorov_quotient_def) +qed + +lemma Kolmogorov_quotient_eq_alt: + "Kolmogorov_quotient X x = Kolmogorov_quotient X y \ + (\U. closedin X U \ (x \ U \ y \ U))" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + by (metis Kolmogorov_quotient_in_closed) +next + assume ?rhs then show ?lhs + by (smt (verit) Diff_iff Kolmogorov_quotient_eq closedin_topspace in_mono openin_closedin_eq) +qed + +lemma Kolmogorov_quotient_continuous_map: + assumes "continuous_map X Y f" "t0_space Y" and x: "x \ topspace X" + shows "f (Kolmogorov_quotient X x) = f x" + using assms unfolding continuous_map_def t0_space_def + by (smt (verit, ccfv_SIG) Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace x mem_Collect_eq) + +lemma t0_space_Kolmogorov_quotient: + "t0_space (subtopology X (Kolmogorov_quotient X ` topspace X))" + apply (clarsimp simp: t0_space_def ) + by (smt (verit, best) Kolmogorov_quotient_eq imageE image_eqI open_map_Kolmogorov_quotient open_map_def) + +lemma Kolmogorov_quotient_id: + "t0_space X \ x \ topspace X \ Kolmogorov_quotient X x = x" + by (metis Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace t0_space_def) + +lemma Kolmogorov_quotient_idemp: + "Kolmogorov_quotient X (Kolmogorov_quotient X x) = Kolmogorov_quotient X x" + by (simp add: Kolmogorov_quotient_eq Kolmogorov_quotient_in_open) + +lemma retraction_maps_Kolmogorov_quotient: + "retraction_maps X + (subtopology X (Kolmogorov_quotient X ` topspace X)) + (Kolmogorov_quotient X) id" + unfolding retraction_maps_def continuous_map_in_subtopology + using Kolmogorov_quotient_idemp continuous_map_Kolmogorov_quotient by force + +lemma retraction_map_Kolmogorov_quotient: + "retraction_map X + (subtopology X (Kolmogorov_quotient X ` topspace X)) + (Kolmogorov_quotient X)" + using retraction_map_def retraction_maps_Kolmogorov_quotient by blast + +lemma retract_of_space_Kolmogorov_quotient_image: + "Kolmogorov_quotient X ` topspace X retract_of_space X" +proof - + have "continuous_map X X (Kolmogorov_quotient X)" + by (simp add: continuous_map_Kolmogorov_quotient) + then have "Kolmogorov_quotient X ` topspace X \ topspace X" + by (simp add: continuous_map_image_subset_topspace) + then show ?thesis + by (meson retract_of_space_retraction_maps retraction_maps_Kolmogorov_quotient) +qed + +lemma Kolmogorov_quotient_lift_exists: + assumes "S \ topspace X" "t0_space Y" and f: "continuous_map (subtopology X S) Y f" + obtains g where "continuous_map (subtopology X (image (Kolmogorov_quotient X) S)) Y g" + "\x. x \ S \ g(Kolmogorov_quotient X x) = f x" +proof - + have "\x y. \x \ S; y \ S; Kolmogorov_quotient X x = Kolmogorov_quotient X y\ + \ f x = f y" + using assms + apply (simp add: Kolmogorov_quotient_eq t0_space_def continuous_map_def Int_absorb1 openin_subtopology) + by (smt (verit, del_insts) Int_iff mem_Collect_eq) + then obtain g where g: "continuous_map (subtopology X (Kolmogorov_quotient X ` S)) Y g" + "g ` (topspace X \ Kolmogorov_quotient X ` S) = f ` S" + "\x. x \ S \ g (Kolmogorov_quotient X x) = f x" + using quotient_map_lift_exists [OF quotient_map_Kolmogorov_quotient_gen [of X S] f] + by (metis assms(1) topspace_subtopology topspace_subtopology_subset) + show ?thesis + proof qed (use g in auto) +qed + +subsection\Closed diagonals and graphs\ + +lemma Hausdorff_space_closedin_diagonal: + "Hausdorff_space X \ + closedin (prod_topology X X) ((\x. (x,x)) ` topspace X)" +proof - + have \
: "((\x. (x, x)) ` topspace X) \ topspace X \ topspace X" + by auto + show ?thesis + apply (simp add: closedin_def openin_prod_topology_alt Hausdorff_space_def disjnt_iff \
) + apply (intro all_cong1 imp_cong ex_cong1 conj_cong refl) + by (force dest!: openin_subset)+ +qed + +lemma closed_map_diag_eq: + "closed_map X (prod_topology X X) (\x. (x,x)) \ Hausdorff_space X" +proof - + have "section_map X (prod_topology X X) (\x. (x, x))" + unfolding section_map_def retraction_maps_def + by (smt (verit) continuous_map_fst continuous_map_of_fst continuous_map_on_empty continuous_map_pairwise fst_conv fst_diag_fst snd_diag_fst) + then have "embedding_map X (prod_topology X X) (\x. (x, x))" + by (rule section_imp_embedding_map) + then show ?thesis + using Hausdorff_space_closedin_diagonal embedding_imp_closed_map_eq by blast +qed + +lemma closedin_continuous_maps_eq: + assumes "Hausdorff_space Y" and f: "continuous_map X Y f" and g: "continuous_map X Y g" + shows "closedin X {x \ topspace X. f x = g x}" +proof - + have \
:"{x \ topspace X. f x = g x} = {x \ topspace X. (f x,g x) \ ((\y.(y,y)) ` topspace Y)}" + using f continuous_map_image_subset_topspace by fastforce + show ?thesis + unfolding \
+ proof (intro closedin_continuous_map_preimage) + show "continuous_map X (prod_topology Y Y) (\x. (f x, g x))" + by (simp add: continuous_map_pairedI f g) + show "closedin (prod_topology Y Y) ((\y. (y, y)) ` topspace Y)" + using Hausdorff_space_closedin_diagonal assms by blast + qed +qed + +lemma retract_of_space_imp_closedin: + assumes "Hausdorff_space X" and S: "S retract_of_space X" + shows "closedin X S" +proof - + obtain r where r: "continuous_map X (subtopology X S) r" "\x\S. r x = x" + using assms by (meson retract_of_space_def) + then have \
: "S = {x \ topspace X. r x = x}" + using S retract_of_space_imp_subset by (force simp: continuous_map_def) + show ?thesis + unfolding \
+ using r continuous_map_into_fulltopology assms + by (force intro: closedin_continuous_maps_eq) +qed + +lemma homeomorphic_maps_graph: + "homeomorphic_maps X (subtopology (prod_topology X Y) ((\x. (x, f x)) ` (topspace X))) + (\x. (x, f x)) fst \ continuous_map X Y f" + (is "?lhs=?rhs") +proof + assume ?lhs + then + have h: "homeomorphic_map X (subtopology (prod_topology X Y) ((\x. (x, f x)) ` topspace X)) (\x. (x, f x))" + by (auto simp: homeomorphic_maps_map) + have "f = snd \ (\x. (x, f x))" + by force + then show ?rhs + by (metis (no_types, lifting) h continuous_map_in_subtopology continuous_map_snd_of homeomorphic_eq_everything_map) +next + assume ?rhs + then show ?lhs + unfolding homeomorphic_maps_def + by (smt (verit, ccfv_threshold) continuous_map_eq continuous_map_subtopology_fst embedding_map_def embedding_map_graph homeomorphic_eq_everything_map image_cong image_iff prod.collapse prod.inject) +qed + + +subsection \ KC spaces, those where all compact sets are closed.\ + +definition kc_space + where "kc_space X \ \S. compactin X S \ closedin X S" + +lemma kc_space_euclidean: "kc_space (euclidean :: 'a::metric_space topology)" + by (simp add: compact_imp_closed kc_space_def) + +lemma kc_space_expansive: + "\kc_space X; topspace Y = topspace X; \U. openin X U \ openin Y U\ + \ kc_space Y" + by (meson compactin_contractive kc_space_def topology_finer_closedin) + +lemma compactin_imp_closedin_gen: + "\kc_space X; compactin X S\ \ closedin X S" + using kc_space_def by blast + +lemma Hausdorff_imp_kc_space: "Hausdorff_space X \ kc_space X" + by (simp add: compactin_imp_closedin kc_space_def) + +lemma kc_imp_t1_space: "kc_space X \ t1_space X" + by (simp add: finite_imp_compactin kc_space_def t1_space_closedin_finite) + +lemma kc_space_subtopology: + "kc_space X \ kc_space(subtopology X S)" + by (metis closedin_Int_closure_of closure_of_eq compactin_subtopology inf.absorb2 kc_space_def) + +lemma kc_space_discrete_topology: + "kc_space(discrete_topology U)" + using Hausdorff_space_discrete_topology compactin_imp_closedin kc_space_def by blast + +lemma kc_space_continuous_injective_map_preimage: + assumes "kc_space Y" "continuous_map X Y f" and injf: "inj_on f (topspace X)" + shows "kc_space X" + unfolding kc_space_def +proof (intro strip) + fix S + assume S: "compactin X S" + have "S = {x \ topspace X. f x \ f ` S}" + using S compactin_subset_topspace inj_onD [OF injf] by blast + with assms S show "closedin X S" + by (metis (no_types, lifting) Collect_cong closedin_continuous_map_preimage compactin_imp_closedin_gen image_compactin) +qed + +lemma kc_space_retraction_map_image: + assumes "retraction_map X Y r" "kc_space X" + shows "kc_space Y" +proof - + obtain s where s: "continuous_map X Y r" "continuous_map Y X s" "\x. x \ topspace Y \ r (s x) = x" + using assms by (force simp: retraction_map_def retraction_maps_def) + then have inj: "inj_on s (topspace Y)" + by (metis inj_on_def) + show ?thesis + unfolding kc_space_def + proof (intro strip) + fix S + assume S: "compactin Y S" + have "S = {x \ topspace Y. s x \ s ` S}" + using S compactin_subset_topspace inj_onD [OF inj] by blast + with assms S show "closedin Y S" + by (meson compactin_imp_closedin_gen inj kc_space_continuous_injective_map_preimage s(2)) + qed +qed + +lemma homeomorphic_kc_space: + "X homeomorphic_space Y \ kc_space X \ kc_space Y" + by (meson homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym kc_space_continuous_injective_map_preimage) + +lemma compact_kc_eq_maximal_compact_space: + assumes "compact_space X" + shows "kc_space X \ + (\Y. topspace Y = topspace X \ (\S. openin X S \ openin Y S) \ compact_space Y \ Y = X)" (is "?lhs=?rhs") +proof + assume ?lhs + then show ?rhs + by (metis closedin_compact_space compactin_contractive kc_space_def topology_eq topology_finer_closedin) +next + assume R: ?rhs + show ?lhs + unfolding kc_space_def + proof (intro strip) + fix S + assume S: "compactin X S" + define Y where + "Y \ topology (arbitrary union_of (finite intersection_of (\A. A = topspace X - S \ openin X A) + relative_to (topspace X)))" + have "topspace Y = topspace X" + by (auto simp: Y_def) + have "openin X T \ openin Y T" for T + by (simp add: Y_def arbitrary_union_of_inc finite_intersection_of_inc openin_subbase openin_subset relative_to_subset) + have "compact_space Y" + proof (rule Alexander_subbase_alt) + show "\\'. finite \' \ \' \ \ \ topspace X \ \ \'" + if \: "\ \ insert (topspace X - S) (Collect (openin X))" and sub: "topspace X \ \\" for \ + proof - + consider "\ \ Collect (openin X)" | \ where "\ = insert (topspace X - S) \" "\ \ Collect (openin X)" + using \ by (metis insert_Diff subset_insert_iff) + then show ?thesis + proof cases + case 1 + then show ?thesis + by (metis assms compact_space_alt mem_Collect_eq subsetD that(2)) + next + case 2 + then have "S \ \\" + using S sub compactin_subset_topspace by blast + with 2 obtain \ where "finite \ \ \ \ \ \ S \ \\" + using S unfolding compactin_def by (metis Ball_Collect) + with 2 show ?thesis + by (rule_tac x="insert (topspace X - S) \" in exI) blast + qed + qed + qed (auto simp: Y_def) + have "Y = X" + using R \\S. openin X S \ openin Y S\ \compact_space Y\ \topspace Y = topspace X\ by blast + moreover have "openin Y (topspace X - S)" + by (simp add: Y_def arbitrary_union_of_inc finite_intersection_of_inc openin_subbase relative_to_subset) + ultimately show "closedin X S" + unfolding closedin_def using S compactin_subset_topspace by blast + qed +qed + +lemma continuous_imp_closed_map_gen: + "\compact_space X; kc_space Y; continuous_map X Y f\ \ closed_map X Y f" + by (meson closed_map_def closedin_compact_space compactin_imp_closedin_gen image_compactin) + +lemma kc_space_compact_subtopologies: + "kc_space X \ (\K. compactin X K \ kc_space(subtopology X K))" (is "?lhs=?rhs") +proof + assume ?lhs + then show ?rhs + by (auto simp: kc_space_def closedin_closed_subtopology compactin_subtopology) +next + assume R: ?rhs + show ?lhs + unfolding kc_space_def + proof (intro strip) + fix K + assume K: "compactin X K" + then have "K \ topspace X" + by (simp add: compactin_subset_topspace) + moreover have "X closure_of K \ K" + proof + fix x + assume x: "x \ X closure_of K" + have "kc_space (subtopology X K)" + by (simp add: R \compactin X K\) + have "compactin X (insert x K)" + by (metis K x compactin_Un compactin_sing in_closure_of insert_is_Un) + then show "x \ K" + by (metis R x K Int_insert_left_if1 closedin_Int_closure_of compact_imp_compactin_subtopology + insertCI kc_space_def subset_insertI) + qed + ultimately show "closedin X K" + using closure_of_subset_eq by blast + qed +qed + +lemma kc_space_compact_prod_topology: + assumes "compact_space X" + shows "kc_space(prod_topology X X) \ Hausdorff_space X" (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + unfolding closed_map_diag_eq [symmetric] + proof (intro continuous_imp_closed_map_gen) + show "continuous_map X (prod_topology X X) (\x. (x, x))" + by (intro continuous_intros) + qed (use L assms in auto) +next + assume ?rhs then show ?lhs + by (simp add: Hausdorff_imp_kc_space Hausdorff_space_prod_topology) +qed + +lemma kc_space_prod_topology: + "kc_space(prod_topology X X) \ (\K. compactin X K \ Hausdorff_space(subtopology X K))" (is "?lhs=?rhs") +proof + assume ?lhs + then show ?rhs + by (metis compactin_subspace kc_space_compact_prod_topology kc_space_subtopology subtopology_Times) +next + assume R: ?rhs + have "kc_space (subtopology (prod_topology X X) L)" if "compactin (prod_topology X X) L" for L + proof - + define K where "K \ fst ` L \ snd ` L" + have "L \ K \ K" + by (force simp: K_def) + have "compactin X K" + by (metis K_def compactin_Un continuous_map_fst continuous_map_snd image_compactin that) + then have "Hausdorff_space (subtopology X K)" + by (simp add: R) + then have "kc_space (prod_topology (subtopology X K) (subtopology X K))" + by (simp add: \compactin X K\ compact_space_subtopology kc_space_compact_prod_topology) + then have "kc_space (subtopology (prod_topology (subtopology X K) (subtopology X K)) L)" + using kc_space_subtopology by blast + then show ?thesis + using \L \ K \ K\ subtopology_Times subtopology_subtopology + by (metis (no_types, lifting) Sigma_cong inf.absorb_iff2) + qed + then show ?lhs + using kc_space_compact_subtopologies by blast +qed + +lemma kc_space_prod_topology_alt: + "kc_space(prod_topology X X) \ + kc_space X \ + (\K. compactin X K \ Hausdorff_space(subtopology X K))" + using Hausdorff_imp_kc_space kc_space_compact_subtopologies kc_space_prod_topology by blast + +proposition kc_space_prod_topology_left: + assumes X: "kc_space X" and Y: "Hausdorff_space Y" + shows "kc_space (prod_topology X Y)" + unfolding kc_space_def +proof (intro strip) + fix K + assume K: "compactin (prod_topology X Y) K" + then have "K \ topspace X \ topspace Y" + using compactin_subset_topspace topspace_prod_topology by blast + moreover have "\T. openin (prod_topology X Y) T \ (a,b) \ T \ T \ (topspace X \ topspace Y) - K" + if ab: "(a,b) \ (topspace X \ topspace Y) - K" for a b + proof - + have "compactin Y {b}" + using that by force + moreover + have "compactin Y {y \ topspace Y. (a,y) \ K}" + proof - + have "compactin (prod_topology X Y) (K \ {a} \ topspace Y)" + using that compact_Int_closedin [OF K] + by (simp add: X closedin_prod_Times_iff compactin_imp_closedin_gen) + moreover have "subtopology (prod_topology X Y) (K \ {a} \ topspace Y) homeomorphic_space + subtopology Y {y \ topspace Y. (a, y) \ K}" + unfolding homeomorphic_space_def homeomorphic_maps_def + using that + apply (rule_tac x="snd" in exI) + apply (rule_tac x="Pair a" in exI) + by (force simp: continuous_map_in_subtopology continuous_map_from_subtopology continuous_map_subtopology_snd continuous_map_paired) + ultimately show ?thesis + by (simp add: compactin_subspace homeomorphic_compact_space) + qed + moreover have "disjnt {b} {y \ topspace Y. (a,y) \ K}" + using ab by force + ultimately obtain V U where VU: "openin Y V" "openin Y U" "{b} \ V" "{y \ topspace Y. (a,y) \ K} \ U" "disjnt V U" + using Hausdorff_space_compact_separation [OF Y] by blast + define V' where "V' \ topspace Y - U" + have W: "closedin Y V'" "{y \ topspace Y. (a,y) \ K} \ topspace Y - V'" "disjnt V (topspace Y - V')" + using VU by (auto simp: V'_def disjnt_iff) + with VU obtain "V \ topspace Y" "V' \ topspace Y" + by (meson closedin_subset openin_closedin_eq) + then obtain "b \ V" "disjnt {y \ topspace Y. (a,y) \ K} V'" "V \ V'" + using VU unfolding disjnt_iff V'_def by force + define C where "C \ image fst (K \ {z \ topspace(prod_topology X Y). snd z \ V'})" + have "closedin (prod_topology X Y) {z \ topspace (prod_topology X Y). snd z \ V'}" + using closedin_continuous_map_preimage \closedin Y V'\ continuous_map_snd by blast + then have "compactin X C" + unfolding C_def by (meson K compact_Int_closedin continuous_map_fst image_compactin) + then have "closedin X C" + using assms by (auto simp: kc_space_def) + show ?thesis + proof (intro exI conjI) + show "openin (prod_topology X Y) ((topspace X - C) \ V)" + by (simp add: VU \closedin X C\ openin_diff openin_prod_Times_iff) + have "a \ C" + using VU by (auto simp: C_def V'_def) + then show "(a, b) \ (topspace X - C) \ V" + using \a \ C\ \b \ V\ that by blast + show "(topspace X - C) \ V \ topspace X \ topspace Y - K" + using \V \ V'\ \V \ topspace Y\ + apply (simp add: C_def ) + by (smt (verit, ccfv_threshold) DiffE DiffI IntI SigmaE SigmaI image_eqI mem_Collect_eq prod.sel(1) snd_conv subset_iff) + qed + qed + ultimately show "closedin (prod_topology X Y) K" + by (metis surj_pair closedin_def openin_subopen topspace_prod_topology) +qed + +lemma kc_space_prod_topology_right: + "\Hausdorff_space X; kc_space Y\ \ kc_space (prod_topology X Y)" + using kc_space_prod_topology_left homeomorphic_kc_space homeomorphic_space_prod_topology_swap by blast + + +subsection \Regular spaces\ + +text \Regular spaces are *not* a priori assumed to be Hausdorff or $T_1$\ + +definition regular_space + where "regular_space X \ + \C a. closedin X C \ a \ topspace X - C + \ (\U V. openin X U \ openin X V \ a \ U \ C \ V \ disjnt U V)" + +lemma homeomorphic_regular_space_aux: + assumes hom: "X homeomorphic_space Y" and X: "regular_space X" + shows "regular_space Y" +proof - + obtain f g where hmf: "homeomorphic_map X Y f" and hmg: "homeomorphic_map Y X g" + and fg: "(\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" + using assms X homeomorphic_maps_map homeomorphic_space_def by fastforce + show ?thesis + unfolding regular_space_def + proof clarify + fix C a + assume Y: "closedin Y C" "a \ topspace Y" and "a \ C" + then obtain "closedin X (g ` C)" "g a \ topspace X" "g a \ g ` C" + using \closedin Y C\ hmg homeomorphic_map_closedness_eq + by (smt (verit, ccfv_SIG) fg homeomorphic_imp_surjective_map image_iff in_mono) + then obtain S T where ST: "openin X S" "openin X T" "g a \ S" "g`C \ T" "disjnt S T" + using X unfolding regular_space_def by (metis DiffI) + then have "openin Y (f`S)" "openin Y (f`T)" + by (meson hmf homeomorphic_map_openness_eq)+ + moreover have "a \ f`S \ C \ f`T" + using ST by (smt (verit, best) Y closedin_subset fg image_eqI subset_iff) + moreover have "disjnt (f`S) (f`T)" + using ST by (smt (verit, ccfv_SIG) disjnt_iff fg image_iff openin_subset subsetD) + ultimately show "\U V. openin Y U \ openin Y V \ a \ U \ C \ V \ disjnt U V" + by metis + qed +qed + +lemma homeomorphic_regular_space: + "X homeomorphic_space Y + \ (regular_space X \ regular_space Y)" + by (meson homeomorphic_regular_space_aux homeomorphic_space_sym) + +lemma regular_space: + "regular_space X \ + (\C a. closedin X C \ a \ topspace X - C + \ (\U. openin X U \ a \ U \ disjnt C (X closure_of U)))" + unfolding regular_space_def +proof (intro all_cong1 imp_cong refl ex_cong1) + fix C a U + assume C: "closedin X C \ a \ topspace X - C" + show "(\V. openin X U \ openin X V \ a \ U \ C \ V \ disjnt U V) + \ (openin X U \ a \ U \ disjnt C (X closure_of U))" (is "?lhs=?rhs") + proof + assume ?lhs + then show ?rhs + by (smt (verit, best) disjnt_iff in_closure_of subsetD) + next + assume R: ?rhs + then have "disjnt U (topspace X - X closure_of U)" + by (meson DiffD2 closure_of_subset disjnt_iff openin_subset subsetD) + moreover have "C \ topspace X - X closure_of U" + by (meson C DiffI R closedin_subset disjnt_iff subset_eq) + ultimately show ?lhs + using R by (rule_tac x="topspace X - X closure_of U" in exI) auto + qed +qed + +lemma neighbourhood_base_of_closedin: + "neighbourhood_base_of (closedin X) X \ regular_space X" (is "?lhs=?rhs") +proof - + have "?lhs \ (\W x. openin X W \ x \ W \ + (\U. openin X U \ (\V. closedin X V \ x \ U \ U \ V \ V \ W)))" + by (simp add: neighbourhood_base_of) + also have "\ \ (\W x. closedin X W \ x \ topspace X - W \ + (\U. openin X U \ (\V. closedin X V \ x \ U \ U \ V \ V \ topspace X - W)))" + by (smt (verit) Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq) + also have "\ \ ?rhs" + proof - + have \
: "(\V. closedin X V \ x \ U \ U \ V \ V \ topspace X - W) + \ (\V. openin X V \ x \ U \ W \ V \ disjnt U V)" (is "?lhs=?rhs") + if "openin X U" "closedin X W" "x \ topspace X" "x \ W" for W U x + proof + assume ?lhs with \closedin X W\ show ?rhs + unfolding closedin_def by (smt (verit) Diff_mono disjnt_Diff1 double_diff subset_eq) + next + assume ?rhs with \openin X U\ show ?lhs + unfolding openin_closedin_eq disjnt_def + by (smt (verit) Diff_Diff_Int Diff_disjoint Diff_eq_empty_iff Int_Diff inf.orderE) + qed + show ?thesis + unfolding regular_space_def + by (intro all_cong1 ex_cong1 imp_cong refl) (metis \
DiffE) + qed + finally show ?thesis . +qed + +lemma regular_space_discrete_topology: + "regular_space (discrete_topology S)" + using neighbourhood_base_of_closedin neighbourhood_base_of_discrete_topology by fastforce + +lemma regular_space_subtopology: + "regular_space X \ regular_space (subtopology X S)" + unfolding regular_space_def openin_subtopology_alt closedin_subtopology_alt disjnt_iff + by clarsimp (smt (verit, best) inf.orderE inf_le1 le_inf_iff) + + +lemma regular_space_retraction_map_image: + "\retraction_map X Y r; regular_space X\ \ regular_space Y" + using hereditary_imp_retractive_property homeomorphic_regular_space regular_space_subtopology by blast + +lemma regular_t0_imp_Hausdorff_space: + "\regular_space X; t0_space X\ \ Hausdorff_space X" + apply (clarsimp simp: regular_space_def t0_space Hausdorff_space_def) + by (metis disjnt_sym subsetD) + +lemma regular_t0_eq_Hausdorff_space: + "regular_space X \ (t0_space X \ Hausdorff_space X)" + using Hausdorff_imp_t0_space regular_t0_imp_Hausdorff_space by blast + +lemma regular_t1_imp_Hausdorff_space: + "\regular_space X; t1_space X\ \ Hausdorff_space X" + by (simp add: regular_t0_imp_Hausdorff_space t1_imp_t0_space) + +lemma regular_t1_eq_Hausdorff_space: + "regular_space X \ t1_space X \ Hausdorff_space X" + using regular_t0_imp_Hausdorff_space t1_imp_t0_space t1_or_Hausdorff_space by blast + +lemma compact_Hausdorff_imp_regular_space: + assumes "compact_space X" "Hausdorff_space X" + shows "regular_space X" + unfolding regular_space_def +proof clarify + fix S a + assume "closedin X S" and "a \ topspace X" and "a \ S" + then show "\U V. openin X U \ openin X V \ a \ U \ S \ V \ disjnt U V" + using assms unfolding Hausdorff_space_compact_sets + by (metis closedin_compact_space compactin_sing disjnt_empty1 insert_subset disjnt_insert1) +qed + +lemma regular_space_topspace_empty: "topspace X = {} \ regular_space X" + by (simp add: Hausdorff_space_topspace_empty compact_Hausdorff_imp_regular_space compact_space_topspace_empty) + +lemma neighbourhood_base_of_closed_Hausdorff_space: + "regular_space X \ Hausdorff_space X \ + neighbourhood_base_of (\C. closedin X C \ Hausdorff_space(subtopology X C)) X" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + by (simp add: Hausdorff_space_subtopology neighbourhood_base_of_closedin) +next + assume ?rhs then show ?lhs + by (metis (mono_tags, lifting) Hausdorff_space_closed_neighbourhood neighbourhood_base_of neighbourhood_base_of_closedin openin_topspace) +qed + +lemma locally_compact_imp_kc_eq_Hausdorff_space: + "neighbourhood_base_of (compactin X) X \ kc_space X \ Hausdorff_space X" + by (metis Hausdorff_imp_kc_space kc_imp_t1_space kc_space_def neighbourhood_base_of_closedin neighbourhood_base_of_mono regular_t1_imp_Hausdorff_space) + +lemma regular_space_compact_closed_separation: + assumes X: "regular_space X" + and S: "compactin X S" + and T: "closedin X T" + and "disjnt S T" + shows "\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V" +proof (cases "S={}") + case True + then show ?thesis + by (meson T closedin_def disjnt_empty1 empty_subsetI openin_empty openin_topspace) +next + case False + then have "\U V. x \ S \ openin X U \ openin X V \ x \ U \ T \ V \ disjnt U V" for x + using assms unfolding regular_space_def + by (smt (verit) Diff_iff compactin_subset_topspace disjnt_iff subsetD) + then obtain U V where UV: "\x. x \ S \ openin X (U x) \ openin X (V x) \ x \ (U x) \ T \ (V x) \ disjnt (U x) (V x)" + by metis + then obtain \ where "finite \" "\ \ U ` S" "S \ \ \" + using S unfolding compactin_def by (smt (verit) UN_iff image_iff subsetI) + then obtain K where "finite K" "K \ S" and K: "S \ \(U ` K)" + by (metis finite_subset_image) + show ?thesis + proof (intro exI conjI) + show "openin X (\(U ` K))" + using \K \ S\ UV by blast + show "openin X (\(V ` K))" + using False K UV \K \ S\ \finite K\ by blast + show "S \ \(U ` K)" + by (simp add: K) + show "T \ \(V ` K)" + using UV \K \ S\ by blast + show "disjnt (\(U ` K)) (\(V ` K))" + by (smt (verit) Inter_iff UN_E UV \K \ S\ disjnt_iff image_eqI subset_iff) + qed +qed + +lemma regular_space_compact_closed_sets: + "regular_space X \ + (\S T. compactin X S \ closedin X T \ disjnt S T + \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V))" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + using regular_space_compact_closed_separation by fastforce +next + assume R: ?rhs + show ?lhs + unfolding regular_space + proof clarify + fix S x + assume "closedin X S" and "x \ topspace X" and "x \ S" + then obtain U V where "openin X U \ openin X V \ {x} \ U \ S \ V \ disjnt U V" + by (metis R compactin_sing disjnt_empty1 disjnt_insert1) + then show "\U. openin X U \ x \ U \ disjnt S (X closure_of U)" + by (smt (verit, best) disjnt_iff in_closure_of insert_subset subsetD) + qed +qed + + +lemma regular_space_prod_topology: + "regular_space (prod_topology X Y) \ + topspace X = {} \ topspace Y = {} \ regular_space X \ regular_space Y" (is "?lhs=?rhs") +proof + assume ?lhs + then show ?rhs + by (metis regular_space_retraction_map_image retraction_map_fst retraction_map_snd) +next + assume R: ?rhs + show ?lhs + proof (cases "topspace X = {} \ topspace Y = {}") + case True + then show ?thesis + by (simp add: regular_space_topspace_empty) + next + case False + then have "regular_space X" "regular_space Y" + using R by auto + show ?thesis + unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of + proof clarify + fix W x y + assume W: "openin (prod_topology X Y) W" and "(x,y) \ W" + then obtain U V where U: "openin X U" "x \ U" and V: "openin Y V" "y \ V" + and "U \ V \ W" + by (metis openin_prod_topology_alt) + obtain D1 C1 where 1: "openin X D1" "closedin X C1" "x \ D1" "D1 \ C1" "C1 \ U" + by (metis \regular_space X\ U neighbourhood_base_of neighbourhood_base_of_closedin) + obtain D2 C2 where 2: "openin Y D2" "closedin Y C2" "y \ D2" "D2 \ C2" "C2 \ V" + by (metis \regular_space Y\ V neighbourhood_base_of neighbourhood_base_of_closedin) + show "\U V. openin (prod_topology X Y) U \ closedin (prod_topology X Y) V \ + (x,y) \ U \ U \ V \ V \ W" + proof (intro conjI exI) + show "openin (prod_topology X Y) (D1 \ D2)" + by (simp add: 1 2 openin_prod_Times_iff) + show "closedin (prod_topology X Y) (C1 \ C2)" + by (simp add: 1 2 closedin_prod_Times_iff) + qed (use 1 2 \U \ V \ W\ in auto) + qed + qed +qed + +lemma regular_space_product_topology: + "regular_space (product_topology X I) \ + topspace (product_topology X I) = {} \ (\i \ I. regular_space (X i))" (is "?lhs=?rhs") +proof + assume ?lhs + then show ?rhs + by (meson regular_space_retraction_map_image retraction_map_product_projection) +next + assume R: ?rhs + show ?lhs + proof (cases "topspace(product_topology X I) = {}") + case True + then show ?thesis + by (simp add: regular_space_topspace_empty) + next + case False + then obtain x where x: "x \ topspace (product_topology X I)" + by blast + define \ where "\ \ {Pi\<^sub>E I U |U. finite {i \ I. U i \ topspace (X i)} + \ (\i\I. openin (X i) (U i))}" + have oo: "openin (product_topology X I) = arbitrary union_of (\W. W \ \)" + by (simp add: \_def openin_product_topology product_topology_base_alt) + have "\U V. openin (product_topology X I) U \ closedin (product_topology X I) V \ x \ U \ U \ V \ V \ Pi\<^sub>E I W" + if fin: "finite {i \ I. W i \ topspace (X i)}" + and opeW: "\k. k \ I \ openin (X k) (W k)" and x: "x \ PiE I W" for W x + proof - + have "\i. i \ I \ \U V. openin (X i) U \ closedin (X i) V \ x i \ U \ U \ V \ V \ W i" + by (metis False PiE_iff R neighbourhood_base_of neighbourhood_base_of_closedin opeW x) + then obtain U C where UC: + "\i. i \ I \ openin (X i) (U i) \ closedin (X i) (C i) \ x i \ U i \ U i \ C i \ C i \ W i" + by metis + define PI where "PI \ \V. PiE I (\i. if W i = topspace(X i) then topspace(X i) else V i)" + show ?thesis + proof (intro conjI exI) + have "\i\I. W i \ topspace (X i) \ openin (X i) (U i)" + using UC by force + with fin show "openin (product_topology X I) (PI U)" + by (simp add: Collect_mono_iff PI_def openin_PiE_gen rev_finite_subset) + show "closedin (product_topology X I) (PI C)" + by (simp add: UC closedin_product_topology PI_def) + show "x \ PI U" + using UC x by (fastforce simp: PI_def) + show "PI U \ PI C" + by (smt (verit) UC Orderings.order_eq_iff PiE_mono PI_def) + show "PI C \ Pi\<^sub>E I W" + by (simp add: UC PI_def subset_PiE) + qed + qed + then have "neighbourhood_base_of (closedin (product_topology X I)) (product_topology X I)" + unfolding neighbourhood_base_of_topology_base [OF oo] by (force simp: \_def) + then show ?thesis + by (simp add: neighbourhood_base_of_closedin) + qed +qed + +lemma closed_map_paired_gen_aux: + assumes "regular_space Y" and f: "closed_map Z X f" and g: "closed_map Z Y g" + and clo: "\y. y \ topspace X \ closedin Z {x \ topspace Z. f x = y}" + and contg: "continuous_map Z Y g" + shows "closed_map Z (prod_topology X Y) (\x. (f x, g x))" + unfolding closed_map_def +proof (intro strip) + fix C assume "closedin Z C" + then have "C \ topspace Z" + by (simp add: closedin_subset) + have "f ` topspace Z \ topspace X" "g ` topspace Z \ topspace Y" + by (simp_all add: assms closed_map_imp_subset_topspace) + show "closedin (prod_topology X Y) ((\x. (f x, g x)) ` C)" + unfolding closedin_def topspace_prod_topology + proof (intro conjI) + have "closedin Y (g ` C)" + using \closedin Z C\ assms(3) closed_map_def by blast + with assms show "(\x. (f x, g x)) ` C \ topspace X \ topspace Y" + using \C \ topspace Z\ \f ` topspace Z \ topspace X\ continuous_map_closedin subsetD by fastforce + have *: "\T. openin (prod_topology X Y) T \ (y1,y2) \ T \ T \ topspace X \ topspace Y - (\x. (f x, g x)) ` C" + if "(y1,y2) \ (\x. (f x, g x)) ` C" and y1: "y1 \ topspace X" and y2: "y2 \ topspace Y" + for y1 y2 + proof - + define A where "A \ topspace Z - (C \ {x \ topspace Z. f x = y1})" + have A: "openin Z A" "{x \ topspace Z. g x = y2} \ A" + using that \closedin Z C\ clo that(2) by (auto simp: A_def) + obtain V0 where "openin Y V0 \ y2 \ V0" and UA: "{x \ topspace Z. g x \ V0} \ A" + using g A y2 unfolding closed_map_fibre_neighbourhood by blast + then obtain V V' where VV: "openin Y V \ closedin Y V' \ y2 \ V \ V \ V'" and "V' \ V0" + by (metis (no_types, lifting) \regular_space Y\ neighbourhood_base_of neighbourhood_base_of_closedin) + with UA have subA: "{x \ topspace Z. g x \ V'} \ A" + by blast + show ?thesis + proof - + define B where "B \ topspace Z - (C \ {x \ topspace Z. g x \ V'})" + have "openin Z B" + using VV \closedin Z C\ contg by (fastforce simp: B_def continuous_map_closedin) + have "{x \ topspace Z. f x = y1} \ B" + using A_def subA by (auto simp: A_def B_def) + then obtain U where "openin X U" "y1 \ U" and subB: "{x \ topspace Z. f x \ U} \ B" + using \openin Z B\ y1 f unfolding closed_map_fibre_neighbourhood by meson + show ?thesis + proof (intro conjI exI) + show "openin (prod_topology X Y) (U \ V)" + by (metis VV \openin X U\ openin_prod_Times_iff) + show "(y1, y2) \ U \ V" + by (simp add: VV \y1 \ U\) + show "U \ V \ topspace X \ topspace Y - (\x. (f x, g x)) ` C" + using VV \C \ topspace Z\ \openin X U\ subB + by (force simp: image_iff B_def subset_iff dest: openin_subset) + qed + qed + qed + then show "openin (prod_topology X Y) (topspace X \ topspace Y - (\x. (f x, g x)) ` C)" + by (smt (verit, ccfv_threshold) Diff_iff SigmaE openin_subopen) + qed +qed + + +lemma closed_map_paired_gen: + assumes f: "closed_map Z X f" and g: "closed_map Z Y g" + and D: "(regular_space X \ continuous_map Z X f \ (\z \ topspace Y. closedin Z {x \ topspace Z. g x = z}) + \ regular_space Y \ continuous_map Z Y g \ (\y \ topspace X. closedin Z {x \ topspace Z. f x = y}))" + (is "?RSX \ ?RSY") + shows "closed_map Z (prod_topology X Y) (\x. (f x, g x))" + using D +proof + assume RSX: ?RSX + have eq: "(\x. (f x, g x)) = (\(x,y). (y,x)) \ (\x. (g x, f x))" + by auto + show ?thesis + unfolding eq + proof (rule closed_map_compose) + show "closed_map Z (prod_topology Y X) (\x. (g x, f x))" + using RSX closed_map_paired_gen_aux f g by fastforce + show "closed_map (prod_topology Y X) (prod_topology X Y) (\(x, y). (y, x))" + using homeomorphic_imp_closed_map homeomorphic_map_swap by blast + qed +qed (blast intro: assms closed_map_paired_gen_aux) + +lemma closed_map_paired: + assumes "closed_map Z X f" and contf: "continuous_map Z X f" + "closed_map Z Y g" and contg: "continuous_map Z Y g" + and D: "t1_space X \ regular_space Y \ regular_space X \ t1_space Y" + shows "closed_map Z (prod_topology X Y) (\x. (f x, g x))" +proof (rule closed_map_paired_gen) + show "regular_space X \ continuous_map Z X f \ (\z\topspace Y. closedin Z {x \ topspace Z. g x = z}) \ regular_space Y \ continuous_map Z Y g \ (\y\topspace X. closedin Z {x \ topspace Z. f x = y})" + using D contf contg + by (smt (verit, del_insts) Collect_cong closedin_continuous_map_preimage t1_space_closedin_singleton singleton_iff) +qed (use assms in auto) + +lemma closed_map_pairwise: + assumes "closed_map Z X (fst \ f)" "continuous_map Z X (fst \ f)" + "closed_map Z Y (snd \ f)" "continuous_map Z Y (snd \ f)" + "t1_space X \ regular_space Y \ regular_space X \ t1_space Y" + shows "closed_map Z (prod_topology X Y) f" +proof - + have "closed_map Z (prod_topology X Y) (\a. ((fst \ f) a, (snd \ f) a))" + using assms closed_map_paired by blast + then show ?thesis + by auto +qed + + +lemma tube_lemma_right: + assumes W: "openin (prod_topology X Y) W" and C: "compactin Y C" + and x: "x \ topspace X" and subW: "{x} \ C \ W" + shows "\U V. openin X U \ openin Y V \ x \ U \ C \ V \ U \ V \ W" +proof (cases "C = {}") + case True + with x show ?thesis by auto +next + case False + have "\U V. openin X U \ openin Y V \ x \ U \ y \ V \ U \ V \ W" + if "y \ C" for y + using W openin_prod_topology_alt subW subsetD that by fastforce + then obtain U V where UV: "\y. y \ C \ openin X (U y) \ openin Y (V y) \ x \ U y \ y \ V y \ U y \ V y \ W" + by metis + then obtain D where D: "finite D" "D \ C" "C \ \ (V ` D)" + using compactinD [OF C, of "V`C"] + by (smt (verit) UN_I finite_subset_image imageE subsetI) + show ?thesis + proof (intro exI conjI) + show "openin X (\ (U ` D))" "openin Y (\ (V ` D))" + using D False UV by blast+ + show "x \ \ (U ` D)" "C \ \ (V ` D)" "\ (U ` D) \ \ (V ` D) \ W" + using D UV by force+ + qed +qed + + +lemma closed_map_fst: + assumes "compact_space Y" + shows "closed_map (prod_topology X Y) X fst" +proof - + have *: "{x \ topspace X \ topspace Y. fst x \ U} = U \ topspace Y" + if "U \ topspace X" for U + using that by force + have **: "\U y. \openin (prod_topology X Y) U; y \ topspace X; + {x \ topspace X \ topspace Y. fst x = y} \ U\ + \ \V. openin X V \ y \ V \ V \ topspace Y \ U" + using tube_lemma_right[of X Y _ "topspace Y"] assms compact_space_def + by force + show ?thesis + unfolding closed_map_fibre_neighbourhood + by (force simp: * openin_subset cong: conj_cong intro: **) +qed + +lemma closed_map_snd: + assumes "compact_space X" + shows "closed_map (prod_topology X Y) Y snd" +proof - + have "snd = fst o prod.swap" + by force + moreover have "closed_map (prod_topology X Y) Y (fst o prod.swap)" + proof (rule closed_map_compose) + show "closed_map (prod_topology X Y) (prod_topology Y X) prod.swap" + by (metis (no_types, lifting) homeomorphic_imp_closed_map homeomorphic_map_eq homeomorphic_map_swap prod.swap_def split_beta) + show "closed_map (prod_topology Y X) Y fst" + by (simp add: closed_map_fst assms) + qed + ultimately show ?thesis + by metis +qed + +lemma closed_map_paired_closed_map_right: + "\closed_map X Y f; regular_space X; + \y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}\ + \ closed_map X (prod_topology X Y) (\x. (x, f x))" + by (rule closed_map_paired_gen [OF closed_map_id, unfolded id_def]) auto + +lemma closed_map_paired_closed_map_left: + assumes "closed_map X Y f" "regular_space X" + "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" + shows "closed_map X (prod_topology Y X) (\x. (f x, x))" +proof - + have eq: "(\x. (f x, x)) = (\(x,y). (y,x)) \ (\x. (x, f x))" + by auto + show ?thesis + unfolding eq + proof (rule closed_map_compose) + show "closed_map X (prod_topology X Y) (\x. (x, f x))" + by (simp add: assms closed_map_paired_closed_map_right) + show "closed_map (prod_topology X Y) (prod_topology Y X) (\(x, y). (y, x))" + using homeomorphic_imp_closed_map homeomorphic_map_swap by blast + qed +qed + +lemma closed_map_imp_closed_graph: + assumes "closed_map X Y f" "regular_space X" + "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" + shows "closedin (prod_topology X Y) ((\x. (x, f x)) ` topspace X)" + using assms closed_map_def closed_map_paired_closed_map_right by blast + +lemma proper_map_paired_closed_map_right: + assumes "closed_map X Y f" "regular_space X" + "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" + shows "proper_map X (prod_topology X Y) (\x. (x, f x))" + by (simp add: assms closed_injective_imp_proper_map inj_on_def closed_map_paired_closed_map_right) + +lemma proper_map_paired_closed_map_left: + assumes "closed_map X Y f" "regular_space X" + "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" + shows "proper_map X (prod_topology Y X) (\x. (f x, x))" + by (simp add: assms closed_injective_imp_proper_map inj_on_def closed_map_paired_closed_map_left) + +proposition regular_space_continuous_proper_map_image: + assumes "regular_space X" and contf: "continuous_map X Y f" and pmapf: "proper_map X Y f" + and fim: "f ` (topspace X) = topspace Y" + shows "regular_space Y" + unfolding regular_space_def +proof clarify + fix C y + assume "closedin Y C" and "y \ topspace Y" and "y \ C" + have "closed_map X Y f" "(\y \ topspace Y. compactin X {x \ topspace X. f x = y})" + using pmapf proper_map_def by force+ + moreover have "closedin X {z \ topspace X. f z \ C}" + using \closedin Y C\ contf continuous_map_closedin by fastforce + moreover have "disjnt {z \ topspace X. f z = y} {z \ topspace X. f z \ C}" + using \y \ C\ disjnt_iff by blast + ultimately + obtain U V where UV: "openin X U" "openin X V" "{z \ topspace X. f z = y} \ U" "{z \ topspace X. f z \ C} \ V" + and dUV: "disjnt U V" + using \y \ topspace Y\ \regular_space X\ unfolding regular_space_compact_closed_sets + by meson + + have *: "\U T. openin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U \ + (\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U)" + using \closed_map X Y f\ unfolding closed_map_preimage_neighbourhood by blast + obtain V1 where V1: "openin Y V1 \ y \ V1" and sub1: "{x \ topspace X. f x \ V1} \ U" + using * [of U "{y}"] UV \y \ topspace Y\ by auto + moreover + obtain V2 where "openin Y V2 \ C \ V2" and sub2: "{x \ topspace X. f x \ V2} \ V" + by (smt (verit, ccfv_SIG) * UV \closedin Y C\ closedin_subset mem_Collect_eq subset_iff) + moreover have "disjnt V1 V2" + proof - + have "\x. \\x. x \ U \ x \ V; x \ V1; x \ V2\ \ False" + by (smt (verit) V1 fim image_iff mem_Collect_eq openin_subset sub1 sub2 subsetD) + with dUV show ?thesis by (auto simp: disjnt_iff) + qed + ultimately show "\U V. openin Y U \ openin Y V \ y \ U \ C \ V \ disjnt U V" + by meson +qed + +lemma regular_space_perfect_map_image: + "\regular_space X; perfect_map X Y f\ \ regular_space Y" + by (meson perfect_map_def regular_space_continuous_proper_map_image) + +proposition regular_space_perfect_map_image_eq: + assumes "Hausdorff_space X" and perf: "perfect_map X Y f" + shows "regular_space X \ regular_space Y" (is "?lhs=?rhs") +proof + assume ?lhs + then show ?rhs + using perf regular_space_perfect_map_image by blast +next + assume R: ?rhs + have "continuous_map X Y f" "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" + using perf by (auto simp: perfect_map_def) + then have "closed_map X Y f" and preYf: "(\y \ topspace Y. compactin X {x \ topspace X. f x = y})" + by (simp_all add: proper_map_def) + show ?lhs + unfolding regular_space_def + proof clarify + fix C x + assume "closedin X C" and "x \ topspace X" and "x \ C" + obtain U1 U2 where "openin X U1" "openin X U2" "{x} \ U1" and "disjnt U1 U2" + and subV: "C \ {z \ topspace X. f z = f x} \ U2" + proof (rule Hausdorff_space_compact_separation [of X "{x}" "C \ {z \ topspace X. f z = f x}", OF \Hausdorff_space X\]) + show "compactin X {x}" + by (simp add: \x \ topspace X\) + show "compactin X (C \ {z \ topspace X. f z = f x})" + using \closedin X C\ fim \x \ topspace X\ closed_Int_compactin preYf by fastforce + show "disjnt {x} (C \ {z \ topspace X. f z = f x})" + using \x \ C\ by force + qed + have "closedin Y (f ` (C - U2))" + using \closed_map X Y f\ \closedin X C\ \openin X U2\ closed_map_def by blast + moreover + have "f x \ topspace Y - f ` (C - U2)" + using \closedin X C\ \continuous_map X Y f\ \x \ topspace X\ closedin_subset continuous_map_def subV by fastforce + ultimately + obtain V1 V2 where VV: "openin Y V1" "openin Y V2" "f x \ V1" + and subV2: "f ` (C - U2) \ V2" and "disjnt V1 V2" + by (meson R regular_space_def) + show "\U U'. openin X U \ openin X U' \ x \ U \ C \ U' \ disjnt U U'" + proof (intro exI conjI) + show "openin X (U1 \ {x \ topspace X. f x \ V1})" + using VV(1) \continuous_map X Y f\ \openin X U1\ continuous_map by fastforce + show "openin X (U2 \ {x \ topspace X. f x \ V2})" + using VV(2) \continuous_map X Y f\ \openin X U2\ continuous_map by fastforce + show "x \ U1 \ {x \ topspace X. f x \ V1}" + using VV(3) \x \ topspace X\ \{x} \ U1\ by auto + show "C \ U2 \ {x \ topspace X. f x \ V2}" + using \closedin X C\ closedin_subset subV2 by auto + show "disjnt (U1 \ {x \ topspace X. f x \ V1}) (U2 \ {x \ topspace X. f x \ V2})" + using \disjnt U1 U2\ \disjnt V1 V2\ by (auto simp: disjnt_iff) + qed + qed +qed + + + +subsection\Locally compact spaces\ + +definition locally_compact_space + where "locally_compact_space X \ + \x \ topspace X. \U K. openin X U \ compactin X K \ x \ U \ U \ K" + +lemma homeomorphic_locally_compact_spaceD: + assumes X: "locally_compact_space X" and "X homeomorphic_space Y" + shows "locally_compact_space Y" +proof - + obtain f where hmf: "homeomorphic_map X Y f" + using assms homeomorphic_space by blast + then have eq: "topspace Y = f ` (topspace X)" + by (simp add: homeomorphic_imp_surjective_map) + have "\V K. openin Y V \ compactin Y K \ f x \ V \ V \ K" + if "x \ topspace X" "openin X U" "compactin X K" "x \ U" "U \ K" for x U K + using that + by (meson hmf homeomorphic_map_compactness_eq homeomorphic_map_openness_eq image_mono image_eqI) + with X show ?thesis + by (smt (verit) eq image_iff locally_compact_space_def) +qed + +lemma homeomorphic_locally_compact_space: + assumes "X homeomorphic_space Y" + shows "locally_compact_space X \ locally_compact_space Y" + by (meson assms homeomorphic_locally_compact_spaceD homeomorphic_space_sym) + +lemma locally_compact_space_retraction_map_image: + assumes "retraction_map X Y r" and X: "locally_compact_space X" + shows "locally_compact_space Y" +proof - + obtain s where s: "retraction_maps X Y r s" + using assms retraction_map_def by blast + obtain T where "T retract_of_space X" and Teq: "T = s ` topspace Y" + using retraction_maps_section_image1 s by blast + then obtain r where r: "continuous_map X (subtopology X T) r" "\x\T. r x = x" + by (meson retract_of_space_def) + have "locally_compact_space (subtopology X T)" + unfolding locally_compact_space_def openin_subtopology_alt + proof clarsimp + fix x + assume "x \ topspace X" "x \ T" + obtain U K where UK: "openin X U \ compactin X K \ x \ U \ U \ K" + by (meson X \x \ topspace X\ locally_compact_space_def) + then have "compactin (subtopology X T) (r ` K) \ T \ U \ r ` K" + by (smt (verit) IntD1 image_compactin image_iff inf_le2 r subset_iff) + then show "\U. openin X U \ (\K. compactin (subtopology X T) K \ x \ U \ T \ U \ K)" + using UK by auto + qed + with Teq show ?thesis + using homeomorphic_locally_compact_space retraction_maps_section_image2 s by blast +qed + +lemma compact_imp_locally_compact_space: + "compact_space X \ locally_compact_space X" + using compact_space_def locally_compact_space_def by blast + +lemma neighbourhood_base_imp_locally_compact_space: + "neighbourhood_base_of (compactin X) X \ locally_compact_space X" + by (metis locally_compact_space_def neighbourhood_base_of openin_topspace) + +lemma locally_compact_imp_neighbourhood_base: + assumes loc: "locally_compact_space X" and reg: "regular_space X" + shows "neighbourhood_base_of (compactin X) X" + unfolding neighbourhood_base_of +proof clarify + fix W x + assume "openin X W" and "x \ W" + then obtain U K where "openin X U" "compactin X K" "x \ U" "U \ K" + by (metis loc locally_compact_space_def openin_subset subsetD) + moreover have "openin X (U \ W) \ x \ U \ W" + using \openin X W\ \x \ W\ \openin X U\ \x \ U\ by blast + then have "\u' v. openin X u' \ closedin X v \ x \ u' \ u' \ v \ v \ U \ v \ W" + using reg + by (metis le_infE neighbourhood_base_of neighbourhood_base_of_closedin) + then show "\U V. openin X U \ compactin X V \ x \ U \ U \ V \ V \ W" + by (meson \U \ K\ \compactin X K\ closed_compactin subset_trans) +qed + +lemma Hausdorff_regular: "\Hausdorff_space X; neighbourhood_base_of (compactin X) X\ \ regular_space X" + by (metis compactin_imp_closedin neighbourhood_base_of_closedin neighbourhood_base_of_mono) + +lemma locally_compact_Hausdorff_imp_regular_space: + assumes loc: "locally_compact_space X" and "Hausdorff_space X" + shows "regular_space X" + unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of +proof clarify + fix W x + assume "openin X W" and "x \ W" + then have "x \ topspace X" + using openin_subset by blast + then obtain U K where "openin X U" "compactin X K" and UK: "x \ U" "U \ K" + by (meson loc locally_compact_space_def) + with \Hausdorff_space X\ have "regular_space (subtopology X K)" + using Hausdorff_space_subtopology compact_Hausdorff_imp_regular_space compact_space_subtopology by blast + then have "\U' V'. openin (subtopology X K) U' \ closedin (subtopology X K) V' \ x \ U' \ U' \ V' \ V' \ K \ W" + unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of + by (meson IntI \U \ K\ \openin X W\ \x \ U\ \x \ W\ openin_subtopology_Int2 subsetD) + then obtain V C where "openin X V" "closedin X C" and VC: "x \ K \ V" "K \ V \ K \ C" "K \ C \ K \ W" + by (metis Int_commute closedin_subtopology openin_subtopology) + show "\U V. openin X U \ closedin X V \ x \ U \ U \ V \ V \ W" + proof (intro conjI exI) + show "openin X (U \ V)" + using \openin X U\ \openin X V\ by blast + show "closedin X (K \ C)" + using \closedin X C\ \compactin X K\ compactin_imp_closedin \Hausdorff_space X\ by blast + qed (use UK VC in auto) +qed + +lemma locally_compact_space_neighbourhood_base: + "Hausdorff_space X \ regular_space X + \ locally_compact_space X \ neighbourhood_base_of (compactin X) X" + by (metis locally_compact_imp_neighbourhood_base locally_compact_Hausdorff_imp_regular_space + neighbourhood_base_imp_locally_compact_space) + +lemma locally_compact_Hausdorff_or_regular: + "locally_compact_space X \ (Hausdorff_space X \ regular_space X) \ locally_compact_space X \ regular_space X" + using locally_compact_Hausdorff_imp_regular_space by blast + +lemma locally_compact_space_compact_closedin: + assumes "Hausdorff_space X \ regular_space X" + shows "locally_compact_space X \ + (\x \ topspace X. \U K. openin X U \ compactin X K \ closedin X K \ x \ U \ U \ K)" + using locally_compact_Hausdorff_or_regular unfolding locally_compact_space_def + by (metis assms closed_compactin inf.absorb_iff2 le_infE neighbourhood_base_of neighbourhood_base_of_closedin) + +lemma locally_compact_space_compact_closure_of: + assumes "Hausdorff_space X \ regular_space X" + shows "locally_compact_space X \ + (\x \ topspace X. \U. openin X U \ compactin X (X closure_of U) \ x \ U)" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + by (metis assms closed_compactin closedin_closure_of closure_of_eq closure_of_mono locally_compact_space_compact_closedin) +next + assume ?rhs then show ?lhs + by (meson closure_of_subset locally_compact_space_def openin_subset) +qed + +lemma locally_compact_space_neighbourhood_base_closedin: + assumes "Hausdorff_space X \ regular_space X" + shows "locally_compact_space X \ neighbourhood_base_of (\C. compactin X C \ closedin X C) X" (is "?lhs=?rhs") +proof + assume L: ?lhs + then have "regular_space X" + using assms locally_compact_Hausdorff_imp_regular_space by blast + with L have "neighbourhood_base_of (compactin X) X" + by (simp add: locally_compact_imp_neighbourhood_base) + with \regular_space X\ show ?rhs + by (smt (verit, ccfv_threshold) closed_compactin neighbourhood_base_of subset_trans neighbourhood_base_of_closedin) +next + assume ?rhs then show ?lhs + using neighbourhood_base_imp_locally_compact_space neighbourhood_base_of_mono by blast +qed + +lemma locally_compact_space_neighbourhood_base_closure_of: + assumes "Hausdorff_space X \ regular_space X" + shows "locally_compact_space X \ neighbourhood_base_of (\T. compactin X (X closure_of T)) X" + (is "?lhs=?rhs") +proof + assume L: ?lhs + then have "regular_space X" + using assms locally_compact_Hausdorff_imp_regular_space by blast + with L have "neighbourhood_base_of (\A. compactin X A \ closedin X A) X" + using locally_compact_space_neighbourhood_base_closedin by blast + then show ?rhs + by (simp add: closure_of_closedin neighbourhood_base_of_mono) +next + assume ?rhs then show ?lhs + unfolding locally_compact_space_def neighbourhood_base_of + by (meson closure_of_subset openin_topspace subset_trans) +qed + +lemma locally_compact_space_neighbourhood_base_open_closure_of: + assumes "Hausdorff_space X \ regular_space X" + shows "locally_compact_space X \ + neighbourhood_base_of (\U. openin X U \ compactin X (X closure_of U)) X" + (is "?lhs=?rhs") +proof + assume L: ?lhs + then have "regular_space X" + using assms locally_compact_Hausdorff_imp_regular_space by blast + then have "neighbourhood_base_of (\T. compactin X (X closure_of T)) X" + using L locally_compact_space_neighbourhood_base_closure_of by auto + with L show ?rhs + unfolding neighbourhood_base_of + by (meson closed_compactin closure_of_closure_of closure_of_eq closure_of_mono subset_trans) +next + assume ?rhs then show ?lhs + unfolding locally_compact_space_def neighbourhood_base_of + by (meson closure_of_subset openin_topspace subset_trans) +qed + +lemma locally_compact_space_compact_closed_compact: + assumes "Hausdorff_space X \ regular_space X" + shows "locally_compact_space X \ + (\K. compactin X K + \ (\U L. openin X U \ compactin X L \ closedin X L \ K \ U \ U \ L))" + (is "?lhs=?rhs") +proof + assume L: ?lhs + then obtain U L where UL: "\x \ topspace X. openin X (U x) \ compactin X (L x) \ closedin X (L x) \ x \ U x \ U x \ L x" + unfolding locally_compact_space_compact_closedin [OF assms] + by metis + show ?rhs + proof clarify + fix K + assume "compactin X K" + then have "K \ topspace X" + by (simp add: compactin_subset_topspace) + then have *: "(\U\U ` K. openin X U) \ K \ \ (U ` K)" + using UL by blast + with \compactin X K\ obtain KF where KF: "finite KF" "KF \ K" "K \ \(U ` KF)" + by (metis compactinD finite_subset_image) + show "\U L. openin X U \ compactin X L \ closedin X L \ K \ U \ U \ L" + proof (intro conjI exI) + show "openin X (\ (U ` KF))" + using "*" \KF \ K\ by fastforce + show "compactin X (\ (L ` KF))" + by (smt (verit) UL \K \ topspace X\ KF compactin_Union finite_imageI imageE subset_iff) + show "closedin X (\ (L ` KF))" + by (smt (verit) UL \K \ topspace X\ KF closedin_Union finite_imageI imageE subsetD) + qed (use UL \K \ topspace X\ KF in auto) + qed +next + assume ?rhs then show ?lhs + by (metis compactin_sing insert_subset locally_compact_space_def) +qed + +lemma locally_compact_regular_space_neighbourhood_base: + "locally_compact_space X \ regular_space X \ + neighbourhood_base_of (\C. compactin X C \ closedin X C) X" + using locally_compact_space_neighbourhood_base_closedin neighbourhood_base_of_closedin neighbourhood_base_of_mono by blast + +lemma locally_compact_kc_space: + "neighbourhood_base_of (compactin X) X \ kc_space X \ + locally_compact_space X \ Hausdorff_space X" + using Hausdorff_imp_kc_space locally_compact_imp_kc_eq_Hausdorff_space locally_compact_space_neighbourhood_base by blast + +lemma locally_compact_kc_space_alt: + "neighbourhood_base_of (compactin X) X \ kc_space X \ + locally_compact_space X \ Hausdorff_space X \ regular_space X" + using Hausdorff_regular locally_compact_kc_space by blast + +lemma locally_compact_kc_imp_regular_space: + "\neighbourhood_base_of (compactin X) X; kc_space X\ \ regular_space X" + using Hausdorff_regular locally_compact_imp_kc_eq_Hausdorff_space by blast + +lemma kc_locally_compact_space: + "kc_space X + \ neighbourhood_base_of (compactin X) X \ locally_compact_space X \ Hausdorff_space X \ regular_space X" + using Hausdorff_regular locally_compact_kc_space by blast + +lemma locally_compact_space_closed_subset: + assumes loc: "locally_compact_space X" and "closedin X S" + shows "locally_compact_space (subtopology X S)" +proof (clarsimp simp: locally_compact_space_def) + fix x assume x: "x \ topspace X" "x \ S" + then obtain U K where UK: "openin X U \ compactin X K \ x \ U \ U \ K" + by (meson loc locally_compact_space_def) + show "\U. openin (subtopology X S) U \ + (\K. compactin (subtopology X S) K \ x \ U \ U \ K)" + proof (intro conjI exI) + show "openin (subtopology X S) (S \ U)" + by (simp add: UK openin_subtopology_Int2) + show "compactin (subtopology X S) (S \ K)" + by (simp add: UK assms(2) closed_Int_compactin compactin_subtopology) + qed (use UK x in auto) +qed + +lemma locally_compact_space_open_subset: + assumes reg: "regular_space X" and loc: "locally_compact_space X" and "openin X S" + shows "locally_compact_space (subtopology X S)" +proof (clarsimp simp: locally_compact_space_def) + fix x assume x: "x \ topspace X" "x \ S" + then obtain U K where UK: "openin X U" "compactin X K" "x \ U" "U \ K" + by (meson loc locally_compact_space_def) + have "openin X (U \ S)" + by (simp add: UK \openin X S\ openin_Int) + with UK reg x obtain V C + where VC: "openin X V" "closedin X C" "x \ V" "V \ C" "C \ U" "C \ S" + by (metis IntI le_inf_iff neighbourhood_base_of neighbourhood_base_of_closedin) + show "\U. openin (subtopology X S) U \ + (\K. compactin (subtopology X S) K \ x \ U \ U \ K)" + proof (intro conjI exI) + show "openin (subtopology X S) V" + using VC by (meson \openin X S\ openin_open_subtopology order_trans) + show "compactin (subtopology X S) (C \ K)" + using UK VC closed_Int_compactin compactin_subtopology by fastforce + qed (use UK VC x in auto) +qed + +lemma locally_compact_space_discrete_topology: + "locally_compact_space (discrete_topology U)" + by (simp add: neighbourhood_base_imp_locally_compact_space neighbourhood_base_of_discrete_topology) + +lemma locally_compact_space_continuous_open_map_image: + "\continuous_map X X' f; open_map X X' f; + f ` topspace X = topspace X'; locally_compact_space X\ \ locally_compact_space X'" +unfolding locally_compact_space_def open_map_def + by (smt (verit, ccfv_SIG) image_compactin image_iff image_mono) + +lemma locally_compact_subspace_openin_closure_of: + assumes "Hausdorff_space X" and S: "S \ topspace X" + and loc: "locally_compact_space (subtopology X S)" + shows "openin (subtopology X (X closure_of S)) S" + unfolding openin_subopen [where S=S] +proof clarify + fix a assume "a \ S" + then obtain T K where *: "openin X T" "compactin X K" "K \ S" "a \ S" "a \ T" "S \ T \ K" + using loc unfolding locally_compact_space_def + by (metis IntE S compactin_subtopology inf_commute openin_subtopology topspace_subtopology_subset) + have "T \ X closure_of S \ X closure_of (T \ S)" + by (simp add: "*"(1) openin_Int_closure_of_subset) + also have "... \ S" + using * \Hausdorff_space X\ by (metis closure_of_minimal compactin_imp_closedin order.trans inf_commute) + finally have "T \ X closure_of S \ T \ S" by simp + then have "openin (subtopology X (X closure_of S)) (T \ S)" + unfolding openin_subtopology using \openin X T\ S closure_of_subset by fastforce + with * show "\T. openin (subtopology X (X closure_of S)) T \ a \ T \ T \ S" + by blast +qed + +lemma locally_compact_subspace_closed_Int_openin: + "\Hausdorff_space X \ S \ topspace X \ locally_compact_space(subtopology X S)\ + \ \C U. closedin X C \ openin X U \ C \ U = S" + by (metis closedin_closure_of inf_commute locally_compact_subspace_openin_closure_of openin_subtopology) + +lemma locally_compact_subspace_open_in_closure_of_eq: + assumes "Hausdorff_space X" and loc: "locally_compact_space X" + shows "openin (subtopology X (X closure_of S)) S \ S \ topspace X \ locally_compact_space(subtopology X S)" (is "?lhs=?rhs") +proof + assume L: ?lhs + then obtain "S \ topspace X" "regular_space X" + using assms locally_compact_Hausdorff_imp_regular_space openin_subset by fastforce + then have "locally_compact_space (subtopology (subtopology X (X closure_of S)) S)" + by (simp add: L loc locally_compact_space_closed_subset locally_compact_space_open_subset regular_space_subtopology) + then show ?rhs + by (metis L inf.orderE inf_commute le_inf_iff openin_subset subtopology_subtopology topspace_subtopology) +next + assume ?rhs then show ?lhs + using assms locally_compact_subspace_openin_closure_of by blast +qed + +lemma locally_compact_subspace_closed_Int_openin_eq: + assumes "Hausdorff_space X" and loc: "locally_compact_space X" + shows "(\C U. closedin X C \ openin X U \ C \ U = S) \ S \ topspace X \ locally_compact_space(subtopology X S)" (is "?lhs=?rhs") +proof + assume L: ?lhs + then obtain C U where "closedin X C" "openin X U" and Seq: "S = C \ U" + by blast + then have "C \ topspace X" + by (simp add: closedin_subset) + have "locally_compact_space (subtopology (subtopology X C) (topspace (subtopology X C) \ U))" + proof (rule locally_compact_space_open_subset) + show "regular_space (subtopology X C)" + by (simp add: \Hausdorff_space X\ loc locally_compact_Hausdorff_imp_regular_space regular_space_subtopology) + show "locally_compact_space (subtopology X C)" + by (simp add: \closedin X C\ loc locally_compact_space_closed_subset) + show "openin (subtopology X C) (topspace (subtopology X C) \ U)" + by (simp add: \openin X U\ Int_left_commute inf_commute openin_Int openin_subtopology_Int2) +qed + then show ?rhs + by (metis Seq \C \ topspace X\ inf.coboundedI1 subtopology_subtopology subtopology_topspace) +next + assume ?rhs then show ?lhs + using assms locally_compact_subspace_closed_Int_openin by blast +qed + +lemma dense_locally_compact_openin_Hausdorff_space: + "\Hausdorff_space X; S \ topspace X; X closure_of S = topspace X; + locally_compact_space (subtopology X S)\ \ openin X S" + by (metis locally_compact_subspace_openin_closure_of subtopology_topspace) + +lemma locally_compact_space_prod_topology: + "locally_compact_space (prod_topology X Y) \ + topspace (prod_topology X Y) = {} \ + locally_compact_space X \ locally_compact_space Y" (is "?lhs=?rhs") +proof (cases "topspace (prod_topology X Y) = {}") + case True + then show ?thesis + unfolding locally_compact_space_def by blast +next + case False + then obtain w z where wz: "w \ topspace X" "z \ topspace Y" + by auto + show ?thesis + proof + assume L: ?lhs then show ?rhs + by (metis wz empty_iff locally_compact_space_retraction_map_image retraction_map_fst retraction_map_snd) + next + assume R: ?rhs + show ?lhs + unfolding locally_compact_space_def + proof clarsimp + fix x y + assume "x \ topspace X" and "y \ topspace Y" + obtain U C where "openin X U" "compactin X C" "x \ U" "U \ C" + by (meson False R \x \ topspace X\ locally_compact_space_def) + obtain V D where "openin Y V" "compactin Y D" "y \ V" "V \ D" + by (meson False R \y \ topspace Y\ locally_compact_space_def) + show "\U. openin (prod_topology X Y) U \ (\K. compactin (prod_topology X Y) K \ (x, y) \ U \ U \ K)" + proof (intro exI conjI) + show "openin (prod_topology X Y) (U \ V)" + by (simp add: \openin X U\ \openin Y V\ openin_prod_Times_iff) + show "compactin (prod_topology X Y) (C \ D)" + by (simp add: \compactin X C\ \compactin Y D\ compactin_Times) + show "(x, y) \ U \ V" + by (simp add: \x \ U\ \y \ V\) + show "U \ V \ C \ D" + by (simp add: Sigma_mono \U \ C\ \V \ D\) + qed + qed + qed +qed + +lemma locally_compact_space_product_topology: + "locally_compact_space(product_topology X I) \ + topspace(product_topology X I) = {} \ + finite {i \ I. \ compact_space(X i)} \ (\i \ I. locally_compact_space(X i))" (is "?lhs=?rhs") +proof (cases "topspace (product_topology X I) = {}") + case True + then show ?thesis + by (simp add: locally_compact_space_def) +next + case False + show ?thesis + proof + assume L: ?lhs + obtain z where z: "z \ topspace (product_topology X I)" + using False by auto + with L z obtain U C where "openin (product_topology X I) U" "compactin (product_topology X I) C" "z \ U" "U \ C" + by (meson locally_compact_space_def) + then obtain V where finV: "finite {i \ I. V i \ topspace (X i)}" and "\i \ I. openin (X i) (V i)" + and "z \ PiE I V" "PiE I V \ U" + by (auto simp: openin_product_topology_alt) + have "compact_space (X i)" if "i \ I" "V i = topspace (X i)" for i + proof - + have "compactin (X i) ((\x. x i) ` C)" + using \compactin (product_topology X I) C\ image_compactin + by (metis continuous_map_product_projection \i \ I\) + moreover have "V i \ (\x. x i) ` C" + proof - + have "V i \ (\x. x i) ` Pi\<^sub>E I V" + by (metis \z \ Pi\<^sub>E I V\ empty_iff image_projection_PiE order_refl \i \ I\) + also have "\ \ (\x. x i) ` C" + using \U \ C\ \Pi\<^sub>E I V \ U\ by blast + finally show ?thesis . + qed + ultimately show ?thesis + by (metis closed_compactin closedin_topspace compact_space_def that(2)) + qed + with finV have "finite {i \ I. \ compact_space (X i)}" + by (metis (mono_tags, lifting) mem_Collect_eq finite_subset subsetI) + moreover have "locally_compact_space (X i)" if "i \ I" for i + by (meson False L locally_compact_space_retraction_map_image retraction_map_product_projection that) + ultimately show ?rhs by metis + next + assume R: ?rhs + show ?lhs + unfolding locally_compact_space_def + proof clarsimp + fix z + assume z: "z \ (\\<^sub>E i\I. topspace (X i))" + have "\U C. openin (X i) U \ compactin (X i) C \ z i \ U \ U \ C \ + (compact_space(X i) \ U = topspace(X i) \ C = topspace(X i))" + if "i \ I" for i + using that R z unfolding locally_compact_space_def compact_space_def + by (metis (no_types, lifting) False PiE_mem openin_topspace set_eq_subset) + then obtain U C where UC: "\i. i \ I \ + openin (X i) (U i) \ compactin (X i) (C i) \ z i \ U i \ U i \ C i \ + (compact_space(X i) \ U i = topspace(X i) \ C i = topspace(X i))" + by metis + show "\U. openin (product_topology X I) U \ (\K. compactin (product_topology X I) K \ z \ U \ U \ K)" + proof (intro exI conjI) + show "openin (product_topology X I) (Pi\<^sub>E I U)" + by (smt (verit) Collect_cong False R UC compactin_subspace openin_PiE_gen subset_antisym subtopology_topspace) + show "compactin (product_topology X I) (Pi\<^sub>E I C)" + by (simp add: UC compactin_PiE) + qed (use UC z in blast)+ + qed + qed +qed + +lemma locally_compact_space_sum_topology: + "locally_compact_space (sum_topology X I) \ (\i \ I. locally_compact_space (X i))" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + by (metis closed_map_component_injection embedding_map_imp_homeomorphic_space embedding_map_component_injection + embedding_imp_closed_map_eq homeomorphic_locally_compact_space locally_compact_space_closed_subset) +next + assume R: ?rhs + show ?lhs + unfolding locally_compact_space_def + proof clarsimp + fix i y + assume "i \ I" and y: "y \ topspace (X i)" + then obtain U K where UK: "openin (X i) U" "compactin (X i) K" "y \ U" "U \ K" + using R by (fastforce simp: locally_compact_space_def) + then show "\U. openin (sum_topology X I) U \ (\K. compactin (sum_topology X I) K \ (i, y) \ U \ U \ K)" + by (metis \i \ I\ continuous_map_component_injection image_compactin image_mono + imageI open_map_component_injection open_map_def) + qed +qed + +proposition quotient_map_prod_right: + assumes loc: "locally_compact_space Z" + and reg: "Hausdorff_space Z \ regular_space Z" + and f: "quotient_map X Y f" + shows "quotient_map (prod_topology Z X) (prod_topology Z Y) (\(x,y). (x,f y))" +proof - + define h where "h \ (\(x::'a,y). (x,f y))" + have "continuous_map (prod_topology Z X) Y (f o snd)" + by (simp add: continuous_map_of_snd f quotient_imp_continuous_map) + then have cmh: "continuous_map (prod_topology Z X) (prod_topology Z Y) h" + by (simp add: h_def continuous_map_paired split_def continuous_map_fst o_def) + have fim: "f ` topspace X = topspace Y" + by (simp add: f quotient_imp_surjective_map) + moreover + have "openin (prod_topology Z X) {u \ topspace Z \ topspace X. h u \ W} + \ openin (prod_topology Z Y) W" (is "?lhs=?rhs") + if W: "W \ topspace Z \ topspace Y" for W + proof + define S where "S \ {u \ topspace Z \ topspace X. h u \ W}" + assume ?lhs + then have L: "openin (prod_topology Z X) S" + using S_def by blast + have "\T. openin (prod_topology Z Y) T \ (x0, z0) \ T \ T \ W" + if \
: "(x0,z0) \ W" for x0 z0 + proof - + have x0: "x0 \ topspace Z" + using W that by blast + obtain y0 where y0: "y0 \ topspace X" "f y0 = z0" + by (metis W fim imageE insert_absorb insert_subset mem_Sigma_iff \
) + then have "(x0, y0) \ S" + by (simp add: S_def h_def that x0) + have "continuous_map Z (prod_topology Z X) (\x. (x, y0))" + by (simp add: continuous_map_paired y0) + with openin_continuous_map_preimage [OF _ L] + have ope_ZS: "openin Z {x \ topspace Z. (x,y0) \ S}" + by blast + obtain U U' where "openin Z U" "compactin Z U'" "closedin Z U'" + "x0 \ U" "U \ U'" "U' \ {x \ topspace Z. (x,y0) \ S}" + using loc ope_ZS x0 \(x0, y0) \ S\ + by (force simp: locally_compact_space_neighbourhood_base_closedin [OF reg] + neighbourhood_base_of) + then have D: "U' \ {y0} \ S" + by (auto simp: ) + define V where "V \ {z \ topspace Y. U' \ {y \ topspace X. f y = z} \ S}" + have "z0 \ V" + using D y0 Int_Collect fim by (fastforce simp: h_def V_def S_def) + have "openin X {x \ topspace X. f x \ V} \ openin Y V" + using f unfolding V_def quotient_map_def subset_iff + by (smt (verit, del_insts) Collect_cong mem_Collect_eq) + moreover have "openin X {x \ topspace X. f x \ V}" + proof - + let ?Z = "subtopology Z U'" + have *: "{x \ topspace X. f x \ V} = topspace X - snd ` (U' \ topspace X - S)" + by (force simp: V_def S_def h_def simp flip: fim) + have "compact_space ?Z" + using \compactin Z U'\ compactin_subspace by auto + moreover have "closedin (prod_topology ?Z X) (U' \ topspace X - S)" + by (simp add: L \closedin Z U'\ closedin_closed_subtopology closedin_diff closedin_prod_Times_iff + prod_topology_subtopology(1)) + ultimately show ?thesis + using "*" closed_map_snd closed_map_def by fastforce + qed + ultimately have "openin Y V" + by metis + show ?thesis + proof (intro conjI exI) + show "openin (prod_topology Z Y) (U \ V)" + by (simp add: openin_prod_Times_iff \openin Z U\ \openin Y V\) + show "(x0, z0) \ U \ V" + by (simp add: \x0 \ U\ \z0 \ V\) + show "U \ V \ W" + using \U \ U'\ by (force simp: V_def S_def h_def simp flip: fim) + qed + qed + with openin_subopen show ?rhs by force + next + assume ?rhs then show ?lhs + using openin_continuous_map_preimage cmh by fastforce + qed + ultimately show ?thesis + by (fastforce simp: image_iff quotient_map_def h_def) +qed + +lemma quotient_map_prod_left: + assumes loc: "locally_compact_space Z" + and reg: "Hausdorff_space Z \ regular_space Z" + and f: "quotient_map X Y f" + shows "quotient_map (prod_topology X Z) (prod_topology Y Z) (\(x,y). (f x,y))" +proof - + have "(\(x,y). (f x,y)) = prod.swap \ (\(x,y). (x,f y)) \ prod.swap" + by force + then + show ?thesis + apply (rule ssubst) + proof (intro quotient_map_compose) + show "quotient_map (prod_topology X Z) (prod_topology Z X) prod.swap" + "quotient_map (prod_topology Z Y) (prod_topology Y Z) prod.swap" + using homeomorphic_map_def homeomorphic_map_swap quotient_map_eq by fastforce+ + show "quotient_map (prod_topology Z X) (prod_topology Z Y) (\(x, y). (x, f y))" + by (simp add: f loc quotient_map_prod_right reg) + qed +qed + +lemma locally_compact_space_perfect_map_preimage: + assumes "locally_compact_space X'" and f: "perfect_map X X' f" + shows "locally_compact_space X" + unfolding locally_compact_space_def +proof (intro strip) + fix x + assume x: "x \ topspace X" + then obtain U K where "openin X' U" "compactin X' K" "f x \ U" "U \ K" + using assms unfolding locally_compact_space_def perfect_map_def + by (metis (no_types, lifting) continuous_map_closedin) + show "\U K. openin X U \ compactin X K \ x \ U \ U \ K" + proof (intro exI conjI) + have "continuous_map X X' f" + using f perfect_map_def by blast + then show "openin X {x \ topspace X. f x \ U}" + by (simp add: \openin X' U\ continuous_map) + show "compactin X {x \ topspace X. f x \ K}" + using \compactin X' K\ f perfect_imp_proper_map proper_map_alt by blast + qed (use x \f x \ U\ \U \ K\ in auto) +qed + + +subsection\Special characterizations of classes of functions into and out of R\ + +lemma monotone_map_into_euclideanreal_alt: + assumes "continuous_map X euclideanreal f" + shows "(\k. is_interval k \ connectedin X {x \ topspace X. f x \ k}) \ + connected_space X \ monotone_map X euclideanreal f" (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + proof + show "connected_space X" + using L connected_space_subconnected by blast + have "connectedin X {x \ topspace X. f x \ {y}}" for y + by (metis L is_interval_1 nle_le singletonD) + then show "monotone_map X euclideanreal f" + by (simp add: monotone_map) + qed +next + assume R: ?rhs + then + have *: False + if "a < b" "closedin X U" "closedin X V" "U \ {}" "V \ {}" "disjnt U V" + and UV: "{x \ topspace X. f x \ {a..b}} = U \ V" + and dis: "disjnt U {x \ topspace X. f x = b}" "disjnt V {x \ topspace X. f x = a}" + for a b U V + proof - + define E1 where "E1 \ U \ {x \ topspace X. f x \ {c. c \ a}}" + define E2 where "E2 \ V \ {x \ topspace X. f x \ {c. b \ c}}" + have "closedin X {x \ topspace X. f x \ a}" "closedin X {x \ topspace X. b \ f x}" + using assms continuous_map_upper_lower_semicontinuous_le by blast+ + then have "closedin X E1" "closedin X E2" + unfolding E1_def E2_def using that by auto + moreover + have "E1 \ E2 = {}" + unfolding E1_def E2_def using \a \disjnt U V\ dis UV + by (simp add: disjnt_def set_eq_iff) (smt (verit)) + have "topspace X \ E1 \ E2" + unfolding E1_def E2_def using UV by fastforce + have "E1 = {} \ E2 = {}" + using R connected_space_closedin + using \E1 \ E2 = {}\ \closedin X E1\ \closedin X E2\ \topspace X \ E1 \ E2\ by blast + then show False + using E1_def E2_def \U \ {}\ \V \ {}\ by fastforce + qed + show ?lhs + proof (intro strip) + fix K :: "real set" + assume "is_interval K" + have False + if "a \ K" "b \ K" and clo: "closedin X U" "closedin X V" + and UV: "{x. x \ topspace X \ f x \ K} \ U \ V" + "U \ V \ {x. x \ topspace X \ f x \ K} = {}" + and nondis: "\ disjnt U {x. x \ topspace X \ f x = a}" + "\ disjnt V {x. x \ topspace X \ f x = b}" + for a b U V + proof - + have "\y. connectedin X {x. x \ topspace X \ f x = y}" + using R monotone_map by fastforce + then have **: False if "p \ U \ q \ V \ f p = f q \ f q \ K" for p q + unfolding connectedin_closedin + using \a \ K\ \b \ K\ UV clo that + by (smt (verit, ccfv_threshold) closedin_subset disjoint_iff mem_Collect_eq subset_iff) + consider "a < b" | "a = b" | "b < a" + by linarith + then show ?thesis + proof cases + case 1 + define W where "W \ {x \ topspace X. f x \ {a..b}}" + have "closedin X W" + unfolding W_def + by (metis (no_types) assms closed_real_atLeastAtMost closed_closedin continuous_map_closedin) + show ?thesis + proof (rule * [OF 1 , of "U \ W" "V \ W"]) + show "closedin X (U \ W)" "closedin X (V \ W)" + using \closedin X W\ clo by auto + show "U \ W \ {}" "V \ W \ {}" + using nondis 1 by (auto simp: disjnt_iff W_def) + show "disjnt (U \ W) (V \ W)" + using \is_interval K\ unfolding is_interval_1 disjnt_iff W_def + by (metis (mono_tags, lifting) \a \ K\ \b \ K\ ** Int_Collect atLeastAtMost_iff) + have "\x. \x \ topspace X; a \ f x; f x \ b\ \ x \ U \ x \ V" + using \a \ K\ \b \ K\ \is_interval K\ UV unfolding is_interval_1 disjnt_iff + by blast + then show "{x \ topspace X. f x \ {a..b}} = U \ W \ V \ W" + by (auto simp: W_def) + show "disjnt (U \ W) {x \ topspace X. f x = b}" "disjnt (V \ W) {x \ topspace X. f x = a}" + using ** \a \ K\ \b \ K\ nondis by (force simp: disjnt_iff)+ + qed + next + case 2 + then show ?thesis + using ** nondis \b \ K\ by (force simp add: disjnt_iff) + next + case 3 + define W where "W \ {x \ topspace X. f x \ {b..a}}" + have "closedin X W" + unfolding W_def + by (metis (no_types) assms closed_real_atLeastAtMost closed_closedin continuous_map_closedin) + show ?thesis + proof (rule * [OF 3, of "V \ W" "U \ W"]) + show "closedin X (U \ W)" "closedin X (V \ W)" + using \closedin X W\ clo by auto + show "U \ W \ {}" "V \ W \ {}" + using nondis 3 by (auto simp: disjnt_iff W_def) + show "disjnt (V \ W) (U \ W)" + using \is_interval K\ unfolding is_interval_1 disjnt_iff W_def + by (metis (mono_tags, lifting) \a \ K\ \b \ K\ ** Int_Collect atLeastAtMost_iff) + have "\x. \x \ topspace X; b \ f x; f x \ a\ \ x \ U \ x \ V" + using \a \ K\ \b \ K\ \is_interval K\ UV unfolding is_interval_1 disjnt_iff + by blast + then show "{x \ topspace X. f x \ {b..a}} = V \ W \ U \ W" + by (auto simp: W_def) + show "disjnt (V \ W) {x \ topspace X. f x = a}" "disjnt (U \ W) {x \ topspace X. f x = b}" + using ** \a \ K\ \b \ K\ nondis by (force simp: disjnt_iff)+ + qed + qed + qed + then show "connectedin X {x \ topspace X. f x \ K}" + unfolding connectedin_closedin disjnt_iff by blast + qed +qed + +lemma monotone_map_into_euclideanreal: + "\connected_space X; continuous_map X euclideanreal f\ + \ monotone_map X euclideanreal f \ + (\k. is_interval k \ connectedin X {x \ topspace X. f x \ k})" + by (simp add: monotone_map_into_euclideanreal_alt) + +lemma monotone_map_euclideanreal_alt: + "(\I::real set. is_interval I \ is_interval {x::real. x \ S \ f x \ I}) \ + is_interval S \ (mono_on S f \ antimono_on S f)" (is "?lhs=?rhs") +proof + assume L [rule_format]: ?lhs + show ?rhs + proof + show "is_interval S" + using L is_interval_1 by auto + have False if "a \ S" "b \ S" "c \ S" "a f c < f b \ f a > f b \ f c > f b" for a b c + using d + proof + assume "f a < f b \ f c < f b" + then show False + using L [of "{y. y < f b}"] unfolding is_interval_1 + by (smt (verit, best) mem_Collect_eq that) + next + assume "f b < f a \ f b < f c" + then show False + using L [of "{y. y > f b}"] unfolding is_interval_1 + by (smt (verit, best) mem_Collect_eq that) + qed + then show "mono_on S f \ monotone_on S (\) (\) f" + unfolding monotone_on_def by (smt (verit)) + qed +next + assume ?rhs then show ?lhs + unfolding is_interval_1 monotone_on_def by simp meson +qed + + +lemma monotone_map_euclideanreal: + fixes S :: "real set" + shows + "\is_interval S; continuous_on S f\ \ + monotone_map (top_of_set S) euclideanreal f \ (mono_on S f \ monotone_on S (\) (\) f)" + using monotone_map_euclideanreal_alt + by (simp add: monotone_map_into_euclideanreal connectedin_subtopology is_interval_connected_1) + +lemma injective_eq_monotone_map: + fixes f :: "real \ real" + assumes "is_interval S" "continuous_on S f" + shows "inj_on f S \ strict_mono_on S f \ strict_antimono_on S f" + by (metis assms injective_imp_monotone_map monotone_map_euclideanreal strict_antimono_iff_antimono + strict_mono_iff_mono top_greatest topspace_euclidean topspace_euclidean_subtopology) + + +subsection\Normal spaces including Urysohn's lemma and the Tietze extension theorem\ + +definition normal_space + where "normal_space X \ + \S T. closedin X S \ closedin X T \ disjnt S T + \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V)" + +lemma normal_space_retraction_map_image: + assumes r: "retraction_map X Y r" and X: "normal_space X" + shows "normal_space Y" + unfolding normal_space_def +proof clarify + fix S T + assume "closedin Y S" and "closedin Y T" and "disjnt S T" + obtain r' where r': "retraction_maps X Y r r'" + using r retraction_map_def by blast + have "closedin X {x \ topspace X. r x \ S}" "closedin X {x \ topspace X. r x \ T}" + using closedin_continuous_map_preimage \closedin Y S\ \closedin Y T\ r' + by (auto simp: retraction_maps_def) + moreover + have "disjnt {x \ topspace X. r x \ S} {x \ topspace X. r x \ T}" + using \disjnt S T\ by (auto simp: disjnt_def) + ultimately + obtain U V where UV: "openin X U \ openin X V \ {x \ topspace X. r x \ S} \ U \ {x \ topspace X. r x \ T} \ V" "disjnt U V" + by (meson X normal_space_def) + show "\U V. openin Y U \ openin Y V \ S \ U \ T \ V \ disjnt U V" + proof (intro exI conjI) + show "openin Y {x \ topspace Y. r' x \ U}" "openin Y {x \ topspace Y. r' x \ V}" + using openin_continuous_map_preimage UV r' + by (auto simp: retraction_maps_def) + show "S \ {x \ topspace Y. r' x \ U}" "T \ {x \ topspace Y. r' x \ V}" + using openin_continuous_map_preimage UV r' \closedin Y S\ \closedin Y T\ + by (auto simp add: closedin_def continuous_map_closedin retraction_maps_def subset_iff) + show "disjnt {x \ topspace Y. r' x \ U} {x \ topspace Y. r' x \ V}" + using \disjnt U V\ by (auto simp: disjnt_def) + qed +qed + +lemma homeomorphic_normal_space: + "X homeomorphic_space Y \ normal_space X \ normal_space Y" + unfolding homeomorphic_space_def + by (meson homeomorphic_imp_retraction_maps homeomorphic_maps_sym normal_space_retraction_map_image retraction_map_def) + +lemma normal_space: + "normal_space X \ + (\S T. closedin X S \ closedin X T \ disjnt S T + \ (\U. openin X U \ S \ U \ disjnt T (X closure_of U)))" +proof - + have "(\V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V) \ openin X U \ S \ U \ disjnt T (X closure_of U)" + (is "?lhs=?rhs") + if "closedin X S" "closedin X T" "disjnt S T" for S T U + proof + show "?lhs \ ?rhs" + by (smt (verit, best) disjnt_iff in_closure_of subsetD) + assume R: ?rhs + then have "(U \ S) \ (topspace X - X closure_of U) = {}" + by (metis Diff_eq_empty_iff Int_Diff Int_Un_eq(4) closure_of_subset inf.orderE openin_subset) + moreover have "T \ topspace X - X closure_of U" + by (meson DiffI R closedin_subset disjnt_iff subsetD subsetI that(2)) + ultimately show ?lhs + by (metis R closedin_closure_of closedin_def disjnt_def sup.orderE) + qed + then show ?thesis + unfolding normal_space_def by meson +qed + +lemma normal_space_alt: + "normal_space X \ + (\S U. closedin X S \ openin X U \ S \ U \ (\V. openin X V \ S \ V \ X closure_of V \ U))" +proof - + have "\V. openin X V \ S \ V \ X closure_of V \ U" + if "\T. closedin X T \ disjnt S T \ (\U. openin X U \ S \ U \ disjnt T (X closure_of U))" + "closedin X S" "openin X U" "S \ U" + for S U + using that + by (smt (verit) Diff_eq_empty_iff Int_Diff closure_of_subset_topspace disjnt_def inf.orderE inf_commute openin_closedin_eq) + moreover have "\U. openin X U \ S \ U \ disjnt T (X closure_of U)" + if "\U. openin X U \ S \ U \ (\V. openin X V \ S \ V \ X closure_of V \ U)" + and "closedin X S" "closedin X T" "disjnt S T" + for S T + using that + by (smt (verit) Diff_Diff_Int Diff_eq_empty_iff Int_Diff closedin_def disjnt_def inf.absorb_iff2 inf.orderE) + ultimately show ?thesis + by (fastforce simp: normal_space) +qed + +lemma normal_space_closures: + "normal_space X \ + (\S T. S \ topspace X \ T \ topspace X \ + disjnt (X closure_of S) (X closure_of T) + \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V))" + (is "?lhs=?rhs") +proof + show "?lhs \ ?rhs" + by (meson closedin_closure_of closure_of_subset normal_space_def order.trans) + show "?rhs \ ?lhs" + by (metis closedin_subset closure_of_eq normal_space_def) +qed + +lemma normal_space_disjoint_closures: + "normal_space X \ + (\S T. closedin X S \ closedin X T \ disjnt S T + \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ + disjnt (X closure_of U) (X closure_of V)))" + (is "?lhs=?rhs") +proof + show "?lhs \ ?rhs" + by (metis closedin_closure_of normal_space) + show "?rhs \ ?lhs" + by (smt (verit) closure_of_subset disjnt_iff normal_space openin_subset subset_eq) +qed + +lemma normal_space_dual: + "normal_space X \ + (\U V. openin X U \ openin X V \ U \ V = topspace X + \ (\S T. closedin X S \ closedin X T \ S \ U \ T \ V \ S \ T = topspace X))" + (is "_ = ?rhs") +proof - + have "normal_space X \ + (\U V. closedin X U \ closedin X V \ disjnt U V \ + (\S T. \ (openin X S \ openin X T \ + \ (U \ S \ V \ T \ disjnt S T))))" + unfolding normal_space_def by meson + also have "... \ (\U V. openin X U \ openin X V \ disjnt (topspace X - U) (topspace X - V) \ + (\S T. \ (openin X S \ openin X T \ + \ (topspace X - U \ S \ topspace X - V \ T \ disjnt S T))))" + by (auto simp: all_closedin) + also have "... \ ?rhs" + proof - + have *: "disjnt (topspace X - U) (topspace X - V) \ U \ V = topspace X" + if "U \ topspace X" "V \ topspace X" for U V + using that by (auto simp: disjnt_iff) + show ?thesis + using ex_closedin * + apply (simp add: ex_closedin * [OF openin_subset openin_subset] cong: conj_cong) + apply (intro all_cong1 ex_cong1 imp_cong refl) + by (smt (verit, best) "*" Diff_Diff_Int Diff_subset Diff_subset_conv inf.orderE inf_commute openin_subset sup_commute) + qed + finally show ?thesis . +qed + + +lemma normal_t1_imp_Hausdorff_space: + assumes "normal_space X" "t1_space X" + shows "Hausdorff_space X" + unfolding Hausdorff_space_def +proof clarify + fix x y + assume xy: "x \ topspace X" "y \ topspace X" "x \ y" + then have "disjnt {x} {y}" + by (auto simp: disjnt_iff) + then show "\U V. openin X U \ openin X V \ x \ U \ y \ V \ disjnt U V" + using assms xy closedin_t1_singleton normal_space_def + by (metis singletonI subsetD) +qed + +lemma normal_t1_eq_Hausdorff_space: + "normal_space X \ t1_space X \ Hausdorff_space X" + using normal_t1_imp_Hausdorff_space t1_or_Hausdorff_space by blast + +lemma normal_t1_imp_regular_space: + "\normal_space X; t1_space X\ \ regular_space X" + by (metis compactin_imp_closedin normal_space_def normal_t1_eq_Hausdorff_space regular_space_compact_closed_sets) + +lemma compact_Hausdorff_or_regular_imp_normal_space: + "\compact_space X; Hausdorff_space X \ regular_space X\ + \ normal_space X" + by (metis Hausdorff_space_compact_sets closedin_compact_space normal_space_def regular_space_compact_closed_sets) + +lemma normal_space_discrete_topology: + "normal_space(discrete_topology U)" + by (metis discrete_topology_closure_of inf_le2 normal_space_alt) + +lemma normal_space_fsigmas: + "normal_space X \ + (\S T. fsigma_in X S \ fsigma_in X T \ separatedin X S T + \ (\U B. openin X U \ openin X B \ S \ U \ T \ B \ disjnt U B))" (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + proof clarify + fix S T + assume "fsigma_in X S" + then obtain C where C: "\n. closedin X (C n)" "\n. C n \ C (Suc n)" "\ (range C) = S" + by (meson fsigma_in_ascending) + assume "fsigma_in X T" + then obtain D where D: "\n. closedin X (D n)" "\n. D n \ D (Suc n)" "\ (range D) = T" + by (meson fsigma_in_ascending) + assume "separatedin X S T" + have "\n. disjnt (D n) (X closure_of S)" + by (metis D(3) \separatedin X S T\ disjnt_Union1 disjnt_def rangeI separatedin_def) + then have "\n. \V V'. openin X V \ openin X V' \ D n \ V \ X closure_of S \ V' \ disjnt V V'" + by (metis D(1) L closedin_closure_of normal_space_def) + then obtain V V' where V: "\n. openin X (V n)" and "\n. openin X (V' n)" "\n. disjnt (V n) (V' n)" + and DV: "\n. D n \ V n" + and subV': "\n. X closure_of S \ V' n" + by metis + then have VV: "V' n \ X closure_of V n = {}" for n + using openin_Int_closure_of_eq_empty [of X "V' n" "V n"] by (simp add: Int_commute disjnt_def) + have "\n. disjnt (C n) (X closure_of T)" + by (metis C(3) \separatedin X S T\ disjnt_Union1 disjnt_def rangeI separatedin_def) + then have "\n. \U U'. openin X U \ openin X U' \ C n \ U \ X closure_of T \ U' \ disjnt U U'" + by (metis C(1) L closedin_closure_of normal_space_def) + then obtain U U' where U: "\n. openin X (U n)" and "\n. openin X (U' n)" "\n. disjnt (U n) (U' n)" + and CU: "\n. C n \ U n" + and subU': "\n. X closure_of T \ U' n" + by metis + then have UU: "U' n \ X closure_of U n = {}" for n + using openin_Int_closure_of_eq_empty [of X "U' n" "U n"] by (simp add: Int_commute disjnt_def) + show "\U B. openin X U \ openin X B \ S \ U \ T \ B \ disjnt U B" + proof (intro conjI exI) + have "\S n. closedin X (\m\n. X closure_of V m)" + by (force intro: closedin_Union) + then show "openin X (\n. U n - (\m\n. X closure_of V m))" + using U by blast + have "\S n. closedin X (\m\n. X closure_of U m)" + by (force intro: closedin_Union) + then show "openin X (\n. V n - (\m\n. X closure_of U m))" + using V by blast + have "S \ topspace X" + by (simp add: \fsigma_in X S\ fsigma_in_subset) + then show "S \ (\n. U n - (\m\n. X closure_of V m))" + apply (clarsimp simp: Ball_def) + by (metis VV C(3) CU IntI UN_E closure_of_subset empty_iff subV' subsetD) + have "T \ topspace X" + by (simp add: \fsigma_in X T\ fsigma_in_subset) + then show "T \ (\n. V n - (\m\n. X closure_of U m))" + apply (clarsimp simp: Ball_def) + by (metis UU D(3) DV IntI UN_E closure_of_subset empty_iff subU' subsetD) + have "\x m n. \x \ U n; x \ V m; \k\m. x \ X closure_of U k\ \ \k\n. x \ X closure_of V k" + by (meson U V closure_of_subset nat_le_linear openin_subset subsetD) + then show "disjnt (\n. U n - (\m\n. X closure_of V m)) (\n. V n - (\m\n. X closure_of U m))" + by (force simp: disjnt_iff) + qed + qed +next + show "?rhs \ ?lhs" + by (simp add: closed_imp_fsigma_in normal_space_def separatedin_closed_sets) +qed + +lemma normal_space_fsigma_subtopology: + assumes "normal_space X" "fsigma_in X S" + shows "normal_space (subtopology X S)" + unfolding normal_space_fsigmas +proof clarify + fix T U + assume "fsigma_in (subtopology X S) T" + and "fsigma_in (subtopology X S) U" + and TU: "separatedin (subtopology X S) T U" + then obtain A B where "openin X A \ openin X B \ T \ A \ U \ B \ disjnt A B" + by (metis assms fsigma_in_fsigma_subtopology normal_space_fsigmas separatedin_subtopology) + then + show "\A B. openin (subtopology X S) A \ openin (subtopology X S) B \ T \ A \ + U \ B \ disjnt A B" + using TU + by (force simp add: separatedin_subtopology openin_subtopology_alt disjnt_iff) +qed + +lemma normal_space_closed_subtopology: + assumes "normal_space X" "closedin X S" + shows "normal_space (subtopology X S)" + by (simp add: assms closed_imp_fsigma_in normal_space_fsigma_subtopology) + +lemma normal_space_continuous_closed_map_image: + assumes "normal_space X" and contf: "continuous_map X Y f" + and clof: "closed_map X Y f" and fim: "f ` topspace X = topspace Y" +shows "normal_space Y" + unfolding normal_space_def +proof clarify + fix S T + assume "closedin Y S" and "closedin Y T" and "disjnt S T" + have "closedin X {x \ topspace X. f x \ S}" "closedin X {x \ topspace X. f x \ T}" + using \closedin Y S\ \closedin Y T\ closedin_continuous_map_preimage contf by auto + moreover + have "disjnt {x \ topspace X. f x \ S} {x \ topspace X. f x \ T}" + using \disjnt S T\ by (auto simp: disjnt_iff) + ultimately + obtain U V where "closedin X U" "closedin X V" + and subXU: "{x \ topspace X. f x \ S} \ topspace X - U" + and subXV: "{x \ topspace X. f x \ T} \ topspace X - V" + and dis: "disjnt (topspace X - U) (topspace X -V)" + using \normal_space X\ by (force simp add: normal_space_def ex_openin) + have "closedin Y (f ` U)" "closedin Y (f ` V)" + using \closedin X U\ \closedin X V\ clof closed_map_def by blast+ + moreover have "S \ topspace Y - f ` U" + using \closedin Y S\ \closedin X U\ subXU by (force dest: closedin_subset) + moreover have "T \ topspace Y - f ` V" + using \closedin Y T\ \closedin X V\ subXV by (force dest: closedin_subset) + moreover have "disjnt (topspace Y - f ` U) (topspace Y - f ` V)" + using fim dis by (force simp add: disjnt_iff) + ultimately show "\U V. openin Y U \ openin Y V \ S \ U \ T \ V \ disjnt U V" + by (force simp add: ex_openin) +qed + + +subsection \Hereditary topological properties\ + +definition hereditarily + where "hereditarily P X \ + \S. S \ topspace X \ P(subtopology X S)" + +lemma hereditarily: + "hereditarily P X \ (\S. P(subtopology X S))" + by (metis Int_lower1 hereditarily_def subtopology_restrict) + +lemma hereditarily_mono: + "\hereditarily P X; \x. P x \ Q x\ \ hereditarily Q X" + by (simp add: hereditarily) + +lemma hereditarily_inc: + "hereditarily P X \ P X" + by (metis hereditarily subtopology_topspace) + +lemma hereditarily_subtopology: + "hereditarily P X \ hereditarily P (subtopology X S)" + by (simp add: hereditarily subtopology_subtopology) + +lemma hereditarily_normal_space_continuous_closed_map_image: + assumes X: "hereditarily normal_space X" and contf: "continuous_map X Y f" + and clof: "closed_map X Y f" and fim: "f ` (topspace X) = topspace Y" + shows "hereditarily normal_space Y" + unfolding hereditarily_def +proof (intro strip) + fix T + assume "T \ topspace Y" + then have nx: "normal_space (subtopology X {x \ topspace X. f x \ T})" + by (meson X hereditarily) + moreover have "continuous_map (subtopology X {x \ topspace X. f x \ T}) (subtopology Y T) f" + by (simp add: contf continuous_map_from_subtopology continuous_map_in_subtopology image_subset_iff) + moreover have "closed_map (subtopology X {x \ topspace X. f x \ T}) (subtopology Y T) f" + by (simp add: clof closed_map_restriction) + ultimately show "normal_space (subtopology Y T)" + using fim normal_space_continuous_closed_map_image by fastforce +qed + +lemma homeomorphic_hereditarily_normal_space: + "X homeomorphic_space Y + \ (hereditarily normal_space X \ hereditarily normal_space Y)" + by (meson hereditarily_normal_space_continuous_closed_map_image homeomorphic_eq_everything_map + homeomorphic_space homeomorphic_space_sym) + +lemma hereditarily_normal_space_retraction_map_image: + "\retraction_map X Y r; hereditarily normal_space X\ \ hereditarily normal_space Y" + by (smt (verit) hereditarily_subtopology hereditary_imp_retractive_property homeomorphic_hereditarily_normal_space) + +subsection\Limits in a topological space\ + +lemma limitin_const_iff: + assumes "t1_space X" "\ trivial_limit F" + shows "limitin X (\k. a) l F \ l \ topspace X \ a = l" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + using assms unfolding limitin_def t1_space_def by (metis eventually_const openin_topspace) +next + assume ?rhs then show ?lhs + using assms by (auto simp: limitin_def t1_space_def) +qed + +lemma compactin_sequence_with_limit: + assumes lim: "limitin X \ l sequentially" and "S \ range \" and SX: "S \ topspace X" + shows "compactin X (insert l S)" +unfolding compactin_def +proof (intro conjI strip) + show "insert l S \ topspace X" + by (meson SX insert_subset lim limitin_topspace) + fix \ + assume \
: "Ball \ (openin X) \ insert l S \ \ \" + have "\V. finite V \ V \ \ \ (\t \ V. l \ t) \ S \ \ V" + if *: "\x \ S. \T \ \. x \ T" and "T \ \" "l \ T" for T + proof - + obtain V where V: "\x. x \ S \ V x \ \ \ x \ V x" + using * by metis + obtain N where N: "\n. N \ n \ \ n \ T" + by (meson "\
" \T \ \\ \l \ T\ lim limitin_sequentially) + show ?thesis + proof (intro conjI exI) + have "x \ T" + if "x \ S" and "\A. (\x \ S. (\n\N. x \ \ n) \ A \ V x) \ x \ A" for x + by (metis (no_types) N V that assms(2) imageE nle_le subsetD) + then show "S \ \ (insert T (V ` (S \ \ ` {0..N})))" + by force + qed (use V that in auto) + qed + then show "\\. finite \ \ \ \ \ \ insert l S \ \ \" + by (smt (verit, best) Union_iff \
insert_subset subsetD) +qed + +lemma limitin_Hausdorff_unique: + assumes "limitin X f l1 F" "limitin X f l2 F" "\ trivial_limit F" "Hausdorff_space X" + shows "l1 = l2" +proof (rule ccontr) + assume "l1 \ l2" + with assms obtain U V where "openin X U" "openin X V" "l1 \ U" "l2 \ V" "disjnt U V" + by (metis Hausdorff_space_def limitin_topspace) + then have "eventually (\x. f x \ U) F" "eventually (\x. f x \ V) F" + using assms by (fastforce simp: limitin_def)+ + then have "\x. f x \ U \ f x \ V" + using assms eventually_elim2 filter_eq_iff by fastforce + with assms \disjnt U V\ show False + by (meson disjnt_iff) +qed + +lemma limitin_kc_unique: + assumes "kc_space X" and lim1: "limitin X f l1 sequentially" and lim2: "limitin X f l2 sequentially" + shows "l1 = l2" +proof (rule ccontr) + assume "l1 \ l2" + define A where "A \ insert l1 (range f - {l2})" + have "l1 \ topspace X" + using lim1 limitin_def by fastforce + moreover have "compactin X (insert l1 (topspace X \ (range f - {l2})))" + by (meson Diff_subset compactin_sequence_with_limit inf_le1 inf_le2 lim1 subset_trans) + ultimately have "compactin X (topspace X \ A)" + by (simp add: A_def) + then have OXA: "openin X (topspace X - A)" + by (metis Diff_Diff_Int Diff_subset \kc_space X\ kc_space_def openin_closedin_eq) + have "l2 \ topspace X - A" + using \l1 \ l2\ A_def lim2 limitin_topspace by fastforce + then have "\\<^sub>F x in sequentially. f x = l2" + using limitinD [OF lim2 OXA] by (auto simp: A_def eventually_conj_iff) + then show False + using limitin_transform_eventually [OF _ lim1] + limitin_const_iff [OF kc_imp_t1_space trivial_limit_sequentially] + using \l1 \ l2\ \kc_space X\ by fastforce +qed + +lemma limitin_closedin: + assumes lim: "limitin X f l F" + and "closedin X S" and ev: "eventually (\x. f x \ S) F" "\ trivial_limit F" + shows "l \ S" +proof (rule ccontr) + assume "l \ S" + have "\\<^sub>F x in F. f x \ topspace X - S" + by (metis Diff_iff \l \ S\ \closedin X S\ closedin_def lim limitin_def) + with ev eventually_elim2 trivial_limit_def show False + by force +qed + +end + diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Abstract_Topology.thy --- a/src/HOL/Analysis/Abstract_Topology.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Analysis/Abstract_Topology.thy Sun May 07 22:51:23 2023 +0200 @@ -1,5 +1,4 @@ -(* Author: L C Paulson, University of Cambridge [ported from HOL Light] -*) +(* Author: L C Paulson, University of Cambridge [ported from HOL Light] *) section \Operators involving abstract topology\ @@ -174,6 +173,18 @@ shows "closedin U (S - T)" by (metis Int_Diff cT closedin_Int closedin_subset inf.orderE oS openin_closedin_eq) +lemma all_openin: "(\U. openin X U \ P U) \ (\U. closedin X U \ P (topspace X - U))" + by (metis Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq) + +lemma all_closedin: "(\U. closedin X U \ P U) \ (\U. openin X U \ P (topspace X - U))" + by (metis Diff_Diff_Int closedin_subset inf.absorb_iff2 openin_closedin_eq) + +lemma ex_openin: "(\U. openin X U \ P U) \ (\U. closedin X U \ P (topspace X - U))" + by (metis Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq) + +lemma ex_closedin: "(\U. closedin X U \ P U) \ (\U. openin X U \ P (topspace X - U))" + by (metis Diff_Diff_Int closedin_subset inf.absorb_iff2 openin_closedin_eq) + subsection\The discrete topology\ @@ -285,6 +296,10 @@ unfolding closedin_def topspace_subtopology by (auto simp: openin_subtopology) +lemma closedin_relative_to: + "(closedin X relative_to S) = closedin (subtopology X S)" + by (force simp: relative_to_def closedin_subtopology) + lemma openin_subtopology_refl: "openin (subtopology U V) V \ V \ topspace U" unfolding openin_subtopology by auto (metis IntD1 in_mono openin_subset) @@ -368,6 +383,10 @@ "closedin X S \ (closedin (subtopology X S) T \ closedin X T \ T \ S)" by (metis closedin_Int closedin_imp_subset closedin_subtopology inf.orderE) +lemma closedin_trans_full: + "\closedin (subtopology X U) S; closedin X U\ \ closedin X S" + using closedin_closed_subtopology by blast + lemma openin_subtopology_Un: "\openin (subtopology X T) S; openin (subtopology X U) S\ \ openin (subtopology X (T \ U)) S" @@ -1371,7 +1390,7 @@ by (metis (mono_tags) closedin_closure_of closedin_locally_finite_Union imageE locally_finite_in_closure) lemma closure_of_Union_subset: "\((\S. X closure_of S) ` \) \ X closure_of (\\)" - by clarify (meson Union_upper closure_of_mono subsetD) + by (simp add: SUP_le_iff Sup_upper closure_of_mono) lemma closure_of_locally_finite_Union: assumes "locally_finite_in X \" @@ -1867,6 +1886,250 @@ by (metis \closedin X T\ closed_map_def closedin_subtopology f) qed +lemma closed_map_closure_of_image: + "closed_map X Y f \ + f ` topspace X \ topspace Y \ + (\S. S \ topspace X \ Y closure_of (f ` S) \ image f (X closure_of S))" (is "?lhs=?rhs") +proof + assume ?lhs + then show ?rhs + by (simp add: closed_map_def closed_map_imp_subset_topspace closure_of_minimal closure_of_subset image_mono) +next + assume ?rhs + then show ?lhs + by (metis closed_map_def closed_map_into_discrete_topology closure_of_eq + closure_of_subset_eq topspace_discrete_topology) +qed + +lemma open_map_interior_of_image_subset: + "open_map X Y f \ (\S. image f (X interior_of S) \ Y interior_of (f ` S))" + by (metis image_mono interior_of_eq interior_of_maximal interior_of_subset open_map_def openin_interior_of subset_antisym) + +lemma open_map_interior_of_image_subset_alt: + "open_map X Y f \ (\S\topspace X. f ` (X interior_of S) \ Y interior_of f ` S)" + by (metis interior_of_eq open_map_def open_map_interior_of_image_subset openin_subset subset_interior_of_eq) + +lemma open_map_interior_of_image_subset_gen: + "open_map X Y f \ + (f ` topspace X \ topspace Y \ (\S. f ` (X interior_of S) \ Y interior_of f ` S))" + by (meson open_map_imp_subset_topspace open_map_interior_of_image_subset) + +lemma open_map_preimage_neighbourhood: + "open_map X Y f \ + (f ` topspace X \ topspace Y \ + (\U T. closedin X U \ T \ topspace Y \ + {x \ topspace X. f x \ T} \ U \ + (\V. closedin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U)))" (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + proof (intro conjI strip) + show "f ` topspace X \ topspace Y" + by (simp add: \open_map X Y f\ open_map_imp_subset_topspace) + next + fix U T + assume UT: "closedin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U" + have "closedin Y (topspace Y - f ` (topspace X - U))" + by (meson UT L open_map_def openin_closedin_eq openin_diff openin_topspace) + with UT + show "\V. closedin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U" + using image_iff by auto + qed +next + assume R: ?rhs + show ?lhs + unfolding open_map_def + proof (intro strip) + fix U assume "openin X U" + have "{x \ topspace X. f x \ topspace Y - f ` U} \ topspace X - U" + by blast + then obtain V where V: "closedin Y V" + and sub: "topspace Y - f ` U \ V" "{x \ topspace X. f x \ V} \ topspace X - U" + using R \openin X U\ by (meson Diff_subset openin_closedin_eq) + then have "f ` U \ topspace Y - V" + using R \openin X U\ openin_subset by fastforce + with sub have "f ` U = topspace Y - V" + by auto + then show "openin Y (f ` U)" + using V(1) by auto + qed +qed + + +lemma closed_map_preimage_neighbourhood: + "closed_map X Y f \ + image f (topspace X) \ topspace Y \ + (\U T. openin X U \ T \ topspace Y \ + {x \ topspace X. f x \ T} \ U + \ (\V. openin Y V \ T \ V \ + {x \ topspace X. f x \ V} \ U))" (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + proof (intro conjI strip) + show "f ` topspace X \ topspace Y" + by (simp add: L closed_map_imp_subset_topspace) + next + fix U T + assume UT: "openin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U" + then have "openin Y (topspace Y - f ` (topspace X - U))" + by (meson L closed_map_def closedin_def closedin_diff closedin_topspace) + then show "\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U" + using UT image_iff by auto + qed +next + assume R: ?rhs + show ?lhs + unfolding closed_map_def + proof (intro strip) + fix U assume "closedin X U" + have "{x \ topspace X. f x \ topspace Y - f ` U} \ topspace X - U" + by blast + then obtain V where V: "openin Y V" + and sub: "topspace Y - f ` U \ V" "{x \ topspace X. f x \ V} \ topspace X - U" + using R Diff_subset \closedin X U\ unfolding closedin_def + by (smt (verit, ccfv_threshold) Collect_mem_eq Collect_mono_iff) + then have "f ` U \ topspace Y - V" + using R \closedin X U\ closedin_subset by fastforce + with sub have "f ` U = topspace Y - V" + by auto + with V show "closedin Y (f ` U)" + by auto + qed +qed + +lemma closed_map_fibre_neighbourhood: + "closed_map X Y f \ + f ` (topspace X) \ topspace Y \ + (\U y. openin X U \ y \ topspace Y \ {x \ topspace X. f x = y} \ U + \ (\V. openin Y V \ y \ V \ {x \ topspace X. f x \ V} \ U))" + unfolding closed_map_preimage_neighbourhood +proof (intro conj_cong refl all_cong1) + fix U + assume "f ` topspace X \ topspace Y" + show "(\T. openin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U + \ (\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U)) + = (\y. openin X U \ y \ topspace Y \ {x \ topspace X. f x = y} \ U + \ (\V. openin Y V \ y \ V \ {x \ topspace X. f x \ V} \ U))" + (is "(\T. ?P T) \ (\y. ?Q y)") + proof + assume L [rule_format]: "\T. ?P T" + show "\y. ?Q y" + proof + fix y show "?Q y" + using L [of "{y}"] by blast + qed + next + assume R: "\y. ?Q y" + show "\T. ?P T" + proof (cases "openin X U") + case True + note [[unify_search_bound=3]] + obtain V where V: "\y. \y \ topspace Y; {x \ topspace X. f x = y} \ U\ \ + openin Y (V y) \ y \ V y \ {x \ topspace X. f x \ V y} \ U" + using R by (simp add: True) meson + show ?thesis + proof clarify + fix T + assume "openin X U" and "T \ topspace Y" and "{x \ topspace X. f x \ T} \ U" + with V show "\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U" + by (rule_tac x="\y\T. V y" in exI) fastforce + qed + qed auto + qed +qed + +lemma open_map_in_subtopology: + "openin Y S + \ (open_map X (subtopology Y S) f \ open_map X Y f \ f ` (topspace X) \ S)" + by (metis le_inf_iff open_map_def open_map_imp_subset_topspace open_map_into_subtopology openin_trans_full topspace_subtopology) + +lemma open_map_from_open_subtopology: + "\openin Y S; open_map X (subtopology Y S) f\ \ open_map X Y f" + using open_map_in_subtopology by blast + +lemma closed_map_in_subtopology: + "closedin Y S + \ closed_map X (subtopology Y S) f \ (closed_map X Y f \ f ` topspace X \ S)" + by (metis closed_map_def closed_map_imp_subset_topspace closed_map_into_subtopology + closedin_closed_subtopology closedin_subset topspace_subtopology_subset) + +lemma closed_map_from_closed_subtopology: + "\closedin Y S; closed_map X (subtopology Y S) f\ \ closed_map X Y f" + using closed_map_in_subtopology by blast + +lemma closed_map_from_composition_left: + assumes cmf: "closed_map X Z (g \ f)" and contf: "continuous_map X Y f" and fim: "f ` topspace X = topspace Y" + shows "closed_map Y Z g" + unfolding closed_map_def +proof (intro strip) + fix U assume "closedin Y U" + then have "closedin X {x \ topspace X. f x \ U}" + using contf closedin_continuous_map_preimage by blast + then have "closedin Z ((g \ f) ` {x \ topspace X. f x \ U})" + using cmf closed_map_def by blast + moreover + have "\y. y \ U \ \x \ topspace X. f x \ U \ g y = g (f x)" + by (smt (verit, ccfv_SIG) \closedin Y U\ closedin_subset fim image_iff subsetD) + then have "(g \ f) ` {x \ topspace X. f x \ U} = g`U" by auto + ultimately show "closedin Z (g ` U)" + by metis +qed + +text \identical proof as the above\ +lemma open_map_from_composition_left: + assumes cmf: "open_map X Z (g \ f)" and contf: "continuous_map X Y f" and fim: "f ` topspace X = topspace Y" + shows "open_map Y Z g" + unfolding open_map_def +proof (intro strip) + fix U assume "openin Y U" + then have "openin X {x \ topspace X. f x \ U}" + using contf openin_continuous_map_preimage by blast + then have "openin Z ((g \ f) ` {x \ topspace X. f x \ U})" + using cmf open_map_def by blast + moreover + have "\y. y \ U \ \x \ topspace X. f x \ U \ g y = g (f x)" + by (smt (verit, ccfv_SIG) \openin Y U\ openin_subset fim image_iff subsetD) + then have "(g \ f) ` {x \ topspace X. f x \ U} = g`U" by auto + ultimately show "openin Z (g ` U)" + by metis +qed + +lemma closed_map_from_composition_right: + assumes cmf: "closed_map X Z (g \ f)" "f ` topspace X \ topspace Y" "continuous_map Y Z g" "inj_on g (topspace Y)" + shows "closed_map X Y f" + unfolding closed_map_def +proof (intro strip) + fix C assume "closedin X C" + have "\y c. \y \ topspace Y; g y = g (f c); c \ C\ \ y \ f ` C" + using \closedin X C\ assms closedin_subset inj_onD by fastforce + then + have "f ` C = {x \ topspace Y. g x \ (g \ f) ` C}" + using \closedin X C\ assms(2) closedin_subset by fastforce + moreover have "closedin Z ((g \ f) ` C)" + using \closedin X C\ cmf closed_map_def by blast + ultimately show "closedin Y (f ` C)" + using assms(3) closedin_continuous_map_preimage by fastforce +qed + +text \identical proof as the above\ +lemma open_map_from_composition_right: + assumes "open_map X Z (g \ f)" "f ` topspace X \ topspace Y" "continuous_map Y Z g" "inj_on g (topspace Y)" + shows "open_map X Y f" + unfolding open_map_def +proof (intro strip) + fix C assume "openin X C" + have "\y c. \y \ topspace Y; g y = g (f c); c \ C\ \ y \ f ` C" + using \openin X C\ assms openin_subset inj_onD by fastforce + then + have "f ` C = {x \ topspace Y. g x \ (g \ f) ` C}" + using \openin X C\ assms(2) openin_subset by fastforce + moreover have "openin Z ((g \ f) ` C)" + using \openin X C\ assms(1) open_map_def by blast + ultimately show "openin Y (f ` C)" + using assms(3) openin_continuous_map_preimage by fastforce +qed + subsection\Quotient maps\ definition quotient_map where @@ -2151,6 +2414,84 @@ qed qed +lemma quotient_map_saturated_closed: + "quotient_map X Y f \ + continuous_map X Y f \ f ` (topspace X) = topspace Y \ + (\U. closedin X U \ {x \ topspace X. f x \ f ` U} \ U \ closedin Y (f ` U))" + (is "?lhs = ?rhs") +proof + assume L: ?lhs + then obtain fim: "f ` topspace X = topspace Y" + and Y: "\U. U \ topspace Y \ closedin Y U = closedin X {x \ topspace X. f x \ U}" + by (simp add: quotient_map_closedin) + show ?rhs + proof (intro conjI allI impI) + show "continuous_map X Y f" + by (simp add: L quotient_imp_continuous_map) + show "f ` topspace X = topspace Y" + by (simp add: fim) + next + fix U :: "'a set" + assume U: "closedin X U \ {x \ topspace X. f x \ f ` U} \ U" + then have sub: "f ` U \ topspace Y" and eq: "{x \ topspace X. f x \ f ` U} = U" + using fim closedin_subset by fastforce+ + show "closedin Y (f ` U)" + by (simp add: sub Y eq U) + qed +next + assume ?rhs + then obtain YX: "\U. closedin Y U \ closedin X {x \ topspace X. f x \ U}" + and fim: "f ` topspace X = topspace Y" + and XY: "\U. \closedin X U; {x \ topspace X. f x \ f ` U} \ U\ \ closedin Y (f ` U)" + by (simp add: continuous_map_closedin) + show ?lhs + proof (simp add: quotient_map_closedin fim, intro allI impI iffI) + fix U :: "'b set" + assume "U \ topspace Y" and X: "closedin X {x \ topspace X. f x \ U}" + have feq: "f ` {x \ topspace X. f x \ U} = U" + using \U \ topspace Y\ fim by auto + show "closedin Y U" + using XY [OF X] by (simp add: feq) + next + fix U :: "'b set" + assume "U \ topspace Y" and Y: "closedin Y U" + show "closedin X {x \ topspace X. f x \ U}" + by (metis YX [OF Y]) + qed +qed + +lemma quotient_map_onto_image: + assumes "f ` topspace X \ topspace Y" and U: "\U. U \ topspace Y \ openin X {x \ topspace X. f x \ U} = openin Y U" + shows "quotient_map X (subtopology Y (f ` topspace X)) f" + unfolding quotient_map_def topspace_subtopology +proof (intro conjI strip) + fix U + assume "U \ topspace Y \ f ` topspace X" + with U have "openin X {x \ topspace X. f x \ U} \ \x. openin Y x \ U = f ` topspace X \ x" + by fastforce + moreover have "\x. openin Y x \ U = f ` topspace X \ x \ openin X {x \ topspace X. f x \ U}" + by (metis (mono_tags, lifting) Collect_cong IntE IntI U image_eqI openin_subset) + ultimately show "openin X {x \ topspace X. f x \ U} = openin (subtopology Y (f ` topspace X)) U" + by (force simp: openin_subtopology_alt image_iff) +qed (use assms in auto) + +lemma quotient_map_lift_exists: + assumes f: "quotient_map X Y f" and h: "continuous_map X Z h" + and "\x y. \x \ topspace X; y \ topspace X; f x = f y\ \ h x = h y" + obtains g where "continuous_map Y Z g" "g ` topspace Y = h ` topspace X" + "\x. x \ topspace X \ g(f x) = h x" +proof - + obtain g where g: "\x. x \ topspace X \ h x = g(f x)" + using function_factors_left_gen[of "\x. x \ topspace X" f h] assms by blast + show ?thesis + proof + show "g ` topspace Y = h ` topspace X" + using f g by (force dest!: quotient_imp_surjective_map) + show "continuous_map Y Z g" + by (smt (verit) f g h continuous_compose_quotient_map_eq continuous_map_eq o_def) + qed (simp add: g) +qed + subsection\ Separated Sets\ definition separatedin :: "'a topology \ 'a set \ 'a set \ bool" @@ -2245,6 +2586,11 @@ unfolding openin_closedin_eq topspace_subtopology separation_closedin_Un_gen disjnt_def by (auto simp: Diff_triv Int_commute Un_Diff inf_absorb1 topspace_def) +lemma separatedin_full: + "S \ T = topspace X + \ separatedin X S T \ disjnt S T \ closedin X S \ openin X S \ closedin X T \ openin X T" + by (metis separatedin_open_sets separation_closedin_Un_gen separation_openin_Un_gen subtopology_topspace) + subsection\Homeomorphisms\ text\(1-way and 2-way versions may be useful in places)\ @@ -2766,6 +3112,10 @@ by (intro conj_cong arg_cong [where f=Not] ex_cong1; blast dest!: openin_subset) qed +lemma connectedinD: + "\connectedin X S; openin X E1; openin X E2; S \ E1 \ E2; E1 \ E2 \ S = {}; E1 \ S \ {}; E2 \ S \ {}\ \ False" + by (meson connectedin) + lemma connectedin_iff_connected [simp]: "connectedin euclidean S \ connected S" by (simp add: connected_def connectedin) @@ -2966,7 +3316,7 @@ by (metis closure_of_eq) qed -lemma connectedin_inter_frontier_of: +lemma connectedin_Int_frontier_of: assumes "connectedin X S" "S \ T \ {}" "S - T \ {}" shows "S \ X frontier_of T \ {}" proof - @@ -3068,7 +3418,7 @@ moreover have "connectedin (discrete_topology U) S \ (\a. S = {a})" proof show "connectedin (discrete_topology U) S \ \a. S = {a}" - using False connectedin_inter_frontier_of insert_Diff by fastforce + using False connectedin_Int_frontier_of insert_Diff by fastforce qed (use True in auto) ultimately show ?thesis by auto @@ -3360,7 +3710,7 @@ corollary compact_space_imp_nest: fixes C :: "nat \ 'a set" assumes "compact_space X" and clo: "\n. closedin X (C n)" - and ne: "\n. C n \ {}" and inc: "\m n. m \ n \ C n \ C m" + and ne: "\n. C n \ {}" and dec: "decseq C" shows "(\n. C n) \ {}" proof - let ?\ = "range (\n. \m \ n. C m)" @@ -3370,8 +3720,8 @@ proof - obtain n where "\k. k \ K \ k \ n" using Max.coboundedI \finite K\ by blast - with inc have "C n \ (\n\K. \m \ n. C m)" - by blast + with dec have "C n \ (\n\K. \m \ n. C m)" + unfolding decseq_def by blast with ne [of n] show ?thesis by blast qed @@ -3556,6 +3906,10 @@ unfolding closed_map_def by (auto simp: closedin_closed_subtopology embedding_map_def homeomorphic_map_closedness_eq) +lemma embedding_imp_closed_map_eq: + "embedding_map X Y f \ (closed_map X Y f \ closedin Y (f ` topspace X))" + using closed_map_def embedding_imp_closed_map by blast + subsection\Retraction and section maps\ @@ -3688,6 +4042,7 @@ shows "continuous_map euclidean (top_of_set S) f" by (simp add: cont continuous_map_into_subtopology eq image_subset_iff_subset_vimage) + subsection\<^marker>\tag unimportant\ \Half-global and completely global cases\ lemma continuous_openin_preimage_gen: @@ -4131,6 +4486,188 @@ using assms continuous_on_generated_topo_iff by blast +subsection\Continuity via bases/subbases, hence upper and lower semicontinuity\ + +lemma continuous_map_into_topology_base: + assumes P: "openin Y = arbitrary union_of P" + and f: "\x. x \ topspace X \ f x \ topspace Y" + and ope: "\U. P U \ openin X {x \ topspace X. f x \ U}" + shows "continuous_map X Y f" +proof - + have *: "\\. (\t. t \ \ \ P t) \ openin X {x \ topspace X. \U\\. f x \ U}" + by (smt (verit) Ball_Collect ope mem_Collect_eq openin_subopen) + show ?thesis + using P by (auto simp: continuous_map_def arbitrary_def union_of_def intro!: f *) +qed + +lemma continuous_map_into_topology_base_eq: + assumes P: "openin Y = arbitrary union_of P" + shows + "continuous_map X Y f \ + (\x \ topspace X. f x \ topspace Y) \ (\U. P U \ openin X {x \ topspace X. f x \ U})" + (is "?lhs=?rhs") +proof + assume L: ?lhs + then have "\x. x \ topspace X \ f x \ topspace Y" + by (meson continuous_map_def) + moreover have "\U. P U \ openin X {x \ topspace X. f x \ U}" + using L assms continuous_map openin_topology_base_unique by fastforce + ultimately show ?rhs by auto +qed (simp add: assms continuous_map_into_topology_base) + +lemma continuous_map_into_topology_subbase: + fixes U P + defines "Y \ topology(arbitrary union_of (finite intersection_of P relative_to U))" + assumes f: "\x. x \ topspace X \ f x \ topspace Y" + and ope: "\U. P U \ openin X {x \ topspace X. f x \ U}" + shows "continuous_map X Y f" +proof (intro continuous_map_into_topology_base) + show "openin Y = arbitrary union_of (finite intersection_of P relative_to U)" + unfolding Y_def using istopology_subbase topology_inverse' by blast + show "openin X {x \ topspace X. f x \ V}" + if \
: "(finite intersection_of P relative_to U) V" for V + proof - + define finv where "finv \ \V. {x \ topspace X. f x \ V}" + obtain \ where \: "finite \" "\V. V \ \ \ P V" + "{x \ topspace X. f x \ V} = (\V \ insert U \. finv V)" + using \
by (fastforce simp: finv_def intersection_of_def relative_to_def) + show ?thesis + unfolding \ + proof (intro openin_Inter ope) + have U: "U = topspace Y" + unfolding Y_def using topspace_subbase by fastforce + fix V + assume V: "V \ finv ` insert U \" + with U f have "openin X {x \ topspace X. f x \ U}" + by (auto simp: openin_subopen [of X "Collect _"]) + then show "openin X V" + using V \(2) ope by (fastforce simp: finv_def) + qed (use \finite \\ in auto) + qed +qed (use f in auto) + +lemma continuous_map_into_topology_subbase_eq: + assumes "Y = topology(arbitrary union_of (finite intersection_of P relative_to U))" + shows + "continuous_map X Y f \ + (\x \ topspace X. f x \ topspace Y) \ (\U. P U \ openin X {x \ topspace X. f x \ U})" + (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + proof (intro conjI strip) + show "\x. x \ topspace X \ f x \ topspace Y" + using L continuous_map_def by fastforce + fix V + assume "P V" + have "U = topspace Y" + unfolding assms using topspace_subbase by fastforce + then have eq: "{x \ topspace X. f x \ V} = {x \ topspace X. f x \ U \ V}" + using L by (auto simp: continuous_map) + have "openin Y (U \ V)" + unfolding assms openin_subbase + by (meson \P V\ arbitrary_union_of_inc finite_intersection_of_inc relative_to_inc) + show "openin X {x \ topspace X. f x \ V}" + using L \openin Y (U \ V)\ continuous_map eq by fastforce + qed +next + show "?rhs \ ?lhs" + unfolding assms + by (intro continuous_map_into_topology_subbase) auto +qed + +lemma subbase_subtopology_euclidean: + fixes U :: "'a::order_topology set" + shows + "topology + (arbitrary union_of + (finite intersection_of (\x. x \ range greaterThan \ range lessThan) relative_to U)) + = subtopology euclidean U" +proof - + have "\V. (finite intersection_of (\x. x \ range greaterThan \ x \ range lessThan)) V \ x \ V \ V \ W" + if "open W" "x \ W" for W and x::'a + using \open W\ [unfolded open_generated_order] \x \ W\ + proof (induct rule: generate_topology.induct) + case UNIV + then show ?case + using finite_intersection_of_empty by blast + next + case (Int a b) + then show ?case + by (meson Int_iff finite_intersection_of_Int inf_mono) + next + case (UN K) + then show ?case + by (meson Union_iff subset_iff) + next + case (Basis s) + then show ?case + by (metis (no_types, lifting) Un_iff finite_intersection_of_inc order_refl) + qed + moreover + have "\V::'a set. (finite intersection_of (\x. x \ range greaterThan \ x \ range lessThan)) V \ open V" + by (force simp: intersection_of_def subset_iff) + ultimately have *: "openin (euclidean::'a topology) = + (arbitrary union_of (finite intersection_of (\x. x \ range greaterThan \ x \ range lessThan)))" + by (smt (verit, best) openin_topology_base_unique open_openin) + show ?thesis + unfolding subtopology_def arbitrary_union_of_relative_to [symmetric] + apply (simp add: relative_to_def flip: *) + by (metis Int_commute) +qed + +lemma continuous_map_upper_lower_semicontinuous_lt_gen: + fixes U :: "'a::order_topology set" + shows "continuous_map X (subtopology euclidean U) f \ + (\x \ topspace X. f x \ U) \ + (\a. openin X {x \ topspace X. f x > a}) \ + (\a. openin X {x \ topspace X. f x < a})" + by (auto simp: continuous_map_into_topology_subbase_eq [OF subbase_subtopology_euclidean [symmetric, of U]] + greaterThan_def lessThan_def image_iff simp flip: all_simps) + +lemma continuous_map_upper_lower_semicontinuous_lt: + fixes f :: "'a \ 'b::order_topology" + shows "continuous_map X euclidean f \ + (\a. openin X {x \ topspace X. f x > a}) \ + (\a. openin X {x \ topspace X. f x < a})" + using continuous_map_upper_lower_semicontinuous_lt_gen [where U=UNIV] + by auto + +lemma Int_Collect_imp_eq: "A \ {x. x\A \ P x} = {x\A. P x}" + by blast + +lemma continuous_map_upper_lower_semicontinuous_le_gen: + shows + "continuous_map X (subtopology euclideanreal U) f \ + (\x \ topspace X. f x \ U) \ + (\a. closedin X {x \ topspace X. f x \ a}) \ + (\a. closedin X {x \ topspace X. f x \ a})" + unfolding continuous_map_upper_lower_semicontinuous_lt_gen + by (auto simp: closedin_def Diff_eq Compl_eq not_le Int_Collect_imp_eq) + +lemma continuous_map_upper_lower_semicontinuous_le: + "continuous_map X euclideanreal f \ + (\a. closedin X {x \ topspace X. f x \ a}) \ + (\a. closedin X {x \ topspace X. f x \ a})" + using continuous_map_upper_lower_semicontinuous_le_gen [where U=UNIV] + by auto + +lemma continuous_map_upper_lower_semicontinuous_lte_gen: + "continuous_map X (subtopology euclideanreal U) f \ + (\x \ topspace X. f x \ U) \ + (\a. openin X {x \ topspace X. f x < a}) \ + (\a. closedin X {x \ topspace X. f x \ a})" + unfolding continuous_map_upper_lower_semicontinuous_lt_gen + by (auto simp: closedin_def Diff_eq Compl_eq not_le Int_Collect_imp_eq) + +lemma continuous_map_upper_lower_semicontinuous_lte: + "continuous_map X euclideanreal f \ + (\a. openin X {x \ topspace X. f x < a}) \ + (\a. closedin X {x \ topspace X. f x \ a})" + using continuous_map_upper_lower_semicontinuous_lte_gen [where U=UNIV] + by auto + + subsection\<^marker>\tag important\ \Pullback topology\ text \Pulling back a topology by map gives again a topology. \subtopology\ is diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Abstract_Topology_2.thy --- a/src/HOL/Analysis/Abstract_Topology_2.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Analysis/Abstract_Topology_2.thy Sun May 07 22:51:23 2023 +0200 @@ -347,7 +347,7 @@ continuous_on S f; continuous_on T f\ \ continuous_on (S \ T) f" unfolding continuous_on closedin_limpt - by (metis Lim_trivial_limit Lim_within_union Un_iff trivial_limit_within) + by (metis Lim_trivial_limit Lim_within_Un Un_iff trivial_limit_within) lemma continuous_on_cases_local: "\closedin (top_of_set (S \ T)) S; closedin (top_of_set (S \ T)) T; @@ -1138,6 +1138,20 @@ by (force simp: retract_of_space_def) qed +lemma retract_of_space_trans: + assumes "S retract_of_space X" "T retract_of_space subtopology X S" + shows "T retract_of_space X" + using assms + apply (simp add: retract_of_space_retraction_maps) + by (metis id_comp inf.absorb_iff2 retraction_maps_compose subtopology_subtopology) + +lemma retract_of_space_subtopology: + assumes "S retract_of_space X" "S \ U" + shows "S retract_of_space subtopology X U" + using assms + apply (clarsimp simp: retract_of_space_def) + by (metis continuous_map_from_subtopology inf.absorb2 subtopology_subtopology) + lemma retract_of_space_clopen: assumes "openin X S" "closedin X S" "S = {} \ topspace X = {}" shows "S retract_of_space X" @@ -1194,6 +1208,13 @@ using embedding_map_imp_homeomorphic_space homeomorphic_space_sym section_imp_embedding_map section_map_def by blast +lemma hereditary_imp_retractive_property: + assumes "\X S. P X \ P(subtopology X S)" + "\X X'. X homeomorphic_space X' \ (P X \ Q X')" + assumes "retraction_map X X' r" "P X" + shows "Q X'" + by (meson assms retraction_map_def retraction_maps_section_image2) + subsection\Paths and path-connectedness\ definition pathin :: "'a topology \ (real \ 'a) \ bool" where @@ -1591,4 +1612,31 @@ connected_component_of_set X x = connected_component_of_set X y" by (meson connected_component_of_nonoverlap) +subsection\Combining theorems for continuous functions into the reals\ + +text \The homeomorphism between the real line and the open interval $(-1,1)$\ + +lemma continuous_map_real_shrink: + "continuous_map euclideanreal (top_of_set {-1<..<1}) (\x. x / (1 + \x\))" +proof - + have "continuous_on UNIV (\x::real. x / (1 + \x\))" + by (intro continuous_intros) auto + then show ?thesis + by (auto simp: continuous_map_in_subtopology divide_simps) +qed + +lemma continuous_on_real_grow: + "continuous_on {-1<..<1} (\x::real. x / (1 - \x\))" + by (intro continuous_intros) auto + +lemma real_grow_shrink: + fixes x::real + shows "x / (1 + \x\) / (1 - \x / (1 + \x\)\) = x" + by (force simp: divide_simps) + +lemma homeomorphic_maps_real_shrink: + "homeomorphic_maps euclideanreal (subtopology euclideanreal {-1<..<1}) + (\x. x / (1 + \x\)) (\y. y / (1 - \y\))" + by (force simp: homeomorphic_maps_def continuous_map_real_shrink continuous_on_real_grow divide_simps) + end \ No newline at end of file diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Analysis.thy --- a/src/HOL/Analysis/Analysis.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Analysis/Analysis.thy Sun May 07 22:51:23 2023 +0200 @@ -4,6 +4,9 @@ Convex Determinants (* Topology *) + FSigma + Sum_Topology + Abstract_Topological_Spaces Connected Abstract_Limits Isolated diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Arcwise_Connected.thy --- a/src/HOL/Analysis/Arcwise_Connected.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Analysis/Arcwise_Connected.thy Sun May 07 22:51:23 2023 +0200 @@ -533,6 +533,216 @@ by (auto simp: field_split_simps) +lemma padic_rational_approximation_straddle: + assumes "\ > 0" "p > 1" + obtains n q r + where "of_int q / p^n < x" "x < of_int r / p^n" "\q / p^n - r / p^n \ < \" +proof - + obtain n where n: "2 / \ < p ^ n" + using \p>1\ real_arch_pow by blast + define q where "q \ \p ^ n * x\ - 1" + show thesis + proof + show "q / p ^ n < x" "x < real_of_int (q+2) / p ^ n" + using assms by (simp_all add: q_def divide_simps floor_less_cancel mult.commute) + show "\q / p ^ n - real_of_int (q+2) / p ^ n\ < \" + using assms n by (simp add: q_def divide_simps mult.commute) + qed +qed + +lemma padic_rational_approximation_straddle_pos: + assumes "\ > 0" "p > 1" "x > 0" + obtains n q r + where "of_nat q / p^n < x" "x < of_nat r / p^n" "\q / p^n - r / p^n \ < \" +proof - + obtain n q r + where *: "of_int q / p^n < x" "x < of_int r / p^n" "\q / p^n - r / p^n \ < \" + using padic_rational_approximation_straddle assms by metis + then have "r \ 0" + using assms by (smt (verit, best) divide_nonpos_pos of_int_0_le_iff zero_less_power) + show thesis + proof + show "real (max 0 (nat q)) / p ^ n < x" + using * by (metis assms(3) div_0 max_nat.left_neutral nat_eq_iff2 of_nat_0 of_nat_nat) + show "x < real (nat r) / p ^ n" + using \r \ 0\ * by force + show "\real (max 0 (nat q)) / p ^ n - real (nat r) / p ^ n\ < \" + using * assms by (simp add: divide_simps) + qed +qed + +lemma padic_rational_approximation_straddle_pos_le: + assumes "\ > 0" "p > 1" "x \ 0" + obtains n q r + where "of_nat q / p^n \ x" "x < of_nat r / p^n" "\q / p^n - r / p^n \ < \" +proof - + obtain n q r + where *: "of_int q / p^n < x" "x < of_int r / p^n" "\q / p^n - r / p^n \ < \" + using padic_rational_approximation_straddle assms by metis + then have "r \ 0" + using assms by (smt (verit, best) divide_nonpos_pos of_int_0_le_iff zero_less_power) + show thesis + proof + show "real (max 0 (nat q)) / p ^ n \ x" + using * assms(3) nle_le by fastforce + show "x < real (nat r) / p ^ n" + using \r \ 0\ * by force + show "\real (max 0 (nat q)) / p ^ n - real (nat r) / p ^ n\ < \" + using * assms by (simp add: divide_simps) + qed +qed + + +subsubsection \Definition by recursion on dyadic rationals in [0,1]\ + +lemma recursion_on_dyadic_fractions: + assumes base: "R a b" + and step: "\x y. R x y \ \z. R x z \ R z y" and trans: "\x y z. \R x y; R y z\ \ R x z" + shows "\f :: real \ 'a. f 0 = a \ f 1 = b \ + (\x \ dyadics \ {0..1}. \y \ dyadics \ {0..1}. x < y \ R (f x) (f y))" +proof - + obtain mid where mid: "R x y \ R x (mid x y)" "R x y \ R (mid x y) y" for x y + using step by metis + define g where "g \ rec_nat (\k. if k = 0 then a else b) (\n r k. if even k then r (k div 2) else mid (r ((k - 1) div 2)) (r ((Suc k) div 2)))" + have g0 [simp]: "g 0 = (\k. if k = 0 then a else b)" + by (simp add: g_def) + have gSuc [simp]: "\n. g(Suc n) = (\k. if even k then g n (k div 2) else mid (g n ((k - 1) div 2)) (g n ((Suc k) div 2)))" + by (auto simp: g_def) + have g_eq_g: "2 ^ d * k = k' \ g n k = g (n + d) k'" for n d k k' + by (induction d arbitrary: k k') auto + have "g n k = g n' k'" if "real k / 2^n = real k' / 2^n'" "n' \ n" for k n k' n' + proof - + have "real k = real k' * 2 ^ (n-n')" + using that by (simp add: power_diff divide_simps) + then have "k = k' * 2 ^ (n-n')" + using of_nat_eq_iff by fastforce + with g_eq_g show ?thesis + by (metis le_add_diff_inverse mult.commute that(2)) + qed + then have g_eq_g: "g n k = g n' k'" if "real k / 2 ^ n = real k' / 2 ^ n'" for k n k' n' + by (metis nat_le_linear that) + then obtain f where "(\(k,n). g n k) = f \ (\(k,n). k / 2 ^ n)" + using function_factors_left by (smt (verit, del_insts) case_prod_beta') + then have f_eq_g: "\k n. f(real k / 2 ^ n) = g n k" + by (simp add: fun_eq_iff) + show ?thesis + proof (intro exI conjI strip) + show "f 0 = a" + by (metis f_eq_g g0 div_0 of_nat_0) + show "f 1 = b" + by (metis f_eq_g g0 div_by_1 of_nat_1_eq_iff power_0 zero_neq_one) + show "R (f x) (f y)" + if x: "x \ dyadics \ {0..1}" and y: "y \ dyadics \ {0..1}" and "x < y" for x y + proof - + obtain n1 k1 where xeq: "x = real k1 / 2^n1" "k1 \ 2^n1" + using x by (auto simp: dyadics_def) + obtain n2 k2 where yeq: "y = real k2 / 2^n2" "k2 \ 2^n2" + using y by (auto simp: dyadics_def) + have xcommon: "x = real(2^n2 * k1) / 2 ^ (n1+n2)" + using xeq by (simp add: power_add) + have ycommon: "y = real(2^n1 * k2) / 2 ^ (n1+n2)" + using yeq by (simp add: power_add) + have *: "R (g n j) (g n k)" if "j < k" "k \ 2^n" for n j k + using that + proof (induction n arbitrary: j k) + case 0 + then show ?case + by (simp add: base) + next + case (Suc n) + show ?case + proof (cases "even j") + case True + then obtain a where [simp]: "j = 2*a" + by blast + show ?thesis + proof (cases "even k") + case True + with Suc show ?thesis + by (auto elim!: evenE) + next + case False + then obtain b where [simp]: "k = Suc (2*b)" + using oddE by fastforce + show ?thesis + using Suc + apply simp + by (smt (verit, ccfv_SIG) less_Suc_eq linorder_not_le local.trans mid(1) nat_mult_less_cancel1 pos2) + qed + next + case False + then obtain a where [simp]: "j = Suc (2*a)" + using oddE by fastforce + show ?thesis + proof (cases "even k") + case True + then obtain b where [simp]: "k = 2*b" + by blast + show ?thesis + using Suc + apply simp + by (smt (verit, ccfv_SIG) Suc_leI Suc_lessD le_trans lessI linorder_neqE_nat linorder_not_le local.trans mid(2) nat_mult_less_cancel1 pos2) + next + case False + then obtain b where [simp]: "k = Suc (2*b)" + using oddE by fastforce + show ?thesis + using Suc + apply simp + by (smt (verit) Suc_leI le_trans lessI less_or_eq_imp_le linorder_neqE_nat linorder_not_le local.trans mid(1) mid(2) nat_mult_less_cancel1 pos2) + qed + qed + qed + show ?thesis + unfolding xcommon ycommon f_eq_g + proof (rule *) + show "2 ^ n2 * k1 < 2 ^ n1 * k2" + using of_nat_less_iff \x < y\ by (fastforce simp: xeq yeq field_simps) + show "2 ^ n1 * k2 \ 2 ^ (n1 + n2)" + by (simp add: power_add yeq) + qed + qed + qed +qed + +lemma dyadics_add: + assumes "x \ dyadics" "y \ dyadics" + shows "x+y \ dyadics" +proof - + obtain i j m n where x: "x = of_nat i / 2 ^ m" and y: "y = of_nat j / 2 ^ n" + using assms by (auto simp: dyadics_def) + have xcommon: "x = of_nat(2^n * i) / 2 ^ (m+n)" + using x by (simp add: power_add) + moreover + have ycommon: "y = of_nat(2^m * j) / 2 ^ (m+n)" + using y by (simp add: power_add) + ultimately have "x+y = (of_nat(2^n * i + 2^m * j)) / 2 ^ (m+n)" + by (simp add: field_simps) + then show ?thesis + unfolding dyadics_def by blast +qed + +lemma dyadics_diff: + fixes x :: "'a::linordered_field" + assumes "x \ dyadics" "y \ dyadics" "y \ x" + shows "x-y \ dyadics" +proof - + obtain i j m n where x: "x = of_nat i / 2 ^ m" and y: "y = of_nat j / 2 ^ n" + using assms by (auto simp: dyadics_def) + have j_le_i: "j * 2 ^ m \ i * 2 ^ n" + using of_nat_le_iff \y \ x\ unfolding x y by (fastforce simp add: divide_simps) + have xcommon: "x = of_nat(2^n * i) / 2 ^ (m+n)" + using x by (simp add: power_add) + moreover + have ycommon: "y = of_nat(2^m * j) / 2 ^ (m+n)" + using y by (simp add: power_add) + ultimately have "x-y = (of_nat(2^n * i - 2^m * j)) / 2 ^ (m+n)" + by (simp add: xcommon ycommon field_simps j_le_i of_nat_diff) + then show ?thesis + unfolding dyadics_def by blast +qed + + theorem homeomorphic_monotone_image_interval: fixes f :: "real \ 'a::{real_normed_vector,complete_space}" diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Sun May 07 22:51:23 2023 +0200 @@ -175,39 +175,6 @@ finally show "norm (A *v x) \ CARD('m) * real (CARD('n)) * B * norm x" . qed -lemma rational_approximation: - assumes "e > 0" - obtains r::real where "r \ \" "\r - x\ < e" - using Rats_dense_in_real [of "x - e/2" "x + e/2"] assms by auto - -lemma Rats_closure_real: "closure \ = (UNIV::real set)" -proof - - have "\x::real. x \ closure \" - by (metis closure_approachable dist_real_def rational_approximation) - then show ?thesis by auto -qed - -proposition matrix_rational_approximation: - fixes A :: "real^'n^'m" - assumes "e > 0" - obtains B where "\i j. B$i$j \ \" "onorm(\x. (A - B) *v x) < e" -proof - - have "\i j. \q \ \. \q - A $ i $ j\ < e / (2 * CARD('m) * CARD('n))" - using assms by (force intro: rational_approximation [of "e / (2 * CARD('m) * CARD('n))"]) - then obtain B where B: "\i j. B$i$j \ \" and Bclo: "\i j. \B$i$j - A $ i $ j\ < e / (2 * CARD('m) * CARD('n))" - by (auto simp: lambda_skolem Bex_def) - show ?thesis - proof - have "onorm ((*v) (A - B)) \ real CARD('m) * real CARD('n) * - (e / (2 * real CARD('m) * real CARD('n)))" - apply (rule onorm_le_matrix_component) - using Bclo by (simp add: abs_minus_commute less_imp_le) - also have "\ < e" - using \0 < e\ by (simp add: field_split_simps) - finally show "onorm ((*v) (A - B)) < e" . - qed (use B in auto) -qed - lemma vector_sub_project_orthogonal_cart: "(b::real^'n) \ (x - ((b \ x) / (b \ b)) *s b) = 0" unfolding inner_simps scalar_mult_eq_scaleR by auto @@ -478,6 +445,42 @@ qed *) +subsection\Arbitrarily good rational approximations\ + +lemma rational_approximation: + assumes "e > 0" + obtains r::real where "r \ \" "\r - x\ < e" + using Rats_dense_in_real [of "x - e/2" "x + e/2"] assms by auto + +lemma Rats_closure_real: "closure \ = (UNIV::real set)" +proof - + have "\x::real. x \ closure \" + by (metis closure_approachable dist_real_def rational_approximation) + then show ?thesis by auto +qed + +proposition matrix_rational_approximation: + fixes A :: "real^'n^'m" + assumes "e > 0" + obtains B where "\i j. B$i$j \ \" "onorm(\x. (A - B) *v x) < e" +proof - + have "\i j. \q \ \. \q - A $ i $ j\ < e / (2 * CARD('m) * CARD('n))" + using assms by (force intro: rational_approximation [of "e / (2 * CARD('m) * CARD('n))"]) + then obtain B where B: "\i j. B$i$j \ \" and Bclo: "\i j. \B$i$j - A $ i $ j\ < e / (2 * CARD('m) * CARD('n))" + by (auto simp: lambda_skolem Bex_def) + show ?thesis + proof + have "onorm ((*v) (A - B)) \ real CARD('m) * real CARD('n) * + (e / (2 * real CARD('m) * real CARD('n)))" + apply (rule onorm_le_matrix_component) + using Bclo by (simp add: abs_minus_commute less_imp_le) + also have "\ < e" + using \0 < e\ by (simp add: field_split_simps) + finally show "onorm ((*v) (A - B)) < e" . + qed (use B in auto) +qed + + subsection "Derivative" definition\<^marker>\tag important\ "jacobian f net = matrix(frechet_derivative f net)" @@ -547,8 +550,7 @@ lemma interval_split_cart: "{a..b::real^'n} \ {x. x$k \ c} = {a .. (\ i. if i = k then min (b$k) c else b$i)}" "cbox a b \ {x. x$k \ c} = {(\ i. if i = k then max (a$k) c else a$i) .. b}" - apply (rule_tac[!] set_eqI) - unfolding Int_iff mem_box_cart mem_Collect_eq interval_cbox_cart + unfolding Int_iff mem_box_cart mem_Collect_eq interval_cbox_cart set_eq_iff unfolding vec_lambda_beta by auto diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Analysis/Convex_Euclidean_Space.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Sun May 07 22:51:23 2023 +0200 @@ -1716,6 +1716,11 @@ shows "connected S \ convex S" by (metis is_interval_convex convex_connected is_interval_connected_1) +lemma connected_space_iff_is_interval_1 [iff]: + fixes S :: "real set" + shows "connected_space (top_of_set S) \ is_interval S" + using connectedin_topspace is_interval_connected_1 by force + lemma connected_convex_1_gen: fixes S :: "'a :: euclidean_space set" assumes "DIM('a) = 1" diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Elementary_Topology.thy --- a/src/HOL/Analysis/Elementary_Topology.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Analysis/Elementary_Topology.thy Sun May 07 22:51:23 2023 +0200 @@ -1323,32 +1323,6 @@ lemma at_within_eq_bot_iff: "at c within A = bot \ c \ closure (A - {c})" using not_trivial_limit_within[of c A] by blast -text \Some property holds "sufficiently close" to the limit point.\ - -lemma trivial_limit_eventually: "trivial_limit net \ eventually P net" - by simp - -lemma trivial_limit_eq: "trivial_limit net \ (\P. eventually P net)" - by (simp add: filter_eq_iff) - -lemma Lim_topological: - "(f \ l) net \ - trivial_limit net \ (\S. open S \ l \ S \ eventually (\x. f x \ S) net)" - unfolding tendsto_def trivial_limit_eq by auto - -lemma eventually_within_Un: - "eventually P (at x within (s \ t)) \ - eventually P (at x within s) \ eventually P (at x within t)" - unfolding eventually_at_filter - by (auto elim!: eventually_rev_mp) - -lemma Lim_within_union: - "(f \ l) (at x within (s \ t)) \ - (f \ l) (at x within s) \ (f \ l) (at x within t)" - unfolding tendsto_def - by (auto simp: eventually_within_Un) - - subsection \Limits\ text \The expected monotonicity property.\ diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/FSigma.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/FSigma.thy Sun May 07 22:51:23 2023 +0200 @@ -0,0 +1,236 @@ +(* Author: L C Paulson, University of Cambridge [ported from HOL Light, metric.ml] *) + +section \$F$-Sigma and $G$-Delta sets in a Topological Space\ + +text \An $F$-sigma set is a countable union of closed sets; a $G$-delta set is the dual.\ + +theory FSigma + imports Abstract_Topology +begin + + +definition fsigma_in + where "fsigma_in X \ countable union_of closedin X" + +definition gdelta_in + where "gdelta_in X \ (countable intersection_of openin X) relative_to topspace X" + +lemma fsigma_in_ascending: + "fsigma_in X S \ (\C. (\n. closedin X (C n)) \ (\n. C n \ C(Suc n)) \ \ (range C) = S)" + unfolding fsigma_in_def + by (metis closedin_Un countable_union_of_ascending closedin_empty) + +lemma gdelta_in_alt: + "gdelta_in X S \ S \ topspace X \ (countable intersection_of openin X) S" +proof - + have "(countable intersection_of openin X) (topspace X)" + by (simp add: countable_intersection_of_inc) + then show ?thesis + unfolding gdelta_in_def + by (metis countable_intersection_of_inter relative_to_def relative_to_imp_subset relative_to_subset) +qed + +lemma fsigma_in_subset: "fsigma_in X S \ S \ topspace X" + using closedin_subset by (fastforce simp: fsigma_in_def union_of_def subset_iff) + +lemma gdelta_in_subset: "gdelta_in X S \ S \ topspace X" + by (simp add: gdelta_in_alt) + +lemma closed_imp_fsigma_in: "closedin X S \ fsigma_in X S" + by (simp add: countable_union_of_inc fsigma_in_def) + +lemma open_imp_gdelta_in: "openin X S \ gdelta_in X S" + by (simp add: countable_intersection_of_inc gdelta_in_alt openin_subset) + +lemma fsigma_in_empty [iff]: "fsigma_in X {}" + by (simp add: closed_imp_fsigma_in) + +lemma gdelta_in_empty [iff]: "gdelta_in X {}" + by (simp add: open_imp_gdelta_in) + +lemma fsigma_in_topspace [iff]: "fsigma_in X (topspace X)" + by (simp add: closed_imp_fsigma_in) + +lemma gdelta_in_topspace [iff]: "gdelta_in X (topspace X)" + by (simp add: open_imp_gdelta_in) + +lemma fsigma_in_Union: + "\countable T; \S. S \ T \ fsigma_in X S\ \ fsigma_in X (\ T)" + by (simp add: countable_union_of_Union fsigma_in_def) + +lemma fsigma_in_Un: + "\fsigma_in X S; fsigma_in X T\ \ fsigma_in X (S \ T)" + by (simp add: countable_union_of_Un fsigma_in_def) + +lemma fsigma_in_Int: + "\fsigma_in X S; fsigma_in X T\ \ fsigma_in X (S \ T)" + by (simp add: closedin_Int countable_union_of_Int fsigma_in_def) + +lemma gdelta_in_Inter: + "\countable T; T\{}; \S. S \ T \ gdelta_in X S\ \ gdelta_in X (\ T)" + by (simp add: Inf_less_eq countable_intersection_of_Inter gdelta_in_alt) + +lemma gdelta_in_Int: + "\gdelta_in X S; gdelta_in X T\ \ gdelta_in X (S \ T)" + by (simp add: countable_intersection_of_inter gdelta_in_alt le_infI2) + +lemma gdelta_in_Un: + "\gdelta_in X S; gdelta_in X T\ \ gdelta_in X (S \ T)" + by (simp add: countable_intersection_of_union gdelta_in_alt openin_Un) + +lemma fsigma_in_diff: + assumes "fsigma_in X S" "gdelta_in X T" + shows "fsigma_in X (S - T)" +proof - + have [simp]: "S - (S \ T) = S - T" for S T :: "'a set" + by auto + have [simp]: "topspace X - \\ = (\T\\. topspace X - T)" for \ + by auto + have "\\. \countable \; \ \ Collect (openin X)\ \ + (countable union_of closedin X) (\ ((-) (topspace X) ` \))" + by (metis Ball_Collect countable_union_of_UN countable_union_of_inc openin_closedin_eq) + then have "\S. gdelta_in X S \ fsigma_in X (topspace X - S)" + by (simp add: fsigma_in_def gdelta_in_def all_relative_to all_intersection_of del: UN_simps) + then show ?thesis + by (metis Diff_Int2 Diff_Int_distrib2 assms fsigma_in_Int fsigma_in_subset inf.absorb_iff2) +qed + +lemma gdelta_in_diff: + assumes "gdelta_in X S" "fsigma_in X T" + shows "gdelta_in X (S - T)" +proof - + have [simp]: "topspace X - \\ = topspace X \ (\T\\. topspace X - T)" for \ + by auto + have "\\. \countable \; \ \ Collect (closedin X)\ + \ (countable intersection_of openin X relative_to topspace X) + (topspace X \ \ ((-) (topspace X) ` \))" + by (metis Ball_Collect closedin_def countable_intersection_of_INT countable_intersection_of_inc relative_to_inc) + then have "\S. fsigma_in X S \ gdelta_in X (topspace X - S)" + by (simp add: fsigma_in_def gdelta_in_def all_union_of del: INT_simps) + then show ?thesis + by (metis Diff_Int2 Diff_Int_distrib2 assms gdelta_in_Int gdelta_in_alt inf.orderE inf_commute) +qed + +lemma gdelta_in_fsigma_in: + "gdelta_in X S \ S \ topspace X \ fsigma_in X (topspace X - S)" + by (metis double_diff fsigma_in_diff fsigma_in_topspace gdelta_in_alt gdelta_in_diff gdelta_in_topspace) + +lemma fsigma_in_gdelta_in: + "fsigma_in X S \ S \ topspace X \ gdelta_in X (topspace X - S)" + by (metis Diff_Diff_Int fsigma_in_subset gdelta_in_fsigma_in inf.absorb_iff2) + +lemma gdelta_in_descending: + "gdelta_in X S \ (\C. (\n. openin X (C n)) \ (\n. C(Suc n) \ C n) \ \(range C) = S)" (is "?lhs=?rhs") +proof + assume ?lhs + then obtain C where C: "S \ topspace X" "\n. closedin X (C n)" + "\n. C n \ C(Suc n)" "\(range C) = topspace X - S" + by (meson fsigma_in_ascending gdelta_in_fsigma_in) + define D where "D \ \n. topspace X - C n" + have "\n. openin X (D n) \ D (Suc n) \ D n" + by (simp add: Diff_mono C D_def openin_diff) + moreover have "\(range D) = S" + by (simp add: Diff_Diff_Int Int_absorb1 C D_def) + ultimately show ?rhs + by metis +next + assume ?rhs + then obtain C where "S \ topspace X" + and C: "\n. openin X (C n)" "\n. C(Suc n) \ C n" "\(range C) = S" + using openin_subset by fastforce + define D where "D \ \n. topspace X - C n" + have "\n. closedin X (D n) \ D n \ D(Suc n)" + by (simp add: Diff_mono C closedin_diff D_def) + moreover have "\(range D) = topspace X - S" + using C D_def by blast + ultimately show ?lhs + by (metis \S \ topspace X\ fsigma_in_ascending gdelta_in_fsigma_in) +qed + +lemma homeomorphic_map_fsigmaness: + assumes f: "homeomorphic_map X Y f" and "U \ topspace X" + shows "fsigma_in Y (f ` U) \ fsigma_in X U" (is "?lhs=?rhs") +proof - + obtain g where g: "homeomorphic_maps X Y f g" and g: "homeomorphic_map Y X g" + and 1: "(\x \ topspace X. g(f x) = x)" and 2: "(\y \ topspace Y. f(g y) = y)" + using assms homeomorphic_map_maps by (metis homeomorphic_maps_map) + show ?thesis + proof + assume ?lhs + then obtain \ where "countable \" and \: "\ \ Collect (closedin Y)" "\\ = f`U" + by (force simp: fsigma_in_def union_of_def) + define \ where "\ \ image (image g) \" + have "countable \" + using \_def \countable \\ by blast + moreover have "\ \ Collect (closedin X)" + using f g homeomorphic_map_closedness_eq \ unfolding \_def by blast + moreover have "\\ \ U" + unfolding \_def by (smt (verit) assms 1 \ image_Union image_iff in_mono subsetI) + moreover have "U \ \\" + unfolding \_def using assms \ + by (smt (verit, del_insts) "1" imageI image_Union subset_iff) + ultimately show ?rhs + by (metis fsigma_in_def subset_antisym union_of_def) + next + assume ?rhs + then obtain \ where "countable \" and \: "\ \ Collect (closedin X)" "\\ = U" + by (auto simp: fsigma_in_def union_of_def) + define \ where "\ \ image (image f) \" + have "countable \" + using \_def \countable \\ by blast + moreover have "\ \ Collect (closedin Y)" + using f g homeomorphic_map_closedness_eq \ unfolding \_def by blast + moreover have "\\ = f`U" + unfolding \_def using \ by blast + ultimately show ?lhs + by (metis fsigma_in_def union_of_def) + qed +qed + +lemma homeomorphic_map_fsigmaness_eq: + "homeomorphic_map X Y f + \ (fsigma_in X U \ U \ topspace X \ fsigma_in Y (f ` U))" + by (metis fsigma_in_subset homeomorphic_map_fsigmaness) + +lemma homeomorphic_map_gdeltaness: + assumes "homeomorphic_map X Y f" "U \ topspace X" + shows "gdelta_in Y (f ` U) \ gdelta_in X U" +proof - + have "topspace Y - f ` U = f ` (topspace X - U)" + by (metis (no_types, lifting) Diff_subset assms homeomorphic_eq_everything_map inj_on_image_set_diff) + then show ?thesis + using assms homeomorphic_imp_surjective_map + by (fastforce simp: gdelta_in_fsigma_in homeomorphic_map_fsigmaness_eq) +qed + +lemma homeomorphic_map_gdeltaness_eq: + "homeomorphic_map X Y f + \ gdelta_in X U \ U \ topspace X \ gdelta_in Y (f ` U)" + by (meson gdelta_in_subset homeomorphic_map_gdeltaness) + +lemma fsigma_in_relative_to: + "(fsigma_in X relative_to S) = fsigma_in (subtopology X S)" + by (simp add: fsigma_in_def countable_union_of_relative_to closedin_relative_to) + +lemma fsigma_in_subtopology: + "fsigma_in (subtopology X U) S \ (\T. fsigma_in X T \ S = T \ U)" + by (metis fsigma_in_relative_to inf_commute relative_to_def) + +lemma gdelta_in_relative_to: + "(gdelta_in X relative_to S) = gdelta_in (subtopology X S)" + apply (simp add: gdelta_in_def) + by (metis countable_intersection_of_relative_to openin_relative_to subtopology_restrict) + +lemma gdelta_in_subtopology: + "gdelta_in (subtopology X U) S \ (\T. gdelta_in X T \ S = T \ U)" + by (metis gdelta_in_relative_to inf_commute relative_to_def) + +lemma fsigma_in_fsigma_subtopology: + "fsigma_in X S \ (fsigma_in (subtopology X S) T \ fsigma_in X T \ T \ S)" + by (metis fsigma_in_Int fsigma_in_gdelta_in fsigma_in_subtopology inf.orderE topspace_subtopology_subset) + +lemma gdelta_in_gdelta_subtopology: + "gdelta_in X S \ (gdelta_in (subtopology X S) T \ gdelta_in X T \ T \ S)" + by (metis gdelta_in_Int gdelta_in_subset gdelta_in_subtopology inf.orderE topspace_subtopology_subset) + +end diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Further_Topology.thy --- a/src/HOL/Analysis/Further_Topology.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Analysis/Further_Topology.thy Sun May 07 22:51:23 2023 +0200 @@ -969,9 +969,6 @@ subsection\ Special cases and corollaries involving spheres\ -lemma disjnt_Diff1: "X \ Y' \ disjnt (X - Y) (X' - Y')" - by (auto simp: disjnt_def) - proposition extend_map_affine_to_sphere_cofinite_simple: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact S" "convex U" "bounded U" diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Homotopy.thy --- a/src/HOL/Analysis/Homotopy.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Analysis/Homotopy.thy Sun May 07 22:51:23 2023 +0200 @@ -156,7 +156,7 @@ qed show ?thesis using assms - apply (clarsimp simp add: homotopic_with_def) + apply (clarsimp simp: homotopic_with_def) subgoal for h by (rule_tac x="h \ (\y. (1 - fst y, snd y))" in exI) (simp add: continuous_map_compose [OF *]) done @@ -211,7 +211,7 @@ show "\t\{0..1}. P (\x. k (t, x))" proof fix t show "t\{0..1} \ P (\x. k (t, x))" - by (cases "t \ 1/2") (auto simp add: k_def P) + by (cases "t \ 1/2") (auto simp: k_def P) qed qed qed @@ -578,7 +578,7 @@ proposition homotopic_paths_join: "\homotopic_paths S p p'; homotopic_paths S q q'; pathfinish p = pathstart q\ \ homotopic_paths S (p +++ q) (p' +++ q')" - apply (clarsimp simp add: homotopic_paths_def homotopic_with_def) + apply (clarsimp simp: homotopic_paths_def homotopic_with_def) apply (rename_tac k1 k2) apply (rule_tac x="(\y. ((k1 \ Pair (fst y)) +++ (k2 \ Pair (fst y))) (snd y))" in exI) apply (intro conjI continuous_intros continuous_on_homotopic_join_lemma; force simp: joinpaths_def pathstart_def pathfinish_def path_image_def) @@ -639,8 +639,8 @@ show "continuous_on {x \ ?A. snd x \ 1/2} (\t. p (fst t * (2 * snd t)))" "continuous_on {x \ ?A. 1/2 \ snd x} (\t. p (fst t * (1 - (2 * snd t - 1))))" "continuous_on ?A snd" - by (intro continuous_on_compose2 [OF p] continuous_intros; auto simp add: mult_le_one)+ - qed (auto simp add: algebra_simps) + by (intro continuous_on_compose2 [OF p] continuous_intros; auto simp: mult_le_one)+ + qed (auto simp: algebra_simps) then show ?thesis using assms apply (subst homotopic_paths_sym_eq) @@ -1004,7 +1004,7 @@ by (force simp: field_simps not_le mult_left_mono affine_ineq dest!: 1 2) show "(subpath u v g +++ subpath v w g) t = subpath u w g (?f t)" if "t \ {0..1}" for t using assms - unfolding joinpaths_def subpath_def by (auto simp add: divide_simps add.commute mult.commute mult.left_commute) + unfolding joinpaths_def subpath_def by (auto simp: divide_simps add.commute mult.commute mult.left_commute) qed (use False in auto) qed then show ?thesis @@ -1062,26 +1062,24 @@ then have "continuous_on ({0..1} \ {0..1}) (g \ fst)" by (fastforce intro!: continuous_intros)+ with g show ?thesis - by (auto simp add: homotopic_loops_def homotopic_with_def path_defs image_subset_iff) + by (auto simp: homotopic_loops_def homotopic_with_def path_defs image_subset_iff) qed lemma homotopic_loops_imp_path_component_value: - "\homotopic_loops S p q; 0 \ t; t \ 1\ - \ path_component S (p t) (q t)" -apply (clarsimp simp add: homotopic_loops_def homotopic_with_def path_defs) + "\homotopic_loops S p q; 0 \ t; t \ 1\ \ path_component S (p t) (q t)" +apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs) apply (rule_tac x="h \ (\u. (u, t))" in exI) apply (fastforce elim!: continuous_on_subset intro!: continuous_intros) done lemma homotopic_points_eq_path_component: "homotopic_loops S (linepath a a) (linepath b b) \ path_component S a b" -by (auto simp: path_component_imp_homotopic_points - dest: homotopic_loops_imp_path_component_value [where t=1]) + using homotopic_loops_imp_path_component_value path_component_imp_homotopic_points by fastforce lemma path_connected_eq_homotopic_points: - "path_connected S \ + "path_connected S \ (\a b. a \ S \ b \ S \ homotopic_loops S (linepath a a) (linepath b b))" -by (auto simp: path_connected_def path_component_def homotopic_points_eq_path_component) + by (auto simp: path_connected_def path_component_def homotopic_points_eq_path_component) subsection\Simply connected sets\ @@ -1134,16 +1132,9 @@ then show ?rhs using simply_connected_eq_contractible_loop_any by (blast intro: simply_connected_imp_path_connected) next - assume r: ?rhs - note pa = r [THEN conjunct2, rule_format] - show ?lhs - proof (clarsimp simp add: simply_connected_eq_contractible_loop_any) - fix p a - assume "path p" and "path_image p \ S" "pathfinish p = pathstart p" - and "a \ S" - with pa [of p] show "homotopic_loops S p (linepath a a)" - using homotopic_loops_trans path_connected_eq_homotopic_points r by blast - qed + assume ?rhs + then show ?lhs + by (meson homotopic_loops_trans path_connected_eq_homotopic_points simply_connected_eq_contractible_loop_any) qed lemma simply_connected_eq_contractible_loop_all: @@ -1158,19 +1149,8 @@ next case False then obtain a where "a \ S" by blast - show ?thesis - proof - assume "simply_connected S" - then show ?rhs - using \a \ S\ \simply_connected S\ simply_connected_eq_contractible_loop_any - by blast - next - assume ?rhs - then show "simply_connected S" - unfolding simply_connected_eq_contractible_loop_any - by (meson False homotopic_loops_refl homotopic_loops_sym homotopic_loops_trans - path_component_imp_homotopic_points path_component_refl) - qed + then show ?thesis + by (meson False homotopic_loops_sym homotopic_loops_trans simply_connected_def simply_connected_eq_contractible_loop_any) qed lemma simply_connected_eq_contractible_path: @@ -1212,19 +1192,17 @@ "path_image q \ S" "pathstart q = pathstart p" "pathfinish q = pathfinish p" for p q proof - - have "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))" - by (simp add: homotopic_paths_rid homotopic_paths_sym that) - also have "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p)) - (p +++ reversepath q +++ q)" + have "homotopic_paths S p (p +++ reversepath q +++ q)" using that - by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl homotopic_paths_sym_eq pathstart_linepath) + by (smt (verit, best) homotopic_paths_join homotopic_paths_linv homotopic_paths_rid homotopic_paths_sym + homotopic_paths_trans pathstart_linepath) also have "homotopic_paths S (p +++ reversepath q +++ q) ((p +++ reversepath q) +++ q)" by (simp add: that homotopic_paths_assoc) also have "homotopic_paths S ((p +++ reversepath q) +++ q) (linepath (pathstart q) (pathstart q) +++ q)" using * [of "p +++ reversepath q"] that - by (simp add: homotopic_paths_join path_image_join) + by (simp add: homotopic_paths_assoc homotopic_paths_join path_image_join) also have "homotopic_paths S (linepath (pathstart q) (pathstart q) +++ q) q" using that homotopic_paths_lid by blast finally show ?thesis . @@ -1291,12 +1269,12 @@ have "\p. path p \ path_image p \ S \ pathfinish p = pathstart p \ homotopic_loops S p (linepath a a)" - using a apply (clarsimp simp add: homotopic_loops_def homotopic_with_def path_defs) + using a apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs) apply (rule_tac x="(h \ (\y. (fst y, (p \ snd) y)))" in exI) apply (intro conjI continuous_on_compose continuous_intros; force elim: continuous_on_subset) done with \a \ S\ show ?thesis - by (auto simp add: simply_connected_eq_contractible_loop_all False) + by (auto simp: simply_connected_eq_contractible_loop_all False) qed corollary contractible_imp_connected: @@ -1369,20 +1347,20 @@ lemma homotopic_into_contractible: fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" assumes f: "continuous_on S f" "f ` S \ T" - and g: "continuous_on S g" "g ` S \ T" - and T: "contractible T" - shows "homotopic_with_canon (\h. True) S T f g" -using homotopic_through_contractible [of S f T id T g id] -by (simp add: assms contractible_imp_path_connected) + and g: "continuous_on S g" "g ` S \ T" + and T: "contractible T" + shows "homotopic_with_canon (\h. True) S T f g" + using homotopic_through_contractible [of S f T id T g id] + by (simp add: assms contractible_imp_path_connected) lemma homotopic_from_contractible: fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" assumes f: "continuous_on S f" "f ` S \ T" - and g: "continuous_on S g" "g ` S \ T" - and "contractible S" "path_connected T" - shows "homotopic_with_canon (\h. True) S T f g" -using homotopic_through_contractible [of S id S f T id g] -by (simp add: assms contractible_imp_path_connected) + and g: "continuous_on S g" "g ` S \ T" + and "contractible S" "path_connected T" + shows "homotopic_with_canon (\h. True) S T f g" + using homotopic_through_contractible [of S id S f T id g] + by (simp add: assms contractible_imp_path_connected) subsection\Starlike sets\ @@ -1402,8 +1380,7 @@ proof - have "rel_interior S \ {}" by (simp add: assms rel_interior_eq_empty) - then obtain a where a: "a \ rel_interior S" by blast - with ST have "a \ T" by blast + with ST obtain a where a: "a \ rel_interior S" and "a \ T" by blast have "\x. x \ T \ open_segment a x \ rel_interior S" by (rule rel_interior_closure_convex_segment [OF \convex S\ a]) (use assms in auto) then have "\x\T. a \ T \ open_segment a x \ T" @@ -1439,7 +1416,7 @@ lemma starlike_imp_contractible: fixes S :: "'a::real_normed_vector set" shows "starlike S \ contractible S" -using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def) + using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def) lemma contractible_UNIV [simp]: "contractible (UNIV :: 'a::real_normed_vector set)" by (simp add: starlike_imp_contractible) @@ -1447,27 +1424,27 @@ lemma starlike_imp_simply_connected: fixes S :: "'a::real_normed_vector set" shows "starlike S \ simply_connected S" -by (simp add: contractible_imp_simply_connected starlike_imp_contractible) + by (simp add: contractible_imp_simply_connected starlike_imp_contractible) lemma convex_imp_simply_connected: fixes S :: "'a::real_normed_vector set" shows "convex S \ simply_connected S" -using convex_imp_starlike starlike_imp_simply_connected by blast + using convex_imp_starlike starlike_imp_simply_connected by blast lemma starlike_imp_path_connected: fixes S :: "'a::real_normed_vector set" shows "starlike S \ path_connected S" -by (simp add: simply_connected_imp_path_connected starlike_imp_simply_connected) + by (simp add: simply_connected_imp_path_connected starlike_imp_simply_connected) lemma starlike_imp_connected: fixes S :: "'a::real_normed_vector set" shows "starlike S \ connected S" -by (simp add: path_connected_imp_connected starlike_imp_path_connected) + by (simp add: path_connected_imp_connected starlike_imp_path_connected) lemma is_interval_simply_connected_1: fixes S :: "real set" shows "is_interval S \ simply_connected S" -using convex_imp_simply_connected is_interval_convex_1 is_interval_path_connected_1 simply_connected_imp_path_connected by auto + by (meson convex_imp_simply_connected is_interval_connected_1 is_interval_convex_1 simply_connected_imp_connected) lemma contractible_empty [simp]: "contractible {}" by (simp add: contractible_def homotopic_on_emptyI) @@ -1476,15 +1453,8 @@ fixes S :: "'a::euclidean_space set" assumes "convex S" and TS: "rel_interior S \ T" "T \ closure S" shows "contractible T" -proof (cases "S = {}") - case True - with assms show ?thesis - by (simp add: subsetCE) -next - case False - show ?thesis - by (meson False assms starlike_convex_tweak_boundary_points starlike_imp_contractible) -qed + by (metis assms closure_eq_empty contractible_empty empty_subsetI + starlike_convex_tweak_boundary_points starlike_imp_contractible subset_antisym) lemma convex_imp_contractible: fixes S :: "'a::real_normed_vector set" @@ -1494,13 +1464,13 @@ lemma contractible_sing [simp]: fixes a :: "'a::real_normed_vector" shows "contractible {a}" -by (rule convex_imp_contractible [OF convex_singleton]) + by (rule convex_imp_contractible [OF convex_singleton]) lemma is_interval_contractible_1: fixes S :: "real set" shows "is_interval S \ contractible S" -using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1 - is_interval_simply_connected_1 by auto + using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1 + is_interval_simply_connected_1 by auto lemma contractible_Times: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" @@ -1556,9 +1526,7 @@ lemma locally_open_subset: assumes "locally P S" "openin (top_of_set S) t" shows "locally P t" - using assms - unfolding locally_def - by (elim all_forward) (meson dual_order.trans openin_imp_subset openin_subset_trans openin_trans) + by (smt (verit, ccfv_SIG) assms order.trans locally_def openin_imp_subset openin_subset_trans openin_trans) lemma locally_diff_closed: "\locally P S; closedin (top_of_set S) t\ \ locally P (S - t)" @@ -1575,16 +1543,13 @@ using zero_less_one by blast then show ?thesis unfolding locally_def - by (auto simp add: openin_euclidean_subtopology_iff subset_singleton_iff conj_disj_distribR) + by (auto simp: openin_euclidean_subtopology_iff subset_singleton_iff conj_disj_distribR) qed lemma locally_iff: "locally P S \ (\T x. open T \ x \ S \ T \ (\U. open U \ (\V. P V \ x \ S \ U \ S \ U \ V \ V \ S \ T)))" - apply (simp add: le_inf_iff locally_def openin_open, safe) - apply (metis IntE IntI le_inf_iff) - apply (metis IntI Int_subset_iff) - done + by (smt (verit) locally_def openin_open) lemma locally_Int: assumes S: "locally P S" and T: "locally P T" @@ -1649,34 +1614,29 @@ using hom by (auto simp: homeomorphism_def) have gw: "g ` W = S \ f -` W" using \W \ T\ g by force - have \: "openin (top_of_set S) (g ` W)" - proof - - have "continuous_on S f" - using f(3) by blast - then show "openin (top_of_set S) (g ` W)" - by (simp add: gw Collect_conj_eq \openin (top_of_set T) W\ continuous_on_open f(2)) - qed + have "openin (top_of_set S) (g ` W)" + using \openin (top_of_set T) W\ continuous_on_open f gw by auto then obtain U V where osu: "openin (top_of_set S) U" and uv: "P V" "g y \ U" "U \ V" "V \ g ` W" - using S [unfolded locally_def, rule_format, of "g ` W" "g y"] \y \ W\ by force + by (metis S \y \ W\ image_eqI locallyE) have "V \ S" using uv by (simp add: gw) have fv: "f ` V = T \ {x. g x \ V}" using \f ` S = T\ f \V \ S\ by auto + have contvf: "continuous_on V f" + using \V \ S\ continuous_on_subset f(3) by blast have "f ` V \ W" using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto - have contvf: "continuous_on V f" - using \V \ S\ continuous_on_subset f(3) by blast - have contvg: "continuous_on (f ` V) g" - using \f ` V \ W\ \W \ T\ continuous_on_subset [OF g(3)] by blast + then have contvg: "continuous_on (f ` V) g" + using \W \ T\ continuous_on_subset [OF g(3)] by blast have "V \ g ` f ` V" by (metis \V \ S\ hom homeomorphism_def homeomorphism_of_subsets order_refl) then have homv: "homeomorphism V (f ` V) f g" - using \V \ S\ f by (auto simp add: homeomorphism_def contvf contvg) + using \V \ S\ f by (auto simp: homeomorphism_def contvf contvg) have "openin (top_of_set (g ` T)) U" using \g ` T = S\ by (simp add: osu) - then have 1: "openin (top_of_set T) (T \ g -` U)" + then have "openin (top_of_set T) (T \ g -` U)" using \continuous_on T g\ continuous_on_open [THEN iffD1] by blast - have 2: "\V. Q V \ y \ (T \ g -` U) \ (T \ g -` U) \ V \ V \ W" + moreover have "\V. Q V \ y \ (T \ g -` U) \ (T \ g -` U) \ V \ V \ W" proof (intro exI conjI) show "Q (f ` V)" using Q homv \P V\ by blast @@ -1687,44 +1647,28 @@ show "f ` V \ W" using \f ` V \ W\ by blast qed - show "\U. openin (top_of_set T) U \ (\v. Q v \ y \ U \ U \ v \ v \ W)" - by (meson 1 2) + ultimately show "\U. openin (top_of_set T) U \ (\v. Q v \ y \ U \ U \ v \ v \ W)" + by meson qed lemma homeomorphism_locally: fixes f:: "'a::metric_space \ 'b::metric_space" - assumes hom: "homeomorphism S T f g" - and eq: "\S T. homeomorphism S T f g \ (P S \ Q T)" + assumes "homeomorphism S T f g" + and "\S T. homeomorphism S T f g \ (P S \ Q T)" shows "locally P S \ locally Q T" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - using eq hom homeomorphism_locally_imp by blast -next - assume ?rhs - then show ?lhs - using eq homeomorphism_sym homeomorphism_symD [OF hom] - by (blast intro: homeomorphism_locally_imp) -qed + by (smt (verit) assms homeomorphism_locally_imp homeomorphism_symD) lemma homeomorphic_locally: fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" assumes hom: "S homeomorphic T" and iff: "\X Y. X homeomorphic Y \ (P X \ Q Y)" shows "locally P S \ locally Q T" -proof - - obtain f g where hom: "homeomorphism S T f g" - using assms by (force simp: homeomorphic_def) - then show ?thesis - using homeomorphic_def local.iff - by (blast intro!: homeomorphism_locally) -qed + by (smt (verit, ccfv_SIG) hom homeomorphic_def homeomorphism_locally homeomorphism_locally_imp iff) lemma homeomorphic_local_compactness: fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" shows "S homeomorphic T \ locally compact S \ locally compact T" -by (simp add: homeomorphic_compactness homeomorphic_locally) + by (simp add: homeomorphic_compactness homeomorphic_locally) lemma locally_translation: fixes P :: "'a :: real_normed_vector set \ bool" @@ -1746,7 +1690,7 @@ and oo: "\T. openin (top_of_set S) T \ openin (top_of_set (f ` S)) (f ` T)" and Q: "\T. \T \ S; P T\ \ Q(f ` T)" shows "locally Q (f ` S)" -proof (clarsimp simp add: locally_def) +proof (clarsimp simp: locally_def) fix W y assume oiw: "openin (top_of_set (f ` S)) W" and "y \ W" then have "W \ f ` S" by (simp add: openin_euclidean_subtopology_iff) @@ -1756,8 +1700,7 @@ using \W \ f ` S\ \y \ W\ by blast then obtain U V where "openin (top_of_set S) U" "P V" "x \ U" "U \ V" "V \ S \ f -` W" - using P [unfolded locally_def, rule_format, of "(S \ f -` W)" x] oivf \y \ W\ - by auto + by (metis IntI P \y \ W\ locallyE oivf vimageI) then have "openin (top_of_set (f ` S)) (f ` U)" by (simp add: oo) then show "\X. openin (top_of_set (f ` S)) X \ (\Y. Q Y \ y \ X \ X \ Y \ Y \ W)" @@ -1778,17 +1721,16 @@ proof - let ?A = "{b. \T. openin (top_of_set S) T \ b \ T \ (\x\T. P x \ Q x)}" let ?B = "{b. \T. openin (top_of_set S) T \ b \ T \ (\x\T. P x \ \ Q x)}" - have 1: "openin (top_of_set S) ?A" - by (subst openin_subopen, blast) - have 2: "openin (top_of_set S) ?B" - by (subst openin_subopen, blast) - have \
: "?A \ ?B = {}" + have "?A \ ?B = {}" by (clarsimp simp: set_eq_iff) (metis (no_types, opaque_lifting) Int_iff opD openin_Int) - have *: "S \ ?A \ ?B" + moreover have "S \ ?A \ ?B" by clarsimp (meson opI) - have "?A = {} \ ?B = {}" - using \connected S\ [unfolded connected_openin, simplified, rule_format, OF 1 \
* 2] - by blast + moreover have "openin (top_of_set S) ?A" + by (subst openin_subopen, blast) + moreover have "openin (top_of_set S) ?B" + by (subst openin_subopen, blast) + ultimately have "?A = {} \ ?B = {}" + by (metis (no_types, lifting) \connected S\ connected_openin) then show ?thesis by clarsimp (meson opI etc) qed @@ -1851,15 +1793,11 @@ shows "locally (\U. f constant_on U) S \ f constant_on S" (is "?lhs = ?rhs") proof assume ?lhs - then have "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x\T. f x = f a)" - unfolding locally_def - by (metis (mono_tags, opaque_lifting) constant_on_def constant_on_subset openin_subtopology_self) then show ?rhs - using assms - by (simp add: locally_constant_imp_constant) + by (smt (verit, del_insts) assms constant_on_def locally_constant_imp_constant locally_def openin_subtopology_self subset_iff) next assume ?rhs then show ?lhs - using assms by (metis constant_on_subset locallyI openin_imp_subset order_refl) + by (metis constant_on_subset locallyI openin_imp_subset order_refl) qed @@ -1938,10 +1876,8 @@ proof fix x assume "x \ S" - then obtain e where "e>0" and e: "closed (cball x e \ S)" - using R by blast - then have "compact (cball x e \ S)" - by (simp add: bounded_Int compact_eq_bounded_closed) + then obtain e where "e>0" and "compact (cball x e \ S)" + by (metis Int_commute compact_Int_closed compact_cball inf.right_idem R) moreover have "\y\ball x e \ S. \\>0. cball y \ \ S \ ball x e" by (meson Elementary_Metric_Spaces.open_ball IntD1 le_infI1 open_contains_cball_eq) moreover have "openin (top_of_set S) (ball x e \ S)" @@ -1984,10 +1920,8 @@ have "openin (top_of_set S) (\ (u ` T))" using T that uv by fastforce moreover - have "compact (\ (v ` T))" - by (meson T compact_UN subset_eq that(1) uv) - moreover have "\ (v ` T) \ S" - by (metis SUP_least T(1) subset_eq that(1) uv) + obtain "compact (\ (v ` T))" "\ (v ` T) \ S" + by (metis T UN_subset_iff \K \ S\ compact_UN subset_iff uv) ultimately show ?thesis using T by auto qed @@ -1996,7 +1930,7 @@ next assume ?rhs then show ?lhs - apply (clarsimp simp add: locally_compact) + apply (clarsimp simp: locally_compact) apply (drule_tac x="{x}" in spec, simp) done qed @@ -2014,14 +1948,7 @@ have ope: "openin (top_of_set S) (ball x e)" by (meson e open_ball ball_subset_cball dual_order.trans open_subset) show ?thesis - proof (intro exI conjI) - let ?U = "ball x e" - let ?V = "cball x e" - show "x \ ?U" "?U \ ?V" "?V \ S" "compact ?V" - using \e > 0\ e by auto - show "openin (top_of_set S) ?U" - using ope by blast - qed + by (meson \0 < e\ ball_subset_cball centre_in_ball compact_cball e ope) qed show ?thesis unfolding locally_compact by (blast intro: *) @@ -2045,13 +1972,13 @@ lemma locally_compact_Int: fixes S :: "'a :: t2_space set" - shows "\locally compact S; locally compact t\ \ locally compact (S \ t)" -by (simp add: compact_Int locally_Int) + shows "\locally compact S; locally compact T\ \ locally compact (S \ T)" + by (simp add: compact_Int locally_Int) lemma locally_compact_closedin: fixes S :: "'a :: heine_borel set" - shows "\closedin (top_of_set S) t; locally compact S\ - \ locally compact t" + shows "\closedin (top_of_set S) T; locally compact S\ + \ locally compact T" unfolding closedin_closed using closed_imp_locally_compact locally_compact_Int by blast @@ -2130,12 +2057,7 @@ obtain e2 where "e2 > 0" and e2: "closed (cball x e2 \ T)" using LCT \x \ T\ unfolding locally_compact_Int_cball by blast moreover have "closed (cball x (min e1 e2) \ (S \ T))" - proof - - have "closed (cball x e1 \ (cball x e2 \ S))" - by (metis closed_Int closed_cball e1 inf_left_commute) - then show ?thesis - by (simp add: Int_Un_distrib cball_min_Int closed_Int closed_Un e2 inf_assoc) - qed + by (smt (verit) Int_Un_distrib2 Int_commute cball_min_Int closed_Int closed_Un closed_cball e1 e2 inf_left_commute) ultimately show ?thesis by (rule_tac x="min e1 e2" in exI) linarith qed @@ -2298,7 +2220,7 @@ by (simp_all add: closedin_closed_Int) moreover have "D \ closure V1 = D \ V1" "D \ closure V2 = D \ V2" using \D \ V1 \ V2\ \open V1\ \open V2\ V12 - by (auto simp add: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+) + by (auto simp: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+) ultimately have cloDV1: "closedin (top_of_set D) (D \ V1)" and cloDV2: "closedin (top_of_set D) (D \ V2)" by metis+ @@ -2721,7 +2643,7 @@ fixes S :: "'a:: real_normed_vector set" assumes "convex S" shows "locally path_connected S" -proof (clarsimp simp add: locally_path_connected) +proof (clarsimp simp: locally_path_connected) fix V x assume "openin (top_of_set S) V" and "x \ V" then obtain T e where "V = S \ T" "x \ S" "0 < e" "ball x e \ T" @@ -3824,7 +3746,7 @@ show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set S)) (top_of_set S) (\(t,x). (1 - t) * x + t * z)" using \z \ S\ - by (auto simp add: case_prod_unfold intro!: continuous_intros \
) + by (auto simp: case_prod_unfold intro!: continuous_intros \
) qed auto qed (simp add: contractible_space_empty) qed @@ -3946,7 +3868,7 @@ by (rule homotopic_with_compose_continuous_right [where X=S]) (use k in auto) moreover have "homotopic_with_canon (\x. True) T U (f \ id) (f \ (h \ k))" by (rule homotopic_with_compose_continuous_left [where Y=T]) - (use f in \auto simp add: hom homotopic_with_symD\) + (use f in \auto simp: hom homotopic_with_symD\) ultimately show ?thesis using that homotopic_with_trans by (fastforce simp add: o_def) qed @@ -3987,7 +3909,7 @@ by (rule homotopic_with_compose_continuous_left [where Y=S]) (use h in auto) moreover have "homotopic_with_canon (\x. True) U T (id \ f) ((h \ k) \ f)" by (rule homotopic_with_compose_continuous_right [where X=T]) - (use f in \auto simp add: hom homotopic_with_symD\) + (use f in \auto simp: hom homotopic_with_symD\) ultimately show ?thesis using homotopic_with_trans by (fastforce simp add: o_def) qed @@ -4123,7 +4045,7 @@ fixes U :: "'a::euclidean_space set" assumes "convex U" "\ collinear U" "countable S" shows "path_connected(U - S)" -proof (clarsimp simp add: path_connected_def) +proof (clarsimp simp: path_connected_def) fix a b assume "a \ U" "a \ S" "b \ U" "b \ S" let ?m = "midpoint a b" @@ -4245,7 +4167,7 @@ next assume ge2: "aff_dim S \ 2" then have "\ collinear S" - proof (clarsimp simp add: collinear_affine_hull) + proof (clarsimp simp: collinear_affine_hull) fix u v assume "S \ affine hull {u, v}" then have "aff_dim S \ aff_dim {u, v}" @@ -4282,7 +4204,7 @@ assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S" and "\ collinear S" "countable T" shows "path_connected(S - T)" -proof (clarsimp simp add: path_connected_component) +proof (clarsimp simp: path_connected_component) fix x y assume xy: "x \ S" "x \ T" "y \ S" "y \ T" show "path_component (S - T) x y" @@ -4483,7 +4405,7 @@ by blast next show "cball a r \ T \ f ` (cball a r \ T)" - proof (clarsimp simp add: dist_norm norm_minus_commute) + proof (clarsimp simp: dist_norm norm_minus_commute) fix x assume x: "norm (x - a) \ r" and "x \ T" have "\v \ {0..1}. ((1 - v) * r - norm ((x - a) - v *\<^sub>R (u - a))) \ 1 = 0" @@ -4575,7 +4497,7 @@ then show "continuous_on S gg" by (rule continuous_on_subset) (use ST in auto) show "ff ` S \ S" - proof (clarsimp simp add: ff_def) + proof (clarsimp simp: ff_def) fix x assume "x \ S" and x: "dist a x < r" and "x \ T" then have "f x \ cball a r \ T" @@ -4584,7 +4506,7 @@ using ST(1) \x \ T\ gid hom homeomorphism_def x by fastforce qed show "gg ` S \ S" - proof (clarsimp simp add: gg_def) + proof (clarsimp simp: gg_def) fix x assume "x \ S" and x: "dist a x < r" and "x \ T" then have "g x \ cball a r \ T" diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Ordered_Euclidean_Space.thy --- a/src/HOL/Analysis/Ordered_Euclidean_Space.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Analysis/Ordered_Euclidean_Space.thy Sun May 07 22:51:23 2023 +0200 @@ -2,7 +2,7 @@ theory Ordered_Euclidean_Space imports - Convex_Euclidean_Space + Convex_Euclidean_Space Abstract_Limits "HOL-Library.Product_Order" begin @@ -157,6 +157,32 @@ shows "((\i. inf (X i) (Y i)) \ inf x y) net" unfolding inf_min eucl_inf by (intro assms tendsto_intros) +lemma tendsto_Inf: + fixes f :: "'a \ 'b \ 'c::ordered_euclidean_space" + assumes "finite K" "\i. i \ K \ ((\x. f x i) \ l i) F" + shows "((\x. Inf (f x ` K)) \ Inf (l ` K)) F" + using assms + by (induction K rule: finite_induct) (auto simp: cInf_insert_If tendsto_inf) + +lemma tendsto_Sup: + fixes f :: "'a \ 'b \ 'c::ordered_euclidean_space" + assumes "finite K" "\i. i \ K \ ((\x. f x i) \ l i) F" + shows "((\x. Sup (f x ` K)) \ Sup (l ` K)) F" + using assms + by (induction K rule: finite_induct) (auto simp: cSup_insert_If tendsto_sup) + +lemma continuous_map_Inf: + fixes f :: "'a \ 'b \ 'c::ordered_euclidean_space" + assumes "finite K" "\i. i \ K \ continuous_map X euclidean (\x. f x i)" + shows "continuous_map X euclidean (\x. INF i\K. f x i)" + using assms by (simp add: continuous_map_atin tendsto_Inf) + +lemma continuous_map_Sup: + fixes f :: "'a \ 'b \ 'c::ordered_euclidean_space" + assumes "finite K" "\i. i \ K \ continuous_map X euclidean (\x. f x i)" + shows "continuous_map X euclidean (\x. SUP i\K. f x i)" + using assms by (simp add: continuous_map_atin tendsto_Sup) + lemma tendsto_componentwise_max: assumes f: "(f \ l) F" and g: "(g \ m) F" shows "((\x. (\i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i)) \ (\i\Basis. max (l \ i) (m \ i) *\<^sub>R i)) F" @@ -167,10 +193,6 @@ shows "((\x. (\i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i)) \ (\i\Basis. min (l \ i) (m \ i) *\<^sub>R i)) F" by (intro tendsto_intros assms) -lemma (in order) atLeastatMost_empty'[simp]: - "(\ a \ b) \ {a..b} = {}" - by (auto) - instance real :: ordered_euclidean_space by standard auto diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Analysis/Path_Connected.thy Sun May 07 22:51:23 2023 +0200 @@ -2313,6 +2313,174 @@ using assms homeomorphic_map_path_component_of by fastforce +subsection\Paths and path-connectedness\ + +lemma path_connected_space_quotient_map_image: + "\quotient_map X Y q; path_connected_space X\ \ path_connected_space Y" + by (metis path_connectedin_continuous_map_image path_connectedin_topspace quotient_imp_continuous_map quotient_imp_surjective_map) + +lemma path_connected_space_retraction_map_image: + "\retraction_map X Y r; path_connected_space X\ \ path_connected_space Y" + using path_connected_space_quotient_map_image retraction_imp_quotient_map by blast + +lemma path_connected_space_prod_topology: + "path_connected_space(prod_topology X Y) \ + topspace(prod_topology X Y) = {} \ path_connected_space X \ path_connected_space Y" +proof (cases "topspace(prod_topology X Y) = {}") + case True + then show ?thesis + by (simp add: path_connected_space_topspace_empty) +next + case False + have "path_connected_space (prod_topology X Y)" + if X: "path_connected_space X" and Y: "path_connected_space Y" + proof (clarsimp simp: path_connected_space_def) + fix x y x' y' + assume "x \ topspace X" and "y \ topspace Y" and "x' \ topspace X" and "y' \ topspace Y" + obtain f where "pathin X f" "f 0 = x" "f 1 = x'" + by (meson X \x \ topspace X\ \x' \ topspace X\ path_connected_space_def) + obtain g where "pathin Y g" "g 0 = y" "g 1 = y'" + by (meson Y \y \ topspace Y\ \y' \ topspace Y\ path_connected_space_def) + show "\h. pathin (prod_topology X Y) h \ h 0 = (x,y) \ h 1 = (x',y')" + proof (intro exI conjI) + show "pathin (prod_topology X Y) (\t. (f t, g t))" + using \pathin X f\ \pathin Y g\ by (simp add: continuous_map_paired pathin_def) + show "(\t. (f t, g t)) 0 = (x, y)" + using \f 0 = x\ \g 0 = y\ by blast + show "(\t. (f t, g t)) 1 = (x', y')" + using \f 1 = x'\ \g 1 = y'\ by blast + qed + qed + then show ?thesis + by (metis False Sigma_empty1 Sigma_empty2 topspace_prod_topology path_connected_space_retraction_map_image + retraction_map_fst retraction_map_snd) +qed + +lemma path_connectedin_Times: + "path_connectedin (prod_topology X Y) (S \ T) \ + S = {} \ T = {} \ path_connectedin X S \ path_connectedin Y T" + by (auto simp add: path_connectedin_def subtopology_Times path_connected_space_prod_topology) + + +subsection\Path components\ + +lemma path_component_of_subtopology_eq: + "path_component_of (subtopology X U) x = path_component_of X x \ path_component_of_set X x \ U" + (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + by (metis path_connectedin_path_component_of path_connectedin_subtopology) +next + show "?rhs \ ?lhs" + unfolding fun_eq_iff + by (metis path_connectedin_subtopology path_component_of path_component_of_aux path_component_of_mono) +qed + +lemma path_components_of_subtopology: + assumes "C \ path_components_of X" "C \ U" + shows "C \ path_components_of (subtopology X U)" + using assms path_component_of_refl path_component_of_subtopology_eq topspace_subtopology + by (smt (verit) imageE path_component_in_path_components_of path_components_of_def) + +lemma path_imp_connected_component_of: + "path_component_of X x y \ connected_component_of X x y" + by (metis in_mono mem_Collect_eq path_component_subset_connected_component_of) + +lemma finite_path_components_of_finite: + "finite(topspace X) \ finite(path_components_of X)" + by (simp add: Union_path_components_of finite_UnionD) + +lemma path_component_of_continuous_image: + "\continuous_map X X' f; path_component_of X x y\ \ path_component_of X' (f x) (f y)" + by (meson image_eqI path_component_of path_connectedin_continuous_map_image) + +lemma path_component_of_pair [simp]: + "path_component_of_set (prod_topology X Y) (x,y) = + path_component_of_set X x \ path_component_of_set Y y" (is "?lhs = ?rhs") +proof (cases "?lhs = {}") + case True + then show ?thesis + by (metis Sigma_empty1 Sigma_empty2 mem_Sigma_iff path_component_of_eq_empty topspace_prod_topology) +next + case False + then have "path_component_of X x x" "path_component_of Y y y" + using path_component_of_eq_empty path_component_of_refl by fastforce+ + moreover + have "path_connectedin (prod_topology X Y) (path_component_of_set X x \ path_component_of_set Y y)" + by (metis path_connectedin_Times path_connectedin_path_component_of) + moreover have "path_component_of X x a" "path_component_of Y y b" + if "(x, y) \ C'" "(a,b) \ C'" and "path_connectedin (prod_topology X Y) C'" for C' a b + by (smt (verit, best) that continuous_map_fst continuous_map_snd fst_conv snd_conv path_component_of path_component_of_continuous_image)+ + ultimately show ?thesis + by (intro path_component_of_unique) auto +qed + +lemma path_components_of_prod_topology: + "path_components_of (prod_topology X Y) = + (\(C,D). C \ D) ` (path_components_of X \ path_components_of Y)" + by (force simp add: image_iff path_components_of_def) + +lemma path_components_of_prod_topology': + "path_components_of (prod_topology X Y) = + {C \ D |C D. C \ path_components_of X \ D \ path_components_of Y}" + by (auto simp: path_components_of_prod_topology) + +lemma path_component_of_product_topology: + "path_component_of_set (product_topology X I) f = + (if f \ extensional I then PiE I (\i. path_component_of_set (X i) (f i)) else {})" + (is "?lhs = ?rhs") +proof (cases "path_component_of_set (product_topology X I) f = {}") + case True + then show ?thesis + by (smt (verit) PiE_eq_empty_iff PiE_iff path_component_of_eq_empty topspace_product_topology) +next + case False + then have [simp]: "f \ extensional I" + by (auto simp: path_component_of_eq_empty PiE_iff path_component_of_equiv) + show ?thesis + proof (intro path_component_of_unique) + show "f \ ?rhs" + using False path_component_of_eq_empty path_component_of_refl by force + show "path_connectedin (product_topology X I) (if f \ extensional I then \\<^sub>E i\I. path_component_of_set (X i) (f i) else {})" + by (simp add: path_connectedin_PiE path_connectedin_path_component_of) + fix C' + assume "f \ C'" and C': "path_connectedin (product_topology X I) C'" + show "C' \ ?rhs" + proof - + have "C' \ extensional I" + using PiE_def C' path_connectedin_subset_topspace by fastforce + with \f \ C'\ C' show ?thesis + apply (clarsimp simp: PiE_iff subset_iff) + by (smt (verit, ccfv_threshold) continuous_map_product_projection path_component_of path_component_of_continuous_image) + qed + qed +qed + +lemma path_components_of_product_topology: + "path_components_of (product_topology X I) = + {PiE I B |B. \i \ I. B i \ path_components_of(X i)}" (is "?lhs=?rhs") +proof + show "?lhs \ ?rhs" + apply (simp add: path_components_of_def image_subset_iff) + by (smt (verit, best) PiE_iff image_eqI path_component_of_product_topology) +next + show "?rhs \ ?lhs" + proof + fix F + assume "F \ ?rhs" + then obtain B where B: "F = Pi\<^sub>E I B" + and "\i\I. \x\topspace (X i). B i = path_component_of_set (X i) x" + by (force simp add: path_components_of_def image_iff) + then obtain f where ftop: "\i. i \ I \ f i \ topspace (X i)" + and BF: "\i. i \ I \ B i = path_component_of_set (X i) (f i)" + by metis + then have "F = path_component_of_set (product_topology X I) (restrict f I)" + by (metis (mono_tags, lifting) B PiE_cong path_component_of_product_topology restrict_apply' restrict_extensional) + then show "F \ ?lhs" + by (simp add: ftop path_component_in_path_components_of) + qed +qed + subsection \Sphere is path-connected\ lemma path_connected_punctured_universe: diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Product_Topology.thy --- a/src/HOL/Analysis/Product_Topology.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Analysis/Product_Topology.thy Sun May 07 22:51:23 2023 +0200 @@ -1,7 +1,7 @@ section\The binary product topology\ theory Product_Topology -imports Function_Topology + imports Function_Topology begin section \Product Topology\ @@ -464,6 +464,10 @@ "homeomorphic_map (prod_topology X Y) (prod_topology Y X) (\(x,y). (y,x))" using homeomorphic_map_maps homeomorphic_maps_swap by metis +lemma homeomorphic_space_prod_topology_swap: + "(prod_topology X Y) homeomorphic_space (prod_topology Y X)" + using homeomorphic_map_swap homeomorphic_space by blast + lemma embedding_map_graph: "embedding_map X (prod_topology X Y) (\x. (x, f x)) \ continuous_map X Y f" (is "?lhs = ?rhs") @@ -472,8 +476,7 @@ have "snd \ (\x. (x, f x)) = f" by force moreover have "continuous_map X Y (snd \ (\x. (x, f x)))" - using L - unfolding embedding_map_def + using L unfolding embedding_map_def by (meson continuous_map_in_subtopology continuous_map_snd_of homeomorphic_imp_continuous_map) ultimately show ?rhs by simp diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/Sum_Topology.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Sum_Topology.thy Sun May 07 22:51:23 2023 +0200 @@ -0,0 +1,146 @@ +section\Disjoint sum of arbitarily many spaces\ + +theory Sum_Topology + imports Abstract_Topology +begin + + +definition sum_topology :: "('a \ 'b topology) \ 'a set \ ('a \ 'b) topology" where + "sum_topology X I \ + topology (\U. U \ Sigma I (topspace \ X) \ (\i \ I. openin (X i) {x. (i,x) \ U}))" + +lemma is_sum_topology: "istopology (\U. U \ Sigma I (topspace \ X) \ (\i\I. openin (X i) {x. (i, x) \ U}))" +proof - + have 1: "{x. (i, x) \ S \ T} = {x. (i, x) \ S} \ {x::'b. (i, x) \ T}" for S T and i::'a + by auto + have 2: "{x. (i, x) \ \ \} = (\K\\. {x::'b. (i, x) \ K})" for \ and i::'a + by auto + show ?thesis + unfolding istopology_def 1 2 by blast +qed + +lemma openin_sum_topology: + "openin (sum_topology X I) U \ + U \ Sigma I (topspace \ X) \ (\i \ I. openin (X i) {x. (i,x) \ U})" + by (auto simp: sum_topology_def is_sum_topology) + +lemma openin_disjoint_union: + "openin (sum_topology X I) (Sigma I U) \ (\i \ I. openin (X i) (U i))" + using openin_subset by (force simp: openin_sum_topology) + +lemma topspace_sum_topology [simp]: + "topspace(sum_topology X I) = Sigma I (topspace \ X)" + by (metis comp_apply openin_disjoint_union openin_subset openin_sum_topology openin_topspace subset_antisym) + +lemma openin_sum_topology_alt: + "openin (sum_topology X I) U \ (\T. U = Sigma I T \ (\i \ I. openin (X i) (T i)))" + by (bestsimp simp: openin_sum_topology dest: openin_subset) + +lemma forall_openin_sum_topology: + "(\U. openin (sum_topology X I) U \ P U) \ (\T. (\i \ I. openin (X i) (T i)) \ P(Sigma I T))" + by (auto simp: openin_sum_topology_alt) + +lemma exists_openin_sum_topology: + "(\U. openin (sum_topology X I) U \ P U) \ + (\T. (\i \ I. openin (X i) (T i)) \ P(Sigma I T))" + by (auto simp: openin_sum_topology_alt) + +lemma closedin_sum_topology: + "closedin (sum_topology X I) U \ U \ Sigma I (topspace \ X) \ (\i \ I. closedin (X i) {x. (i,x) \ U})" + (is "?lhs = ?rhs") +proof + assume L: ?lhs + then have U: "U \ Sigma I (topspace \ X)" + using closedin_subset by fastforce + then have "\i\I. {x. (i, x) \ U} \ topspace (X i)" + by fastforce + moreover have "openin (X i) (topspace (X i) - {x. (i, x) \ U})" if "i\I" for i + apply (subst openin_subopen) + using L that unfolding closedin_def openin_sum_topology + by (bestsimp simp: o_def subset_iff) + ultimately show ?rhs + by (simp add: U closedin_def) +next + assume R: ?rhs + then have "openin (X i) {x. (i, x) \ topspace (sum_topology X I) - U}" if "i\I" for i + apply (subst openin_subopen) + using that unfolding closedin_def openin_sum_topology + by (bestsimp simp: o_def subset_iff) + then show ?lhs + by (simp add: R closedin_def openin_sum_topology) +qed + +lemma closedin_disjoint_union: + "closedin (sum_topology X I) (Sigma I U) \ (\i \ I. closedin (X i) (U i))" + using closedin_subset by (force simp: closedin_sum_topology) + +lemma closedin_sum_topology_alt: + "closedin (sum_topology X I) U \ (\T. U = Sigma I T \ (\i \ I. closedin (X i) (T i)))" + using closedin_subset unfolding closedin_sum_topology by auto blast + +lemma forall_closedin_sum_topology: + "(\U. closedin (sum_topology X I) U \ P U) \ + (\t. (\i \ I. closedin (X i) (t i)) \ P(Sigma I t))" + by (metis closedin_sum_topology_alt) + +lemma exists_closedin_sum_topology: + "(\U. closedin (sum_topology X I) U \ P U) \ + (\T. (\i \ I. closedin (X i) (T i)) \ P(Sigma I T))" + by (simp add: closedin_sum_topology_alt) blast + +lemma open_map_component_injection: + "i \ I \ open_map (X i) (sum_topology X I) (\x. (i,x))" + unfolding open_map_def openin_sum_topology + using openin_subset openin_subopen by (force simp: image_iff) + +lemma closed_map_component_injection: + assumes "i \ I" + shows "closed_map(X i) (sum_topology X I) (\x. (i,x))" +proof - + have "closedin (X j) {x. j = i \ x \ U}" + if "\U S. closedin U S \ S \ topspace U" and "closedin (X i) U" and "j \ I" for U j + using that by (cases "j=i") auto + then show ?thesis + using closedin_subset assms by (force simp: closed_map_def closedin_sum_topology image_iff) +qed + +lemma continuous_map_component_injection: + "i \ I \ continuous_map(X i) (sum_topology X I) (\x. (i,x))" + apply (clarsimp simp: continuous_map_def openin_sum_topology) + by (smt (verit, best) Collect_cong mem_Collect_eq openin_subset subsetD) + +lemma subtopology_sum_topology: + "subtopology (sum_topology X I) (Sigma I S) = + sum_topology (\i. subtopology (X i) (S i)) I" + unfolding topology_eq +proof (intro strip iffI) + fix T + assume *: "openin (subtopology (sum_topology X I) (Sigma I S)) T" + then obtain U where U: "\i\I. openin (X i) (U i) \ T = Sigma I U \ Sigma I S" + by (auto simp: openin_subtopology openin_sum_topology_alt) + have "T = (SIGMA i:I. U i \ S i)" + proof + show "T \ (SIGMA i:I. U i \ S i)" + by (metis "*" SigmaE Sigma_Int_distrib2 U openin_imp_subset subset_iff) + show "(SIGMA i:I. U i \ S i) \ T" + using U by fastforce + qed + then show "openin (sum_topology (\i. subtopology (X i) (S i)) I) T" + by (simp add: U openin_disjoint_union openin_subtopology_Int) +next + fix T + assume *: "openin (sum_topology (\i. subtopology (X i) (S i)) I) T" + then obtain U where "\i\I. \T. openin (X i) T \ U i = T \ S i" and Teq: "T = Sigma I U" + by (auto simp: openin_subtopology openin_sum_topology_alt) + then obtain B where "\i. i\I \ openin (X i) (B i) \ U i = B i \ S i" + by metis + then show "openin (subtopology (sum_topology X I) (Sigma I S)) T" + by (auto simp: openin_subtopology openin_sum_topology_alt Teq) +qed + +lemma embedding_map_component_injection: + "i \ I \ embedding_map (X i) (sum_topology X I) (\x. (i,x))" + by (metis injective_open_imp_embedding_map continuous_map_component_injection + open_map_component_injection inj_onI prod.inject) + +end diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Analysis/T1_Spaces.thy --- a/src/HOL/Analysis/T1_Spaces.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Analysis/T1_Spaces.thy Sun May 07 22:51:23 2023 +0200 @@ -49,6 +49,9 @@ apply (simp add: closedin_contains_derived_set t1_space_derived_set_of_singleton) using t1_space_alt by auto +lemma t1_space_euclidean: "t1_space (euclidean :: 'a::metric_space topology)" + by (simp add: t1_space_closedin_singleton) + lemma closedin_t1_singleton: "\t1_space X; a \ topspace X\ \ closedin X {a}" by (simp add: t1_space_closedin_singleton) @@ -380,8 +383,7 @@ lemma Hausdorff_space_discrete_topology [simp]: "Hausdorff_space (discrete_topology U)" unfolding Hausdorff_space_def - apply safe - by (metis discrete_topology_unique_alt disjnt_empty2 disjnt_insert2 insert_iff mk_disjoint_insert topspace_discrete_topology) + by (metis Hausdorff_space_compact_sets Hausdorff_space_def compactin_discrete_topology equalityE openin_discrete_topology) lemma compactin_Int: "\Hausdorff_space X; compactin X S; compactin X T\ \ compactin X (S \ T)" @@ -395,6 +397,10 @@ "\Hausdorff_space X; finite S\ \ X derived_set_of S = {}" using Hausdorff_imp_t1_space t1_space_derived_set_of_finite by auto +lemma infinite_perfect_set: + "\Hausdorff_space X; S \ X derived_set_of S; S \ {}\ \ infinite S" + using derived_set_of_finite by blast + lemma derived_set_of_singleton: "Hausdorff_space X \ X derived_set_of {x} = {}" by (simp add: derived_set_of_finite) @@ -698,4 +704,47 @@ qed qed +lemma Hausdorff_space_closed_neighbourhood: + "Hausdorff_space X \ + (\x \ topspace X. \U C. openin X U \ closedin X C \ + Hausdorff_space(subtopology X C) \ x \ U \ U \ C)" (is "_ = ?rhs") +proof + assume R: ?rhs + show "Hausdorff_space X" + unfolding Hausdorff_space_def + proof clarify + fix x y + assume x: "x \ topspace X" and y: "y \ topspace X" and "x \ y" + obtain T C where *: "openin X T" "closedin X C" "x \ T" "T \ C" + and C: "Hausdorff_space (subtopology X C)" + by (meson R \x \ topspace X\) + show "\U V. openin X U \ openin X V \ x \ U \ y \ V \ disjnt U V" + proof (cases "y \ C") + case True + with * C obtain U V where U: "openin (subtopology X C) U" + and V: "openin (subtopology X C) V" + and "x \ U" "y \ V" "disjnt U V" + unfolding Hausdorff_space_def + by (smt (verit, best) \x \ y\ closedin_subset subsetD topspace_subtopology_subset) + then obtain U' V' where UV': "U = U' \ C" "openin X U'" "V = V' \ C" "openin X V'" + by (meson openin_subtopology) + have "disjnt (T \ U') V'" + using \disjnt U V\ UV' \T \ C\ by (force simp: disjnt_iff) + with \T \ C\ have "disjnt (T \ U') (V' \ (topspace X - C))" + unfolding disjnt_def by blast + moreover + have "openin X (T \ U')" + by (simp add: \openin X T\ \openin X U'\ openin_Int) + moreover have "openin X (V' \ (topspace X - C))" + using \closedin X C\ \openin X V'\ by auto + ultimately show ?thesis + using UV' \x \ T\ \x \ U\ \y \ V\ by blast + next + case False + with * y show ?thesis + by (force simp: closedin_def disjnt_def) + qed + qed +qed fastforce + end diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Archimedean_Field.thy --- a/src/HOL/Archimedean_Field.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Archimedean_Field.thy Sun May 07 22:51:23 2023 +0200 @@ -16,15 +16,10 @@ proof - have "Sup (uminus ` S) = - (Inf S)" proof (rule antisym) - show "- (Inf S) \ Sup (uminus ` S)" - apply (subst minus_le_iff) - apply (rule cInf_greatest [OF \S \ {}\]) - apply (subst minus_le_iff) - apply (rule cSup_upper) - apply force - using bdd - apply (force simp: abs_le_iff bdd_above_def) - done + have "\x. x \ S \ bdd_above (uminus ` S)" + using bdd by (force simp: abs_le_iff bdd_above_def) + then show "- (Inf S) \ Sup (uminus ` S)" + by (meson cInf_greatest [OF \S \ {}\] cSUP_upper minus_le_iff) next have *: "\x. x \ S \ Inf S \ x" by (meson abs_le_iff bdd bdd_below_def cInf_lower minus_le_iff) diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Conditionally_Complete_Lattices.thy --- a/src/HOL/Conditionally_Complete_Lattices.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Conditionally_Complete_Lattices.thy Sun May 07 22:51:23 2023 +0200 @@ -481,6 +481,12 @@ assume "X = {}" with lt_ex[of a] least show ?thesis by (auto simp: less_le_not_le) qed (intro cSup_eq_non_empty assms) +lemma cSup_unique: + fixes b :: "'a :: {conditionally_complete_lattice, no_bot}" + assumes "\c. (\x \ s. x \ c) \ b \ c" + shows "Sup s = b" + by (metis assms cSup_eq order.refl) + lemma cInf_eq: fixes a :: "'a :: {conditionally_complete_lattice, no_top}" assumes upper: "\x. x \ X \ a \ x" @@ -490,6 +496,12 @@ assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le) qed (intro cInf_eq_non_empty assms) +lemma cInf_unique: + fixes b :: "'a :: {conditionally_complete_lattice, no_top}" + assumes "\c. (\x \ s. x \ c) \ b \ c" + shows "Inf s = b" + by (meson assms cInf_eq order.refl) + class conditionally_complete_linorder = conditionally_complete_lattice + linorder begin diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Data_Structures/Leftist_Heap.thy --- a/src/HOL/Data_Structures/Leftist_Heap.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Data_Structures/Leftist_Heap.thy Sun May 07 22:51:23 2023 +0200 @@ -158,6 +158,8 @@ subsection "Complexity" +text \We count only the calls of the only recursive function: @{const merge}\ + text\Explicit termination argument: sum of sizes\ fun T_merge :: "'a::ord lheap \ 'a lheap \ nat" where @@ -168,11 +170,11 @@ else T_merge t1 r2) + 1" definition T_insert :: "'a::ord \ 'a lheap \ nat" where -"T_insert x t = T_merge (Node Leaf (x, 1) Leaf) t + 1" +"T_insert x t = T_merge (Node Leaf (x, 1) Leaf) t" fun T_del_min :: "'a::ord lheap \ nat" where -"T_del_min Leaf = 1" | -"T_del_min (Node l _ r) = T_merge l r + 1" +"T_del_min Leaf = 0" | +"T_del_min (Node l _ r) = T_merge l r" lemma T_merge_min_height: "ltree l \ ltree r \ T_merge l r \ min_height l + min_height r + 1" proof(induction l r rule: merge.induct) @@ -185,7 +187,7 @@ le_log2_of_power[OF min_height_size1[of r]] T_merge_min_height[of l r] assms by linarith -corollary T_insert_log: "ltree t \ T_insert x t \ log 2 (size1 t) + 3" +corollary T_insert_log: "ltree t \ T_insert x t \ log 2 (size1 t) + 2" using T_merge_log[of "Node Leaf (x, 1) Leaf" t] by(simp add: T_insert_def split: tree.split) @@ -203,15 +205,15 @@ qed corollary T_del_min_log: assumes "ltree t" - shows "T_del_min t \ 2 * log 2 (size1 t) + 1" + shows "T_del_min t \ 2 * log 2 (size1 t)" proof(cases t rule: tree2_cases) case Leaf thus ?thesis using assms by simp next case [simp]: (Node l _ _ r) - have "T_del_min t = T_merge l r + 1" by simp - also have "\ \ log 2 (size1 l) + log 2 (size1 r) + 2" + have "T_del_min t = T_merge l r" by simp + also have "\ \ log 2 (size1 l) + log 2 (size1 r) + 1" using \ltree t\ T_merge_log[of l r] by (auto simp del: T_merge.simps) - also have "\ \ 2 * log 2 (size1 t) + 1" + also have "\ \ 2 * log 2 (size1 t)" using ld_ld_1_less[of "size1 l" "size1 r"] by (simp) finally show ?thesis . qed diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Filter.thy --- a/src/HOL/Filter.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Filter.thy Sun May 07 22:51:23 2023 +0200 @@ -458,6 +458,12 @@ lemma False_imp_not_eventually: "(\x. \ P x ) \ \ trivial_limit net \ \ eventually (\x. P x) net" by (simp add: eventually_False) +lemma trivial_limit_eventually: "trivial_limit net \ eventually P net" + by simp + +lemma trivial_limit_eq: "trivial_limit net \ (\P. eventually P net)" + by (simp add: filter_eq_iff) + lemma eventually_Inf: "eventually P (Inf B) \ (\X\B. finite X \ eventually P (Inf X))" proof - let ?F = "\P. \X\B. finite X \ eventually P (Inf X)" diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Fun.thy --- a/src/HOL/Fun.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Fun.thy Sun May 07 22:51:23 2023 +0200 @@ -1031,6 +1031,12 @@ abbreviation strict_mono_on :: "'a set \ ('a \ 'b :: ord) \ bool" where "strict_mono_on A \ monotone_on A (<) (<)" +abbreviation antimono_on :: "'a set \ ('a \ 'b :: ord) \ bool" + where "antimono_on A \ monotone_on A (\) (\)" + +abbreviation strict_antimono_on :: "'a set \ ('a \ 'b :: ord) \ bool" + where "strict_antimono_on A \ monotone_on A (<) (>)" + lemma mono_on_def[no_atp]: "mono_on A f \ (\r s. r \ A \ s \ A \ r \ s \ f r \ f s)" by (auto simp add: monotone_on_def) @@ -1265,6 +1271,32 @@ "strict_mono_on A (f :: (_ :: linorder) \ _ :: preorder) \ mono_on A f" by (rule mono_onI, rule strict_mono_on_leD) +lemma mono_imp_strict_mono: + fixes f :: "'a::order \ 'b::order" + shows "\mono_on S f; inj_on f S\ \ strict_mono_on S f" + by (auto simp add: monotone_on_def order_less_le inj_on_eq_iff) + +lemma strict_mono_iff_mono: + fixes f :: "'a::linorder \ 'b::order" + shows "strict_mono_on S f \ mono_on S f \ inj_on f S" +proof + show "strict_mono_on S f \ mono_on S f \ inj_on f S" + by (simp add: strict_mono_on_imp_inj_on strict_mono_on_imp_mono_on) +qed (auto intro: mono_imp_strict_mono) + +lemma antimono_imp_strict_antimono: + fixes f :: "'a::order \ 'b::order" + shows "\antimono_on S f; inj_on f S\ \ strict_antimono_on S f" + by (auto simp add: monotone_on_def order_less_le inj_on_eq_iff) + +lemma strict_antimono_iff_antimono: + fixes f :: "'a::linorder \ 'b::order" + shows "strict_antimono_on S f \ antimono_on S f \ inj_on f S" +proof + show "strict_antimono_on S f \ antimono_on S f \ inj_on f S" + by (force simp add: monotone_on_def intro: linorder_inj_onI) +qed (auto intro: antimono_imp_strict_antimono) + lemma mono_compose: "mono Q \ mono (\i x. Q i (f x))" unfolding mono_def le_fun_def by auto diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Library/Set_Idioms.thy --- a/src/HOL/Library/Set_Idioms.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Library/Set_Idioms.thy Sun May 07 22:51:23 2023 +0200 @@ -557,4 +557,162 @@ by blast qed +lemma countable_union_of_empty [simp]: "(countable union_of P) {}" + by (simp add: union_of_empty) + +lemma countable_intersection_of_empty [simp]: "(countable intersection_of P) UNIV" + by (simp add: intersection_of_empty) + +lemma countable_union_of_inc: "P S \ (countable union_of P) S" + by (simp add: union_of_inc) + +lemma countable_intersection_of_inc: "P S \ (countable intersection_of P) S" + by (simp add: intersection_of_inc) + +lemma countable_union_of_complement: + "(countable union_of P) S \ (countable intersection_of (\S. P(-S))) (-S)" + (is "?lhs=?rhs") +proof + assume ?lhs + then obtain \ where "countable \" and \: "\ \ Collect P" "\\ = S" + by (metis union_of_def) + define \' where "\' \ (\C. -C) ` \" + have "\' \ {S. P (- S)}" "\\' = -S" + using \'_def \ by auto + then show ?rhs + unfolding intersection_of_def by (metis \'_def \countable \\ countable_image) +next + assume ?rhs + then obtain \ where "countable \" and \: "\ \ {S. P (- S)}" "\\ = -S" + by (metis intersection_of_def) + define \' where "\' \ (\C. -C) ` \" + have "\' \ Collect P" "\ \' = S" + using \'_def \ by auto + then show ?lhs + unfolding union_of_def + by (metis \'_def \countable \\ countable_image) +qed + +lemma countable_intersection_of_complement: + "(countable intersection_of P) S \ (countable union_of (\S. P(- S))) (- S)" + by (simp add: countable_union_of_complement) + +lemma countable_union_of_explicit: + assumes "P {}" + shows "(countable union_of P) S \ + (\T. (\n::nat. P(T n)) \ \(range T) = S)" (is "?lhs=?rhs") +proof + assume ?lhs + then obtain \ where "countable \" and \: "\ \ Collect P" "\\ = S" + by (metis union_of_def) + then show ?rhs + by (metis SUP_bot Sup_empty assms from_nat_into mem_Collect_eq range_from_nat_into subsetD) +next + assume ?rhs + then show ?lhs + by (metis countableI_type countable_image image_subset_iff mem_Collect_eq union_of_def) +qed + +lemma countable_union_of_ascending: + assumes empty: "P {}" and Un: "\T U. \P T; P U\ \ P(T \ U)" + shows "(countable union_of P) S \ + (\T. (\n. P(T n)) \ (\n. T n \ T(Suc n)) \ \(range T) = S)" (is "?lhs=?rhs") +proof + assume ?lhs + then obtain T where T: "\n::nat. P(T n)" "\(range T) = S" + by (meson empty countable_union_of_explicit) + have "P (\ (T ` {..n}))" for n + by (induction n) (auto simp: atMost_Suc Un T) + with T show ?rhs + by (rule_tac x="\n. \k\n. T k" in exI) force +next + assume ?rhs + then show ?lhs + using empty countable_union_of_explicit by auto +qed + +lemma countable_union_of_idem [simp]: + "countable union_of countable union_of P = countable union_of P" (is "?lhs=?rhs") +proof + fix S + show "(countable union_of countable union_of P) S = (countable union_of P) S" + proof + assume L: "?lhs S" + then obtain \ where "countable \" and \: "\ \ Collect (countable union_of P)" "\\ = S" + by (metis union_of_def) + then have "\U \ \. \\. countable \ \ \ \ Collect P \ U = \\" + by (metis Ball_Collect union_of_def) + then obtain \ where \: "\U \ \. countable (\ U) \ \ U \ Collect P \ U = \(\ U)" + by metis + have "countable (\ (\ ` \))" + using \ \countable \\ by blast + moreover have "\ (\ ` \) \ Collect P" + by (simp add: Sup_le_iff \) + moreover have "\ (\ (\ ` \)) = S" + by auto (metis Union_iff \ \(2))+ + ultimately show "?rhs S" + by (meson union_of_def) + qed (simp add: countable_union_of_inc) +qed + +lemma countable_intersection_of_idem [simp]: + "countable intersection_of countable intersection_of P = + countable intersection_of P" + by (force simp: countable_intersection_of_complement) + +lemma countable_union_of_Union: + "\countable \; \S. S \ \ \ (countable union_of P) S\ + \ (countable union_of P) (\ \)" + by (metis Ball_Collect countable_union_of_idem union_of_def) + +lemma countable_union_of_UN: + "\countable I; \i. i \ I \ (countable union_of P) (U i)\ + \ (countable union_of P) (\i\I. U i)" + by (metis (mono_tags, lifting) countable_image countable_union_of_Union imageE) + +lemma countable_union_of_Un: + "\(countable union_of P) S; (countable union_of P) T\ + \ (countable union_of P) (S \ T)" + by (smt (verit) Union_Un_distrib countable_Un le_sup_iff union_of_def) + +lemma countable_intersection_of_Inter: + "\countable \; \S. S \ \ \ (countable intersection_of P) S\ + \ (countable intersection_of P) (\ \)" + by (metis countable_intersection_of_idem intersection_of_def mem_Collect_eq subsetI) + +lemma countable_intersection_of_INT: + "\countable I; \i. i \ I \ (countable intersection_of P) (U i)\ + \ (countable intersection_of P) (\i\I. U i)" + by (metis (mono_tags, lifting) countable_image countable_intersection_of_Inter imageE) + +lemma countable_intersection_of_inter: + "\(countable intersection_of P) S; (countable intersection_of P) T\ + \ (countable intersection_of P) (S \ T)" + by (simp add: countable_intersection_of_complement countable_union_of_Un) + +lemma countable_union_of_Int: + assumes S: "(countable union_of P) S" and T: "(countable union_of P) T" + and Int: "\S T. P S \ P T \ P(S \ T)" + shows "(countable union_of P) (S \ T)" +proof - + obtain \ where "countable \" and \: "\ \ Collect P" "\\ = S" + using S by (metis union_of_def) + obtain \ where "countable \" and \: "\ \ Collect P" "\\ = T" + using T by (metis union_of_def) + have "\U V. \U \ \; V \ \\ \ (countable union_of P) (U \ V)" + using \ \ by (metis Ball_Collect countable_union_of_inc local.Int) + then have "(countable union_of P) (\U\\. \V\\. U \ V)" + by (meson \countable \\ \countable \\ countable_union_of_UN) + moreover have "S \ T = (\U\\. \V\\. U \ V)" + by (simp add: \ \) + ultimately show ?thesis + by presburger +qed + +lemma countable_intersection_of_union: + assumes S: "(countable intersection_of P) S" and T: "(countable intersection_of P) T" + and Un: "\S T. P S \ P T \ P(S \ T)" + shows "(countable intersection_of P) (S \ T)" + by (metis (mono_tags, lifting) Compl_Int S T Un compl_sup countable_intersection_of_complement countable_union_of_Int) + end diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Limits.thy --- a/src/HOL/Limits.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Limits.thy Sun May 07 22:51:23 2023 +0200 @@ -3251,4 +3251,21 @@ for f :: "real \ real" using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp: neq_iff) +lemma Lim_topological: + "(f \ l) net \ + trivial_limit net \ (\S. open S \ l \ S \ eventually (\x. f x \ S) net)" + unfolding tendsto_def trivial_limit_eq by auto + +lemma eventually_within_Un: + "eventually P (at x within (s \ t)) \ + eventually P (at x within s) \ eventually P (at x within t)" + unfolding eventually_at_filter + by (auto elim!: eventually_rev_mp) + +lemma Lim_within_Un: + "(f \ l) (at x within (s \ t)) \ + (f \ l) (at x within s) \ (f \ l) (at x within t)" + unfolding tendsto_def + by (auto simp: eventually_within_Un) + end diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Real.thy --- a/src/HOL/Real.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Real.thy Sun May 07 22:51:23 2023 +0200 @@ -1096,6 +1096,27 @@ by (metis reals_Archimedean3 dual_order.order_iff_strict le0 le_less_trans not_le x0 xc) +lemma inverse_Suc: "inverse (Suc n) > 0" + using of_nat_0_less_iff positive_imp_inverse_positive zero_less_Suc by blast + +lemma Archimedean_eventually_inverse: + fixes \::real shows "(\\<^sub>F n in sequentially. inverse (real (Suc n)) < \) \ 0 < \" + (is "?lhs=?rhs") +proof + assume ?lhs + then show ?rhs + unfolding eventually_at_top_dense using inverse_Suc order_less_trans by blast +next + assume ?rhs + then obtain N where "inverse (Suc N) < \" + using reals_Archimedean by blast + moreover have "inverse (Suc n) \ inverse (Suc N)" if "n \ N" for n + using inverse_Suc that by fastforce + ultimately show ?lhs + unfolding eventually_sequentially + using order_le_less_trans by blast +qed + subsection \Rationals\ lemma Rats_abs_iff[simp]: @@ -1381,10 +1402,34 @@ lemma forall_pos_mono_1: "(\d e::real. d < e \ P d \ P e) \ (\n. P (inverse (real (Suc n)))) \ 0 < e \ P e" - apply (rule forall_pos_mono) - apply auto - apply (metis Suc_pred of_nat_Suc) - done + using reals_Archimedean by blast + +lemma Archimedean_eventually_pow: + fixes x::real + assumes "1 < x" + shows "\\<^sub>F n in sequentially. b < x ^ n" +proof - + obtain N where "\n. n\N \ b < x ^ n" + by (metis assms le_less order_less_trans power_strict_increasing_iff real_arch_pow) + then show ?thesis + using eventually_sequentially by blast +qed + +lemma Archimedean_eventually_pow_inverse: + fixes x::real + assumes "\x\ < 1" "\ > 0" + shows "\\<^sub>F n in sequentially. \x^n\ < \" +proof (cases "x = 0") + case True + then show ?thesis + by (simp add: assms eventually_at_top_dense zero_power) +next + case False + then have "\\<^sub>F n in sequentially. inverse \ < inverse \x\ ^ n" + by (simp add: Archimedean_eventually_pow assms(1) one_less_inverse) + then show ?thesis + by eventually_elim (metis \\ > 0\ inverse_less_imp_less power_abs power_inverse) +qed subsection \Floor and Ceiling Functions from the Reals to the Integers\ diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Set.thy --- a/src/HOL/Set.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Set.thy Sun May 07 22:51:23 2023 +0200 @@ -1946,6 +1946,9 @@ lemma disjnt_Un2 [simp]: "disjnt C (A \ B) \ disjnt C A \ disjnt C B" by (auto simp: disjnt_def) +lemma disjnt_Diff1: "disjnt (X-Y) (U-V)" and disjnt_Diff2: "disjnt (U-V) (X-Y)" if "X \ V" + using that by (auto simp: disjnt_def) + lemma disjoint_image_subset: "\pairwise disjnt \; \X. X \ \ \ f X \ X\ \ pairwise disjnt (f `\)" unfolding disjnt_def pairwise_def by fast diff -r c1b8fdd6588a -r e30a56be9c7b src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy Sun May 07 14:25:41 2023 +0200 +++ b/src/HOL/Set_Interval.thy Sun May 07 22:51:23 2023 +0200 @@ -280,8 +280,8 @@ context order begin -lemma atLeastatMost_empty[simp]: - "b < a \ {a..b} = {}" +lemma atLeastatMost_empty[simp]: "b < a \ {a..b} = {}" + and atLeastatMost_empty'[simp]: "\ a \ b \ {a..b} = {}" by(auto simp: atLeastAtMost_def atLeast_def atMost_def) lemma atLeastLessThan_empty[simp]: diff -r c1b8fdd6588a -r e30a56be9c7b src/Pure/Isar/code.ML --- a/src/Pure/Isar/code.ML Sun May 07 14:25:41 2023 +0200 +++ b/src/Pure/Isar/code.ML Sun May 07 22:51:23 2023 +0200 @@ -758,29 +758,24 @@ fun meta_rewrite thy = Local_Defs.meta_rewrite_rule (Proof_Context.init_global thy); -fun expand_eta thy k thm = - let - val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm; - val (_, args) = strip_comb lhs; - val l = if k = ~1 - then (length o fst o strip_abs) rhs - else Int.max (0, k - length args); - val (raw_vars, _) = Term.strip_abs_eta l rhs; - val vars = burrow_fst (Name.variant_list (map (fst o fst) (Term.add_vars lhs []))) - raw_vars; - fun expand (v, ty) thm = Drule.fun_cong_rule thm - (Thm.global_cterm_of thy (Var ((v, 0), ty))); - in - thm - |> fold expand vars - |> Conv.fconv_rule Drule.beta_eta_conversion - end; - fun same_arity thy thms = let - val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals; - val k = fold (Integer.max o num_args_of o Thm.prop_of) thms 0; - in map (expand_eta thy k) thms end; + val lhs_rhss = map (Logic.dest_equals o Thm.plain_prop_of) thms; + val k = fold (Integer.max o length o snd o strip_comb o fst) lhs_rhss 0; + fun expand_eta (lhs, rhs) thm = + let + val l = k - length (snd (strip_comb lhs)); + val (raw_vars, _) = Term.strip_abs_eta l rhs; + val vars = burrow_fst (Name.variant_list (map (fst o fst) (Term.add_vars lhs []))) + raw_vars; + fun expand (v, ty) thm = Drule.fun_cong_rule thm + (Thm.global_cterm_of thy (Var ((v, 0), ty))); + in + thm + |> fold expand vars + |> Conv.fconv_rule Drule.beta_eta_conversion + end; + in map2 expand_eta lhs_rhss thms end; fun mk_desymbolization pre post mk vs = let diff -r c1b8fdd6588a -r e30a56be9c7b src/Tools/Code/code_thingol.ML --- a/src/Tools/Code/code_thingol.ML Sun May 07 14:25:41 2023 +0200 +++ b/src/Tools/Code/code_thingol.ML Sun May 07 22:51:23 2023 +0200 @@ -587,25 +587,6 @@ | is_undefined_clause ctxt _ = false; -fun satisfied_app wanted (ty, ts) = - let - val given = length ts; - val delta = wanted - given; - val rty = (drop wanted o binder_types) ty ---> body_type ty; - in - if delta = 0 then - (([], (ts, rty)), []) - else if delta < 0 then - let - val (ts1, ts2) = chop wanted ts - in (([], (ts1, rty)), ts2) end - else - let - val tys = (take delta o drop given o binder_types) ty; - val vs_tys = invent_params ((fold o fold_aterms) Term.declare_term_frees ts) tys; - in ((vs_tys, (ts @ map Free vs_tys, rty)), []) end - end - fun ensure_tyco ctxt algbr eqngr permissive tyco = let val thy = Proof_Context.theory_of ctxt; @@ -751,41 +732,34 @@ then translation_error ctxt permissive some_thm deps "Abstraction violation" ("constant " ^ Code.string_of_const thy c) else () - in translate_const_proper ctxt algbr eqngr permissive some_thm (c, ty) (deps, program) end -and translate_const_proper ctxt algbr eqngr permissive some_thm (c, ty) = - let - val thy = Proof_Context.theory_of ctxt; val (annotate, ty') = dest_tagged_type ty; val typargs = Sign.const_typargs thy (c, ty'); val sorts = Code_Preproc.sortargs eqngr c; val (dom, range) = Term.strip_type ty'; in - ensure_const ctxt algbr eqngr permissive c - ##>> fold_map (translate_typ ctxt algbr eqngr permissive) typargs - ##>> fold_map (translate_dicts ctxt algbr eqngr permissive some_thm) (typargs ~~ sorts) - ##>> fold_map (translate_typ ctxt algbr eqngr permissive) (range :: dom) - #>> (fn (((c, typargs), dss), range :: dom) => + (deps, program) + |> ensure_const ctxt algbr eqngr permissive c + ||>> fold_map (translate_typ ctxt algbr eqngr permissive) typargs + ||>> fold_map (translate_dicts ctxt algbr eqngr permissive some_thm) (typargs ~~ sorts) + ||>> fold_map (translate_typ ctxt algbr eqngr permissive) (range :: dom) + |>> (fn (((c, typargs), dss), range :: dom) => { sym = Constant c, typargs = typargs, dicts = dss, dom = dom, range = range, annotation = if annotate then SOME (dom `--> range) else NONE }) end -and translate_case ctxt algbr eqngr permissive some_thm (t_pos, []) (c_ty, ts) = +and translate_case ctxt algbr eqngr permissive some_thm (t_pos, []) _ (const as { dom, ... }, ts) = let fun project_term xs = nth xs t_pos; val project_clause = the_single o nth_drop t_pos; - val ty_case = project_term (binder_types (snd c_ty)); - fun distill_clauses ty_case t = + val ty_case = project_term dom; + fun distill_clauses t = map (fn ([pat], body) => (pat, body)) (distill_minimized_clause [ty_case] t) in - translate_const ctxt algbr eqngr permissive some_thm NONE c_ty - ##>> fold_map (translate_term ctxt algbr eqngr permissive some_thm NONE) ts - ##>> translate_typ ctxt algbr eqngr permissive ty_case - #>> (fn ((const, ts), ty_case) => - ICase { term = project_term ts, typ = ty_case, - clauses = (filter_out (is_undefined_clause ctxt) o distill_clauses ty_case o project_clause) ts, - primitive = IConst const `$$ ts }) + pair (ICase { term = project_term ts, typ = ty_case, + clauses = (filter_out (is_undefined_clause ctxt) o distill_clauses o project_clause) ts, + primitive = IConst const `$$ ts }) end - | translate_case ctxt algbr eqngr permissive some_thm (t_pos, case_pats) (c_ty, ts) = + | translate_case ctxt algbr eqngr permissive some_thm (t_pos, case_pats) ty (const as { dom, ... }, ts) = let fun project_term xs = nth xs t_pos; fun project_cases xs = @@ -793,40 +767,39 @@ |> nth_drop t_pos |> curry (op ~~) case_pats |> map_filter (fn (NONE, _) => NONE | (SOME _, x) => SOME x); - val ty_case = project_term (binder_types (snd c_ty)); - val constrs = map_filter I case_pats ~~ project_cases ts - |> map (fn ((c, n), t) => - ((c, (take n o binder_types o fastype_of_tagged_term) t ---> ty_case), n)); + val tys = take (length case_pats + 1) (binder_types ty); + val ty_case = project_term tys; + val ty_case' = project_term dom; + val constrs = map_filter I case_pats ~~ project_cases tys + |> map (fn ((c, n), ty) => + ((c, (take n o binder_types) ty ---> ty_case), n)); fun distill_clauses constrs ts_clause = maps (fn ((constr as { dom = tys, ... }, n), t) => map (fn (pat_args, body) => (IConst constr `$$ pat_args, body)) (distill_minimized_clause (take n tys) t)) (constrs ~~ ts_clause); in - translate_const ctxt algbr eqngr permissive some_thm NONE c_ty - ##>> fold_map (translate_term ctxt algbr eqngr permissive some_thm NONE) ts - ##>> translate_typ ctxt algbr eqngr permissive ty_case - ##>> fold_map (fn (c_ty, n) => + fold_map (fn (c_ty, n) => translate_const ctxt algbr eqngr permissive some_thm NONE c_ty #>> rpair n) constrs - #>> (fn (((const, ts), ty_case), constrs) => - ICase { term = project_term ts, typ = ty_case, + #>> (fn constrs => + ICase { term = project_term ts, typ = ty_case', clauses = (filter_out (is_undefined_clause ctxt) o distill_clauses constrs o project_cases) ts, primitive = IConst const `$$ ts }) end -and translate_app_case ctxt algbr eqngr permissive some_thm pattern_schema c_ty ((vs_tys, (ts1, rty)), ts2) = - fold_map (fn (v, ty) => translate_typ ctxt algbr eqngr permissive ty #>> pair (SOME v)) vs_tys - ##>> translate_case ctxt algbr eqngr permissive some_thm pattern_schema (c_ty, ts1) - ##>> translate_typ ctxt algbr eqngr permissive rty - ##>> fold_map (translate_term ctxt algbr eqngr permissive some_thm NONE) ts2 - #>> (fn (((vs_tys, t), rty), ts) => (vs_tys `|==> (t, rty)) `$$ ts) and translate_app ctxt algbr eqngr permissive some_thm some_abs (c_ty as (c, ty), ts) = - case Code.get_case_schema (Proof_Context.theory_of ctxt) c of - SOME (wanted, pattern_schema) => - translate_app_case ctxt algbr eqngr permissive some_thm pattern_schema c_ty (satisfied_app wanted (ty, ts)) - | NONE => - translate_const ctxt algbr eqngr permissive some_thm some_abs c_ty - ##>> fold_map (translate_term ctxt algbr eqngr permissive some_thm NONE) ts - #>> (fn (const, ts) => IConst const `$$ ts) + translate_const ctxt algbr eqngr permissive some_thm some_abs c_ty + ##>> fold_map (translate_term ctxt algbr eqngr permissive some_thm NONE) ts + #-> (fn (const, ts) => case Code.get_case_schema (Proof_Context.theory_of ctxt) c of + SOME (wanted, pattern_schema) => + let + val ((vs_tys, (ts1, rty)), ts2) = satisfied_application wanted (const, ts); + val (_, ty') = dest_tagged_type ty; + in + translate_case ctxt algbr eqngr permissive some_thm pattern_schema ty' (const, ts1) + #>> (fn t => (vs_tys `|==> (t, rty)) `$$ ts2) + end + | NONE => + pair (IConst const `$$ ts)) and translate_tyvar_sort ctxt (algbr as (proj_sort, _)) eqngr permissive (v, sort) = fold_map (ensure_class ctxt algbr eqngr permissive) (proj_sort sort) #>> (fn sort => (unprefix "'" v, sort))