# HG changeset patch # User immler # Date 1546856974 -3600 # Node ID 1331e57b54c61d022494122470c566054c4fbc85 # Parent d692ef26021ef69836cecd9511af43cd4acab385 moved material from Connected.thy to more appropriate places diff -r d692ef26021e -r 1331e57b54c6 src/HOL/Analysis/Connected.thy --- a/src/HOL/Analysis/Connected.thy Mon Jan 07 10:22:22 2019 +0100 +++ b/src/HOL/Analysis/Connected.thy Mon Jan 07 11:29:34 2019 +0100 @@ -605,7 +605,7 @@ done -text \Proving a function is constant by proving that a level set is open\ +subsection%unimportant \Proving a function is constant by proving that a level set is open\ lemma continuous_levelset_openin_cases: fixes f :: "_ \ 'b::t1_space" @@ -634,1574 +634,8 @@ using assms (3,4) by fast -text \Some arithmetical combinations (more to prove).\ -lemma open_scaling[intro]: - fixes s :: "'a::real_normed_vector set" - assumes "c \ 0" - and "open s" - shows "open((\x. c *\<^sub>R x) ` s)" -proof - - { - fix x - assume "x \ s" - then obtain e where "e>0" - and e:"\x'. dist x' x < e \ x' \ s" using assms(2)[unfolded open_dist, THEN bspec[where x=x]] - by auto - have "e * \c\ > 0" - using assms(1)[unfolded zero_less_abs_iff[symmetric]] \e>0\ by auto - moreover - { - fix y - assume "dist y (c *\<^sub>R x) < e * \c\" - 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) - 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 - } - 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 - } - then show ?thesis unfolding open_dist by auto -qed - -lemma minus_image_eq_vimage: - fixes A :: "'a::ab_group_add set" - shows "(\x. - x) ` A = (\x. - x) -` A" - by (auto intro!: image_eqI [where f="\x. - x"]) - -lemma open_negations: - fixes S :: "'a::real_normed_vector set" - shows "open S \ open ((\x. - x) ` S)" - using open_scaling [of "- 1" S] by simp - -lemma open_translation: - fixes S :: "'a::real_normed_vector set" - assumes "open S" - shows "open((\x. a + x) ` S)" -proof - - { - fix x - have "continuous (at x) (\x. x - a)" - by (intro continuous_diff continuous_ident continuous_const) - } - moreover have "{x. x - a \ S} = (+) a ` S" - by force - ultimately show ?thesis - by (metis assms continuous_open_vimage vimage_def) -qed - -lemma open_neg_translation: - fixes s :: "'a::real_normed_vector set" - assumes "open s" - shows "open((\x. a - x) ` s)" - using open_translation[OF open_negations[OF assms], of a] - by (auto simp: image_image) - -lemma open_affinity: - fixes S :: "'a::real_normed_vector set" - assumes "open S" "c \ 0" - shows "open ((\x. a + c *\<^sub>R x) ` S)" -proof - - have *: "(\x. a + c *\<^sub>R x) = (\x. a + x) \ (\x. c *\<^sub>R x)" - unfolding o_def .. - have "(+) a ` (*\<^sub>R) c ` S = ((+) a \ (*\<^sub>R) c) ` S" - by auto - then show ?thesis - using assms open_translation[of "(*\<^sub>R) c ` S" a] - unfolding * - by auto -qed - -lemma interior_translation: - fixes S :: "'a::real_normed_vector set" - shows "interior ((\x. a + x) ` S) = (\x. a + x) ` (interior S)" -proof (rule set_eqI, rule) - fix x - assume "x \ interior ((+) a ` S)" - then obtain e where "e > 0" and e: "ball x e \ (+) a ` S" - unfolding mem_interior by auto - then have "ball (x - a) e \ S" - unfolding subset_eq Ball_def mem_ball dist_norm - 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 -next - fix x - assume "x \ (+) a ` interior S" - then obtain y e where "e > 0" and e: "ball y e \ S" and y: "x = a + y" - unfolding image_iff Bex_def mem_interior by auto - { - fix z - have *: "a + y - z = y + a - z" by auto - assume "z \ ball x e" - then have "z - a \ S" - using e[unfolded subset_eq, THEN bspec[where x="z - a"]] - unfolding mem_ball dist_norm y group_add_class.diff_diff_eq2 * - by auto - then have "z \ (+) a ` S" - unfolding image_iff by (auto intro!: bexI[where x="z - a"]) - } - then have "ball x e \ (+) a ` S" - unfolding subset_eq by auto - then show "x \ interior ((+) a ` S)" - unfolding mem_interior using \e > 0\ by auto -qed - -subsection \Continuity implies uniform continuity on a compact domain\ - -text\From the proof of the Heine-Borel theorem: Lemma 2 in section 3.7, page 69 of -J. C. Burkill and H. Burkill. A Second Course in Mathematical Analysis (CUP, 2002)\ - -lemma Heine_Borel_lemma: - assumes "compact S" and Ssub: "S \ \\" and opn: "\G. G \ \ \ open G" - obtains e where "0 < e" "\x. x \ S \ \G \ \. ball x e \ G" -proof - - have False if neg: "\e. 0 < e \ \x \ S. \G \ \. \ ball x e \ G" - proof - - have "\x \ S. \G \ \. \ ball x (1 / Suc n) \ G" for n - using neg by simp - then obtain f where "\n. f n \ S" and fG: "\G n. G \ \ \ \ ball (f n) (1 / Suc n) \ G" - by metis - then obtain l r where "l \ S" "strict_mono r" and to_l: "(f \ r) \ l" - using \compact S\ compact_def that by metis - then obtain G where "l \ G" "G \ \" - using Ssub by auto - then obtain e where "0 < e" and e: "\z. dist z l < e \ z \ G" - using opn open_dist by blast - obtain N1 where N1: "\n. n \ N1 \ dist (f (r n)) l < e/2" - using to_l apply (simp add: lim_sequentially) - using \0 < e\ half_gt_zero that by blast - obtain N2 where N2: "of_nat N2 > 2/e" - using reals_Archimedean2 by blast - obtain x where "x \ ball (f (r (max N1 N2))) (1 / real (Suc (r (max N1 N2))))" and "x \ G" - using fG [OF \G \ \\, of "r (max N1 N2)"] by blast - then have "dist (f (r (max N1 N2))) x < 1 / real (Suc (r (max N1 N2)))" - by simp - also have "... \ 1 / real (Suc (max N1 N2))" - apply (simp add: divide_simps del: max.bounded_iff) - using \strict_mono r\ seq_suble by blast - also have "... \ 1 / real (Suc N2)" - by (simp add: field_simps) - also have "... < e/2" - using N2 \0 < e\ by (simp add: field_simps) - finally have "dist (f (r (max N1 N2))) x < e / 2" . - moreover have "dist (f (r (max N1 N2))) l < e/2" - using N1 max.cobounded1 by blast - ultimately have "dist x l < e" - using dist_triangle_half_r by blast - then show ?thesis - using e \x \ G\ by blast - qed - then show ?thesis - by (meson that) -qed - -lemma compact_uniformly_equicontinuous: - assumes "compact S" - and cont: "\x e. \x \ S; 0 < e\ - \ \d. 0 < d \ - (\f \ \. \x' \ S. dist x' x < d \ dist (f x') (f x) < e)" - and "0 < e" - obtains d where "0 < d" - "\f x x'. \f \ \; x \ S; x' \ S; dist x' x < d\ \ dist (f x') (f x) < e" -proof - - obtain d where d_pos: "\x e. \x \ S; 0 < e\ \ 0 < d x e" - and d_dist : "\x x' e f. \dist x' x < d x e; x \ S; x' \ S; 0 < e; f \ \\ \ dist (f x') (f x) < e" - using cont by metis - let ?\ = "((\x. ball x (d x (e / 2))) ` S)" - have Ssub: "S \ \ ?\" - by clarsimp (metis d_pos \0 < e\ dist_self half_gt_zero_iff) - then obtain k where "0 < k" and k: "\x. x \ S \ \G \ ?\. ball x k \ G" - by (rule Heine_Borel_lemma [OF \compact S\]) auto - moreover have "dist (f v) (f u) < e" if "f \ \" "u \ S" "v \ S" "dist v u < k" for f u v - proof - - obtain G where "G \ ?\" "u \ G" "v \ G" - using k that - by (metis \dist v u < k\ \u \ S\ \0 < k\ centre_in_ball subsetD dist_commute mem_ball) - then obtain w where w: "dist w u < d w (e / 2)" "dist w v < d w (e / 2)" "w \ S" - by auto - with that d_dist have "dist (f w) (f v) < e/2" - by (metis \0 < e\ dist_commute half_gt_zero) - moreover - have "dist (f w) (f u) < e/2" - using that d_dist w by (metis \0 < e\ dist_commute divide_pos_pos zero_less_numeral) - ultimately show ?thesis - using dist_triangle_half_r by blast - qed - ultimately show ?thesis using that by blast -qed - -corollary compact_uniformly_continuous: - fixes f :: "'a :: metric_space \ 'b :: metric_space" - assumes f: "continuous_on S f" and S: "compact S" - shows "uniformly_continuous_on S f" - using f - unfolding continuous_on_iff uniformly_continuous_on_def - by (force intro: compact_uniformly_equicontinuous [OF S, of "{f}"]) - -subsection%unimportant \Topological stuff about the set of Reals\ - -lemma open_real: - fixes s :: "real set" - shows "open s \ (\x \ s. \e>0. \x'. \x' - x\ < e --> x' \ s)" - unfolding open_dist dist_norm by simp - -lemma islimpt_approachable_real: - fixes s :: "real set" - shows "x islimpt s \ (\e>0. \x'\ s. x' \ x \ \x' - x\ < e)" - unfolding islimpt_approachable dist_norm by simp - -lemma closed_real: - fixes s :: "real set" - shows "closed s \ (\x. (\e>0. \x' \ s. x' \ x \ \x' - x\ < e) \ x \ s)" - unfolding closed_limpt islimpt_approachable dist_norm by simp - -lemma continuous_at_real_range: - fixes f :: "'a::real_normed_vector \ real" - shows "continuous (at x) f \ (\e>0. \d>0. \x'. norm(x' - x) < d --> \f x' - f x\ < e)" - unfolding continuous_at - unfolding Lim_at - unfolding dist_norm - apply auto - apply (erule_tac x=e in allE, auto) - apply (rule_tac x=d in exI, auto) - apply (erule_tac x=x' in allE, auto) - apply (erule_tac x=e in allE, auto) - done - -lemma continuous_on_real_range: - fixes f :: "'a::real_normed_vector \ real" - shows "continuous_on s f \ - (\x \ s. \e>0. \d>0. (\x' \ s. norm(x' - x) < d \ \f x' - f x\ < e))" - unfolding continuous_on_iff dist_norm by simp - - -subsection%unimportant \Cartesian products\ - -lemma bounded_Times: - assumes "bounded s" "bounded t" - shows "bounded (s \ t)" -proof - - obtain x y a b where "\z\s. dist x z \ a" "\z\t. dist y z \ b" - using assms [unfolded bounded_def] by auto - then have "\z\s \ t. dist (x, y) z \ sqrt (a\<^sup>2 + b\<^sup>2)" - by (auto simp: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono) - then show ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto -qed - -lemma mem_Times_iff: "x \ A \ B \ fst x \ A \ snd x \ B" - by (induct x) simp - -lemma seq_compact_Times: "seq_compact s \ seq_compact t \ seq_compact (s \ t)" - unfolding seq_compact_def - apply clarify - apply (drule_tac x="fst \ f" in spec) - apply (drule mp, simp add: mem_Times_iff) - apply (clarify, rename_tac l1 r1) - apply (drule_tac x="snd \ f \ r1" in spec) - apply (drule mp, simp add: mem_Times_iff) - apply (clarify, rename_tac l2 r2) - apply (rule_tac x="(l1, l2)" in rev_bexI, simp) - apply (rule_tac x="r1 \ r2" in exI) - apply (rule conjI, simp add: strict_mono_def) - apply (drule_tac f=r2 in LIMSEQ_subseq_LIMSEQ, assumption) - apply (drule (1) tendsto_Pair) back - apply (simp add: o_def) - done - -lemma compact_Times: - assumes "compact s" "compact t" - shows "compact (s \ t)" -proof (rule compactI) - fix C - assume C: "\t\C. open t" "s \ t \ \C" - have "\x\s. \a. open a \ x \ a \ (\d\C. finite d \ a \ t \ \d)" - proof - fix x - assume "x \ s" - have "\y\t. \a b c. c \ C \ open a \ open b \ x \ a \ y \ b \ a \ b \ c" (is "\y\t. ?P y") - proof - fix y - assume "y \ t" - with \x \ s\ C obtain c where "c \ C" "(x, y) \ c" "open c" by auto - then show "?P y" by (auto elim!: open_prod_elim) - qed - then obtain a b c where b: "\y. y \ t \ open (b y)" - and c: "\y. y \ t \ c y \ C \ open (a y) \ open (b y) \ x \ a y \ y \ b y \ a y \ b y \ c y" - by metis - then have "\y\t. open (b y)" "t \ (\y\t. b y)" by auto - with compactE_image[OF \compact t\] obtain D where D: "D \ t" "finite D" "t \ (\y\D. b y)" - by metis - moreover from D c have "(\y\D. a y) \ t \ (\y\D. c y)" - by (fastforce simp: subset_eq) - ultimately show "\a. open a \ x \ a \ (\d\C. finite d \ a \ t \ \d)" - using c by (intro exI[of _ "c`D"] exI[of _ "\(a`D)"] conjI) (auto intro!: open_INT) - qed - then obtain a d where a: "\x. x\s \ open (a x)" "s \ (\x\s. a x)" - and d: "\x. x \ s \ d x \ C \ finite (d x) \ a x \ t \ \d x" - unfolding subset_eq UN_iff by metis - moreover - from compactE_image[OF \compact s\ a] - obtain e where e: "e \ s" "finite e" and s: "s \ (\x\e. a x)" - by auto - moreover - { - from s have "s \ t \ (\x\e. a x \ t)" - by auto - also have "\ \ (\x\e. \d x)" - using d \e \ s\ by (intro UN_mono) auto - finally have "s \ t \ (\x\e. \d x)" . - } - ultimately show "\C'\C. finite C' \ s \ t \ \C'" - by (intro exI[of _ "(\x\e. d x)"]) (auto simp: subset_eq) -qed - -text\Hence some useful properties follow quite easily.\ - -lemma compact_scaling: - fixes s :: "'a::real_normed_vector set" - assumes "compact s" - shows "compact ((\x. c *\<^sub>R x) ` s)" -proof - - let ?f = "\x. scaleR c x" - have *: "bounded_linear ?f" by (rule bounded_linear_scaleR_right) - show ?thesis - using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f] - using linear_continuous_at[OF *] assms - by auto -qed - -lemma compact_negations: - fixes s :: "'a::real_normed_vector set" - assumes "compact s" - shows "compact ((\x. - x) ` s)" - using compact_scaling [OF assms, of "- 1"] by auto - -lemma compact_sums: - fixes s t :: "'a::real_normed_vector set" - assumes "compact s" - and "compact t" - 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 - have "continuous_on (s \ t) (\z. fst z + snd z)" - unfolding continuous_on by (rule ballI) (intro tendsto_intros) - then show ?thesis - unfolding * using compact_continuous_image compact_Times [OF assms] by auto -qed - -lemma compact_differences: - fixes s t :: "'a::real_normed_vector set" - assumes "compact s" - 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 - then show ?thesis - using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto -qed - -lemma compact_translation: - fixes s :: "'a::real_normed_vector set" - assumes "compact s" - shows "compact ((\x. a + x) ` s)" -proof - - have "{x + y |x y. x \ s \ y \ {a}} = (\x. a + x) ` s" - by auto - then show ?thesis - using compact_sums[OF assms compact_sing[of a]] by auto -qed - -lemma compact_affinity: - fixes s :: "'a::real_normed_vector set" - assumes "compact s" - shows "compact ((\x. a + c *\<^sub>R x) ` s)" -proof - - have "(+) a ` (*\<^sub>R) c ` s = (\x. a + c *\<^sub>R x) ` s" - by auto - then show ?thesis - using compact_translation[OF compact_scaling[OF assms], of a c] by auto -qed - -text \Hence we get the following.\ - -lemma compact_sup_maxdistance: - fixes s :: "'a::metric_space set" - assumes "compact s" - and "s \ {}" - shows "\x\s. \y\s. \u\s. \v\s. dist u v \ dist x y" -proof - - have "compact (s \ s)" - using \compact s\ by (intro compact_Times) - moreover have "s \ s \ {}" - using \s \ {}\ by auto - moreover have "continuous_on (s \ s) (\x. dist (fst x) (snd x))" - by (intro continuous_at_imp_continuous_on ballI continuous_intros) - ultimately show ?thesis - using continuous_attains_sup[of "s \ s" "\x. dist (fst x) (snd x)"] by auto -qed - - -subsection \The diameter of a set\ - -definition%important diameter :: "'a::metric_space set \ real" where - "diameter S = (if S = {} then 0 else SUP (x,y)\S\S. dist x y)" - -lemma diameter_empty [simp]: "diameter{} = 0" - by (auto simp: diameter_def) - -lemma diameter_singleton [simp]: "diameter{x} = 0" - by (auto simp: diameter_def) - -lemma diameter_le: - assumes "S \ {} \ 0 \ d" - and no: "\x y. \x \ S; y \ S\ \ norm(x - y) \ d" - shows "diameter S \ d" -using assms - by (auto simp: dist_norm diameter_def intro: cSUP_least) - -lemma diameter_bounded_bound: - fixes s :: "'a :: metric_space set" - assumes s: "bounded s" "x \ s" "y \ s" - shows "dist x y \ diameter s" -proof - - from s obtain z d where z: "\x. x \ s \ dist z x \ d" - unfolding bounded_def by auto - have "bdd_above (case_prod dist ` (s\s))" - proof (intro bdd_aboveI, safe) - fix a b - assume "a \ s" "b \ s" - with z[of a] z[of b] dist_triangle[of a b z] - show "dist a b \ 2 * d" - by (simp add: dist_commute) - qed - moreover have "(x,y) \ s\s" using s by auto - ultimately have "dist x y \ (SUP (x,y)\s\s. dist x y)" - by (rule cSUP_upper2) simp - with \x \ s\ show ?thesis - by (auto simp: diameter_def) -qed - -lemma diameter_lower_bounded: - fixes s :: "'a :: metric_space set" - assumes s: "bounded s" - and d: "0 < d" "d < diameter s" - shows "\x\s. \y\s. d < dist x y" -proof (rule ccontr) - assume contr: "\ ?thesis" - moreover have "s \ {}" - using d by (auto simp: diameter_def) - ultimately have "diameter s \ d" - by (auto simp: not_less diameter_def intro!: cSUP_least) - with \d < diameter s\ show False by auto -qed - -lemma diameter_bounded: - assumes "bounded s" - shows "\x\s. \y\s. dist x y \ diameter s" - and "\d>0. d < diameter s \ (\x\s. \y\s. dist x y > d)" - using diameter_bounded_bound[of s] diameter_lower_bounded[of s] assms - by auto - -lemma bounded_two_points: - "bounded S \ (\e. \x\S. \y\S. dist x y \ e)" - apply (rule iffI) - subgoal using diameter_bounded(1) by auto - subgoal using bounded_any_center[of S] by meson - done - -lemma diameter_compact_attained: - assumes "compact s" - and "s \ {}" - shows "\x\s. \y\s. dist x y = diameter s" -proof - - have b: "bounded s" using assms(1) - by (rule compact_imp_bounded) - then obtain x y where xys: "x\s" "y\s" - and xy: "\u\s. \v\s. dist u v \ dist x y" - using compact_sup_maxdistance[OF assms] by auto - then have "diameter s \ dist x y" - unfolding diameter_def - apply clarsimp - apply (rule cSUP_least, fast+) - done - then show ?thesis - by (metis b diameter_bounded_bound order_antisym xys) -qed - -lemma diameter_ge_0: - assumes "bounded S" shows "0 \ diameter S" - by (metis all_not_in_conv assms diameter_bounded_bound diameter_empty dist_self order_refl) - -lemma diameter_subset: - assumes "S \ T" "bounded T" - shows "diameter S \ diameter T" -proof (cases "S = {} \ T = {}") - case True - with assms show ?thesis - by (force simp: diameter_ge_0) -next - case False - then have "bdd_above ((\x. case x of (x, xa) \ dist x xa) ` (T \ T))" - using \bounded T\ diameter_bounded_bound by (force simp: bdd_above_def) - with False \S \ T\ show ?thesis - apply (simp add: diameter_def) - apply (rule cSUP_subset_mono, auto) - done -qed - -lemma diameter_closure: - assumes "bounded S" - shows "diameter(closure S) = diameter S" -proof (rule order_antisym) - have "False" if "diameter S < diameter (closure S)" - proof - - define d where "d = diameter(closure S) - diameter(S)" - have "d > 0" - using that by (simp add: d_def) - then have "diameter(closure(S)) - d / 2 < diameter(closure(S))" - by simp - have dd: "diameter (closure S) - d / 2 = (diameter(closure(S)) + diameter(S)) / 2" - by (simp add: d_def divide_simps) - have bocl: "bounded (closure S)" - using assms by blast - moreover have "0 \ diameter S" - using assms diameter_ge_0 by blast - ultimately obtain x y where "x \ closure S" "y \ closure S" and xy: "diameter(closure(S)) - d / 2 < dist x y" - using diameter_bounded(2) [OF bocl, rule_format, of "diameter(closure(S)) - d / 2"] \d > 0\ d_def by auto - then obtain x' y' where x'y': "x' \ S" "dist x' x < d/4" "y' \ S" "dist y' y < d/4" - using closure_approachable - by (metis \0 < d\ zero_less_divide_iff zero_less_numeral) - then have "dist x' y' \ diameter S" - using assms diameter_bounded_bound by blast - with x'y' have "dist x y \ d / 4 + diameter S + d / 4" - by (meson add_mono_thms_linordered_semiring(1) dist_triangle dist_triangle3 less_eq_real_def order_trans) - then show ?thesis - using xy d_def by linarith - qed - then show "diameter (closure S) \ diameter S" - by fastforce - next - show "diameter S \ diameter (closure S)" - by (simp add: assms bounded_closure closure_subset diameter_subset) -qed - -lemma diameter_cball [simp]: - fixes a :: "'a::euclidean_space" - shows "diameter(cball a r) = (if r < 0 then 0 else 2*r)" -proof - - have "diameter(cball a r) = 2*r" if "r \ 0" - proof (rule order_antisym) - show "diameter (cball a r) \ 2*r" - proof (rule diameter_le) - fix x y assume "x \ cball a r" "y \ cball a r" - then have "norm (x - a) \ r" "norm (a - y) \ r" - by (auto simp: dist_norm norm_minus_commute) - then have "norm (x - y) \ r+r" - using norm_diff_triangle_le by blast - then show "norm (x - y) \ 2*r" by simp - qed (simp add: that) - have "2*r = dist (a + r *\<^sub>R (SOME i. i \ Basis)) (a - r *\<^sub>R (SOME i. i \ Basis))" - apply (simp add: dist_norm) - by (metis abs_of_nonneg mult.right_neutral norm_numeral norm_scaleR norm_some_Basis real_norm_def scaleR_2 that) - also have "... \ diameter (cball a r)" - apply (rule diameter_bounded_bound) - using that by (auto simp: dist_norm) - finally show "2*r \ diameter (cball a r)" . - qed - then show ?thesis by simp -qed - -lemma diameter_ball [simp]: - fixes a :: "'a::euclidean_space" - shows "diameter(ball a r) = (if r < 0 then 0 else 2*r)" -proof - - have "diameter(ball a r) = 2*r" if "r > 0" - by (metis bounded_ball diameter_closure closure_ball diameter_cball less_eq_real_def linorder_not_less that) - then show ?thesis - by (simp add: diameter_def) -qed - -lemma diameter_closed_interval [simp]: "diameter {a..b} = (if b < a then 0 else b-a)" -proof - - have "{a .. b} = cball ((a+b)/2) ((b-a)/2)" - by (auto simp: dist_norm abs_if divide_simps split: if_split_asm) - then show ?thesis - by simp -qed - -lemma diameter_open_interval [simp]: "diameter {a<.. \ {}" "S \ \\" and ope: "\B. B \ \ \ open B" - obtains \ where "0 < \" "\T. \T \ S; diameter T < \\ \ \B \ \. T \ B" -proof (cases "S = {}") - case True - then show ?thesis - by (metis \\ \ {}\ zero_less_one empty_subsetI equals0I subset_trans that) -next - case False - { fix x assume "x \ S" - then obtain C where C: "x \ C" "C \ \" - using \S \ \\\ by blast - then obtain r where r: "r>0" "ball x (2*r) \ C" - by (metis mult.commute mult_2_right not_le ope openE field_sum_of_halves zero_le_numeral zero_less_mult_iff) - then have "\r C. r > 0 \ ball x (2*r) \ C \ C \ \" - using C by blast - } - then obtain r where r: "\x. x \ S \ r x > 0 \ (\C \ \. ball x (2*r x) \ C)" - by metis - then have "S \ (\x \ S. ball x (r x))" - by auto - then obtain \ where "finite \" "S \ \\" and \: "\ \ (\x. ball x (r x)) ` S" - by (rule compactE [OF \compact S\]) auto - then obtain S0 where "S0 \ S" "finite S0" and S0: "\ = (\x. ball x (r x)) ` S0" - by (meson finite_subset_image) - then have "S0 \ {}" - using False \S \ \\\ by auto - define \ where "\ = Inf (r ` S0)" - have "\ > 0" - using \finite S0\ \S0 \ S\ \S0 \ {}\ r by (auto simp: \_def finite_less_Inf_iff) - show ?thesis - proof - show "0 < \" - by (simp add: \0 < \\) - show "\B \ \. T \ B" if "T \ S" and dia: "diameter T < \" for T - proof (cases "T = {}") - case True - then show ?thesis - using \\ \ {}\ by blast - next - case False - then obtain y where "y \ T" by blast - then have "y \ S" - using \T \ S\ by auto - then obtain x where "x \ S0" and x: "y \ ball x (r x)" - using \S \ \\\ S0 that by blast - have "ball y \ \ ball y (r x)" - by (metis \_def \S0 \ {}\ \finite S0\ \x \ S0\ empty_is_image finite_imageI finite_less_Inf_iff imageI less_irrefl not_le subset_ball) - also have "... \ ball x (2*r x)" - by clarsimp (metis dist_commute dist_triangle_less_add mem_ball mult_2 x) - finally obtain C where "C \ \" "ball y \ \ C" - by (meson r \S0 \ S\ \x \ S0\ dual_order.trans subsetCE) - have "bounded T" - using \compact S\ bounded_subset compact_imp_bounded \T \ S\ by blast - then have "T \ ball y \" - using \y \ T\ dia diameter_bounded_bound by fastforce - then show ?thesis - apply (rule_tac x=C in bexI) - using \ball y \ \ C\ \C \ \\ by auto - qed - qed -qed - -lemma diameter_cbox: - fixes a b::"'a::euclidean_space" - shows "(\i \ Basis. a \ i \ b \ i) \ diameter (cbox a b) = dist a b" - by (force simp: diameter_def intro!: cSup_eq_maximum L2_set_mono - simp: euclidean_dist_l2[where 'a='a] cbox_def dist_norm) - -subsection \Separation between points and sets\ - -proposition separate_point_closed: - fixes s :: "'a::heine_borel set" - assumes "closed s" and "a \ s" - shows "\d>0. \x\s. d \ dist a x" -proof (cases "s = {}") - case True - then show ?thesis by(auto intro!: exI[where x=1]) -next - case False - from assms obtain x where "x\s" "\y\s. dist a x \ dist a y" - using \s \ {}\ by (blast intro: distance_attains_inf [of s a]) - with \x\s\ show ?thesis using dist_pos_lt[of a x] and\a \ s\ - by blast -qed - -proposition separate_compact_closed: - fixes s t :: "'a::heine_borel set" - assumes "compact s" - and t: "closed t" "s \ t = {}" - shows "\d>0. \x\s. \y\t. d \ dist x y" -proof cases - assume "s \ {} \ t \ {}" - then have "s \ {}" "t \ {}" by auto - let ?inf = "\x. infdist x t" - have "continuous_on s ?inf" - by (auto intro!: continuous_at_imp_continuous_on continuous_infdist continuous_ident) - then obtain x where x: "x \ s" "\y\s. ?inf x \ ?inf y" - using continuous_attains_inf[OF \compact s\ \s \ {}\] by auto - then have "0 < ?inf x" - using t \t \ {}\ in_closed_iff_infdist_zero by (auto simp: less_le infdist_nonneg) - moreover have "\x'\s. \y\t. ?inf x \ dist x' y" - using x by (auto intro: order_trans infdist_le) - ultimately show ?thesis by auto -qed (auto intro!: exI[of _ 1]) - -proposition separate_closed_compact: - fixes s t :: "'a::heine_borel set" - assumes "closed s" - and "compact t" - and "s \ t = {}" - shows "\d>0. \x\s. \y\t. d \ dist x y" -proof - - have *: "t \ s = {}" - using assms(3) by auto - show ?thesis - using separate_compact_closed[OF assms(2,1) *] by (force simp: dist_commute) -qed - -proposition compact_in_open_separated: - fixes A::"'a::heine_borel set" - assumes "A \ {}" - assumes "compact A" - assumes "open B" - assumes "A \ B" - obtains e where "e > 0" "{x. infdist x A \ e} \ B" -proof atomize_elim - have "closed (- B)" "compact A" "- B \ A = {}" - using assms by (auto simp: open_Diff compact_eq_bounded_closed) - from separate_closed_compact[OF this] - obtain d'::real where d': "d'>0" "\x y. x \ B \ y \ A \ d' \ dist x y" - by auto - define d where "d = d' / 2" - hence "d>0" "d < d'" using d' by auto - with d' have d: "\x y. x \ B \ y \ A \ d < dist x y" - by force - show "\e>0. {x. infdist x A \ e} \ B" - proof (rule ccontr) - assume "\e. 0 < e \ {x. infdist x A \ e} \ B" - with \d > 0\ obtain x where x: "infdist x A \ d" "x \ B" - by auto - from assms have "closed A" "A \ {}" by (auto simp: compact_eq_bounded_closed) - from infdist_attains_inf[OF this] - obtain y where y: "y \ A" "infdist x A = dist x y" - by auto - have "dist x y \ d" using x y by simp - also have "\ < dist x y" using y d x by auto - finally show False by simp - qed -qed - - -subsection%unimportant \Compact sets and the closure operation\ - -lemma closed_scaling: - fixes S :: "'a::real_normed_vector set" - assumes "closed S" - shows "closed ((\x. c *\<^sub>R x) ` S)" -proof (cases "c = 0") - case True then show ?thesis - by (auto simp: image_constant_conv) -next - case False - from assms have "closed ((\x. inverse c *\<^sub>R x) -` S)" - by (simp add: continuous_closed_vimage) - also have "(\x. inverse c *\<^sub>R x) -` S = (\x. c *\<^sub>R x) ` S" - using \c \ 0\ by (auto elim: image_eqI [rotated]) - finally show ?thesis . -qed - -lemma closed_negations: - fixes S :: "'a::real_normed_vector set" - assumes "closed S" - shows "closed ((\x. -x) ` S)" - using closed_scaling[OF assms, of "- 1"] by simp - -lemma compact_closed_sums: - fixes S :: "'a::real_normed_vector set" - assumes "compact S" and "closed T" - shows "closed (\x\ S. \y \ T. {x + y})" -proof - - let ?S = "{x + y |x y. x \ S \ y \ T}" - { - fix x l - assume as: "\n. x n \ ?S" "(x \ l) sequentially" - from as(1) obtain f where f: "\n. x n = fst (f n) + snd (f n)" "\n. fst (f n) \ S" "\n. snd (f n) \ T" - using choice[of "\n y. x n = (fst y) + (snd y) \ fst y \ S \ snd y \ T"] by auto - obtain l' r where "l'\S" and r: "strict_mono r" and lr: "(((\n. fst (f n)) \ r) \ l') sequentially" - using assms(1)[unfolded compact_def, THEN spec[where x="\ n. fst (f n)"]] using f(2) by auto - have "((\n. snd (f (r n))) \ l - l') sequentially" - using tendsto_diff[OF LIMSEQ_subseq_LIMSEQ[OF as(2) r] lr] and f(1) - unfolding o_def - by auto - then have "l - l' \ T" - using assms(2)[unfolded closed_sequential_limits, - THEN spec[where x="\ n. snd (f (r n))"], - THEN spec[where x="l - l'"]] - 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 - } - moreover have "?S = (\x\ S. \y \ T. {x + y})" - by force - ultimately show ?thesis - unfolding closed_sequential_limits - by (metis (no_types, lifting)) -qed - -lemma closed_compact_sums: - fixes S T :: "'a::real_normed_vector set" - assumes "closed S" "compact T" - shows "closed (\x\ S. \y \ T. {x + y})" -proof - - have "(\x\ T. \y \ S. {x + y}) = (\x\ S. \y \ T. {x + y})" - by auto - then show ?thesis - using compact_closed_sums[OF assms(2,1)] by simp -qed - -lemma compact_closed_differences: - fixes S T :: "'a::real_normed_vector set" - assumes "compact S" "closed T" - shows "closed (\x\ S. \y \ T. {x - y})" -proof - - 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 -qed - -lemma closed_compact_differences: - fixes S T :: "'a::real_normed_vector set" - assumes "closed S" "compact T" - shows "closed (\x\ S. \y \ T. {x - y})" -proof - - have "(\x\ S. \y \ uminus ` T. {x + y}) = {x - y |x y. x \ S \ y \ T}" - by auto - then show ?thesis - using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp -qed - -lemma closed_translation: - fixes a :: "'a::real_normed_vector" - assumes "closed S" - shows "closed ((\x. a + x) ` S)" -proof - - have "(\x\ {a}. \y \ S. {x + y}) = ((+) a ` S)" by auto - then show ?thesis - using compact_closed_sums[OF compact_sing[of a] assms] by auto -qed - -lemma closure_translation: - fixes a :: "'a::real_normed_vector" - shows "closure ((\x. a + x) ` s) = (\x. a + x) ` (closure s)" -proof - - have *: "(+) a ` (- s) = - (+) a ` s" - apply auto - unfolding image_iff - apply (rule_tac x="x - a" in bexI, auto) - done - show ?thesis - unfolding closure_interior translation_Compl - using interior_translation[of a "- s"] - unfolding * - by auto -qed - -lemma frontier_translation: - fixes a :: "'a::real_normed_vector" - shows "frontier((\x. a + x) ` s) = (\x. a + x) ` (frontier s)" - unfolding frontier_def translation_diff interior_translation closure_translation - by auto - -lemma sphere_translation: - fixes a :: "'n::real_normed_vector" - shows "sphere (a+c) r = (+) a ` sphere c r" -apply safe -apply (rule_tac x="x-a" in image_eqI) -apply (auto simp: dist_norm algebra_simps) -done - -lemma cball_translation: - fixes a :: "'n::real_normed_vector" - shows "cball (a+c) r = (+) a ` cball c r" -apply safe -apply (rule_tac x="x-a" in image_eqI) -apply (auto simp: dist_norm algebra_simps) -done - -lemma ball_translation: - fixes a :: "'n::real_normed_vector" - shows "ball (a+c) r = (+) a ` ball c r" -apply safe -apply (rule_tac x="x-a" in image_eqI) -apply (auto simp: dist_norm algebra_simps) -done - - -subsection%unimportant \Closure of halfspaces and hyperplanes\ - -lemma continuous_on_closed_Collect_le: - fixes f g :: "'a::t2_space \ real" - assumes f: "continuous_on s f" and g: "continuous_on s g" and s: "closed s" - shows "closed {x \ s. f x \ g x}" -proof - - have "closed ((\x. g x - f x) -` {0..} \ s)" - using closed_real_atLeast continuous_on_diff [OF g f] - by (simp add: continuous_on_closed_vimage [OF s]) - also have "((\x. g x - f x) -` {0..} \ s) = {x\s. f x \ g x}" - by auto - finally show ?thesis . -qed - -lemma continuous_at_inner: "continuous (at x) (inner a)" - unfolding continuous_at by (intro tendsto_intros) - -lemma closed_halfspace_le: "closed {x. inner a x \ b}" - by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) - -lemma closed_halfspace_ge: "closed {x. inner a x \ b}" - by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) - -lemma closed_hyperplane: "closed {x. inner a x = b}" - by (simp add: closed_Collect_eq continuous_on_inner continuous_on_const continuous_on_id) - -lemma closed_halfspace_component_le: "closed {x::'a::euclidean_space. x\i \ a}" - by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) - -lemma closed_halfspace_component_ge: "closed {x::'a::euclidean_space. x\i \ a}" - by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) - -lemma closed_interval_left: - fixes b :: "'a::euclidean_space" - shows "closed {x::'a. \i\Basis. x\i \ b\i}" - by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) - -lemma closed_interval_right: - fixes a :: "'a::euclidean_space" - shows "closed {x::'a. \i\Basis. a\i \ x\i}" - by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) - -lemma continuous_le_on_closure: - fixes a::real - assumes f: "continuous_on (closure s) f" - and x: "x \ closure(s)" - and xlo: "\x. x \ s ==> f(x) \ a" - shows "f(x) \ a" - using image_closure_subset [OF f] - using image_closure_subset [OF f] closed_halfspace_le [of "1::real" a] assms - by force - -lemma continuous_ge_on_closure: - fixes a::real - assumes f: "continuous_on (closure s) f" - and x: "x \ closure(s)" - and xlo: "\x. x \ s ==> f(x) \ a" - shows "f(x) \ a" - using image_closure_subset [OF f] closed_halfspace_ge [of a "1::real"] assms - by force - -lemma Lim_component_le: - fixes f :: "'a \ 'b::euclidean_space" - assumes "(f \ l) net" - and "\ (trivial_limit net)" - and "eventually (\x. f(x)\i \ b) net" - shows "l\i \ b" - by (rule tendsto_le[OF assms(2) tendsto_const tendsto_inner[OF assms(1) tendsto_const] assms(3)]) - -lemma Lim_component_ge: - fixes f :: "'a \ 'b::euclidean_space" - assumes "(f \ l) net" - and "\ (trivial_limit net)" - and "eventually (\x. b \ (f x)\i) net" - shows "b \ l\i" - by (rule tendsto_le[OF assms(2) tendsto_inner[OF assms(1) tendsto_const] tendsto_const assms(3)]) - -lemma Lim_component_eq: - fixes f :: "'a \ 'b::euclidean_space" - assumes net: "(f \ l) net" "\ trivial_limit net" - and ev:"eventually (\x. f(x)\i = b) net" - shows "l\i = b" - using ev[unfolded order_eq_iff eventually_conj_iff] - using Lim_component_ge[OF net, of b i] - using Lim_component_le[OF net, of i b] - by auto - -text \Limits relative to a union.\ - -lemma eventually_within_Un: - "eventually P (at x within (s \ t)) \ - eventually P (at x within s) \ eventually P (at x within t)" - unfolding eventually_at_filter - by (auto elim!: eventually_rev_mp) - -lemma Lim_within_union: - "(f \ l) (at x within (s \ t)) \ - (f \ l) (at x within s) \ (f \ l) (at x within t)" - unfolding tendsto_def - by (auto simp: eventually_within_Un) - -lemma Lim_topological: - "(f \ l) net \ - trivial_limit net \ (\S. open S \ l \ S \ eventually (\x. f x \ S) net)" - unfolding tendsto_def trivial_limit_eq by auto - -text \Continuity relative to a union.\ - -lemma continuous_on_Un_local: - "\closedin (subtopology euclidean (s \ t)) s; closedin (subtopology euclidean (s \ t)) t; - continuous_on s f; continuous_on t f\ - \ continuous_on (s \ t) f" - unfolding continuous_on closedin_limpt - by (metis Lim_trivial_limit Lim_within_union Un_iff trivial_limit_within) - -lemma continuous_on_cases_local: - "\closedin (subtopology euclidean (s \ t)) s; closedin (subtopology euclidean (s \ t)) t; - continuous_on s f; continuous_on t g; - \x. \x \ s \ \P x \ x \ t \ P x\ \ f x = g x\ - \ continuous_on (s \ t) (\x. if P x then f x else g x)" - by (rule continuous_on_Un_local) (auto intro: continuous_on_eq) - -lemma continuous_on_cases_le: - fixes h :: "'a :: topological_space \ real" - assumes "continuous_on {t \ s. h t \ a} f" - and "continuous_on {t \ s. a \ h t} g" - and h: "continuous_on s h" - and "\t. \t \ s; h t = a\ \ f t = g t" - shows "continuous_on s (\t. if h t \ a then f(t) else g(t))" -proof - - have s: "s = (s \ h -` atMost a) \ (s \ h -` atLeast a)" - by force - have 1: "closedin (subtopology euclidean s) (s \ h -` atMost a)" - by (rule continuous_closedin_preimage [OF h closed_atMost]) - have 2: "closedin (subtopology euclidean s) (s \ h -` atLeast a)" - by (rule continuous_closedin_preimage [OF h closed_atLeast]) - have eq: "s \ h -` {..a} = {t \ s. h t \ a}" "s \ h -` {a..} = {t \ s. a \ h t}" - by auto - show ?thesis - apply (rule continuous_on_subset [of s, OF _ order_refl]) - apply (subst s) - apply (rule continuous_on_cases_local) - using 1 2 s assms apply (auto simp: eq) - done -qed - -lemma continuous_on_cases_1: - fixes s :: "real set" - assumes "continuous_on {t \ s. t \ a} f" - and "continuous_on {t \ s. a \ t} g" - and "a \ s \ f a = g a" - shows "continuous_on s (\t. if t \ a then f(t) else g(t))" -using assms -by (auto simp: continuous_on_id intro: continuous_on_cases_le [where h = id, simplified]) - -subsubsection\Some more convenient intermediate-value theorem formulations\ - -lemma connected_ivt_hyperplane: - assumes "connected S" and xy: "x \ S" "y \ S" and b: "inner a x \ b" "b \ inner a y" - shows "\z \ S. inner a z = b" -proof (rule ccontr) - assume as:"\ (\z\S. inner a z = b)" - let ?A = "{x. inner a x < b}" - let ?B = "{x. inner a x > b}" - have "open ?A" "open ?B" - using open_halfspace_lt and open_halfspace_gt by auto - moreover have "?A \ ?B = {}" by auto - moreover have "S \ ?A \ ?B" using as by auto - ultimately show False - using \connected S\[unfolded connected_def not_ex, - THEN spec[where x="?A"], THEN spec[where x="?B"]] - using xy b by auto -qed - -lemma connected_ivt_component: - fixes x::"'a::euclidean_space" - shows "connected S \ x \ S \ y \ S \ x\k \ a \ a \ y\k \ (\z\S. z\k = a)" - using connected_ivt_hyperplane[of S x y "k::'a" a] - by (auto simp: inner_commute) - -lemma image_affinity_cbox: fixes m::real - fixes a b c :: "'a::euclidean_space" - shows "(\x. m *\<^sub>R x + c) ` cbox a b = - (if cbox a b = {} then {} - else (if 0 \ m then cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) - else cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c)))" -proof (cases "m = 0") - case True - { - fix x - assume "\i\Basis. x \ i \ c \ i" "\i\Basis. c \ i \ x \ i" - then have "x = c" - by (simp add: dual_order.antisym euclidean_eqI) - } - moreover have "c \ cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c)" - unfolding True by (auto simp: cbox_sing) - ultimately show ?thesis using True by (auto simp: cbox_def) -next - case False - { - fix y - assume "\i\Basis. a \ i \ y \ i" "\i\Basis. y \ i \ b \ i" "m > 0" - then have "\i\Basis. (m *\<^sub>R a + c) \ i \ (m *\<^sub>R y + c) \ i" and "\i\Basis. (m *\<^sub>R y + c) \ i \ (m *\<^sub>R b + c) \ i" - by (auto simp: inner_distrib) - } - moreover - { - fix y - assume "\i\Basis. a \ i \ y \ i" "\i\Basis. y \ i \ b \ i" "m < 0" - then have "\i\Basis. (m *\<^sub>R b + c) \ i \ (m *\<^sub>R y + c) \ i" and "\i\Basis. (m *\<^sub>R y + c) \ i \ (m *\<^sub>R a + c) \ i" - by (auto simp: mult_left_mono_neg inner_distrib) - } - moreover - { - fix y - assume "m > 0" and "\i\Basis. (m *\<^sub>R a + c) \ i \ y \ i" and "\i\Basis. y \ i \ (m *\<^sub>R b + c) \ i" - then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" - unfolding image_iff Bex_def mem_box - apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) - apply (auto simp: pos_le_divide_eq pos_divide_le_eq mult.commute inner_distrib inner_diff_left) - done - } - moreover - { - fix y - assume "\i\Basis. (m *\<^sub>R b + c) \ i \ y \ i" "\i\Basis. y \ i \ (m *\<^sub>R a + c) \ i" "m < 0" - then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" - unfolding image_iff Bex_def mem_box - apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) - apply (auto simp: neg_le_divide_eq neg_divide_le_eq mult.commute inner_distrib inner_diff_left) - done - } - ultimately show ?thesis using False by (auto simp: cbox_def) -qed - -lemma image_smult_cbox:"(\x. m *\<^sub>R (x::_::euclidean_space)) ` cbox a b = - (if cbox a b = {} then {} else if 0 \ m then cbox (m *\<^sub>R a) (m *\<^sub>R b) else cbox (m *\<^sub>R b) (m *\<^sub>R a))" - using image_affinity_cbox[of m 0 a b] by auto - -lemma islimpt_greaterThanLessThan1: - fixes a b::"'a::{linorder_topology, dense_order}" - assumes "a < b" - shows "a islimpt {a<.. T" - from open_right[OF this \a < b\] - obtain c where c: "a < c" "{a.. T" by auto - with assms dense[of a "min c b"] - show "\y\{a<.. T \ y \ a" - by (metis atLeastLessThan_iff greaterThanLessThan_iff min_less_iff_conj - not_le order.strict_implies_order subset_eq) -qed - -lemma islimpt_greaterThanLessThan2: - fixes a b::"'a::{linorder_topology, dense_order}" - assumes "a < b" - shows "b islimpt {a<.. T" - from open_left[OF this \a < b\] - obtain c where c: "c < b" "{c<..b} \ T" by auto - with assms dense[of "max a c" b] - show "\y\{a<.. T \ y \ b" - by (metis greaterThanAtMost_iff greaterThanLessThan_iff max_less_iff_conj - not_le order.strict_implies_order subset_eq) -qed - -lemma closure_greaterThanLessThan[simp]: - fixes a b::"'a::{linorder_topology, dense_order}" - shows "a < b \ closure {a <..< b} = {a .. b}" (is "_ \ ?l = ?r") -proof - have "?l \ closure ?r" - by (rule closure_mono) auto - thus "closure {a<.. {a..b}" by simp -qed (auto simp: closure_def order.order_iff_strict islimpt_greaterThanLessThan1 - islimpt_greaterThanLessThan2) - -lemma closure_greaterThan[simp]: - fixes a b::"'a::{no_top, linorder_topology, dense_order}" - shows "closure {a<..} = {a..}" -proof - - from gt_ex obtain b where "a < b" by auto - hence "{a<..} = {a<.. {b..}" by auto - also have "closure \ = {a..}" using \a < b\ unfolding closure_Un - by auto - finally show ?thesis . -qed - -lemma closure_lessThan[simp]: - fixes b::"'a::{no_bot, linorder_topology, dense_order}" - shows "closure {.. {..a}" by auto - also have "closure \ = {..b}" using \a < b\ unfolding closure_Un - by auto - finally show ?thesis . -qed - -lemma closure_atLeastLessThan[simp]: - fixes a b::"'a::{linorder_topology, dense_order}" - assumes "a < b" - shows "closure {a ..< b} = {a .. b}" -proof - - from assms have "{a ..< b} = {a} \ {a <..< b}" by auto - also have "closure \ = {a .. b}" unfolding closure_Un - by (auto simp: assms less_imp_le) - finally show ?thesis . -qed - -lemma closure_greaterThanAtMost[simp]: - fixes a b::"'a::{linorder_topology, dense_order}" - assumes "a < b" - shows "closure {a <.. b} = {a .. b}" -proof - - from assms have "{a <.. b} = {b} \ {a <..< b}" by auto - also have "closure \ = {a .. b}" unfolding closure_Un - by (auto simp: assms less_imp_le) - finally show ?thesis . -qed - - -subsection \Homeomorphisms\ - -definition%important "homeomorphism s t f g \ - (\x\s. (g(f x) = x)) \ (f ` s = t) \ continuous_on s f \ - (\y\t. (f(g y) = y)) \ (g ` t = s) \ continuous_on t g" - -lemma homeomorphismI [intro?]: - assumes "continuous_on S f" "continuous_on T g" - "f ` S \ T" "g ` T \ S" "\x. x \ S \ g(f x) = x" "\y. y \ T \ f(g y) = y" - shows "homeomorphism S T f g" - using assms by (force simp: homeomorphism_def) - -lemma homeomorphism_translation: - fixes a :: "'a :: real_normed_vector" - shows "homeomorphism ((+) a ` S) S ((+) (- a)) ((+) a)" -unfolding homeomorphism_def by (auto simp: algebra_simps continuous_intros) - -lemma homeomorphism_ident: "homeomorphism T T (\a. a) (\a. a)" - by (rule homeomorphismI) (auto simp: continuous_on_id) - -lemma homeomorphism_compose: - assumes "homeomorphism S T f g" "homeomorphism T U h k" - shows "homeomorphism S U (h o f) (g o k)" - using assms - unfolding homeomorphism_def - by (intro conjI ballI continuous_on_compose) (auto simp: image_comp [symmetric]) - -lemma homeomorphism_symD: "homeomorphism S t f g \ homeomorphism t S g f" - by (simp add: homeomorphism_def) - -lemma homeomorphism_sym: "homeomorphism S t f g = homeomorphism t S g f" - by (force simp: homeomorphism_def) - -definition%important homeomorphic :: "'a::topological_space set \ 'b::topological_space set \ bool" - (infixr "homeomorphic" 60) - where "s homeomorphic t \ (\f g. homeomorphism s t f g)" - -lemma homeomorphic_empty [iff]: - "S homeomorphic {} \ S = {}" "{} homeomorphic S \ S = {}" - by (auto simp: homeomorphic_def homeomorphism_def) - -lemma homeomorphic_refl: "s homeomorphic s" - unfolding homeomorphic_def homeomorphism_def - using continuous_on_id - apply (rule_tac x = "(\x. x)" in exI) - apply (rule_tac x = "(\x. x)" in exI) - apply blast - done - -lemma homeomorphic_sym: "s homeomorphic t \ t homeomorphic s" - unfolding homeomorphic_def homeomorphism_def - by blast - -lemma homeomorphic_trans [trans]: - assumes "S homeomorphic T" - and "T homeomorphic U" - shows "S homeomorphic U" - using assms - unfolding homeomorphic_def -by (metis homeomorphism_compose) - -lemma homeomorphic_minimal: - "s homeomorphic t \ - (\f g. (\x\s. f(x) \ t \ (g(f(x)) = x)) \ - (\y\t. g(y) \ s \ (f(g(y)) = y)) \ - continuous_on s f \ continuous_on t g)" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - by (fastforce simp: homeomorphic_def homeomorphism_def) -next - assume ?rhs - then show ?lhs - apply clarify - unfolding homeomorphic_def homeomorphism_def - by (metis equalityI image_subset_iff subsetI) - qed - -lemma homeomorphicI [intro?]: - "\f ` S = T; g ` T = S; - continuous_on S f; continuous_on T g; - \x. x \ S \ g(f(x)) = x; - \y. y \ T \ f(g(y)) = y\ \ S homeomorphic T" -unfolding homeomorphic_def homeomorphism_def by metis - -lemma homeomorphism_of_subsets: - "\homeomorphism S T f g; S' \ S; T'' \ T; f ` S' = T'\ - \ homeomorphism S' T' f g" -apply (auto simp: homeomorphism_def elim!: continuous_on_subset) -by (metis subsetD imageI) - -lemma homeomorphism_apply1: "\homeomorphism S T f g; x \ S\ \ g(f x) = x" - by (simp add: homeomorphism_def) - -lemma homeomorphism_apply2: "\homeomorphism S T f g; x \ T\ \ f(g x) = x" - by (simp add: homeomorphism_def) - -lemma homeomorphism_image1: "homeomorphism S T f g \ f ` S = T" - by (simp add: homeomorphism_def) - -lemma homeomorphism_image2: "homeomorphism S T f g \ g ` T = S" - by (simp add: homeomorphism_def) - -lemma homeomorphism_cont1: "homeomorphism S T f g \ continuous_on S f" - by (simp add: homeomorphism_def) - -lemma homeomorphism_cont2: "homeomorphism S T f g \ continuous_on T g" - by (simp add: homeomorphism_def) - -lemma continuous_on_no_limpt: - "(\x. \ x islimpt S) \ continuous_on S f" - unfolding continuous_on_def - by (metis UNIV_I empty_iff eventually_at_topological islimptE open_UNIV tendsto_def trivial_limit_within) - -lemma continuous_on_finite: - fixes S :: "'a::t1_space set" - shows "finite S \ continuous_on S f" -by (metis continuous_on_no_limpt islimpt_finite) - -lemma homeomorphic_finite: - fixes S :: "'a::t1_space set" and T :: "'b::t1_space set" - assumes "finite T" - shows "S homeomorphic T \ finite S \ finite T \ card S = card T" (is "?lhs = ?rhs") -proof - assume "S homeomorphic T" - with assms show ?rhs - apply (auto simp: homeomorphic_def homeomorphism_def) - apply (metis finite_imageI) - by (metis card_image_le finite_imageI le_antisym) -next - assume R: ?rhs - with finite_same_card_bij obtain h where "bij_betw h S T" - by auto - with R show ?lhs - apply (auto simp: homeomorphic_def homeomorphism_def continuous_on_finite) - apply (rule_tac x=h in exI) - apply (rule_tac x="inv_into S h" in exI) - apply (auto simp: bij_betw_inv_into_left bij_betw_inv_into_right bij_betw_imp_surj_on inv_into_into bij_betwE) - apply (metis bij_betw_def bij_betw_inv_into) - done -qed - -text \Relatively weak hypotheses if a set is compact.\ - -lemma homeomorphism_compact: - fixes f :: "'a::topological_space \ 'b::t2_space" - assumes "compact s" "continuous_on s f" "f ` s = t" "inj_on f s" - shows "\g. homeomorphism s t f g" -proof - - define g where "g x = (SOME y. y\s \ f y = x)" for x - have g: "\x\s. g (f x) = x" - using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto - { - fix y - assume "y \ t" - then obtain x where x:"f x = y" "x\s" - using assms(3) by auto - then have "g (f x) = x" using g by auto - then have "f (g y) = y" unfolding x(1)[symmetric] by auto - } - then have g':"\x\t. f (g x) = x" by auto - moreover - { - fix x - have "x\s \ x \ g ` t" - using g[THEN bspec[where x=x]] - unfolding image_iff - using assms(3) - by (auto intro!: bexI[where x="f x"]) - moreover - { - assume "x\g ` t" - then obtain y where y:"y\t" "g y = x" by auto - then obtain x' where x':"x'\s" "f x' = y" - using assms(3) by auto - then have "x \ s" - unfolding g_def - using someI2[of "\b. b\s \ f b = y" x' "\x. x\s"] - unfolding y(2)[symmetric] and g_def - by auto - } - ultimately have "x\s \ x \ g ` t" .. - } - then have "g ` t = s" by auto - ultimately show ?thesis - unfolding homeomorphism_def homeomorphic_def - apply (rule_tac x=g in exI) - using g and assms(3) and continuous_on_inv[OF assms(2,1), of g, unfolded assms(3)] and assms(2) - apply auto - done -qed - -lemma homeomorphic_compact: - fixes f :: "'a::topological_space \ 'b::t2_space" - shows "compact s \ continuous_on s f \ (f ` s = t) \ inj_on f s \ s homeomorphic t" - unfolding homeomorphic_def by (metis homeomorphism_compact) - -text\Preservation of topological properties.\ - -lemma homeomorphic_compactness: "s homeomorphic t \ (compact s \ compact t)" - unfolding homeomorphic_def homeomorphism_def - by (metis compact_continuous_image) - -text\Results on translation, scaling etc.\ - -lemma homeomorphic_scaling: - fixes s :: "'a::real_normed_vector set" - assumes "c \ 0" - 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 - -lemma homeomorphic_translation: - 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 - -lemma homeomorphic_affinity: - fixes s :: "'a::real_normed_vector set" - assumes "c \ 0" - 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 - 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 -qed - -lemma homeomorphic_balls: - fixes a b ::"'a::real_normed_vector" - assumes "0 < d" "0 < e" - shows "(ball a d) homeomorphic (ball b e)" (is ?th) - and "(cball a d) homeomorphic (cball b e)" (is ?cth) -proof - - show ?th 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_less_eq mult_strict_left_mono) - done - 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 -qed - -lemma homeomorphic_spheres: - fixes a b ::"'a::real_normed_vector" - assumes "0 < d" "0 < e" - shows "(sphere a d) homeomorphic (sphere b e)" -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_less_eq mult_strict_left_mono) - done - -lemma homeomorphic_ball01_UNIV: - "ball (0::'a::real_normed_vector) 1 homeomorphic (UNIV:: 'a set)" - (is "?B homeomorphic ?U") -proof - have "x \ (\z. z /\<^sub>R (1 - norm z)) ` ball 0 1" for x::'a - apply (rule_tac x="x /\<^sub>R (1 + norm x)" in image_eqI) - apply (auto simp: divide_simps) - using norm_ge_zero [of x] apply linarith+ - done - 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: divide_simps) - done - then show "(\z::'a. z /\<^sub>R (1 + norm z)) ` ?U = ?B" - by (force simp: divide_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 - 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: divide_simps) - show "\y. y /\<^sub>R (1 + norm y) /\<^sub>R (1 - norm (y /\<^sub>R (1 + norm y))) = y" - apply (auto simp: divide_simps) - apply (metis le_add_same_cancel1 norm_ge_zero not_le zero_less_one) - done -qed - -proposition homeomorphic_ball_UNIV: - fixes a ::"'a::real_normed_vector" - assumes "0 < r" shows "ball a r homeomorphic (UNIV:: 'a set)" - using assms homeomorphic_ball01_UNIV homeomorphic_balls(1) homeomorphic_trans zero_less_one by blast - - -text \Connectedness is invariant under homeomorphisms.\ +subsection%unimportant \Connectedness is invariant under homeomorphisms.\ lemma homeomorphic_connectedness: assumes "s homeomorphic t" @@ -2209,806 +643,6 @@ using assms unfolding homeomorphic_def homeomorphism_def by (metis connected_continuous_image) -subsection%unimportant\Inverse function property for open/closed maps\ - -lemma continuous_on_inverse_open_map: - assumes contf: "continuous_on S f" - and imf: "f ` S = T" - and injf: "\x. x \ S \ g (f x) = x" - and oo: "\U. openin (subtopology euclidean S) U \ openin (subtopology euclidean T) (f ` U)" - shows "continuous_on T g" -proof - - from imf injf have gTS: "g ` T = S" - by force - from imf injf have fU: "U \ S \ (f ` U) = T \ g -` U" for U - by force - show ?thesis - by (simp add: continuous_on_open [of T g] gTS) (metis openin_imp_subset fU oo) -qed - -lemma continuous_on_inverse_closed_map: - assumes contf: "continuous_on S f" - and imf: "f ` S = T" - and injf: "\x. x \ S \ g(f x) = x" - and oo: "\U. closedin (subtopology euclidean S) U \ closedin (subtopology euclidean T) (f ` U)" - shows "continuous_on T g" -proof - - from imf injf have gTS: "g ` T = S" - by force - from imf injf have fU: "U \ S \ (f ` U) = T \ g -` U" for U - by force - show ?thesis - by (simp add: continuous_on_closed [of T g] gTS) (metis closedin_imp_subset fU oo) -qed - -lemma homeomorphism_injective_open_map: - assumes contf: "continuous_on S f" - and imf: "f ` S = T" - and injf: "inj_on f S" - and oo: "\U. openin (subtopology euclidean S) U \ openin (subtopology euclidean T) (f ` U)" - obtains g where "homeomorphism S T f g" -proof - have "continuous_on T (inv_into S f)" - by (metis contf continuous_on_inverse_open_map imf injf inv_into_f_f oo) - with imf injf contf show "homeomorphism S T f (inv_into S f)" - by (auto simp: homeomorphism_def) -qed - -lemma homeomorphism_injective_closed_map: - assumes contf: "continuous_on S f" - and imf: "f ` S = T" - and injf: "inj_on f S" - and oo: "\U. closedin (subtopology euclidean S) U \ closedin (subtopology euclidean T) (f ` U)" - obtains g where "homeomorphism S T f g" -proof - have "continuous_on T (inv_into S f)" - by (metis contf continuous_on_inverse_closed_map imf injf inv_into_f_f oo) - with imf injf contf show "homeomorphism S T f (inv_into S f)" - by (auto simp: homeomorphism_def) -qed - -lemma homeomorphism_imp_open_map: - assumes hom: "homeomorphism S T f g" - and oo: "openin (subtopology euclidean S) U" - shows "openin (subtopology euclidean T) (f ` U)" -proof - - from hom oo have [simp]: "f ` U = T \ g -` U" - using openin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) - from hom have "continuous_on T g" - unfolding homeomorphism_def by blast - moreover have "g ` T = S" - by (metis hom homeomorphism_def) - ultimately show ?thesis - by (simp add: continuous_on_open oo) -qed - -lemma homeomorphism_imp_closed_map: - assumes hom: "homeomorphism S T f g" - and oo: "closedin (subtopology euclidean S) U" - shows "closedin (subtopology euclidean T) (f ` U)" -proof - - from hom oo have [simp]: "f ` U = T \ g -` U" - using closedin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) - from hom have "continuous_on T g" - unfolding homeomorphism_def by blast - moreover have "g ` T = S" - by (metis hom homeomorphism_def) - ultimately show ?thesis - by (simp add: continuous_on_closed oo) -qed - - -subsection \"Isometry" (up to constant bounds) of injective linear map etc\ - -lemma cauchy_isometric: - assumes e: "e > 0" - and s: "subspace s" - and f: "bounded_linear f" - and normf: "\x\s. norm (f x) \ e * norm x" - and xs: "\n. x n \ s" - and cf: "Cauchy (f \ x)" - shows "Cauchy x" -proof - - interpret f: bounded_linear f by fact - have "\N. \n\N. norm (x n - x N) < d" if "d > 0" for d :: real - proof - - from that obtain N where N: "\n\N. norm (f (x n) - f (x N)) < e * d" - using cf[unfolded Cauchy_def o_def dist_norm, THEN spec[where x="e*d"]] e - by auto - have "norm (x n - x N) < d" if "n \ N" for n - proof - - have "e * norm (x n - x N) \ norm (f (x n - x N))" - using subspace_diff[OF s, of "x n" "x N"] - using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]] - using normf[THEN bspec[where x="x n - x N"]] - by auto - also have "norm (f (x n - x N)) < e * d" - using \N \ n\ N unfolding f.diff[symmetric] by auto - finally show ?thesis - using \e>0\ by simp - qed - then show ?thesis by auto - qed - then show ?thesis - by (simp add: Cauchy_altdef2 dist_norm) -qed - -lemma complete_isometric_image: - assumes "0 < e" - and s: "subspace s" - and f: "bounded_linear f" - and normf: "\x\s. norm(f x) \ e * norm(x)" - and cs: "complete s" - shows "complete (f ` s)" -proof - - have "\l\f ` s. (g \ l) sequentially" - if as:"\n::nat. g n \ f ` s" and cfg:"Cauchy g" for g - proof - - from that obtain x where "\n. x n \ s \ g n = f (x n)" - using choice[of "\ n xa. xa \ s \ g n = f xa"] by auto - then have x: "\n. x n \ s" "\n. g n = f (x n)" by auto - then have "f \ x = g" by (simp add: fun_eq_iff) - then obtain l where "l\s" and l:"(x \ l) sequentially" - using cs[unfolded complete_def, THEN spec[where x=x]] - using cauchy_isometric[OF \0 < e\ s f normf] and cfg and x(1) - by auto - then show ?thesis - using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l] - by (auto simp: \f \ x = g\) - qed - then show ?thesis - unfolding complete_def by auto -qed - -proposition injective_imp_isometric: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes s: "closed s" "subspace s" - and f: "bounded_linear f" "\x\s. f x = 0 \ x = 0" - shows "\e>0. \x\s. norm (f x) \ e * norm x" -proof (cases "s \ {0::'a}") - case True - have "norm x \ norm (f x)" if "x \ s" for x - proof - - from True that have "x = 0" by auto - then show ?thesis by simp - qed - then show ?thesis - by (auto intro!: exI[where x=1]) -next - case False - interpret f: bounded_linear f by fact - from False obtain a where a: "a \ 0" "a \ s" - by auto - from False have "s \ {}" - by auto - let ?S = "{f x| x. x \ s \ norm x = norm a}" - let ?S' = "{x::'a. x\s \ norm x = norm a}" - let ?S'' = "{x::'a. norm x = norm a}" - - have "?S'' = frontier (cball 0 (norm a))" - by (simp add: sphere_def dist_norm) - then have "compact ?S''" by (metis compact_cball compact_frontier) - moreover have "?S' = s \ ?S''" by auto - ultimately have "compact ?S'" - using closed_Int_compact[of s ?S''] using s(1) by auto - moreover have *:"f ` ?S' = ?S" by auto - ultimately have "compact ?S" - using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto - then have "closed ?S" - using compact_imp_closed by auto - moreover from a have "?S \ {}" by auto - ultimately obtain b' where "b'\?S" "\y\?S. norm b' \ norm y" - using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto - then obtain b where "b\s" - and ba: "norm b = norm a" - and b: "\x\{x \ s. norm x = norm a}. norm (f b) \ norm (f x)" - unfolding *[symmetric] unfolding image_iff by auto - - let ?e = "norm (f b) / norm b" - have "norm b > 0" - using ba and a and norm_ge_zero by auto - moreover have "norm (f b) > 0" - using f(2)[THEN bspec[where x=b], OF \b\s\] - using \norm b >0\ by simp - ultimately have "0 < norm (f b) / norm b" by simp - moreover - have "norm (f b) / norm b * norm x \ norm (f x)" if "x\s" for x - proof (cases "x = 0") - case True - then show "norm (f b) / norm b * norm x \ norm (f x)" - by auto - next - case False - with \a \ 0\ have *: "0 < norm a / norm x" - unfolding zero_less_norm_iff[symmetric] by simp - have "\x\s. c *\<^sub>R x \ s" for c - using s[unfolded subspace_def] by simp - with \x \ s\ \x \ 0\ have "(norm a / norm x) *\<^sub>R x \ {x \ s. norm x = norm a}" - by simp - with \x \ 0\ \a \ 0\ show "norm (f b) / norm b * norm x \ norm (f x)" - using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]] - unfolding f.scaleR and ba - by (auto simp: mult.commute pos_le_divide_eq pos_divide_le_eq) - qed - ultimately show ?thesis by auto -qed - -proposition closed_injective_image_subspace: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "subspace s" "bounded_linear f" "\x\s. f x = 0 \ x = 0" "closed s" - shows "closed(f ` s)" -proof - - obtain e where "e > 0" and e: "\x\s. e * norm x \ norm (f x)" - using injective_imp_isometric[OF assms(4,1,2,3)] by auto - show ?thesis - using complete_isometric_image[OF \e>0\ assms(1,2) e] and assms(4) - unfolding complete_eq_closed[symmetric] by auto -qed - - -subsection%unimportant \Some properties of a canonical subspace\ - -lemma subspace_substandard: "subspace {x::'a::euclidean_space. (\i\Basis. P i \ x\i = 0)}" - by (auto simp: subspace_def inner_add_left) - -lemma closed_substandard: "closed {x::'a::euclidean_space. \i\Basis. P i \ x\i = 0}" - (is "closed ?A") -proof - - let ?D = "{i\Basis. P i}" - have "closed (\i\?D. {x::'a. x\i = 0})" - by (simp add: closed_INT closed_Collect_eq continuous_on_inner - continuous_on_const continuous_on_id) - also have "(\i\?D. {x::'a. x\i = 0}) = ?A" - by auto - finally show "closed ?A" . -qed - -lemma dim_substandard: - assumes d: "d \ Basis" - shows "dim {x::'a::euclidean_space. \i\Basis. i \ d \ x\i = 0} = card d" (is "dim ?A = _") -proof (rule dim_unique) - from d show "d \ ?A" - by (auto simp: inner_Basis) - from d show "independent d" - by (rule independent_mono [OF independent_Basis]) - have "x \ span d" if "\i\Basis. i \ d \ x \ i = 0" for x - proof - - have "finite d" - by (rule finite_subset [OF d finite_Basis]) - then have "(\i\d. (x \ i) *\<^sub>R i) \ span d" - by (simp add: span_sum span_clauses) - also have "(\i\d. (x \ i) *\<^sub>R i) = (\i\Basis. (x \ i) *\<^sub>R i)" - by (rule sum.mono_neutral_cong_left [OF finite_Basis d]) (auto simp: that) - finally show "x \ span d" - by (simp only: euclidean_representation) - qed - then show "?A \ span d" by auto -qed simp - -text \Hence closure and completeness of all subspaces.\ -lemma ex_card: - assumes "n \ card A" - shows "\S\A. card S = n" -proof (cases "finite A") - case True - from ex_bij_betw_nat_finite[OF this] obtain f where f: "bij_betw f {0..n \ card A\ have "{..< n} \ {..< card A}" "inj_on f {..< n}" - by (auto simp: bij_betw_def intro: subset_inj_on) - ultimately have "f ` {..< n} \ A" "card (f ` {..< n}) = n" - by (auto simp: bij_betw_def card_image) - then show ?thesis by blast -next - case False - with \n \ card A\ show ?thesis by force -qed - -lemma closed_subspace: - fixes s :: "'a::euclidean_space set" - assumes "subspace s" - shows "closed s" -proof - - have "dim s \ card (Basis :: 'a set)" - using dim_subset_UNIV by auto - with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \ Basis" - by auto - let ?t = "{x::'a. \i\Basis. i \ d \ x\i = 0}" - have "\f. linear f \ f ` {x::'a. \i\Basis. i \ d \ x \ i = 0} = s \ - inj_on f {x::'a. \i\Basis. i \ d \ x \ i = 0}" - using dim_substandard[of d] t d assms - by (intro subspace_isomorphism[OF subspace_substandard[of "\i. i \ d"]]) (auto simp: inner_Basis) - then obtain f where f: - "linear f" - "f ` {x. \i\Basis. i \ d \ x \ i = 0} = s" - "inj_on f {x. \i\Basis. i \ d \ x \ i = 0}" - by blast - interpret f: bounded_linear f - using f by (simp add: linear_conv_bounded_linear) - have "x \ ?t \ f x = 0 \ x = 0" for x - using f.zero d f(3)[THEN inj_onD, of x 0] by auto - moreover have "closed ?t" by (rule closed_substandard) - moreover have "subspace ?t" by (rule subspace_substandard) - ultimately show ?thesis - using closed_injective_image_subspace[of ?t f] - unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto -qed - -lemma complete_subspace: "subspace s \ complete s" - for s :: "'a::euclidean_space set" - using complete_eq_closed closed_subspace by auto - -lemma closed_span [iff]: "closed (span s)" - for s :: "'a::euclidean_space set" - by (simp add: closed_subspace subspace_span) - -lemma dim_closure [simp]: "dim (closure s) = dim s" (is "?dc = ?d") - for s :: "'a::euclidean_space set" -proof - - have "?dc \ ?d" - using closure_minimal[OF span_superset, of s] - using closed_subspace[OF subspace_span, of s] - using dim_subset[of "closure s" "span s"] - by simp - then show ?thesis - using dim_subset[OF closure_subset, of s] - by simp -qed - - -subsection%unimportant \Affine transformations of intervals\ - -lemma real_affinity_le: "0 < m \ m * x + c \ y \ x \ inverse m * y + - (c / m)" - for m :: "'a::linordered_field" - by (simp add: field_simps) - -lemma real_le_affinity: "0 < m \ y \ m * x + c \ inverse m * y + - (c / m) \ x" - for m :: "'a::linordered_field" - by (simp add: field_simps) - -lemma real_affinity_lt: "0 < m \ m * x + c < y \ x < inverse m * y + - (c / m)" - for m :: "'a::linordered_field" - by (simp add: field_simps) - -lemma real_lt_affinity: "0 < m \ y < m * x + c \ inverse m * y + - (c / m) < x" - for m :: "'a::linordered_field" - by (simp add: field_simps) - -lemma real_affinity_eq: "m \ 0 \ m * x + c = y \ x = inverse m * y + - (c / m)" - for m :: "'a::linordered_field" - by (simp add: field_simps) - -lemma real_eq_affinity: "m \ 0 \ y = m * x + c \ inverse m * y + - (c / m) = x" - for m :: "'a::linordered_field" - by (simp add: field_simps) - - -subsection \Banach fixed point theorem (not really topological ...)\ - -theorem banach_fix: - assumes s: "complete s" "s \ {}" - and c: "0 \ c" "c < 1" - and f: "f ` s \ s" - and lipschitz: "\x\s. \y\s. dist (f x) (f y) \ c * dist x y" - shows "\!x\s. f x = x" -proof - - from c have "1 - c > 0" by simp - - from s(2) obtain z0 where z0: "z0 \ s" by blast - define z where "z n = (f ^^ n) z0" for n - with f z0 have z_in_s: "z n \ s" for n :: nat - by (induct n) auto - define d where "d = dist (z 0) (z 1)" - - have fzn: "f (z n) = z (Suc n)" for n - by (simp add: z_def) - have cf_z: "dist (z n) (z (Suc n)) \ (c ^ n) * d" for n :: nat - proof (induct n) - case 0 - then show ?case - by (simp add: d_def) - next - case (Suc m) - with \0 \ c\ have "c * dist (z m) (z (Suc m)) \ c ^ Suc m * d" - using mult_left_mono[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by simp - then show ?case - using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s] - by (simp add: fzn mult_le_cancel_left) - qed - - have cf_z2: "(1 - c) * dist (z m) (z (m + n)) \ (c ^ m) * d * (1 - c ^ n)" for n m :: nat - proof (induct n) - case 0 - show ?case by simp - next - case (Suc k) - from c have "(1 - c) * dist (z m) (z (m + Suc k)) \ - (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))" - by (simp add: dist_triangle) - also from c cf_z[of "m + k"] have "\ \ (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)" - by simp - also from Suc have "\ \ c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d" - by (simp add: field_simps) - also have "\ = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)" - by (simp add: power_add field_simps) - also from c have "\ \ (c ^ m) * d * (1 - c ^ Suc k)" - by (simp add: field_simps) - finally show ?case by simp - qed - - have "\N. \m n. N \ m \ N \ n \ dist (z m) (z n) < e" if "e > 0" for e - proof (cases "d = 0") - case True - from \1 - c > 0\ have "(1 - c) * x \ 0 \ x \ 0" for x - by (metis mult_zero_left mult.commute real_mult_le_cancel_iff1) - with c cf_z2[of 0] True have "z n = z0" for n - by (simp add: z_def) - with \e > 0\ show ?thesis by simp - next - case False - with zero_le_dist[of "z 0" "z 1"] have "d > 0" - by (metis d_def less_le) - with \1 - c > 0\ \e > 0\ have "0 < e * (1 - c) / d" - by simp - with c obtain N where N: "c ^ N < e * (1 - c) / d" - using real_arch_pow_inv[of "e * (1 - c) / d" c] by auto - have *: "dist (z m) (z n) < e" if "m > n" and as: "m \ N" "n \ N" for m n :: nat - proof - - from c \n \ N\ have *: "c ^ n \ c ^ N" - using power_decreasing[OF \n\N\, of c] by simp - from c \m > n\ have "1 - c ^ (m - n) > 0" - using power_strict_mono[of c 1 "m - n"] by simp - with \d > 0\ \0 < 1 - c\ have **: "d * (1 - c ^ (m - n)) / (1 - c) > 0" - by simp - from cf_z2[of n "m - n"] \m > n\ - have "dist (z m) (z n) \ c ^ n * d * (1 - c ^ (m - n)) / (1 - c)" - by (simp add: pos_le_divide_eq[OF \1 - c > 0\] mult.commute dist_commute) - also have "\ \ c ^ N * d * (1 - c ^ (m - n)) / (1 - c)" - using mult_right_mono[OF * order_less_imp_le[OF **]] - by (simp add: mult.assoc) - also have "\ < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)" - using mult_strict_right_mono[OF N **] by (auto simp: mult.assoc) - also from c \d > 0\ \1 - c > 0\ have "\ = e * (1 - c ^ (m - n))" - by simp - also from c \1 - c ^ (m - n) > 0\ \e > 0\ have "\ \ e" - using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto - finally show ?thesis by simp - qed - have "dist (z n) (z m) < e" if "N \ m" "N \ n" for m n :: nat - proof (cases "n = m") - case True - with \e > 0\ show ?thesis by simp - next - case False - with *[of n m] *[of m n] and that show ?thesis - by (auto simp: dist_commute nat_neq_iff) - qed - then show ?thesis by auto - qed - then have "Cauchy z" - by (simp add: cauchy_def) - then obtain x where "x\s" and x:"(z \ x) sequentially" - using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto - - define e where "e = dist (f x) x" - have "e = 0" - proof (rule ccontr) - assume "e \ 0" - then have "e > 0" - unfolding e_def using zero_le_dist[of "f x" x] - by (metis dist_eq_0_iff dist_nz e_def) - then obtain N where N:"\n\N. dist (z n) x < e / 2" - using x[unfolded lim_sequentially, THEN spec[where x="e/2"]] by auto - then have N':"dist (z N) x < e / 2" by auto - have *: "c * dist (z N) x \ dist (z N) x" - unfolding mult_le_cancel_right2 - using zero_le_dist[of "z N" x] and c - by (metis dist_eq_0_iff dist_nz order_less_asym less_le) - have "dist (f (z N)) (f x) \ c * dist (z N) x" - using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]] - using z_in_s[of N] \x\s\ - using c - by auto - also have "\ < e / 2" - using N' and c using * by auto - finally show False - unfolding fzn - using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x] - unfolding e_def - by auto - qed - then have "f x = x" by (auto simp: e_def) - moreover have "y = x" if "f y = y" "y \ s" for y - proof - - from \x \ s\ \f x = x\ that have "dist x y \ c * dist x y" - using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]] by simp - with c and zero_le_dist[of x y] have "dist x y = 0" - by (simp add: mult_le_cancel_right1) - then show ?thesis by simp - qed - ultimately show ?thesis - using \x\s\ by blast -qed - -lemma banach_fix_type: - fixes f::"'a::complete_space\'a" - assumes c:"0 \ c" "c < 1" - and lipschitz:"\x. \y. dist (f x) (f y) \ c * dist x y" - shows "\!x. (f x = x)" - using assms banach_fix[OF complete_UNIV UNIV_not_empty assms(1,2) subset_UNIV, of f] - by auto - - -subsection \Edelstein fixed point theorem\ - -theorem edelstein_fix: - fixes s :: "'a::metric_space set" - assumes s: "compact s" "s \ {}" - and gs: "(g ` s) \ s" - and dist: "\x\s. \y\s. x \ y \ dist (g x) (g y) < dist x y" - shows "\!x\s. g x = x" -proof - - let ?D = "(\x. (x, x)) ` s" - have D: "compact ?D" "?D \ {}" - by (rule compact_continuous_image) - (auto intro!: s continuous_Pair continuous_ident simp: continuous_on_eq_continuous_within) - - have "\x y e. x \ s \ y \ s \ 0 < e \ dist y x < e \ dist (g y) (g x) < e" - using dist by fastforce - then have "continuous_on s g" - by (auto simp: continuous_on_iff) - then have cont: "continuous_on ?D (\x. dist ((g \ fst) x) (snd x))" - unfolding continuous_on_eq_continuous_within - by (intro continuous_dist ballI continuous_within_compose) - (auto intro!: continuous_fst continuous_snd continuous_ident simp: image_image) - - obtain a where "a \ s" and le: "\x. x \ s \ dist (g a) a \ dist (g x) x" - using continuous_attains_inf[OF D cont] by auto - - have "g a = a" - proof (rule ccontr) - assume "g a \ a" - with \a \ s\ gs have "dist (g (g a)) (g a) < dist (g a) a" - by (intro dist[rule_format]) auto - moreover have "dist (g a) a \ dist (g (g a)) (g a)" - using \a \ s\ gs by (intro le) auto - ultimately show False by auto - qed - moreover have "\x. x \ s \ g x = x \ x = a" - using dist[THEN bspec[where x=a]] \g a = a\ and \a\s\ by auto - ultimately show "\!x\s. g x = x" - using \a \ s\ by blast -qed - - -lemma cball_subset_cball_iff: - fixes a :: "'a :: euclidean_space" - shows "cball a r \ cball a' r' \ dist a a' + r \ r' \ r < 0" - (is "?lhs \ ?rhs") -proof - assume ?lhs - then show ?rhs - proof (cases "r < 0") - case True - then show ?rhs by simp - next - case False - then have [simp]: "r \ 0" by simp - have "norm (a - a') + r \ r'" - proof (cases "a = a'") - case True - then show ?thesis - using subsetD [where c = "a + r *\<^sub>R (SOME i. i \ Basis)", OF \?lhs\] subsetD [where c = a, OF \?lhs\] - by (force simp: SOME_Basis dist_norm) - next - case False - have "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = norm (a' - a - (r / norm (a - a')) *\<^sub>R (a - a'))" - by (simp add: algebra_simps) - also have "... = norm ((-1 - (r / norm (a - a'))) *\<^sub>R (a - a'))" - by (simp add: algebra_simps) - also from \a \ a'\ have "... = \- norm (a - a') - r\" - by (simp add: abs_mult_pos field_simps) - finally have [simp]: "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = \norm (a - a') + r\" - by linarith - from \a \ a'\ show ?thesis - using subsetD [where c = "a' + (1 + r / norm(a - a')) *\<^sub>R (a - a')", OF \?lhs\] - by (simp add: dist_norm scaleR_add_left) - qed - then show ?rhs - by (simp add: dist_norm) - qed -next - assume ?rhs - then show ?lhs - by (auto simp: ball_def dist_norm) - (metis add.commute add_le_cancel_right dist_norm dist_triangle3 order_trans) -qed - -lemma cball_subset_ball_iff: "cball a r \ ball a' r' \ dist a a' + r < r' \ r < 0" - (is "?lhs \ ?rhs") - for a :: "'a::euclidean_space" -proof - assume ?lhs - then show ?rhs - proof (cases "r < 0") - case True then - show ?rhs by simp - next - case False - then have [simp]: "r \ 0" by simp - have "norm (a - a') + r < r'" - proof (cases "a = a'") - case True - then show ?thesis - using subsetD [where c = "a + r *\<^sub>R (SOME i. i \ Basis)", OF \?lhs\] subsetD [where c = a, OF \?lhs\] - by (force simp: SOME_Basis dist_norm) - next - case False - have False if "norm (a - a') + r \ r'" - proof - - from that have "\r' - norm (a - a')\ \ r" - by (simp split: abs_split) - (metis \0 \ r\ \?lhs\ centre_in_cball dist_commute dist_norm less_asym mem_ball subset_eq) - then show ?thesis - using subsetD [where c = "a + (r' / norm(a - a') - 1) *\<^sub>R (a - a')", OF \?lhs\] \a \ a'\ - by (simp add: dist_norm field_simps) - (simp add: diff_divide_distrib scaleR_left_diff_distrib) - qed - then show ?thesis by force - qed - then show ?rhs by (simp add: dist_norm) - qed -next - assume ?rhs - then show ?lhs - by (auto simp: ball_def dist_norm) - (metis add.commute add_le_cancel_right dist_norm dist_triangle3 le_less_trans) -qed - -lemma ball_subset_cball_iff: "ball a r \ cball a' r' \ dist a a' + r \ r' \ r \ 0" - (is "?lhs = ?rhs") - for a :: "'a::euclidean_space" -proof (cases "r \ 0") - case True - then show ?thesis - using dist_not_less_zero less_le_trans by force -next - case False - show ?thesis - proof - assume ?lhs - then have "(cball a r \ cball a' r')" - by (metis False closed_cball closure_ball closure_closed closure_mono not_less) - with False show ?rhs - by (fastforce iff: cball_subset_cball_iff) - next - assume ?rhs - with False show ?lhs - using ball_subset_cball cball_subset_cball_iff by blast - qed -qed - -lemma ball_subset_ball_iff: - fixes a :: "'a :: euclidean_space" - shows "ball a r \ ball a' r' \ dist a a' + r \ r' \ r \ 0" - (is "?lhs = ?rhs") -proof (cases "r \ 0") - case True then show ?thesis - using dist_not_less_zero less_le_trans by force -next - case False show ?thesis - proof - assume ?lhs - then have "0 < r'" - by (metis (no_types) False \?lhs\ centre_in_ball dist_norm le_less_trans mem_ball norm_ge_zero not_less set_mp) - then have "(cball a r \ cball a' r')" - by (metis False\?lhs\ closure_ball closure_mono not_less) - then show ?rhs - using False cball_subset_cball_iff by fastforce - next - assume ?rhs then show ?lhs - apply (auto simp: ball_def) - apply (metis add.commute add_le_cancel_right dist_commute dist_triangle_lt not_le order_trans) - using dist_not_less_zero order.strict_trans2 apply blast - done - qed -qed - - -lemma ball_eq_ball_iff: - fixes x :: "'a :: euclidean_space" - shows "ball x d = ball y e \ d \ 0 \ e \ 0 \ x=y \ d=e" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - proof (cases "d \ 0 \ e \ 0") - case True - with \?lhs\ show ?rhs - by safe (simp_all only: ball_eq_empty [of y e, symmetric] ball_eq_empty [of x d, symmetric]) - next - case False - with \?lhs\ show ?rhs - apply (auto simp: set_eq_subset ball_subset_ball_iff dist_norm norm_minus_commute algebra_simps) - apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) - apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) - done - qed -next - assume ?rhs then show ?lhs - by (auto simp: set_eq_subset ball_subset_ball_iff) -qed - -lemma cball_eq_cball_iff: - fixes x :: "'a :: euclidean_space" - shows "cball x d = cball y e \ d < 0 \ e < 0 \ x=y \ d=e" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - proof (cases "d < 0 \ e < 0") - case True - with \?lhs\ show ?rhs - by safe (simp_all only: cball_eq_empty [of y e, symmetric] cball_eq_empty [of x d, symmetric]) - next - case False - with \?lhs\ show ?rhs - apply (auto simp: set_eq_subset cball_subset_cball_iff dist_norm norm_minus_commute algebra_simps) - apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) - apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) - done - qed -next - assume ?rhs then show ?lhs - by (auto simp: set_eq_subset cball_subset_cball_iff) -qed - -lemma ball_eq_cball_iff: - fixes x :: "'a :: euclidean_space" - shows "ball x d = cball y e \ d \ 0 \ e < 0" (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - apply (auto simp: set_eq_subset ball_subset_cball_iff cball_subset_ball_iff algebra_simps) - apply (metis add_increasing2 add_le_cancel_right add_less_same_cancel1 dist_not_less_zero less_le_trans zero_le_dist) - apply (metis add_less_same_cancel1 dist_not_less_zero less_le_trans not_le) - using \?lhs\ ball_eq_empty cball_eq_empty apply blast+ - done -next - assume ?rhs then show ?lhs by auto -qed - -lemma cball_eq_ball_iff: - fixes x :: "'a :: euclidean_space" - shows "cball x d = ball y e \ d < 0 \ e \ 0" - using ball_eq_cball_iff by blast - -lemma finite_ball_avoid: - fixes S :: "'a :: euclidean_space set" - assumes "open S" "finite X" "p \ S" - shows "\e>0. \w\ball p e. w\S \ (w\p \ w\X)" -proof - - obtain e1 where "0 < e1" and e1_b:"ball p e1 \ S" - using open_contains_ball_eq[OF \open S\] assms by auto - obtain e2 where "0 < e2" and "\x\X. x \ p \ e2 \ dist p x" - using finite_set_avoid[OF \finite X\,of p] by auto - hence "\w\ball p (min e1 e2). w\S \ (w\p \ w\X)" using e1_b by auto - thus "\e>0. \w\ball p e. w \ S \ (w \ p \ w \ X)" using \e2>0\ \e1>0\ - apply (rule_tac x="min e1 e2" in exI) - by auto -qed - -lemma finite_cball_avoid: - fixes S :: "'a :: euclidean_space set" - assumes "open S" "finite X" "p \ S" - shows "\e>0. \w\cball p e. w\S \ (w\p \ w\X)" -proof - - obtain e1 where "e1>0" and e1: "\w\ball p e1. w\S \ (w\p \ w\X)" - using finite_ball_avoid[OF assms] by auto - define e2 where "e2 \ e1/2" - have "e2>0" and "e2 < e1" unfolding e2_def using \e1>0\ by auto - then have "cball p e2 \ ball p e1" by (subst cball_subset_ball_iff,auto) - then show "\e>0. \w\cball p e. w \ S \ (w \ p \ w \ X)" using \e2>0\ e1 by auto -qed - subsection\Various separability-type properties\ lemma univ_second_countable: diff -r d692ef26021e -r 1331e57b54c6 src/HOL/Analysis/Elementary_Metric_Spaces.thy --- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy Mon Jan 07 10:22:22 2019 +0100 +++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy Mon Jan 07 11:29:34 2019 +0100 @@ -44,7 +44,7 @@ qed -subsection \Combination of Elementary and Abstract Topology\ +subsection \Combination of Elementary and Abstract Topology (TODO: this might be a separate theory?)\ lemma closedin_limpt: "closedin (subtopology euclidean T) S \ S \ T \ (\x. x islimpt S \ x \ T \ x \ S)" @@ -257,6 +257,145 @@ by metis +subsubsection%unimportant \Continuity relative to a union.\ + +lemma continuous_on_Un_local: + "\closedin (subtopology euclidean (s \ t)) s; closedin (subtopology euclidean (s \ t)) t; + continuous_on s f; continuous_on t f\ + \ continuous_on (s \ t) f" + unfolding continuous_on closedin_limpt + by (metis Lim_trivial_limit Lim_within_union Un_iff trivial_limit_within) + +lemma continuous_on_cases_local: + "\closedin (subtopology euclidean (s \ t)) s; closedin (subtopology euclidean (s \ t)) t; + continuous_on s f; continuous_on t g; + \x. \x \ s \ \P x \ x \ t \ P x\ \ f x = g x\ + \ continuous_on (s \ t) (\x. if P x then f x else g x)" + by (rule continuous_on_Un_local) (auto intro: continuous_on_eq) + +lemma continuous_on_cases_le: + fixes h :: "'a :: topological_space \ real" + assumes "continuous_on {t \ s. h t \ a} f" + and "continuous_on {t \ s. a \ h t} g" + and h: "continuous_on s h" + and "\t. \t \ s; h t = a\ \ f t = g t" + shows "continuous_on s (\t. if h t \ a then f(t) else g(t))" +proof - + have s: "s = (s \ h -` atMost a) \ (s \ h -` atLeast a)" + by force + have 1: "closedin (subtopology euclidean s) (s \ h -` atMost a)" + by (rule continuous_closedin_preimage [OF h closed_atMost]) + have 2: "closedin (subtopology euclidean s) (s \ h -` atLeast a)" + by (rule continuous_closedin_preimage [OF h closed_atLeast]) + have eq: "s \ h -` {..a} = {t \ s. h t \ a}" "s \ h -` {a..} = {t \ s. a \ h t}" + by auto + show ?thesis + apply (rule continuous_on_subset [of s, OF _ order_refl]) + apply (subst s) + apply (rule continuous_on_cases_local) + using 1 2 s assms apply (auto simp: eq) + done +qed + +lemma continuous_on_cases_1: + fixes s :: "real set" + assumes "continuous_on {t \ s. t \ a} f" + and "continuous_on {t \ s. a \ t} g" + and "a \ s \ f a = g a" + shows "continuous_on s (\t. if t \ a then f(t) else g(t))" +using assms +by (auto simp: continuous_on_id intro: continuous_on_cases_le [where h = id, simplified]) + + +subsubsection%unimportant\Inverse function property for open/closed maps\ + +lemma continuous_on_inverse_open_map: + assumes contf: "continuous_on S f" + and imf: "f ` S = T" + and injf: "\x. x \ S \ g (f x) = x" + and oo: "\U. openin (subtopology euclidean S) U \ openin (subtopology euclidean T) (f ` U)" + shows "continuous_on T g" +proof - + from imf injf have gTS: "g ` T = S" + by force + from imf injf have fU: "U \ S \ (f ` U) = T \ g -` U" for U + by force + show ?thesis + by (simp add: continuous_on_open [of T g] gTS) (metis openin_imp_subset fU oo) +qed + +lemma continuous_on_inverse_closed_map: + assumes contf: "continuous_on S f" + and imf: "f ` S = T" + and injf: "\x. x \ S \ g(f x) = x" + and oo: "\U. closedin (subtopology euclidean S) U \ closedin (subtopology euclidean T) (f ` U)" + shows "continuous_on T g" +proof - + from imf injf have gTS: "g ` T = S" + by force + from imf injf have fU: "U \ S \ (f ` U) = T \ g -` U" for U + by force + show ?thesis + by (simp add: continuous_on_closed [of T g] gTS) (metis closedin_imp_subset fU oo) +qed + +lemma homeomorphism_injective_open_map: + assumes contf: "continuous_on S f" + and imf: "f ` S = T" + and injf: "inj_on f S" + and oo: "\U. openin (subtopology euclidean S) U \ openin (subtopology euclidean T) (f ` U)" + obtains g where "homeomorphism S T f g" +proof + have "continuous_on T (inv_into S f)" + by (metis contf continuous_on_inverse_open_map imf injf inv_into_f_f oo) + with imf injf contf show "homeomorphism S T f (inv_into S f)" + by (auto simp: homeomorphism_def) +qed + +lemma homeomorphism_injective_closed_map: + assumes contf: "continuous_on S f" + and imf: "f ` S = T" + and injf: "inj_on f S" + and oo: "\U. closedin (subtopology euclidean S) U \ closedin (subtopology euclidean T) (f ` U)" + obtains g where "homeomorphism S T f g" +proof + have "continuous_on T (inv_into S f)" + by (metis contf continuous_on_inverse_closed_map imf injf inv_into_f_f oo) + with imf injf contf show "homeomorphism S T f (inv_into S f)" + by (auto simp: homeomorphism_def) +qed + +lemma homeomorphism_imp_open_map: + assumes hom: "homeomorphism S T f g" + and oo: "openin (subtopology euclidean S) U" + shows "openin (subtopology euclidean T) (f ` U)" +proof - + from hom oo have [simp]: "f ` U = T \ g -` U" + using openin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) + from hom have "continuous_on T g" + unfolding homeomorphism_def by blast + moreover have "g ` T = S" + by (metis hom homeomorphism_def) + ultimately show ?thesis + by (simp add: continuous_on_open oo) +qed + +lemma homeomorphism_imp_closed_map: + assumes hom: "homeomorphism S T f g" + and oo: "closedin (subtopology euclidean S) U" + shows "closedin (subtopology euclidean T) (f ` U)" +proof - + from hom oo have [simp]: "f ` U = T \ g -` U" + using closedin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) + from hom have "continuous_on T g" + unfolding homeomorphism_def by blast + moreover have "g ` T = S" + by (metis hom homeomorphism_def) + ultimately show ?thesis + by (simp add: continuous_on_closed oo) +qed + + subsection \Open and closed balls\ definition%important ball :: "'a::metric_space \ real \ 'a set" @@ -734,6 +873,149 @@ using assms by (fast intro: tendsto_le tendsto_intros) +subsection \Continuity\ + +text\Derive the epsilon-delta forms, which we often use as "definitions"\ + +proposition continuous_within_eps_delta: + "continuous (at x within s) f \ (\e>0. \d>0. \x'\ s. dist x' x < d --> dist (f x') (f x) < e)" + unfolding continuous_within and Lim_within by fastforce + +corollary continuous_at_eps_delta: + "continuous (at x) f \ (\e > 0. \d > 0. \x'. dist x' x < d \ dist (f x') (f x) < e)" + using continuous_within_eps_delta [of x UNIV f] by simp + +lemma continuous_at_right_real_increasing: + fixes f :: "real \ real" + assumes nondecF: "\x y. x \ y \ f x \ f y" + shows "continuous (at_right a) f \ (\e>0. \d>0. f (a + d) - f a < e)" + apply (simp add: greaterThan_def dist_real_def continuous_within Lim_within_le) + apply (intro all_cong ex_cong, safe) + apply (erule_tac x="a + d" in allE, simp) + apply (simp add: nondecF field_simps) + apply (drule nondecF, simp) + done + +lemma continuous_at_left_real_increasing: + assumes nondecF: "\ x y. x \ y \ f x \ ((f y) :: real)" + shows "(continuous (at_left (a :: real)) f) = (\e > 0. \delta > 0. f a - f (a - delta) < e)" + apply (simp add: lessThan_def dist_real_def continuous_within Lim_within_le) + apply (intro all_cong ex_cong, safe) + apply (erule_tac x="a - d" in allE, simp) + apply (simp add: nondecF field_simps) + apply (cut_tac x="a - d" and y=x in nondecF, simp_all) + done + +text\Versions in terms of open balls.\ + +lemma continuous_within_ball: + "continuous (at x within s) f \ + (\e > 0. \d > 0. f ` (ball x d \ s) \ ball (f x) e)" + (is "?lhs = ?rhs") +proof + assume ?lhs + { + fix e :: real + assume "e > 0" + then obtain d where d: "d>0" "\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" + using \?lhs\[unfolded continuous_within Lim_within] by auto + { + fix y + assume "y \ f ` (ball x d \ s)" + then have "y \ ball (f x) e" + using d(2) + apply (auto simp: dist_commute) + apply (erule_tac x=xa in ballE, auto) + using \e > 0\ + apply auto + done + } + then have "\d>0. f ` (ball x d \ s) \ ball (f x) e" + using \d > 0\ + unfolding subset_eq ball_def by (auto simp: dist_commute) + } + then show ?rhs by auto +next + assume ?rhs + then show ?lhs + unfolding continuous_within Lim_within ball_def subset_eq + apply (auto simp: dist_commute) + apply (erule_tac x=e in allE, auto) + done +qed + +lemma continuous_at_ball: + "continuous (at x) f \ (\e>0. \d>0. f ` (ball x d) \ ball (f x) e)" (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball + apply auto + apply (erule_tac x=e in allE, auto) + apply (rule_tac x=d in exI, auto) + apply (erule_tac x=xa in allE) + apply (auto simp: dist_commute) + done +next + assume ?rhs + then show ?lhs + unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball + apply auto + apply (erule_tac x=e in allE, auto) + apply (rule_tac x=d in exI, auto) + apply (erule_tac x="f xa" in allE) + apply (auto simp: dist_commute) + done +qed + +text\Define setwise continuity in terms of limits within the set.\ + +lemma continuous_on_iff: + "continuous_on s f \ + (\x\s. \e>0. \d>0. \x'\s. dist x' x < d \ dist (f x') (f x) < e)" + unfolding continuous_on_def Lim_within + by (metis dist_pos_lt dist_self) + +lemma continuous_within_E: + assumes "continuous (at x within s) f" "e>0" + obtains d where "d>0" "\x'. \x'\ s; dist x' x \ d\ \ dist (f x') (f x) < e" + using assms apply (simp add: continuous_within_eps_delta) + apply (drule spec [of _ e], clarify) + apply (rule_tac d="d/2" in that, auto) + done + +lemma continuous_onI [intro?]: + assumes "\x e. \e > 0; x \ s\ \ \d>0. \x'\s. dist x' x < d \ dist (f x') (f x) \ e" + shows "continuous_on s f" +apply (simp add: continuous_on_iff, clarify) +apply (rule ex_forward [OF assms [OF half_gt_zero]], auto) +done + +text\Some simple consequential lemmas.\ + +lemma continuous_onE: + assumes "continuous_on s f" "x\s" "e>0" + obtains d where "d>0" "\x'. \x' \ s; dist x' x \ d\ \ dist (f x') (f x) < e" + using assms + apply (simp add: continuous_on_iff) + apply (elim ballE allE) + apply (auto intro: that [where d="d/2" for d]) + done + +text\The usual transformation theorems.\ + +lemma continuous_transform_within: + fixes f g :: "'a::metric_space \ 'b::topological_space" + assumes "continuous (at x within s) f" + and "0 < d" + and "x \ s" + and "\x'. \x' \ s; dist x' x < d\ \ f x' = g x'" + shows "continuous (at x within s) g" + using assms + unfolding continuous_within + by (force intro: Lim_transform_within) + + subsection \Closure and Limit Characterization\ lemma closure_approachable: @@ -814,6 +1096,7 @@ qed qed + subsection \Boundedness\ (* FIXME: This has to be unified with BSEQ!! *) @@ -945,67 +1228,16 @@ by (auto intro!: boundedI) qed - -subsection \Consequences for Real Numbers\ - -lemma closed_contains_Inf: - fixes S :: "real set" - shows "S \ {} \ bdd_below S \ closed S \ Inf S \ S" - by (metis closure_contains_Inf closure_closed) - -lemma closed_subset_contains_Inf: - fixes A C :: "real set" - shows "closed C \ A \ C \ A \ {} \ bdd_below A \ Inf A \ C" - by (metis closure_contains_Inf closure_minimal subset_eq) - -lemma atLeastAtMost_subset_contains_Inf: - fixes A :: "real set" and a b :: real - shows "A \ {} \ a \ b \ A \ {a..b} \ Inf A \ {a..b}" - by (rule closed_subset_contains_Inf) - (auto intro: closed_real_atLeastAtMost intro!: bdd_belowI[of A a]) - -lemma bounded_real: "bounded (S::real set) \ (\a. \x\S. \x\ \ a)" - by (simp add: bounded_iff) - -lemma bounded_imp_bdd_above: "bounded S \ bdd_above (S :: real set)" - by (auto simp: bounded_def bdd_above_def dist_real_def) - (metis abs_le_D1 abs_minus_commute diff_le_eq) - -lemma bounded_imp_bdd_below: "bounded S \ bdd_below (S :: real set)" - by (auto simp: bounded_def bdd_below_def dist_real_def) - (metis abs_le_D1 add.commute diff_le_eq) - -lemma bounded_has_Sup: - fixes S :: "real set" - assumes "bounded S" - and "S \ {}" - shows "\x\S. x \ Sup S" - and "\b. (\x\S. x \ b) \ Sup S \ b" -proof - show "\b. (\x\S. x \ b) \ Sup S \ b" - using assms by (metis cSup_least) -qed (metis cSup_upper assms(1) bounded_imp_bdd_above) - -lemma Sup_insert: - fixes S :: "real set" - shows "bounded S \ Sup (insert x S) = (if S = {} then x else max x (Sup S))" - by (auto simp: bounded_imp_bdd_above sup_max cSup_insert_If) - -lemma bounded_has_Inf: - fixes S :: "real set" - assumes "bounded S" - and "S \ {}" - shows "\x\S. x \ Inf S" - and "\b. (\x\S. x \ b) \ Inf S \ b" -proof - show "\b. (\x\S. x \ b) \ Inf S \ b" - using assms by (metis cInf_greatest) -qed (metis cInf_lower assms(1) bounded_imp_bdd_below) - -lemma Inf_insert: - fixes S :: "real set" - shows "bounded S \ Inf (insert x S) = (if S = {} then x else min x (Inf S))" - by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If) +lemma bounded_Times: + assumes "bounded s" "bounded t" + shows "bounded (s \ t)" +proof - + obtain x y a b where "\z\s. dist x z \ a" "\z\t. dist y z \ b" + using assms [unfolded bounded_def] by auto + then have "\z\s \ t. dist (x, y) z \ sqrt (a\<^sup>2 + b\<^sup>2)" + by (auto simp: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono) + then show ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto +qed subsection \Compactness\ @@ -1029,6 +1261,23 @@ shows "T \ ball x r \ {}" using assms centre_in_ball closure_iff_nhds_not_empty by blast +lemma compact_sup_maxdistance: + fixes s :: "'a::metric_space set" + assumes "compact s" + and "s \ {}" + shows "\x\s. \y\s. \u\s. \v\s. dist u v \ dist x y" +proof - + have "compact (s \ s)" + using \compact s\ by (intro compact_Times) + moreover have "s \ s \ {}" + using \s \ {}\ by auto + moreover have "continuous_on (s \ s) (\x. dist (fst x) (snd x))" + by (intro continuous_at_imp_continuous_on ballI continuous_intros) + ultimately show ?thesis + using continuous_attains_sup[of "s \ s" "\x. dist (fst x) (snd x)"] by auto +qed + + subsubsection\Totally bounded\ lemma cauchy_def: "Cauchy s \ (\e>0. \N. \m n. m \ N \ n \ N \ dist (s m) (s n) < e)" @@ -1141,6 +1390,403 @@ using compact_imp_bounded unfolding compact_eq_Bolzano_Weierstrass . +subsection \Banach fixed point theorem\ + +theorem banach_fix:\ \TODO: rename to \Banach_fix\\ + assumes s: "complete s" "s \ {}" + and c: "0 \ c" "c < 1" + and f: "f ` s \ s" + and lipschitz: "\x\s. \y\s. dist (f x) (f y) \ c * dist x y" + shows "\!x\s. f x = x" +proof - + from c have "1 - c > 0" by simp + + from s(2) obtain z0 where z0: "z0 \ s" by blast + define z where "z n = (f ^^ n) z0" for n + with f z0 have z_in_s: "z n \ s" for n :: nat + by (induct n) auto + define d where "d = dist (z 0) (z 1)" + + have fzn: "f (z n) = z (Suc n)" for n + by (simp add: z_def) + have cf_z: "dist (z n) (z (Suc n)) \ (c ^ n) * d" for n :: nat + proof (induct n) + case 0 + then show ?case + by (simp add: d_def) + next + case (Suc m) + with \0 \ c\ have "c * dist (z m) (z (Suc m)) \ c ^ Suc m * d" + using mult_left_mono[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by simp + then show ?case + using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s] + by (simp add: fzn mult_le_cancel_left) + qed + + have cf_z2: "(1 - c) * dist (z m) (z (m + n)) \ (c ^ m) * d * (1 - c ^ n)" for n m :: nat + proof (induct n) + case 0 + show ?case by simp + next + case (Suc k) + from c have "(1 - c) * dist (z m) (z (m + Suc k)) \ + (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))" + by (simp add: dist_triangle) + also from c cf_z[of "m + k"] have "\ \ (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)" + by simp + also from Suc have "\ \ c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d" + by (simp add: field_simps) + also have "\ = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)" + by (simp add: power_add field_simps) + also from c have "\ \ (c ^ m) * d * (1 - c ^ Suc k)" + by (simp add: field_simps) + finally show ?case by simp + qed + + have "\N. \m n. N \ m \ N \ n \ dist (z m) (z n) < e" if "e > 0" for e + proof (cases "d = 0") + case True + from \1 - c > 0\ have "(1 - c) * x \ 0 \ x \ 0" for x + by (metis mult_zero_left mult.commute real_mult_le_cancel_iff1) + with c cf_z2[of 0] True have "z n = z0" for n + by (simp add: z_def) + with \e > 0\ show ?thesis by simp + next + case False + with zero_le_dist[of "z 0" "z 1"] have "d > 0" + by (metis d_def less_le) + with \1 - c > 0\ \e > 0\ have "0 < e * (1 - c) / d" + by simp + with c obtain N where N: "c ^ N < e * (1 - c) / d" + using real_arch_pow_inv[of "e * (1 - c) / d" c] by auto + have *: "dist (z m) (z n) < e" if "m > n" and as: "m \ N" "n \ N" for m n :: nat + proof - + from c \n \ N\ have *: "c ^ n \ c ^ N" + using power_decreasing[OF \n\N\, of c] by simp + from c \m > n\ have "1 - c ^ (m - n) > 0" + using power_strict_mono[of c 1 "m - n"] by simp + with \d > 0\ \0 < 1 - c\ have **: "d * (1 - c ^ (m - n)) / (1 - c) > 0" + by simp + from cf_z2[of n "m - n"] \m > n\ + have "dist (z m) (z n) \ c ^ n * d * (1 - c ^ (m - n)) / (1 - c)" + by (simp add: pos_le_divide_eq[OF \1 - c > 0\] mult.commute dist_commute) + also have "\ \ c ^ N * d * (1 - c ^ (m - n)) / (1 - c)" + using mult_right_mono[OF * order_less_imp_le[OF **]] + by (simp add: mult.assoc) + also have "\ < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)" + using mult_strict_right_mono[OF N **] by (auto simp: mult.assoc) + also from c \d > 0\ \1 - c > 0\ have "\ = e * (1 - c ^ (m - n))" + by simp + also from c \1 - c ^ (m - n) > 0\ \e > 0\ have "\ \ e" + using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto + finally show ?thesis by simp + qed + have "dist (z n) (z m) < e" if "N \ m" "N \ n" for m n :: nat + proof (cases "n = m") + case True + with \e > 0\ show ?thesis by simp + next + case False + with *[of n m] *[of m n] and that show ?thesis + by (auto simp: dist_commute nat_neq_iff) + qed + then show ?thesis by auto + qed + then have "Cauchy z" + by (simp add: cauchy_def) + then obtain x where "x\s" and x:"(z \ x) sequentially" + using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto + + define e where "e = dist (f x) x" + have "e = 0" + proof (rule ccontr) + assume "e \ 0" + then have "e > 0" + unfolding e_def using zero_le_dist[of "f x" x] + by (metis dist_eq_0_iff dist_nz e_def) + then obtain N where N:"\n\N. dist (z n) x < e / 2" + using x[unfolded lim_sequentially, THEN spec[where x="e/2"]] by auto + then have N':"dist (z N) x < e / 2" by auto + have *: "c * dist (z N) x \ dist (z N) x" + unfolding mult_le_cancel_right2 + using zero_le_dist[of "z N" x] and c + by (metis dist_eq_0_iff dist_nz order_less_asym less_le) + have "dist (f (z N)) (f x) \ c * dist (z N) x" + using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]] + using z_in_s[of N] \x\s\ + using c + by auto + also have "\ < e / 2" + using N' and c using * by auto + finally show False + unfolding fzn + using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x] + unfolding e_def + by auto + qed + then have "f x = x" by (auto simp: e_def) + moreover have "y = x" if "f y = y" "y \ s" for y + proof - + from \x \ s\ \f x = x\ that have "dist x y \ c * dist x y" + using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]] by simp + with c and zero_le_dist[of x y] have "dist x y = 0" + by (simp add: mult_le_cancel_right1) + then show ?thesis by simp + qed + ultimately show ?thesis + using \x\s\ by blast +qed + + +subsection \Edelstein fixed point theorem\ + +theorem edelstein_fix:\ \TODO: rename to \Edelstein_fix\\ + fixes s :: "'a::metric_space set" + assumes s: "compact s" "s \ {}" + and gs: "(g ` s) \ s" + and dist: "\x\s. \y\s. x \ y \ dist (g x) (g y) < dist x y" + shows "\!x\s. g x = x" +proof - + let ?D = "(\x. (x, x)) ` s" + have D: "compact ?D" "?D \ {}" + by (rule compact_continuous_image) + (auto intro!: s continuous_Pair continuous_ident simp: continuous_on_eq_continuous_within) + + have "\x y e. x \ s \ y \ s \ 0 < e \ dist y x < e \ dist (g y) (g x) < e" + using dist by fastforce + then have "continuous_on s g" + by (auto simp: continuous_on_iff) + then have cont: "continuous_on ?D (\x. dist ((g \ fst) x) (snd x))" + unfolding continuous_on_eq_continuous_within + by (intro continuous_dist ballI continuous_within_compose) + (auto intro!: continuous_fst continuous_snd continuous_ident simp: image_image) + + obtain a where "a \ s" and le: "\x. x \ s \ dist (g a) a \ dist (g x) x" + using continuous_attains_inf[OF D cont] by auto + + have "g a = a" + proof (rule ccontr) + assume "g a \ a" + with \a \ s\ gs have "dist (g (g a)) (g a) < dist (g a) a" + by (intro dist[rule_format]) auto + moreover have "dist (g a) a \ dist (g (g a)) (g a)" + using \a \ s\ gs by (intro le) auto + ultimately show False by auto + qed + moreover have "\x. x \ s \ g x = x \ x = a" + using dist[THEN bspec[where x=a]] \g a = a\ and \a\s\ by auto + ultimately show "\!x\s. g x = x" + using \a \ s\ by blast +qed + +subsection \The diameter of a set\ + +definition%important diameter :: "'a::metric_space set \ real" where + "diameter S = (if S = {} then 0 else SUP (x,y)\S\S. dist x y)" + +lemma diameter_empty [simp]: "diameter{} = 0" + by (auto simp: diameter_def) + +lemma diameter_singleton [simp]: "diameter{x} = 0" + by (auto simp: diameter_def) + +lemma diameter_le: + assumes "S \ {} \ 0 \ d" + and no: "\x y. \x \ S; y \ S\ \ norm(x - y) \ d" + shows "diameter S \ d" +using assms + by (auto simp: dist_norm diameter_def intro: cSUP_least) + +lemma diameter_bounded_bound: + fixes s :: "'a :: metric_space set" + assumes s: "bounded s" "x \ s" "y \ s" + shows "dist x y \ diameter s" +proof - + from s obtain z d where z: "\x. x \ s \ dist z x \ d" + unfolding bounded_def by auto + have "bdd_above (case_prod dist ` (s\s))" + proof (intro bdd_aboveI, safe) + fix a b + assume "a \ s" "b \ s" + with z[of a] z[of b] dist_triangle[of a b z] + show "dist a b \ 2 * d" + by (simp add: dist_commute) + qed + moreover have "(x,y) \ s\s" using s by auto + ultimately have "dist x y \ (SUP (x,y)\s\s. dist x y)" + by (rule cSUP_upper2) simp + with \x \ s\ show ?thesis + by (auto simp: diameter_def) +qed + +lemma diameter_lower_bounded: + fixes s :: "'a :: metric_space set" + assumes s: "bounded s" + and d: "0 < d" "d < diameter s" + shows "\x\s. \y\s. d < dist x y" +proof (rule ccontr) + assume contr: "\ ?thesis" + moreover have "s \ {}" + using d by (auto simp: diameter_def) + ultimately have "diameter s \ d" + by (auto simp: not_less diameter_def intro!: cSUP_least) + with \d < diameter s\ show False by auto +qed + +lemma diameter_bounded: + assumes "bounded s" + shows "\x\s. \y\s. dist x y \ diameter s" + and "\d>0. d < diameter s \ (\x\s. \y\s. dist x y > d)" + using diameter_bounded_bound[of s] diameter_lower_bounded[of s] assms + by auto + +lemma bounded_two_points: + "bounded S \ (\e. \x\S. \y\S. dist x y \ e)" + apply (rule iffI) + subgoal using diameter_bounded(1) by auto + subgoal using bounded_any_center[of S] by meson + done + +lemma diameter_compact_attained: + assumes "compact s" + and "s \ {}" + shows "\x\s. \y\s. dist x y = diameter s" +proof - + have b: "bounded s" using assms(1) + by (rule compact_imp_bounded) + then obtain x y where xys: "x\s" "y\s" + and xy: "\u\s. \v\s. dist u v \ dist x y" + using compact_sup_maxdistance[OF assms] by auto + then have "diameter s \ dist x y" + unfolding diameter_def + apply clarsimp + apply (rule cSUP_least, fast+) + done + then show ?thesis + by (metis b diameter_bounded_bound order_antisym xys) +qed + +lemma diameter_ge_0: + assumes "bounded S" shows "0 \ diameter S" + by (metis all_not_in_conv assms diameter_bounded_bound diameter_empty dist_self order_refl) + +lemma diameter_subset: + assumes "S \ T" "bounded T" + shows "diameter S \ diameter T" +proof (cases "S = {} \ T = {}") + case True + with assms show ?thesis + by (force simp: diameter_ge_0) +next + case False + then have "bdd_above ((\x. case x of (x, xa) \ dist x xa) ` (T \ T))" + using \bounded T\ diameter_bounded_bound by (force simp: bdd_above_def) + with False \S \ T\ show ?thesis + apply (simp add: diameter_def) + apply (rule cSUP_subset_mono, auto) + done +qed + +lemma diameter_closure: + assumes "bounded S" + shows "diameter(closure S) = diameter S" +proof (rule order_antisym) + have "False" if "diameter S < diameter (closure S)" + proof - + define d where "d = diameter(closure S) - diameter(S)" + have "d > 0" + using that by (simp add: d_def) + then have "diameter(closure(S)) - d / 2 < diameter(closure(S))" + by simp + have dd: "diameter (closure S) - d / 2 = (diameter(closure(S)) + diameter(S)) / 2" + by (simp add: d_def divide_simps) + have bocl: "bounded (closure S)" + using assms by blast + moreover have "0 \ diameter S" + using assms diameter_ge_0 by blast + ultimately obtain x y where "x \ closure S" "y \ closure S" and xy: "diameter(closure(S)) - d / 2 < dist x y" + using diameter_bounded(2) [OF bocl, rule_format, of "diameter(closure(S)) - d / 2"] \d > 0\ d_def by auto + then obtain x' y' where x'y': "x' \ S" "dist x' x < d/4" "y' \ S" "dist y' y < d/4" + using closure_approachable + by (metis \0 < d\ zero_less_divide_iff zero_less_numeral) + then have "dist x' y' \ diameter S" + using assms diameter_bounded_bound by blast + with x'y' have "dist x y \ d / 4 + diameter S + d / 4" + by (meson add_mono_thms_linordered_semiring(1) dist_triangle dist_triangle3 less_eq_real_def order_trans) + then show ?thesis + using xy d_def by linarith + qed + then show "diameter (closure S) \ diameter S" + by fastforce + next + show "diameter S \ diameter (closure S)" + by (simp add: assms bounded_closure closure_subset diameter_subset) +qed + +proposition Lebesgue_number_lemma: + assumes "compact S" "\ \ {}" "S \ \\" and ope: "\B. B \ \ \ open B" + obtains \ where "0 < \" "\T. \T \ S; diameter T < \\ \ \B \ \. T \ B" +proof (cases "S = {}") + case True + then show ?thesis + by (metis \\ \ {}\ zero_less_one empty_subsetI equals0I subset_trans that) +next + case False + { fix x assume "x \ S" + then obtain C where C: "x \ C" "C \ \" + using \S \ \\\ by blast + then obtain r where r: "r>0" "ball x (2*r) \ C" + by (metis mult.commute mult_2_right not_le ope openE field_sum_of_halves zero_le_numeral zero_less_mult_iff) + then have "\r C. r > 0 \ ball x (2*r) \ C \ C \ \" + using C by blast + } + then obtain r where r: "\x. x \ S \ r x > 0 \ (\C \ \. ball x (2*r x) \ C)" + by metis + then have "S \ (\x \ S. ball x (r x))" + by auto + then obtain \ where "finite \" "S \ \\" and \: "\ \ (\x. ball x (r x)) ` S" + by (rule compactE [OF \compact S\]) auto + then obtain S0 where "S0 \ S" "finite S0" and S0: "\ = (\x. ball x (r x)) ` S0" + by (meson finite_subset_image) + then have "S0 \ {}" + using False \S \ \\\ by auto + define \ where "\ = Inf (r ` S0)" + have "\ > 0" + using \finite S0\ \S0 \ S\ \S0 \ {}\ r by (auto simp: \_def finite_less_Inf_iff) + show ?thesis + proof + show "0 < \" + by (simp add: \0 < \\) + show "\B \ \. T \ B" if "T \ S" and dia: "diameter T < \" for T + proof (cases "T = {}") + case True + then show ?thesis + using \\ \ {}\ by blast + next + case False + then obtain y where "y \ T" by blast + then have "y \ S" + using \T \ S\ by auto + then obtain x where "x \ S0" and x: "y \ ball x (r x)" + using \S \ \\\ S0 that by blast + have "ball y \ \ ball y (r x)" + by (metis \_def \S0 \ {}\ \finite S0\ \x \ S0\ empty_is_image finite_imageI finite_less_Inf_iff imageI less_irrefl not_le subset_ball) + also have "... \ ball x (2*r x)" + by clarsimp (metis dist_commute dist_triangle_less_add mem_ball mult_2 x) + finally obtain C where "C \ \" "ball y \ \ C" + by (meson r \S0 \ S\ \x \ S0\ dual_order.trans subsetCE) + have "bounded T" + using \compact S\ bounded_subset compact_imp_bounded \T \ S\ by blast + then have "T \ ball y \" + using \y \ T\ dia diameter_bounded_bound by fastforce + then show ?thesis + apply (rule_tac x=C in bexI) + using \ball y \ \ C\ \C \ \\ by auto + qed + qed +qed + + subsection \Metric spaces with the Heine-Borel property\ text \ @@ -1308,7 +1954,7 @@ qed -subsubsection \Completeness\ +subsection \Completeness\ proposition (in metric_space) completeI: assumes "\f. \n. f n \ s \ Cauchy f \ \l\s. f \ l" @@ -1566,8 +2212,22 @@ using frontier_subset_closed compact_eq_bounded_closed by blast - -subsubsection \Properties of Balls and Spheres\ +lemma continuous_closed_imp_Cauchy_continuous: + fixes S :: "('a::complete_space) set" + shows "\continuous_on S f; closed S; Cauchy \; \n. (\ n) \ S\ \ Cauchy(f \ \)" + apply (simp add: complete_eq_closed [symmetric] continuous_on_sequentially) + by (meson LIMSEQ_imp_Cauchy complete_def) + +lemma banach_fix_type: + fixes f::"'a::complete_space\'a" + assumes c:"0 \ c" "c < 1" + and lipschitz:"\x. \y. dist (f x) (f y) \ c * dist x y" + shows "\!x. (f x = x)" + using assms banach_fix[OF complete_UNIV UNIV_not_empty assms(1,2) subset_UNIV, of f] + by auto + + +subsection \Properties of Balls and Spheres\ lemma compact_cball[simp]: fixes x :: "'a::heine_borel" @@ -1589,7 +2249,7 @@ by blast -subsubsection \Distance from a Set\ +subsection \Distance from a Set\ lemma distance_attains_sup: assumes "compact s" "s \ {}" @@ -1625,6 +2285,7 @@ with that show ?thesis by fastforce qed + subsection \Infimum Distance\ definition%important "infdist x A = (if A = {} then 0 else INF a\A. dist x a)" @@ -1857,134 +2518,90 @@ qed -subsection \Continuity\ - -text\Derive the epsilon-delta forms, which we often use as "definitions"\ - -proposition continuous_within_eps_delta: - "continuous (at x within s) f \ (\e>0. \d>0. \x'\ s. dist x' x < d --> dist (f x') (f x) < e)" - unfolding continuous_within and Lim_within by fastforce - -corollary continuous_at_eps_delta: - "continuous (at x) f \ (\e > 0. \d > 0. \x'. dist x' x < d \ dist (f x') (f x) < e)" - using continuous_within_eps_delta [of x UNIV f] by simp - -lemma continuous_at_right_real_increasing: - fixes f :: "real \ real" - assumes nondecF: "\x y. x \ y \ f x \ f y" - shows "continuous (at_right a) f \ (\e>0. \d>0. f (a + d) - f a < e)" - apply (simp add: greaterThan_def dist_real_def continuous_within Lim_within_le) - apply (intro all_cong ex_cong, safe) - apply (erule_tac x="a + d" in allE, simp) - apply (simp add: nondecF field_simps) - apply (drule nondecF, simp) - done - -lemma continuous_at_left_real_increasing: - assumes nondecF: "\ x y. x \ y \ f x \ ((f y) :: real)" - shows "(continuous (at_left (a :: real)) f) = (\e > 0. \delta > 0. f a - f (a - delta) < e)" - apply (simp add: lessThan_def dist_real_def continuous_within Lim_within_le) - apply (intro all_cong ex_cong, safe) - apply (erule_tac x="a - d" in allE, simp) - apply (simp add: nondecF field_simps) - apply (cut_tac x="a - d" and y=x in nondecF, simp_all) - done - -text\Versions in terms of open balls.\ - -lemma continuous_within_ball: - "continuous (at x within s) f \ - (\e > 0. \d > 0. f ` (ball x d \ s) \ ball (f x) e)" - (is "?lhs = ?rhs") -proof - assume ?lhs - { - fix e :: real - assume "e > 0" - then obtain d where d: "d>0" "\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" - using \?lhs\[unfolded continuous_within Lim_within] by auto - { - fix y - assume "y \ f ` (ball x d \ s)" - then have "y \ ball (f x) e" - using d(2) - apply (auto simp: dist_commute) - apply (erule_tac x=xa in ballE, auto) - using \e > 0\ - apply auto - done - } - then have "\d>0. f ` (ball x d \ s) \ ball (f x) e" - using \d > 0\ - unfolding subset_eq ball_def by (auto simp: dist_commute) - } - then show ?rhs by auto +subsection \Separation between Points and Sets\ + +proposition separate_point_closed: + fixes s :: "'a::heine_borel set" + assumes "closed s" and "a \ s" + shows "\d>0. \x\s. d \ dist a x" +proof (cases "s = {}") + case True + then show ?thesis by(auto intro!: exI[where x=1]) next - assume ?rhs - then show ?lhs - unfolding continuous_within Lim_within ball_def subset_eq - apply (auto simp: dist_commute) - apply (erule_tac x=e in allE, auto) - done -qed - -lemma continuous_at_ball: - "continuous (at x) f \ (\e>0. \d>0. f ` (ball x d) \ ball (f x) e)" (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball - apply auto - apply (erule_tac x=e in allE, auto) - apply (rule_tac x=d in exI, auto) - apply (erule_tac x=xa in allE) - apply (auto simp: dist_commute) - done -next - assume ?rhs - then show ?lhs - unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball - apply auto - apply (erule_tac x=e in allE, auto) - apply (rule_tac x=d in exI, auto) - apply (erule_tac x="f xa" in allE) - apply (auto simp: dist_commute) - done + case False + from assms obtain x where "x\s" "\y\s. dist a x \ dist a y" + using \s \ {}\ by (blast intro: distance_attains_inf [of s a]) + with \x\s\ show ?thesis using dist_pos_lt[of a x] and\a \ s\ + by blast qed -text\Define setwise continuity in terms of limits within the set.\ - -lemma continuous_on_iff: - "continuous_on s f \ - (\x\s. \e>0. \d>0. \x'\s. dist x' x < d \ dist (f x') (f x) < e)" - unfolding continuous_on_def Lim_within - by (metis dist_pos_lt dist_self) - -lemma continuous_within_E: - assumes "continuous (at x within s) f" "e>0" - obtains d where "d>0" "\x'. \x'\ s; dist x' x \ d\ \ dist (f x') (f x) < e" - using assms apply (simp add: continuous_within_eps_delta) - apply (drule spec [of _ e], clarify) - apply (rule_tac d="d/2" in that, auto) - done - -lemma continuous_onI [intro?]: - assumes "\x e. \e > 0; x \ s\ \ \d>0. \x'\s. dist x' x < d \ dist (f x') (f x) \ e" - shows "continuous_on s f" -apply (simp add: continuous_on_iff, clarify) -apply (rule ex_forward [OF assms [OF half_gt_zero]], auto) -done - -text\Some simple consequential lemmas.\ - -lemma continuous_onE: - assumes "continuous_on s f" "x\s" "e>0" - obtains d where "d>0" "\x'. \x' \ s; dist x' x \ d\ \ dist (f x') (f x) < e" - using assms - apply (simp add: continuous_on_iff) - apply (elim ballE allE) - apply (auto intro: that [where d="d/2" for d]) - done +proposition separate_compact_closed: + fixes s t :: "'a::heine_borel set" + assumes "compact s" + and t: "closed t" "s \ t = {}" + shows "\d>0. \x\s. \y\t. d \ dist x y" +proof cases + assume "s \ {} \ t \ {}" + then have "s \ {}" "t \ {}" by auto + let ?inf = "\x. infdist x t" + have "continuous_on s ?inf" + by (auto intro!: continuous_at_imp_continuous_on continuous_infdist continuous_ident) + then obtain x where x: "x \ s" "\y\s. ?inf x \ ?inf y" + using continuous_attains_inf[OF \compact s\ \s \ {}\] by auto + then have "0 < ?inf x" + using t \t \ {}\ in_closed_iff_infdist_zero by (auto simp: less_le infdist_nonneg) + moreover have "\x'\s. \y\t. ?inf x \ dist x' y" + using x by (auto intro: order_trans infdist_le) + ultimately show ?thesis by auto +qed (auto intro!: exI[of _ 1]) + +proposition separate_closed_compact: + fixes s t :: "'a::heine_borel set" + assumes "closed s" + and "compact t" + and "s \ t = {}" + shows "\d>0. \x\s. \y\t. d \ dist x y" +proof - + have *: "t \ s = {}" + using assms(3) by auto + show ?thesis + using separate_compact_closed[OF assms(2,1) *] by (force simp: dist_commute) +qed + +proposition compact_in_open_separated: + fixes A::"'a::heine_borel set" + assumes "A \ {}" + assumes "compact A" + assumes "open B" + assumes "A \ B" + obtains e where "e > 0" "{x. infdist x A \ e} \ B" +proof atomize_elim + have "closed (- B)" "compact A" "- B \ A = {}" + using assms by (auto simp: open_Diff compact_eq_bounded_closed) + from separate_closed_compact[OF this] + obtain d'::real where d': "d'>0" "\x y. x \ B \ y \ A \ d' \ dist x y" + by auto + define d where "d = d' / 2" + hence "d>0" "d < d'" using d' by auto + with d' have d: "\x y. x \ B \ y \ A \ d < dist x y" + by force + show "\e>0. {x. infdist x A \ e} \ B" + proof (rule ccontr) + assume "\e. 0 < e \ {x. infdist x A \ e} \ B" + with \d > 0\ obtain x where x: "infdist x A \ d" "x \ B" + by auto + from assms have "closed A" "A \ {}" by (auto simp: compact_eq_bounded_closed) + from infdist_attains_inf[OF this] + obtain y where y: "y \ A" "infdist x A = dist x y" + by auto + have "dist x y \ d" using x y by simp + also have "\ < dist x y" using y d x by auto + finally show False by simp + qed +qed + + +subsection \Uniform Continuity\ lemma uniformly_continuous_onE: assumes "uniformly_continuous_on s f" "0 < e" @@ -2069,34 +2686,99 @@ unfolding uniformly_continuous_on_def by blast qed -lemma continuous_closed_imp_Cauchy_continuous: - fixes S :: "('a::complete_space) set" - shows "\continuous_on S f; closed S; Cauchy \; \n. (\ n) \ S\ \ Cauchy(f \ \)" - apply (simp add: complete_eq_closed [symmetric] continuous_on_sequentially) - by (meson LIMSEQ_imp_Cauchy complete_def) - -text\The usual transformation theorems.\ - -lemma continuous_transform_within: - fixes f g :: "'a::metric_space \ 'b::topological_space" - assumes "continuous (at x within s) f" - and "0 < d" - and "x \ s" - and "\x'. \x' \ s; dist x' x < d\ \ f x' = g x'" - shows "continuous (at x within s) g" - using assms - unfolding continuous_within - by (force intro: Lim_transform_within) - -subsubsection%unimportant \Structural rules for uniform continuity\ - -lemma (in bounded_linear) uniformly_continuous_on[continuous_intros]: - fixes g :: "_::metric_space \ _" - assumes "uniformly_continuous_on s g" - shows "uniformly_continuous_on s (\x. f (g x))" - using assms unfolding uniformly_continuous_on_sequentially - unfolding dist_norm tendsto_norm_zero_iff diff[symmetric] - by (auto intro: tendsto_zero) + +subsection \Continuity on a Compact Domain Implies Uniform Continuity\ + +text\From the proof of the Heine-Borel theorem: Lemma 2 in section 3.7, page 69 of +J. C. Burkill and H. Burkill. A Second Course in Mathematical Analysis (CUP, 2002)\ + +lemma Heine_Borel_lemma: + assumes "compact S" and Ssub: "S \ \\" and opn: "\G. G \ \ \ open G" + obtains e where "0 < e" "\x. x \ S \ \G \ \. ball x e \ G" +proof - + have False if neg: "\e. 0 < e \ \x \ S. \G \ \. \ ball x e \ G" + proof - + have "\x \ S. \G \ \. \ ball x (1 / Suc n) \ G" for n + using neg by simp + then obtain f where "\n. f n \ S" and fG: "\G n. G \ \ \ \ ball (f n) (1 / Suc n) \ G" + by metis + then obtain l r where "l \ S" "strict_mono r" and to_l: "(f \ r) \ l" + using \compact S\ compact_def that by metis + then obtain G where "l \ G" "G \ \" + using Ssub by auto + then obtain e where "0 < e" and e: "\z. dist z l < e \ z \ G" + using opn open_dist by blast + obtain N1 where N1: "\n. n \ N1 \ dist (f (r n)) l < e/2" + using to_l apply (simp add: lim_sequentially) + using \0 < e\ half_gt_zero that by blast + obtain N2 where N2: "of_nat N2 > 2/e" + using reals_Archimedean2 by blast + obtain x where "x \ ball (f (r (max N1 N2))) (1 / real (Suc (r (max N1 N2))))" and "x \ G" + using fG [OF \G \ \\, of "r (max N1 N2)"] by blast + then have "dist (f (r (max N1 N2))) x < 1 / real (Suc (r (max N1 N2)))" + by simp + also have "... \ 1 / real (Suc (max N1 N2))" + apply (simp add: divide_simps del: max.bounded_iff) + using \strict_mono r\ seq_suble by blast + also have "... \ 1 / real (Suc N2)" + by (simp add: field_simps) + also have "... < e/2" + using N2 \0 < e\ by (simp add: field_simps) + finally have "dist (f (r (max N1 N2))) x < e / 2" . + moreover have "dist (f (r (max N1 N2))) l < e/2" + using N1 max.cobounded1 by blast + ultimately have "dist x l < e" + using dist_triangle_half_r by blast + then show ?thesis + using e \x \ G\ by blast + qed + then show ?thesis + by (meson that) +qed + +lemma compact_uniformly_equicontinuous: + assumes "compact S" + and cont: "\x e. \x \ S; 0 < e\ + \ \d. 0 < d \ + (\f \ \. \x' \ S. dist x' x < d \ dist (f x') (f x) < e)" + and "0 < e" + obtains d where "0 < d" + "\f x x'. \f \ \; x \ S; x' \ S; dist x' x < d\ \ dist (f x') (f x) < e" +proof - + obtain d where d_pos: "\x e. \x \ S; 0 < e\ \ 0 < d x e" + and d_dist : "\x x' e f. \dist x' x < d x e; x \ S; x' \ S; 0 < e; f \ \\ \ dist (f x') (f x) < e" + using cont by metis + let ?\ = "((\x. ball x (d x (e / 2))) ` S)" + have Ssub: "S \ \ ?\" + by clarsimp (metis d_pos \0 < e\ dist_self half_gt_zero_iff) + then obtain k where "0 < k" and k: "\x. x \ S \ \G \ ?\. ball x k \ G" + by (rule Heine_Borel_lemma [OF \compact S\]) auto + moreover have "dist (f v) (f u) < e" if "f \ \" "u \ S" "v \ S" "dist v u < k" for f u v + proof - + obtain G where "G \ ?\" "u \ G" "v \ G" + using k that + by (metis \dist v u < k\ \u \ S\ \0 < k\ centre_in_ball subsetD dist_commute mem_ball) + then obtain w where w: "dist w u < d w (e / 2)" "dist w v < d w (e / 2)" "w \ S" + by auto + with that d_dist have "dist (f w) (f v) < e/2" + by (metis \0 < e\ dist_commute half_gt_zero) + moreover + have "dist (f w) (f u) < e/2" + using that d_dist w by (metis \0 < e\ dist_commute divide_pos_pos zero_less_numeral) + ultimately show ?thesis + using dist_triangle_half_r by blast + qed + ultimately show ?thesis using that by blast +qed + +corollary compact_uniformly_continuous: + fixes f :: "'a :: metric_space \ 'b :: metric_space" + assumes f: "continuous_on S f" and S: "compact S" + shows "uniformly_continuous_on S f" + using f + unfolding continuous_on_iff uniformly_continuous_on_def + by (force intro: compact_uniformly_equicontinuous [OF S, of "{f}"]) + subsection%unimportant\ Theorems relating continuity and uniform continuity to closures\ @@ -2389,8 +3071,6 @@ apply (rule_tac x="e/2" in exI, force+) done -subsection \With abstract Topology\ - lemma Times_in_interior_subtopology: fixes U :: "('a::metric_space \ 'b::metric_space) set" assumes "(x, y) \ U" "openin (subtopology euclidean (S \ T)) U" @@ -2729,5 +3409,112 @@ using continuous_at_avoid[of x f a] assms(4) by auto +subsection \Consequences for Real Numbers\ + +lemma closed_contains_Inf: + fixes S :: "real set" + shows "S \ {} \ bdd_below S \ closed S \ Inf S \ S" + by (metis closure_contains_Inf closure_closed) + +lemma closed_subset_contains_Inf: + fixes A C :: "real set" + shows "closed C \ A \ C \ A \ {} \ bdd_below A \ Inf A \ C" + by (metis closure_contains_Inf closure_minimal subset_eq) + +lemma atLeastAtMost_subset_contains_Inf: + fixes A :: "real set" and a b :: real + shows "A \ {} \ a \ b \ A \ {a..b} \ Inf A \ {a..b}" + by (rule closed_subset_contains_Inf) + (auto intro: closed_real_atLeastAtMost intro!: bdd_belowI[of A a]) + +lemma bounded_real: "bounded (S::real set) \ (\a. \x\S. \x\ \ a)" + by (simp add: bounded_iff) + +lemma bounded_imp_bdd_above: "bounded S \ bdd_above (S :: real set)" + by (auto simp: bounded_def bdd_above_def dist_real_def) + (metis abs_le_D1 abs_minus_commute diff_le_eq) + +lemma bounded_imp_bdd_below: "bounded S \ bdd_below (S :: real set)" + by (auto simp: bounded_def bdd_below_def dist_real_def) + (metis abs_le_D1 add.commute diff_le_eq) + +lemma bounded_has_Sup: + fixes S :: "real set" + assumes "bounded S" + and "S \ {}" + shows "\x\S. x \ Sup S" + and "\b. (\x\S. x \ b) \ Sup S \ b" +proof + show "\b. (\x\S. x \ b) \ Sup S \ b" + using assms by (metis cSup_least) +qed (metis cSup_upper assms(1) bounded_imp_bdd_above) + +lemma Sup_insert: + fixes S :: "real set" + shows "bounded S \ Sup (insert x S) = (if S = {} then x else max x (Sup S))" + by (auto simp: bounded_imp_bdd_above sup_max cSup_insert_If) + +lemma bounded_has_Inf: + fixes S :: "real set" + assumes "bounded S" + and "S \ {}" + shows "\x\S. x \ Inf S" + and "\b. (\x\S. x \ b) \ Inf S \ b" +proof + show "\b. (\x\S. x \ b) \ Inf S \ b" + using assms by (metis cInf_greatest) +qed (metis cInf_lower assms(1) bounded_imp_bdd_below) + +lemma Inf_insert: + fixes S :: "real set" + shows "bounded S \ Inf (insert x S) = (if S = {} then x else min x (Inf S))" + by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If) + +lemma open_real: + fixes s :: "real set" + shows "open s \ (\x \ s. \e>0. \x'. \x' - x\ < e --> x' \ s)" + unfolding open_dist dist_norm by simp + +lemma islimpt_approachable_real: + fixes s :: "real set" + shows "x islimpt s \ (\e>0. \x'\ s. x' \ x \ \x' - x\ < e)" + unfolding islimpt_approachable dist_norm by simp + +lemma closed_real: + fixes s :: "real set" + shows "closed s \ (\x. (\e>0. \x' \ s. x' \ x \ \x' - x\ < e) \ x \ s)" + unfolding closed_limpt islimpt_approachable dist_norm by simp + +lemma continuous_at_real_range: + fixes f :: "'a::real_normed_vector \ real" + shows "continuous (at x) f \ (\e>0. \d>0. \x'. norm(x' - x) < d --> \f x' - f x\ < e)" + unfolding continuous_at + unfolding Lim_at + unfolding dist_norm + apply auto + apply (erule_tac x=e in allE, auto) + apply (rule_tac x=d in exI, auto) + apply (erule_tac x=x' in allE, auto) + apply (erule_tac x=e in allE, auto) + done + +lemma continuous_on_real_range: + fixes f :: "'a::real_normed_vector \ real" + shows "continuous_on s f \ + (\x \ s. \e>0. \d>0. (\x' \ s. norm(x' - x) < d \ \f x' - f x\ < e))" + unfolding continuous_on_iff dist_norm by simp + +lemma continuous_on_closed_Collect_le: + fixes f g :: "'a::t2_space \ real" + assumes f: "continuous_on s f" and g: "continuous_on s g" and s: "closed s" + shows "closed {x \ s. f x \ g x}" +proof - + have "closed ((\x. g x - f x) -` {0..} \ s)" + using closed_real_atLeast continuous_on_diff [OF g f] + by (simp add: continuous_on_closed_vimage [OF s]) + also have "((\x. g x - f x) -` {0..} \ s) = {x\s. f x \ g x}" + by auto + finally show ?thesis . +qed end \ No newline at end of file diff -r d692ef26021e -r 1331e57b54c6 src/HOL/Analysis/Elementary_Normed_Spaces.thy --- a/src/HOL/Analysis/Elementary_Normed_Spaces.thy Mon Jan 07 10:22:22 2019 +0100 +++ b/src/HOL/Analysis/Elementary_Normed_Spaces.thy Mon Jan 07 11:29:34 2019 +0100 @@ -1022,6 +1022,14 @@ subsubsection%unimportant \Structural rules for uniform continuity\ +lemma (in bounded_linear) uniformly_continuous_on[continuous_intros]: + fixes g :: "_::metric_space \ _" + assumes "uniformly_continuous_on s g" + shows "uniformly_continuous_on s (\x. f (g x))" + using assms unfolding uniformly_continuous_on_sequentially + unfolding dist_norm tendsto_norm_zero_iff diff[symmetric] + by (auto intro: tendsto_zero) + lemma uniformly_continuous_on_dist[continuous_intros]: fixes f g :: "'a::metric_space \ 'b::metric_space" assumes "uniformly_continuous_on s f" @@ -1119,4 +1127,513 @@ "bounded_linear f \ continuous_on s f" using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto +subsection%unimportant \Arithmetic Preserves Topological Properties\ + +lemma open_scaling[intro]: + fixes s :: "'a::real_normed_vector set" + assumes "c \ 0" + and "open s" + shows "open((\x. c *\<^sub>R x) ` s)" +proof - + { + fix x + assume "x \ s" + then obtain e where "e>0" + and e:"\x'. dist x' x < e \ x' \ s" using assms(2)[unfolded open_dist, THEN bspec[where x=x]] + by auto + have "e * \c\ > 0" + using assms(1)[unfolded zero_less_abs_iff[symmetric]] \e>0\ by auto + moreover + { + fix y + assume "dist y (c *\<^sub>R x) < e * \c\" + 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) + 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 + } + 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 + } + then show ?thesis unfolding open_dist by auto +qed + +lemma minus_image_eq_vimage: + fixes A :: "'a::ab_group_add set" + shows "(\x. - x) ` A = (\x. - x) -` A" + by (auto intro!: image_eqI [where f="\x. - x"]) + +lemma open_negations: + fixes S :: "'a::real_normed_vector set" + shows "open S \ open ((\x. - x) ` S)" + using open_scaling [of "- 1" S] by simp + +lemma open_translation: + fixes S :: "'a::real_normed_vector set" + assumes "open S" + shows "open((\x. a + x) ` S)" +proof - + { + fix x + have "continuous (at x) (\x. x - a)" + by (intro continuous_diff continuous_ident continuous_const) + } + moreover have "{x. x - a \ S} = (+) a ` S" + by force + ultimately show ?thesis + by (metis assms continuous_open_vimage vimage_def) +qed + +lemma open_neg_translation: + fixes s :: "'a::real_normed_vector set" + assumes "open s" + shows "open((\x. a - x) ` s)" + using open_translation[OF open_negations[OF assms], of a] + by (auto simp: image_image) + +lemma open_affinity: + fixes S :: "'a::real_normed_vector set" + assumes "open S" "c \ 0" + shows "open ((\x. a + c *\<^sub>R x) ` S)" +proof - + have *: "(\x. a + c *\<^sub>R x) = (\x. a + x) \ (\x. c *\<^sub>R x)" + unfolding o_def .. + have "(+) a ` (*\<^sub>R) c ` S = ((+) a \ (*\<^sub>R) c) ` S" + by auto + then show ?thesis + using assms open_translation[of "(*\<^sub>R) c ` S" a] + unfolding * + by auto +qed + +lemma interior_translation: + fixes S :: "'a::real_normed_vector set" + shows "interior ((\x. a + x) ` S) = (\x. a + x) ` (interior S)" +proof (rule set_eqI, rule) + fix x + assume "x \ interior ((+) a ` S)" + then obtain e where "e > 0" and e: "ball x e \ (+) a ` S" + unfolding mem_interior by auto + then have "ball (x - a) e \ S" + unfolding subset_eq Ball_def mem_ball dist_norm + 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 +next + fix x + assume "x \ (+) a ` interior S" + then obtain y e where "e > 0" and e: "ball y e \ S" and y: "x = a + y" + unfolding image_iff Bex_def mem_interior by auto + { + fix z + have *: "a + y - z = y + a - z" by auto + assume "z \ ball x e" + then have "z - a \ S" + using e[unfolded subset_eq, THEN bspec[where x="z - a"]] + unfolding mem_ball dist_norm y group_add_class.diff_diff_eq2 * + by auto + then have "z \ (+) a ` S" + unfolding image_iff by (auto intro!: bexI[where x="z - a"]) + } + then have "ball x e \ (+) a ` S" + unfolding subset_eq by auto + then show "x \ interior ((+) a ` S)" + unfolding mem_interior using \e > 0\ by auto +qed + +lemma compact_scaling: + fixes s :: "'a::real_normed_vector set" + assumes "compact s" + shows "compact ((\x. c *\<^sub>R x) ` s)" +proof - + let ?f = "\x. scaleR c x" + have *: "bounded_linear ?f" by (rule bounded_linear_scaleR_right) + show ?thesis + using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f] + using linear_continuous_at[OF *] assms + by auto +qed + +lemma compact_negations: + fixes s :: "'a::real_normed_vector set" + assumes "compact s" + shows "compact ((\x. - x) ` s)" + using compact_scaling [OF assms, of "- 1"] by auto + +lemma compact_sums: + fixes s t :: "'a::real_normed_vector set" + assumes "compact s" + and "compact t" + 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 + have "continuous_on (s \ t) (\z. fst z + snd z)" + unfolding continuous_on by (rule ballI) (intro tendsto_intros) + then show ?thesis + unfolding * using compact_continuous_image compact_Times [OF assms] by auto +qed + +lemma compact_differences: + fixes s t :: "'a::real_normed_vector set" + assumes "compact s" + 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 + then show ?thesis + using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto +qed + +lemma compact_translation: + fixes s :: "'a::real_normed_vector set" + assumes "compact s" + shows "compact ((\x. a + x) ` s)" +proof - + have "{x + y |x y. x \ s \ y \ {a}} = (\x. a + x) ` s" + by auto + then show ?thesis + using compact_sums[OF assms compact_sing[of a]] by auto +qed + +lemma compact_affinity: + fixes s :: "'a::real_normed_vector set" + assumes "compact s" + shows "compact ((\x. a + c *\<^sub>R x) ` s)" +proof - + have "(+) a ` (*\<^sub>R) c ` s = (\x. a + c *\<^sub>R x) ` s" + by auto + then show ?thesis + using compact_translation[OF compact_scaling[OF assms], of a c] by auto +qed + +lemma closed_scaling: + fixes S :: "'a::real_normed_vector set" + assumes "closed S" + shows "closed ((\x. c *\<^sub>R x) ` S)" +proof (cases "c = 0") + case True then show ?thesis + by (auto simp: image_constant_conv) +next + case False + from assms have "closed ((\x. inverse c *\<^sub>R x) -` S)" + by (simp add: continuous_closed_vimage) + also have "(\x. inverse c *\<^sub>R x) -` S = (\x. c *\<^sub>R x) ` S" + using \c \ 0\ by (auto elim: image_eqI [rotated]) + finally show ?thesis . +qed + +lemma closed_negations: + fixes S :: "'a::real_normed_vector set" + assumes "closed S" + shows "closed ((\x. -x) ` S)" + using closed_scaling[OF assms, of "- 1"] by simp + +lemma compact_closed_sums: + fixes S :: "'a::real_normed_vector set" + assumes "compact S" and "closed T" + shows "closed (\x\ S. \y \ T. {x + y})" +proof - + let ?S = "{x + y |x y. x \ S \ y \ T}" + { + fix x l + assume as: "\n. x n \ ?S" "(x \ l) sequentially" + from as(1) obtain f where f: "\n. x n = fst (f n) + snd (f n)" "\n. fst (f n) \ S" "\n. snd (f n) \ T" + using choice[of "\n y. x n = (fst y) + (snd y) \ fst y \ S \ snd y \ T"] by auto + obtain l' r where "l'\S" and r: "strict_mono r" and lr: "(((\n. fst (f n)) \ r) \ l') sequentially" + using assms(1)[unfolded compact_def, THEN spec[where x="\ n. fst (f n)"]] using f(2) by auto + have "((\n. snd (f (r n))) \ l - l') sequentially" + using tendsto_diff[OF LIMSEQ_subseq_LIMSEQ[OF as(2) r] lr] and f(1) + unfolding o_def + by auto + then have "l - l' \ T" + using assms(2)[unfolded closed_sequential_limits, + THEN spec[where x="\ n. snd (f (r n))"], + THEN spec[where x="l - l'"]] + 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 + } + moreover have "?S = (\x\ S. \y \ T. {x + y})" + by force + ultimately show ?thesis + unfolding closed_sequential_limits + by (metis (no_types, lifting)) +qed + +lemma closed_compact_sums: + fixes S T :: "'a::real_normed_vector set" + assumes "closed S" "compact T" + shows "closed (\x\ S. \y \ T. {x + y})" +proof - + have "(\x\ T. \y \ S. {x + y}) = (\x\ S. \y \ T. {x + y})" + by auto + then show ?thesis + using compact_closed_sums[OF assms(2,1)] by simp +qed + +lemma compact_closed_differences: + fixes S T :: "'a::real_normed_vector set" + assumes "compact S" "closed T" + shows "closed (\x\ S. \y \ T. {x - y})" +proof - + 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 +qed + +lemma closed_compact_differences: + fixes S T :: "'a::real_normed_vector set" + assumes "closed S" "compact T" + shows "closed (\x\ S. \y \ T. {x - y})" +proof - + have "(\x\ S. \y \ uminus ` T. {x + y}) = {x - y |x y. x \ S \ y \ T}" + by auto + then show ?thesis + using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp +qed + +lemma closed_translation: + fixes a :: "'a::real_normed_vector" + assumes "closed S" + shows "closed ((\x. a + x) ` S)" +proof - + have "(\x\ {a}. \y \ S. {x + y}) = ((+) a ` S)" by auto + then show ?thesis + using compact_closed_sums[OF compact_sing[of a] assms] by auto +qed + +lemma closure_translation: + fixes a :: "'a::real_normed_vector" + shows "closure ((\x. a + x) ` s) = (\x. a + x) ` (closure s)" +proof - + have *: "(+) a ` (- s) = - (+) a ` s" + by (auto intro!: image_eqI[where x="x - a" for x]) + show ?thesis + unfolding closure_interior translation_Compl + using interior_translation[of a "- s"] + unfolding * + by auto +qed + +lemma frontier_translation: + fixes a :: "'a::real_normed_vector" + shows "frontier((\x. a + x) ` s) = (\x. a + x) ` (frontier s)" + unfolding frontier_def translation_diff interior_translation closure_translation + by auto + +lemma sphere_translation: + fixes a :: "'n::real_normed_vector" + shows "sphere (a+c) r = (+) a ` sphere c r" + by (auto simp: dist_norm algebra_simps intro!: image_eqI[where x="x - a" for x]) + +lemma cball_translation: + fixes a :: "'n::real_normed_vector" + shows "cball (a+c) r = (+) a ` cball c r" + by (auto simp: dist_norm algebra_simps intro!: image_eqI[where x="x - a" for x]) + +lemma ball_translation: + fixes a :: "'n::real_normed_vector" + shows "ball (a+c) r = (+) a ` ball c r" + by (auto simp: dist_norm algebra_simps intro!: image_eqI[where x="x - a" for x]) + + +subsection%unimportant\Homeomorphisms\ + +lemma homeomorphic_scaling: + fixes s :: "'a::real_normed_vector set" + assumes "c \ 0" + 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 + +lemma homeomorphic_translation: + 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 + +lemma homeomorphic_affinity: + fixes s :: "'a::real_normed_vector set" + assumes "c \ 0" + 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 + 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 +qed + +lemma homeomorphic_balls: + fixes a b ::"'a::real_normed_vector" + assumes "0 < d" "0 < e" + shows "(ball a d) homeomorphic (ball b e)" (is ?th) + and "(cball a d) homeomorphic (cball b e)" (is ?cth) +proof - + show ?th 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_less_eq mult_strict_left_mono) + done + 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 +qed + +lemma homeomorphic_spheres: + fixes a b ::"'a::real_normed_vector" + assumes "0 < d" "0 < e" + shows "(sphere a d) homeomorphic (sphere b e)" +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_less_eq mult_strict_left_mono) + done + +lemma homeomorphic_ball01_UNIV: + "ball (0::'a::real_normed_vector) 1 homeomorphic (UNIV:: 'a set)" + (is "?B homeomorphic ?U") +proof + have "x \ (\z. z /\<^sub>R (1 - norm z)) ` ball 0 1" for x::'a + apply (rule_tac x="x /\<^sub>R (1 + norm x)" in image_eqI) + apply (auto simp: divide_simps) + using norm_ge_zero [of x] apply linarith+ + done + 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: divide_simps) + done + then show "(\z::'a. z /\<^sub>R (1 + norm z)) ` ?U = ?B" + by (force simp: divide_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 + 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: divide_simps) + show "\y. y /\<^sub>R (1 + norm y) /\<^sub>R (1 - norm (y /\<^sub>R (1 + norm y))) = y" + apply (auto simp: divide_simps) + apply (metis le_add_same_cancel1 norm_ge_zero not_le zero_less_one) + done +qed + +proposition homeomorphic_ball_UNIV: + fixes a ::"'a::real_normed_vector" + assumes "0 < r" shows "ball a r homeomorphic (UNIV:: 'a set)" + using assms homeomorphic_ball01_UNIV homeomorphic_balls(1) homeomorphic_trans zero_less_one by blast + + +subsection%unimportant \Completeness of "Isometry" (up to constant bounds)\ + +lemma cauchy_isometric:\ \TODO: rename lemma to \Cauchy_isometric\\ + assumes e: "e > 0" + and s: "subspace s" + and f: "bounded_linear f" + and normf: "\x\s. norm (f x) \ e * norm x" + and xs: "\n. x n \ s" + and cf: "Cauchy (f \ x)" + shows "Cauchy x" +proof - + interpret f: bounded_linear f by fact + have "\N. \n\N. norm (x n - x N) < d" if "d > 0" for d :: real + proof - + from that obtain N where N: "\n\N. norm (f (x n) - f (x N)) < e * d" + using cf[unfolded Cauchy_def o_def dist_norm, THEN spec[where x="e*d"]] e + by auto + have "norm (x n - x N) < d" if "n \ N" for n + proof - + have "e * norm (x n - x N) \ norm (f (x n - x N))" + using subspace_diff[OF s, of "x n" "x N"] + using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]] + using normf[THEN bspec[where x="x n - x N"]] + by auto + also have "norm (f (x n - x N)) < e * d" + using \N \ n\ N unfolding f.diff[symmetric] by auto + finally show ?thesis + using \e>0\ by simp + qed + then show ?thesis by auto + qed + then show ?thesis + by (simp add: Cauchy_altdef2 dist_norm) +qed + +lemma complete_isometric_image: + assumes "0 < e" + and s: "subspace s" + and f: "bounded_linear f" + and normf: "\x\s. norm(f x) \ e * norm(x)" + and cs: "complete s" + shows "complete (f ` s)" +proof - + have "\l\f ` s. (g \ l) sequentially" + if as:"\n::nat. g n \ f ` s" and cfg:"Cauchy g" for g + proof - + from that obtain x where "\n. x n \ s \ g n = f (x n)" + using choice[of "\ n xa. xa \ s \ g n = f xa"] by auto + then have x: "\n. x n \ s" "\n. g n = f (x n)" by auto + then have "f \ x = g" by (simp add: fun_eq_iff) + then obtain l where "l\s" and l:"(x \ l) sequentially" + using cs[unfolded complete_def, THEN spec[where x=x]] + using cauchy_isometric[OF \0 < e\ s f normf] and cfg and x(1) + by auto + then show ?thesis + using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l] + by (auto simp: \f \ x = g\) + qed + then show ?thesis + unfolding complete_def by auto +qed + + end \ No newline at end of file diff -r d692ef26021e -r 1331e57b54c6 src/HOL/Analysis/Elementary_Topology.thy --- a/src/HOL/Analysis/Elementary_Topology.thy Mon Jan 07 10:22:22 2019 +0100 +++ b/src/HOL/Analysis/Elementary_Topology.thy Mon Jan 07 11:29:34 2019 +0100 @@ -19,6 +19,123 @@ using openI by auto +subsubsection%unimportant \Archimedean properties and useful consequences\ + +text\Bernoulli's inequality\ +proposition Bernoulli_inequality: + fixes x :: real + assumes "-1 \ x" + shows "1 + n * x \ (1 + x) ^ n" +proof (induct n) + case 0 + then show ?case by simp +next + case (Suc n) + have "1 + Suc n * x \ 1 + (Suc n)*x + n * x^2" + by (simp add: algebra_simps) + also have "... = (1 + x) * (1 + n*x)" + by (auto simp: power2_eq_square algebra_simps of_nat_Suc) + also have "... \ (1 + x) ^ Suc n" + using Suc.hyps assms mult_left_mono by fastforce + finally show ?case . +qed + +corollary Bernoulli_inequality_even: + fixes x :: real + assumes "even n" + shows "1 + n * x \ (1 + x) ^ n" +proof (cases "-1 \ x \ n=0") + case True + then show ?thesis + by (auto simp: Bernoulli_inequality) +next + case False + then have "real n \ 1" + by simp + with False have "n * x \ -1" + by (metis linear minus_zero mult.commute mult.left_neutral mult_left_mono_neg neg_le_iff_le order_trans zero_le_one) + then have "1 + n * x \ 0" + by auto + also have "... \ (1 + x) ^ n" + using assms + using zero_le_even_power by blast + finally show ?thesis . +qed + +corollary real_arch_pow: + fixes x :: real + assumes x: "1 < x" + shows "\n. y < x^n" +proof - + from x have x0: "x - 1 > 0" + by arith + from reals_Archimedean3[OF x0, rule_format, of y] + obtain n :: nat where n: "y < real n * (x - 1)" by metis + from x0 have x00: "x- 1 \ -1" by arith + from Bernoulli_inequality[OF x00, of n] n + have "y < x^n" by auto + then show ?thesis by metis +qed + +corollary real_arch_pow_inv: + fixes x y :: real + assumes y: "y > 0" + and x1: "x < 1" + shows "\n. x^n < y" +proof (cases "x > 0") + case True + with x1 have ix: "1 < 1/x" by (simp add: field_simps) + from real_arch_pow[OF ix, of "1/y"] + obtain n where n: "1/y < (1/x)^n" by blast + then show ?thesis using y \x > 0\ + by (auto simp add: field_simps) +next + case False + with y x1 show ?thesis + by (metis less_le_trans not_less power_one_right) +qed + +lemma forall_pos_mono: + "(\d e::real. d < e \ P d \ P e) \ + (\n::nat. n \ 0 \ P (inverse (real n))) \ (\e. 0 < e \ P e)" + by (metis real_arch_inverse) + +lemma forall_pos_mono_1: + "(\d e::real. d < e \ P d \ P e) \ + (\n. P (inverse (real (Suc n)))) \ 0 < e \ P e" + apply (rule forall_pos_mono) + apply auto + apply (metis Suc_pred of_nat_Suc) + done + +subsubsection%unimportant \Affine transformations of intervals\ + +lemma real_affinity_le: "0 < m \ m * x + c \ y \ x \ inverse m * y + - (c / m)" + for m :: "'a::linordered_field" + by (simp add: field_simps) + +lemma real_le_affinity: "0 < m \ y \ m * x + c \ inverse m * y + - (c / m) \ x" + for m :: "'a::linordered_field" + by (simp add: field_simps) + +lemma real_affinity_lt: "0 < m \ m * x + c < y \ x < inverse m * y + - (c / m)" + for m :: "'a::linordered_field" + by (simp add: field_simps) + +lemma real_lt_affinity: "0 < m \ y < m * x + c \ inverse m * y + - (c / m) < x" + for m :: "'a::linordered_field" + by (simp add: field_simps) + +lemma real_affinity_eq: "m \ 0 \ m * x + c = y \ x = inverse m * y + - (c / m)" + for m :: "'a::linordered_field" + by (simp add: field_simps) + +lemma real_eq_affinity: "m \ 0 \ y = m * x + c \ inverse m * y + - (c / m) = x" + for m :: "'a::linordered_field" + by (simp add: field_simps) + + + subsection \Topological Basis\ context topological_space @@ -1112,6 +1229,23 @@ lemma trivial_limit_eq: "trivial_limit net \ (\P. eventually P net)" by (simp add: filter_eq_iff) +lemma Lim_topological: + "(f \ l) net \ + trivial_limit net \ (\S. open S \ l \ S \ eventually (\x. f x \ S) net)" + unfolding tendsto_def trivial_limit_eq by auto + +lemma eventually_within_Un: + "eventually P (at x within (s \ t)) \ + eventually P (at x within s) \ eventually P (at x within t)" + unfolding eventually_at_filter + by (auto elim!: eventually_rev_mp) + +lemma Lim_within_union: + "(f \ l) (at x within (s \ t)) \ + (f \ l) (at x within s) \ (f \ l) (at x within t)" + unfolding tendsto_def + by (auto simp: eventually_within_Un) + subsection \Limits\ @@ -1971,6 +2105,73 @@ by (rule countable_acc_point_imp_seq_compact) (metis islimpt_eq_acc_point) +subsection%unimportant \Cartesian products\ + +lemma seq_compact_Times: "seq_compact s \ seq_compact t \ seq_compact (s \ t)" + unfolding seq_compact_def + apply clarify + apply (drule_tac x="fst \ f" in spec) + apply (drule mp, simp add: mem_Times_iff) + apply (clarify, rename_tac l1 r1) + apply (drule_tac x="snd \ f \ r1" in spec) + apply (drule mp, simp add: mem_Times_iff) + apply (clarify, rename_tac l2 r2) + apply (rule_tac x="(l1, l2)" in rev_bexI, simp) + apply (rule_tac x="r1 \ r2" in exI) + apply (rule conjI, simp add: strict_mono_def) + apply (drule_tac f=r2 in LIMSEQ_subseq_LIMSEQ, assumption) + apply (drule (1) tendsto_Pair) back + apply (simp add: o_def) + done + +lemma compact_Times: + assumes "compact s" "compact t" + shows "compact (s \ t)" +proof (rule compactI) + fix C + assume C: "\t\C. open t" "s \ t \ \C" + have "\x\s. \a. open a \ x \ a \ (\d\C. finite d \ a \ t \ \d)" + proof + fix x + assume "x \ s" + have "\y\t. \a b c. c \ C \ open a \ open b \ x \ a \ y \ b \ a \ b \ c" (is "\y\t. ?P y") + proof + fix y + assume "y \ t" + with \x \ s\ C obtain c where "c \ C" "(x, y) \ c" "open c" by auto + then show "?P y" by (auto elim!: open_prod_elim) + qed + then obtain a b c where b: "\y. y \ t \ open (b y)" + and c: "\y. y \ t \ c y \ C \ open (a y) \ open (b y) \ x \ a y \ y \ b y \ a y \ b y \ c y" + by metis + then have "\y\t. open (b y)" "t \ (\y\t. b y)" by auto + with compactE_image[OF \compact t\] obtain D where D: "D \ t" "finite D" "t \ (\y\D. b y)" + by metis + moreover from D c have "(\y\D. a y) \ t \ (\y\D. c y)" + by (fastforce simp: subset_eq) + ultimately show "\a. open a \ x \ a \ (\d\C. finite d \ a \ t \ \d)" + using c by (intro exI[of _ "c`D"] exI[of _ "\(a`D)"] conjI) (auto intro!: open_INT) + qed + then obtain a d where a: "\x. x\s \ open (a x)" "s \ (\x\s. a x)" + and d: "\x. x \ s \ d x \ C \ finite (d x) \ a x \ t \ \d x" + unfolding subset_eq UN_iff by metis + moreover + from compactE_image[OF \compact s\ a] + obtain e where e: "e \ s" "finite e" and s: "s \ (\x\e. a x)" + by auto + moreover + { + from s have "s \ t \ (\x\e. a x \ t)" + by auto + also have "\ \ (\x\e. \d x)" + using d \e \ s\ by (intro UN_mono) auto + finally have "s \ t \ (\x\e. \d x)" . + } + ultimately show "\C'\C. finite C' \ s \ t \ \C'" + by (intro exI[of _ "(\x\e. d x)"]) (auto simp: subset_eq) +qed + + subsection \Continuity\ lemma continuous_at_imp_continuous_within: @@ -2096,5 +2297,296 @@ using T_def by (auto elim!: eventually_mono) qed +subsection \Homeomorphisms\ + +definition%important "homeomorphism s t f g \ + (\x\s. (g(f x) = x)) \ (f ` s = t) \ continuous_on s f \ + (\y\t. (f(g y) = y)) \ (g ` t = s) \ continuous_on t g" + +lemma homeomorphismI [intro?]: + assumes "continuous_on S f" "continuous_on T g" + "f ` S \ T" "g ` T \ S" "\x. x \ S \ g(f x) = x" "\y. y \ T \ f(g y) = y" + shows "homeomorphism S T f g" + using assms by (force simp: homeomorphism_def) + +lemma homeomorphism_translation: + fixes a :: "'a :: real_normed_vector" + shows "homeomorphism ((+) a ` S) S ((+) (- a)) ((+) a)" +unfolding homeomorphism_def by (auto simp: algebra_simps continuous_intros) + +lemma homeomorphism_ident: "homeomorphism T T (\a. a) (\a. a)" + by (rule homeomorphismI) (auto simp: continuous_on_id) + +lemma homeomorphism_compose: + assumes "homeomorphism S T f g" "homeomorphism T U h k" + shows "homeomorphism S U (h o f) (g o k)" + using assms + unfolding homeomorphism_def + by (intro conjI ballI continuous_on_compose) (auto simp: image_comp [symmetric]) + +lemma homeomorphism_symD: "homeomorphism S t f g \ homeomorphism t S g f" + by (simp add: homeomorphism_def) + +lemma homeomorphism_sym: "homeomorphism S t f g = homeomorphism t S g f" + by (force simp: homeomorphism_def) + +definition%important homeomorphic :: "'a::topological_space set \ 'b::topological_space set \ bool" + (infixr "homeomorphic" 60) + where "s homeomorphic t \ (\f g. homeomorphism s t f g)" + +lemma homeomorphic_empty [iff]: + "S homeomorphic {} \ S = {}" "{} homeomorphic S \ S = {}" + by (auto simp: homeomorphic_def homeomorphism_def) + +lemma homeomorphic_refl: "s homeomorphic s" + unfolding homeomorphic_def homeomorphism_def + using continuous_on_id + apply (rule_tac x = "(\x. x)" in exI) + apply (rule_tac x = "(\x. x)" in exI) + apply blast + done + +lemma homeomorphic_sym: "s homeomorphic t \ t homeomorphic s" + unfolding homeomorphic_def homeomorphism_def + by blast + +lemma homeomorphic_trans [trans]: + assumes "S homeomorphic T" + and "T homeomorphic U" + shows "S homeomorphic U" + using assms + unfolding homeomorphic_def +by (metis homeomorphism_compose) + +lemma homeomorphic_minimal: + "s homeomorphic t \ + (\f g. (\x\s. f(x) \ t \ (g(f(x)) = x)) \ + (\y\t. g(y) \ s \ (f(g(y)) = y)) \ + continuous_on s f \ continuous_on t g)" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + by (fastforce simp: homeomorphic_def homeomorphism_def) +next + assume ?rhs + then show ?lhs + apply clarify + unfolding homeomorphic_def homeomorphism_def + by (metis equalityI image_subset_iff subsetI) + qed + +lemma homeomorphicI [intro?]: + "\f ` S = T; g ` T = S; + continuous_on S f; continuous_on T g; + \x. x \ S \ g(f(x)) = x; + \y. y \ T \ f(g(y)) = y\ \ S homeomorphic T" +unfolding homeomorphic_def homeomorphism_def by metis + +lemma homeomorphism_of_subsets: + "\homeomorphism S T f g; S' \ S; T'' \ T; f ` S' = T'\ + \ homeomorphism S' T' f g" +apply (auto simp: homeomorphism_def elim!: continuous_on_subset) +by (metis subsetD imageI) + +lemma homeomorphism_apply1: "\homeomorphism S T f g; x \ S\ \ g(f x) = x" + by (simp add: homeomorphism_def) + +lemma homeomorphism_apply2: "\homeomorphism S T f g; x \ T\ \ f(g x) = x" + by (simp add: homeomorphism_def) + +lemma homeomorphism_image1: "homeomorphism S T f g \ f ` S = T" + by (simp add: homeomorphism_def) + +lemma homeomorphism_image2: "homeomorphism S T f g \ g ` T = S" + by (simp add: homeomorphism_def) + +lemma homeomorphism_cont1: "homeomorphism S T f g \ continuous_on S f" + by (simp add: homeomorphism_def) + +lemma homeomorphism_cont2: "homeomorphism S T f g \ continuous_on T g" + by (simp add: homeomorphism_def) + +lemma continuous_on_no_limpt: + "(\x. \ x islimpt S) \ continuous_on S f" + unfolding continuous_on_def + by (metis UNIV_I empty_iff eventually_at_topological islimptE open_UNIV tendsto_def trivial_limit_within) + +lemma continuous_on_finite: + fixes S :: "'a::t1_space set" + shows "finite S \ continuous_on S f" +by (metis continuous_on_no_limpt islimpt_finite) + +lemma homeomorphic_finite: + fixes S :: "'a::t1_space set" and T :: "'b::t1_space set" + assumes "finite T" + shows "S homeomorphic T \ finite S \ finite T \ card S = card T" (is "?lhs = ?rhs") +proof + assume "S homeomorphic T" + with assms show ?rhs + apply (auto simp: homeomorphic_def homeomorphism_def) + apply (metis finite_imageI) + by (metis card_image_le finite_imageI le_antisym) +next + assume R: ?rhs + with finite_same_card_bij obtain h where "bij_betw h S T" + by auto + with R show ?lhs + apply (auto simp: homeomorphic_def homeomorphism_def continuous_on_finite) + apply (rule_tac x=h in exI) + apply (rule_tac x="inv_into S h" in exI) + apply (auto simp: bij_betw_inv_into_left bij_betw_inv_into_right bij_betw_imp_surj_on inv_into_into bij_betwE) + apply (metis bij_betw_def bij_betw_inv_into) + done +qed + +text \Relatively weak hypotheses if a set is compact.\ + +lemma homeomorphism_compact: + fixes f :: "'a::topological_space \ 'b::t2_space" + assumes "compact s" "continuous_on s f" "f ` s = t" "inj_on f s" + shows "\g. homeomorphism s t f g" +proof - + define g where "g x = (SOME y. y\s \ f y = x)" for x + have g: "\x\s. g (f x) = x" + using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto + { + fix y + assume "y \ t" + then obtain x where x:"f x = y" "x\s" + using assms(3) by auto + then have "g (f x) = x" using g by auto + then have "f (g y) = y" unfolding x(1)[symmetric] by auto + } + then have g':"\x\t. f (g x) = x" by auto + moreover + { + fix x + have "x\s \ x \ g ` t" + using g[THEN bspec[where x=x]] + unfolding image_iff + using assms(3) + by (auto intro!: bexI[where x="f x"]) + moreover + { + assume "x\g ` t" + then obtain y where y:"y\t" "g y = x" by auto + then obtain x' where x':"x'\s" "f x' = y" + using assms(3) by auto + then have "x \ s" + unfolding g_def + using someI2[of "\b. b\s \ f b = y" x' "\x. x\s"] + unfolding y(2)[symmetric] and g_def + by auto + } + ultimately have "x\s \ x \ g ` t" .. + } + then have "g ` t = s" by auto + ultimately show ?thesis + unfolding homeomorphism_def homeomorphic_def + apply (rule_tac x=g in exI) + using g and assms(3) and continuous_on_inv[OF assms(2,1), of g, unfolded assms(3)] and assms(2) + apply auto + done +qed + +lemma homeomorphic_compact: + fixes f :: "'a::topological_space \ 'b::t2_space" + shows "compact s \ continuous_on s f \ (f ` s = t) \ inj_on f s \ s homeomorphic t" + unfolding homeomorphic_def by (metis homeomorphism_compact) + +text\Preservation of topological properties.\ + +lemma homeomorphic_compactness: "s homeomorphic t \ (compact s \ compact t)" + unfolding homeomorphic_def homeomorphism_def + by (metis compact_continuous_image) + + +subsection%unimportant \On Linorder Topologies\ + +lemma islimpt_greaterThanLessThan1: + fixes a b::"'a::{linorder_topology, dense_order}" + assumes "a < b" + shows "a islimpt {a<.. T" + from open_right[OF this \a < b\] + obtain c where c: "a < c" "{a.. T" by auto + with assms dense[of a "min c b"] + show "\y\{a<.. T \ y \ a" + by (metis atLeastLessThan_iff greaterThanLessThan_iff min_less_iff_conj + not_le order.strict_implies_order subset_eq) +qed + +lemma islimpt_greaterThanLessThan2: + fixes a b::"'a::{linorder_topology, dense_order}" + assumes "a < b" + shows "b islimpt {a<.. T" + from open_left[OF this \a < b\] + obtain c where c: "c < b" "{c<..b} \ T" by auto + with assms dense[of "max a c" b] + show "\y\{a<.. T \ y \ b" + by (metis greaterThanAtMost_iff greaterThanLessThan_iff max_less_iff_conj + not_le order.strict_implies_order subset_eq) +qed + +lemma closure_greaterThanLessThan[simp]: + fixes a b::"'a::{linorder_topology, dense_order}" + shows "a < b \ closure {a <..< b} = {a .. b}" (is "_ \ ?l = ?r") +proof + have "?l \ closure ?r" + by (rule closure_mono) auto + thus "closure {a<.. {a..b}" by simp +qed (auto simp: closure_def order.order_iff_strict islimpt_greaterThanLessThan1 + islimpt_greaterThanLessThan2) + +lemma closure_greaterThan[simp]: + fixes a b::"'a::{no_top, linorder_topology, dense_order}" + shows "closure {a<..} = {a..}" +proof - + from gt_ex obtain b where "a < b" by auto + hence "{a<..} = {a<.. {b..}" by auto + also have "closure \ = {a..}" using \a < b\ unfolding closure_Un + by auto + finally show ?thesis . +qed + +lemma closure_lessThan[simp]: + fixes b::"'a::{no_bot, linorder_topology, dense_order}" + shows "closure {.. {..a}" by auto + also have "closure \ = {..b}" using \a < b\ unfolding closure_Un + by auto + finally show ?thesis . +qed + +lemma closure_atLeastLessThan[simp]: + fixes a b::"'a::{linorder_topology, dense_order}" + assumes "a < b" + shows "closure {a ..< b} = {a .. b}" +proof - + from assms have "{a ..< b} = {a} \ {a <..< b}" by auto + also have "closure \ = {a .. b}" unfolding closure_Un + by (auto simp: assms less_imp_le) + finally show ?thesis . +qed + +lemma closure_greaterThanAtMost[simp]: + fixes a b::"'a::{linorder_topology, dense_order}" + assumes "a < b" + shows "closure {a <.. b} = {a .. b}" +proof - + from assms have "{a <.. b} = {b} \ {a <..< b}" by auto + also have "closure \ = {a .. b}" unfolding closure_Un + by (auto simp: assms less_imp_le) + finally show ?thesis . +qed + end \ No newline at end of file diff -r d692ef26021e -r 1331e57b54c6 src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Mon Jan 07 10:22:22 2019 +0100 +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Mon Jan 07 11:29:34 2019 +0100 @@ -31,6 +31,239 @@ qed +subsection%unimportant\Balls in Euclidean Space\ + +lemma cball_subset_cball_iff: + fixes a :: "'a :: euclidean_space" + shows "cball a r \ cball a' r' \ dist a a' + r \ r' \ r < 0" + (is "?lhs \ ?rhs") +proof + assume ?lhs + then show ?rhs + proof (cases "r < 0") + case True + then show ?rhs by simp + next + case False + then have [simp]: "r \ 0" by simp + have "norm (a - a') + r \ r'" + proof (cases "a = a'") + case True + then show ?thesis + using subsetD [where c = "a + r *\<^sub>R (SOME i. i \ Basis)", OF \?lhs\] subsetD [where c = a, OF \?lhs\] + by (force simp: SOME_Basis dist_norm) + next + case False + have "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = norm (a' - a - (r / norm (a - a')) *\<^sub>R (a - a'))" + by (simp add: algebra_simps) + also have "... = norm ((-1 - (r / norm (a - a'))) *\<^sub>R (a - a'))" + by (simp add: algebra_simps) + also from \a \ a'\ have "... = \- norm (a - a') - r\" + by (simp add: abs_mult_pos field_simps) + finally have [simp]: "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = \norm (a - a') + r\" + by linarith + from \a \ a'\ show ?thesis + using subsetD [where c = "a' + (1 + r / norm(a - a')) *\<^sub>R (a - a')", OF \?lhs\] + by (simp add: dist_norm scaleR_add_left) + qed + then show ?rhs + by (simp add: dist_norm) + qed +next + assume ?rhs + then show ?lhs + by (auto simp: ball_def dist_norm) + (metis add.commute add_le_cancel_right dist_norm dist_triangle3 order_trans) +qed + +lemma cball_subset_ball_iff: "cball a r \ ball a' r' \ dist a a' + r < r' \ r < 0" + (is "?lhs \ ?rhs") + for a :: "'a::euclidean_space" +proof + assume ?lhs + then show ?rhs + proof (cases "r < 0") + case True then + show ?rhs by simp + next + case False + then have [simp]: "r \ 0" by simp + have "norm (a - a') + r < r'" + proof (cases "a = a'") + case True + then show ?thesis + using subsetD [where c = "a + r *\<^sub>R (SOME i. i \ Basis)", OF \?lhs\] subsetD [where c = a, OF \?lhs\] + by (force simp: SOME_Basis dist_norm) + next + case False + have False if "norm (a - a') + r \ r'" + proof - + from that have "\r' - norm (a - a')\ \ r" + by (simp split: abs_split) + (metis \0 \ r\ \?lhs\ centre_in_cball dist_commute dist_norm less_asym mem_ball subset_eq) + then show ?thesis + using subsetD [where c = "a + (r' / norm(a - a') - 1) *\<^sub>R (a - a')", OF \?lhs\] \a \ a'\ + by (simp add: dist_norm field_simps) + (simp add: diff_divide_distrib scaleR_left_diff_distrib) + qed + then show ?thesis by force + qed + then show ?rhs by (simp add: dist_norm) + qed +next + assume ?rhs + then show ?lhs + by (auto simp: ball_def dist_norm) + (metis add.commute add_le_cancel_right dist_norm dist_triangle3 le_less_trans) +qed + +lemma ball_subset_cball_iff: "ball a r \ cball a' r' \ dist a a' + r \ r' \ r \ 0" + (is "?lhs = ?rhs") + for a :: "'a::euclidean_space" +proof (cases "r \ 0") + case True + then show ?thesis + using dist_not_less_zero less_le_trans by force +next + case False + show ?thesis + proof + assume ?lhs + then have "(cball a r \ cball a' r')" + by (metis False closed_cball closure_ball closure_closed closure_mono not_less) + with False show ?rhs + by (fastforce iff: cball_subset_cball_iff) + next + assume ?rhs + with False show ?lhs + using ball_subset_cball cball_subset_cball_iff by blast + qed +qed + +lemma ball_subset_ball_iff: + fixes a :: "'a :: euclidean_space" + shows "ball a r \ ball a' r' \ dist a a' + r \ r' \ r \ 0" + (is "?lhs = ?rhs") +proof (cases "r \ 0") + case True then show ?thesis + using dist_not_less_zero less_le_trans by force +next + case False show ?thesis + proof + assume ?lhs + then have "0 < r'" + by (metis (no_types) False \?lhs\ centre_in_ball dist_norm le_less_trans mem_ball norm_ge_zero not_less set_mp) + then have "(cball a r \ cball a' r')" + by (metis False\?lhs\ closure_ball closure_mono not_less) + then show ?rhs + using False cball_subset_cball_iff by fastforce + next + assume ?rhs then show ?lhs + apply (auto simp: ball_def) + apply (metis add.commute add_le_cancel_right dist_commute dist_triangle_lt not_le order_trans) + using dist_not_less_zero order.strict_trans2 apply blast + done + qed +qed + + +lemma ball_eq_ball_iff: + fixes x :: "'a :: euclidean_space" + shows "ball x d = ball y e \ d \ 0 \ e \ 0 \ x=y \ d=e" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + proof (cases "d \ 0 \ e \ 0") + case True + with \?lhs\ show ?rhs + by safe (simp_all only: ball_eq_empty [of y e, symmetric] ball_eq_empty [of x d, symmetric]) + next + case False + with \?lhs\ show ?rhs + apply (auto simp: set_eq_subset ball_subset_ball_iff dist_norm norm_minus_commute algebra_simps) + apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) + apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) + done + qed +next + assume ?rhs then show ?lhs + by (auto simp: set_eq_subset ball_subset_ball_iff) +qed + +lemma cball_eq_cball_iff: + fixes x :: "'a :: euclidean_space" + shows "cball x d = cball y e \ d < 0 \ e < 0 \ x=y \ d=e" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + proof (cases "d < 0 \ e < 0") + case True + with \?lhs\ show ?rhs + by safe (simp_all only: cball_eq_empty [of y e, symmetric] cball_eq_empty [of x d, symmetric]) + next + case False + with \?lhs\ show ?rhs + apply (auto simp: set_eq_subset cball_subset_cball_iff dist_norm norm_minus_commute algebra_simps) + apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) + apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) + done + qed +next + assume ?rhs then show ?lhs + by (auto simp: set_eq_subset cball_subset_cball_iff) +qed + +lemma ball_eq_cball_iff: + fixes x :: "'a :: euclidean_space" + shows "ball x d = cball y e \ d \ 0 \ e < 0" (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + apply (auto simp: set_eq_subset ball_subset_cball_iff cball_subset_ball_iff algebra_simps) + apply (metis add_increasing2 add_le_cancel_right add_less_same_cancel1 dist_not_less_zero less_le_trans zero_le_dist) + apply (metis add_less_same_cancel1 dist_not_less_zero less_le_trans not_le) + using \?lhs\ ball_eq_empty cball_eq_empty apply blast+ + done +next + assume ?rhs then show ?lhs by auto +qed + +lemma cball_eq_ball_iff: + fixes x :: "'a :: euclidean_space" + shows "cball x d = ball y e \ d < 0 \ e \ 0" + using ball_eq_cball_iff by blast + +lemma finite_ball_avoid: + fixes S :: "'a :: euclidean_space set" + assumes "open S" "finite X" "p \ S" + shows "\e>0. \w\ball p e. w\S \ (w\p \ w\X)" +proof - + obtain e1 where "0 < e1" and e1_b:"ball p e1 \ S" + using open_contains_ball_eq[OF \open S\] assms by auto + obtain e2 where "0 < e2" and "\x\X. x \ p \ e2 \ dist p x" + using finite_set_avoid[OF \finite X\,of p] by auto + hence "\w\ball p (min e1 e2). w\S \ (w\p \ w\X)" using e1_b by auto + thus "\e>0. \w\ball p e. w \ S \ (w \ p \ w \ X)" using \e2>0\ \e1>0\ + apply (rule_tac x="min e1 e2" in exI) + by auto +qed + +lemma finite_cball_avoid: + fixes S :: "'a :: euclidean_space set" + assumes "open S" "finite X" "p \ S" + shows "\e>0. \w\cball p e. w\S \ (w\p \ w\X)" +proof - + obtain e1 where "e1>0" and e1: "\w\ball p e1. w\S \ (w\p \ w\X)" + using finite_ball_avoid[OF assms] by auto + define e2 where "e2 \ e1/2" + have "e2>0" and "e2 < e1" unfolding e2_def using \e1>0\ by auto + then have "cball p e2 \ ball p e1" by (subst cball_subset_ball_iff,auto) + then show "\e>0. \w\cball p e. w \ S \ (w \ p \ w \ X)" using \e2>0\ e1 by auto +qed + + subsection \Boxes\ abbreviation One :: "'a::euclidean_space" @@ -537,6 +770,65 @@ by (auto simp: box_def inner_sum_left inner_Basis sum.If_cases) qed +lemma image_affinity_cbox: fixes m::real + fixes a b c :: "'a::euclidean_space" + shows "(\x. m *\<^sub>R x + c) ` cbox a b = + (if cbox a b = {} then {} + else (if 0 \ m then cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) + else cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c)))" +proof (cases "m = 0") + case True + { + fix x + assume "\i\Basis. x \ i \ c \ i" "\i\Basis. c \ i \ x \ i" + then have "x = c" + by (simp add: dual_order.antisym euclidean_eqI) + } + moreover have "c \ cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c)" + unfolding True by (auto simp: cbox_sing) + ultimately show ?thesis using True by (auto simp: cbox_def) +next + case False + { + fix y + assume "\i\Basis. a \ i \ y \ i" "\i\Basis. y \ i \ b \ i" "m > 0" + then have "\i\Basis. (m *\<^sub>R a + c) \ i \ (m *\<^sub>R y + c) \ i" and "\i\Basis. (m *\<^sub>R y + c) \ i \ (m *\<^sub>R b + c) \ i" + by (auto simp: inner_distrib) + } + moreover + { + fix y + assume "\i\Basis. a \ i \ y \ i" "\i\Basis. y \ i \ b \ i" "m < 0" + then have "\i\Basis. (m *\<^sub>R b + c) \ i \ (m *\<^sub>R y + c) \ i" and "\i\Basis. (m *\<^sub>R y + c) \ i \ (m *\<^sub>R a + c) \ i" + by (auto simp: mult_left_mono_neg inner_distrib) + } + moreover + { + fix y + assume "m > 0" and "\i\Basis. (m *\<^sub>R a + c) \ i \ y \ i" and "\i\Basis. y \ i \ (m *\<^sub>R b + c) \ i" + then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" + unfolding image_iff Bex_def mem_box + apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) + apply (auto simp: pos_le_divide_eq pos_divide_le_eq mult.commute inner_distrib inner_diff_left) + done + } + moreover + { + fix y + assume "\i\Basis. (m *\<^sub>R b + c) \ i \ y \ i" "\i\Basis. y \ i \ (m *\<^sub>R a + c) \ i" "m < 0" + then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" + unfolding image_iff Bex_def mem_box + apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) + apply (auto simp: neg_le_divide_eq neg_divide_le_eq mult.commute inner_distrib inner_diff_left) + done + } + ultimately show ?thesis using False by (auto simp: cbox_def) +qed + +lemma image_smult_cbox:"(\x. m *\<^sub>R (x::_::euclidean_space)) ` cbox a b = + (if cbox a b = {} then {} else if 0 \ m then cbox (m *\<^sub>R a) (m *\<^sub>R b) else cbox (m *\<^sub>R b) (m *\<^sub>R a))" + using image_affinity_cbox[of m 0 a b] by auto + subsection \General Intervals\ @@ -756,7 +1048,8 @@ using bounded_bilinear_inner assms by (rule bounded_bilinear.continuous_on) -subsection \Openness of halfspaces.\ + +subsection%unimportant \Openness of halfspaces.\ lemma open_halfspace_lt: "open {x. inner a x < b}" by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) @@ -781,9 +1074,110 @@ shows "open {x. x Closure of halfspaces and hyperplanes\ + +lemma continuous_at_inner: "continuous (at x) (inner a)" + unfolding continuous_at by (intro tendsto_intros) + +lemma closed_halfspace_le: "closed {x. inner a x \ b}" + by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) + +lemma closed_halfspace_ge: "closed {x. inner a x \ b}" + by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) + +lemma closed_hyperplane: "closed {x. inner a x = b}" + by (simp add: closed_Collect_eq continuous_on_inner continuous_on_const continuous_on_id) + +lemma closed_halfspace_component_le: "closed {x::'a::euclidean_space. x\i \ a}" + by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) + +lemma closed_halfspace_component_ge: "closed {x::'a::euclidean_space. x\i \ a}" + by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) + +lemma closed_interval_left: + fixes b :: "'a::euclidean_space" + shows "closed {x::'a. \i\Basis. x\i \ b\i}" + by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) + +lemma closed_interval_right: + fixes a :: "'a::euclidean_space" + shows "closed {x::'a. \i\Basis. a\i \ x\i}" + by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) + +lemma continuous_le_on_closure: + fixes a::real + assumes f: "continuous_on (closure s) f" + and x: "x \ closure(s)" + and xlo: "\x. x \ s ==> f(x) \ a" + shows "f(x) \ a" + using image_closure_subset [OF f] + using image_closure_subset [OF f] closed_halfspace_le [of "1::real" a] assms + by force + +lemma continuous_ge_on_closure: + fixes a::real + assumes f: "continuous_on (closure s) f" + and x: "x \ closure(s)" + and xlo: "\x. x \ s ==> f(x) \ a" + shows "f(x) \ a" + using image_closure_subset [OF f] closed_halfspace_ge [of a "1::real"] assms + by force + + +subsection%unimportant\Some more convenient intermediate-value theorem formulations\ + +lemma connected_ivt_hyperplane: + assumes "connected S" and xy: "x \ S" "y \ S" and b: "inner a x \ b" "b \ inner a y" + shows "\z \ S. inner a z = b" +proof (rule ccontr) + assume as:"\ (\z\S. inner a z = b)" + let ?A = "{x. inner a x < b}" + let ?B = "{x. inner a x > b}" + have "open ?A" "open ?B" + using open_halfspace_lt and open_halfspace_gt by auto + moreover have "?A \ ?B = {}" by auto + moreover have "S \ ?A \ ?B" using as by auto + ultimately show False + using \connected S\[unfolded connected_def not_ex, + THEN spec[where x="?A"], THEN spec[where x="?B"]] + using xy b by auto +qed + +lemma connected_ivt_component: + fixes x::"'a::euclidean_space" + shows "connected S \ x \ S \ y \ S \ x\k \ a \ a \ y\k \ (\z\S. z\k = a)" + using connected_ivt_hyperplane[of S x y "k::'a" a] + by (auto simp: inner_commute) + subsection \Limit Component Bounds\ +lemma Lim_component_le: + fixes f :: "'a \ 'b::euclidean_space" + assumes "(f \ l) net" + and "\ (trivial_limit net)" + and "eventually (\x. f(x)\i \ b) net" + shows "l\i \ b" + by (rule tendsto_le[OF assms(2) tendsto_const tendsto_inner[OF assms(1) tendsto_const] assms(3)]) + +lemma Lim_component_ge: + fixes f :: "'a \ 'b::euclidean_space" + assumes "(f \ l) net" + and "\ (trivial_limit net)" + and "eventually (\x. b \ (f x)\i) net" + shows "b \ l\i" + by (rule tendsto_le[OF assms(2) tendsto_inner[OF assms(1) tendsto_const] tendsto_const assms(3)]) + +lemma Lim_component_eq: + fixes f :: "'a \ 'b::euclidean_space" + assumes net: "(f \ l) net" "\ trivial_limit net" + and ev:"eventually (\x. f(x)\i = b) net" + shows "l\i = b" + using ev[unfolded order_eq_iff eventually_conj_iff] + using Lim_component_ge[OF net, of b i] + using Lim_component_le[OF net, of i b] + by auto + lemma open_box[intro]: "open (box a b)" proof - have "open (\i\Basis. ((\) i) -` {a \ i <..< b \ i})" @@ -1192,6 +1586,67 @@ qed +subsection%unimportant \Diameter\ + +lemma diameter_cball [simp]: + fixes a :: "'a::euclidean_space" + shows "diameter(cball a r) = (if r < 0 then 0 else 2*r)" +proof - + have "diameter(cball a r) = 2*r" if "r \ 0" + proof (rule order_antisym) + show "diameter (cball a r) \ 2*r" + proof (rule diameter_le) + fix x y assume "x \ cball a r" "y \ cball a r" + then have "norm (x - a) \ r" "norm (a - y) \ r" + by (auto simp: dist_norm norm_minus_commute) + then have "norm (x - y) \ r+r" + using norm_diff_triangle_le by blast + then show "norm (x - y) \ 2*r" by simp + qed (simp add: that) + have "2*r = dist (a + r *\<^sub>R (SOME i. i \ Basis)) (a - r *\<^sub>R (SOME i. i \ Basis))" + apply (simp add: dist_norm) + by (metis abs_of_nonneg mult.right_neutral norm_numeral norm_scaleR norm_some_Basis real_norm_def scaleR_2 that) + also have "... \ diameter (cball a r)" + apply (rule diameter_bounded_bound) + using that by (auto simp: dist_norm) + finally show "2*r \ diameter (cball a r)" . + qed + then show ?thesis by simp +qed + +lemma diameter_ball [simp]: + fixes a :: "'a::euclidean_space" + shows "diameter(ball a r) = (if r < 0 then 0 else 2*r)" +proof - + have "diameter(ball a r) = 2*r" if "r > 0" + by (metis bounded_ball diameter_closure closure_ball diameter_cball less_eq_real_def linorder_not_less that) + then show ?thesis + by (simp add: diameter_def) +qed + +lemma diameter_closed_interval [simp]: "diameter {a..b} = (if b < a then 0 else b-a)" +proof - + have "{a .. b} = cball ((a+b)/2) ((b-a)/2)" + by (auto simp: dist_norm abs_if divide_simps split: if_split_asm) + then show ?thesis + by simp +qed + +lemma diameter_open_interval [simp]: "diameter {a<..i \ Basis. a \ i \ b \ i) \ diameter (cbox a b) = dist a b" + by (force simp: diameter_def intro!: cSup_eq_maximum L2_set_mono + simp: euclidean_dist_l2[where 'a='a] cbox_def dist_norm) + + subsection%unimportant\Relating linear images to open/closed/interior/closure\ proposition open_surjective_linear_image: @@ -1294,6 +1749,202 @@ shows "interior(uminus ` S) = image uminus (interior S)" by (simp add: bij_uminus interior_bijective_linear_image linear_uminus) +subsection%unimportant \"Isometry" (up to constant bounds) of Injective Linear Map\ + +proposition injective_imp_isometric: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes s: "closed s" "subspace s" + and f: "bounded_linear f" "\x\s. f x = 0 \ x = 0" + shows "\e>0. \x\s. norm (f x) \ e * norm x" +proof (cases "s \ {0::'a}") + case True + have "norm x \ norm (f x)" if "x \ s" for x + proof - + from True that have "x = 0" by auto + then show ?thesis by simp + qed + then show ?thesis + by (auto intro!: exI[where x=1]) +next + case False + interpret f: bounded_linear f by fact + from False obtain a where a: "a \ 0" "a \ s" + by auto + from False have "s \ {}" + by auto + let ?S = "{f x| x. x \ s \ norm x = norm a}" + let ?S' = "{x::'a. x\s \ norm x = norm a}" + let ?S'' = "{x::'a. norm x = norm a}" + + have "?S'' = frontier (cball 0 (norm a))" + by (simp add: sphere_def dist_norm) + then have "compact ?S''" by (metis compact_cball compact_frontier) + moreover have "?S' = s \ ?S''" by auto + ultimately have "compact ?S'" + using closed_Int_compact[of s ?S''] using s(1) by auto + moreover have *:"f ` ?S' = ?S" by auto + ultimately have "compact ?S" + using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto + then have "closed ?S" + using compact_imp_closed by auto + moreover from a have "?S \ {}" by auto + ultimately obtain b' where "b'\?S" "\y\?S. norm b' \ norm y" + using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto + then obtain b where "b\s" + and ba: "norm b = norm a" + and b: "\x\{x \ s. norm x = norm a}. norm (f b) \ norm (f x)" + unfolding *[symmetric] unfolding image_iff by auto + + let ?e = "norm (f b) / norm b" + have "norm b > 0" + using ba and a and norm_ge_zero by auto + moreover have "norm (f b) > 0" + using f(2)[THEN bspec[where x=b], OF \b\s\] + using \norm b >0\ by simp + ultimately have "0 < norm (f b) / norm b" by simp + moreover + have "norm (f b) / norm b * norm x \ norm (f x)" if "x\s" for x + proof (cases "x = 0") + case True + then show "norm (f b) / norm b * norm x \ norm (f x)" + by auto + next + case False + with \a \ 0\ have *: "0 < norm a / norm x" + unfolding zero_less_norm_iff[symmetric] by simp + have "\x\s. c *\<^sub>R x \ s" for c + using s[unfolded subspace_def] by simp + with \x \ s\ \x \ 0\ have "(norm a / norm x) *\<^sub>R x \ {x \ s. norm x = norm a}" + by simp + with \x \ 0\ \a \ 0\ show "norm (f b) / norm b * norm x \ norm (f x)" + using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]] + unfolding f.scaleR and ba + by (auto simp: mult.commute pos_le_divide_eq pos_divide_le_eq) + qed + ultimately show ?thesis by auto +qed + +proposition closed_injective_image_subspace: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "subspace s" "bounded_linear f" "\x\s. f x = 0 \ x = 0" "closed s" + shows "closed(f ` s)" +proof - + obtain e where "e > 0" and e: "\x\s. e * norm x \ norm (f x)" + using injective_imp_isometric[OF assms(4,1,2,3)] by auto + show ?thesis + using complete_isometric_image[OF \e>0\ assms(1,2) e] and assms(4) + unfolding complete_eq_closed[symmetric] by auto +qed + + +subsection%unimportant \Some properties of a canonical subspace\ + +lemma subspace_substandard: "subspace {x::'a::euclidean_space. (\i\Basis. P i \ x\i = 0)}" + by (auto simp: subspace_def inner_add_left) + +lemma closed_substandard: "closed {x::'a::euclidean_space. \i\Basis. P i \ x\i = 0}" + (is "closed ?A") +proof - + let ?D = "{i\Basis. P i}" + have "closed (\i\?D. {x::'a. x\i = 0})" + by (simp add: closed_INT closed_Collect_eq continuous_on_inner + continuous_on_const continuous_on_id) + also have "(\i\?D. {x::'a. x\i = 0}) = ?A" + by auto + finally show "closed ?A" . +qed + +lemma dim_substandard: + assumes d: "d \ Basis" + shows "dim {x::'a::euclidean_space. \i\Basis. i \ d \ x\i = 0} = card d" (is "dim ?A = _") +proof (rule dim_unique) + from d show "d \ ?A" + by (auto simp: inner_Basis) + from d show "independent d" + by (rule independent_mono [OF independent_Basis]) + have "x \ span d" if "\i\Basis. i \ d \ x \ i = 0" for x + proof - + have "finite d" + by (rule finite_subset [OF d finite_Basis]) + then have "(\i\d. (x \ i) *\<^sub>R i) \ span d" + by (simp add: span_sum span_clauses) + also have "(\i\d. (x \ i) *\<^sub>R i) = (\i\Basis. (x \ i) *\<^sub>R i)" + by (rule sum.mono_neutral_cong_left [OF finite_Basis d]) (auto simp: that) + finally show "x \ span d" + by (simp only: euclidean_representation) + qed + then show "?A \ span d" by auto +qed simp + +text \Hence closure and completeness of all subspaces.\ +lemma ex_card: + assumes "n \ card A" + shows "\S\A. card S = n" +proof (cases "finite A") + case True + from ex_bij_betw_nat_finite[OF this] obtain f where f: "bij_betw f {0..n \ card A\ have "{..< n} \ {..< card A}" "inj_on f {..< n}" + by (auto simp: bij_betw_def intro: subset_inj_on) + ultimately have "f ` {..< n} \ A" "card (f ` {..< n}) = n" + by (auto simp: bij_betw_def card_image) + then show ?thesis by blast +next + case False + with \n \ card A\ show ?thesis by force +qed + +lemma closed_subspace: + fixes s :: "'a::euclidean_space set" + assumes "subspace s" + shows "closed s" +proof - + have "dim s \ card (Basis :: 'a set)" + using dim_subset_UNIV by auto + with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \ Basis" + by auto + let ?t = "{x::'a. \i\Basis. i \ d \ x\i = 0}" + have "\f. linear f \ f ` {x::'a. \i\Basis. i \ d \ x \ i = 0} = s \ + inj_on f {x::'a. \i\Basis. i \ d \ x \ i = 0}" + using dim_substandard[of d] t d assms + by (intro subspace_isomorphism[OF subspace_substandard[of "\i. i \ d"]]) (auto simp: inner_Basis) + then obtain f where f: + "linear f" + "f ` {x. \i\Basis. i \ d \ x \ i = 0} = s" + "inj_on f {x. \i\Basis. i \ d \ x \ i = 0}" + by blast + interpret f: bounded_linear f + using f by (simp add: linear_conv_bounded_linear) + have "x \ ?t \ f x = 0 \ x = 0" for x + using f.zero d f(3)[THEN inj_onD, of x 0] by auto + moreover have "closed ?t" by (rule closed_substandard) + moreover have "subspace ?t" by (rule subspace_substandard) + ultimately show ?thesis + using closed_injective_image_subspace[of ?t f] + unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto +qed + +lemma complete_subspace: "subspace s \ complete s" + for s :: "'a::euclidean_space set" + using complete_eq_closed closed_subspace by auto + +lemma closed_span [iff]: "closed (span s)" + for s :: "'a::euclidean_space set" + by (simp add: closed_subspace subspace_span) + +lemma dim_closure [simp]: "dim (closure s) = dim s" (is "?dc = ?d") + for s :: "'a::euclidean_space set" +proof - + have "?dc \ ?d" + using closure_minimal[OF span_superset, of s] + using closed_subspace[OF subspace_span, of s] + using dim_subset[of "closure s" "span s"] + by simp + then show ?thesis + using dim_subset[OF closure_subset, of s] + by simp +qed + + no_notation eucl_less (infix "