# HG changeset patch # User hoelzl # Date 1358425217 -3600 # Node ID 3690724028b182daf3667b69c61446c2633ec17e # Parent a7c273a83d2783d550e65eaf15f310af002483ef add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite diff -r a7c273a83d27 -r 3690724028b1 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Thu Jan 17 12:26:54 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Thu Jan 17 13:20:17 2013 +0100 @@ -2414,32 +2414,6 @@ 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" @@ -2808,37 +2782,60 @@ with `U \ \A = {}` show False by auto qed -lemma countable_compact: - fixes U :: "'a :: second_countable_topology set" - shows "compact U \ +definition "countably_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]) + +lemma countably_compactE: + assumes "countably_compact s" and "\t\C. open t" and "s \ \C" "countable C" + obtains C' where "C' \ C" and "finite C'" and "s \ \C'" + using assms unfolding countably_compact_def by metis + +lemma countably_compactI: + assumes "\C. \t\C. open t \ s \ \C \ countable C \ (\C'\C. finite C' \ s \ \C')" + shows "countably_compact s" + using assms unfolding countably_compact_def by metis + +lemma compact_imp_countably_compact: "compact U \ countably_compact U" + by (auto simp: compact_eq_heine_borel countably_compact_def) + +lemma countably_compact_imp_compact: + assumes "countably_compact U" + assumes ccover: "countable B" "\b\B. open b" + assumes basis: "\T x. open T \ x \ T \ x \ U \ \b\B. x \ b \ b \ U \ T" + shows "compact U" + using `countably_compact U` unfolding compact_eq_heine_borel countably_compact_def +proof safe 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}" + + moreover def C \ "{b\B. \a\A. b \ U \ a}" ultimately have "countable C" "\a\C. open a" - unfolding C_def by (auto simp: topological_basis_open) + unfolding C_def using ccover by auto moreover - have "\A \ \C" + have "\A \ U \ \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 + fix x a assume "x \ U" "x \ a" "a \ A" + with basis[of a x] A obtain b where "b \ B" "x \ b" "b \ U \ a" by blast + 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" + moreover then have "\t\T. \a\A. t \ U \ 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) +qed + +lemma countably_compact_imp_compact_second_countable: + "countably_compact U \ compact (U :: 'a :: second_countable_topology set)" +proof (rule countably_compact_imp_compact) + fix T and x :: 'a assume "open T" "x \ T" + from topological_basisE[OF is_basis this] guess b . + then show "\b\SOME B. countable B \ topological_basis B. x \ b \ b \ U \ T" by auto +qed (insert countable_basis topological_basis_open[OF is_basis], auto) subsubsection{* Sequential compactness *} @@ -2848,12 +2845,11 @@ (\f. (\n. f n \ S) \ (\l\S. \r. subseq r \ ((f o r) ---> l) sequentially))" -lemma seq_compact_imp_compact: - fixes U :: "'a :: second_countable_topology set" +lemma seq_compact_imp_countably_compact: + fixes U :: "'a :: first_countable_topology set" assumes "seq_compact U" - shows "compact U" - unfolding countable_compact -proof safe + shows "countably_compact U" +proof (safe intro!: countably_compactI) 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) @@ -2944,7 +2940,8 @@ 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 + using compact_imp_seq_compact seq_compact_imp_countably_compact countably_compact_imp_compact_second_countable + by blast lemma seq_compactI: assumes "\f. \n. f n \ S \ \l\S. \r. subseq r \ ((f o r) ---> l) sequentially" @@ -2962,31 +2959,25 @@ 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" + have "\l\s. \r. subseq r \ (f \ r) ----> l" 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 + obtain l where "infinite {n. f n = f l}" + using pigeonhole_infinite[OF _ True] by auto + then obtain r where "subseq r" and fr: "\n. f (r n) = f l" + using infinite_enumerate by blast + hence "subseq r \ (f \ r) ----> f l" + by (simp add: fr tendsto_const o_def) + with f show "\l\s. \r. subseq r \ (f \ r) ----> l" + by auto 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" + have "\r. subseq r \ (f \ r) ----> l" 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" .. + with `l \ s` show "\l\s. \r. subseq r \ (f \ r) ----> l" .. qed } thus ?thesis unfolding seq_compact_def by auto