# HG changeset patch # User hoelzl # Date 1362494601 -3600 # Node ID 490f34774a9ae14809d3e396c4a3f492af3694ac # Parent 166170c5f8a2895a2bbc06ae030248d0bf399dd6 eventually nhds represented using sequentially diff -r 166170c5f8a2 -r 490f34774a9a src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Mar 05 15:43:20 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Mar 05 15:43:21 2013 +0100 @@ -245,6 +245,70 @@ qed qed + +lemma countable_basis: + obtains A :: "nat \ 'a::first_countable_topology set" where + "\i. open (A i)" "\i. x \ A i" + "\F. (\n. F n \ A n) \ F ----> x" +proof atomize_elim + from countable_basis_at_decseq[of x] guess A . note A = this + { fix F S assume "\n. F n \ A n" "open S" "x \ S" + with A(3)[of S] have "eventually (\n. F n \ S) sequentially" + by (auto elim: eventually_elim1 simp: subset_eq) } + with A show "\A. (\i. open (A i)) \ (\i. x \ A i) \ (\F. (\n. F n \ A n) \ F ----> x)" + by (intro exI[of _ A]) (auto simp: tendsto_def) +qed + +lemma sequentially_imp_eventually_nhds_within: + fixes a :: "'a::first_countable_topology" + assumes "\f. (\n. f n \ s) \ f ----> a \ eventually (\n. P (f n)) sequentially" + shows "eventually P (nhds a within s)" +proof (rule ccontr) + from countable_basis[of a] guess A . note A = this + assume "\ eventually P (nhds a within s)" + with A have P: "\F. \n. F n \ s \ F n \ A n \ \ P (F n)" + unfolding Limits.eventually_within eventually_nhds by (intro choice) fastforce + then guess F .. + hence F0: "\n. F n \ s" and F2: "\n. F n \ A n" and F3: "\n. \ P (F n)" + by fast+ + with A have "F ----> a" by auto + hence "eventually (\n. P (F n)) sequentially" + using assms F0 by simp + thus "False" by (simp add: F3) +qed + +lemma eventually_nhds_within_iff_sequentially: + fixes a :: "'a::first_countable_topology" + shows "eventually P (nhds a within s) \ + (\f. (\n. f n \ s) \ f ----> a \ eventually (\n. P (f n)) sequentially)" +proof (safe intro!: sequentially_imp_eventually_nhds_within) + assume "eventually P (nhds a within s)" + then obtain S where "open S" "a \ S" "\x\S. x \ s \ P x" + by (auto simp: Limits.eventually_within eventually_nhds) + moreover fix f assume "\n. f n \ s" "f ----> a" + ultimately show "eventually (\n. P (f n)) sequentially" + by (auto dest!: topological_tendstoD elim: eventually_elim1) +qed + +lemma eventually_nhds_iff_sequentially: + fixes a :: "'a::first_countable_topology" + shows "eventually P (nhds a) \ (\f. f ----> a \ eventually (\n. P (f n)) sequentially)" + using eventually_nhds_within_iff_sequentially[of P a UNIV] by simp + +lemma not_eventually_sequentiallyD: + assumes P: "\ eventually P sequentially" + shows "\r. subseq r \ (\n. \ P (r n))" +proof - + from P have "\n. \m\n. \ P m" + unfolding eventually_sequentially by (simp add: not_less) + then obtain r where "\n. r n \ n" "\n. \ P (r n)" + by (auto simp: choice_iff) + then show ?thesis + by (auto intro!: exI[of _ "\n. r (((Suc \ r) ^^ Suc n) 0)"] + simp: less_eq_Suc_le subseq_Suc_iff) +qed + + instance prod :: (first_countable_topology, first_countable_topology) first_countable_topology proof fix x :: "'a \ 'b"