move first_countable_topology to the HOL image
authorhoelzl
Fri, 22 Mar 2013 10:41:43 +0100
changeset 51473 1210309fddab
parent 51472 adb441e4b9e9
child 51474 1e9e68247ad1
move first_countable_topology to the HOL image
src/HOL/Lim.thy
src/HOL/Metric_Spaces.thy
src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy
src/HOL/Probability/Fin_Map.thy
src/HOL/Topological_Spaces.thy
--- a/src/HOL/Lim.thy	Fri Mar 22 10:41:43 2013 +0100
+++ b/src/HOL/Lim.thy	Fri Mar 22 10:41:43 2013 +0100
@@ -226,58 +226,6 @@
 by (rule isUCont [THEN isUCont_Cauchy])
 
 
-subsection {* Relation of LIM and LIMSEQ *}
-
-lemma sequentially_imp_eventually_within:
-  fixes a :: "'a::metric_space"
-  assumes "\<forall>f. (\<forall>n. f n \<in> s \<and> f n \<noteq> a) \<and> f ----> a \<longrightarrow>
-    eventually (\<lambda>n. P (f n)) sequentially"
-  shows "eventually P (at a within s)"
-proof (rule ccontr)
-  let ?I = "\<lambda>n. inverse (real (Suc n))"
-  def F \<equiv> "\<lambda>n::nat. SOME x. x \<in> s \<and> x \<noteq> a \<and> dist x a < ?I n \<and> \<not> P x"
-  assume "\<not> eventually P (at a within s)"
-  hence P: "\<forall>d>0. \<exists>x. x \<in> s \<and> x \<noteq> a \<and> dist x a < d \<and> \<not> P x"
-    unfolding eventually_within eventually_at by fast
-  hence "\<And>n. \<exists>x. x \<in> s \<and> x \<noteq> a \<and> dist x a < ?I n \<and> \<not> P x" by simp
-  hence F: "\<And>n. F n \<in> s \<and> F n \<noteq> a \<and> dist (F n) a < ?I n \<and> \<not> P (F n)"
-    unfolding F_def by (rule someI_ex)
-  hence F0: "\<forall>n. F n \<in> s" and F1: "\<forall>n. F n \<noteq> a"
-    and F2: "\<forall>n. dist (F n) a < ?I n" and F3: "\<forall>n. \<not> P (F n)"
-    by fast+
-  from LIMSEQ_inverse_real_of_nat have "F ----> a"
-    by (rule metric_tendsto_imp_tendsto,
-      simp add: dist_norm F2 less_imp_le)
-  hence "eventually (\<lambda>n. P (F n)) sequentially"
-    using assms F0 F1 by simp
-  thus "False" by (simp add: F3)
-qed
-
-lemma sequentially_imp_eventually_at:
-  fixes a :: "'a::metric_space"
-  assumes "\<forall>f. (\<forall>n. f n \<noteq> a) \<and> f ----> a \<longrightarrow>
-    eventually (\<lambda>n. P (f n)) sequentially"
-  shows "eventually P (at a)"
-  using assms sequentially_imp_eventually_within [where s=UNIV] by simp
-
-lemma LIMSEQ_SEQ_conv1:
-  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
-  assumes f: "f -- a --> l"
-  shows "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. f (S n)) ----> l"
-  using tendsto_compose_eventually [OF f, where F=sequentially] by simp
-
-lemma LIMSEQ_SEQ_conv2:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
-  assumes "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. f (S n)) ----> l"
-  shows "f -- a --> l"
-  using assms unfolding tendsto_def [where l=l]
-  by (simp add: sequentially_imp_eventually_at)
-
-lemma LIMSEQ_SEQ_conv:
-  "(\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> (a::'a::metric_space) \<longrightarrow> (\<lambda>n. X (S n)) ----> L) =
-   (X -- a --> (L::'b::topological_space))"
-  using LIMSEQ_SEQ_conv2 LIMSEQ_SEQ_conv1 ..
-
 lemma LIM_less_bound: 
   fixes f :: "real \<Rightarrow> real"
   assumes ev: "b < x" "\<forall> x' \<in> { b <..< x}. 0 \<le> f x'" and "isCont f x"
--- a/src/HOL/Metric_Spaces.thy	Fri Mar 22 10:41:43 2013 +0100
+++ b/src/HOL/Metric_Spaces.thy	Fri Mar 22 10:41:43 2013 +0100
@@ -100,13 +100,31 @@
     unfolding open_dist by fast
 qed
 
-lemma (in metric_space) open_ball: "open {y. dist x y < d}"
+lemma open_ball: "open {y. dist x y < d}"
 proof (unfold open_dist, intro ballI)
   fix y assume *: "y \<in> {y. dist x y < d}"
   then show "\<exists>e>0. \<forall>z. dist z y < e \<longrightarrow> z \<in> {y. dist x y < d}"
     by (auto intro!: exI[of _ "d - dist x y"] simp: field_simps dist_triangle_lt)
 qed
 
+subclass first_countable_topology
+proof
+  fix x 
+  show "\<exists>A::nat \<Rightarrow> 'a set. (\<forall>i. x \<in> A i \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))"
+  proof (safe intro!: exI[of _ "\<lambda>n. {y. dist x y < inverse (Suc n)}"])
+    fix S assume "open S" "x \<in> S"
+    then obtain e where "0 < e" "{y. dist x y < e} \<subseteq> S"
+      by (auto simp: open_dist subset_eq dist_commute)
+    moreover
+    then obtain i where "inverse (Suc i) < e"
+      by (auto dest!: reals_Archimedean)
+    then have "{y. dist x y < inverse (Suc i)} \<subseteq> {y. dist x y < e}"
+      by auto
+    ultimately show "\<exists>i. {y. dist x y < inverse (Suc i)} \<subseteq> S"
+      by blast
+  qed (auto intro: open_ball)
+qed
+
 end
 
 instance metric_space \<subseteq> t2_space
--- 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
@@ -181,44 +181,15 @@
 
 end
 
-class first_countable_topology = topological_space +
-  assumes first_countable_basis:
-    "\<exists>A. countable A \<and> (\<forall>a\<in>A. x \<in> a \<and> open a) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>a\<in>A. a \<subseteq> S))"
-
-lemma (in first_countable_topology) countable_basis_at_decseq:
-  obtains A :: "nat \<Rightarrow> 'a set" where
-    "\<And>i. open (A i)" "\<And>i. x \<in> (A i)"
-    "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially"
-proof atomize_elim
-  from first_countable_basis[of x] obtain A
-    where "countable A"
-    and nhds: "\<And>a. a \<in> A \<Longrightarrow> open a" "\<And>a. a \<in> A \<Longrightarrow> x \<in> a"
-    and incl: "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> \<exists>a\<in>A. a \<subseteq> S"  by auto
-  then have "A \<noteq> {}" by auto
-  with `countable A` have r: "A = range (from_nat_into A)" by auto
-  def F \<equiv> "\<lambda>n. \<Inter>i\<le>n. from_nat_into A i"
-  show "\<exists>A. (\<forall>i. open (A i)) \<and> (\<forall>i. x \<in> A i) \<and>
-      (\<forall>S. open S \<longrightarrow> x \<in> S \<longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially)"
-  proof (safe intro!: exI[of _ F])
-    fix i
-    show "open (F i)" using nhds(1) r by (auto simp: F_def intro!: open_INT)
-    show "x \<in> F i" using nhds(2) r by (auto simp: F_def)
-  next
-    fix S assume "open S" "x \<in> S"
-    from incl[OF this] obtain i where "F i \<subseteq> S"
-      by (subst (asm) r) (auto simp: F_def)
-    moreover have "\<And>j. i \<le> j \<Longrightarrow> F j \<subseteq> F i"
-      by (auto simp: F_def)
-    ultimately show "eventually (\<lambda>i. F i \<subseteq> S) sequentially"
-      by (auto simp: eventually_sequentially)
-  qed
-qed
-
 lemma (in first_countable_topology) first_countable_basisE:
   obtains A where "countable A" "\<And>a. a \<in> A \<Longrightarrow> x \<in> a" "\<And>a. a \<in> A \<Longrightarrow> open a"
     "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> (\<exists>a\<in>A. a \<subseteq> S)"
   using first_countable_basis[of x]
-  by atomize_elim auto
+  apply atomize_elim
+  apply (elim exE)
+  apply (rule_tac x="range A" in exI)
+  apply auto
+  done
 
 lemma (in first_countable_topology) first_countable_basis_Int_stableE:
   obtains A where "countable A" "\<And>a. a \<in> A \<Longrightarrow> x \<in> a" "\<And>a. a \<in> A \<Longrightarrow> open a"
@@ -245,77 +216,25 @@
   qed
 qed
 
-
-lemma countable_basis:
-  obtains A :: "nat \<Rightarrow> 'a::first_countable_topology set" where
-    "\<And>i. open (A i)" "\<And>i. x \<in> A i"
-    "\<And>F. (\<forall>n. F n \<in> A n) \<Longrightarrow> F ----> x"
-proof atomize_elim
-  from countable_basis_at_decseq[of x] guess A . note A = this
-  { fix F S assume "\<forall>n. F n \<in> A n" "open S" "x \<in> S"
-    with A(3)[of S] have "eventually (\<lambda>n. F n \<in> S) sequentially"
-      by (auto elim: eventually_elim1 simp: subset_eq) }
-  with A show "\<exists>A. (\<forall>i. open (A i)) \<and> (\<forall>i. x \<in> A i) \<and> (\<forall>F. (\<forall>n. F n \<in> A n) \<longrightarrow> 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 "\<forall>f. (\<forall>n. f n \<in> s) \<and> f ----> a \<longrightarrow> eventually (\<lambda>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 "\<not> eventually P (nhds a within s)"
-  with A have P: "\<exists>F. \<forall>n. F n \<in> s \<and> F n \<in> A n \<and> \<not> P (F n)"
-    unfolding Limits.eventually_within eventually_nhds by (intro choice) fastforce
-  then guess F ..
-  hence F0: "\<forall>n. F n \<in> s" and F2: "\<forall>n. F n \<in> A n" and F3: "\<forall>n. \<not> P (F n)"
-    by fast+
-  with A have "F ----> a" by auto
-  hence "eventually (\<lambda>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) \<longleftrightarrow> 
-    (\<forall>f. (\<forall>n. f n \<in> s) \<and> f ----> a \<longrightarrow> eventually (\<lambda>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 \<in> S" "\<forall>x\<in>S. x \<in> s \<longrightarrow> P x"
-    by (auto simp: Limits.eventually_within eventually_nhds)
-  moreover fix f assume "\<forall>n. f n \<in> s" "f ----> a"
-  ultimately show "eventually (\<lambda>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) \<longleftrightarrow> (\<forall>f. f ----> a \<longrightarrow> eventually (\<lambda>n. P (f n)) sequentially)"
-  using eventually_nhds_within_iff_sequentially[of P a UNIV] by simp
-
-lemma not_eventually_sequentiallyD:
-  assumes P: "\<not> eventually P sequentially"
-  shows "\<exists>r. subseq r \<and> (\<forall>n. \<not> P (r n))"
-proof -
-  from P have "\<forall>n. \<exists>m\<ge>n. \<not> P m"
-    unfolding eventually_sequentially by (simp add: not_less)
-  then obtain r where "\<And>n. r n \<ge> n" "\<And>n. \<not> P (r n)"
-    by (auto simp: choice_iff)
-  then show ?thesis
-    by (auto intro!: exI[of _ "\<lambda>n. r (((Suc \<circ> r) ^^ Suc n) 0)"]
-             simp: less_eq_Suc_le subseq_Suc_iff)
-qed
-
+lemma (in topological_space) first_countableI:
+  assumes "countable A" and 1: "\<And>a. a \<in> A \<Longrightarrow> x \<in> a" "\<And>a. a \<in> A \<Longrightarrow> open a"
+   and 2: "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> \<exists>a\<in>A. a \<subseteq> S"
+  shows "\<exists>A::nat \<Rightarrow> 'a set. (\<forall>i. x \<in> A i \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))"
+proof (safe intro!: exI[of _ "from_nat_into A"])
+  have "A \<noteq> {}" using 2[of UNIV] by auto
+  { fix i show "x \<in> from_nat_into A i" "open (from_nat_into A i)"
+      using range_from_nat_into_subset[OF `A \<noteq> {}`] 1 by auto }
+  { fix S assume "open S" "x\<in>S" from 2[OF this] show "\<exists>i. from_nat_into A i \<subseteq> S"
+      using subset_range_from_nat_into[OF `countable A`] by auto }
+qed
 
 instance prod :: (first_countable_topology, first_countable_topology) first_countable_topology
 proof
   fix x :: "'a \<times> 'b"
   from first_countable_basisE[of "fst x"] guess A :: "'a set set" . note A = this
   from first_countable_basisE[of "snd x"] guess B :: "'b set set" . note B = this
-  show "\<exists>A::('a\<times>'b) set set. countable A \<and> (\<forall>a\<in>A. x \<in> a \<and> open a) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>a\<in>A. a \<subseteq> S))"
-  proof (intro exI[of _ "(\<lambda>(a, b). a \<times> b) ` (A \<times> B)"], safe)
+  show "\<exists>A::nat \<Rightarrow> ('a \<times> 'b) set. (\<forall>i. x \<in> A i \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))"
+  proof (rule first_countableI[of "(\<lambda>(a, b). a \<times> b) ` (A \<times> B)"], safe)
     fix a b assume x: "a \<in> A" "b \<in> B"
     with A(2, 3)[of a] B(2, 3)[of b] show "x \<in> a \<times> b" "open (a \<times> b)"
       unfolding mem_Times_iff by (auto intro: open_Times)
@@ -329,23 +248,6 @@
   qed (simp add: A B)
 qed
 
-instance metric_space \<subseteq> first_countable_topology
-proof
-  fix x :: 'a
-  show "\<exists>A. countable A \<and> (\<forall>a\<in>A. x \<in> a \<and> open a) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>a\<in>A. a \<subseteq> S))"
-  proof (intro exI, safe)
-    fix S assume "open S" "x \<in> S"
-    then obtain r where "0 < r" "{y. dist x y < r} \<subseteq> S"
-      by (auto simp: open_dist dist_commute subset_eq)
-    moreover from reals_Archimedean[OF `0 < r`] guess n ..
-    moreover
-    then have "{y. dist x y < inverse (Suc n)} \<subseteq> {y. dist x y < r}"
-      by (auto simp: inverse_eq_divide)
-    ultimately show "\<exists>a\<in>range (\<lambda>n. {y. dist x y < inverse (Suc n)}). a \<subseteq> S"
-      by auto
-  qed (auto intro: open_ball)
-qed
-
 class second_countable_topology = topological_space +
   assumes ex_countable_subbasis: "\<exists>B::'a::topological_space set set. countable B \<and> open = generate_topology B"
 begin
@@ -417,9 +319,9 @@
   then have B: "countable B" "topological_basis B"
     using countable_basis is_basis
     by (auto simp: countable_basis is_basis)
-  then show "\<exists>A. countable A \<and> (\<forall>a\<in>A. x \<in> a \<and> open a) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>a\<in>A. a \<subseteq> S))"
-    by (intro exI[of _ "{b\<in>B. x \<in> b}"])
-       (fastforce simp: topological_space_class.topological_basis_def)
+  then show "\<exists>A::nat \<Rightarrow> 'a set. (\<forall>i. x \<in> A i \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))"
+    by (intro first_countableI[of "{b\<in>B. x \<in> b}"])
+       (fastforce simp: topological_space_class.topological_basis_def)+
 qed
 
 subsection {* Polish spaces *}
@@ -2353,7 +2255,7 @@
   apply (subgoal_tac "\<And>x (y::real). 0 < 1 + abs y \<and> (x <= y \<longrightarrow> x <= 1 + abs y)")
   by metis arith
 
-lemma Bseq_eq_bounded: "Bseq f \<longleftrightarrow> bounded (range f)"
+lemma Bseq_eq_bounded: "Bseq f \<longleftrightarrow> bounded (range f::_::real_normed_vector set)"
   unfolding Bseq_def bounded_pos by auto
 
 lemma bounded_Int[intro]: "bounded S \<or> bounded T \<Longrightarrow> bounded (S \<inter> T)"
--- a/src/HOL/Probability/Fin_Map.thy	Fri Mar 22 10:41:43 2013 +0100
+++ b/src/HOL/Probability/Fin_Map.thy	Fri Mar 22 10:41:43 2013 +0100
@@ -222,8 +222,8 @@
   hence open_sub: "\<And>i S. i\<in>domain x \<Longrightarrow> open (S i) \<Longrightarrow> x i\<in>(S i) \<Longrightarrow> (\<exists>a\<in>A i. a\<subseteq>(S i))" by auto
   have A_notempty: "\<And>i. i \<in> domain x \<Longrightarrow> A i \<noteq> {}" using open_sub[of _ "\<lambda>_. UNIV"] by auto
   let ?A = "(\<lambda>f. Pi' (domain x) f) ` (Pi\<^isub>E (domain x) A)"
-  show "\<exists>A. countable A \<and> (\<forall>a\<in>A. x \<in> a \<and> open a) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>a\<in>A. a \<subseteq> S))"
-  proof (rule exI[where x="?A"], safe)
+  show "\<exists>A::nat \<Rightarrow> ('a\<Rightarrow>\<^isub>F'b) set. (\<forall>i. x \<in> (A i) \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))"
+  proof (rule first_countableI[where A="?A"], safe)
     show "countable ?A" using A by (simp add: countable_PiE)
   next
     fix S::"('a \<Rightarrow>\<^isub>F 'b) set" assume "open S" "x \<in> S"
--- 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
@@ -681,17 +681,17 @@
 definition (in topological_space) at :: "'a \<Rightarrow> 'a filter"
   where "at a = nhds a within - {a}"
 
-abbreviation at_right :: "'a\<Colon>{topological_space, order} \<Rightarrow> 'a filter" where
+abbreviation (in order_topology) at_right :: "'a \<Rightarrow> 'a filter" where
   "at_right x \<equiv> at x within {x <..}"
 
-abbreviation at_left :: "'a\<Colon>{topological_space, order} \<Rightarrow> 'a filter" where
+abbreviation (in order_topology) at_left :: "'a \<Rightarrow> 'a filter" where
   "at_left x \<equiv> at x within {..< x}"
 
-lemma eventually_nhds:
+lemma (in topological_space) eventually_nhds:
   "eventually P (nhds a) \<longleftrightarrow> (\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x))"
   unfolding nhds_def
 proof (rule eventually_Abs_filter, rule is_filter.intro)
-  have "open (UNIV :: 'a :: topological_space set) \<and> a \<in> UNIV \<and> (\<forall>x\<in>UNIV. True)" by simp
+  have "open UNIV \<and> a \<in> UNIV \<and> (\<forall>x\<in>UNIV. True)" by simp
   thus "\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. True)" ..
 next
   fix P Q
@@ -843,7 +843,8 @@
     map (fn thm => @{thm tendsto_eq_rhs} OF [thm]) o Tendsto_Intros.get o Context.proof_of);
 *}
 
-lemma tendsto_def: "(f ---> l) F \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) F)"
+lemma (in topological_space) tendsto_def:
+   "(f ---> l) F \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) F)"
   unfolding filterlim_def
 proof safe
   fix S assume "open S" "l \<in> S" "filtermap f F \<le> nhds l"
@@ -859,12 +860,12 @@
 lemma tendsto_mono: "F \<le> F' \<Longrightarrow> (f ---> l) F' \<Longrightarrow> (f ---> l) F"
   unfolding tendsto_def le_filter_def by fast
 
-lemma topological_tendstoI:
+lemma (in topological_space) topological_tendstoI:
   "(\<And>S. open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>x. f x \<in> S) F)
     \<Longrightarrow> (f ---> l) F"
   unfolding tendsto_def by auto
 
-lemma topological_tendstoD:
+lemma (in topological_space) topological_tendstoD:
   "(f ---> l) F \<Longrightarrow> open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>x. f x \<in> S) F"
   unfolding tendsto_def by auto
 
@@ -1290,6 +1291,19 @@
   "subseq r \<Longrightarrow> eventually P sequentially \<Longrightarrow> eventually (\<lambda>n. P (r n)) sequentially"
   unfolding eventually_sequentially by (metis seq_suble le_trans)
 
+lemma not_eventually_sequentiallyD:
+  assumes P: "\<not> eventually P sequentially"
+  shows "\<exists>r. subseq r \<and> (\<forall>n. \<not> P (r n))"
+proof -
+  from P have "\<forall>n. \<exists>m\<ge>n. \<not> P m"
+    unfolding eventually_sequentially by (simp add: not_less)
+  then obtain r where "\<And>n. r n \<ge> n" "\<And>n. \<not> P (r n)"
+    by (auto simp: choice_iff)
+  then show ?thesis
+    by (auto intro!: exI[of _ "\<lambda>n. r (((Suc \<circ> r) ^^ Suc n) 0)"]
+             simp: less_eq_Suc_le subseq_Suc_iff)
+qed
+
 lemma filterlim_subseq: "subseq f \<Longrightarrow> filterlim f sequentially sequentially"
   unfolding filterlim_iff by (metis eventually_subseq)
 
@@ -1427,6 +1441,83 @@
 lemma decseq_le: "decseq X \<Longrightarrow> X ----> L \<Longrightarrow> (L::'a::linorder_topology) \<le> X n"
   by (metis decseq_def LIMSEQ_le_const2)
 
+subsection {* First countable topologies *}
+
+class first_countable_topology = topological_space +
+  assumes first_countable_basis:
+    "\<exists>A::nat \<Rightarrow> 'a set. (\<forall>i. x \<in> A i \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))"
+
+lemma (in first_countable_topology) countable_basis_at_decseq:
+  obtains A :: "nat \<Rightarrow> 'a set" where
+    "\<And>i. open (A i)" "\<And>i. x \<in> (A i)"
+    "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially"
+proof atomize_elim
+  from first_countable_basis[of x] obtain A :: "nat \<Rightarrow> 'a set" where
+    nhds: "\<And>i. open (A i)" "\<And>i. x \<in> A i"
+    and incl: "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> \<exists>i. A i \<subseteq> S"  by auto
+  def F \<equiv> "\<lambda>n. \<Inter>i\<le>n. A i"
+  show "\<exists>A. (\<forall>i. open (A i)) \<and> (\<forall>i. x \<in> A i) \<and>
+      (\<forall>S. open S \<longrightarrow> x \<in> S \<longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially)"
+  proof (safe intro!: exI[of _ F])
+    fix i
+    show "open (F i)" using nhds(1) by (auto simp: F_def intro!: open_INT)
+    show "x \<in> F i" using nhds(2) by (auto simp: F_def)
+  next
+    fix S assume "open S" "x \<in> S"
+    from incl[OF this] obtain i where "F i \<subseteq> S" unfolding F_def by auto
+    moreover have "\<And>j. i \<le> j \<Longrightarrow> F j \<subseteq> F i"
+      by (auto simp: F_def)
+    ultimately show "eventually (\<lambda>i. F i \<subseteq> S) sequentially"
+      by (auto simp: eventually_sequentially)
+  qed
+qed
+
+lemma (in first_countable_topology) countable_basis:
+  obtains A :: "nat \<Rightarrow> 'a set" where
+    "\<And>i. open (A i)" "\<And>i. x \<in> A i"
+    "\<And>F. (\<forall>n. F n \<in> A n) \<Longrightarrow> F ----> x"
+proof atomize_elim
+  from countable_basis_at_decseq[of x] guess A . note A = this
+  { fix F S assume "\<forall>n. F n \<in> A n" "open S" "x \<in> S"
+    with A(3)[of S] have "eventually (\<lambda>n. F n \<in> S) sequentially"
+      by (auto elim: eventually_elim1 simp: subset_eq) }
+  with A show "\<exists>A. (\<forall>i. open (A i)) \<and> (\<forall>i. x \<in> A i) \<and> (\<forall>F. (\<forall>n. F n \<in> A n) \<longrightarrow> F ----> x)"
+    by (intro exI[of _ A]) (auto simp: tendsto_def)
+qed
+
+lemma (in first_countable_topology) sequentially_imp_eventually_nhds_within:
+  assumes "\<forall>f. (\<forall>n. f n \<in> s) \<and> f ----> a \<longrightarrow> eventually (\<lambda>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 "\<not> eventually P (nhds a within s)"
+  with A have P: "\<exists>F. \<forall>n. F n \<in> s \<and> F n \<in> A n \<and> \<not> P (F n)"
+    unfolding eventually_within eventually_nhds by (intro choice) fastforce
+  then guess F ..
+  hence F0: "\<forall>n. F n \<in> s" and F2: "\<forall>n. F n \<in> A n" and F3: "\<forall>n. \<not> P (F n)"
+    by fast+
+  with A have "F ----> a" by auto
+  hence "eventually (\<lambda>n. P (F n)) sequentially"
+    using assms F0 by simp
+  thus "False" by (simp add: F3)
+qed
+
+lemma (in first_countable_topology) eventually_nhds_within_iff_sequentially:
+  "eventually P (nhds a within s) \<longleftrightarrow> 
+    (\<forall>f. (\<forall>n. f n \<in> s) \<and> f ----> a \<longrightarrow> eventually (\<lambda>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 \<in> S" "\<forall>x\<in>S. x \<in> s \<longrightarrow> P x"
+    by (auto simp: eventually_within eventually_nhds)
+  moreover fix f assume "\<forall>n. f n \<in> s" "f ----> a"
+  ultimately show "eventually (\<lambda>n. P (f n)) sequentially"
+    by (auto dest!: topological_tendstoD elim: eventually_elim1)
+qed
+
+lemma (in first_countable_topology) eventually_nhds_iff_sequentially:
+  "eventually P (nhds a) \<longleftrightarrow> (\<forall>f. f ----> a \<longrightarrow> eventually (\<lambda>n. P (f n)) sequentially)"
+  using eventually_nhds_within_iff_sequentially[of P a UNIV] by simp
+
 subsection {* Function limit at a point *}
 
 abbreviation
@@ -1487,6 +1578,35 @@
   shows "(\<lambda>x. g (f x)) -- a --> c"
   using g f inj by (rule tendsto_compose_eventually)
 
+subsubsection {* Relation of LIM and LIMSEQ *}
+
+lemma (in first_countable_topology) sequentially_imp_eventually_within:
+  "(\<forall>f. (\<forall>n. f n \<in> s \<and> f n \<noteq> a) \<and> f ----> a \<longrightarrow> eventually (\<lambda>n. P (f n)) sequentially) \<Longrightarrow>
+    eventually P (at a within s)"
+  unfolding at_def within_within_eq
+  by (intro sequentially_imp_eventually_nhds_within) auto
+
+lemma (in first_countable_topology) sequentially_imp_eventually_at:
+  "(\<forall>f. (\<forall>n. f n \<noteq> a) \<and> f ----> a \<longrightarrow> eventually (\<lambda>n. P (f n)) sequentially) \<Longrightarrow> eventually P (at a)"
+  using assms sequentially_imp_eventually_within [where s=UNIV] by simp
+
+lemma LIMSEQ_SEQ_conv1:
+  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
+  assumes f: "f -- a --> l"
+  shows "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. f (S n)) ----> l"
+  using tendsto_compose_eventually [OF f, where F=sequentially] by simp
+
+lemma LIMSEQ_SEQ_conv2:
+  fixes f :: "'a::first_countable_topology \<Rightarrow> 'b::topological_space"
+  assumes "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. f (S n)) ----> l"
+  shows "f -- a --> l"
+  using assms unfolding tendsto_def [where l=l] by (simp add: sequentially_imp_eventually_at)
+
+lemma LIMSEQ_SEQ_conv:
+  "(\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> (a::'a::first_countable_topology) \<longrightarrow> (\<lambda>n. X (S n)) ----> L) =
+   (X -- a --> (L::'b::topological_space))"
+  using LIMSEQ_SEQ_conv2 LIMSEQ_SEQ_conv1 ..
+
 subsection {* Continuity *}
 
 definition isCont :: "('a::topological_space \<Rightarrow> 'b::topological_space) \<Rightarrow> 'a \<Rightarrow> bool" where