# HG changeset patch # User paulson # Date 1605717320 0 # Node ID 0e422e806ef3eda4263d879fdfead8854531cee6 # Parent 4eea17b3ac5851c2798199fb71b39cf55aac9bff# Parent 6b3599ff06876682939a00bd429d74c2bdf0b8ca merged diff -r 4eea17b3ac58 -r 0e422e806ef3 src/HOL/Analysis/Elementary_Normed_Spaces.thy --- a/src/HOL/Analysis/Elementary_Normed_Spaces.thy Wed Nov 18 13:44:34 2020 +0100 +++ b/src/HOL/Analysis/Elementary_Normed_Spaces.thy Wed Nov 18 16:35:20 2020 +0000 @@ -85,23 +85,23 @@ subsection \Support\ definition (in monoid_add) support_on :: "'b set \ ('b \ 'a) \ 'b set" - where "support_on s f = {x\s. f x \ 0}" + where "support_on S f = {x\S. f x \ 0}" -lemma in_support_on: "x \ support_on s f \ x \ s \ f x \ 0" +lemma in_support_on: "x \ support_on S f \ x \ S \ f x \ 0" by (simp add: support_on_def) lemma support_on_simps[simp]: "support_on {} f = {}" - "support_on (insert x s) f = - (if f x = 0 then support_on s f else insert x (support_on s f))" - "support_on (s \ t) f = support_on s f \ support_on t f" - "support_on (s \ t) f = support_on s f \ support_on t f" - "support_on (s - t) f = support_on s f - support_on t f" - "support_on (f ` s) g = f ` (support_on s (g \ f))" + "support_on (insert x S) f = + (if f x = 0 then support_on S f else insert x (support_on S f))" + "support_on (S \ T) f = support_on S f \ support_on T f" + "support_on (S \ T) f = support_on S f \ support_on T f" + "support_on (S - T) f = support_on S f - support_on T f" + "support_on (f ` S) g = f ` (support_on S (g \ f))" unfolding support_on_def by auto lemma support_on_cong: - "(\x. x \ s \ f x = 0 \ g x = 0) \ support_on s f = support_on s g" + "(\x. x \ S \ f x = 0 \ g x = 0) \ support_on S f = support_on S g" by (auto simp: support_on_def) lemma support_on_if: "a \ 0 \ support_on A (\x. if P x then a else 0) = {x\A. P x}" @@ -153,14 +153,13 @@ proof (clarsimp, intro conjI impI subsetI) show "\0 \ m; a \ b; x \ {m *\<^sub>R a + c..m *\<^sub>R b + c}\ \ x \ (\x. m *\<^sub>R x + c) ` {a..b}" for x - apply (rule_tac x="inverse m *\<^sub>R (x-c)" in rev_image_eqI) - using False apply (auto simp: pos_le_divideR_eq pos_divideR_le_eq le_diff_eq diff_le_eq) - done + using False + by (rule_tac x="inverse m *\<^sub>R (x-c)" in image_eqI) + (auto simp: pos_le_divideR_eq pos_divideR_le_eq le_diff_eq diff_le_eq) show "\\ 0 \ m; a \ b; x \ {m *\<^sub>R b + c..m *\<^sub>R a + c}\ \ x \ (\x. m *\<^sub>R x + c) ` {a..b}" for x - apply (rule_tac x="inverse m *\<^sub>R (x-c)" in rev_image_eqI) - apply (auto simp add: neg_le_divideR_eq neg_divideR_le_eq not_le le_diff_eq diff_le_eq) - done + by (rule_tac x="inverse m *\<^sub>R (x-c)" in image_eqI) + (auto simp add: neg_le_divideR_eq neg_divideR_le_eq le_diff_eq diff_le_eq) qed qed qed @@ -196,12 +195,12 @@ have "\x'\ball x e. x' \ y \ dist x' y < d" proof (cases "d \ dist x y") case True - then show "\x'\ball x e. x' \ y \ dist x' y < d" + then show ?thesis proof (cases "x = y") case True then have False using \d \ dist x y\ \d>0\ by auto - then show "\x'\ball x e. x' \ y \ dist x' y < d" + then show ?thesis by auto next case False @@ -226,17 +225,9 @@ by (auto simp: dist_commute) moreover have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d" - unfolding dist_norm - apply simp - unfolding norm_minus_cancel - using \d > 0\ \x\y\[unfolded dist_nz] dist_commute[of x y] - unfolding dist_norm - apply auto - done - ultimately show "\x'\ball x e. x' \ y \ dist x' y < d" - apply (rule_tac x = "y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI) - apply auto - done + using \0 < d\ by (fastforce simp: dist_norm) + ultimately show ?thesis + by (rule_tac x = "y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI) auto qed next case False @@ -244,21 +235,15 @@ show "\x' \ ball x e. x' \ y \ dist x' y < d" proof (cases "x = y") case True - obtain z where **: "z \ y" "dist z y < min e d" + obtain z where z: "z \ y" "dist z y < min e d" using perfect_choose_dist[of "min e d" y] using \d > 0\ \e>0\ by auto - show "\x'\ball x e. x' \ y \ dist x' y < d" - unfolding \x = y\ - using \z \ y\ ** - apply (rule_tac x=z in bexI) - apply (auto simp: dist_commute) - done + show ?thesis + by (metis True z dist_commute mem_ball min_less_iff_conj) next case False - then show "\x'\ball x e. x' \ y \ dist x' y < d" - using \d>0\ \d > dist x y\ \?rhs\ - apply (rule_tac x=x in bexI, auto) - done + then show ?thesis + using \d>0\ \d > dist x y\ \?rhs\ by force qed qed } @@ -276,7 +261,7 @@ assume "y \ T" "open T" then obtain r where "0 < r" "\z. dist z y < r \ z \ T" unfolding open_dist by fast - (* choose point between x and y, within distance r of y. *) + \\choose point between @{term x} and @{term y}, within distance @{term r} of @{term y}.\ define k where "k = min 1 (r / (2 * dist x y))" define z where "z = y + scaleR k (x - y)" have z_def2: "z = x + scaleR (1 - k) (y - x)" @@ -287,14 +272,7 @@ then have "z \ T" using \\z. dist z y < r \ z \ T\ by simp have "dist x z < dist x y" - unfolding z_def2 dist_norm - apply (simp add: norm_minus_commute) - apply (simp only: dist_norm [symmetric]) - apply (subgoal_tac "\1 - k\ * dist x y < 1 * dist x y", simp) - apply (rule mult_strict_right_mono) - apply (simp add: k_def \0 < r\ \x \ y\) - apply (simp add: \x \ y\) - done + using \0 < r\ assms by (simp add: z_def2 k_def dist_norm norm_minus_commute) then have "z \ ball x (dist x y)" by simp have "z \ y" @@ -310,31 +288,28 @@ lemma mem_ball_0 [simp]: "x \ ball 0 e \ norm x < e" for x :: "'a::real_normed_vector" - by (simp add: dist_norm) + by simp lemma mem_cball_0 [simp]: "x \ cball 0 e \ norm x \ e" for x :: "'a::real_normed_vector" - by (simp add: dist_norm) + by simp lemma closure_ball [simp]: fixes x :: "'a::real_normed_vector" - shows "0 < e \ closure (ball x e) = cball x e" - apply (rule equalityI) - apply (rule closure_minimal) - apply (rule ball_subset_cball) - apply (rule closed_cball) - apply (rule subsetI, rename_tac y) - apply (simp add: le_less [where 'a=real]) - apply (erule disjE) - apply (rule subsetD [OF closure_subset], simp) - apply (simp add: closure_def, clarify) - apply (rule closure_ball_lemma) - apply simp - done + assumes "0 < e" + shows "closure (ball x e) = cball x e" +proof + show "closure (ball x e) \ cball x e" + using closed_cball closure_minimal by blast + have "\y. dist x y < e \ dist x y = e \ y \ closure (ball x e)" + by (metis Un_iff assms closure_ball_lemma closure_def dist_eq_0_iff mem_Collect_eq mem_ball) + then show "cball x e \ closure (ball x e)" + by force +qed lemma mem_sphere_0 [simp]: "x \ sphere 0 e \ norm x = e" for x :: "'a::real_normed_vector" - by (simp add: dist_norm) + by simp (* In a trivial vector space, this fails for e = 0. *) lemma interior_cball [simp]: @@ -440,20 +415,27 @@ lemma image_add_ball [simp]: fixes a :: "'a::real_normed_vector" shows "(+) b ` ball a r = ball (a+b) r" -apply (intro equalityI subsetI) -apply (force simp: dist_norm) -apply (rule_tac x="x-b" in image_eqI) -apply (auto simp: dist_norm algebra_simps) -done +proof - + { fix x :: 'a + assume "dist (a + b) x < r" + moreover + have "b + (x - b) = x" + by simp + ultimately have "x \ (+) b ` ball a r" + by (metis add.commute dist_add_cancel image_eqI mem_ball) } + then show ?thesis + by (auto simp: add.commute) +qed lemma image_add_cball [simp]: fixes a :: "'a::real_normed_vector" shows "(+) b ` cball a r = cball (a+b) r" -apply (intro equalityI subsetI) -apply (force simp: dist_norm) -apply (rule_tac x="x-b" in image_eqI) -apply (auto simp: dist_norm algebra_simps) -done +proof - + have "\x. dist (a + b) x \ r \ \y\cball a r. x = b + y" + by (metis (no_types) add.commute diff_add_cancel dist_add_cancel2 mem_cball) + then show ?thesis + by (force simp: add.commute) +qed subsection\<^marker>\tag unimportant\ \Various Lemmas on Normed Algebras\ @@ -492,20 +474,21 @@ lemma trivial_limit_at_infinity: "\ trivial_limit (at_infinity :: ('a::{real_normed_vector,perfect_space}) filter)" - unfolding trivial_limit_def eventually_at_infinity - apply clarsimp - apply (subgoal_tac "\x::'a. x \ 0", clarify) - apply (rule_tac x="scaleR (b / norm x) x" in exI, simp) - apply (cut_tac islimpt_UNIV [of "0::'a", unfolded islimpt_def]) - apply (drule_tac x=UNIV in spec, simp) - done +proof - + obtain x::'a where "x \ 0" + by (meson perfect_choose_dist zero_less_one) + then have "b \ norm ((b / norm x) *\<^sub>R x)" for b + by simp + then show ?thesis + unfolding trivial_limit_def eventually_at_infinity + by blast +qed lemma at_within_ball_bot_iff: fixes x y :: "'a::{real_normed_vector,perfect_space}" shows "at x within ball y r = bot \ (r=0 \ x \ cball y r)" - unfolding trivial_limit_within - apply (auto simp add:trivial_limit_within islimpt_ball ) - by (metis le_less_trans less_eq_real_def zero_le_dist) + unfolding trivial_limit_within + by (metis (no_types) cball_empty equals0D islimpt_ball less_linear) subsection \Limits\ @@ -516,9 +499,12 @@ corollary Lim_at_infinityI [intro?]: assumes "\e. e > 0 \ \B. \x. norm x \ B \ dist (f x) l \ e" shows "(f \ l) at_infinity" - apply (simp add: Lim_at_infinity, clarify) - apply (rule ex_forward [OF assms [OF half_gt_zero]], auto) - done +proof - + have "\e. e > 0 \ \B. \x. norm x \ B \ dist (f x) l < e" + by (meson assms dense le_less_trans) + then show ?thesis + using Lim_at_infinity by blast +qed lemma Lim_transform_within_set_eq: fixes a :: "'a::metric_space" and l :: "'b::metric_space" @@ -555,12 +541,13 @@ assumes f: "(f \ 0) F" and g: "eventually (\x. norm(g x) \ B) F" shows "((\z. f z * g z) \ 0) F" proof - - have *: "((\x. norm (f x) * B) \ 0) F" - by (simp add: f tendsto_mult_left_zero tendsto_norm_zero) have "((\x. norm (f x) * norm (g x)) \ 0) F" - apply (rule Lim_null_comparison [OF _ *]) - apply (simp add: eventually_mono [OF g] mult_left_mono) - done + proof (rule Lim_null_comparison) + show "\\<^sub>F x in F. norm (norm (f x) * norm (g x)) \ norm (f x) * B" + by (simp add: eventually_mono [OF g] mult_left_mono) + show "((\x. norm (f x) * B) \ 0) F" + by (simp add: f tendsto_mult_left_zero tendsto_norm_zero) + qed then show ?thesis by (subst tendsto_norm_zero_iff [symmetric]) (simp add: norm_mult) qed @@ -570,12 +557,13 @@ assumes g: "eventually (\x. norm(g x) \ B) F" and f: "(f \ 0) F" shows "((\z. g z * f z) \ 0) F" proof - - have *: "((\x. B * norm (f x)) \ 0) F" - by (simp add: f tendsto_mult_right_zero tendsto_norm_zero) have "((\x. norm (g x) * norm (f x)) \ 0) F" - apply (rule Lim_null_comparison [OF _ *]) - apply (simp add: eventually_mono [OF g] mult_right_mono) - done + proof (rule Lim_null_comparison) + show "\\<^sub>F x in F. norm (norm (g x) * norm (f x)) \ B * norm (f x)" + by (simp add: eventually_mono [OF g] mult_right_mono) + show "((\x. B * norm (f x)) \ 0) F" + by (simp add: f tendsto_mult_right_zero tendsto_norm_zero) + qed then show ?thesis by (subst tendsto_norm_zero_iff [symmetric]) (simp add: norm_mult) qed @@ -597,10 +585,10 @@ by (rule f) finally show ?thesis . qed - show "\\<^sub>F x in net. dist (f x *\<^sub>R g x) 0 < \" - apply (rule eventually_mono [OF eventually_conj [OF tendstoD [OF f B] gB] ]) - apply (auto simp: \0 < \\ field_split_simps * split: if_split_asm) - done + have "\x. \\f x\ < \ / (\B\ + 1); norm (g x) \ B\ \ \f x\ * norm (g x) < \" + by (simp add: "*" pos_less_divide_eq) + then show "\\<^sub>F x in net. dist (f x *\<^sub>R g x) 0 < \" + using \0 < \\ by (auto intro: eventually_mono [OF eventually_conj [OF tendstoD [OF f B] gB]]) qed lemma Lim_norm_ubound: @@ -641,12 +629,10 @@ shows "netlimit (at a) = a" proof (cases "\x. x \ a") case True then obtain x where x: "x \ a" .. - have "\ trivial_limit (at a)" - unfolding trivial_limit_def eventually_at dist_norm - apply clarsimp - apply (rule_tac x="a + scaleR (d / 2) (sgn (x - a))" in exI) - apply (simp add: norm_sgn sgn_zero_iff x) - done + have "\d. 0 < d \ \x. x \ a \ norm (x - a) < d" + by (rule_tac x="a + scaleR (d / 2) (sgn (x - a))" in exI) (simp add: norm_sgn sgn_zero_iff x) + then have "\ trivial_limit (at a)" + by (auto simp: trivial_limit_def eventually_at dist_norm) then show ?thesis by (rule Lim_ident_at [of a UNIV]) qed simp @@ -671,9 +657,7 @@ by (meson less_imp_le not_le order_trans zero_less_one) lemma bounded_pos_less: "bounded S \ (\b>0. \x\ S. norm x < b)" - apply (simp add: bounded_pos) - apply (safe; rule_tac x="b+1" in exI; force) - done + by (metis bounded_pos le_less_trans less_imp_le linordered_field_no_ub) lemma Bseq_eq_bounded: fixes f :: "nat \ 'a::real_normed_vector" @@ -700,9 +684,7 @@ lemma bounded_scaling: fixes S :: "'a::real_normed_vector set" shows "bounded S \ bounded ((\x. c *\<^sub>R x) ` S)" - apply (rule bounded_linear_image, assumption) - apply (rule bounded_linear_scaleR_right) - done + by (simp add: bounded_linear_image bounded_linear_scaleR_right) lemma bounded_scaleR_comp: fixes f :: "'a \ 'b::real_normed_vector" @@ -818,13 +800,13 @@ proof - obtain M where M: "\x. norm (a x * z ^ x) \ M" using summable_imp_bounded [OF sum] by (force simp: bounded_iff) - then have *: "summable (\n. norm (a n) * norm w ^ n)" - by (rule_tac M=M in Abel_lemma) (auto simp: norm_mult norm_power intro: no) show ?thesis - apply (rule series_comparison_complex [of "(\n. of_real(norm(a n) * norm w ^ n))"]) - apply (simp only: summable_complex_of_real *) - apply (auto simp: norm_mult norm_power) - done + proof (rule series_comparison_complex) + have "\n. norm (a n) * norm z ^ n \ M" + by (metis (no_types) M norm_mult norm_power) + then show "summable (\n. complex_of_real (norm (a n) * norm w ^ n))" + using Abel_lemma no norm_ge_zero summable_of_real by blast + qed (auto simp: norm_mult norm_power) qed @@ -931,9 +913,7 @@ have *: "\S T. S \ range F \ T \ range F \ S \ T \ T \ S" by (metis mono image_iff le_cases) show ?thesis - apply (rule compact_chain [OF _ _ *]) - using F apply (blast intro: dest: *)+ - done + using F by (intro compact_chain [OF _ _ *]; blast dest: *) qed text\The Baire property of dense sets\ @@ -1006,17 +986,17 @@ qed let ?\ = "\n X. openin (top_of_set S) X \ X \ {} \ S \ closure X \ ?g n \ closure X \ ball x e" - have "closure (S \ ball x (e / 2)) \ closure(ball x (e/2))" + have "closure (S \ ball x (e/2)) \ closure(ball x (e/2))" by (simp add: closure_mono) also have "... \ ball x e" using \e > 0\ by auto - finally have "closure (S \ ball x (e / 2)) \ ball x e" . - moreover have"openin (top_of_set S) (S \ ball x (e / 2))" "S \ ball x (e / 2) \ {}" + finally have "closure (S \ ball x (e/2)) \ ball x e" . + moreover have"openin (top_of_set S) (S \ ball x (e/2))" "S \ ball x (e/2) \ {}" using \0 < e\ \x \ S\ by auto - ultimately obtain Y where Y: "?\ 0 Y \ Y \ S \ ball x (e / 2)" + ultimately obtain Y where Y: "?\ 0 Y \ Y \ S \ ball x (e/2)" using * [of "S \ ball x (e/2)" 0] by metis show thesis - proof (rule exE [OF dependent_nat_choice [of ?\ "\n X Y. Y \ X"]]) + proof (rule exE [OF dependent_nat_choice]) show "\x. ?\ 0 x" using Y by auto show "\Y. ?\ (Suc n) Y \ Y \ X" if "?\ n X" for X n @@ -1159,20 +1139,16 @@ { fix y assume "dist y (c *\<^sub>R x) < e * \c\" + then have "norm (c *\<^sub>R ((1 / c) *\<^sub>R y - x)) < e * norm c" + by (simp add: \c \ 0\ dist_norm scale_right_diff_distrib) then have "norm ((1 / c) *\<^sub>R y - x) < e" - unfolding dist_norm - using norm_scaleR[of c "(1 / c) *\<^sub>R y - x", unfolded scaleR_right_diff_distrib, unfolded scaleR_scaleR] assms(1) - assms(1)[unfolded zero_less_abs_iff[symmetric]] by (simp del:zero_less_abs_iff) + by (simp add: \c \ 0\) then have "y \ (*\<^sub>R) c ` s" using rev_image_eqI[of "(1 / c) *\<^sub>R y" s y "(*\<^sub>R) c"] - using e[THEN spec[where x="(1 / c) *\<^sub>R y"]] - using assms(1) - unfolding dist_norm scaleR_scaleR - by auto + by (simp add: \c \ 0\ dist_norm e) } ultimately have "\e>0. \x'. dist x' (c *\<^sub>R x) < e \ x' \ (*\<^sub>R) c ` s" - apply (rule_tac x="e * \c\" in exI, auto) - done + by (rule_tac x="e * \c\" in exI, auto) } then show ?thesis unfolding open_dist by auto qed @@ -1243,11 +1219,7 @@ by (auto simp: diff_diff_eq) then show "x \ (+) a ` interior S" unfolding image_iff - apply (rule_tac x="x - a" in bexI) - unfolding mem_interior - using \e > 0\ - apply auto - done + by (metis \0 < e\ add.commute diff_add_cancel mem_interior) next fix x assume "x \ (+) a ` interior S" @@ -1301,11 +1273,7 @@ shows "compact {x + y | x y. x \ s \ y \ t}" proof - have *: "{x + y | x y. x \ s \ y \ t} = (\z. fst z + snd z) ` (s \ t)" - apply auto - unfolding image_iff - apply (rule_tac x="(xa, y)" in bexI) - apply auto - done + by (fastforce simp: image_iff) have "continuous_on (s \ t) (\z. fst z + snd z)" unfolding continuous_on by (rule ballI) (intro tendsto_intros) then show ?thesis @@ -1318,10 +1286,8 @@ and "compact t" shows "compact {x - y | x y. x \ s \ y \ t}" proof- - have "{x - y | x y. x\s \ y \ t} = {x + y | x y. x \ s \ y \ (uminus ` t)}" - apply auto - apply (rule_tac x= xa in exI, auto) - done + have "{x - y | x y. x\s \ y \ t} = {x + y | x y. x \ s \ y \ (uminus ` t)}" + using diff_conv_add_uminus by force then show ?thesis using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto qed @@ -1396,11 +1362,7 @@ using f(3) by auto then have "l \ ?S" - using \l' \ S\ - apply auto - apply (rule_tac x=l' in exI) - apply (rule_tac x="l - l'" in exI, auto) - done + using \l' \ S\ by force } moreover have "?S = (\x\ S. \y \ T. {x + y})" by force @@ -1428,7 +1390,7 @@ have "(\x\ S. \y \ uminus ` T. {x + y}) = (\x\ S. \y \ T. {x - y})" by force then show ?thesis - using compact_closed_sums[OF assms(1) closed_negations[OF assms(2)]] by auto + by (metis assms closed_negations compact_closed_sums) qed lemma closed_compact_differences: @@ -1504,39 +1466,30 @@ subsection\<^marker>\tag unimportant\\Homeomorphisms\ lemma homeomorphic_scaling: - fixes s :: "'a::real_normed_vector set" + fixes S :: "'a::real_normed_vector set" assumes "c \ 0" - shows "s homeomorphic ((\x. c *\<^sub>R x) ` s)" + shows "S homeomorphic ((\x. c *\<^sub>R x) ` S)" unfolding homeomorphic_minimal apply (rule_tac x="\x. c *\<^sub>R x" in exI) apply (rule_tac x="\x. (1 / c) *\<^sub>R x" in exI) - using assms - apply (auto simp: continuous_intros) - done + using assms by (auto simp: continuous_intros) lemma homeomorphic_translation: - fixes s :: "'a::real_normed_vector set" - shows "s homeomorphic ((\x. a + x) ` s)" + fixes S :: "'a::real_normed_vector set" + shows "S homeomorphic ((\x. a + x) ` S)" unfolding homeomorphic_minimal apply (rule_tac x="\x. a + x" in exI) apply (rule_tac x="\x. -a + x" in exI) - using continuous_on_add [OF continuous_on_const continuous_on_id, of s a] - continuous_on_add [OF continuous_on_const continuous_on_id, of "plus a ` s" "- a"] - apply auto - done + by (auto simp: continuous_intros) lemma homeomorphic_affinity: - fixes s :: "'a::real_normed_vector set" + fixes S :: "'a::real_normed_vector set" assumes "c \ 0" - shows "s homeomorphic ((\x. a + c *\<^sub>R x) ` s)" + shows "S homeomorphic ((\x. a + c *\<^sub>R x) ` S)" proof - - have *: "(+) a ` (*\<^sub>R) c ` s = (\x. a + c *\<^sub>R x) ` s" by auto + have *: "(+) a ` (*\<^sub>R) c ` S = (\x. a + c *\<^sub>R x) ` S" by auto show ?thesis - using homeomorphic_trans - using homeomorphic_scaling[OF assms, of s] - using homeomorphic_translation[of "(\x. c *\<^sub>R x) ` s" a] - unfolding * - by auto + by (metis "*" assms homeomorphic_scaling homeomorphic_trans homeomorphic_translation) qed lemma homeomorphic_balls: @@ -1549,16 +1502,12 @@ apply(rule_tac x="\x. b + (e/d) *\<^sub>R (x - a)" in exI) apply(rule_tac x="\x. a + (d/e) *\<^sub>R (x - b)" in exI) using assms - apply (auto intro!: continuous_intros - simp: dist_commute dist_norm pos_divide_less_eq mult_strict_left_mono) - done + by (auto intro!: continuous_intros simp: dist_commute dist_norm pos_divide_less_eq) show ?cth unfolding homeomorphic_minimal apply(rule_tac x="\x. b + (e/d) *\<^sub>R (x - a)" in exI) apply(rule_tac x="\x. a + (d/e) *\<^sub>R (x - b)" in exI) using assms - apply (auto intro!: continuous_intros - simp: dist_commute dist_norm pos_divide_le_eq mult_strict_left_mono) - done + by (auto intro!: continuous_intros simp: dist_commute dist_norm pos_divide_le_eq) qed lemma homeomorphic_spheres: @@ -1569,9 +1518,7 @@ apply(rule_tac x="\x. b + (e/d) *\<^sub>R (x - a)" in exI) apply(rule_tac x="\x. a + (d/e) *\<^sub>R (x - b)" in exI) using assms - apply (auto intro!: continuous_intros - simp: dist_commute dist_norm pos_divide_less_eq mult_strict_left_mono) - done + by (auto intro!: continuous_intros simp: dist_commute dist_norm pos_divide_less_eq) lemma homeomorphic_ball01_UNIV: "ball (0::'a::real_normed_vector) 1 homeomorphic (UNIV:: 'a set)" @@ -1585,24 +1532,21 @@ then show "(\z::'a. z /\<^sub>R (1 - norm z)) ` ?B = ?U" by blast have "x \ range (\z. (1 / (1 + norm z)) *\<^sub>R z)" if "norm x < 1" for x::'a - apply (rule_tac x="x /\<^sub>R (1 - norm x)" in image_eqI) - using that apply (auto simp: field_split_simps) - done + using that + by (rule_tac x="x /\<^sub>R (1 - norm x)" in image_eqI) (auto simp: field_split_simps) then show "(\z::'a. z /\<^sub>R (1 + norm z)) ` ?U = ?B" by (force simp: field_split_simps dest: add_less_zeroD) show "continuous_on (ball 0 1) (\z. z /\<^sub>R (1 - norm z))" by (rule continuous_intros | force)+ - show "continuous_on UNIV (\z. z /\<^sub>R (1 + norm z))" - apply (intro continuous_intros) - apply (metis le_add_same_cancel1 norm_ge_zero not_le zero_less_one) - done + have 0: "\z. 1 + norm z \ 0" + by (metis (no_types) le_add_same_cancel1 norm_ge_zero not_one_le_zero) + then show "continuous_on UNIV (\z. z /\<^sub>R (1 + norm z))" + by (auto intro!: continuous_intros) show "\x. x \ ball 0 1 \ x /\<^sub>R (1 - norm x) /\<^sub>R (1 + norm (x /\<^sub>R (1 - norm x))) = x" by (auto simp: field_split_simps) show "\y. y /\<^sub>R (1 + norm y) /\<^sub>R (1 - norm (y /\<^sub>R (1 + norm y))) = y" - apply (auto simp: field_split_simps) - apply (metis le_add_same_cancel1 norm_ge_zero not_le zero_less_one) - done + using 0 by (auto simp: field_split_simps) qed proposition homeomorphic_ball_UNIV: @@ -1625,14 +1569,12 @@ by (fastforce simp add: Set.image_subset_iff cong: conj_cong) next case False - then obtain z where z: "z \ S" "f z \ f x" + then obtain z where "z \ S" "f z \ f x" by blast - have finn: "finite {norm (z - f x) |z. z \ f ` S - {f x}}" + moreover have finn: "finite {norm (z - f x) |z. z \ f ` S - {f x}}" using assms by simp - then have *: "0 < Inf{norm(z - f x) | z. z \ f ` S - {f x}}" - apply (rule finite_imp_less_Inf) - using z apply force+ - done + ultimately have *: "0 < Inf{norm(z - f x) | z. z \ f ` S - {f x}}" + by (force intro: finite_imp_less_Inf) show ?thesis by (force intro!: * cInf_le_finite [OF finn]) qed @@ -1720,24 +1662,43 @@ { fix x assume x: "x \ S" then obtain e where "e>0" and ele: "\y. \y \ S; f y \ f x\ \ e \ norm (f y - f x)" using conf no [OF x] by auto - then have e2: "0 \ e / 2" + then have e2: "0 \ e/2" by simp - have "f y = f x" if "y \ S" and ccs: "f y \ connected_component_set (f ` S) (f x)" for y - apply (rule ccontr) - using connected_closed [of "connected_component_set (f ` S) (f x)"] \e>0\ - apply (simp add: del: ex_simps) - apply (drule spec [where x="cball (f x) (e / 2)"]) - apply (drule spec [where x="- ball(f x) e"]) - apply (auto simp: dist_norm open_closed [symmetric] simp del: le_divide_eq_numeral1 dest!: connected_component_in) - apply (metis diff_self e2 ele norm_minus_commute norm_zero not_less) - using centre_in_cball connected_component_refl_eq e2 x apply blast - using ccs - apply (force simp: cball_def dist_norm norm_minus_commute dest: ele [OF \y \ S\]) - done + define F where "F \ connected_component_set (f ` S) (f x)" + have False if "y \ S" and ccs: "f y \ F" and not: "f y \ f x" for y + proof - + define C where "C \ cball (f x) (e/2)" + define D where "D \ - ball (f x) e" + have disj: "C \ D = {}" + unfolding C_def D_def using \0 < e\ by fastforce + moreover have FCD: "F \ C \ D" + proof - + have "t \ C \ t \ D" if "t \ F" for t + proof - + obtain y where "y \ S" "t = f y" + using F_def \t \ F\ connected_component_in by blast + then show ?thesis + by (metis C_def ComplI D_def centre_in_cball dist_norm e2 ele mem_ball norm_minus_commute not_le) + qed + then show ?thesis + by auto + qed + ultimately have "C \ F = {} \ D \ F = {}" + using connected_closed [of "F"] \e>0\ not + unfolding C_def D_def + by (metis Elementary_Metric_Spaces.open_ball F_def closed_cball connected_connected_component inf_bot_left open_closed) + moreover have "C \ F \ {}" + unfolding disjoint_iff + by (metis FCD ComplD image_eqI mem_Collect_eq subsetD x D_def F_def Un_iff \0 < e\ centre_in_ball connected_component_refl_eq) + moreover have "D \ F \ {}" + unfolding disjoint_iff + by (metis ComplI D_def ccs dist_norm ele mem_ball norm_minus_commute not not_le that(1)) + ultimately show ?thesis by metis + qed moreover have "connected_component_set (f ` S) (f x) \ f ` S" by (auto simp: connected_component_in) ultimately have "connected_component_set (f ` S) (f x) = {f x}" - by (auto simp: x) + by (auto simp: x F_def) } with assms show ?thesis by blast