# HG changeset patch # User hoelzl # Date 1358184636 -3600 # Node ID 2b21b4e2d7cb2af6dbbc9c9a1af3011057dac282 # Parent 1421884baf5b871cc9914a7e147a3960bc1f8839 differentiate (cover) compactness and sequential compactness diff -r 1421884baf5b -r 2b21b4e2d7cb src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Mon Jan 14 17:53:37 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Mon Jan 14 18:30:36 2013 +0100 @@ -35,7 +35,8 @@ qed (auto intro!: continuous_intros) lemma brouwer_compactness_lemma: - assumes "compact s" "continuous_on s f" "\ (\x\s. (f x = (0::_::euclidean_space)))" + fixes f :: "'a::metric_space \ 'b::euclidean_space" + assumes "compact s" "continuous_on s f" "\ (\x\s. (f x = 0))" obtains d where "0 < d" "\x\s. d \ norm(f x)" proof (cases "s={}") case False diff -r 1421884baf5b -r 2b21b4e2d7cb src/HOL/Multivariate_Analysis/Path_Connected.thy --- a/src/HOL/Multivariate_Analysis/Path_Connected.thy Mon Jan 14 17:53:37 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy Mon Jan 14 18:30:36 2013 +0100 @@ -57,7 +57,7 @@ apply (rule convex_connected, rule convex_real_interval) done -lemma compact_path_image[intro]: "path g \ compact(path_image g)" +lemma compact_path_image[intro]: "path g \ compact(path_image g :: 'a::metric_space set)" unfolding path_def path_image_def by (erule compact_continuous_image, rule compact_interval) diff -r 1421884baf5b -r 2b21b4e2d7cb src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Jan 14 17:53:37 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Jan 14 18:30:36 2013 +0100 @@ -2344,26 +2344,638 @@ apply (blast intro!: order_antisym dest!: isGlb_le_isLb) done - -subsection {* Equivalent versions of compactness *} +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))" + +subsubsection {* Bolzano-Weierstrass property *} + +lemma heine_borel_imp_bolzano_weierstrass: + assumes "compact s" "infinite t" "t \ s" + shows "\x \ s. x islimpt t" +proof(rule ccontr) + assume "\ (\x \ s. x islimpt t)" + then obtain f where f:"\x\s. x \ f x \ open (f x) \ (\y\t. y \ f x \ y = x)" unfolding islimpt_def + using bchoice[of s "\ x T. x \ T \ open T \ (\y\t. y \ T \ y = x)"] by auto + obtain g where g:"g\{t. \x. x \ s \ t = f x}" "finite g" "s \ \g" + using assms(1)[unfolded compact_eq_heine_borel, THEN spec[where x="{t. \x. x\s \ t = f x}"]] using f by auto + from g(1,3) have g':"\x\g. \xa \ s. x = f xa" by auto + { fix x y assume "x\t" "y\t" "f x = f y" + hence "x \ f x" "y \ f x \ y = x" using f[THEN bspec[where x=x]] and `t\s` by auto + hence "x = y" using `f x = f y` and f[THEN bspec[where x=y]] and `y\t` and `t\s` by auto } + hence "inj_on f t" unfolding inj_on_def by simp + hence "infinite (f ` t)" using assms(2) using finite_imageD by auto + moreover + { fix x assume "x\t" "f x \ g" + from g(3) assms(3) `x\t` obtain h where "h\g" and "x\h" by auto + then obtain y where "y\s" "h = f y" using g'[THEN bspec[where x=h]] by auto + hence "y = x" using f[THEN bspec[where x=y]] and `x\t` and `x\h`[unfolded `h = f y`] by auto + hence False using `f x \ g` `h\g` unfolding `h = f y` by auto } + hence "f ` t \ g" by auto + ultimately show False using g(2) using finite_subset by auto +qed + +lemma islimpt_range_imp_convergent_subsequence: + fixes l :: "'a :: {t1_space, first_countable_topology}" + assumes l: "l islimpt (range f)" + shows "\r. subseq r \ ((f \ r) ---> l) sequentially" +proof - + from first_countable_topology_class.countable_basis_at_decseq[of l] guess A . note A = this + + def s \ "\n i. SOME j. i < j \ f j \ A (Suc n)" + { fix n i + have "\a. i < a \ f a \ A (Suc n) - (f ` {.. i} - {l})" (is "\a. _ \ _ \ ?A") + apply (rule l[THEN islimptE, of "?A"]) + using A(2) apply fastforce + using A(1) + apply (intro open_Diff finite_imp_closed) + apply auto + apply (rule_tac x=x in exI) + apply auto + done + then have "\a. i < a \ f a \ A (Suc n)" by blast + then have "i < s n i" "f (s n i) \ A (Suc n)" + unfolding s_def by (auto intro: someI2_ex) } + note s = this + def r \ "nat_rec (s 0 0) s" + have "subseq r" + by (auto simp: r_def s subseq_Suc_iff) + moreover + have "(\n. f (r n)) ----> l" + proof (rule topological_tendstoI) + fix S assume "open S" "l \ S" + with A(3) have "eventually (\i. A i \ S) sequentially" by auto + moreover + { fix i assume "Suc 0 \ i" then have "f (r i) \ A i" + by (cases i) (simp_all add: r_def s) } + then have "eventually (\i. f (r i) \ A i) sequentially" by (auto simp: eventually_sequentially) + ultimately show "eventually (\i. f (r i) \ S) sequentially" + by eventually_elim auto + qed + ultimately show "\r. subseq r \ (f \ r) ----> l" + by (auto simp: convergent_def comp_def) +qed + +lemma finite_range_imp_infinite_repeats: + fixes f :: "nat \ 'a" + assumes "finite (range f)" + shows "\k. infinite {n. f n = k}" +proof - + { fix A :: "'a set" assume "finite A" + hence "\f. infinite {n. f n \ A} \ \k. infinite {n. f n = k}" + proof (induct) + case empty thus ?case by simp + next + case (insert x A) + show ?case + proof (cases "finite {n. f n = x}") + case True + with `infinite {n. f n \ insert x A}` + have "infinite {n. f n \ A}" by simp + thus "\k. infinite {n. f n = k}" by (rule insert) + next + case False thus "\k. infinite {n. f n = k}" .. + qed + qed + } note H = this + from assms show "\k. infinite {n. f n = k}" + by (rule H) simp +qed + +lemma sequence_infinite_lemma: + fixes f :: "nat \ 'a::t1_space" + assumes "\n. f n \ l" and "(f ---> l) sequentially" + shows "infinite (range f)" +proof + assume "finite (range f)" + hence "closed (range f)" by (rule finite_imp_closed) + hence "open (- range f)" by (rule open_Compl) + from assms(1) have "l \ - range f" by auto + from assms(2) have "eventually (\n. f n \ - range f) sequentially" + using `open (- range f)` `l \ - range f` by (rule topological_tendstoD) + thus False unfolding eventually_sequentially by auto +qed + +lemma closure_insert: + fixes x :: "'a::t1_space" + shows "closure (insert x s) = insert x (closure s)" +apply (rule closure_unique) +apply (rule insert_mono [OF closure_subset]) +apply (rule closed_insert [OF closed_closure]) +apply (simp add: closure_minimal) +done + +lemma islimpt_insert: + fixes x :: "'a::t1_space" + shows "x islimpt (insert a s) \ x islimpt s" +proof + assume *: "x islimpt (insert a s)" + show "x islimpt s" + proof (rule islimptI) + fix t assume t: "x \ t" "open t" + show "\y\s. y \ t \ y \ x" + proof (cases "x = a") + case True + obtain y where "y \ insert a s" "y \ t" "y \ x" + using * t by (rule islimptE) + with `x = a` show ?thesis by auto + next + case False + with t have t': "x \ t - {a}" "open (t - {a})" + by (simp_all add: open_Diff) + obtain y where "y \ insert a s" "y \ t - {a}" "y \ x" + using * t' by (rule islimptE) + thus ?thesis by auto + qed + qed +next + assume "x islimpt s" thus "x islimpt (insert a s)" + by (rule islimpt_subset) auto +qed + +lemma islimpt_union_finite: + fixes x :: "'a::t1_space" + shows "finite s \ x islimpt (s \ t) \ x islimpt t" +by (induct set: finite, simp_all add: islimpt_insert) + +lemma sequence_unique_limpt: + fixes f :: "nat \ 'a::t2_space" + assumes "(f ---> l) sequentially" and "l' islimpt (range f)" + shows "l' = l" +proof (rule ccontr) + assume "l' \ l" + obtain s t where "open s" "open t" "l' \ s" "l \ t" "s \ t = {}" + using hausdorff [OF `l' \ l`] by auto + have "eventually (\n. f n \ t) sequentially" + using assms(1) `open t` `l \ t` by (rule topological_tendstoD) + then obtain N where "\n\N. f n \ t" + unfolding eventually_sequentially by auto + + have "UNIV = {.. {N..}" by auto + hence "l' islimpt (f ` ({.. {N..}))" using assms(2) by simp + hence "l' islimpt (f ` {.. f ` {N..})" by (simp add: image_Un) + hence "l' islimpt (f ` {N..})" by (simp add: islimpt_union_finite) + then obtain y where "y \ f ` {N..}" "y \ s" "y \ l'" + using `l' \ s` `open s` by (rule islimptE) + then obtain n where "N \ n" "f n \ s" "f n \ l'" by auto + with `\n\N. f n \ t` have "f n \ s \ t" by simp + with `s \ t = {}` show False by simp +qed + +lemma bolzano_weierstrass_imp_closed: + fixes s :: "'a::{first_countable_topology, t2_space} set" + assumes "\t. infinite t \ t \ s --> (\x \ s. x islimpt t)" + shows "closed s" +proof- + { fix x l assume as: "\n::nat. x n \ s" "(x ---> l) sequentially" + hence "l \ s" + proof(cases "\n. x n \ l") + case False thus "l\s" using as(1) by auto + next + case True note cas = this + with as(2) have "infinite (range x)" using sequence_infinite_lemma[of x l] by auto + then obtain l' where "l'\s" "l' islimpt (range x)" using assms[THEN spec[where x="range x"]] as(1) by auto + thus "l\s" using sequence_unique_limpt[of x l l'] using as cas by auto + qed } + thus ?thesis unfolding closed_sequential_limits by fast +qed + +lemma compact_imp_closed: + fixes s :: "'a::{first_countable_topology, t2_space} set" + shows "compact s \ closed s" +proof - + assume "compact s" + hence "\t. infinite t \ t \ s \ (\x \ s. x islimpt t)" + using heine_borel_imp_bolzano_weierstrass[of s] by auto + thus "closed s" + by (rule bolzano_weierstrass_imp_closed) +qed + +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)" + unfolding compact_eq_heine_borel +proof (intro allI impI) + fix f assume *: "Ball f open \ s \ t \ \f" + from * `compact s` obtain s' where "s' \ f \ finite s' \ s \ \s'" + unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) metis + moreover from * `compact t` obtain t' where "t' \ f \ finite t' \ t \ \t'" + unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) metis + ultimately show "\f'\f. finite f' \ s \ t \ \f'" + by (auto intro!: exI[of _ "s' \ t'"]) +qed + +lemma compact_Union [intro]: "finite S \ (\T. T \ S \ compact T) \ compact (\S)" + by (induct set: finite) auto + +lemma compact_UN [intro]: + "finite A \ (\x. x \ A \ compact (B x)) \ compact (\x\A. B x)" + unfolding SUP_def by (rule compact_Union) auto + +lemma compact_inter_closed [intro]: + assumes "compact s" and "closed t" + shows "compact (s \ t)" + unfolding compact_eq_heine_borel +proof safe + fix C assume C: "\c\C. open c" and cover: "s \ t \ \C" + from C `closed t` have "\c\C \ {-t}. open c" by auto + moreover from cover have "s \ \(C \ {-t})" by auto + ultimately have "\D\C \ {-t}. finite D \ s \ \D" + using `compact s` unfolding compact_eq_heine_borel by auto + then guess D .. + then show "\D\C. finite D \ s \ t \ \D" + by (intro exI[of _ "D - {-t}"]) auto +qed + +lemma closed_inter_compact [intro]: + assumes "closed s" and "compact t" + shows "compact (s \ t)" + using compact_inter_closed [of t s] assms + by (simp add: Int_commute) + +lemma compact_inter [intro]: + fixes s t :: "'a :: {t2_space, first_countable_topology} set" + assumes "compact s" and "compact t" + shows "compact (s \ t)" + using assms by (intro compact_inter_closed compact_imp_closed) + +lemma compact_sing [simp]: "compact {a}" + unfolding compact_eq_heine_borel by auto + +lemma compact_insert [simp]: + assumes "compact s" shows "compact (insert x s)" +proof - + have "compact ({x} \ s)" + using compact_sing assms by (rule compact_union) + thus ?thesis by simp +qed + +lemma finite_imp_compact: + shows "finite s \ compact s" + by (induct set: finite) simp_all + +lemma open_delete: + fixes s :: "'a::t1_space set" + shows "open s \ open (s - {x})" + by (simp add: open_Diff) + +text{* Finite intersection property *} + +lemma inj_setminus: "inj_on uminus (A::'a set set)" + by (auto simp: inj_on_def) + +lemma compact_fip: + "compact U \ + (\A. (\a\A. closed a) \ (\B \ A. finite B \ U \ \B \ {}) \ U \ \A \ {})" + (is "_ \ ?R") +proof (safe intro!: compact_eq_heine_borel[THEN iffD2]) + fix A assume "compact U" and A: "\a\A. closed a" "U \ \A = {}" + and fi: "\B \ A. finite B \ U \ \B \ {}" + from A have "(\a\uminus`A. open a) \ U \ \uminus`A" + by auto + with `compact U` obtain B where "B \ A" "finite (uminus`B)" "U \ \(uminus`B)" + unfolding compact_eq_heine_borel by (metis subset_image_iff) + with fi[THEN spec, of B] show False + by (auto dest: finite_imageD intro: inj_setminus) +next + fix A assume ?R and cover: "\a\A. open a" "U \ \A" + from cover have "U \ \(uminus`A) = {}" "\a\uminus`A. closed a" + by auto + with `?R` obtain B where "B \ A" "finite (uminus`B)" "U \ \uminus`B = {}" + by (metis subset_image_iff) + then show "\T\A. finite T \ U \ \T" + by (auto intro!: exI[of _ B] inj_setminus dest: finite_imageD) +qed + +lemma compact_imp_fip: + "compact s \ \t \ f. closed t \ \f'. finite f' \ f' \ f \ (s \ (\ f') \ {}) \ + s \ (\ f) \ {}" + unfolding compact_fip by auto + +text{*Compactness expressed with filters*} + +definition "filter_from_subbase B = Abs_filter (\P. \X \ B. finite X \ Inf X \ P)" + +lemma eventually_filter_from_subbase: + "eventually P (filter_from_subbase B) \ (\X \ B. finite X \ Inf X \ P)" + (is "_ \ ?R P") + unfolding filter_from_subbase_def +proof (rule eventually_Abs_filter is_filter.intro)+ + show "?R (\x. True)" + by (rule exI[of _ "{}"]) (simp add: le_fun_def) +next + fix P Q assume "?R P" then guess X .. + moreover assume "?R Q" then guess Y .. + ultimately show "?R (\x. P x \ Q x)" + by (intro exI[of _ "X \ Y"]) auto +next + fix P Q + assume "?R P" then guess X .. + moreover assume "\x. P x \ Q x" + ultimately show "?R Q" + by (intro exI[of _ X]) auto +qed + +lemma eventually_filter_from_subbaseI: "P \ B \ eventually P (filter_from_subbase B)" + by (subst eventually_filter_from_subbase) (auto intro!: exI[of _ "{P}"]) + +lemma filter_from_subbase_not_bot: + "\X \ B. finite X \ Inf X \ bot \ filter_from_subbase B \ bot" + unfolding trivial_limit_def eventually_filter_from_subbase by auto + +lemma closure_iff_nhds_not_empty: + "x \ closure X \ (\A. \S\A. open S \ x \ S \ X \ A \ {})" +proof safe + assume x: "x \ closure X" + fix S A assume "open S" "x \ S" "X \ A = {}" "S \ A" + then have "x \ closure (-S)" + by (auto simp: closure_complement subset_eq[symmetric] intro: interiorI) + with x have "x \ closure X - closure (-S)" + by auto + also have "\ \ closure (X \ S)" + using `open S` open_inter_closure_subset[of S X] by (simp add: closed_Compl ac_simps) + finally have "X \ S \ {}" by auto + then show False using `X \ A = {}` `S \ A` by auto +next + assume "\A S. S \ A \ open S \ x \ S \ X \ A \ {}" + from this[THEN spec, of "- X", THEN spec, of "- closure X"] + show "x \ closure X" + by (simp add: closure_subset open_Compl) +qed + +lemma compact_filter: + "compact U \ (\F. F \ bot \ eventually (\x. x \ U) F \ (\x\U. inf (nhds x) F \ bot))" +proof (intro allI iffI impI compact_fip[THEN iffD2] notI) + fix F assume "compact U" and F: "F \ bot" "eventually (\x. x \ U) F" + from F have "U \ {}" + by (auto simp: eventually_False) + + def Z \ "closure ` {A. eventually (\x. x \ A) F}" + then have "\z\Z. closed z" + by auto + moreover + have ev_Z: "\z. z \ Z \ eventually (\x. x \ z) F" + unfolding Z_def by (auto elim: eventually_elim1 intro: set_mp[OF closure_subset]) + have "(\B \ Z. finite B \ U \ \B \ {})" + proof (intro allI impI) + fix B assume "finite B" "B \ Z" + with `finite B` ev_Z have "eventually (\x. \b\B. x \ b) F" + by (auto intro!: eventually_Ball_finite) + with F(2) have "eventually (\x. x \ U \ (\B)) F" + by eventually_elim auto + with F show "U \ \B \ {}" + by (intro notI) (simp add: eventually_False) + qed + ultimately have "U \ \Z \ {}" + using `compact U` unfolding compact_fip by blast + then obtain x where "x \ U" and x: "\z. z \ Z \ x \ z" by auto + + have "\P. eventually P (inf (nhds x) F) \ P \ bot" + unfolding eventually_inf eventually_nhds + proof safe + fix P Q R S + assume "eventually R F" "open S" "x \ S" + with open_inter_closure_eq_empty[of S "{x. R x}"] x[of "closure {x. R x}"] + have "S \ {x. R x} \ {}" by (auto simp: Z_def) + moreover assume "Ball S Q" "\x. Q x \ R x \ bot x" + ultimately show False by (auto simp: set_eq_iff) + qed + with `x \ U` show "\x\U. inf (nhds x) F \ bot" + by (metis eventually_bot) +next + fix A assume A: "\a\A. closed a" "\B\A. finite B \ U \ \B \ {}" "U \ \A = {}" + + def P' \ "(\a (x::'a). x \ a)" + then have inj_P': "\A. inj_on P' A" + by (auto intro!: inj_onI simp: fun_eq_iff) + def F \ "filter_from_subbase (P' ` insert U A)" + have "F \ bot" + unfolding F_def + proof (safe intro!: filter_from_subbase_not_bot) + fix X assume "X \ P' ` insert U A" "finite X" "Inf X = bot" + then obtain B where "B \ insert U A" "finite B" and B: "Inf (P' ` B) = bot" + unfolding subset_image_iff by (auto intro: inj_P' finite_imageD) + with A(2)[THEN spec, of "B - {U}"] have "U \ \(B - {U}) \ {}" by auto + with B show False by (auto simp: P'_def fun_eq_iff) + qed + moreover have "eventually (\x. x \ U) F" + unfolding F_def by (rule eventually_filter_from_subbaseI) (auto simp: P'_def) + moreover assume "\F. F \ bot \ eventually (\x. x \ U) F \ (\x\U. inf (nhds x) F \ bot)" + ultimately obtain x where "x \ U" and x: "inf (nhds x) F \ bot" + by auto + + { fix V assume "V \ A" + then have V: "eventually (\x. x \ V) F" + by (auto simp add: F_def image_iff P'_def intro!: eventually_filter_from_subbaseI) + have "x \ closure V" + unfolding closure_iff_nhds_not_empty + proof (intro impI allI) + fix S A assume "open S" "x \ S" "S \ A" + then have "eventually (\x. x \ A) (nhds x)" by (auto simp: eventually_nhds) + with V have "eventually (\x. x \ V \ A) (inf (nhds x) F)" + by (auto simp: eventually_inf) + with x show "V \ A \ {}" + by (auto simp del: Int_iff simp add: trivial_limit_def) + qed + then have "x \ V" + using `V \ A` A(1) by simp } + with `x\U` have "x \ U \ \A" by auto + with `U \ \A = {}` show False by auto +qed + +lemma countable_compact: + fixes U :: "'a :: second_countable_topology set" + shows "compact U \ + (\A. countable A \ (\a\A. open a) \ U \ \A \ (\T\A. finite T \ U \ \T))" +proof (safe intro!: compact_eq_heine_borel[THEN iffD2]) + fix A assume A: "\a\A. open a" "U \ \A" + assume *: "\A. countable A \ (\a\A. open a) \ U \ \A \ (\T\A. finite T \ U \ \T)" + def B \ "SOME B::'a set set. countable B \ topological_basis B" + then have B: "countable B" "topological_basis B" + by (auto simp: countable_basis is_basis) + + moreover def C \ "{b\B. \a\A. b \ a}" + ultimately have "countable C" "\a\C. open a" + unfolding C_def by (auto simp: topological_basis_open) + moreover + have "\A \ \C" + proof safe + fix x a assume "x \ a" "a \ A" + with topological_basisE[of B a x] B A + obtain b where "x \ b" "b \ B" "b \ a" by metis + with `a \ A` show "x \ \C" unfolding C_def by auto + qed + then have "U \ \C" using `U \ \A` by auto + ultimately obtain T where "T\C" "finite T" "U \ \T" + using * by metis + moreover then have "\t\T. \a\A. t \ a" + by (auto simp: C_def) + then guess f unfolding bchoice_iff Bex_def .. + ultimately show "\T\A. finite T \ U \ \T" + unfolding C_def by (intro exI[of _ "f`T"]) fastforce +qed (auto simp: compact_eq_heine_borel) subsubsection{* Sequential compactness *} definition - compact :: "'a::metric_space set \ bool" where (* TODO: generalize *) - "compact S \ + seq_compact :: "'a::topological_space set \ bool" where + "seq_compact S \ (\f. (\n. f n \ S) \ (\l\S. \r. subseq r \ ((f o r) ---> l) sequentially))" -lemma compactI: +lemma seq_compact_imp_compact: + fixes U :: "'a :: second_countable_topology set" + assumes "seq_compact U" + shows "compact U" + unfolding countable_compact +proof safe + fix A assume A: "\a\A. open a" "U \ \A" "countable A" + have subseq: "\X. range X \ U \ \r x. x \ U \ subseq r \ (X \ r) ----> x" + using `seq_compact U` by (fastforce simp: seq_compact_def subset_eq) + show "\T\A. finite T \ U \ \T" + proof cases + assume "finite A" with A show ?thesis by auto + next + assume "infinite A" + then have "A \ {}" by auto + show ?thesis + proof (rule ccontr) + assume "\ (\T\A. finite T \ U \ \T)" + then have "\T. \x. T \ A \ finite T \ (x \ U - \T)" by auto + then obtain X' where T: "\T. T \ A \ finite T \ X' T \ U - \T" by metis + def X \ "\n. X' (from_nat_into A ` {.. n})" + have X: "\n. X n \ U - (\i\n. from_nat_into A i)" + using `A \ {}` unfolding X_def SUP_def by (intro T) (auto intro: from_nat_into) + then have "range X \ U" by auto + with subseq[of X] obtain r x where "x \ U" and r: "subseq r" "(X \ r) ----> x" by auto + from `x\U` `U \ \A` from_nat_into_surj[OF `countable A`] + obtain n where "x \ from_nat_into A n" by auto + with r(2) A(1) from_nat_into[OF `A \ {}`, of n] + have "eventually (\i. X (r i) \ from_nat_into A n) sequentially" + unfolding tendsto_def by (auto simp: comp_def) + then obtain N where "\i. N \ i \ X (r i) \ from_nat_into A n" + by (auto simp: eventually_sequentially) + moreover from X have "\i. n \ r i \ X (r i) \ from_nat_into A n" + by auto + moreover from `subseq r`[THEN seq_suble, of "max n N"] have "\i. n \ r i \ N \ i" + by (auto intro!: exI[of _ "max n N"]) + ultimately show False + by auto + qed + qed +qed + +lemma compact_imp_seq_compact: + fixes U :: "'a :: first_countable_topology set" + assumes "compact U" shows "seq_compact U" + unfolding seq_compact_def +proof safe + fix X :: "nat \ 'a" assume "\n. X n \ U" + then have "eventually (\x. x \ U) (filtermap X sequentially)" + by (auto simp: eventually_filtermap) + moreover have "filtermap X sequentially \ bot" + by (simp add: trivial_limit_def eventually_filtermap) + ultimately obtain x where "x \ U" and x: "inf (nhds x) (filtermap X sequentially) \ bot" (is "?F \ _") + using `compact U` by (auto simp: compact_filter) + + from countable_basis_at_decseq[of x] guess A . note A = this + def s \ "\n i. SOME j. i < j \ X j \ A (Suc n)" + { fix n i + have "\a. i < a \ X a \ A (Suc n)" + proof (rule ccontr) + assume "\ (\a>i. X a \ A (Suc n))" + then have "\a. Suc i \ a \ X a \ A (Suc n)" by auto + then have "eventually (\x. x \ A (Suc n)) (filtermap X sequentially)" + by (auto simp: eventually_filtermap eventually_sequentially) + moreover have "eventually (\x. x \ A (Suc n)) (nhds x)" + using A(1,2)[of "Suc n"] by (auto simp: eventually_nhds) + ultimately have "eventually (\x. False) ?F" + by (auto simp add: eventually_inf) + with x show False + by (simp add: eventually_False) + qed + then have "i < s n i" "X (s n i) \ A (Suc n)" + unfolding s_def by (auto intro: someI2_ex) } + note s = this + def r \ "nat_rec (s 0 0) s" + have "subseq r" + by (auto simp: r_def s subseq_Suc_iff) + moreover + have "(\n. X (r n)) ----> x" + proof (rule topological_tendstoI) + fix S assume "open S" "x \ S" + with A(3) have "eventually (\i. A i \ S) sequentially" by auto + moreover + { fix i assume "Suc 0 \ i" then have "X (r i) \ A i" + by (cases i) (simp_all add: r_def s) } + then have "eventually (\i. X (r i) \ A i) sequentially" by (auto simp: eventually_sequentially) + ultimately show "eventually (\i. X (r i) \ S) sequentially" + by eventually_elim auto + qed + ultimately show "\x \ U. \r. subseq r \ (X \ r) ----> x" + using `x \ U` by (auto simp: convergent_def comp_def) +qed + +lemma seq_compact_eq_compact: + fixes U :: "'a :: second_countable_topology set" + shows "seq_compact U \ compact U" + using compact_imp_seq_compact seq_compact_imp_compact by blast + +lemma seq_compactI: assumes "\f. \n. f n \ S \ \l\S. \r. subseq r \ ((f o r) ---> l) sequentially" - shows "compact S" - unfolding compact_def using assms by fast - -lemma compactE: - assumes "compact S" "\n. f n \ S" + shows "seq_compact S" + unfolding seq_compact_def using assms by fast + +lemma seq_compactE: + assumes "seq_compact S" "\n. f n \ S" obtains l r where "l \ S" "subseq r" "((f \ r) ---> l) sequentially" - using assms unfolding compact_def by fast + using assms unfolding seq_compact_def by fast + +lemma bolzano_weierstrass_imp_seq_compact: + fixes s :: "'a::{t1_space, first_countable_topology} set" + assumes "\t. infinite t \ t \ s --> (\x \ s. x islimpt t)" + shows "seq_compact s" +proof - + { fix f :: "nat \ 'a" assume f: "\n. f n \ s" + have "\l\s. \r. subseq r \ ((f \ r) ---> l) sequentially" + proof (cases "finite (range f)") + case True + hence "\l. infinite {n. f n = l}" + by (rule finite_range_imp_infinite_repeats) + then obtain l where "infinite {n. f n = l}" .. + hence "\r. subseq r \ (\n. r n \ {n. f n = l})" + by (rule infinite_enumerate) + then obtain r where "subseq r" and fr: "\n. f (r n) = l" by auto + hence "subseq r \ ((f \ r) ---> l) sequentially" + unfolding o_def by (simp add: fr tendsto_const) + hence "\r. subseq r \ ((f \ r) ---> l) sequentially" + by - (rule exI) + from f have "\n. f (r n) \ s" by simp + hence "l \ s" by (simp add: fr) + thus "\l\s. \r. subseq r \ ((f \ r) ---> l) sequentially" + by (rule rev_bexI) fact + next + case False + with f assms have "\x\s. x islimpt (range f)" by auto + then obtain l where "l \ s" "l islimpt (range f)" .. + have "\r. subseq r \ ((f \ r) ---> l) sequentially" + using `l islimpt (range f)` + by (rule islimpt_range_imp_convergent_subsequence) + with `l \ s` show "\l\s. \r. subseq r \ ((f \ r) ---> l) sequentially" .. + qed + } + thus ?thesis unfolding seq_compact_def by auto +qed text {* A metric space (or topological vector space) is said to have the @@ -2375,10 +2987,10 @@ "bounded s \ \n. f n \ s \ \l r. subseq r \ ((f \ r) ---> l) sequentially" -lemma bounded_closed_imp_compact: +lemma bounded_closed_imp_seq_compact: fixes s::"'a::heine_borel set" - assumes "bounded s" and "closed s" shows "compact s" -proof (unfold compact_def, clarify) + assumes "bounded s" and "closed s" shows "seq_compact s" +proof (unfold seq_compact_def, clarify) fix f :: "nat \ 'a" assume f: "\n. f n \ s" obtain l r where r: "subseq r" and l: "((f \ r) ---> l) sequentially" using bounded_imp_convergent_subsequence [OF `bounded s` `\n. f n \ s`] by auto @@ -2660,10 +3272,10 @@ apply(erule_tac x=y in allE) apply(erule_tac x=y in ballE) by auto qed -lemma compact_imp_complete: assumes "compact s" shows "complete s" +lemma seq_compact_imp_complete: assumes "seq_compact s" shows "complete s" proof- { fix f assume as: "(\n::nat. f n \ s)" "Cauchy f" - from as(1) obtain l r where lr: "l\s" "subseq r" "((f \ r) ---> l) sequentially" using assms unfolding compact_def by blast + from as(1) obtain l r where lr: "l\s" "subseq r" "((f \ r) ---> l) sequentially" using assms unfolding seq_compact_def by blast note lr' = subseq_bigger [OF lr(2)] @@ -2685,10 +3297,10 @@ fix f :: "nat \ 'a" assume "Cauchy f" hence "bounded (range f)" by (rule cauchy_imp_bounded) - hence "compact (closure (range f))" - using bounded_closed_imp_compact [of "closure (range f)"] by auto + hence "seq_compact (closure (range f))" + using bounded_closed_imp_seq_compact [of "closure (range f)"] by auto hence "complete (closure (range f))" - by (rule compact_imp_complete) + by (rule seq_compact_imp_complete) moreover have "\n. f n \ closure (range f)" using closure_subset [of "range f"] by auto ultimately have "\l\closure (range f). (f ---> l) sequentially" @@ -2747,8 +3359,8 @@ "helper_1 s e n = (SOME y::'a. y \ s \ (\m (dist (helper_1 s e m) y < e)))" declare helper_1.simps[simp del] -lemma compact_imp_totally_bounded: - assumes "compact s" +lemma seq_compact_imp_totally_bounded: + assumes "seq_compact s" shows "\e>0. \k. finite k \ k \ s \ s \ (\((\x. ball x e) ` k))" proof(rule, rule, rule ccontr) fix e::real assume "e>0" and assm:"\ (\k. finite k \ k \ s \ s \ \(\x. ball x e) ` k)" @@ -2765,7 +3377,7 @@ thus "x n \ s \ (\m dist (x m) (x n) < e)" unfolding Q_def by auto qed } hence "\n::nat. x n \ s" and x:"\n. \m < n. \ (dist (x m) (x n) < e)" by blast+ - then obtain l r where "l\s" and r:"subseq r" and "((x \ r) ---> l) sequentially" using assms(1)[unfolded compact_def, THEN spec[where x=x]] by auto + then obtain l r where "l\s" and r:"subseq r" and "((x \ r) ---> l) sequentially" using assms(1)[unfolded seq_compact_def, THEN spec[where x=x]] by auto from this(3) have "Cauchy (x \ r)" using convergent_imp_cauchy by auto then obtain N::nat where N:"\m n. N \ m \ N \ n \ dist ((x \ r) m) ((x \ r) n) < e" unfolding cauchy_def using `e>0` by auto show False @@ -2779,7 +3391,7 @@ text {* Following Burkill \& Burkill vol. 2. *} lemma heine_borel_lemma: fixes s::"'a::metric_space set" - assumes "compact s" "s \ (\ t)" "\b \ t. open b" + assumes "seq_compact s" "s \ (\ t)" "\b \ t. open b" shows "\e>0. \x \ s. \b \ t. ball x e \ b" proof(rule ccontr) assume "\ (\e>0. \x\s. \b\t. ball x e \ b)" @@ -2792,7 +3404,7 @@ using choice[of "\n::nat. \x. x\s \ (\xa\t. \ ball x (inverse (real (n + 1))) \ xa)"] by auto then obtain l r where l:"l\s" and r:"subseq r" and lr:"((f \ r) ---> l) sequentially" - using assms(1)[unfolded compact_def, THEN spec[where x=f]] by auto + using assms(1)[unfolded seq_compact_def, THEN spec[where x=f]] by auto obtain b where "l\b" "b\t" using assms(2) and l by auto then obtain e where "e>0" and e:"\z. dist z l < e \ z\b" @@ -2818,16 +3430,19 @@ thus False using e and `y\b` by auto qed -lemma compact_imp_heine_borel: "compact s ==> (\f. (\t \ f. open t) \ s \ (\ f) - \ (\f'. f' \ f \ finite f' \ s \ (\ f')))" +lemma seq_compact_imp_heine_borel: + fixes s :: "'a :: metric_space set" + shows "seq_compact s \ compact s" + unfolding compact_eq_heine_borel proof clarify - fix f assume "compact s" " \t\f. open t" "s \ \f" + fix f assume "seq_compact s" " \t\f. open t" "s \ \f" then obtain e::real where "e>0" and "\x\s. \b\f. ball x e \ b" using heine_borel_lemma[of s f] by auto hence "\x\s. \b. b\f \ ball x e \ b" by auto hence "\bb. \x\s. bb x \f \ ball x e \ bb x" using bchoice[of s "\x b. b\f \ ball x e \ b"] by auto then obtain bb where bb:"\x\s. (bb x) \ f \ ball x e \ (bb x)" by blast - from `compact s` have "\ k. finite k \ k \ s \ s \ \(\x. ball x e) ` k" using compact_imp_totally_bounded[of s] `e>0` by auto + from `seq_compact s` have "\ k. finite k \ k \ s \ s \ \(\x. ball x e) ` k" + using seq_compact_imp_totally_bounded[of s] `e>0` by auto then obtain k where k:"finite k" "k \ s" "s \ \(\x. ball x e) ` k" by auto have "finite (bb ` k)" using k(1) by auto @@ -2840,132 +3455,8 @@ ultimately show "\f'\f. finite f' \ s \ \f'" using bb k(2) by (rule_tac x="bb ` k" in exI) auto qed -subsubsection {* Bolzano-Weierstrass property *} - -lemma heine_borel_imp_bolzano_weierstrass: - assumes "\f. (\t \ f. open t) \ s \ (\ f) --> (\f'. f' \ f \ finite f' \ s \ (\ f'))" - "infinite t" "t \ s" - shows "\x \ s. x islimpt t" -proof(rule ccontr) - assume "\ (\x \ s. x islimpt t)" - then obtain f where f:"\x\s. x \ f x \ open (f x) \ (\y\t. y \ f x \ y = x)" unfolding islimpt_def - using bchoice[of s "\ x T. x \ T \ open T \ (\y\t. y \ T \ y = x)"] by auto - obtain g where g:"g\{t. \x. x \ s \ t = f x}" "finite g" "s \ \g" - using assms(1)[THEN spec[where x="{t. \x. x\s \ t = f x}"]] using f by auto - from g(1,3) have g':"\x\g. \xa \ s. x = f xa" by auto - { fix x y assume "x\t" "y\t" "f x = f y" - hence "x \ f x" "y \ f x \ y = x" using f[THEN bspec[where x=x]] and `t\s` by auto - hence "x = y" using `f x = f y` and f[THEN bspec[where x=y]] and `y\t` and `t\s` by auto } - hence "inj_on f t" unfolding inj_on_def by simp - hence "infinite (f ` t)" using assms(2) using finite_imageD by auto - moreover - { fix x assume "x\t" "f x \ g" - from g(3) assms(3) `x\t` obtain h where "h\g" and "x\h" by auto - then obtain y where "y\s" "h = f y" using g'[THEN bspec[where x=h]] by auto - hence "y = x" using f[THEN bspec[where x=y]] and `x\t` and `x\h`[unfolded `h = f y`] by auto - hence False using `f x \ g` `h\g` unfolding `h = f y` by auto } - hence "f ` t \ g" by auto - ultimately show False using g(2) using finite_subset by auto -qed - subsubsection {* Complete the chain of compactness variants *} -lemma islimpt_range_imp_convergent_subsequence: - fixes l :: "'a :: {t1_space, first_countable_topology}" - assumes l: "l islimpt (range f)" - shows "\r. subseq r \ ((f \ r) ---> l) sequentially" -proof - - from first_countable_topology_class.countable_basis_at_decseq[of l] guess A . note A = this - - def s \ "\n i. SOME j. i < j \ f j \ A (Suc n)" - { fix n i - have "\a. i < a \ f a \ A (Suc n)" - by (rule l[THEN islimptE, of "A (Suc n) - (f ` {.. i} - {l})"]) - (auto simp: not_le[symmetric] finite_imp_closed A(1,2)) - then have "i < s n i" "f (s n i) \ A (Suc n)" - unfolding s_def by (auto intro: someI2_ex) } - note s = this - def r \ "nat_rec (s 0 0) s" - have "subseq r" - by (auto simp: r_def s subseq_Suc_iff) - moreover - have "(\n. f (r n)) ----> l" - proof (rule topological_tendstoI) - fix S assume "open S" "l \ S" - with A(3) have "eventually (\i. A i \ S) sequentially" by auto - moreover - { fix i assume "Suc 0 \ i" then have "f (r i) \ A i" - by (cases i) (simp_all add: r_def s) } - then have "eventually (\i. f (r i) \ A i) sequentially" by (auto simp: eventually_sequentially) - ultimately show "eventually (\i. f (r i) \ S) sequentially" - by eventually_elim auto - qed - ultimately show "\r. subseq r \ (f \ r) ----> l" - by (auto simp: convergent_def comp_def) -qed - -lemma finite_range_imp_infinite_repeats: - fixes f :: "nat \ 'a" - assumes "finite (range f)" - shows "\k. infinite {n. f n = k}" -proof - - { fix A :: "'a set" assume "finite A" - hence "\f. infinite {n. f n \ A} \ \k. infinite {n. f n = k}" - proof (induct) - case empty thus ?case by simp - next - case (insert x A) - show ?case - proof (cases "finite {n. f n = x}") - case True - with `infinite {n. f n \ insert x A}` - have "infinite {n. f n \ A}" by simp - thus "\k. infinite {n. f n = k}" by (rule insert) - next - case False thus "\k. infinite {n. f n = k}" .. - qed - qed - } note H = this - from assms show "\k. infinite {n. f n = k}" - by (rule H) simp -qed - -lemma bolzano_weierstrass_imp_compact: - fixes s :: "'a::metric_space set" - assumes "\t. infinite t \ t \ s --> (\x \ s. x islimpt t)" - shows "compact s" -proof - - { fix f :: "nat \ 'a" assume f: "\n. f n \ s" - have "\l\s. \r. subseq r \ ((f \ r) ---> l) sequentially" - proof (cases "finite (range f)") - case True - hence "\l. infinite {n. f n = l}" - by (rule finite_range_imp_infinite_repeats) - then obtain l where "infinite {n. f n = l}" .. - hence "\r. subseq r \ (\n. r n \ {n. f n = l})" - by (rule infinite_enumerate) - then obtain r where "subseq r" and fr: "\n. f (r n) = l" by auto - hence "subseq r \ ((f \ r) ---> l) sequentially" - unfolding o_def by (simp add: fr tendsto_const) - hence "\r. subseq r \ ((f \ r) ---> l) sequentially" - by - (rule exI) - from f have "\n. f (r n) \ s" by simp - hence "l \ s" by (simp add: fr) - thus "\l\s. \r. subseq r \ ((f \ r) ---> l) sequentially" - by (rule rev_bexI) fact - next - case False - with f assms have "\x\s. x islimpt (range f)" by auto - then obtain l where "l \ s" "l islimpt (range f)" .. - have "\r. subseq r \ ((f \ r) ---> l) sequentially" - using `l islimpt (range f)` - by (rule islimpt_range_imp_convergent_subsequence) - with `l \ s` show "\l\s. \r. subseq r \ ((f \ r) ---> l) sequentially" .. - qed - } - thus ?thesis unfolding compact_def by auto -qed - primrec helper_2::"(real \ 'a::metric_space) \ nat \ 'a" where "helper_2 beyond 0 = beyond 0" | "helper_2 beyond (Suc n) = beyond (dist undefined (helper_2 beyond n) + 1 )" @@ -3030,127 +3521,26 @@ show False using y and z and dist_triangle_half_l[of "x y" l 1 "x z"] and **[of y z] by auto qed -lemma sequence_infinite_lemma: - fixes f :: "nat \ 'a::t1_space" - assumes "\n. f n \ l" and "(f ---> l) sequentially" - shows "infinite (range f)" -proof - assume "finite (range f)" - hence "closed (range f)" by (rule finite_imp_closed) - hence "open (- range f)" by (rule open_Compl) - from assms(1) have "l \ - range f" by auto - from assms(2) have "eventually (\n. f n \ - range f) sequentially" - using `open (- range f)` `l \ - range f` by (rule topological_tendstoD) - thus False unfolding eventually_sequentially by auto -qed - -lemma closure_insert: - fixes x :: "'a::t1_space" - shows "closure (insert x s) = insert x (closure s)" -apply (rule closure_unique) -apply (rule insert_mono [OF closure_subset]) -apply (rule closed_insert [OF closed_closure]) -apply (simp add: closure_minimal) -done - -lemma islimpt_insert: - fixes x :: "'a::t1_space" - shows "x islimpt (insert a s) \ x islimpt s" -proof - assume *: "x islimpt (insert a s)" - show "x islimpt s" - proof (rule islimptI) - fix t assume t: "x \ t" "open t" - show "\y\s. y \ t \ y \ x" - proof (cases "x = a") - case True - obtain y where "y \ insert a s" "y \ t" "y \ x" - using * t by (rule islimptE) - with `x = a` show ?thesis by auto - next - case False - with t have t': "x \ t - {a}" "open (t - {a})" - by (simp_all add: open_Diff) - obtain y where "y \ insert a s" "y \ t - {a}" "y \ x" - using * t' by (rule islimptE) - thus ?thesis by auto - qed - qed -next - assume "x islimpt s" thus "x islimpt (insert a s)" - by (rule islimpt_subset) auto -qed - -lemma islimpt_union_finite: - fixes x :: "'a::t1_space" - shows "finite s \ x islimpt (s \ t) \ x islimpt t" -by (induct set: finite, simp_all add: islimpt_insert) - -lemma sequence_unique_limpt: - fixes f :: "nat \ 'a::t2_space" - assumes "(f ---> l) sequentially" and "l' islimpt (range f)" - shows "l' = l" -proof (rule ccontr) - assume "l' \ l" - obtain s t where "open s" "open t" "l' \ s" "l \ t" "s \ t = {}" - using hausdorff [OF `l' \ l`] by auto - have "eventually (\n. f n \ t) sequentially" - using assms(1) `open t` `l \ t` by (rule topological_tendstoD) - then obtain N where "\n\N. f n \ t" - unfolding eventually_sequentially by auto - - have "UNIV = {.. {N..}" by auto - hence "l' islimpt (f ` ({.. {N..}))" using assms(2) by simp - hence "l' islimpt (f ` {.. f ` {N..})" by (simp add: image_Un) - hence "l' islimpt (f ` {N..})" by (simp add: islimpt_union_finite) - then obtain y where "y \ f ` {N..}" "y \ s" "y \ l'" - using `l' \ s` `open s` by (rule islimptE) - then obtain n where "N \ n" "f n \ s" "f n \ l'" by auto - with `\n\N. f n \ t` have "f n \ s \ t" by simp - with `s \ t = {}` show False by simp -qed - -lemma bolzano_weierstrass_imp_closed: - fixes s :: "'a::metric_space set" (* TODO: can this be generalized? *) - assumes "\t. infinite t \ t \ s --> (\x \ s. x islimpt t)" - shows "closed s" -proof- - { fix x l assume as: "\n::nat. x n \ s" "(x ---> l) sequentially" - hence "l \ s" - proof(cases "\n. x n \ l") - case False thus "l\s" using as(1) by auto - next - case True note cas = this - with as(2) have "infinite (range x)" using sequence_infinite_lemma[of x l] by auto - then obtain l' where "l'\s" "l' islimpt (range x)" using assms[THEN spec[where x="range x"]] as(1) by auto - thus "l\s" using sequence_unique_limpt[of x l l'] using as cas by auto - qed } - thus ?thesis unfolding closed_sequential_limits by fast -qed - text {* Hence express everything as an equivalence. *} -lemma compact_eq_heine_borel: - fixes s :: "'a::metric_space set" - shows "compact s \ - (\f. (\t \ f. open t) \ s \ (\ f) - --> (\f'. f' \ f \ finite f' \ s \ (\ f')))" (is "?lhs = ?rhs") -proof - assume ?lhs thus ?rhs by (rule compact_imp_heine_borel) -next - assume ?rhs - hence "\t. infinite t \ t \ s \ (\x\s. x islimpt t)" - by (blast intro: heine_borel_imp_bolzano_weierstrass[of s]) - thus ?lhs by (rule bolzano_weierstrass_imp_compact) -qed +lemma compact_eq_seq_compact_metric: + "compact (s :: 'a::metric_space set) \ seq_compact s" + using compact_imp_seq_compact seq_compact_imp_heine_borel by blast + +lemma compact_def: + "compact (S :: 'a::metric_space set) \ + (\f. (\n. f n \ S) \ + (\l\S. \r. subseq r \ ((f o r) ---> l) sequentially)) " + unfolding compact_eq_seq_compact_metric seq_compact_def by auto lemma compact_eq_bolzano_weierstrass: fixes s :: "'a::metric_space set" shows "compact s \ (\t. infinite t \ t \ s --> (\x \ s. x islimpt t))" (is "?lhs = ?rhs") proof - assume ?lhs thus ?rhs unfolding compact_eq_heine_borel using heine_borel_imp_bolzano_weierstrass[of s] by auto + assume ?lhs thus ?rhs using heine_borel_imp_bolzano_weierstrass[of s] by auto next - assume ?rhs thus ?lhs by (rule bolzano_weierstrass_imp_compact) + assume ?rhs thus ?lhs + unfolding compact_eq_seq_compact_metric by (rule bolzano_weierstrass_imp_seq_compact) qed lemma nat_approx_posE: @@ -3167,8 +3557,8 @@ qed lemma compact_eq_totally_bounded: - shows "compact s \ complete s \ (\e>0. \k. finite k \ s \ (\((\x. ball x e) ` k)))" -proof (safe intro!: compact_imp_complete) + "compact s \ complete s \ (\e>0. \k. finite k \ s \ (\((\x. ball x e) ` k)))" +proof (safe intro!: seq_compact_imp_complete[unfolded compact_eq_seq_compact_metric[symmetric]]) fix e::real def f \ "(\x::'a. ball x e) ` UNIV" assume "0 < e" "compact s" @@ -3305,124 +3695,24 @@ fixes s :: "'a::heine_borel set" shows "compact s \ bounded s \ closed s" (is "?lhs = ?rhs") proof - assume ?lhs thus ?rhs unfolding compact_eq_bolzano_weierstrass using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed by auto + assume ?lhs thus ?rhs + unfolding compact_eq_bolzano_weierstrass using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed by auto next - assume ?rhs thus ?lhs using bounded_closed_imp_compact by auto + assume ?rhs thus ?lhs + using bounded_closed_imp_seq_compact[of s] unfolding compact_eq_seq_compact_metric by auto qed lemma compact_imp_bounded: fixes s :: "'a::metric_space set" - shows "compact s ==> bounded s" + shows "compact s \ bounded s" proof - assume "compact s" - hence "\f. (\t\f. open t) \ s \ \f \ (\f'\f. finite f' \ s \ \f')" - by (rule compact_imp_heine_borel) hence "\t. infinite t \ t \ s \ (\x \ s. x islimpt t)" using heine_borel_imp_bolzano_weierstrass[of s] by auto thus "bounded s" by (rule bolzano_weierstrass_imp_bounded) qed -lemma compact_imp_closed: - fixes s :: "'a::metric_space set" - shows "compact s ==> closed s" -proof - - assume "compact s" - hence "\f. (\t\f. open t) \ s \ \f \ (\f'\f. finite f' \ s \ \f')" - by (rule compact_imp_heine_borel) - hence "\t. infinite t \ t \ s \ (\x \ s. x islimpt t)" - using heine_borel_imp_bolzano_weierstrass[of s] by auto - thus "closed s" - by (rule bolzano_weierstrass_imp_closed) -qed - -text{* In particular, some common special cases. *} - -lemma compact_empty[simp]: - "compact {}" - unfolding compact_def - by simp - -lemma compact_union [intro]: - assumes "compact s" and "compact t" - shows "compact (s \ t)" -proof (rule compactI) - fix f :: "nat \ 'a" - assume "\n. f n \ s \ t" - hence "infinite {n. f n \ s \ t}" by simp - hence "infinite {n. f n \ s} \ infinite {n. f n \ t}" by simp - thus "\l\s \ t. \r. subseq r \ ((f \ r) ---> l) sequentially" - proof - assume "infinite {n. f n \ s}" - from infinite_enumerate [OF this] - obtain q where "subseq q" "\n. (f \ q) n \ s" by auto - obtain r l where "l \ s" "subseq r" "((f \ q \ r) ---> l) sequentially" - using `compact s` `\n. (f \ q) n \ s` by (rule compactE) - hence "l \ s \ t" "subseq (q \ r)" "((f \ (q \ r)) ---> l) sequentially" - using `subseq q` by (simp_all add: subseq_o o_assoc) - thus ?thesis by auto - next - assume "infinite {n. f n \ t}" - from infinite_enumerate [OF this] - obtain q where "subseq q" "\n. (f \ q) n \ t" by auto - obtain r l where "l \ t" "subseq r" "((f \ q \ r) ---> l) sequentially" - using `compact t` `\n. (f \ q) n \ t` by (rule compactE) - hence "l \ s \ t" "subseq (q \ r)" "((f \ (q \ r)) ---> l) sequentially" - using `subseq q` by (simp_all add: subseq_o o_assoc) - thus ?thesis by auto - qed -qed - -lemma compact_Union [intro]: "finite S \ (\T. T \ S \ compact T) \ compact (\S)" - by (induct set: finite) auto - -lemma compact_UN [intro]: - "finite A \ (\x. x \ A \ compact (B x)) \ compact (\x\A. B x)" - unfolding SUP_def by (rule compact_Union) auto - -lemma compact_inter_closed [intro]: - assumes "compact s" and "closed t" - shows "compact (s \ t)" -proof (rule compactI) - fix f :: "nat \ 'a" - assume "\n. f n \ s \ t" - hence "\n. f n \ s" and "\n. f n \ t" by simp_all - obtain l r where "l \ s" "subseq r" "((f \ r) ---> l) sequentially" - using `compact s` `\n. f n \ s` by (rule compactE) - moreover - from `closed t` `\n. f n \ t` `((f \ r) ---> l) sequentially` have "l \ t" - unfolding closed_sequential_limits o_def by fast - ultimately show "\l\s \ t. \r. subseq r \ ((f \ r) ---> l) sequentially" - by auto -qed - -lemma closed_inter_compact [intro]: - assumes "closed s" and "compact t" - shows "compact (s \ t)" - using compact_inter_closed [of t s] assms - by (simp add: Int_commute) - -lemma compact_inter [intro]: - assumes "compact s" and "compact t" - shows "compact (s \ t)" - using assms by (intro compact_inter_closed compact_imp_closed) - -lemma compact_sing [simp]: "compact {a}" - unfolding compact_def o_def subseq_def - by (auto simp add: tendsto_const) - -lemma compact_insert [simp]: - assumes "compact s" shows "compact (insert x s)" -proof - - have "compact ({x} \ s)" - using compact_sing assms by (rule compact_union) - thus ?thesis by simp -qed - -lemma finite_imp_compact: - shows "finite s \ compact s" - by (induct set: finite) simp_all - lemma compact_cball[simp]: fixes x :: "'a::heine_borel" shows "compact(cball x e)" @@ -3448,28 +3738,6 @@ using frontier_subset_closed compact_eq_bounded_closed by blast -lemma open_delete: - fixes s :: "'a::t1_space set" - shows "open s \ open (s - {x})" - by (simp add: open_Diff) - -text{* Finite intersection property. I could make it an equivalence in fact. *} - -lemma compact_imp_fip: - assumes "compact s" "\t \ f. closed t" - "\f'. finite f' \ f' \ f --> (s \ (\ f') \ {})" - shows "s \ (\ f) \ {}" -proof - assume as:"s \ (\ f) = {}" - hence "s \ \ uminus ` f" by auto - moreover have "Ball (uminus ` f) open" using open_Diff closed_Diff using assms(2) by auto - ultimately obtain f' where f':"f' \ uminus ` f" "finite f'" "s \ \f'" using assms(1)[unfolded compact_eq_heine_borel, THEN spec[where x="(\t. - t) ` f"]] by auto - hence "finite (uminus ` f') \ uminus ` f' \ f" by(auto simp add: Diff_Diff_Int) - hence "s \ \uminus ` f' \ {}" using assms(3)[THEN spec[where x="uminus ` f'"]] by auto - thus False using f'(3) unfolding subset_eq and Union_iff by blast -qed - - subsection {* Bounded closed nest property (proof does not use Heine-Borel) *} lemma bounded_closed_nest: @@ -3478,10 +3746,10 @@ shows "\a::'a::heine_borel. \n::nat. a \ s(n)" proof- from assms(2) obtain x where x:"\n::nat. x n \ s n" using choice[of "\n x. x\ s n"] by auto - from assms(4,1) have *:"compact (s 0)" using bounded_closed_imp_compact[of "s 0"] by auto + from assms(4,1) have *:"seq_compact (s 0)" using bounded_closed_imp_seq_compact[of "s 0"] by auto then obtain l r where lr:"l\s 0" "subseq r" "((x \ r) ---> l) sequentially" - unfolding compact_def apply(erule_tac x=x in allE) using x using assms(3) by blast + unfolding seq_compact_def apply(erule_tac x=x in allE) using x using assms(3) by blast { fix n::nat { fix e::real assume "e>0" @@ -3492,7 +3760,7 @@ hence "(x \ r) (max N n) \ s n" using x apply(erule_tac x=n in allE) using x apply(erule_tac x="r (max N n)" in allE) - using assms(3) apply(erule_tac x=n in allE)apply( erule_tac x="r (max N n)" in allE) by auto + using assms(3) apply(erule_tac x=n in allE) apply(erule_tac x="r (max N n)" in allE) by auto ultimately have "\y\s n. dist y l < e" by auto } hence "l \ s n" using closed_approachable[of "s n" l] assms(1) by blast @@ -4531,6 +4799,7 @@ text {* Preservation of compactness and connectedness under continuous function. *} lemma compact_continuous_image: + fixes f :: "'a :: metric_space \ 'b :: metric_space" assumes "continuous_on s f" "compact s" shows "compact(f ` s)" proof- @@ -4538,7 +4807,8 @@ then obtain y where y:"\n. y n \ s \ x n = f (y n)" unfolding image_iff Bex_def using choice[of "\n xa. xa \ s \ x n = f xa"] by auto then obtain l r where "l\s" and r:"subseq r" and lr:"((y \ r) ---> l) sequentially" using assms(2)[unfolded compact_def, THEN spec[where x=y]] by auto { fix e::real assume "e>0" - then obtain d where "d>0" and d:"\x'\s. dist x' l < d \ dist (f x') (f l) < e" using assms(1)[unfolded continuous_on_iff, THEN bspec[where x=l], OF `l\s`] by auto + then obtain d where "d>0" and d:"\x'\s. dist x' l < d \ dist (f x') (f l) < e" + using assms(1)[unfolded continuous_on_iff, THEN bspec[where x=l], OF `l\s`] by auto then obtain N::nat where N:"\n\N. dist ((y \ r) n) l < d" using lr[unfolded LIMSEQ_def, THEN spec[where x=d]] by auto { fix n::nat assume "n\N" hence "dist ((x \ r) n) (f l) < e" using N[THEN spec[where x=n]] d[THEN bspec[where x="y (r n)"]] y[THEN spec[where x="r n"]] by auto } hence "\N. \n\N. dist ((x \ r) n) (f l) < e" by auto } @@ -4579,7 +4849,8 @@ hence "s \ \{ball x (d x (e / 2)) |x. x \ s}" unfolding subset_eq by auto moreover { fix b assume "b\{ball x (d x (e / 2)) |x. x \ s}" hence "open b" by auto } - ultimately obtain ea where "ea>0" and ea:"\x\s. \b\{ball x (d x (e / 2)) |x. x \ s}. ball x ea \ b" using heine_borel_lemma[OF assms(2), of "{ball x (d x (e / 2)) | x. x\s }"] by auto + ultimately obtain ea where "ea>0" and ea:"\x\s. \b\{ball x (d x (e / 2)) |x. x \ s}. ball x ea \ b" + using heine_borel_lemma[OF assms(2)[unfolded compact_eq_seq_compact_metric], of "{ball x (d x (e / 2)) | x. x\s }"] by auto { fix x y assume "x\s" "y\s" and as:"dist y x < ea" obtain z where "z\s" and z:"ball x ea \ ball z (d z (e / 2))" using ea[THEN bspec[where x=x]] and `x\s` by auto @@ -4806,8 +5077,8 @@ lemma mem_Times_iff: "x \ A \ B \ fst x \ A \ snd x \ B" by (induct x) simp -lemma compact_Times: "compact s \ compact t \ compact (s \ t)" -unfolding compact_def +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) @@ -4823,6 +5094,12 @@ apply (simp add: o_def) done +text {* Generalize to @{class topological_space} *} +lemma compact_Times: + fixes s :: "'a::metric_space set" and t :: "'b::metric_space set" + shows "compact s \ compact t \ compact (s \ t)" + unfolding compact_eq_seq_compact_metric by (rule seq_compact_Times) + text{* Hence some useful properties follow quite easily. *} lemma compact_scaling: @@ -5342,8 +5619,8 @@ using bounded_interval[of a b] by auto lemma compact_interval: fixes a :: "'a::ordered_euclidean_space" shows "compact {a .. b}" - using bounded_closed_imp_compact[of "{a..b}"] using bounded_interval[of a b] - by auto + using bounded_closed_imp_seq_compact[of "{a..b}"] using bounded_interval[of a b] + by (auto simp: compact_eq_seq_compact_metric) lemma open_interval_midpoint: fixes a :: "'a::ordered_euclidean_space" assumes "{a<.. {}" shows "((1/2) *\<^sub>R (a + b)) \ {a<..n\m. f n \ S" obtains l r where "l \ S" "subseq r" "((f \ r) ---> l) sequentially" proof atomize_elim have "subseq (op + m)" by (simp add: subseq_def) have "\n. (f o (\i. m + i)) n \ S" using assms by auto - from compactE[OF `compact S` this] guess l r . + from seq_compactE[OF `compact S`[unfolded compact_eq_seq_compact_metric] this] guess l r . hence "l \ S" "subseq ((\i. m + i) o r) \ (f \ ((\i. m + i) o r)) ----> l" using subseq_o[OF `subseq (op + m)` `subseq r`] by (auto simp: o_def) thus "\l r. l \ S \ subseq r \ (f \ r) ----> l" by blast