# HG changeset patch # User immler # Date 1546882969 -3600 # Node ID 9c22ff18125bc063fbe016da9bf542e746d5fddf # Parent 19d8a59481dbe4a60c6de14aa48a696986a26332 generalized diff -r 19d8a59481db -r 9c22ff18125b src/HOL/Analysis/Elementary_Metric_Spaces.thy --- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy Mon Jan 07 14:57:45 2019 +0100 +++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy Mon Jan 07 18:42:49 2019 +0100 @@ -2743,33 +2743,31 @@ done lemma Times_in_interior_subtopology: - fixes U :: "('a::metric_space \ 'b::metric_space) set" assumes "(x, y) \ U" "openin (subtopology euclidean (S \ T)) U" obtains V W where "openin (subtopology euclidean S) V" "x \ V" "openin (subtopology euclidean T) W" "y \ W" "(V \ W) \ U" proof - - from assms obtain e where "e > 0" and "U \ S \ T" - and e: "\x' y'. \x'\S; y'\T; dist (x', y') (x, y) < e\ \ (x', y') \ U" - by (force simp: openin_euclidean_subtopology_iff) - with assms have "x \ S" "y \ T" - by auto + from assms obtain E where "open E" "U = S \ T \ E" "(x, y) \ E" "x \ S" "y \ T" + by (auto simp: openin_open) + from open_prod_elim[OF \open E\ \(x, y) \ E\] + obtain E1 E2 where "open E1" "open E2" "(x, y) \ E1 \ E2" "E1 \ E2 \ E" + by blast show ?thesis proof - show "openin (subtopology euclidean S) (ball x (e/2) \ S)" - by (simp add: Int_commute openin_open_Int) - show "x \ ball x (e / 2) \ S" - by (simp add: \0 < e\ \x \ S\) - show "openin (subtopology euclidean T) (ball y (e/2) \ T)" - by (simp add: Int_commute openin_open_Int) - show "y \ ball y (e / 2) \ T" - by (simp add: \0 < e\ \y \ T\) - show "(ball x (e / 2) \ S) \ (ball y (e / 2) \ T) \ U" - by clarify (simp add: e dist_Pair_Pair \0 < e\ dist_commute sqrt_sum_squares_half_less) + show "openin (subtopology euclidean S) (E1 \ S)" + "openin (subtopology euclidean T) (E2 \ T)" + using \open E1\ \open E2\ + by (auto simp: openin_open) + show "x \ E1 \ S" "y \ E2 \ T" + using \(x, y) \ E1 \ E2\ \x \ S\ \y \ T\ by auto + show "(E1 \ S) \ (E2 \ T) \ U" + using \E1 \ E2 \ E\ \U = _\ + by (auto simp: ) qed qed lemma openin_Times_eq: - fixes S :: "'a::metric_space set" and T :: "'b::metric_space set" + fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" shows "openin (subtopology euclidean (S \ T)) (S' \ T') \ S' = {} \ T' = {} \ openin (subtopology euclidean S) S' \ openin (subtopology euclidean T) T'" @@ -2811,33 +2809,42 @@ unfolding closedin_closed using closed_Times by blast lemma Lim_transform_within_openin: - fixes a :: "'a::metric_space" assumes f: "(f \ l) (at a within T)" and "openin (subtopology euclidean T) S" "a \ S" and eq: "\x. \x \ S; x \ a\ \ f x = g x" shows "(g \ l) (at a within T)" proof - - obtain \ where "0 < \" and \: "ball a \ \ T \ S" - using assms by (force simp: openin_contains_ball) - then have "a \ ball a \" - by simp - show ?thesis - by (rule Lim_transform_within [OF f \0 < \\ eq]) (use \ in \auto simp: dist_commute subset_iff\) + have "\\<^sub>F x in at a within T. x \ T \ x \ a" + by (simp add: eventually_at_filter) + moreover + from \openin _ _\ obtain U where "open U" "S = T \ U" + by (auto simp: openin_open) + then have "a \ U" using \a \ S\ by auto + from topological_tendstoD[OF tendsto_ident_at \open U\ \a \ U\] + have "\\<^sub>F x in at a within T. x \ U" by auto + ultimately + have "\\<^sub>F x in at a within T. f x = g x" + by eventually_elim (auto simp: \S = _\ eq) + then show ?thesis using f + by (rule Lim_transform_eventually) qed lemma closure_Int_ballI: - fixes S :: "'a :: metric_space set" assumes "\U. \openin (subtopology euclidean S) U; U \ {}\ \ T \ U \ {}" - shows "S \ closure T" -proof (clarsimp simp: closure_approachable dist_commute) - fix x and e::real - assume "x \ S" "0 < e" - with assms [of "S \ ball x e"] show "\y\T. dist x y < e" - by force + shows "S \ closure T" +proof (clarsimp simp: closure_iff_nhds_not_empty) + fix x and A and V + assume "x \ S" "V \ A" "open V" "x \ V" "T \ A = {}" + then have "openin (subtopology euclidean S) (A \ V \ S)" + by (auto simp: openin_open intro!: exI[where x="V"]) + moreover have "A \ V \ S \ {}" using \x \ V\ \V \ A\ \x \ S\ + by auto + ultimately have "T \ (A \ V \ S) \ {}" + by (rule assms) + with \T \ A = {}\ show False by auto qed lemma continuous_on_open_gen: - fixes f :: "'a::metric_space \ 'b::metric_space" assumes "f ` S \ T" shows "continuous_on S f \ (\U. openin (subtopology euclidean T) U @@ -2846,38 +2853,32 @@ proof assume ?lhs then show ?rhs - apply (clarsimp simp: openin_euclidean_subtopology_iff continuous_on_iff) - by (metis assms image_subset_iff) + by (clarsimp simp add: continuous_openin_preimage_eq openin_open) + (metis Int_assoc assms image_subset_iff_subset_vimage inf.absorb_iff1) next - have ope: "openin (subtopology euclidean T) (ball y e \ T)" for y e - by (simp add: Int_commute openin_open_Int) assume R [rule_format]: ?rhs show ?lhs - proof (clarsimp simp add: continuous_on_iff) - fix x and e::real - assume "x \ S" and "0 < e" - then have x: "x \ S \ (f -` ball (f x) e \ f -` T)" - using assms by auto - show "\d>0. \x'\S. dist x' x < d \ dist (f x') (f x) < e" - using R [of "ball (f x) e \ T"] x - by (fastforce simp add: ope openin_euclidean_subtopology_iff [of S] dist_commute) + proof (clarsimp simp add: continuous_openin_preimage_eq) + fix U::"'a set" + assume "open U" + then have "openin (subtopology euclidean S) (S \ f -` (U \ T))" + by (metis R inf_commute openin_open) + then show "openin (subtopology euclidean S) (S \ f -` U)" + by (metis Int_assoc Int_commute assms image_subset_iff_subset_vimage inf.absorb_iff2 vimage_Int) qed qed lemma continuous_openin_preimage: - fixes f :: "'a::metric_space \ 'b::metric_space" - shows - "\continuous_on S f; f ` S \ T; openin (subtopology euclidean T) U\ + "\continuous_on S f; f ` S \ T; openin (subtopology euclidean T) U\ \ openin (subtopology euclidean S) (S \ f -` U)" -by (simp add: continuous_on_open_gen) + by (simp add: continuous_on_open_gen) lemma continuous_on_closed_gen: - fixes f :: "'a::metric_space \ 'b::metric_space" assumes "f ` S \ T" - shows "continuous_on S f \ + shows "continuous_on S f \ (\U. closedin (subtopology euclidean T) U \ closedin (subtopology euclidean S) (S \ f -` U))" - (is "?lhs = ?rhs") + (is "?lhs = ?rhs") proof - have *: "U \ T \ S \ f -` (T - U) = S - (S \ f -` U)" for U using assms by blast @@ -2901,13 +2902,11 @@ qed lemma continuous_closedin_preimage_gen: - fixes f :: "'a::metric_space \ 'b::metric_space" assumes "continuous_on S f" "f ` S \ T" "closedin (subtopology euclidean T) U" shows "closedin (subtopology euclidean S) (S \ f -` U)" using assms continuous_on_closed_gen by blast lemma continuous_transform_within_openin: - fixes a :: "'a::metric_space" assumes "continuous (at a within T) f" and "openin (subtopology euclidean T) S" "a \ S" and eq: "\x. x \ S \ f x = g x"