# HG changeset patch # User hoelzl # Date 1363945303 -3600 # Node ID 1210309fddab2ff9f8ad3dec3c8b925abc236cbc # Parent adb441e4b9e992cc93b5c6adfbe4afe795fd1952 move first_countable_topology to the HOL image diff -r adb441e4b9e9 -r 1210309fddab src/HOL/Lim.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 "\f. (\n. f n \ s \ f n \ a) \ f ----> a \ - eventually (\n. P (f n)) sequentially" - shows "eventually P (at a within s)" -proof (rule ccontr) - let ?I = "\n. inverse (real (Suc n))" - def F \ "\n::nat. SOME x. x \ s \ x \ a \ dist x a < ?I n \ \ P x" - assume "\ eventually P (at a within s)" - hence P: "\d>0. \x. x \ s \ x \ a \ dist x a < d \ \ P x" - unfolding eventually_within eventually_at by fast - hence "\n. \x. x \ s \ x \ a \ dist x a < ?I n \ \ P x" by simp - hence F: "\n. F n \ s \ F n \ a \ dist (F n) a < ?I n \ \ P (F n)" - unfolding F_def by (rule someI_ex) - hence F0: "\n. F n \ s" and F1: "\n. F n \ a" - and F2: "\n. dist (F n) a < ?I n" and F3: "\n. \ 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 (\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 "\f. (\n. f n \ a) \ f ----> a \ - eventually (\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 \ 'b::topological_space" - assumes f: "f -- a --> l" - shows "\S. (\n. S n \ a) \ S ----> a \ (\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 \ 'b::topological_space" - assumes "\S. (\n. S n \ a) \ S ----> a \ (\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: - "(\S. (\n. S n \ a) \ S ----> (a::'a::metric_space) \ (\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 \ real" assumes ev: "b < x" "\ x' \ { b <..< x}. 0 \ f x'" and "isCont f x" diff -r adb441e4b9e9 -r 1210309fddab src/HOL/Metric_Spaces.thy --- 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 \ {y. dist x y < d}" then show "\e>0. \z. dist z y < e \ z \ {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 "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" + proof (safe intro!: exI[of _ "\n. {y. dist x y < inverse (Suc n)}"]) + fix S assume "open S" "x \ S" + then obtain e where "0 < e" "{y. dist x y < e} \ 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)} \ {y. dist x y < e}" + by auto + ultimately show "\i. {y. dist x y < inverse (Suc i)} \ S" + by blast + qed (auto intro: open_ball) +qed + end instance metric_space \ t2_space diff -r adb441e4b9e9 -r 1210309fddab 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 @@ -181,44 +181,15 @@ end -class first_countable_topology = topological_space + - assumes first_countable_basis: - "\A. countable A \ (\a\A. x \ a \ open a) \ (\S. open S \ x \ S \ (\a\A. a \ S))" - -lemma (in first_countable_topology) countable_basis_at_decseq: - obtains A :: "nat \ 'a set" where - "\i. open (A i)" "\i. x \ (A i)" - "\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially" -proof atomize_elim - from first_countable_basis[of x] obtain A - where "countable A" - and nhds: "\a. a \ A \ open a" "\a. a \ A \ x \ a" - and incl: "\S. open S \ x \ S \ \a\A. a \ S" by auto - then have "A \ {}" by auto - with `countable A` have r: "A = range (from_nat_into A)" by auto - def F \ "\n. \i\n. from_nat_into A i" - show "\A. (\i. open (A i)) \ (\i. x \ A i) \ - (\S. open S \ x \ S \ eventually (\i. A i \ 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 \ F i" using nhds(2) r by (auto simp: F_def) - next - fix S assume "open S" "x \ S" - from incl[OF this] obtain i where "F i \ S" - by (subst (asm) r) (auto simp: F_def) - moreover have "\j. i \ j \ F j \ F i" - by (auto simp: F_def) - ultimately show "eventually (\i. F i \ S) sequentially" - by (auto simp: eventually_sequentially) - qed -qed - lemma (in first_countable_topology) first_countable_basisE: obtains A where "countable A" "\a. a \ A \ x \ a" "\a. a \ A \ open a" "\S. open S \ x \ S \ (\a\A. a \ 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" "\a. a \ A \ x \ a" "\a. a \ A \ open a" @@ -245,77 +216,25 @@ 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 - +lemma (in topological_space) first_countableI: + assumes "countable A" and 1: "\a. a \ A \ x \ a" "\a. a \ A \ open a" + and 2: "\S. open S \ x \ S \ \a\A. a \ S" + shows "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" +proof (safe intro!: exI[of _ "from_nat_into A"]) + have "A \ {}" using 2[of UNIV] by auto + { fix i show "x \ from_nat_into A i" "open (from_nat_into A i)" + using range_from_nat_into_subset[OF `A \ {}`] 1 by auto } + { fix S assume "open S" "x\S" from 2[OF this] show "\i. from_nat_into A i \ 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 \ '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 "\A::('a\'b) set set. countable A \ (\a\A. x \ a \ open a) \ (\S. open S \ x \ S \ (\a\A. a \ S))" - proof (intro exI[of _ "(\(a, b). a \ b) ` (A \ B)"], safe) + show "\A::nat \ ('a \ 'b) set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" + proof (rule first_countableI[of "(\(a, b). a \ b) ` (A \ B)"], safe) fix a b assume x: "a \ A" "b \ B" with A(2, 3)[of a] B(2, 3)[of b] show "x \ a \ b" "open (a \ b)" unfolding mem_Times_iff by (auto intro: open_Times) @@ -329,23 +248,6 @@ qed (simp add: A B) qed -instance metric_space \ first_countable_topology -proof - fix x :: 'a - show "\A. countable A \ (\a\A. x \ a \ open a) \ (\S. open S \ x \ S \ (\a\A. a \ S))" - proof (intro exI, safe) - fix S assume "open S" "x \ S" - then obtain r where "0 < r" "{y. dist x y < r} \ 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)} \ {y. dist x y < r}" - by (auto simp: inverse_eq_divide) - ultimately show "\a\range (\n. {y. dist x y < inverse (Suc n)}). a \ S" - by auto - qed (auto intro: open_ball) -qed - class second_countable_topology = topological_space + assumes ex_countable_subbasis: "\B::'a::topological_space set set. countable B \ 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 "\A. countable A \ (\a\A. x \ a \ open a) \ (\S. open S \ x \ S \ (\a\A. a \ S))" - by (intro exI[of _ "{b\B. x \ b}"]) - (fastforce simp: topological_space_class.topological_basis_def) + then show "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" + by (intro first_countableI[of "{b\B. x \ b}"]) + (fastforce simp: topological_space_class.topological_basis_def)+ qed subsection {* Polish spaces *} @@ -2353,7 +2255,7 @@ apply (subgoal_tac "\x (y::real). 0 < 1 + abs y \ (x <= y \ x <= 1 + abs y)") by metis arith -lemma Bseq_eq_bounded: "Bseq f \ bounded (range f)" +lemma Bseq_eq_bounded: "Bseq f \ bounded (range f::_::real_normed_vector set)" unfolding Bseq_def bounded_pos by auto lemma bounded_Int[intro]: "bounded S \ bounded T \ bounded (S \ T)" diff -r adb441e4b9e9 -r 1210309fddab src/HOL/Probability/Fin_Map.thy --- 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: "\i S. i\domain x \ open (S i) \ x i\(S i) \ (\a\A i. a\(S i))" by auto have A_notempty: "\i. i \ domain x \ A i \ {}" using open_sub[of _ "\_. UNIV"] by auto let ?A = "(\f. Pi' (domain x) f) ` (Pi\<^isub>E (domain x) A)" - show "\A. countable A \ (\a\A. x \ a \ open a) \ (\S. open S \ x \ S \ (\a\A. a \ S))" - proof (rule exI[where x="?A"], safe) + show "\A::nat \ ('a\\<^isub>F'b) set. (\i. x \ (A i) \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" + proof (rule first_countableI[where A="?A"], safe) show "countable ?A" using A by (simp add: countable_PiE) next fix S::"('a \\<^isub>F 'b) set" assume "open S" "x \ S" diff -r adb441e4b9e9 -r 1210309fddab 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 @@ -681,17 +681,17 @@ definition (in topological_space) at :: "'a \ 'a filter" where "at a = nhds a within - {a}" -abbreviation at_right :: "'a\{topological_space, order} \ 'a filter" where +abbreviation (in order_topology) at_right :: "'a \ 'a filter" where "at_right x \ at x within {x <..}" -abbreviation at_left :: "'a\{topological_space, order} \ 'a filter" where +abbreviation (in order_topology) at_left :: "'a \ 'a filter" where "at_left x \ at x within {..< x}" -lemma eventually_nhds: +lemma (in topological_space) eventually_nhds: "eventually P (nhds a) \ (\S. open S \ a \ S \ (\x\S. P x))" unfolding nhds_def proof (rule eventually_Abs_filter, rule is_filter.intro) - have "open (UNIV :: 'a :: topological_space set) \ a \ UNIV \ (\x\UNIV. True)" by simp + have "open UNIV \ a \ UNIV \ (\x\UNIV. True)" by simp thus "\S. open S \ a \ S \ (\x\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 \ (\S. open S \ l \ S \ eventually (\x. f x \ S) F)" +lemma (in topological_space) tendsto_def: + "(f ---> l) F \ (\S. open S \ l \ S \ eventually (\x. f x \ S) F)" unfolding filterlim_def proof safe fix S assume "open S" "l \ S" "filtermap f F \ nhds l" @@ -859,12 +860,12 @@ lemma tendsto_mono: "F \ F' \ (f ---> l) F' \ (f ---> l) F" unfolding tendsto_def le_filter_def by fast -lemma topological_tendstoI: +lemma (in topological_space) topological_tendstoI: "(\S. open S \ l \ S \ eventually (\x. f x \ S) F) \ (f ---> l) F" unfolding tendsto_def by auto -lemma topological_tendstoD: +lemma (in topological_space) topological_tendstoD: "(f ---> l) F \ open S \ l \ S \ eventually (\x. f x \ S) F" unfolding tendsto_def by auto @@ -1290,6 +1291,19 @@ "subseq r \ eventually P sequentially \ eventually (\n. P (r n)) sequentially" unfolding eventually_sequentially by (metis seq_suble le_trans) +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 + lemma filterlim_subseq: "subseq f \ filterlim f sequentially sequentially" unfolding filterlim_iff by (metis eventually_subseq) @@ -1427,6 +1441,83 @@ lemma decseq_le: "decseq X \ X ----> L \ (L::'a::linorder_topology) \ X n" by (metis decseq_def LIMSEQ_le_const2) +subsection {* First countable topologies *} + +class first_countable_topology = topological_space + + assumes first_countable_basis: + "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" + +lemma (in first_countable_topology) countable_basis_at_decseq: + obtains A :: "nat \ 'a set" where + "\i. open (A i)" "\i. x \ (A i)" + "\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially" +proof atomize_elim + from first_countable_basis[of x] obtain A :: "nat \ 'a set" where + nhds: "\i. open (A i)" "\i. x \ A i" + and incl: "\S. open S \ x \ S \ \i. A i \ S" by auto + def F \ "\n. \i\n. A i" + show "\A. (\i. open (A i)) \ (\i. x \ A i) \ + (\S. open S \ x \ S \ eventually (\i. A i \ 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 \ F i" using nhds(2) by (auto simp: F_def) + next + fix S assume "open S" "x \ S" + from incl[OF this] obtain i where "F i \ S" unfolding F_def by auto + moreover have "\j. i \ j \ F j \ F i" + by (auto simp: F_def) + ultimately show "eventually (\i. F i \ S) sequentially" + by (auto simp: eventually_sequentially) + qed +qed + +lemma (in first_countable_topology) countable_basis: + obtains A :: "nat \ 'a 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 (in first_countable_topology) sequentially_imp_eventually_nhds_within: + 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 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 (in first_countable_topology) eventually_nhds_within_iff_sequentially: + "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: 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 (in first_countable_topology) eventually_nhds_iff_sequentially: + "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 + subsection {* Function limit at a point *} abbreviation @@ -1487,6 +1578,35 @@ shows "(\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: + "(\f. (\n. f n \ s \ f n \ a) \ f ----> a \ eventually (\n. P (f n)) sequentially) \ + 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: + "(\f. (\n. f n \ a) \ f ----> a \ eventually (\n. P (f n)) sequentially) \ eventually P (at a)" + using assms sequentially_imp_eventually_within [where s=UNIV] by simp + +lemma LIMSEQ_SEQ_conv1: + fixes f :: "'a::topological_space \ 'b::topological_space" + assumes f: "f -- a --> l" + shows "\S. (\n. S n \ a) \ S ----> a \ (\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 \ 'b::topological_space" + assumes "\S. (\n. S n \ a) \ S ----> a \ (\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: + "(\S. (\n. S n \ a) \ S ----> (a::'a::first_countable_topology) \ (\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 \ 'b::topological_space) \ 'a \ bool" where