# HG changeset patch # User paulson # Date 1687863391 -3600 # Node ID 759c71cdaf2ab60a2c280a712bb331aaca995649 # Parent b0ad3aba48f1fed35635054ccf96935998e1d4b4 More metric space material diff -r b0ad3aba48f1 -r 759c71cdaf2a src/HOL/Analysis/Urysohn.thy --- a/src/HOL/Analysis/Urysohn.thy Mon Jun 26 14:38:26 2023 +0100 +++ b/src/HOL/Analysis/Urysohn.thy Tue Jun 27 11:56:31 2023 +0100 @@ -1,8 +1,8 @@ -(* Title: HOL/Analysis/Arcwise_Connected.thy +(* Title: HOL/Analysis/Urysohn.thy Authors: LC Paulson, based on material from HOL Light *) -section \Urysohn lemma and its Consequences\ +section \The Urysohn lemma, its consequences and other advanced material about metric spaces\ theory Urysohn imports Abstract_Topological_Spaces Abstract_Metric_Spaces Infinite_Sum Arcwise_Connected @@ -71,7 +71,7 @@ have **: "r \ f x" if r: "r \ dyadics \ {0..1}" "x \ G r" for r x proof (rule *) show "r \ s" if "s \ dyadics \ {0..1}" and "x \ G s" for s :: real - using that r G [of s r] by (force simp add: dest: closure_of_subset openin_subset) + using that r G [of s r] by (force simp: dest: closure_of_subset openin_subset) qed (use that in force) have "\U. openin X U \ x \ U \ (\y \ U. \f y - f x\ < \)" @@ -133,7 +133,7 @@ then show "x \ G'" by (simp add: G'_def that) show "\y \ G'. \f y - f x\ < \" - using ** f_le1 in_closure_of r by (fastforce simp add: True G'_def) + using ** f_le1 in_closure_of r by (fastforce simp: True G'_def) qed next case False @@ -161,7 +161,7 @@ qed qed then have contf: "continuous_map X (top_of_set {0..1}) f" - by (force simp add: Met_TC.continuous_map_to_metric dist_real_def continuous_map_in_subtopology fim simp flip: mtopology_is_euclidean) + by (force simp: Met_TC.continuous_map_to_metric dist_real_def continuous_map_in_subtopology fim simp flip: mtopology_is_euclidean) define g where "g \ \x. a + (b - a) * f x" show thesis proof @@ -328,7 +328,7 @@ \ \i X closure_of k i" by (metis in_closure_of linorder_not_less opek openin_subset subsetD) then show "disjnt U V" - by (force simp add: U_def V_def disjnt_iff) + by (force simp: U_def V_def disjnt_iff) qed qed qed @@ -1117,7 +1117,7 @@ using Urysohn_lemma [of "subtopology X M" K "M-U" 0 1] by auto then obtain g::"'a\real" where contg: "continuous_map (subtopology X M) euclidean g" and gim: "g ` M \ {0..1}" and g0: "\x. x \ K \ g x = 0" and g1: "\x. \x \ M; x \ U\ \ g x = 1" - using \M \ topspace X\ by (force simp add: continuous_map_in_subtopology image_subset_iff) + using \M \ topspace X\ by (force simp: continuous_map_in_subtopology image_subset_iff) show "\f::'a\real. continuous_map X (top_of_set {0..1}) f \ f x = 0 \ f ` S \ {1}" proof (intro exI conjI) show "continuous_map X (top_of_set {0..1}) (\x. if x \ M then g x else 1)" @@ -1703,7 +1703,7 @@ (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" - by (force simp add: topspace_def openin_Alexandroff_compactification) + by (force simp: topspace_def openin_Alexandroff_compactification) have "None \ topspace (Alexandroff_compactification X)" by (meson closedin_empty compactin_empty insert_subset openin_Alexandroff_compactification openin_subset) moreover have "Some x \ topspace (Alexandroff_compactification X)" @@ -1877,7 +1877,7 @@ then show ?thesis using t1_space_one_point_compactification [of "Alexandroff_compactification X" None] using embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_t1_space - by (fastforce simp add: compactin_subtopology closedin_Alexandroff_I closedin_subtopology) + by (fastforce simp: compactin_subtopology closedin_Alexandroff_I closedin_subtopology) qed @@ -2130,7 +2130,7 @@ qed auto then show ?thesis unfolding Hausdorff_space_closedin_diagonal closure_of_subset_eq [symmetric] - by (force simp add: closure_of_def) + by (force simp: closure_of_def) qed qed next @@ -2237,4 +2237,2065 @@ using homeomorphic_space_sym homeomorphic_space homeomorphic_map_def by blast qed +subsection\Extending continuous maps "pointwise" in a regular space\ + +lemma continuous_map_on_intermediate_closure_of: + assumes Y: "regular_space Y" + and T: "T \ X closure_of S" + and f: "\t. t \ T \ limitin Y f (f t) (atin_within X t S)" + shows "continuous_map (subtopology X T) Y f" +proof (clarsimp simp add: continuous_map_atin) + fix a + assume "a \ topspace X" and "a \ T" + have "f ` T \ topspace Y" + by (metis f image_subsetI limitin_topspace) + have "\\<^sub>F x in atin_within X a T. f x \ W" + if W: "openin Y W" "f a \ W" for W + proof - + obtain V C where "openin Y V" "closedin Y C" "f a \ V" "V \ C" "C \ W" + by (metis Y W neighbourhood_base_of neighbourhood_base_of_closedin) + have "\\<^sub>F x in atin_within X a S. f x \ V" + by (metis \a \ T\ \f a \ V\ \openin Y V\ f limitin_def) + then obtain U where "openin X U" "a \ U" and U: "\x \ U - {a}. x \ S \ f x \ V" + by (smt (verit) Diff_iff \a \ topspace X\ eventually_atin_within insert_iff) + moreover have "f z \ W" if "z \ U" "z \ a" "z \ T" for z + proof - + have "z \ topspace X" + using \openin X U\ openin_subset \z \ U\ by blast + then have "f z \ topspace Y" + using \f ` T \ topspace Y\ \z \ T\ by blast + { assume "f z \ topspace Y" "f z \ C" + then have "\\<^sub>F x in atin_within X z S. f x \ topspace Y - C" + by (metis Diff_iff \closedin Y C\ closedin_def f limitinD \z \ T\) + then obtain U' where U': "openin X U'" "z \ U'" + "\x. x \ U' - {z} \ x \ S \ f x \ C" + by (smt (verit) Diff_iff \z \ topspace X\ eventually_atin_within insertCI) + then have *: "\D. z \ D \ openin X D \ \y. y \ S \ y \ D" + by (meson T in_closure_of subsetD \z \ T\) + have False + using * [of "U \ U'"] U' U \V \ C\ \f a \ V\ \f z \ C\ \openin X U\ that + by blast + } + then show ?thesis + using \C \ W\ \f z \ topspace Y\ by auto + qed + ultimately have "\U. openin X U \ a \ U \ (\x \ U - {a}. x \ T \ f x \ W)" + by blast + then show ?thesis + using eventually_atin_within by fastforce + qed + then show "limitin Y f (f a) (atin (subtopology X T) a)" + by (metis \a \ T\ atin_subtopology_within f limitin_def) +qed + + +lemma continuous_map_on_intermediate_closure_of_eq: + assumes "regular_space Y" "S \ T" and Tsub: "T \ X closure_of S" + shows "continuous_map (subtopology X T) Y f \ (\t \ T. limitin Y f (f t) (atin_within X t S))" + (is "?lhs \ ?rhs") +proof + assume L: ?lhs + show ?rhs + proof (clarsimp simp add: continuous_map_atin) + fix x + assume "x \ T" + with L Tsub closure_of_subset_topspace + have "limitin Y f (f x) (atin (subtopology X T) x)" + by (fastforce simp: continuous_map_atin) + then show "limitin Y f (f x) (atin_within X x S)" + using \x \ T\ \S \ T\ + by (force simp: limitin_def atin_subtopology_within eventually_atin_within) + qed +next + show "?rhs \ ?lhs" + using assms continuous_map_on_intermediate_closure_of by blast +qed + +lemma continuous_map_extension_pointwise_alt: + assumes \
: "regular_space Y" "S \ T" "T \ X closure_of S" + and f: "continuous_map (subtopology X S) Y f" + and lim: "\t. t \ T-S \ \l. limitin Y f l (atin_within X t S)" + obtains g where "continuous_map (subtopology X T) Y g" "\x. x \ S \ g x = f x" +proof - + obtain g where g: "\t. t \ T \ t \ S \ limitin Y f (g t) (atin_within X t S)" + by (metis Diff_iff lim) + let ?h = "\x. if x \ S then f x else g x" + show thesis + proof + have T: "T \ topspace X" + using \
closure_of_subset_topspace by fastforce + have "limitin Y ?h (f t) (atin_within X t S)" if "t \ T" "t \ S" for t + proof - + have "limitin Y f (f t) (atin_within X t S)" + by (meson T f limit_continuous_map_within subset_eq that) + then show ?thesis + by (simp add: eventually_atin_within limitin_def) + qed + moreover have "limitin Y ?h (g t) (atin_within X t S)" if "t \ T" "t \ S" for t + by (smt (verit, del_insts) eventually_atin_within g limitin_def that) + ultimately show "continuous_map (subtopology X T) Y ?h" + unfolding continuous_map_on_intermediate_closure_of_eq [OF \
] + by (auto simp: \
atin_subtopology_within) + qed auto +qed + + +lemma continuous_map_extension_pointwise: + assumes "regular_space Y" "S \ T" and Tsub: "T \ X closure_of S" + and ex: " \x. x \ T \ \g. continuous_map (subtopology X (insert x S)) Y g \ + (\x \ S. g x = f x)" + obtains g where "continuous_map (subtopology X T) Y g" "\x. x \ S \ g x = f x" +proof (rule continuous_map_extension_pointwise_alt) + show "continuous_map (subtopology X S) Y f" + proof (clarsimp simp add: continuous_map_atin) + fix t + assume "t \ topspace X" and "t \ S" + then obtain g where g: "limitin Y g (g t) (atin (subtopology X (insert t S)) t)" and gf: "\x \ S. g x = f x" + by (metis Int_iff \S \ T\ continuous_map_atin ex inf.orderE insert_absorb topspace_subtopology) + with \t \ S\ show "limitin Y f (f t) (atin (subtopology X S) t)" + by (simp add: limitin_def atin_subtopology_within_if eventually_atin_within gf insert_absorb) + qed + show "\l. limitin Y f l (atin_within X t S)" if "t \ T - S" for t + proof - + obtain g where g: "continuous_map (subtopology X (insert t S)) Y g" and gf: "\x \ S. g x = f x" + using \S \ T\ ex \t \ T - S\ by force + then have "limitin Y g (g t) (atin_within X t (insert t S))" + using Tsub in_closure_of limit_continuous_map_within that by fastforce + then show ?thesis + unfolding limitin_def + by (smt (verit) eventually_atin_within gf subsetD subset_insertI) + qed +qed (use assms in auto) + + +subsection\Extending Cauchy continuous functions to the closure\ + +lemma Cauchy_continuous_map_extends_to_continuous_closure_of_aux: + assumes m2: "mcomplete_of m2" and f: "Cauchy_continuous_map (submetric m1 S) m2 f" + and "S \ mspace m1" + obtains g + where "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S)) + (mtopology_of m2) g" "\x. x \ S \ g x = f x" +proof (rule continuous_map_extension_pointwise_alt) + interpret L: Metric_space12 "mspace m1" "mdist m1" "mspace m2" "mdist m2" + by (simp add: Metric_space12_mspace_mdist) + interpret S: Metric_space "S \ mspace m1" "mdist m1" + by (simp add: L.M1.subspace) + show "regular_space (mtopology_of m2)" + by (simp add: Metric_space.regular_space_mtopology mtopology_of_def) + show "S \ mtopology_of m1 closure_of S" + by (simp add: assms(3) closure_of_subset) + show "continuous_map (subtopology (mtopology_of m1) S) (mtopology_of m2) f" + by (metis Cauchy_continuous_imp_continuous_map f mtopology_of_submetric) + fix a + assume a: "a \ mtopology_of m1 closure_of S - S" + then obtain \ where ran\: "range \ \ S" "range \ \ mspace m1" + and lim\: "limitin L.M1.mtopology \ a sequentially" + by (force simp: mtopology_of_def L.M1.closure_of_sequentially) + then have "L.M1.MCauchy \" + by (simp add: L.M1.convergent_imp_MCauchy mtopology_of_def) + then have "L.M2.MCauchy (f \ \)" + using f ran\ by (simp add: Cauchy_continuous_map_def L.M1.subspace Metric_space.MCauchy_def) + then obtain l where l: "limitin L.M2.mtopology (f \ \) l sequentially" + by (meson L.M2.mcomplete_def m2 mcomplete_of_def) + have "limitin L.M2.mtopology f l (atin_within L.M1.mtopology a S)" + unfolding L.limit_atin_sequentially_within imp_conjL + proof (intro conjI strip) + show "l \ mspace m2" + using L.M2.limitin_mspace l by blast + fix \ + assume "range \ \ S \ mspace m1 - {a}" and lim\: "limitin L.M1.mtopology \ a sequentially" + then have ran\: "range \ \ S" "range \ \ mspace m1" "\n. \ n \ a" + by auto + have "a \ mspace m1" + using L.M1.limitin_mspace lim\ by auto + have "S.MCauchy \" "S.MCauchy \" + using L.M1.convergent_imp_MCauchy L.M1.MCauchy_def S.MCauchy_def lim\ ran\ lim\ ran\ by force+ + then have "L.M2.MCauchy (f \ \)" "L.M2.MCauchy (f \ \)" + using f by (auto simp: Cauchy_continuous_map_def) + then have ran_f: "range (\x. f (\ x)) \ mspace m2" "range (\x. f (\ x)) \ mspace m2" + by (auto simp: L.M2.MCauchy_def) + have "(\n. mdist m2 (f (\ n)) l) \ 0" + proof (rule Lim_null_comparison) + have "mdist m2 (f (\ n)) l \ mdist m2 (f (\ n)) l + mdist m2 (f (\ n)) (f (\ n))" for n + using \l \ mspace m2\ ran_f L.M2.triangle'' by (smt (verit, best) range_subsetD) + then show "\\<^sub>F n in sequentially. norm (mdist m2 (f (\ n)) l) \ mdist m2 (f (\ n)) l + mdist m2 (f (\ n)) (f (\ n))" + by force + define \ where "\ \ \n. if even n then \ (n div 2) else \ (n div 2)" + have "(\n. mdist m1 (\ n) (\ n)) \ 0" + proof (rule Lim_null_comparison) + show "\\<^sub>F n in sequentially. norm (mdist m1 (\ n) (\ n)) \ mdist m1 (\ n) a + mdist m1 (\ n) a" + using L.M1.triangle' [of _ a] ran\ ran\ \a \ mspace m1\ by (simp add: range_subsetD) + have "(\n. mdist m1 (\ n) a) \ 0" + using L.M1.limitin_metric_dist_null lim\ by blast + moreover have "(\n. mdist m1 (\ n) a) \ 0" + using L.M1.limitin_metric_dist_null lim\ by blast + ultimately show "(\n. mdist m1 (\ n) a + mdist m1 (\ n) a) \ 0" + by (simp add: tendsto_add_zero) + qed + with \S.MCauchy \\ \S.MCauchy \\ have "S.MCauchy \" + by (simp add: S.MCauchy_interleaving_gen \_def) + then have "L.M2.MCauchy (f \ \)" + by (metis Cauchy_continuous_map_def f mdist_submetric mspace_submetric) + then have "(\n. mdist m2 (f (\ n)) (f (\ n))) \ 0" + using L.M2.MCauchy_interleaving_gen [of "f \ \" "f \ \"] + by (simp add: if_distrib \_def o_def cong: if_cong) + moreover have "\\<^sub>F n in sequentially. f (\ n) \ mspace m2 \ (\x. mdist m2 (f (\ x)) l) \ 0" + using l by (auto simp: L.M2.limitin_metric_dist_null \l \ mspace m2\) + ultimately show "(\n. mdist m2 (f (\ n)) l + mdist m2 (f (\ n)) (f (\ n))) \ 0" + by (metis (mono_tags) tendsto_add_zero eventually_sequentially order_refl) + qed + with ran_f show "limitin L.M2.mtopology (f \ \) l sequentially" + by (auto simp: L.M2.limitin_metric_dist_null eventually_sequentially \l \ mspace m2\) + qed + then show "\l. limitin (mtopology_of m2) f l (atin_within (mtopology_of m1) a S)" + by (force simp: mtopology_of_def) +qed auto + + +lemma Cauchy_continuous_map_extends_to_continuous_closure_of: + assumes "mcomplete_of m2" + and f: "Cauchy_continuous_map (submetric m1 S) m2 f" + obtains g + where "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S)) + (mtopology_of m2) g" "\x. x \ S \ g x = f x" +proof - + obtain g where cmg: + "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of (mspace m1 \ S))) + (mtopology_of m2) g" + and gf: "(\x \ mspace m1 \ S. g x = f x)" + using Cauchy_continuous_map_extends_to_continuous_closure_of_aux assms + by (metis inf_commute inf_le2 submetric_restrict) + define h where "h \ \x. if x \ topspace(mtopology_of m1) then g x else f x" + show thesis + proof + show "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S)) + (mtopology_of m2) h" + unfolding h_def + proof (rule continuous_map_eq) + show "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S)) (mtopology_of m2) g" + by (metis closure_of_restrict cmg topspace_mtopology_of) + qed auto + qed (auto simp: gf h_def) +qed + + +lemma Cauchy_continuous_map_extends_to_continuous_intermediate_closure_of: + assumes "mcomplete_of m2" + and f: "Cauchy_continuous_map (submetric m1 S) m2 f" + and T: "T \ mtopology_of m1 closure_of S" + obtains g + where "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) g" + "(\x \ S. g x = f x)" + by (metis Cauchy_continuous_map_extends_to_continuous_closure_of T assms(1) continuous_map_from_subtopology_mono f) + +text \Technical lemma helpful for porting particularly ugly HOL Light proofs\ +lemma all_in_closure_of: + assumes P: "\x \ S. P x" and clo: "closedin X {x \ topspace X. P x}" + shows "\x \ X closure_of S. P x" +proof - + have *: "topspace X \ S \ {x \ topspace X. P x}" + using P by auto + show ?thesis + using closure_of_minimal [OF * clo] closure_of_restrict by fastforce +qed + +lemma Lipschitz_continuous_map_on_intermediate_closure_aux: + assumes lcf: "Lipschitz_continuous_map (submetric m1 S) m2 f" + and "S \ T" and Tsub: "T \ (mtopology_of m1) closure_of S" + and cmf: "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f" + and "S \ mspace m1" + shows "Lipschitz_continuous_map (submetric m1 T) m2 f" +proof - + interpret L: Metric_space12 "mspace m1" "mdist m1" "mspace m2" "mdist m2" + by (simp add: Metric_space12_mspace_mdist) + interpret S: Metric_space "S \ mspace m1" "mdist m1" + by (simp add: L.M1.subspace) + have "T \ mspace m1" + using Tsub by (auto simp: mtopology_of_def closure_of_def) + show ?thesis + unfolding Lipschitz_continuous_map_pos + proof + show "f ` mspace (submetric m1 T) \ mspace m2" + by (metis cmf Metric_space.metric_continuous_map Metric_space_mspace_mdist mtopology_of_def mtopology_of_submetric) + define X where "X \ prod_topology (subtopology L.M1.mtopology T) (subtopology L.M1.mtopology T)" + obtain B::real where "B > 0" and B: "\(x,y) \ S\S. mdist m2 (f x) (f y) \ B * mdist m1 x y" + using lcf \S \ mspace m1\ by (force simp: Lipschitz_continuous_map_pos) + have eq: "{z \ A. case z of (x,y) \ p x y \ B * q x y} = {z \ A. ((\(x,y). B * q x y - p x y)z) \ {0..}}" + for p q and A::"('a*'a)set" + by auto + have clo: "closedin X {z \ topspace X. case z of (x, y) \ mdist m2 (f x) (f y) \ B * mdist m1 x y}" + unfolding eq + proof (rule closedin_continuous_map_preimage) + have *: "continuous_map X L.M2.mtopology (f \ fst)" "continuous_map X L.M2.mtopology (f \ snd)" + using cmf by (auto simp: mtopology_of_def X_def intro: continuous_map_compose continuous_map_fst continuous_map_snd) + then show "continuous_map X euclidean (\x. case x of (x, y) \ B * mdist m1 x y - mdist m2 (f x) (f y))" + unfolding case_prod_unfold + proof (intro continuous_intros; simp add: mtopology_of_def o_def) + show "continuous_map X L.M1.mtopology fst" "continuous_map X L.M1.mtopology snd" + by (simp_all add: X_def continuous_map_subtopology_fst continuous_map_subtopology_snd flip: subtopology_Times) + qed + qed auto + have "mdist m2 (f x) (f y) \ B * mdist m1 x y" if "x \ T" "y \ T" for x y + using all_in_closure_of [OF B clo] \S \ T\ Tsub + by (fastforce simp: X_def subset_iff closure_of_Times closure_of_subtopology inf.absorb2 + mtopology_of_def that) + then show "\B>0. \x\mspace (submetric m1 T). + \y\mspace (submetric m1 T). + mdist m2 (f x) (f y) \ B * mdist (submetric m1 T) x y" + using \0 < B\ by auto + qed +qed + + +lemma Lipschitz_continuous_map_on_intermediate_closure: + assumes "Lipschitz_continuous_map (submetric m1 S) m2 f" + and "S \ T" "T \ (mtopology_of m1) closure_of S" + and "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f" + shows "Lipschitz_continuous_map (submetric m1 T) m2 f" + by (metis Lipschitz_continuous_map_on_intermediate_closure_aux assms closure_of_subset_topspace subset_trans topspace_mtopology_of) + + +lemma Lipschitz_continuous_map_extends_to_closure_of: + assumes m2: "mcomplete_of m2" + and f: "Lipschitz_continuous_map (submetric m1 S) m2 f" + obtains g + where "Lipschitz_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g" + "\x. x \ S \ g x = f x" +proof - + obtain g + where g: "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S)) + (mtopology_of m2) g" "(\x \ S. g x = f x)" + by (metis Cauchy_continuous_map_extends_to_continuous_closure_of Lipschitz_imp_Cauchy_continuous_map f m2) + have "Lipschitz_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g" + proof (rule Lipschitz_continuous_map_on_intermediate_closure) + show "Lipschitz_continuous_map (submetric m1 (mspace m1 \ S)) m2 g" + by (smt (verit, best) IntD2 Lipschitz_continuous_map_eq f g(2) inf_commute mspace_submetric submetric_restrict) + show "mspace m1 \ S \ mtopology_of m1 closure_of S" + using closure_of_subset_Int by force + show "mtopology_of m1 closure_of S \ mtopology_of m1 closure_of (mspace m1 \ S)" + by (metis closure_of_restrict subset_refl topspace_mtopology_of) + show "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S)) (mtopology_of m2) g" + by (simp add: g) + qed + with g that show thesis + by metis +qed + + +lemma Lipschitz_continuous_map_extends_to_intermediate_closure_of: + assumes "mcomplete_of m2" + and "Lipschitz_continuous_map (submetric m1 S) m2 f" + and "T \ mtopology_of m1 closure_of S" + obtains g + where "Lipschitz_continuous_map (submetric m1 T) m2 g" "\x. x \ S \ g x = f x" + by (metis Lipschitz_continuous_map_extends_to_closure_of Lipschitz_continuous_map_from_submetric_mono assms) + +text \This proof uses the same trick to extend the function's domain to its closure\ +lemma uniformly_continuous_map_on_intermediate_closure_aux: + assumes ucf: "uniformly_continuous_map (submetric m1 S) m2 f" + and "S \ T" and Tsub: "T \ (mtopology_of m1) closure_of S" + and cmf: "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f" + and "S \ mspace m1" + shows "uniformly_continuous_map (submetric m1 T) m2 f" +proof - + interpret L: Metric_space12 "mspace m1" "mdist m1" "mspace m2" "mdist m2" + by (simp add: Metric_space12_mspace_mdist) + interpret S: Metric_space "S \ mspace m1" "mdist m1" + by (simp add: L.M1.subspace) + have "T \ mspace m1" + using Tsub by (auto simp: mtopology_of_def closure_of_def) + show ?thesis + unfolding uniformly_continuous_map_def + proof (intro conjI strip) + show "f ` mspace (submetric m1 T) \ mspace m2" + by (metis cmf Metric_space.metric_continuous_map Metric_space_mspace_mdist mtopology_of_def mtopology_of_submetric) + fix \::real + assume "\ > 0" + then obtain \ where "\>0" and \: "\(x,y) \ S\S. mdist m1 x y < \ \ mdist m2 (f x) (f y) \ \/2" + using ucf \S \ mspace m1\ unfolding uniformly_continuous_map_def mspace_submetric + apply (simp add: Ball_def del: divide_const_simps) + by (metis IntD2 half_gt_zero inf.orderE less_eq_real_def) + define X where "X \ prod_topology (subtopology L.M1.mtopology T) (subtopology L.M1.mtopology T)" + text \A clever construction involving the union of two closed sets\ + have eq: "{z \ A. case z of (x,y) \ p x y < d \ q x y \ e} + = {z \ A. ((\(x,y). p x y - d)z) \ {0..}} \ {z \ A. ((\(x,y). e - q x y)z) \ {0..}}" + for p q and d e::real and A::"('a*'a)set" + by auto + have clo: "closedin X {z \ topspace X. case z of (x, y) \ mdist m1 x y < \ \ mdist m2 (f x) (f y) \ \/2}" + unfolding eq + proof (intro closedin_Un closedin_continuous_map_preimage) + have *: "continuous_map X L.M1.mtopology fst" "continuous_map X L.M1.mtopology snd" + by (metis X_def continuous_map_subtopology_fst subtopology_Times continuous_map_subtopology_snd)+ + show "continuous_map X euclidean (\x. case x of (x, y) \ mdist m1 x y - \)" + unfolding case_prod_unfold + by (intro continuous_intros; simp add: mtopology_of_def *) + have *: "continuous_map X L.M2.mtopology (f \ fst)" "continuous_map X L.M2.mtopology (f \ snd)" + using cmf by (auto simp: mtopology_of_def X_def intro: continuous_map_compose continuous_map_fst continuous_map_snd) + then show "continuous_map X euclidean (\x. case x of (x, y) \ \ / 2 - mdist m2 (f x) (f y))" + unfolding case_prod_unfold + by (intro continuous_intros; simp add: mtopology_of_def o_def) + qed auto + have "mdist m2 (f x) (f y) \ \/2" if "x \ T" "y \ T" "mdist m1 x y < \" for x y + using all_in_closure_of [OF \ clo] \S \ T\ Tsub + by (fastforce simp: X_def subset_iff closure_of_Times closure_of_subtopology inf.absorb2 + mtopology_of_def that) + then show "\\>0. \x\mspace (submetric m1 T). \y\mspace (submetric m1 T). mdist (submetric m1 T) y x < \ \ mdist m2 (f y) (f x) < \" + using \0 < \\ \0 < \\ by fastforce + qed +qed + +lemma uniformly_continuous_map_on_intermediate_closure: + assumes "uniformly_continuous_map (submetric m1 S) m2 f" + and "S \ T" and"T \ (mtopology_of m1) closure_of S" + and "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f" + shows "uniformly_continuous_map (submetric m1 T) m2 f" + by (metis assms closure_of_subset_topspace subset_trans topspace_mtopology_of + uniformly_continuous_map_on_intermediate_closure_aux) + +lemma uniformly_continuous_map_extends_to_closure_of: + assumes m2: "mcomplete_of m2" + and f: "uniformly_continuous_map (submetric m1 S) m2 f" + obtains g + where "uniformly_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g" + "\x. x \ S \ g x = f x" +proof - + obtain g + where g: "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S)) + (mtopology_of m2) g" "(\x \ S. g x = f x)" + by (metis Cauchy_continuous_map_extends_to_continuous_closure_of uniformly_imp_Cauchy_continuous_map f m2) + have "uniformly_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g" + proof (rule uniformly_continuous_map_on_intermediate_closure) + show "uniformly_continuous_map (submetric m1 (mspace m1 \ S)) m2 g" + by (smt (verit, best) IntD2 uniformly_continuous_map_eq f g(2) inf_commute mspace_submetric submetric_restrict) + show "mspace m1 \ S \ mtopology_of m1 closure_of S" + using closure_of_subset_Int by force + show "mtopology_of m1 closure_of S \ mtopology_of m1 closure_of (mspace m1 \ S)" + by (metis closure_of_restrict subset_refl topspace_mtopology_of) + show "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S)) (mtopology_of m2) g" + by (simp add: g) + qed + with g that show thesis + by metis +qed + + +lemma uniformly_continuous_map_extends_to_intermediate_closure_of: + assumes "mcomplete_of m2" + and "uniformly_continuous_map (submetric m1 S) m2 f" + and "T \ mtopology_of m1 closure_of S" + obtains g + where "uniformly_continuous_map (submetric m1 T) m2 g" "\x. x \ S \ g x = f x" + by (metis uniformly_continuous_map_extends_to_closure_of uniformly_continuous_map_from_submetric_mono assms) + + +lemma Cauchy_continuous_map_on_intermediate_closure_aux: + assumes ucf: "Cauchy_continuous_map (submetric m1 S) m2 f" + and "S \ T" and Tsub: "T \ (mtopology_of m1) closure_of S" + and cmf: "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f" + and "S \ mspace m1" + shows "Cauchy_continuous_map (submetric m1 T) m2 f" +proof - + interpret L: Metric_space12 "mspace m1" "mdist m1" "mspace m2" "mdist m2" + by (simp add: Metric_space12_mspace_mdist) + interpret S: Metric_space "S \ mspace m1" "mdist m1" + by (simp add: L.M1.subspace) + interpret T: Metric_space T "mdist m1" + by (metis L.M1.subspace Tsub closure_of_subset_topspace dual_order.trans topspace_mtopology_of) + have "T \ mspace m1" + using Tsub by (auto simp: mtopology_of_def closure_of_def) + then show ?thesis + proof (clarsimp simp: Cauchy_continuous_map_def Int_absorb2) + fix \ + assume \: "T.MCauchy \" + have "\y\S. mdist m1 (\ n) y < inverse (Suc n) \ mdist m2 (f (\ n)) (f y) < inverse (Suc n)" for n + proof - + have "\ n \ T" + using \ by (force simp: T.MCauchy_def) + moreover have "continuous_map (mtopology_of (submetric m1 T)) L.M2.mtopology f" + by (metis cmf mtopology_of_def mtopology_of_submetric) + ultimately obtain \ where "\>0" and \: "\x \ T. mdist m1 (\ n) x < \ \ mdist m2 (f(\ n)) (f x) < inverse (Suc n)" + using \T \ mspace m1\ + apply (simp add: mtopology_of_def Metric_space.metric_continuous_map L.M1.subspace Int_absorb2) + by (metis inverse_Suc of_nat_Suc) + have "\y \ S. mdist m1 (\ n) y < min \ (inverse (Suc n))" + using \\ n \ T\ Tsub \\>0\ + unfolding mtopology_of_def L.M1.metric_closure_of subset_iff mem_Collect_eq L.M1.in_mball + by (smt (verit, del_insts) inverse_Suc ) + with \ \S \ T\ show ?thesis + by auto + qed + then obtain \ where \S: "\n. \ n \ S" and \1: "\n. mdist m1 (\ n) (\ n) < inverse (Suc n)" + and \2: "\n. mdist m2 (f (\ n)) (f (\ n)) < inverse (Suc n)" + by metis + have "S.MCauchy \" + unfolding S.MCauchy_def + proof (intro conjI strip) + show "range \ \ S \ mspace m1" + using \S \ mspace m1\ by (auto simp: \S) + fix \ :: real + assume "\>0" + then obtain M where M: "\n n'. M \ n \ M \ n' \ mdist m1 (\ n) (\ n') < \/2" + using \ unfolding T.MCauchy_def by (meson half_gt_zero) + have "\\<^sub>F n in sequentially. inverse (Suc n) < \/4" + using Archimedean_eventually_inverse \0 < \\ divide_pos_pos zero_less_numeral by blast + then obtain N where N: "\n. N \ n \ inverse (Suc n) < \/4" + by (meson eventually_sequentially) + have "mdist m1 (\ n) (\ n') < \" if "n \ max M N" "n' \ max M N" for n n' + proof - + have "mdist m1 (\ n) (\ n') \ mdist m1 (\ n) (\ n) + mdist m1 (\ n) (\ n')" + by (meson T.MCauchy_def T.triangle \S \ \S \ T\ rangeI subset_iff) + also have "\ \ mdist m1 (\ n) (\ n) + mdist m1 (\ n) (\ n') + mdist m1 (\ n') (\ n')" + by (smt (verit, best) T.MCauchy_def T.triangle \S \ \S \ T\ in_mono rangeI) + also have "\ < \/4 + \/2 + \/4" + using \1[of n] \1[of n'] N[of n] N[of n'] that M[of n n'] by (simp add: T.commute) + also have "... \ \" + by simp + finally show ?thesis . + qed + then show "\N. \n n'. N \ n \ N \ n' \ mdist m1 (\ n) (\ n') < \" + by blast + qed + then have f\: "L.M2.MCauchy (f \ \)" + using ucf by (simp add: Cauchy_continuous_map_def) + show "L.M2.MCauchy (f \ \)" + unfolding L.M2.MCauchy_def + proof (intro conjI strip) + show "range (f \ \) \ mspace m2" + using \T \ mspace m1\ \ cmf + apply (auto simp: ) + by (metis Metric_space.metric_continuous_map Metric_space_mspace_mdist T.MCauchy_def image_eqI inf.absorb1 mspace_submetric mtopology_of_def mtopology_of_submetric range_subsetD subset_iff) + fix \ :: real + assume "\>0" + then obtain M where M: "\n n'. M \ n \ M \ n' \ mdist m2 ((f \ \) n) ((f \ \) n') < \/2" + using f\ unfolding L.M2.MCauchy_def by (meson half_gt_zero) + have "\\<^sub>F n in sequentially. inverse (Suc n) < \/4" + using Archimedean_eventually_inverse \0 < \\ divide_pos_pos zero_less_numeral by blast + then obtain N where N: "\n. N \ n \ inverse (Suc n) < \/4" + by (meson eventually_sequentially) + have "mdist m2 ((f \ \) n) ((f \ \) n') < \" if "n \ max M N" "n' \ max M N" for n n' + proof - + have "mdist m2 ((f \ \) n) ((f \ \) n') \ mdist m2 ((f \ \) n) ((f \ \) n) + mdist m2 ((f \ \) n) ((f \ \) n')" + by (meson L.M2.MCauchy_def \range (f \ \) \ mspace m2\ f\ mdist_triangle rangeI subset_eq) + also have "\ \ mdist m2 ((f \ \) n) ((f \ \) n) + mdist m2 ((f \ \) n) ((f \ \) n') + mdist m2 ((f \ \) n') ((f \ \) n')" + by (smt (verit) L.M2.MCauchy_def L.M2.triangle \range (f \ \) \ mspace m2\ f\ range_subsetD) + also have "\ < \/4 + \/2 + \/4" + using \2[of n] \2[of n'] N[of n] N[of n'] that M[of n n'] by (simp add: L.M2.commute) + also have "... \ \" + by simp + finally show ?thesis . + qed + then show "\N. \n n'. N \ n \ N \ n' \ mdist m2 ((f \ \) n) ((f \ \) n') < \" + by blast + qed + qed +qed + +lemma Cauchy_continuous_map_on_intermediate_closure: + assumes "Cauchy_continuous_map (submetric m1 S) m2 f" + and "S \ T" and "T \ (mtopology_of m1) closure_of S" + and "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f" + shows "Cauchy_continuous_map (submetric m1 T) m2 f" + by (metis Cauchy_continuous_map_on_intermediate_closure_aux assms closure_of_subset_topspace order.trans topspace_mtopology_of) + + +lemma Cauchy_continuous_map_extends_to_closure_of: + assumes m2: "mcomplete_of m2" + and f: "Cauchy_continuous_map (submetric m1 S) m2 f" + obtains g + where "Cauchy_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g" + "\x. x \ S \ g x = f x" +proof - + obtain g + where g: "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S)) + (mtopology_of m2) g" "(\x \ S. g x = f x)" + by (metis Cauchy_continuous_map_extends_to_continuous_closure_of f m2) + have "Cauchy_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g" + proof (rule Cauchy_continuous_map_on_intermediate_closure) + show "Cauchy_continuous_map (submetric m1 (mspace m1 \ S)) m2 g" + by (smt (verit, best) IntD2 Cauchy_continuous_map_eq f g(2) inf_commute mspace_submetric submetric_restrict) + show "mspace m1 \ S \ mtopology_of m1 closure_of S" + using closure_of_subset_Int by force + show "mtopology_of m1 closure_of S \ mtopology_of m1 closure_of (mspace m1 \ S)" + by (metis closure_of_restrict subset_refl topspace_mtopology_of) + show "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S)) (mtopology_of m2) g" + by (simp add: g) + qed + with g that show thesis + by metis +qed + + +lemma Cauchy_continuous_map_extends_to_intermediate_closure_of: + assumes "mcomplete_of m2" + and "Cauchy_continuous_map (submetric m1 S) m2 f" + and "T \ mtopology_of m1 closure_of S" + obtains g + where "Cauchy_continuous_map (submetric m1 T) m2 g" "\x. x \ S \ g x = f x" + by (metis Cauchy_continuous_map_extends_to_closure_of Cauchy_continuous_map_from_submetric_mono assms) + + +subsection\Metric space of bounded functions\ + +context Metric_space +begin + +definition fspace :: "'b set \ ('b \ 'a) set" where + "fspace \ \S. {f. f`S \ M \ f \ extensional S \ mbounded (f`S)}" + +definition fdist :: "['b set, 'b \ 'a, 'b \ 'a] \ real" where + "fdist \ \S f g. if f \ fspace S \ g \ fspace S \ S \ {} + then Sup ((\x. d (f x) (g x)) ` S) else 0" + +lemma fspace_empty [simp]: "fspace {} = {\x. undefined}" + by (auto simp: fspace_def) + +lemma fdist_empty [simp]: "fdist {} = (\x y. 0)" + by (auto simp: fdist_def) + +lemma fspace_in_M: "\f \ fspace S; x \ S\ \ f x \ M" + by (auto simp: fspace_def) + +lemma bdd_above_dist: + assumes f: "f \ fspace S" and g: "g \ fspace S" and "S \ {}" + shows "bdd_above ((\u. d (f u) (g u)) ` S)" +proof - + obtain a where "a \ S" + using \S \ {}\ by blast + obtain B x C y where "B>0" and B: "f`S \ mcball x B" + and "C>0" and C: "g`S \ mcball y C" + using f g mbounded_pos by (auto simp: fspace_def) + have "d (f u) (g u) \ B + d x y + C" if "u\S" for u + proof - + have "f u \ M" + by (meson B image_eqI mbounded_mcball mbounded_subset_mspace subsetD that) + have "g u \ M" + by (meson C image_eqI mbounded_mcball mbounded_subset_mspace subsetD that) + have "x \ M" "y \ M" + using B C that by auto + have "d (f u) (g u) \ d (f u) x + d x (g u)" + by (simp add: \f u \ M\ \g u \ M\ \x \ M\ triangle) + also have "\ \ d (f u) x + d x y + d y (g u)" + by (simp add: \f u \ M\ \g u \ M\ \x \ M\ \y \ M\ triangle) + also have "\ \ B + d x y + C" + using B C commute that by fastforce + finally show ?thesis . + qed + then show ?thesis + by (meson bdd_above.I2) +qed + + +lemma Metric_space_funspace: "Metric_space (fspace S) (fdist S)" +proof + show *: "0 \ fdist S f g" for f g + by (auto simp: fdist_def intro: cSUP_upper2 [OF bdd_above_dist]) + show "fdist S f g = fdist S g f" for f g + by (auto simp: fdist_def commute) + show "(fdist S f g = 0) = (f = g)" + if fg: "f \ fspace S" "g \ fspace S" for f g + proof + assume 0: "fdist S f g = 0" + show "f = g" + proof (cases "S={}") + case True + with 0 that show ?thesis + by (simp add: fdist_def fspace_def) + next + case False + with 0 fg have Sup0: "(SUP x\S. d (f x) (g x)) = 0" + by (simp add: fdist_def) + have "d (f x) (g x) = 0" if "x\S" for x + by (smt (verit) False Sup0 \x\S\ bdd_above_dist [OF fg] less_cSUP_iff nonneg) + with fg show "f=g" + by (simp add: fspace_def extensionalityI image_subset_iff) + qed + next + show "f = g \ fdist S f g = 0" + using fspace_in_M [OF \g \ fspace S\] by (auto simp: fdist_def) + qed + show "fdist S f h \ fdist S f g + fdist S g h" + if fgh: "f \ fspace S" "g \ fspace S" "h \ fspace S" for f g h + proof (clarsimp simp add: fdist_def that) + assume "S \ {}" + have dfh: "d (f x) (h x) \ d (f x) (g x) + d (g x) (h x)" if "x\S" for x + by (meson fgh fspace_in_M that triangle) + have bdd_fgh: "bdd_above ((\x. d (f x) (g x)) ` S)" "bdd_above ((\x. d (g x) (h x)) ` S)" + by (simp_all add: \S \ {}\ bdd_above_dist that) + then obtain B C where B: "\x. x\S \ d (f x) (g x) \ B" and C: "\x. x\S \ d (g x) (h x) \ C" + by (auto simp: bdd_above_def) + then have "\x. x\S \ d (f x) (g x) + d (g x) (h x) \ B+C" + by force + then have bdd: "bdd_above ((\x. d (f x) (g x) + d (g x) (h x)) ` S)" + by (auto simp: bdd_above_def) + then have "(SUP x\S. d (f x) (h x)) \ (SUP x\S. d (f x) (g x) + d (g x) (h x))" + by (metis (mono_tags, lifting) cSUP_mono \S \ {}\ dfh) + also have "\ \ (SUP x\S. d (f x) (g x)) + (SUP x\S. d (g x) (h x))" + by (simp add: \S \ {}\ bdd cSUP_le_iff bdd_fgh add_mono cSup_upper) + finally show "(SUP x\S. d (f x) (h x)) \ (SUP x\S. d (f x) (g x)) + (SUP x\S. d (g x) (h x))" . + qed +qed + end + + +definition funspace where + "funspace S m \ metric (Metric_space.fspace (mspace m) (mdist m) S, + Metric_space.fdist (mspace m) (mdist m) S)" + +lemma mspace_funspace [simp]: + "mspace (funspace S m) = Metric_space.fspace (mspace m) (mdist m) S" + by (simp add: Metric_space.Metric_space_funspace Metric_space.mspace_metric funspace_def) + +lemma mdist_funspace [simp]: + "mdist (funspace S m) = Metric_space.fdist (mspace m) (mdist m) S" + by (simp add: Metric_space.Metric_space_funspace Metric_space.mdist_metric funspace_def) + +lemma funspace_imp_welldefined: + "\f \ mspace (funspace S m); x \ S\ \ f x \ mspace m" + by (simp add: Metric_space.fspace_def subset_iff) + +lemma funspace_imp_extensional: + "f \ mspace (funspace S m) \ f \ extensional S" + by (simp add: Metric_space.fspace_def) + +lemma funspace_imp_bounded_image: + "f \ mspace (funspace S m) \ Metric_space.mbounded (mspace m) (mdist m) (f ` S)" + by (simp add: Metric_space.fspace_def) + +lemma funspace_imp_bounded: + "f \ mspace (funspace S m) \ S = {} \ (\c B. \x \ S. mdist m c (f x) \ B)" + by (auto simp: Metric_space.fspace_def Metric_space.mbounded) + + +lemma (in Metric_space) funspace_imp_bounded2: + assumes "f \ fspace S" "g \ fspace S" + obtains B where "\x. x \ S \ d (f x) (g x) \ B" +proof - + have "mbounded (f ` S \ g ` S)" + using mbounded_Un assms by (force simp: fspace_def) + then show thesis + by (metis UnCI imageI mbounded_alt that) +qed + +lemma funspace_imp_bounded2: + assumes "f \ mspace (funspace S m)" "g \ mspace (funspace S m)" + obtains B where "\x. x \ S \ mdist m (f x) (g x) \ B" + by (metis Metric_space_mspace_mdist assms mspace_funspace Metric_space.funspace_imp_bounded2) + +lemma (in Metric_space) funspace_mdist_le: + assumes fg: "f \ fspace S" "g \ fspace S" and "S \ {}" + shows "fdist S f g \ a \ (\x \ S. d (f x) (g x) \ a)" + using assms bdd_above_dist [OF fg] by (simp add: fdist_def cSUP_le_iff) + +lemma funspace_mdist_le: + assumes "f \ mspace (funspace S m)" "g \ mspace (funspace S m)" and "S \ {}" + shows "mdist (funspace S m) f g \ a \ (\x \ S. mdist m (f x) (g x) \ a)" + using assms by (simp add: Metric_space.funspace_mdist_le) + + +lemma (in Metric_space) mcomplete_funspace: + assumes "mcomplete" + shows "mcomplete_of (funspace S Self)" +proof - + interpret F: Metric_space "fspace S" "fdist S" + by (simp add: Metric_space_funspace) + show ?thesis + proof (cases "S={}") + case True + then show ?thesis + by (simp add: mcomplete_of_def mcomplete_trivial_singleton) + next + case False + show ?thesis + proof (clarsimp simp: mcomplete_of_def Metric_space.mcomplete_def) + fix \ + assume \: "F.MCauchy \" + then have \M: "\n x. x \ S \ \ n x \ M" + by (auto simp: F.MCauchy_def intro: fspace_in_M) + have fdist_less: "\N. \n n'. N \ n \ N \ n' \ fdist S (\ n) (\ n') < \" if "\>0" for \ + using \ that by (auto simp: F.MCauchy_def) + have \ext: "\n. \ n \ extensional S" + using \ unfolding F.MCauchy_def by (auto simp: fspace_def) + have \bd: "\n. mbounded (\ n ` S)" + using \ unfolding F.MCauchy_def by (simp add: fspace_def image_subset_iff) + have \in[simp]: "\ n \ fspace S" for n + using F.MCauchy_def \ by blast + have bd2: "\n n'. \B. \x \ S. d (\ n x) (\ n' x) \ B" + using \ unfolding F.MCauchy_def by (metis range_subsetD funspace_imp_bounded2) + have sup: "\n n' x0. x0 \ S \ d (\ n x0) (\ n' x0) \ Sup ((\x. d (\ n x) (\ n' x)) ` S)" + proof (rule cSup_upper) + show "bdd_above ((\x. d (\ n x) (\ n' x)) ` S)" if "x0 \ S" for n n' x0 + using that bd2 by (meson bdd_above.I2) + qed auto + have pcy: "MCauchy (\n. \ n x)" if "x \ S" for x + unfolding MCauchy_def + proof (intro conjI strip) + show "range (\n. \ n x) \ M" + using \M that by blast + fix \ :: real + assume "\ > 0" + then obtain N where N: "\n n'. N \ n \ N \ n' \ fdist S (\ n) (\ n') < \" + using \ by (force simp: F.MCauchy_def) + { fix n n' + assume n: "N \ n" "N \ n'" + have "d (\ n x) (\ n' x) \ (SUP x\S. d (\ n x) (\ n' x))" + using that sup by presburger + then have "d (\ n x) (\ n' x) \ fdist S (\ n) (\ n')" + by (simp add: fdist_def \S \ {}\) + with N n have "d (\ n x) (\ n' x) < \" + by fastforce + } then show "\N. \n n'. N \ n \ N \ n' \ d (\ n x) (\ n' x) < \" + by blast + qed + have "\l. limitin mtopology (\n. \ n x) l sequentially" if "x \ S" for x + using assms mcomplete_def pcy \x \ S\ by presburger + then obtain g0 where g0: "\x. x \ S \ limitin mtopology (\n. \ n x) (g0 x) sequentially" + by metis + define g where "g \ restrict g0 S" + have gext: "g \ extensional S" + and glim: "\x. x \ S \ limitin mtopology (\n. \ n x) (g x) sequentially" + by (auto simp: g_def g0) + have gwd: "g x \ M" if "x \ S" for x + using glim limitin_metric that by blast + have unif: "\N. \x n. x \ S \ N \ n \ d (\ n x) (g x) < \" if "\>0" for \ + proof - + obtain N where N: "\n n'. N \ n \ N \ n' \ Sup ((\x. d (\ n x) (\ n' x)) ` S) < \/2" + using \S\{}\ \\>0\ fdist_less [of "\/2"] + by (metis (mono_tags) \in fdist_def half_gt_zero) + show ?thesis + proof (intro exI strip) + fix x n + assume "x \ S" and "N \ n" + obtain N' where N': "\n. N' \ n \ \ n x \ M \ d (\ n x) (g x) < \/2" + by (metis \0 < \\ \x \ S\ glim half_gt_zero limit_metric_sequentially) + have "d (\ n x) (g x) \ d (\ n x) (\ (max N N') x) + d (\ (max N N') x) (g x)" + using \x \ S\ \M gwd triangle by presburger + also have "\ < \/2 + \/2" + by (smt (verit) N N' \N \ n\ \x \ S\ max.cobounded1 max.cobounded2 sup) + finally show "d (\ n x) (g x) < \" by simp + qed + qed + have "limitin F.mtopology \ g sequentially" + unfolding F.limit_metric_sequentially + proof (intro conjI strip) + obtain N where N: "\n n'. N \ n \ N \ n' \ Sup ((\x. d (\ n x) (\ n' x)) ` S) < 1" + using fdist_less [of 1] \S\{}\ by (auto simp: fdist_def) + have "\x. x \ \ N ` S \ x \ M" + using \M by blast + obtain a B where "a \ M" and B: "\x. x \ (\ N) ` S \ d a x \ B" + by (metis False \M \bd ex_in_conv imageI mbounded_alt_pos) + have "d a (g x) \ B+1" if "x\S" for x + proof - + have "d a (g x) \ d a (\ N x) + d (\ N x) (g x)" + by (simp add: \a \ M\ \M gwd that triangle) + also have "\ \ B+1" + proof - + have "d a (\ N x) \ B" + by (simp add: B that) + moreover + have False if 1: "d (\ N x) (g x) > 1" + proof - + obtain r where "1 < r" and r: "r < d (\ N x) (g x)" + using 1 dense by blast + then obtain N' where N': "\n. N' \ n \ \ n x \ M \ d (\ n x) (g x) < r-1" + using glim [OF \x\S\] by (fastforce simp: limit_metric_sequentially) + have "d (\ N x) (g x) \ d (\ N x) (\ (max N N') x) + d (\ (max N N') x) (g x)" + by (metis \x \ S\ \M commute gwd triangle') + also have "\ < 1 + (r-1)" + by (smt (verit) N N' \x \ S\ max.cobounded1 max.cobounded2 max.idem sup) + finally have "d (\ N x) (g x) < r" + by simp + with r show False + by linarith + qed + ultimately show ?thesis + by force + qed + finally show ?thesis . + qed + with gwd \a \ M\ have "mbounded (g ` S)" + unfolding mbounded by blast + with gwd gext show "g \ fspace S" + by (auto simp: fspace_def) + fix \::real + assume "\>0" + then obtain N where "\x n. x \ S \ N \ n \ d (\ n x) (g x) < \/2" + by (meson unif half_gt_zero) + then have "fdist S (\ n) g \ \/2" if "N \ n" for n + using \g \ fspace S\ False that + by (force simp: funspace_mdist_le simp del: divide_const_simps) + then show "\N. \n\N. \ n \ fspace S \ fdist S (\ n) g < \" + by (metis \0 < \\ \in add_strict_increasing field_sum_of_halves half_gt_zero) + qed + then show "\x. limitin F.mtopology \ x sequentially" + by blast + qed + qed +qed + + + +subsection\Metric space of continuous bounded functions\ + +definition cfunspace where + "cfunspace X m \ submetric (funspace (topspace X) m) {f. continuous_map X (mtopology_of m) f}" + +lemma mspace_cfunspace [simp]: + "mspace (cfunspace X m) = + {f. f ` topspace X \ mspace m \ f \ extensional (topspace X) \ + Metric_space.mbounded (mspace m) (mdist m) (f ` (topspace X)) \ + continuous_map X (mtopology_of m) f}" + by (auto simp: cfunspace_def Metric_space.fspace_def) + +lemma mdist_cfunspace_eq_mdist_funspace: + "mdist (cfunspace X m) = mdist (funspace (topspace X) m)" + by (auto simp: cfunspace_def) + +lemma cfunspace_subset_funspace: + "mspace (cfunspace X m) \ mspace (funspace (topspace X) m)" + by (simp add: cfunspace_def) + +lemma cfunspace_mdist_le: + "\f \ mspace (cfunspace X m); g \ mspace (cfunspace X m); topspace X \ {}\ + \ mdist (cfunspace X m) f g \ a \ (\x \ topspace X. mdist m (f x) (g x) \ a)" + by (simp add: cfunspace_def Metric_space.funspace_mdist_le) + +lemma cfunspace_imp_bounded2: + assumes "f \ mspace (cfunspace X m)" "g \ mspace (cfunspace X m)" + obtains B where "\x. x \ topspace X \ mdist m (f x) (g x) \ B" + by (metis assms all_not_in_conv cfunspace_mdist_le nle_le) + +lemma cfunspace_mdist_lt: + "\compactin X (topspace X); f \ mspace (cfunspace X m); + g \ mspace (cfunspace X m); mdist (cfunspace X m) f g < a; + x \ topspace X\ + \ mdist m (f x) (g x) < a" + by (metis (full_types) cfunspace_mdist_le empty_iff less_eq_real_def less_le_not_le) + +lemma mdist_cfunspace_le: + assumes "0 \ B" and B: "\x. x \ topspace X \ mdist m (f x) (g x) \ B" + shows "mdist (cfunspace X m) f g \ B" +proof (cases "topspace X = {}") + case True + then show ?thesis + by (simp add: Metric_space.fdist_empty \B\0\ cfunspace_def) +next + case False + have bdd: "bdd_above ((\u. mdist m (f u) (g u)) ` topspace X)" + by (meson B bdd_above.I2) + with assms bdd show ?thesis + by (simp add: mdist_cfunspace_eq_mdist_funspace Metric_space.fdist_def cSUP_le_iff) +qed + + +lemma mdist_cfunspace_imp_mdist_le: + "\f \ mspace (cfunspace X m); g \ mspace (cfunspace X m); + mdist (cfunspace X m) f g \ a; x \ topspace X\ \ mdist m (f x) (g x) \ a" + using cfunspace_mdist_le by blast + +lemma compactin_mspace_cfunspace: + "compactin X (topspace X) + \ mspace (cfunspace X m) = + {f. (\x \ topspace X. f x \ mspace m) \ + f \ extensional (topspace X) \ + continuous_map X (mtopology_of m) f}" + by (auto simp: Metric_space.compactin_imp_mbounded image_compactin mtopology_of_def) + +lemma (in Metric_space) mcomplete_cfunspace: + assumes "mcomplete" + shows "mcomplete_of (cfunspace X Self)" +proof - + interpret F: Metric_space "fspace (topspace X)" "fdist (topspace X)" + by (simp add: Metric_space_funspace) + interpret S: Submetric "fspace (topspace X)" "fdist (topspace X)" "mspace (cfunspace X Self)" + proof + show "mspace (cfunspace X Self) \ fspace (topspace X)" + by (metis cfunspace_subset_funspace mdist_Self mspace_Self mspace_funspace) + qed + show ?thesis + proof (cases "topspace X = {}") + case True + then show ?thesis + by (simp add: mcomplete_of_def mcomplete_trivial_singleton mdist_cfunspace_eq_mdist_funspace cong: conj_cong) + next + case False + have *: "continuous_map X mtopology g" + if "range \ \ mspace (cfunspace X Self)" + and g: "limitin F.mtopology \ g sequentially" for \ g + unfolding continuous_map_to_metric + proof (intro strip) + have \: "\n. continuous_map X mtopology (\ n)" + using that by (auto simp: mtopology_of_def) + fix x and \::real + assume "x \ topspace X" and "0 < \" + then obtain N where N: "\n. N \ n \ \ n \ fspace (topspace X) \ fdist (topspace X) (\ n) g < \/3" + unfolding mtopology_of_def F.limitin_metric + by (metis F.limit_metric_sequentially divide_pos_pos g zero_less_numeral) + then obtain U where "openin X U" "x \ U" + and U: "\y. y \ U \ \ N y \ mball (\ N x) (\/3)" + by (metis Metric_space.continuous_map_to_metric Metric_space_axioms \0 < \\ \x \ topspace X\ \ divide_pos_pos zero_less_numeral) + moreover + have "g y \ mball (g x) \" if "y\U" for y + proof - + have "U \ topspace X" + using \openin X U\ by (simp add: openin_subset) + have gx: "g x \ M" + by (meson F.limitin_mspace \x \ topspace X\ fspace_in_M g) + have "y \ topspace X" + using \U \ topspace X\ that by auto + have gy: "g y \ M" + by (meson F.limitin_mspace[OF g] \U \ topspace X\ fspace_in_M subsetD that) + have "d (g x) (g y) < \" + proof - + have *: "d (\ N x0) (g x0) \ \/3" if "x0 \ topspace X" for x0 + proof - + have "bdd_above ((\x. d (\ N x) (g x)) ` topspace X)" + by (metis F.limit_metric_sequentially False N bdd_above_dist g order_refl) + then have "d (\ N x0) (g x0) \ Sup ((\x. d (\ N x) (g x)) ` topspace X)" + by (simp add: cSup_upper that) + also have "\ \ \/3" + by (smt (verit) F.limit_metric_sequentially False N Sup.SUP_cong fdist_def g order_refl) + finally show ?thesis . + qed + have "d (g x) (g y) \ d (g x) (\ N x) + d (\ N x) (g y)" + using U gx gy that triangle by force + also have "\ < \/3 + \/3 + \/3" + by (smt (verit) "*" U gy \x \ topspace X\ \y \ topspace X\ commute in_mball that triangle) + finally show ?thesis by simp + qed + with gx gy show ?thesis by simp + qed + ultimately show "\U. openin X U \ x \ U \ (\y\U. g y \ mball (g x) \)" + by blast + qed + + have "S.sub.mcomplete" + proof (rule S.sequentially_closedin_mcomplete_imp_mcomplete) + show "F.mcomplete" + by (metis assms mcomplete_funspace mcomplete_of_def mdist_Self mdist_funspace mspace_Self mspace_funspace) + fix \ g + assume g: "range \ \ mspace (cfunspace X Self) \ limitin F.mtopology \ g sequentially" + show "g \ mspace (cfunspace X Self)" + proof (simp add: mtopology_of_def, intro conjI) + show "g ` topspace X \ M" "g \ extensional (topspace X)" "mbounded (g ` topspace X)" + using g F.limitin_mspace by (force simp: fspace_def)+ + show "continuous_map X mtopology g" + using * g by blast + qed + qed + then show ?thesis + by (simp add: mcomplete_of_def mdist_cfunspace_eq_mdist_funspace) + qed +qed + + +subsection\Existence of completion for any metric space M as a subspace of M=>R\ + +lemma (in Metric_space) metric_completion_explicit: + obtains f :: "['a,'a] \ real" and S where + "S \ mspace(funspace M euclidean_metric)" + "mcomplete_of (submetric (funspace M euclidean_metric) S)" + "f ` M \ S" + "mtopology_of(funspace M euclidean_metric) closure_of f ` M = S" + "\x y. \x \ M; y \ M\ + \ mdist (funspace M euclidean_metric) (f x) (f y) = d x y" +proof - + define m':: "('a\real) metric" where "m' \ funspace M euclidean_metric" + show thesis + proof (cases "M = {}") + case True + then show ?thesis + using that by (simp add: mcomplete_of_def mcomplete_trivial) + next + case False + then obtain a where "a \ M" + by auto + define f where "f \ \x. (\u \ M. d x u - d a u)" + define S where "S \ mtopology_of(funspace M euclidean_metric) closure_of (f ` M)" + interpret S: Submetric "Met_TC.fspace M" "Met_TC.fdist M" "S \ Met_TC.fspace M" + by (simp add: Met_TC.Metric_space_funspace Submetric.intro Submetric_axioms_def) + + have fim: "f ` M \ mspace m'" + proof (clarsimp simp: m'_def Met_TC.fspace_def) + fix b + assume "b \ M" + then have "\c. \c \ M\ \ \d b c - d a c\ \ d a b" + by (smt (verit, best) \a \ M\ commute triangle'') + then have "(\x. d b x - d a x) ` M \ cball 0 (d a b)" + by force + then show "f b \ extensional M \ bounded (f b ` M)" + by (metis bounded_cball bounded_subset f_def image_restrict_eq restrict_extensional_sub set_eq_subset) + qed + show thesis + proof + show "S \ mspace (funspace M euclidean_metric)" + by (simp add: S_def in_closure_of subset_iff) + have "closedin S.mtopology (S \ Met_TC.fspace M)" + by (simp add: S_def closedin_Int funspace_def) + moreover have "S.mcomplete" + using Metric_space.mcomplete_funspace Met_TC.Metric_space_axioms by (fastforce simp: mcomplete_of_def) + ultimately show "mcomplete_of (submetric (funspace M euclidean_metric) S)" + by (simp add: S.closedin_eq_mcomplete mcomplete_of_def) + show "f ` M \ S" + using S_def fim in_closure_of m'_def by fastforce + show "mtopology_of (funspace M euclidean_metric) closure_of f ` M = S" + by (auto simp: f_def S_def mtopology_of_def) + show "mdist (funspace M euclidean_metric) (f x) (f y) = d x y" + if "x \ M" "y \ M" for x y + proof - + have "\c\M. dist (f x c) (f y c) \ r \ d x y \ r" for r + using that by (auto simp: f_def dist_real_def) + moreover have "dist (f x z) (f y z) \ r" if "d x y \ r" and "z \ M" for r z + using that \x \ M\ \y \ M\ + apply (simp add: f_def Met_TC.fdist_def dist_real_def) + by (smt (verit, best) commute triangle') + ultimately have "(SUP c \ M. dist (f x c) (f y c)) = d x y" + by (intro cSup_unique) auto + with that fim show ?thesis + using that fim by (simp add: Met_TC.fdist_def False m'_def image_subset_iff) + qed + qed + qed +qed + + +lemma (in Metric_space) metric_completion: + obtains f :: "['a,'a] \ real" and m' where + "mcomplete_of m'" + "f ` M \ mspace m' " + "mtopology_of m' closure_of f ` M = mspace m'" + "\x y. \x \ M; y \ M\ \ mdist m' (f x) (f y) = d x y" +proof - + obtain f :: "['a,'a] \ real" and S where + Ssub: "S \ mspace(funspace M euclidean_metric)" + and mcom: "mcomplete_of (submetric (funspace M euclidean_metric) S)" + and fim: "f ` M \ S" + and eqS: "mtopology_of(funspace M euclidean_metric) closure_of f ` M = S" + and eqd: "\x y. \x \ M; y \ M\ \ mdist (funspace M euclidean_metric) (f x) (f y) = d x y" + using metric_completion_explicit by metis + define m' where "m' \ submetric (funspace M euclidean_metric) S" + show thesis + proof + show "mcomplete_of m'" + by (simp add: mcom m'_def) + show "f ` M \ mspace m'" + using Ssub fim m'_def by auto + show "mtopology_of m' closure_of f ` M = mspace m'" + using eqS fim Ssub + by (force simp: m'_def mtopology_of_submetric closure_of_subtopology Int_absorb1) + show "mdist m' (f x) (f y) = d x y" if "x \ M" and "y \ M" for x y + using that eqd m'_def by force + qed +qed + +lemma metrizable_space_completion: + assumes "metrizable_space X" + obtains f :: "['a,'a] \ real" and Y where + "completely_metrizable_space Y" "embedding_map X Y f" + "Y closure_of (f ` (topspace X)) = topspace Y" +proof - + obtain M d where "Metric_space M d" and Xeq: "X = Metric_space.mtopology M d" + using assms metrizable_space_def by blast + then interpret Metric_space M d by simp + obtain f :: "['a,'a] \ real" and m' where + "mcomplete_of m'" + and fim: "f ` M \ mspace m' " + and m': "mtopology_of m' closure_of f ` M = mspace m'" + and eqd: "\x y. \x \ M; y \ M\ \ mdist m' (f x) (f y) = d x y" + by (meson metric_completion) + show thesis + proof + show "completely_metrizable_space (mtopology_of m')" + using \mcomplete_of m'\ + unfolding completely_metrizable_space_def mcomplete_of_def mtopology_of_def + by (metis Metric_space_mspace_mdist) + show "embedding_map X (mtopology_of m') f" + using Metric_space12.isometry_imp_embedding_map + by (metis Metric_space12_def Metric_space_axioms Metric_space_mspace_mdist Xeq eqd fim mtopology_of_def) + show "(mtopology_of m') closure_of f ` topspace X = topspace (mtopology_of m')" + by (simp add: Xeq m') + qed +qed + + + +subsection\Contractions\ + +lemma (in Metric_space) contraction_imp_unique_fixpoint: + assumes "f x = x" "f y = y" + and "f ` M \ M" + and "k < 1" + and "\x y. \x \ M; y \ M\ \ d (f x) (f y) \ k * d x y" + and "x \ M" "y \ M" + shows "x = y" + by (smt (verit, ccfv_SIG) Metric_space.mdist_pos_less Metric_space_axioms assms mult_le_cancel_right1) + +text \Banach Fixed-Point Theorem (aka, Contraction Mapping Principle)\ +lemma (in Metric_space) Banach_fixedpoint_thm: + assumes mcomplete and "M \ {}" and fim: "f ` M \ M" + and "k < 1" + and con: "\x y. \x \ M; y \ M\ \ d (f x) (f y) \ k * d x y" + obtains x where "x \ M" "f x = x" +proof - + obtain a where "a \ M" + using \M \ {}\ by blast + show thesis + proof (cases "\x \ M. f x = f a") + case True + then show ?thesis + by (metis \a \ M\ fim image_subset_iff that) + next + case False + then obtain b where "b \ M" and b: "f b \ f a" + by blast + have "k>0" + using Lipschitz_coefficient_pos [where f=f] + by (metis False \a \ M\ con fim mdist_Self mspace_Self) + define \ where "\ \ \n. (f^^n) a" + have f_iter: "\ n \ M" for n + unfolding \_def by (induction n) (use \a \ M\ fim in auto) + show ?thesis + proof (cases "f a = a") + case True + then show ?thesis + using \a \ M\ that by blast + next + case False + have "MCauchy \" + proof - + show ?thesis + unfolding MCauchy_def + proof (intro conjI strip) + show "range \ \ M" + using f_iter by blast + fix \::real + assume "\>0" + with \k < 1\ \f a \ a\ \a \ M\ fim have gt0: "((1 - k) * \) / d a (f a) > 0" + by (fastforce simp: divide_simps) + obtain N where "k^N < ((1-k) * \) / d a (f a)" + using real_arch_pow_inv [OF gt0 \k < 1\] by blast + then have N: "\n. n \ N \ k^n < ((1-k) * \) / d a (f a)" + by (smt (verit) \0 < k\ assms(4) power_decreasing) + have "\n n'. n N \ n \ N \ n' \ d (\ n) (\ n') < \" + proof (intro exI strip) + fix n n' + assume "n n" "N \ n'" + have "d (\ n) (\ n') \ (\i=n.. i) (\ (Suc i)))" + proof - + have "n < m \ d (\ n) (\ m) \ (\i=n.. i) (\ (Suc i)))" for m + proof (induction m) + case 0 + then show ?case + by simp + next + case (Suc m) + then consider "n n) (\ (Suc m)) \ d (\ n) (\ m) + d (\ m) (\ (Suc m))" + by (simp add: f_iter triangle) + also have "\ \ (\i=n.. i) (\ (Suc i))) + d (\ m) (\ (Suc m))" + using Suc 1 by linarith + also have "\ = (\i = n.. i) (\ (Suc i)))" + using "1" by force + finally show ?thesis . + qed auto + qed + with \n < n'\ show ?thesis by blast + qed + also have "\ \ (\i=n.. {n.. i) (\ (Suc i)) \ d a (f a) * k ^ i" + proof (induction i) + case 0 + then show ?case + by (auto simp: \_def) + next + case (Suc i) + have "d (\ (Suc i)) (\ (Suc (Suc i))) \ k * d (\ i) (\ (Suc i))" + using con \_def f_iter fim by fastforce + also have "\ \ d a (f a) * k ^ Suc i" + using Suc \0 < k\ by auto + finally show ?case . + qed + qed + also have "\ = d a (f a) * (\i=n.. = d a (f a) * (\i=0..n < n'\ by simp + also have "\ = d a (f a) * k^n * (\i = d a (f a) * (k ^ n - k ^ n') / (1 - k)" + using \k < 1\ \n < n'\ apply (simp add: sum_gp_strict) + by (simp add: algebra_simps flip: power_add) + also have "\ < \" + using N \k < 1\ \0 < \\ \0 < k\ \N \ n\ + apply (simp add: field_simps) + by (smt (verit) nonneg pos_less_divide_eq zero_less_divide_iff zero_less_power) + finally show "d (\ n) (\ n') < \" . + qed + then show "\N. \n n'. N \ n \ N \ n' \ d (\ n) (\ n') < \" + by (metis \0 < \\ commute f_iter linorder_not_le local.mdist_zero nat_less_le) + qed + qed + then obtain l where l: "limitin mtopology \ l sequentially" + using \mcomplete\ mcomplete_def by blast + show ?thesis + proof + show "l \ M" + using l limitin_mspace by blast + show "f l = l" + proof (rule limitin_metric_unique) + have "limitin mtopology (f \ \) (f l) sequentially" + proof (rule continuous_map_limit) + have "Lipschitz_continuous_map Self Self f" + using con by (auto simp: Lipschitz_continuous_map_def fim) + then show "continuous_map mtopology mtopology f" + using Lipschitz_continuous_imp_continuous_map Self_def by force + qed (use l in auto) + moreover have "(f \ \) = (\i. \(i+1))" + by (auto simp: \_def) + ultimately show "limitin mtopology (\n. (f^^n)a) (f l) sequentially" + using limitin_sequentially_offset_rev [of mtopology \ 1] + by (simp add: \_def) + qed (use l in \auto simp: \_def\) + qed + qed + qed +qed + + +subsection\ The Baire Category Theorem\ + +text \Possibly relevant to the theorem "Baire" in Elementary Normed Spaces\ + +lemma (in Metric_space) metric_Baire_category: + assumes "mcomplete" "countable \" + and "\T. T \ \ \ openin mtopology T \ mtopology closure_of T = M" +shows "mtopology closure_of \\ = M" +proof (cases "\={}") + case False + then obtain U :: "nat \ 'a set" where U: "range U = \" + by (metis \countable \\ uncountable_def) + with assms have u_open: "\n. openin mtopology (U n)" and u_dense: "\n. mtopology closure_of (U n) = M" + by auto + have "\(range U) \ W \ {}" if W: "openin mtopology W" "W \ {}" for W + proof - + have "W \ M" + using openin_mtopology W by blast + have "\r' x'. 0 < r' \ r' < r/2 \ x' \ M \ mcball x' r' \ mball x r \ U n" + if "r>0" "x \ M" for x r n + proof - + obtain z where z: "z \ U n \ mball x r" + using u_dense [of n] \r>0\ \x \ M\ + by (metis dense_intersects_open centre_in_mball_iff empty_iff openin_mball topspace_mtopology equals0I) + then have "z \ M" by auto + have "openin mtopology (U n \ mball x r)" + by (simp add: openin_Int u_open) + with \z \ M\ z obtain e where "e>0" and e: "mcball z e \ U n \ mball x r" + by (meson openin_mtopology_mcball) + define r' where "r' \ min e (r/4)" + show ?thesis + proof (intro exI conjI) + show "0 < r'" "r' < r / 2" "z \ M" + using \e>0\ \r>0\ \z \ M\ by (auto simp: r'_def) + show "mcball z r' \ mball x r \ U n" + using Metric_space.mcball_subset_concentric e r'_def by auto + qed + qed + then obtain nextr nextx + where nextr: "\r x n. \r>0; x\M\ \ 0 < nextr r x n \ nextr r x n < r/2" + and nextx: "\r x n. \r>0; x\M\ \ nextx r x n \ M" + and nextsub: "\r x n. \r>0; x\M\ \ mcball (nextx r x n) (nextr r x n) \ mball x r \ U n" + by metis + obtain x0 where x0: "x0 \ U 0 \ W" + by (metis W dense_intersects_open topspace_mtopology all_not_in_conv u_dense) + then have "x0 \ M" + using \W \ M\ by fastforce + obtain r0 where "0 < r0" "r0 < 1" and sub: "mcball x0 r0 \ U 0 \ W" + proof - + have "openin mtopology (U 0 \ W)" + using W u_open by blast + then obtain r where "r>0" and r: "mball x0 r \ U 0" "mball x0 r \ W" + by (meson Int_subset_iff openin_mtopology x0) + define r0 where "r0 \ (min r 1) / 2" + show thesis + proof + show "0 < r0" "r0 < 1" + using \r>0\ by (auto simp: r0_def) + show "mcball x0 r0 \ U 0 \ W" + using r \0 < r0\ r0_def by auto + qed + qed + define b where "b \ rec_nat (x0,r0) (\n (x,r). (nextx r x n, nextr r x n))" + have b0[simp]: "b 0 = (x0,r0)" + by (simp add: b_def) + have bSuc[simp]: "b (Suc n) = (let (x,r) = b n in (nextx r x n, nextr r x n))" for n + by (simp add: b_def) + define xf where "xf \ fst \ b" + define rf where "rf \ snd \ b" + have rfxf: "0 < rf n \ xf n \ M" for n + proof (induction n) + case 0 + with \0 < r0\ \x0 \ M\ show ?case + by (auto simp: rf_def xf_def) + next + case (Suc n) + then show ?case + by (auto simp: rf_def xf_def case_prod_unfold nextr nextx Let_def) + qed + have mcball_sub: "mcball (xf (Suc n)) (rf (Suc n)) \ mball (xf n) (rf n) \ U n" for n + using rfxf nextsub by (auto simp: xf_def rf_def case_prod_unfold Let_def) + have half: "rf (Suc n) < rf n / 2" for n + using rfxf nextr by (auto simp: xf_def rf_def case_prod_unfold Let_def) + then have "decseq rf" + using rfxf by (smt (verit, ccfv_threshold) decseq_SucI field_sum_of_halves) + have nested: "mball (xf n) (rf n) \ mball (xf m) (rf m)" if "m \ n" for m n + using that + proof (induction n) + case (Suc n) + then show ?case + by (metis mcball_sub order.trans inf.boundedE le_Suc_eq mball_subset_mcball order.refl) + qed auto + have "MCauchy xf" + unfolding MCauchy_def + proof (intro conjI strip) + show "range xf \ M" + using rfxf by blast + fix \ :: real + assume "\>0" + then obtain N where N: "inverse (2^N) < \" + using real_arch_pow_inv by (force simp flip: power_inverse) + have "d (xf n) (xf n') < \" if "n \ n'" "N \ n" "N \ n'" for n n' + proof - + have *: "rf n < inverse (2 ^ n)" for n + proof (induction n) + case 0 + then show ?case + by (simp add: \r0 < 1\ rf_def) + next + case (Suc n) + with half show ?case + by simp (smt (verit)) + qed + have "rf n \ rf N" + using \decseq rf\ \N \ n\ by (simp add: decseqD) + moreover + have "xf n' \ mball (xf n) (rf n)" + using nested rfxf \n \ n'\ by blast + ultimately have "d (xf n) (xf n') < rf N" + by auto + also have "\ < \" + using "*" N order.strict_trans by blast + finally show ?thesis . + qed + then show "\N. \n n'. N \ n \ N \ n' \ d (xf n) (xf n') < \" + by (metis commute linorder_le_cases) + qed + then obtain l where l: "limitin mtopology xf l sequentially" + using \mcomplete\ mcomplete_alt by blast + have l_in: "l \ mcball (xf n) (rf n)" for n + proof - + have "\\<^sub>F m in sequentially. xf m \ mcball (xf n) (rf n)" + unfolding eventually_sequentially + by (meson nested rfxf centre_in_mball_iff mball_subset_mcball subset_iff) + with l limitin_closedin show ?thesis + by (metis closedin_mcball trivial_limit_sequentially) + qed + then have "\n. l \ U n" + using mcball_sub by blast + moreover have "l \ W" + using l_in[of 0] sub by (auto simp: xf_def rf_def) + ultimately show ?thesis by auto + qed + with U show ?thesis + by (metis dense_intersects_open topspace_mtopology) +qed auto + + + +lemma (in Metric_space) metric_Baire_category_alt: + assumes "mcomplete" "countable \" + and empty: "\T. T \ \ \ closedin mtopology T \ mtopology interior_of T = {}" + shows "mtopology interior_of \\ = {}" +proof - + have *: "mtopology closure_of \((-)M ` \) = M" + proof (intro metric_Baire_category conjI \mcomplete\) + show "countable ((-) M ` \)" + using \countable \\ by blast + fix T + assume "T \ (-) M ` \" + then obtain U where U: "U \ \" "T = M-U" "U \ M" + using empty metric_closedin_iff_sequentially_closed by force + with empty show "openin mtopology T" by blast + show "mtopology closure_of T = M" + using U by (simp add: closure_of_interior_of double_diff empty) + qed + with closure_of_eq show ?thesis + by (fastforce simp: interior_of_closure_of split: if_split_asm) +qed + + +text \Since all locally compact Hausdorff spaces are regular, the disjunction in the HOL Light version is redundant.\ +lemma Baire_category_aux: + assumes "locally_compact_space X" "regular_space X" + and "countable \" + and empty: "\G. G \ \ \ closedin X G \ X interior_of G = {}" + shows "X interior_of \\ = {}" +proof (cases "\ = {}") + case True + then show ?thesis + by simp +next + case False + then obtain T :: "nat \ 'a set" where T: "\ = range T" + by (metis \countable \\ uncountable_def) + with empty have Tempty: "\n. X interior_of (T n) = {}" + by auto + show ?thesis + proof (clarsimp simp: T interior_of_def) + fix z U + assume "z \ U" and opeA: "openin X U" and Asub: "U \ \ (range T)" + with openin_subset have "z \ topspace X" + by blast + have "neighbourhood_base_of (\C. compactin X C \ closedin X C) X" + using assms locally_compact_regular_space_neighbourhood_base by auto + then obtain V K where "openin X V" "compactin X K" "closedin X K" "z \ V" "V \ K" "K \ U" + by (metis (no_types, lifting) \z \ U\ neighbourhood_base_of opeA) + have nb_closedin: "neighbourhood_base_of (closedin X) X" + using \regular_space X\ neighbourhood_base_of_closedin by auto + have "\\. \n. (\ n \ K \ closedin X (\ n) \ X interior_of \ n \ {} \ disjnt (\ n) (T n)) \ + \ (Suc n) \ \ n" + proof (rule dependent_nat_choice) + show "\x\K. closedin X x \ X interior_of x \ {} \ disjnt x (T 0)" + proof - + have False if "V \ T 0" + using Tempty \openin X V\ \z \ V\ interior_of_maximal that by fastforce + then obtain x where "openin X (V - T 0) \ x \ V - T 0" + using T \openin X V\ empty by blast + with nb_closedin + obtain N C where "openin X N" "closedin X C" "x \ N" "N \ C" "C \ V - T 0" + unfolding neighbourhood_base_of by metis + show ?thesis + proof (intro exI conjI) + show "C \ K" + using \C \ V - T 0\ \V \ K\ by auto + show "X interior_of C \ {}" + by (metis \N \ C\ \openin X N\ \x \ N\ empty_iff interior_of_eq_empty) + show "disjnt C (T 0)" + using \C \ V - T 0\ disjnt_iff by fastforce + qed (use \closedin X C\ in auto) + qed + show "\L. (L \ K \ closedin X L \ X interior_of L \ {} \ disjnt L (T (Suc n))) \ L \ C" + if \
: "C \ K \ closedin X C \ X interior_of C \ {} \ disjnt C (T n)" + for C n + proof - + have False if "X interior_of C \ T (Suc n)" + by (metis Tempty interior_of_eq_empty \
openin_interior_of that) + then obtain x where "openin X (X interior_of C - T (Suc n)) \ x \ X interior_of C - T (Suc n)" + using T empty by fastforce + with nb_closedin + obtain N D where "openin X N" "closedin X D" "x \ N" "N \ D" and D: "D \ X interior_of C - T(Suc n)" + unfolding neighbourhood_base_of by metis + show ?thesis + proof (intro conjI exI) + show "D \ K" + using D interior_of_subset \
by fastforce + show "X interior_of D \ {}" + by (metis \N \ D\ \openin X N\ \x \ N\ empty_iff interior_of_eq_empty) + show "disjnt D (T (Suc n))" + using D disjnt_iff by fastforce + show "D \ C" + using interior_of_subset [of X C] D by blast + qed (use \closedin X D\ in auto) + qed + qed + then obtain \ where \: "\n. \ n \ K \ closedin X (\ n) \ X interior_of \ n \ {} \ disjnt (\ n) (T n)" + and "\n. \ (Suc n) \ \ n" by metis + then have "decseq \" + by (simp add: decseq_SucI) + moreover have "\n. \ n \ {}" + by (metis \ bot.extremum_uniqueI interior_of_subset) + ultimately have "\(range \) \ {}" + by (metis \ compact_space_imp_nest \compactin X K\ compactin_subspace closedin_subset_topspace) + moreover have "U \ {y. \x. y \ T x}" + using Asub by auto + with T have "{a. \n. a \ \ n} \ {}" + by (smt (verit) Asub \ Collect_empty_eq UN_iff \K \ U\ disjnt_iff subset_iff) + ultimately show False + by blast + qed +qed + + +lemma Baire_category_alt: + assumes "completely_metrizable_space X \ locally_compact_space X \ regular_space X" + and "countable \" + and "\T. T \ \ \ closedin X T \ X interior_of T = {}" + shows "X interior_of \\ = {}" + using Baire_category_aux [of X \] Metric_space.metric_Baire_category_alt + by (metis assms completely_metrizable_space_def) + + +lemma Baire_category: + assumes "completely_metrizable_space X \ locally_compact_space X \ regular_space X" + and "countable \" + and top: "\T. T \ \ \ openin X T \ X closure_of T = topspace X" + shows "X closure_of \\ = topspace X" +proof (cases "\={}") + case False + have *: "X interior_of \((-)(topspace X) ` \) = {}" + proof (intro Baire_category_alt conjI assms) + show "countable ((-) (topspace X) ` \)" + using assms by blast + fix T + assume "T \ (-) (topspace X) ` \" + then obtain U where U: "U \ \" "T = (topspace X) - U" "U \ (topspace X)" + by (meson top image_iff openin_subset) + then show "closedin X T" + by (simp add: closedin_diff top) + show "X interior_of T = {}" + using U top by (simp add: interior_of_closure_of double_diff) + qed + then show ?thesis + by (simp add: closure_of_eq_topspace interior_of_complement) +qed auto + + +subsection\Sierpinski-Hausdorff type results about countable closed unions\ + +lemma locally_connected_not_countable_closed_union: + assumes "topspace X \ {}" and csX: "connected_space X" + and lcX: "locally_connected_space X" + and X: "completely_metrizable_space X \ locally_compact_space X \ Hausdorff_space X" + and "countable \" and pwU: "pairwise disjnt \" + and clo: "\C. C \ \ \ closedin X C \ C \ {}" + and UU_eq: "\\ = topspace X" + shows "\ = {topspace X}" +proof - + define \ where "\ \ (frontier_of) X ` \" + define B where "B \ \\" + then have Bsub: "B \ topspace X" + by (simp add: Sup_le_iff \_def closedin_frontier_of closedin_subset) + have allsub: "A \ topspace X" if "A \ \" for A + by (meson clo closedin_def that) + show ?thesis + proof (rule ccontr) + assume "\ \ {topspace X}" + with assms have "\A\\. \ (closedin X A \ openin X A)" + by (metis Union_empty connected_space_clopen_in singletonI subsetI subset_singleton_iff) + then have "B \ {}" + by (auto simp: B_def \_def frontier_of_eq_empty allsub) + moreover + have "subtopology X B interior_of B = B" + by (simp add: Bsub interior_of_openin openin_subtopology_refl) + ultimately have int_B_nonempty: "subtopology X B interior_of B \ {}" + by auto + have "subtopology X B interior_of \\ = {}" + proof (intro Baire_category_alt conjI) + have "\\ \ B \ \((interior_of) X ` \)" + using clo closure_of_closedin by (fastforce simp: B_def \_def frontier_of_def) + moreover have "B \ \((interior_of) X ` \) \ \\" + using allsub clo frontier_of_subset_eq interior_of_subset by (fastforce simp: B_def \_def ) + moreover have "disjnt B (\((interior_of) X ` \))" + using pwU + apply (clarsimp simp: B_def \_def frontier_of_def pairwise_def disjnt_iff) + by (metis clo closure_of_eq interior_of_subset subsetD) + ultimately have "B = topspace X - \ ((interior_of) X ` \)" + by (auto simp: UU_eq disjnt_iff) + then have "closedin X B" + by fastforce + with X show "completely_metrizable_space (subtopology X B) \ locally_compact_space (subtopology X B) \ regular_space (subtopology X B)" + by (metis completely_metrizable_space_closedin locally_compact_Hausdorff_or_regular + locally_compact_space_closed_subset regular_space_subtopology) + show "countable \" + by (simp add: \_def \countable \\) + fix V + assume "V \ \" + then obtain S where S: "S \ \" "V = X frontier_of S" + by (auto simp: \_def) + show "closedin (subtopology X B) V" + by (metis B_def Sup_upper \_def \V \ \\ closedin_frontier_of closedin_subset_topspace image_iff) + have "subtopology X B interior_of (X frontier_of S) = {}" + proof (clarsimp simp: interior_of_def openin_subtopology_alt) + fix a U + assume "a \ B" "a \ U" and opeU: "openin X U" and BUsub: "B \ U \ X frontier_of S" + then have "a \ S" + by (meson IntI \S \ \\ clo frontier_of_subset_closedin subsetD) + then obtain W C where "openin X W" "connectedin X C" "a \ W" "W \ C" "C \ U" + by (metis \a \ U\ lcX locally_connected_space opeU) + have "W \ X frontier_of S \ {}" + using \B \ U \ X frontier_of S\ \a \ B\ \a \ U\ \a \ W\ by auto + with frontier_of_openin_straddle_Int + obtain "W \ S \ {}" "W - S \ {}" "W \ topspace X" + using \openin X W\ by (metis openin_subset) + then obtain b where "b \ topspace X" "b \ W-S" + by blast + with UU_eq obtain T where "T \ \" "T \ S" "W \ T \ {}" + by auto + then have "disjnt S T" + by (metis \S \ \\ pairwise_def pwU) + then have "C - T \ {}" + by (meson Diff_eq_empty_iff \W \ C\ \a \ S\ \a \ W\ disjnt_iff subsetD) + then have "C \ X frontier_of T \ {}" + using \W \ T \ {}\ \W \ C\ \connectedin X C\ connectedin_Int_frontier_of by blast + moreover have "C \ X frontier_of T = {}" + proof - + have "X frontier_of S \ S" "X frontier_of T \ T" + using frontier_of_subset_closedin \S \ \\ \T \ \\ clo by blast+ + moreover have "X frontier_of T \ B = B" + using B_def \_def \T \ \\ by blast + ultimately show ?thesis + using BUsub \C \ U\ \disjnt S T\ unfolding disjnt_def by blast + qed + ultimately show False + by simp + qed + with S show "subtopology X B interior_of V = {}" + by meson + qed + then show False + using B_def int_B_nonempty by blast + qed +qed + +lemma real_Sierpinski_lemma: + fixes a b::real + assumes "a \ b" + and "countable \" and pwU: "pairwise disjnt \" + and clo: "\C. C \ \ \ closed C \ C \ {}" + and "\\ = {a..b}" + shows "\ = {{a..b}}" +proof - + have "locally_connected_space (top_of_set {a..b})" + by (simp add: locally_connected_real_interval) + moreover + have "completely_metrizable_space (top_of_set {a..b})" + by (metis box_real(2) completely_metrizable_space_cbox) + ultimately + show ?thesis + using locally_connected_not_countable_closed_union [of "subtopology euclidean {a..b}"] assms + apply (simp add: closedin_subtopology) + by (metis Union_upper inf.orderE) +qed + + +subsection\The Tychonoff embedding\ + +lemma completely_regular_space_cube_embedding_explicit: + assumes "completely_regular_space X" "Hausdorff_space X" + shows "embedding_map X + (product_topology (\f. top_of_set {0..1::real}) + (mspace (submetric (cfunspace X euclidean_metric) + {f. f ` topspace X \ {0..1}})) ) + (\x. \f \ mspace (submetric (cfunspace X euclidean_metric) {f. f ` topspace X \ {0..1}}). + f x)" +proof - + define K where "K \ mspace(submetric (cfunspace X euclidean_metric) {f. f ` topspace X \ {0..1::real}})" + define e where "e \ \x. \f\K. f x" + have "e x \ e y" if xy: "x \ y" "x \ topspace X" "y \ topspace X" for x y + proof - + have "closedin X {x}" + by (simp add: \Hausdorff_space X\ closedin_Hausdorff_singleton \x \ topspace X\) + then obtain f where contf: "continuous_map X euclideanreal f" + and f01: "\x. x \ topspace X \ f x \ {0..1}" and fxy: "f y = 0" "f x = 1" + using \completely_regular_space X\ xy unfolding completely_regular_space_def + by (smt (verit, ccfv_threshold) Diff_iff continuous_map_in_subtopology image_subset_iff singleton_iff) + then have "bounded (f ` topspace X)" + by (meson bounded_closed_interval bounded_subset image_subset_iff) + with contf f01 have "restrict f (topspace X) \ K" + by (auto simp: K_def) + with fxy xy show ?thesis + unfolding e_def by (metis restrict_apply' zero_neq_one) + qed + then have "inj_on e (topspace X)" + by (meson inj_onI) + then obtain e' where e': "\x. x \ topspace X \ e' (e x) = x" + by (metis inv_into_f_f) + have "continuous_map (subtopology (product_topology (\f. top_of_set {0..1}) K) (e ` topspace X)) X e'" + proof (clarsimp simp add: continuous_map_atin limitin_atin openin_subtopology_alt e') + fix x U + assume "e x \ K \\<^sub>E {0..1}" and "x \ topspace X" and "openin X U" and "x \ U" + then obtain g where contg: "continuous_map X (top_of_set {0..1}) g" and "g x = 0" + and gim: "g ` (topspace X - U) \ {1::real}" + using \completely_regular_space X\ unfolding completely_regular_space_def + by (metis Diff_iff openin_closedin_eq) + then have "bounded (g ` topspace X)" + by (meson bounded_closed_interval bounded_subset continuous_map_in_subtopology) + moreover have "g ` topspace X \ {0..1}" + using contg by (simp add: continuous_map_in_subtopology) + ultimately have g_in_K: "restrict g (topspace X) \ K" + using contg by (simp add: K_def continuous_map_in_subtopology) + have "openin (top_of_set {0..1}) {0..<1::real}" + using open_real_greaterThanLessThan[of "-1" 1] by (force simp: openin_open) + moreover have "e x \ (\\<^sub>E f\K. if f = restrict g (topspace X) then {0..<1} else {0..1})" + using \e x \ K \\<^sub>E {0..1}\ by (simp add: e_def \g x = 0\ \x \ topspace X\ PiE_iff) + moreover have "e y = e x" + if "y \ U" and ey: "e y \ (\\<^sub>E f\K. if f = restrict g (topspace X) then {0..<1} else {0..1})" + and y: "y \ topspace X" for y + proof - + have "e y (restrict g (topspace X)) \ {0..<1}" + using ey by (smt (verit, ccfv_SIG) PiE_mem g_in_K) + with gim g_in_K y \y \ U\ show ?thesis + by (fastforce simp: e_def) + qed + ultimately + show "\W. openin (product_topology (\f. top_of_set {0..1}) K) W \ e x \ W \ e' ` (e ` topspace X \ W - {e x}) \ U" + apply (rule_tac x="PiE K (\f. if f = restrict g (topspace X) then {0..<1} else {0..1})" in exI) + by (auto simp: openin_PiE_gen e') + qed + with e' have "embedding_map X (product_topology (\f. top_of_set {0..1}) K) e" + unfolding embedding_map_def homeomorphic_map_maps homeomorphic_maps_def + by (fastforce simp: e_def K_def continuous_map_in_subtopology continuous_map_componentwise) + then show ?thesis + by (simp add: K_def e_def) +qed + + +lemma completely_regular_space_cube_embedding: + fixes X :: "'a topology" + assumes "completely_regular_space X" "Hausdorff_space X" + obtains K:: "('a\real)set" and e + where "embedding_map X (product_topology (\f. top_of_set {0..1::real}) K) e" + using completely_regular_space_cube_embedding_explicit [OF assms] by metis + + +subsection \Urysohn and Tietze analogs for completely regular spaces\ + +text\"Urysohn and Tietze analogs for completely regular spaces if (()) set is +assumed compact instead of closed. Note that Hausdorffness is *not* required: inside () proof +we factor through the Kolmogorov quotient." -- John Harrison\ + +lemma Urysohn_completely_regular_closed_compact: + fixes a b::real + assumes "a \ b" "completely_regular_space X" "closedin X S" "compactin X T" "disjnt S T" + obtains f where "continuous_map X (subtopology euclidean {a..b}) f" "f ` T \ {a}" "f ` S \ {b}" +proof - + obtain f where contf: "continuous_map X (subtopology euclideanreal {0..1}) f" + and f0: "f ` T \ {0}" and f1: "f ` S \ {1}" + proof (cases "T={}") + case True + show thesis + proof + show "continuous_map X (top_of_set {0..1}) (\x. 1::real)" "(\x. 1::real) ` T \ {0}" "(\x. 1::real) ` S \ {1}" + using True by auto + qed + next + case False + have "\t. t \ T \ \f. continuous_map X (subtopology euclideanreal ({0..1})) f \ f t = 0 \ f ` S \ {1}" + using assms unfolding completely_regular_space_def + by (meson DiffI compactin_subset_topspace disjnt_iff subset_eq) + then obtain g where contg: "\t. t \ T \ continuous_map X (subtopology euclideanreal {0..1}) (g t)" + and g0: "\t. t \ T \ g t t = 0" + and g1: "\t. t \ T \ g t ` S \ {1}" + by metis + then have g01: "\t. t \ T \ g t ` topspace X \ {0..1}" + by (meson continuous_map_in_subtopology) + define G where "G \ \t. {x \ topspace X. g t x \ {..<1/2}}" + have "Ball (G`T) (openin X)" + using contg unfolding G_def continuous_map_in_subtopology + by (smt (verit, best) Collect_cong openin_continuous_map_preimage image_iff open_lessThan open_openin) + moreover have "T \ \(G`T)" + using \compactin X T\ g0 compactin_subset_topspace by (force simp: G_def) + ultimately have "\\. finite \ \ \ \ G`T \ T \ \ \" + using \compactin X T\ unfolding compactin_def by blast + then obtain K where K: "finite K" "K \ T" "T \ \(G`K)" + by (metis finite_subset_image) + with False have "K \ {}" + by fastforce + define f where "f \ \x. 2 * max 0 (Inf ((\t. g t x) ` K) - 1/2)" + have [simp]: "max 0 (x - 1/2) = 0 \ x \ 1/2" for x::real + by force + have [simp]: "2 * max 0 (x - 1/2) = 1 \ x = 1" for x::real + by (simp add: max_def_raw) + show thesis + proof + have "g t s = 1" if "s \ S" "t \ K" for s t + using \K \ T\ g1 that by auto + then show "f ` S \ {1}" + using \K \ {}\ by (simp add: f_def image_subset_iff) + have "(INF t\K. g t x) \ 1/2" if "x \ T" for x + proof - + obtain k where "k \ K" "g k x < 1/2" + using K \x \ T\ by (auto simp: G_def) + then show ?thesis + by (meson \finite K\ cInf_le_finite dual_order.trans finite_imageI imageI less_le_not_le) + qed + then show "f ` T \ {0}" + by (force simp: f_def) + have "\t. t \ K \ continuous_map X euclideanreal (g t)" + using \K \ T\ contg continuous_map_in_subtopology by blast + moreover have "2 * max 0 ((INF t\K. g t x) - 1/2) \ 1" if "x \ topspace X" for x + proof - + obtain k where "k \ K" "g k x \ 1" + using K \x \ topspace X\ \K \ {}\ g01 by (fastforce simp: G_def) + then have "(INF t\K. g t x) \ 1" + by (meson \finite K\ cInf_le_finite dual_order.trans finite_imageI imageI) + then show ?thesis + by (simp add: max_def_raw) + qed + ultimately show "continuous_map X (top_of_set {0..1}) f" + by (force simp: f_def continuous_map_in_subtopology intro!: \finite K\ continuous_intros) + qed + qed + define g where "g \ \x. a + (b - a) * f x" + show thesis + proof + have "\x\topspace X. a + (b - a) * f x \ b" + using contf \a \ b\ apply (simp add: continuous_map_in_subtopology image_subset_iff) + by (smt (verit, best) mult_right_le_one_le) + then show "continuous_map X (top_of_set {a..b}) g" + using contf \a \ b\ unfolding g_def continuous_map_in_subtopology image_subset_iff + by (intro conjI continuous_intros; simp) + show "g ` T \ {a}" "g ` S \ {b}" + using f0 f1 by (auto simp: g_def) + qed +qed + + +lemma Urysohn_completely_regular_compact_closed: + fixes a b::real + assumes "a \ b" "completely_regular_space X" "compactin X S" "closedin X T" "disjnt S T" + obtains f where "continuous_map X (subtopology euclidean {a..b}) f" "f ` T \ {a}" "f ` S \ {b}" +proof - + obtain f where contf: "continuous_map X (subtopology euclidean {-b..-a}) f" and fim: "f ` T \ {-a}" "f ` S \ {-b}" + by (meson Urysohn_completely_regular_closed_compact assms disjnt_sym neg_le_iff_le) + show thesis + proof + show "continuous_map X (top_of_set {a..b}) (uminus \ f)" + using contf by (auto simp: continuous_map_in_subtopology o_def) + show "(uminus o f) ` T \ {a}" "(uminus o f) ` S \ {b}" + using fim by fastforce+ + qed +qed + +lemma Urysohn_completely_regular_compact_closed_alt: + fixes a b::real + assumes "completely_regular_space X" "compactin X S" "closedin X T" "disjnt S T" + obtains f where "continuous_map X euclideanreal f" "f ` T \ {a}" "f ` S \ {b}" +proof (cases a b rule: le_cases) + case le + then show ?thesis + by (meson Urysohn_completely_regular_compact_closed assms continuous_map_into_fulltopology that) +next + case ge + then show ?thesis + using Urysohn_completely_regular_closed_compact assms + by (metis Urysohn_completely_regular_closed_compact assms continuous_map_into_fulltopology disjnt_sym that) +qed + + +lemma Tietze_extension_comp_reg_aux: + fixes T :: "real set" + assumes "completely_regular_space X" "Hausdorff_space X" "compactin X S" + and T: "is_interval T" "T\{}" + and contf: "continuous_map (subtopology X S) euclidean f" and fim: "f`S \ T" + obtains g where "continuous_map X euclidean g" "g ` topspace X \ T" "\x. x \ S \ g x = f x" +proof - + obtain K:: "('a\real)set" and e + where e0: "embedding_map X (product_topology (\f. top_of_set {0..1::real}) K) e" + using assms completely_regular_space_cube_embedding by blast + define cube where "cube \ product_topology (\f. top_of_set {0..1::real}) K" + have e: "embedding_map X cube e" + using e0 by (simp add: cube_def) + obtain e' where e': "homeomorphic_maps X (subtopology cube (e ` topspace X)) e e'" + using e by (force simp: cube_def embedding_map_def homeomorphic_map_maps) + then have conte: "continuous_map X (subtopology cube (e ` topspace X)) e" + and conte': "continuous_map (subtopology cube (e ` topspace X)) X e'" + and e'e: "\x\topspace X. e' (e x) = x" + by (auto simp: homeomorphic_maps_def) + have "Hausdorff_space cube" + unfolding cube_def + using Hausdorff_space_euclidean Hausdorff_space_product_topology Hausdorff_space_subtopology by blast + have "normal_space cube" + proof (rule compact_Hausdorff_or_regular_imp_normal_space) + show "compact_space cube" + unfolding cube_def + using compact_space_product_topology compact_space_subtopology compactin_euclidean_iff by blast + qed (use \Hausdorff_space cube\ in auto) + moreover + have comp: "compactin cube (e ` S)" + by (meson \compactin X S\ conte continuous_map_in_subtopology image_compactin) + then have "closedin cube (e ` S)" + by (intro compactin_imp_closedin \Hausdorff_space cube\) + moreover + have "continuous_map (subtopology cube (e ` S)) euclideanreal (f \ e')" + proof (intro continuous_map_compose) + show "continuous_map (subtopology cube (e ` S)) (subtopology X S) e'" + unfolding continuous_map_in_subtopology + proof + show "continuous_map (subtopology cube (e ` S)) X e'" + by (meson \compactin X S\ compactin_subset_topspace conte' continuous_map_from_subtopology_mono image_mono) + show "e' ` topspace (subtopology cube (e ` S)) \ S" + using \compactin X S\ compactin_subset_topspace e'e by fastforce + qed + qed (simp add: contf) + moreover + have "(f \ e') ` e ` S \ T" + using \compactin X S\ compactin_subset_topspace e'e fim by fastforce + ultimately + obtain g where contg: "continuous_map cube euclidean g" and gsub: "g ` topspace cube \ T" + and gf: "\x. x \ e`S \ g x = (f \ e') x" + using Tietze_extension_realinterval T by metis + show thesis + proof + show "continuous_map X euclideanreal (g \ e)" + by (meson contg conte continuous_map_compose continuous_map_in_subtopology) + show "(g \ e) ` topspace X \ T" + using gsub conte continuous_map_image_subset_topspace by fastforce + fix x + assume "x \ S" + then show "(g \ e) x = f x" + using gf \compactin X S\ compactin_subset_topspace e'e by fastforce + qed +qed + + +lemma Tietze_extension_completely_regular: + assumes "completely_regular_space X" "compactin X S" "is_interval T" "T \ {}" + and contf: "continuous_map (subtopology X S) euclidean f" and fim: "f`S \ T" + obtains g where "continuous_map X euclideanreal g" "g ` topspace X \ T" + "\x. x \ S \ g x = f x" +proof - + define Q where "Q \ Kolmogorov_quotient X ` (topspace X)" + obtain g where contg: "continuous_map (subtopology X (Kolmogorov_quotient X ` S)) euclidean g" + and gf: "\x. x \ S \ g(Kolmogorov_quotient X x) = f x" + using Kolmogorov_quotient_lift_exists + by (metis \compactin X S\ contf compactin_subset_topspace open_openin t0_space_def t1_space) + have "S \ topspace X" + by (simp add: \compactin X S\ compactin_subset_topspace) + then have [simp]: "Q \ Kolmogorov_quotient X ` S = Kolmogorov_quotient X ` S" + using Q_def by blast + have creg: "completely_regular_space (subtopology X Q)" + by (simp add: \completely_regular_space X\ completely_regular_space_subtopology) + then have "regular_space (subtopology X Q)" + by (simp add: completely_regular_imp_regular_space) + then have "Hausdorff_space (subtopology X Q)" + using Q_def regular_t0_eq_Hausdorff_space t0_space_Kolmogorov_quotient by blast + moreover + have "compactin (subtopology X Q) (Kolmogorov_quotient X ` S)" + by (metis Q_def \compactin X S\ image_compactin quotient_imp_continuous_map quotient_map_Kolmogorov_quotient) + ultimately obtain h where conth: "continuous_map (subtopology X Q) euclidean h" + and him: "h ` topspace (subtopology X Q) \ T" + and hg: "\x. x \ Kolmogorov_quotient X ` S \ h x = g x" + using Tietze_extension_comp_reg_aux [of "subtopology X Q" "Kolmogorov_quotient X ` S" T g] + apply (simp add: subtopology_subtopology creg contg assms) + using fim gf by blast + show thesis + proof + show "continuous_map X euclideanreal (h \ Kolmogorov_quotient X)" + by (metis Q_def conth continuous_map_compose quotient_imp_continuous_map quotient_map_Kolmogorov_quotient) + show "(h \ Kolmogorov_quotient X) ` topspace X \ T" + using Q_def continuous_map_Kolmogorov_quotient continuous_map_image_subset_topspace him by fastforce + fix x + assume "x \ S" then show "(h \ Kolmogorov_quotient X) x = f x" + by (simp add: gf hg) + qed +qed + +end