# HG changeset patch # User hoelzl # Date 1363945303 -3600 # Node ID 33db4b7189af09131c208eb48db236660cbcbac3 # Parent 270b21f3ae0ab22b3a8c8bfabf783bc494e9dcf4 move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets diff -r 270b21f3ae0a -r 33db4b7189af src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Fri Mar 22 10:41:43 2013 +0100 +++ b/src/HOL/Deriv.thy Fri Mar 22 10:41:43 2013 +0100 @@ -566,30 +566,65 @@ done +lemma compact_Icc[simp, intro]: "compact {a .. b::real}" +proof (cases "a \ b", rule compactI) + fix C assume C: "a \ b" "\t\C. open t" "{a..b} \ \C" + def T == "{a .. b}" + from C(1,3) show "\C'\C. finite C' \ {a..b} \ \C'" + proof (induct rule: Bolzano) + case (trans a b c) + then have *: "{a .. c} = {a .. b} \ {b .. c}" by auto + from trans obtain C1 C2 where "C1\C \ finite C1 \ {a..b} \ \C1" "C2\C \ finite C2 \ {b..c} \ \C2" + by (auto simp: *) + with trans show ?case + unfolding * by (intro exI[of _ "C1 \ C2"]) auto + next + case (local x) + then have "x \ \C" using C by auto + with C(2) obtain c where "x \ c" "open c" "c \ C" by auto + then obtain e where "0 < e" "{x - e <..< x + e} \ c" + by (auto simp: open_real_def dist_real_def subset_eq Ball_def abs_less_iff) + with `c \ C` show ?case + by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto + qed +qed simp + subsection {* Boundedness of continuous functions *} text{*By bisection, function continuous on closed interval is bounded above*} +lemma isCont_eq_Ub: + fixes f :: "real \ 'a::linorder_topology" + shows "a \ b \ \x::real. a \ x \ x \ b \ isCont f x \ + \M. (\x. a \ x \ x \ b \ f x \ M) \ (\x. a \ x \ x \ b \ f x = M)" + using continuous_attains_sup[of "{a .. b}" f] + apply (simp add: continuous_at_imp_continuous_on Ball_def) + apply safe + apply (rule_tac x="f x" in exI) + apply auto + done + +lemma isCont_eq_Lb: + fixes f :: "real \ 'a::linorder_topology" + shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ + \M. (\x. a \ x \ x \ b \ M \ f x) \ (\x. a \ x \ x \ b \ f x = M)" + using continuous_attains_inf[of "{a .. b}" f] + apply (simp add: continuous_at_imp_continuous_on Ball_def) + apply safe + apply (rule_tac x="f x" in exI) + apply auto + done + lemma isCont_bounded: - "[| a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \M::real. \x::real. a \ x & x \ b --> f(x) \ M" -apply (cut_tac P = "% (u,v) . a \ u & u \ v & v \ b --> (\M. \x. u \ x & x \ v --> f x \ M)" in lemma_BOLZANO2) -apply safe -apply simp_all -apply (rename_tac x xa ya M Ma) -apply (metis linorder_not_less order_le_less order_trans) -apply (case_tac "a \ x & x \ b") - prefer 2 - apply (rule_tac x = 1 in exI, force) -apply (simp add: LIM_eq isCont_iff) -apply (drule_tac x = x in spec, auto) -apply (erule_tac V = "\M. \x. a \ x & x \ b & ~ f x \ M" in thin_rl) -apply (drule_tac x = 1 in spec, auto) -apply (rule_tac x = s in exI, clarify) -apply (rule_tac x = "\f x\ + 1" in exI, clarify) -apply (drule_tac x = "xa-x" in spec) -apply (auto simp add: abs_ge_self) -done + fixes f :: "real \ 'a::linorder_topology" + shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. \x. a \ x \ x \ b \ f x \ M" + using isCont_eq_Ub[of a b f] by auto + +lemma isCont_has_Ub: + fixes f :: "real \ 'a::linorder_topology" + shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ + \M. (\x. a \ x \ x \ b \ f x \ M) \ (\N. N < M \ (\x. a \ x \ x \ b \ N < f x))" + using isCont_eq_Ub[of a b f] by auto text{*Refine the above to existence of least upper bound*} @@ -597,72 +632,6 @@ (\t. isLub UNIV S t)" by (blast intro: reals_complete) -lemma isCont_has_Ub: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \M::real. (\x::real. a \ x & x \ b --> f(x) \ M) & - (\N. N < M --> (\x. a \ x & x \ b & N < f(x)))" -apply (cut_tac S = "Collect (%y. \x. a \ x & x \ b & y = f x)" - in lemma_reals_complete) -apply auto -apply (drule isCont_bounded, assumption) -apply (auto simp add: isUb_def leastP_def isLub_def setge_def setle_def) -apply (rule exI, auto) -apply (auto dest!: spec simp add: linorder_not_less) -done - -text{*Now show that it attains its upper bound*} - -lemma isCont_eq_Ub: - assumes le: "a \ b" - and con: "\x::real. a \ x & x \ b --> isCont f x" - shows "\M::real. (\x. a \ x & x \ b --> f(x) \ M) & - (\x. a \ x & x \ b & f(x) = M)" -proof - - from isCont_has_Ub [OF le con] - obtain M where M1: "\x. a \ x \ x \ b \ f x \ M" - and M2: "!!N. N \x. a \ x \ x \ b \ N < f x" by blast - show ?thesis - proof (intro exI, intro conjI) - show " \x. a \ x \ x \ b \ f x \ M" by (rule M1) - show "\x. a \ x \ x \ b \ f x = M" - proof (rule ccontr) - assume "\ (\x. a \ x \ x \ b \ f x = M)" - with M1 have M3: "\x. a \ x & x \ b --> f x < M" - by (fastforce simp add: linorder_not_le [symmetric]) - hence "\x. a \ x & x \ b --> isCont (%x. inverse (M - f x)) x" - by (auto simp add: con) - from isCont_bounded [OF le this] - obtain k where k: "!!x. a \ x & x \ b --> inverse (M - f x) \ k" by auto - have Minv: "!!x. a \ x & x \ b --> 0 < inverse (M - f (x))" - by (simp add: M3 algebra_simps) - have "!!x. a \ x & x \ b --> inverse (M - f x) < k+1" using k - by (auto intro: order_le_less_trans [of _ k]) - with Minv - have "!!x. a \ x & x \ b --> inverse(k+1) < inverse(inverse(M - f x))" - by (intro strip less_imp_inverse_less, simp_all) - hence invlt: "!!x. a \ x & x \ b --> inverse(k+1) < M - f x" - by simp - have "M - inverse (k+1) < M" using k [of a] Minv [of a] le - by (simp, arith) - from M2 [OF this] - obtain x where ax: "a \ x & x \ b & M - inverse(k+1) < f x" .. - thus False using invlt [of x] by force - qed - qed -qed - - -text{*Same theorem for lower bound*} - -lemma isCont_eq_Lb: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \M::real. (\x::real. a \ x & x \ b --> M \ f(x)) & - (\x. a \ x & x \ b & f(x) = M)" -apply (subgoal_tac "\x. a \ x & x \ b --> isCont (%x. - (f x)) x") -prefer 2 apply (blast intro: isCont_minus) -apply (drule_tac f = "(%x. - (f x))" in isCont_eq_Ub) -apply safe -apply auto -done - text{*Another version.*} diff -r 270b21f3ae0a -r 33db4b7189af src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Mar 22 10:41:43 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Mar 22 10:41:43 2013 +0100 @@ -2362,28 +2362,6 @@ subsection {* Compactness *} -subsubsection{* Open-cover compactness *} - -definition compact :: "'a::topological_space set \ bool" where - compact_eq_heine_borel: -- "This name is used for backwards compatibility" - "compact S \ (\C. (\c\C. open c) \ S \ \C \ (\D\C. finite D \ S \ \D))" - -lemma compactI: - assumes "\C. \t\C. open t \ s \ \ C \ \C'. C' \ C \ finite C' \ s \ \ C'" - shows "compact s" - unfolding compact_eq_heine_borel using assms by metis - -lemma compactE: - assumes "compact s" and "\t\C. open t" and "s \ \C" - obtains C' where "C' \ C" and "finite C'" and "s \ \C'" - using assms unfolding compact_eq_heine_borel by metis - -lemma compactE_image: - assumes "compact s" and "\t\C. open (f t)" and "s \ (\c\C. f c)" - obtains C' where "C' \ C" and "finite C'" and "s \ (\c\C'. f c)" - using assms unfolding ball_simps[symmetric] SUP_def - by (metis (lifting) finite_subset_image compact_eq_heine_borel[of s]) - subsubsection {* Bolzano-Weierstrass property *} lemma heine_borel_imp_bolzano_weierstrass: @@ -2621,11 +2599,6 @@ text{* In particular, some common special cases. *} -lemma compact_empty[simp]: - "compact {}" - unfolding compact_eq_heine_borel - by auto - lemma compact_union [intro]: assumes "compact s" "compact t" shows " compact (s \ t)" proof (rule compactI) @@ -4523,40 +4496,6 @@ qed qed -lemma compact_continuous_image: - assumes "continuous_on s f" and "compact s" - shows "compact (f ` s)" -using assms (* FIXME: long unstructured proof *) -unfolding continuous_on_open -unfolding compact_eq_openin_cover -apply clarify -apply (drule_tac x="image (\t. {x \ s. f x \ t}) C" in spec) -apply (drule mp) -apply (rule conjI) -apply simp -apply clarsimp -apply (drule subsetD) -apply (erule imageI) -apply fast -apply (erule thin_rl) -apply clarify -apply (rule_tac x="image (inv_into C (\t. {x \ s. f x \ t})) D" in exI) -apply (intro conjI) -apply clarify -apply (rule inv_into_into) -apply (erule (1) subsetD) -apply (erule finite_imageI) -apply (clarsimp, rename_tac x) -apply (drule (1) subsetD, clarify) -apply (drule (1) subsetD, clarify) -apply (rule rev_bexI) -apply assumption -apply (subgoal_tac "{x \ s. f x \ t} \ (\t. {x \ s. f x \ t}) ` C") -apply (drule f_inv_into_f) -apply fast -apply (erule imageI) -done - lemma connected_continuous_image: assumes "continuous_on s f" "connected s" shows "connected(f ` s)" @@ -4710,56 +4649,6 @@ shows "continuous_on s f \ (\x \ s. \e>0. \d>0. (\x' \ s. norm(x' - x) < d --> abs(f x' - f x) < e))" unfolding continuous_on_iff dist_norm by simp -lemma compact_attains_sup: - fixes S :: "'a::linorder_topology set" - assumes "compact S" "S \ {}" - shows "\s\S. \t\S. t \ s" -proof (rule classical) - assume "\ (\s\S. \t\S. t \ s)" - then obtain t where t: "\s\S. t s \ S" and "\s\S. s < t s" - by (metis not_le) - then have "\s\S. open {..< t s}" "S \ (\s\S. {..< t s})" - by auto - with `compact S` obtain C where "C \ S" "finite C" and C: "S \ (\s\C. {..< t s})" - by (erule compactE_image) - with `S \ {}` have Max: "Max (t`C) \ t`C" and "\s\t`C. s \ Max (t`C)" - by (auto intro!: Max_in) - with C have "S \ {..< Max (t`C)}" - by (auto intro: less_le_trans simp: subset_eq) - with t Max `C \ S` show ?thesis - by fastforce -qed - -lemma compact_attains_inf: - fixes S :: "'a::linorder_topology set" - assumes "compact S" "S \ {}" - shows "\s\S. \t\S. s \ t" -proof (rule classical) - assume "\ (\s\S. \t\S. s \ t)" - then obtain t where t: "\s\S. t s \ S" and "\s\S. t s < s" - by (metis not_le) - then have "\s\S. open {t s <..}" "S \ (\s\S. {t s <..})" - by auto - with `compact S` obtain C where "C \ S" "finite C" and C: "S \ (\s\C. {t s <..})" - by (erule compactE_image) - with `S \ {}` have Min: "Min (t`C) \ t`C" and "\s\t`C. Min (t`C) \ s" - by (auto intro!: Min_in) - with C have "S \ {Min (t`C) <..}" - by (auto intro: le_less_trans simp: subset_eq) - with t Min `C \ S` show ?thesis - by fastforce -qed - -lemma continuous_attains_sup: - fixes f :: "'a::topological_space \ 'b::linorder_topology" - shows "compact s \ s \ {} \ continuous_on s f \ (\x\s. \y\s. f y \ f x)" - using compact_attains_sup[of "f ` s"] compact_continuous_image[of s f] by auto - -lemma continuous_attains_inf: - fixes f :: "'a::topological_space \ 'b::linorder_topology" - shows "compact s \ s \ {} \ continuous_on s f \ (\x\s. \y\s. f x \ f y)" - using compact_attains_inf[of "f ` s"] compact_continuous_image[of s f] by auto - text {* Hence some handy theorems on distance, diameter etc. of/from a set. *} lemma distance_attains_sup: diff -r 270b21f3ae0a -r 33db4b7189af src/HOL/Topological_Spaces.thy --- a/src/HOL/Topological_Spaces.thy Fri Mar 22 10:41:43 2013 +0100 +++ b/src/HOL/Topological_Spaces.thy Fri Mar 22 10:41:43 2013 +0100 @@ -1780,5 +1780,98 @@ "isCont g (f x) \ continuous (at x within s) f \ continuous (at x within s) (\x. g (f x))" using continuous_within_compose2[of x s f g] by (simp add: continuous_at_within) +subsubsection{* Open-cover compactness *} + +context topological_space +begin + +definition compact :: "'a set \ bool" where + compact_eq_heine_borel: -- "This name is used for backwards compatibility" + "compact S \ (\C. (\c\C. open c) \ S \ \C \ (\D\C. finite D \ S \ \D))" + +lemma compactI: + assumes "\C. \t\C. open t \ s \ \ C \ \C'. C' \ C \ finite C' \ s \ \ C'" + shows "compact s" + unfolding compact_eq_heine_borel using assms by metis + +lemma compact_empty[simp]: "compact {}" + by (auto intro!: compactI) + +lemma compactE: + assumes "compact s" and "\t\C. open t" and "s \ \C" + obtains C' where "C' \ C" and "finite C'" and "s \ \C'" + using assms unfolding compact_eq_heine_borel by metis + +lemma compactE_image: + assumes "compact s" and "\t\C. open (f t)" and "s \ (\c\C. f c)" + obtains C' where "C' \ C" and "finite C'" and "s \ (\c\C'. f c)" + using assms unfolding ball_simps[symmetric] SUP_def + by (metis (lifting) finite_subset_image compact_eq_heine_borel[of s]) + end +lemma (in linorder_topology) compact_attains_sup: + assumes "compact S" "S \ {}" + shows "\s\S. \t\S. t \ s" +proof (rule classical) + assume "\ (\s\S. \t\S. t \ s)" + then obtain t where t: "\s\S. t s \ S" and "\s\S. s < t s" + by (metis not_le) + then have "\s\S. open {..< t s}" "S \ (\s\S. {..< t s})" + by auto + with `compact S` obtain C where "C \ S" "finite C" and C: "S \ (\s\C. {..< t s})" + by (erule compactE_image) + with `S \ {}` have Max: "Max (t`C) \ t`C" and "\s\t`C. s \ Max (t`C)" + by (auto intro!: Max_in) + with C have "S \ {..< Max (t`C)}" + by (auto intro: less_le_trans simp: subset_eq) + with t Max `C \ S` show ?thesis + by fastforce +qed + +lemma (in linorder_topology) compact_attains_inf: + assumes "compact S" "S \ {}" + shows "\s\S. \t\S. s \ t" +proof (rule classical) + assume "\ (\s\S. \t\S. s \ t)" + then obtain t where t: "\s\S. t s \ S" and "\s\S. t s < s" + by (metis not_le) + then have "\s\S. open {t s <..}" "S \ (\s\S. {t s <..})" + by auto + with `compact S` obtain C where "C \ S" "finite C" and C: "S \ (\s\C. {t s <..})" + by (erule compactE_image) + with `S \ {}` have Min: "Min (t`C) \ t`C" and "\s\t`C. Min (t`C) \ s" + by (auto intro!: Min_in) + with C have "S \ {Min (t`C) <..}" + by (auto intro: le_less_trans simp: subset_eq) + with t Min `C \ S` show ?thesis + by fastforce +qed + +lemma compact_continuous_image: + assumes f: "continuous_on s f" and s: "compact s" + shows "compact (f ` s)" +proof (rule compactI) + fix C assume "\c\C. open c" and cover: "f`s \ \C" + with f have "\c\C. \A. open A \ A \ s = f -` c \ s" + unfolding continuous_on_open_invariant by blast + then guess A unfolding bchoice_iff .. note A = this + with cover have "\c\C. open (A c)" "s \ (\c\C. A c)" + by (fastforce simp add: subset_eq set_eq_iff)+ + from compactE_image[OF s this] obtain D where "D \ C" "finite D" "s \ (\c\D. A c)" . + with A show "\D \ C. finite D \ f`s \ \D" + by (intro exI[of _ D]) (fastforce simp add: subset_eq set_eq_iff)+ +qed + +lemma continuous_attains_sup: + fixes f :: "'a::topological_space \ 'b::linorder_topology" + shows "compact s \ s \ {} \ continuous_on s f \ (\x\s. \y\s. f y \ f x)" + using compact_attains_sup[of "f ` s"] compact_continuous_image[of s f] by auto + +lemma continuous_attains_inf: + fixes f :: "'a::topological_space \ 'b::linorder_topology" + shows "compact s \ s \ {} \ continuous_on s f \ (\x\s. \y\s. f x \ f y)" + using compact_attains_inf[of "f ` s"] compact_continuous_image[of s f] by auto + +end +