# HG changeset patch # User immler # Date 1463774502 -7200 # Node ID c445b0924e3ab18f6e57bac46dfac4a63bb3e1f3 # Parent 9505a883403c80fc05cdd3179d693d5b76068bb6 uniformly continuous function extended continuously on closure diff -r 9505a883403c -r c445b0924e3a src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri May 20 21:21:28 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri May 20 22:01:42 2016 +0200 @@ -6086,6 +6086,158 @@ qed qed +lemma uniformly_continuous_on_extension_at_closure: + fixes f::"'a::metric_space \ 'b::complete_space" + assumes uc: "uniformly_continuous_on X f" + assumes "x \ closure X" + obtains l where "(f \ l) (at x within X)" +proof - + from assms obtain xs where xs: "xs \ x" "\n. xs n \ X" + by (auto simp: closure_sequential) + + from uniformly_continuous_on_Cauchy[OF uc LIMSEQ_imp_Cauchy, OF xs] + obtain l where l: "(\n. f (xs n)) \ l" + by atomize_elim (simp only: convergent_eq_cauchy) + + have "(f \ l) (at x within X)" + proof (safe intro!: Lim_within_LIMSEQ) + fix xs' + assume "\n. xs' n \ x \ xs' n \ X" + and xs': "xs' \ x" + then have "xs' n \ x" "xs' n \ X" for n by auto + + from uniformly_continuous_on_Cauchy[OF uc LIMSEQ_imp_Cauchy, OF \xs' \ x\ \xs' _ \ X\] + obtain l' where l': "(\n. f (xs' n)) \ l'" + by atomize_elim (simp only: convergent_eq_cauchy) + + show "(\n. f (xs' n)) \ l" + proof (rule tendstoI) + fix e::real assume "e > 0" + define e' where "e' \ e / 2" + have "e' > 0" using \e > 0\ by (simp add: e'_def) + + have "\\<^sub>F n in sequentially. dist (f (xs n)) l < e'" + by (simp add: \0 < e'\ l tendstoD) + moreover + from uc[unfolded uniformly_continuous_on_def, rule_format, OF \e' > 0\] + obtain d where d: "d > 0" "\x x'. x \ X \ x' \ X \ dist x x' < d \ dist (f x) (f x') < e'" + by auto + have "\\<^sub>F n in sequentially. dist (xs n) (xs' n) < d" + by (auto intro!: \0 < d\ order_tendstoD tendsto_eq_intros xs xs') + ultimately + show "\\<^sub>F n in sequentially. dist (f (xs' n)) l < e" + proof eventually_elim + case (elim n) + have "dist (f (xs' n)) l \ dist (f (xs n)) (f (xs' n)) + dist (f (xs n)) l" + by (metis dist_triangle dist_commute) + also have "dist (f (xs n)) (f (xs' n)) < e'" + by (auto intro!: d xs \xs' _ \ _\ elim) + also note \dist (f (xs n)) l < e'\ + also have "e' + e' = e" by (simp add: e'_def) + finally show ?case by simp + qed + qed + qed + thus ?thesis .. +qed + +lemma uniformly_continuous_on_extension_on_closure: + fixes f::"'a::metric_space \ 'b::complete_space" + assumes uc: "uniformly_continuous_on X f" + obtains g where "uniformly_continuous_on (closure X) g" "\x. x \ X \ f x = g x" + "\Y h x. X \ Y \ Y \ closure X \ continuous_on Y h \ (\x. x \ X \ f x = h x) \ x \ Y \ h x = g x" +proof - + from uc have cont_f: "continuous_on X f" + by (simp add: uniformly_continuous_imp_continuous) + obtain y where y: "(f \ y x) (at x within X)" if "x \ closure X" for x + apply atomize_elim + apply (rule choice) + using uniformly_continuous_on_extension_at_closure[OF assms] + by metis + let ?g = "\x. if x \ X then f x else y x" + + have "uniformly_continuous_on (closure X) ?g" + unfolding uniformly_continuous_on_def + proof safe + fix e::real assume "e > 0" + define e' where "e' \ e / 3" + have "e' > 0" using \e > 0\ by (simp add: e'_def) + from uc[unfolded uniformly_continuous_on_def, rule_format, OF \0 < e'\] + obtain d where "d > 0" and d: "\x x'. x \ X \ x' \ X \ dist x' x < d \ dist (f x') (f x) < e'" + by auto + define d' where "d' = d / 3" + have "d' > 0" using \d > 0\ by (simp add: d'_def) + show "\d>0. \x\closure X. \x'\closure X. dist x' x < d \ dist (?g x') (?g x) < e" + proof (safe intro!: exI[where x=d'] \d' > 0\) + fix x x' assume x: "x \ closure X" and x': "x' \ closure X" and dist: "dist x' x < d'" + then obtain xs xs' where xs: "xs \ x" "\n. xs n \ X" + and xs': "xs' \ x'" "\n. xs' n \ X" + by (auto simp: closure_sequential) + have "\\<^sub>F n in sequentially. dist (xs' n) x' < d'" + and "\\<^sub>F n in sequentially. dist (xs n) x < d'" + by (auto intro!: \0 < d'\ order_tendstoD tendsto_eq_intros xs xs') + moreover + have "(\x. f (xs x)) \ y x" if "x \ closure X" "x \ X" "xs \ x" "\n. xs n \ X" for xs x + using that not_eventuallyD + by (force intro!: filterlim_compose[OF y[OF \x \ closure X\]] simp: filterlim_at) + then have "(\x. f (xs' x)) \ ?g x'" "(\x. f (xs x)) \ ?g x" + using x x' + by (auto intro!: continuous_on_tendsto_compose[OF cont_f] simp: xs' xs) + then have "\\<^sub>F n in sequentially. dist (f (xs' n)) (?g x') < e'" + "\\<^sub>F n in sequentially. dist (f (xs n)) (?g x) < e'" + by (auto intro!: \0 < e'\ order_tendstoD tendsto_eq_intros) + ultimately + have "\\<^sub>F n in sequentially. dist (?g x') (?g x) < e" + proof eventually_elim + case (elim n) + have "dist (?g x') (?g x) \ + dist (f (xs' n)) (?g x') + dist (f (xs' n)) (f (xs n)) + dist (f (xs n)) (?g x)" + by (metis add.commute add_le_cancel_left dist_commute dist_triangle dist_triangle_le) + also + { + have "dist (xs' n) (xs n) \ dist (xs' n) x' + dist x' x + dist (xs n) x" + by (metis add.commute add_le_cancel_left dist_triangle dist_triangle_le) + also note \dist (xs' n) x' < d'\ + also note \dist x' x < d'\ + also note \dist (xs n) x < d'\ + finally have "dist (xs' n) (xs n) < d" by (simp add: d'_def) + } + with \xs _ \ X\ \xs' _ \ X\ have "dist (f (xs' n)) (f (xs n)) < e'" + by (rule d) + also note \dist (f (xs' n)) (?g x') < e'\ + also note \dist (f (xs n)) (?g x) < e'\ + finally show ?case by (simp add: e'_def) + qed + then show "dist (?g x') (?g x) < e" by simp + qed + qed + moreover have "f x = ?g x" if "x \ X" for x using that by simp + moreover + { + fix Y h x + assume Y: "x \ Y" "X \ Y" "Y \ closure X" and cont_h: "continuous_on Y h" + and extension: "(\x. x \ X \ f x = h x)" + { + assume "x \ X" + have "x \ closure X" using Y by auto + then obtain xs where xs: "xs \ x" "\n. xs n \ X" + by (auto simp: closure_sequential) + from continuous_on_tendsto_compose[OF cont_h xs(1)] xs(2) Y + have "(\x. f (xs x)) \ h x" + by (auto simp: set_mp extension) + moreover + then have "(\x. f (xs x)) \ y x" + using \x \ X\ not_eventuallyD xs(2) + by (force intro!: filterlim_compose[OF y[OF \x \ closure X\]] simp: filterlim_at xs) + ultimately have "h x = y x" by (rule LIMSEQ_unique) + } then + have "h x = ?g x" + using extension by auto + } + ultimately show ?thesis .. +qed + + subsection\Quotient maps\ lemma quotient_map_imp_continuous_open: