# HG changeset patch # User huffman # Date 1358306928 28800 # Node ID 8c742f9de9f58b4092edfbfd03e0010e3f104e50 # Parent 8757e6aa50eb8ed7d6044b7fdc725dcc350e7e66 generalize topology lemmas; simplify proofs diff -r 8757e6aa50eb -r 8c742f9de9f5 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Thu Jan 17 15:50:56 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Jan 15 19:28:48 2013 -0800 @@ -2152,7 +2152,9 @@ unfolding bounded_def dist_real_def apply(rule_tac x=0 in exI) using assms by auto -lemma bounded_empty[simp]: "bounded {}" by (simp add: bounded_def) +lemma bounded_empty [simp]: "bounded {}" + by (simp add: bounded_def) + lemma bounded_subset: "bounded T \ S \ T ==> bounded S" by (metis bounded_def subset_eq) @@ -2188,17 +2190,6 @@ lemma bounded_ball[simp,intro]: "bounded(ball x e)" by (metis ball_subset_cball bounded_cball bounded_subset) -lemma finite_imp_bounded[intro]: - fixes S :: "'a::metric_space set" assumes "finite S" shows "bounded S" -proof- - { fix a and F :: "'a set" assume as:"bounded F" - then obtain x e where "\y\F. dist x y \ e" unfolding bounded_def by auto - hence "\y\(insert a F). dist x y \ max e (dist x a)" by auto - hence "bounded (insert a F)" unfolding bounded_def by (intro exI) - } - thus ?thesis using finite_induct[of S bounded] using bounded_empty assms by auto -qed - lemma bounded_Un[simp]: "bounded (S \ T) \ bounded S \ bounded T" apply (auto simp add: bounded_def) apply (rename_tac x y r s) @@ -2214,6 +2205,16 @@ lemma bounded_Union[intro]: "finite F \ (\S\F. bounded S) \ bounded(\F)" by (induct rule: finite_induct[of F], auto) +lemma bounded_insert [simp]: "bounded (insert x S) \ bounded S" +proof - + have "\y\{x}. dist x y \ 0" by simp + hence "bounded {x}" unfolding bounded_def by fast + thus ?thesis by (metis insert_is_Un bounded_Un) +qed + +lemma finite_imp_bounded [intro]: "finite S \ bounded S" + by (induct set: finite, simp_all) + lemma bounded_pos: "bounded S \ (\b>0. \x\ S. norm x <= b)" apply (simp add: bounded_iff) apply (subgoal_tac "\x (y::real). 0 < 1 + abs y \ (x <= y \ x <= 1 + abs y)") @@ -2226,9 +2227,6 @@ apply (metis Diff_subset bounded_subset) done -lemma bounded_insert[intro]:"bounded(insert x S) \ bounded S" - by (metis Diff_cancel Un_empty_right Un_insert_right bounded_Un bounded_subset finite.emptyI finite_imp_bounded infinite_remove subset_insertI) - lemma not_bounded_UNIV[simp, intro]: "\ bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)" proof(auto simp add: bounded_pos not_le) @@ -5063,14 +5061,14 @@ qed lemma continuous_attains_sup: - fixes f :: "'a::metric_space \ real" + fixes f :: "'a::topological_space \ real" shows "compact s \ s \ {} \ continuous_on s f ==> (\x \ s. \y \ s. f y \ f x)" using compact_attains_sup[of "f ` s"] using compact_continuous_image[of s f] by auto lemma continuous_attains_inf: - fixes f :: "'a::metric_space \ real" + fixes f :: "'a::topological_space \ real" shows "compact s \ s \ {} \ continuous_on s f \ (\x \ s. \y \ s. f x \ f y)" using compact_attains_inf[of "f ` s"]