# HG changeset patch # User hoelzl # Date 1358535719 -3600 # Node ID 4a2c82644889b94b9c9b53133e45837c6571ba00 # Parent d2c6a0a7fcdf7626c1405e0063818d51906fdea3 generalized diameter from real_normed_vector to metric_space diff -r d2c6a0a7fcdf -r 4a2c82644889 src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Fri Jan 18 20:01:41 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Fri Jan 18 20:01:59 2013 +0100 @@ -3284,7 +3284,7 @@ have "convex hull s \ {}" using hull_subset[of s convex] and assms(2) by auto then obtain u v where obt:"u\convex hull s" "v\convex hull s" "\x\convex hull s. \y\convex hull s. norm (x - y) \ norm (u - v)" - using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] by auto + using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] by (auto simp: dist_norm) thus ?thesis proof(cases "u\s \ v\s", erule_tac disjE) assume "u\s" then obtain y where "y\convex hull s" "norm (u - v) < norm (y - v)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) by auto diff -r d2c6a0a7fcdf -r 4a2c82644889 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Jan 18 20:01:41 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Jan 18 20:01:59 2013 +0100 @@ -3230,10 +3230,10 @@ using bounded_closed_imp_seq_compact[of s] unfolding compact_eq_seq_compact_metric by auto qed -(* TODO: merge this lemma with the ones above *) +(* TODO: is this lemma necessary? *) lemma bounded_increasing_convergent: fixes s :: "nat \ real" - shows "bounded {s n| n. True} \ \n. s n \ s(Suc n) \ \l. s ----> l" + shows "bounded {s n| n. True} \ \n. s n \ s (Suc n) \ \l. s ----> l" using Bseq_mono_convergent[of s] incseq_Suc_iff[of s] by (auto simp: image_def Bseq_eq_bounded convergent_def incseq_def) @@ -3761,6 +3761,10 @@ lemma continuous_at: "continuous (at x) f \ (f ---> f(x)) (at x)" using continuous_within [of x UNIV f] by simp +lemma continuous_isCont: "isCont f x = continuous (at x) f" + unfolding isCont_def LIM_def + unfolding continuous_at Lim_at unfolding dist_nz by auto + lemma continuous_at_within: assumes "continuous (at x) f" shows "continuous (at x within s) f" using assms unfolding continuous_at continuous_within @@ -5098,66 +5102,74 @@ text {* Hence we get the following. *} lemma compact_sup_maxdistance: - fixes s :: "'a::real_normed_vector set" + fixes s :: "'a::metric_space set" assumes "compact s" "s \ {}" - shows "\x\s. \y\s. \u\s. \v\s. norm(u - v) \ norm(x - y)" + shows "\x\s. \y\s. \u\s. \v\s. dist u v \ dist x y" proof- - have "{x - y | x y . x\s \ y\s} \ {}" using `s \ {}` by auto - then obtain x where x:"x\{x - y |x y. x \ s \ y \ s}" "\y\{x - y |x y. x \ s \ y \ s}. norm y \ norm x" - using compact_differences[OF assms(1) assms(1)] - using distance_attains_sup[where 'a="'a", unfolded dist_norm, of "{x - y | x y . x\s \ y\s}" 0] by auto - from x(1) obtain a b where "a\s" "b\s" "x = a - b" by auto - thus ?thesis using x(2)[unfolded `x = a - b`] by blast + 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_dist + continuous_isCont[THEN iffD1] isCont_fst isCont_snd isCont_ident) + ultimately show ?thesis + using continuous_attains_sup[of "s \ s" "\x. dist (fst x) (snd x)"] by auto qed text {* We can state this in terms of diameter of a set. *} -definition "diameter s = (if s = {} then 0::real else Sup {norm(x - y) | x y. x \ s \ y \ s})" - (* TODO: generalize to class metric_space *) +definition "diameter s = (if s = {} then 0::real else Sup {dist x y | x y. x \ s \ y \ s})" + +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 - + let ?D = "{dist x y |x y. x \ s \ y \ s}" + from s obtain z d where z: "\x. x \ s \ dist z x \ d" + unfolding bounded_def by auto + have "dist x y \ Sup ?D" + proof (rule Sup_upper, 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 (insert s, auto) + with `x \ s` show ?thesis + by (auto simp add: 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) + let ?D = "{dist x y |x y. x \ s \ y \ s}" + assume contr: "\ ?thesis" + moreover + from d have "s \ {}" + by (auto simp: diameter_def) + then have "?D \ {}" by auto + ultimately have "Sup ?D \ d" + by (intro Sup_least) (auto simp: not_less) + with `d < diameter s` `s \ {}` show False + by (auto simp: diameter_def) +qed lemma diameter_bounded: assumes "bounded s" - shows "\x\s. \y\s. norm(x - y) \ diameter s" - "\d>0. d < diameter s --> (\x\s. \y\s. norm(x - y) > d)" -proof- - let ?D = "{norm (x - y) |x y. x \ s \ y \ s}" - obtain a where a:"\x\s. norm x \ a" using assms[unfolded bounded_iff] by auto - { fix x y assume "x \ s" "y \ s" - hence "norm (x - y) \ 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: field_simps) } - note * = this - { fix x y assume "x\s" "y\s" hence "s \ {}" by auto - have "norm(x - y) \ diameter s" unfolding diameter_def using `s\{}` *[OF `x\s` `y\s`] `x\s` `y\s` - by simp (blast del: Sup_upper intro!: * Sup_upper) } - moreover - { fix d::real assume "d>0" "d < diameter s" - hence "s\{}" unfolding diameter_def by auto - have "\d' \ ?D. d' > d" - proof(rule ccontr) - assume "\ (\d'\{norm (x - y) |x y. x \ s \ y \ s}. d < d')" - hence "\d'\?D. d' \ d" by auto (metis not_leE) - thus False using `d < diameter s` `s\{}` - apply (auto simp add: diameter_def) - apply (drule Sup_real_iff [THEN [2] rev_iffD2]) - apply (auto, force) - done - qed - hence "\x\s. \y\s. norm(x - y) > d" by auto } - ultimately show "\x\s. \y\s. norm(x - y) \ diameter s" - "\d>0. d < diameter s --> (\x\s. \y\s. norm(x - y) > d)" by auto -qed - -lemma diameter_bounded_bound: - "bounded s \ x \ s \ y \ s ==> norm(x - y) \ diameter s" - using diameter_bounded by blast + shows "\x\s. \y\s. dist x y \ diameter s" + "\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 diameter_compact_attained: - fixes s :: "'a::real_normed_vector set" assumes "compact s" "s \ {}" - shows "\x\s. \y\s. (norm(x - y) = diameter s)" -proof- + 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. norm (u - v) \ norm (x - y)" using compact_sup_maxdistance[OF assms] by auto - hence "diameter s \ norm (x - y)" + 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 + hence "diameter s \ dist x y" unfolding diameter_def by clarsimp (rule Sup_least, fast+) thus ?thesis by (metis b diameter_bounded_bound order_antisym xys)