Various lemmas for HOL-Analysis
authorManuel Eberl <eberlm@in.tum.de>
Sun, 20 Aug 2017 03:35:20 +0200
changeset 66456 621897f47fab
parent 66455 158c513a39f5
child 66465 86223a532d8e
child 66466 aec5d9c88d69
Various lemmas for HOL-Analysis
src/HOL/Analysis/Conformal_Mappings.thy
src/HOL/Analysis/Extended_Real_Limits.thy
src/HOL/Analysis/Path_Connected.thy
src/HOL/Analysis/Set_Integral.thy
src/HOL/Analysis/Summation_Tests.thy
src/HOL/Limits.thy
src/HOL/Series.thy
--- a/src/HOL/Analysis/Conformal_Mappings.thy	Fri Aug 18 22:55:54 2017 +0200
+++ b/src/HOL/Analysis/Conformal_Mappings.thy	Sun Aug 20 03:35:20 2017 +0200
@@ -278,6 +278,21 @@
     by (metis open_ball \<open>\<xi> islimpt T\<close> centre_in_ball fT0 insertE insert_Diff islimptE)
 qed
 
+corollary analytic_continuation_open:
+  assumes "open s" "open s'" "s \<noteq> {}" "connected s'" "s \<subseteq> s'"
+  assumes "f holomorphic_on s'" "g holomorphic_on s'" "\<And>z. z \<in> s \<Longrightarrow> f z = g z"
+  assumes "z \<in> s'"
+  shows   "f z = g z"
+proof -
+  from \<open>s \<noteq> {}\<close> obtain \<xi> where "\<xi> \<in> s" by auto
+  with \<open>open s\<close> have \<xi>: "\<xi> islimpt s" 
+    by (intro interior_limit_point) (auto simp: interior_open)
+  have "f z - g z = 0"
+    by (rule analytic_continuation[of "\<lambda>z. f z - g z" s' s \<xi>])
+       (insert assms \<open>\<xi> \<in> s\<close> \<xi>, auto intro: holomorphic_intros)
+  thus ?thesis by simp
+qed
+
 
 subsection\<open>Open mapping theorem\<close>
 
@@ -3910,4 +3925,291 @@
   then show ?thesis unfolding c_def using w_def by auto
 qed
 
+
+subsection \<open>More facts about poles and residues\<close>
+
+lemma lhopital_complex_simple:
+  assumes "(f has_field_derivative f') (at z)" 
+  assumes "(g has_field_derivative g') (at z)"
+  assumes "f z = 0" "g z = 0" "g' \<noteq> 0" "f' / g' = c"
+  shows   "((\<lambda>w. f w / g w) \<longlongrightarrow> c) (at z)"
+proof -
+  have "eventually (\<lambda>w. w \<noteq> z) (at z)"
+    by (auto simp: eventually_at_filter)
+  hence "eventually (\<lambda>w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z)) = f w / g w) (at z)"
+    by eventually_elim (simp add: assms divide_simps)
+  moreover have "((\<lambda>w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z))) \<longlongrightarrow> f' / g') (at z)"
+    by (intro tendsto_divide has_field_derivativeD assms)
+  ultimately have "((\<lambda>w. f w / g w) \<longlongrightarrow> f' / g') (at z)"
+    by (rule Lim_transform_eventually)
+  with assms show ?thesis by simp
+qed
+
+lemma porder_eqI:
+  assumes "open s" "z \<in> s" "g holomorphic_on s" "g z \<noteq> 0" "n > 0"
+  assumes "\<And>w. w \<in> s - {z} \<Longrightarrow> f w = g w / (w - z) ^ n"
+  shows   "porder f z = n"
+proof -
+  define f' where "f' = (\<lambda>x. if x = z then 0 else inverse (f x))"
+  define g' where "g' = (\<lambda>x. inverse (g x))"
+  define s' where "s' = (g -` (-{0}) \<inter> s)"
+  have "continuous_on s g"
+    by (intro holomorphic_on_imp_continuous_on) fact
+  hence "open s'"
+    unfolding s'_def using assms by (subst (asm) continuous_on_open_vimage) blast+
+  have s': "z \<in> s'" "g' holomorphic_on s'" "g' z \<noteq> 0" using assms 
+    by (auto simp: s'_def g'_def intro!: holomorphic_intros)
+  have f'_g': "f' w = g' w * (w - z) ^ n" if "w \<in> s'" for w
+    unfolding f'_def g'_def using that \<open>n > 0\<close>
+    by (auto simp: assms(6) field_simps s'_def)
+  have "porder f z = zorder f' z"
+    by (simp add: porder_def f'_def)
+  also have "\<dots> = n" using assms f'_g' 
+    by (intro zorder_eqI[OF \<open>open s'\<close> s']) (auto simp: f'_def g'_def field_simps s'_def)
+  finally show ?thesis .
+qed
+
+lemma simple_poleI':
+  assumes "open s" "connected s" "z \<in> s"
+  assumes "\<And>w. w \<in> s - {z} \<Longrightarrow> 
+             ((\<lambda>w. inverse (f w)) has_field_derivative f' w) (at w)"
+  assumes "f holomorphic_on s - {z}" "f' holomorphic_on s" "is_pole f z" "f' z \<noteq> 0"
+  shows   "porder f z = 1"
+proof -
+  define g where "g = (\<lambda>w. if w = z then 0 else inverse (f w))"
+  from \<open>is_pole f z\<close> have "eventually (\<lambda>w. f w \<noteq> 0) (at z)"
+    unfolding is_pole_def using filterlim_at_infinity_imp_eventually_ne by blast
+  then obtain s'' where s'': "open s''" "z \<in> s''" "\<forall>w\<in>s''-{z}. f w \<noteq> 0"
+    by (auto simp: eventually_at_topological)
+  from assms(1) and s''(1) have "open (s \<inter> s'')" by auto
+  then obtain r where r: "r > 0" "ball z r \<subseteq> s \<inter> s''"
+    using assms(3) s''(2) by (subst (asm) open_contains_ball) blast
+  define s' where "s' = ball z r"
+  hence s': "open s'" "connected s'" "z \<in> s'" "s' \<subseteq> s" "\<forall>w\<in>s'-{z}. f w \<noteq> 0"
+    using r s'' by (auto simp: s'_def)
+  have s'_ne: "s' - {z} \<noteq> {}"
+    using r unfolding s'_def by (intro ball_minus_countable_nonempty) auto
+
+  have "porder f z = zorder g z"
+    by (simp add: porder_def g_def)
+  also have "\<dots> = 1"
+  proof (rule simple_zeroI')
+    fix w assume w: "w \<in> s'"
+    have [holomorphic_intros]: "g holomorphic_on s'" unfolding g_def using assms s'
+      by (intro is_pole_inverse_holomorphic holomorphic_on_subset[OF assms(5)]) auto
+    hence "(g has_field_derivative deriv g w) (at w)"
+      using w s' by (intro holomorphic_derivI)
+    also have deriv_g: "deriv g w = f' w" if "w \<in> s' - {z}" for w
+    proof -
+      from that have ne: "eventually (\<lambda>w. w \<noteq> z) (nhds w)"
+        by (intro t1_space_nhds) auto
+      have "deriv g w = deriv (\<lambda>w. inverse (f w)) w"
+        by (intro deriv_cong_ev refl eventually_mono [OF ne]) (auto simp: g_def)
+      also from assms(4)[of w] that s' have "\<dots> = f' w"
+        by (auto dest: DERIV_imp_deriv)
+      finally show ?thesis .
+    qed
+    have "deriv g w = f' w"
+      by (rule analytic_continuation_open[of "s' - {z}" s' "deriv g" f'])
+         (insert s' assms s'_ne deriv_g w, 
+          auto intro!: holomorphic_intros holomorphic_on_subset[OF assms(6)])
+    finally show "(g has_field_derivative f' w) (at w)" .
+  qed (insert assms s', auto simp: g_def)
+  finally show ?thesis .
+qed
+
+lemma residue_holomorphic_over_power:
+  assumes "open A" "z0 \<in> A" "f holomorphic_on A"
+  shows   "residue (\<lambda>z. f z / (z - z0) ^ Suc n) z0 = (deriv ^^ n) f z0 / fact n"
+proof -
+  let ?f = "\<lambda>z. f z / (z - z0) ^ Suc n"
+  from assms(1,2) obtain r where r: "r > 0" "cball z0 r \<subseteq> A"
+    by (auto simp: open_contains_cball)
+  have "(?f has_contour_integral 2 * pi * \<i> * residue ?f z0) (circlepath z0 r)"
+    using r assms by (intro base_residue[of A]) (auto intro!: holomorphic_intros)
+  moreover have "(?f has_contour_integral 2 * pi * \<i> / fact n * (deriv ^^ n) f z0) (circlepath z0 r)"
+    using assms r
+    by (intro Cauchy_has_contour_integral_higher_derivative_circlepath)
+       (auto intro!: holomorphic_on_subset[OF assms(3)] holomorphic_on_imp_continuous_on)
+  ultimately have "2 * pi * \<i> * residue ?f z0 = 2 * pi * \<i> / fact n * (deriv ^^ n) f z0"  
+    by (rule has_contour_integral_unique)
+  thus ?thesis by (simp add: field_simps)
+qed
+
+lemma residue_holomorphic_over_power':
+  assumes "open A" "0 \<in> A" "f holomorphic_on A"
+  shows   "residue (\<lambda>z. f z / z ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n"
+  using residue_holomorphic_over_power[OF assms] by simp
+
+lemma zer_poly_eqI:
+  fixes f :: "complex \<Rightarrow> complex" and z0 :: complex
+  defines "n \<equiv> zorder f z0"
+  assumes "open A" "connected A" "z0 \<in> A" "f holomorphic_on A" "f z0 = 0" "\<exists>z\<in>A. f z \<noteq> 0"
+  assumes lim: "((\<lambda>x. f (g x) / (g x - z0) ^ n) \<longlongrightarrow> c) F"
+  assumes g: "filterlim g (at z0) F" and "F \<noteq> bot"
+  shows   "zer_poly f z0 z0 = c"
+proof -
+  from zorder_exist[OF assms(2-7)] obtain r where
+    r: "r > 0" "cball z0 r \<subseteq> A" "zer_poly f z0 holomorphic_on cball z0 r"
+       "\<And>w. w \<in> cball z0 r \<Longrightarrow> f w = zer_poly f z0 w * (w - z0) ^ n"
+    unfolding n_def by blast
+  from r(1) have "eventually (\<lambda>w. w \<in> ball z0 r \<and> w \<noteq> z0) (at z0)"
+    using eventually_at_ball'[of r z0 UNIV] by auto
+  hence "eventually (\<lambda>w. zer_poly f z0 w = f w / (w - z0) ^ n) (at z0)"
+    by eventually_elim (insert r, auto simp: field_simps)
+  moreover have "continuous_on (ball z0 r) (zer_poly f z0)"
+    using r by (intro holomorphic_on_imp_continuous_on) auto
+  with r(1,2) have "isCont (zer_poly f z0) z0"
+    by (auto simp: continuous_on_eq_continuous_at)
+  hence "(zer_poly f z0 \<longlongrightarrow> zer_poly f z0 z0) (at z0)"
+    unfolding isCont_def .
+  ultimately have "((\<lambda>w. f w / (w - z0) ^ n) \<longlongrightarrow> zer_poly f z0 z0) (at z0)"
+    by (rule Lim_transform_eventually)
+  hence "((\<lambda>x. f (g x) / (g x - z0) ^ n) \<longlongrightarrow> zer_poly f z0 z0) F"
+    by (rule filterlim_compose[OF _ g])
+  from tendsto_unique[OF \<open>F \<noteq> bot\<close> this lim] show ?thesis .
+qed
+
+lemma pol_poly_eqI:
+  fixes f :: "complex \<Rightarrow> complex" and z0 :: complex
+  defines "n \<equiv> porder f z0"
+  assumes "open A" "z0 \<in> A" "f holomorphic_on A-{z0}" "is_pole f z0"
+  assumes lim: "((\<lambda>x. f (g x) * (g x - z0) ^ n) \<longlongrightarrow> c) F"
+  assumes g: "filterlim g (at z0) F" and "F \<noteq> bot"
+  shows   "pol_poly f z0 z0 = c"
+proof -
+  from porder_exist[OF assms(2-5)] obtain r where
+    r: "r > 0" "cball z0 r \<subseteq> A" "pol_poly f z0 holomorphic_on cball z0 r"
+       "\<And>w. w \<in> cball z0 r - {z0} \<Longrightarrow> f w = pol_poly f z0 w / (w - z0) ^ n"
+    unfolding n_def by blast
+  from r(1) have "eventually (\<lambda>w. w \<in> ball z0 r \<and> w \<noteq> z0) (at z0)"
+    using eventually_at_ball'[of r z0 UNIV] by auto
+  hence "eventually (\<lambda>w. pol_poly f z0 w = f w * (w - z0) ^ n) (at z0)"
+    by eventually_elim (insert r, auto simp: field_simps)
+  moreover have "continuous_on (ball z0 r) (pol_poly f z0)"
+    using r by (intro holomorphic_on_imp_continuous_on) auto
+  with r(1,2) have "isCont (pol_poly f z0) z0"
+    by (auto simp: continuous_on_eq_continuous_at)
+  hence "(pol_poly f z0 \<longlongrightarrow> pol_poly f z0 z0) (at z0)"
+    unfolding isCont_def .
+  ultimately have "((\<lambda>w. f w * (w - z0) ^ n) \<longlongrightarrow> pol_poly f z0 z0) (at z0)"
+    by (rule Lim_transform_eventually)
+  hence "((\<lambda>x. f (g x) * (g x - z0) ^ n) \<longlongrightarrow> pol_poly f z0 z0) F"
+    by (rule filterlim_compose[OF _ g])
+  from tendsto_unique[OF \<open>F \<noteq> bot\<close> this lim] show ?thesis .
+qed
+
+lemma residue_simple_pole:
+  assumes "open A" "z0 \<in> A" "f holomorphic_on A - {z0}" 
+  assumes "is_pole f z0" "porder f z0 = 1"
+  shows   "residue f z0 = pol_poly f z0 z0"
+  using assms by (subst residue_porder[of A]) simp_all
+
+lemma residue_simple_pole_limit:
+  assumes "open A" "z0 \<in> A" "f holomorphic_on A - {z0}" 
+  assumes "is_pole f z0" "porder f z0 = 1"
+  assumes "((\<lambda>x. f (g x) * (g x - z0)) \<longlongrightarrow> c) F"
+  assumes "filterlim g (at z0) F" "F \<noteq> bot"
+  shows   "residue f z0 = c"
+proof -
+  have "residue f z0 = pol_poly f z0 z0"
+    by (rule residue_simple_pole assms)+
+  also have "\<dots> = c"
+    using assms by (intro pol_poly_eqI[of A z0 f g c F]) auto
+  finally show ?thesis .
+qed
+
+(* TODO: This is a mess and could be done much more easily if we had
+   a nice compositional theory of poles and zeros *)
+lemma
+  assumes "open s" "connected s" "z \<in> s" "f holomorphic_on s" "g holomorphic_on s"
+  assumes "(g has_field_derivative g') (at z)"
+  assumes "f z \<noteq> 0" "g z = 0" "g' \<noteq> 0"
+  shows   porder_simple_pole_deriv: "porder (\<lambda>w. f w / g w) z = 1"
+    and   residue_simple_pole_deriv: "residue (\<lambda>w. f w / g w) z = f z / g'"
+proof -
+  have "\<exists>w\<in>s. g w \<noteq> 0"
+  proof (rule ccontr)
+    assume *: "\<not>(\<exists>w\<in>s. g w \<noteq> 0)"
+    have **: "eventually (\<lambda>w. w \<in> s) (nhds z)"
+      by (intro eventually_nhds_in_open assms)
+    from * have "deriv g z = deriv (\<lambda>_. 0) z"
+      by (intro deriv_cong_ev eventually_mono [OF **]) auto
+    also have "\<dots> = 0" by simp
+    also from assms have "deriv g z = g'" by (auto dest: DERIV_imp_deriv)
+    finally show False using \<open>g' \<noteq> 0\<close> by contradiction
+  qed
+  then obtain w where w: "w \<in> s" "g w \<noteq> 0" by blast
+  from isolated_zeros[OF assms(5) assms(1-3,8) w]
+  obtain r where r: "r > 0" "ball z r \<subseteq> s" "\<And>w. w \<in> ball z r - {z} \<Longrightarrow> g w \<noteq> 0"
+    by blast
+  from assms r have holo: "(\<lambda>w. f w / g w) holomorphic_on ball z r - {z}"
+    by (auto intro!: holomorphic_intros)
+
+  have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
+    using eventually_at_ball'[OF r(1), of z UNIV] by auto
+  hence "eventually (\<lambda>w. g w \<noteq> 0) (at z)"
+    by eventually_elim (use r in auto)
+  moreover have "continuous_on s g" 
+    by (intro holomorphic_on_imp_continuous_on) fact
+  with assms have "isCont g z"
+    by (auto simp: continuous_on_eq_continuous_at)
+  ultimately have "filterlim g (at 0) (at z)"
+    using \<open>g z = 0\<close> by (auto simp: filterlim_at isCont_def)
+  moreover have "continuous_on s f" by (intro holomorphic_on_imp_continuous_on) fact
+  with assms have "isCont f z"
+    by (auto simp: continuous_on_eq_continuous_at)
+  ultimately have pole: "is_pole (\<lambda>w. f w / g w) z" 
+    unfolding is_pole_def using \<open>f z \<noteq> 0\<close>
+    by (intro filterlim_divide_at_infinity[of _ "f z"]) (auto simp: isCont_def)
+
+  have "continuous_on s f" by (intro holomorphic_on_imp_continuous_on) fact
+  moreover have "open (-{0::complex})" by auto
+  ultimately have "open (f -` (-{0}) \<inter> s)" using \<open>open s\<close>
+    by (subst (asm) continuous_on_open_vimage) blast+
+  moreover have "z \<in> f -` (-{0}) \<inter> s" using assms by auto
+  ultimately obtain r' where r': "r' > 0" "ball z r' \<subseteq> f -` (-{0}) \<inter> s"
+    unfolding open_contains_ball by blast
+
+  let ?D = "\<lambda>w. (f w * deriv g w - g w * deriv f w) / f w ^ 2"
+  show "porder (\<lambda>w. f w / g w) z = 1"
+  proof (rule simple_poleI')
+    show "open (ball z (min r r'))" "connected (ball z (min r r'))" "z \<in> ball z (min r r')"
+      using r'(1) r(1) by auto
+  next
+    fix w assume "w \<in> ball z (min r r') - {z}"
+    with r' have "w \<in> s" "f w \<noteq> 0" by auto
+    have "((\<lambda>w. g w / f w) has_field_derivative ?D w) (at w)"
+      by (rule derivative_eq_intros holomorphic_derivI[OF assms(4)] 
+            holomorphic_derivI[OF assms(5)] | fact)+ 
+         (simp_all add: algebra_simps power2_eq_square)
+    thus "((\<lambda>w. inverse (f w / g w)) has_field_derivative ?D w) (at w)"
+      by (simp add: divide_simps)
+  next
+    from r' show "?D holomorphic_on ball z (min r r')"
+      by (intro holomorphic_intros holomorphic_on_subset[OF assms(4)]
+                holomorphic_on_subset[OF assms(5)]) auto
+  next
+    from assms have "deriv g z = g'"
+      by (auto dest: DERIV_imp_deriv)
+    with assms r' show "(f z * deriv g z - g z * deriv f z) / (f z)\<^sup>2 \<noteq> 0"
+      by simp
+  qed (insert pole holo, auto)
+
+  show "residue (\<lambda>w. f w / g w) z = f z / g'"
+  proof (rule residue_simple_pole_limit)
+    show "porder (\<lambda>w. f w / g w) z = 1" by fact
+    from r show "open (ball z r)" "z \<in> ball z r" by auto
+    
+    have "((\<lambda>w. (f w * (w - z)) / g w) \<longlongrightarrow> f z / g') (at z)"
+    proof (rule lhopital_complex_simple)
+      show "((\<lambda>w. f w * (w - z)) has_field_derivative f z) (at z)"
+        using assms by (auto intro!: derivative_eq_intros holomorphic_derivI[OF assms(4)])
+      show "(g has_field_derivative g') (at z)" by fact
+    qed (insert assms, auto)
+    thus "((\<lambda>w. (f w / g w) * (w - z)) \<longlongrightarrow> f z / g') (at z)"
+      by (simp add: divide_simps)
+  qed (insert holo pole, auto simp: filterlim_ident)
+qed
+
 end
--- a/src/HOL/Analysis/Extended_Real_Limits.thy	Fri Aug 18 22:55:54 2017 +0200
+++ b/src/HOL/Analysis/Extended_Real_Limits.thy	Sun Aug 20 03:35:20 2017 +0200
@@ -358,6 +358,739 @@
   apply (metis INF_absorb centre_in_ball)
   done
 
+subsection \<open>Fun.thy\<close>
+
+lemma inj_fn:
+  fixes f::"'a \<Rightarrow> 'a"
+  assumes "inj f"
+  shows "inj (f^^n)"
+proof (induction n)
+  case (Suc n)
+  have "inj (f o (f^^n))"
+    using inj_comp[OF assms Suc.IH] by simp
+  then show "inj (f^^(Suc n))"
+    by auto
+qed (auto)
+
+lemma surj_fn:
+  fixes f::"'a \<Rightarrow> 'a"
+  assumes "surj f"
+  shows "surj (f^^n)"
+proof (induction n)
+  case (Suc n)
+  have "surj (f o (f^^n))"
+    using assms Suc.IH by (simp add: comp_surj)
+  then show "surj (f^^(Suc n))"
+    by auto
+qed (auto)
+
+lemma bij_fn:
+  fixes f::"'a \<Rightarrow> 'a"
+  assumes "bij f"
+  shows "bij (f^^n)"
+by (rule bijI[OF inj_fn[OF bij_is_inj[OF assms]] surj_fn[OF bij_is_surj[OF assms]]])
+
+lemma inv_fn_o_fn_is_id:
+  fixes f::"'a \<Rightarrow> 'a"
+  assumes "bij f"
+  shows "((inv f)^^n) o (f^^n) = (\<lambda>x. x)"
+proof -
+  have "((inv f)^^n)((f^^n) x) = x" for x n
+  proof (induction n)
+    case (Suc n)
+    have *: "(inv f) (f y) = y" for y
+      by (simp add: assms bij_is_inj)
+    have "(inv f ^^ Suc n) ((f ^^ Suc n) x) = (inv f^^n) (inv f (f ((f^^n) x)))"
+      by (simp add: funpow_swap1)
+    also have "... = (inv f^^n) ((f^^n) x)"
+      using * by auto
+    also have "... = x" using Suc.IH by auto
+    finally show ?case by simp
+  qed (auto)
+  then show ?thesis unfolding o_def by blast
+qed
+
+lemma bij_inv_eq_iff: "bij p \<Longrightarrow> x = inv p y \<longleftrightarrow> p x = y" (* COPIED FROM Permutations *)
+  using surj_f_inv_f[of p] by (auto simp add: bij_def)
+
+lemma fn_o_inv_fn_is_id:
+  fixes f::"'a \<Rightarrow> 'a"
+  assumes "bij f"
+  shows "(f^^n) o ((inv f)^^n) = (\<lambda>x. x)"
+proof -
+  have "(f^^n) (((inv f)^^n) x) = x" for x n
+  proof (induction n)
+    case (Suc n)
+    have *: "f(inv f y) = y" for y
+      using assms by (meson bij_inv_eq_iff)
+    have "(f ^^ Suc n) ((inv f ^^ Suc n) x) = (f^^n) (f (inv f ((inv f^^n) x)))"
+      by (simp add: funpow_swap1)
+    also have "... = (f^^n) ((inv f^^n) x)"
+      using * by auto
+    also have "... = x" using Suc.IH by auto
+    finally show ?case by simp
+  qed (auto)
+  then show ?thesis unfolding o_def by blast
+qed
+
+lemma inv_fn:
+  fixes f::"'a \<Rightarrow> 'a"
+  assumes "bij f"
+  shows "inv (f^^n) = ((inv f)^^n)"
+proof -
+  have "inv (f^^n) x = ((inv f)^^n) x" for x
+  apply (rule inv_into_f_eq, auto simp add: inj_fn[OF bij_is_inj[OF assms]])
+  using fn_o_inv_fn_is_id[OF assms, of n] by (metis comp_apply)
+  then show ?thesis by auto
+qed
+
+lemma mono_inv:
+  fixes f::"'a::linorder \<Rightarrow> 'b::linorder"
+  assumes "mono f" "bij f"
+  shows "mono (inv f)"
+proof
+  fix x y::'b assume "x \<le> y"
+  then show "inv f x \<le> inv f y"
+    by (metis (no_types, lifting) assms bij_is_surj eq_iff le_cases mono_def surj_f_inv_f)
+qed
+
+lemma mono_bij_Inf:
+  fixes f :: "'a::complete_linorder \<Rightarrow> 'b::complete_linorder"
+  assumes "mono f" "bij f"
+  shows "f (Inf A) = Inf (f`A)"
+proof -
+  have "(inv f) (Inf (f`A)) \<le> Inf ((inv f)`(f`A))"
+    using mono_Inf[OF mono_inv[OF assms], of "f`A"] by simp
+  then have "Inf (f`A) \<le> f (Inf ((inv f)`(f`A)))"
+    by (metis (no_types, lifting) assms mono_def bij_inv_eq_iff)
+  also have "... = f(Inf A)"
+    using assms by (simp add: bij_is_inj)
+  finally show ?thesis using mono_Inf[OF assms(1), of A] by auto
+qed
+
+
+lemma Inf_nat_def1:
+  fixes K::"nat set"
+  assumes "K \<noteq> {}"
+  shows "Inf K \<in> K"
+by (auto simp add: Min_def Inf_nat_def) (meson LeastI assms bot.extremum_unique subsetI)
+
+
+subsection \<open>Extended-Real.thy\<close>
+
+text\<open>The proof of this one is copied from \verb+ereal_add_mono+.\<close>
+lemma ereal_add_strict_mono2:
+  fixes a b c d :: ereal
+  assumes "a < b"
+    and "c < d"
+  shows "a + c < b + d"
+using assms
+apply (cases a)
+apply (cases rule: ereal3_cases[of b c d], auto)
+apply (cases rule: ereal3_cases[of b c d], auto)
+done
+
+text \<open>The next ones are analogues of \verb+mult_mono+ and \verb+mult_mono'+ in ereal.\<close>
+
+lemma ereal_mult_mono:
+  fixes a b c d::ereal
+  assumes "b \<ge> 0" "c \<ge> 0" "a \<le> b" "c \<le> d"
+  shows "a * c \<le> b * d"
+by (metis ereal_mult_right_mono mult.commute order_trans assms)
+
+lemma ereal_mult_mono':
+  fixes a b c d::ereal
+  assumes "a \<ge> 0" "c \<ge> 0" "a \<le> b" "c \<le> d"
+  shows "a * c \<le> b * d"
+by (metis ereal_mult_right_mono mult.commute order_trans assms)
+
+lemma ereal_mult_mono_strict:
+  fixes a b c d::ereal
+  assumes "b > 0" "c > 0" "a < b" "c < d"
+  shows "a * c < b * d"
+proof -
+  have "c < \<infinity>" using \<open>c < d\<close> by auto
+  then have "a * c < b * c" by (metis ereal_mult_strict_left_mono[OF assms(3) assms(2)] mult.commute)
+  moreover have "b * c \<le> b * d" using assms(2) assms(4) by (simp add: assms(1) ereal_mult_left_mono less_imp_le)
+  ultimately show ?thesis by simp
+qed
+
+lemma ereal_mult_mono_strict':
+  fixes a b c d::ereal
+  assumes "a > 0" "c > 0" "a < b" "c < d"
+  shows "a * c < b * d"
+apply (rule ereal_mult_mono_strict, auto simp add: assms) using assms by auto
+
+lemma ereal_abs_add:
+  fixes a b::ereal
+  shows "abs(a+b) \<le> abs a + abs b"
+by (cases rule: ereal2_cases[of a b]) (auto)
+
+lemma ereal_abs_diff:
+  fixes a b::ereal
+  shows "abs(a-b) \<le> abs a + abs b"
+by (cases rule: ereal2_cases[of a b]) (auto)
+
+lemma sum_constant_ereal:
+  fixes a::ereal
+  shows "(\<Sum>i\<in>I. a) = a * card I"
+apply (cases "finite I", induct set: finite, simp_all)
+apply (cases a, auto, metis (no_types, hide_lams) add.commute mult.commute semiring_normalization_rules(3))
+done
+
+lemma real_lim_then_eventually_real:
+  assumes "(u \<longlongrightarrow> ereal l) F"
+  shows "eventually (\<lambda>n. u n = ereal(real_of_ereal(u n))) F"
+proof -
+  have "ereal l \<in> {-\<infinity><..<(\<infinity>::ereal)}" by simp
+  moreover have "open {-\<infinity><..<(\<infinity>::ereal)}" by simp
+  ultimately have "eventually (\<lambda>n. u n \<in> {-\<infinity><..<(\<infinity>::ereal)}) F" using assms tendsto_def by blast
+  moreover have "\<And>x. x \<in> {-\<infinity><..<(\<infinity>::ereal)} \<Longrightarrow> x = ereal(real_of_ereal x)" using ereal_real by auto
+  ultimately show ?thesis by (metis (mono_tags, lifting) eventually_mono)
+qed
+
+lemma ereal_Inf_cmult:
+  assumes "c>(0::real)"
+  shows "Inf {ereal c * x |x. P x} = ereal c * Inf {x. P x}"
+proof -
+  have "(\<lambda>x::ereal. c * x) (Inf {x::ereal. P x}) = Inf ((\<lambda>x::ereal. c * x)`{x::ereal. P x})"
+    apply (rule mono_bij_Inf)
+    apply (simp add: assms ereal_mult_left_mono less_imp_le mono_def)
+    apply (rule bij_betw_byWitness[of _ "\<lambda>x. (x::ereal) / c"], auto simp add: assms ereal_mult_divide)
+    using assms ereal_divide_eq apply auto
+    done
+  then show ?thesis by (simp only: setcompr_eq_image[symmetric])
+qed
+
+
+subsubsection \<open>Continuity of addition\<close>
+
+text \<open>The next few lemmas remove an unnecessary assumption in \verb+tendsto_add_ereal+, culminating
+in \verb+tendsto_add_ereal_general+ which essentially says that the addition
+is continuous on ereal times ereal, except at $(-\infty, \infty)$ and $(\infty, -\infty)$.
+It is much more convenient in many situations, see for instance the proof of
+\verb+tendsto_sum_ereal+ below.\<close>
+
+lemma tendsto_add_ereal_PInf:
+  fixes y :: ereal
+  assumes y: "y \<noteq> -\<infinity>"
+  assumes f: "(f \<longlongrightarrow> \<infinity>) F" and g: "(g \<longlongrightarrow> y) F"
+  shows "((\<lambda>x. f x + g x) \<longlongrightarrow> \<infinity>) F"
+proof -
+  have "\<exists>C. eventually (\<lambda>x. g x > ereal C) F"
+  proof (cases y)
+    case (real r)
+    have "y > y-1" using y real by (simp add: ereal_between(1))
+    then have "eventually (\<lambda>x. g x > y - 1) F" using g y order_tendsto_iff by auto
+    moreover have "y-1 = ereal(real_of_ereal(y-1))"
+      by (metis real ereal_eq_1(1) ereal_minus(1) real_of_ereal.simps(1))
+    ultimately have "eventually (\<lambda>x. g x > ereal(real_of_ereal(y - 1))) F" by simp
+    then show ?thesis by auto
+  next
+    case (PInf)
+    have "eventually (\<lambda>x. g x > ereal 0) F" using g PInf by (simp add: tendsto_PInfty)
+    then show ?thesis by auto
+  qed (simp add: y)
+  then obtain C::real where ge: "eventually (\<lambda>x. g x > ereal C) F" by auto
+
+  {
+    fix M::real
+    have "eventually (\<lambda>x. f x > ereal(M - C)) F" using f by (simp add: tendsto_PInfty)
+    then have "eventually (\<lambda>x. (f x > ereal (M-C)) \<and> (g x > ereal C)) F"
+      by (auto simp add: ge eventually_conj_iff)
+    moreover have "\<And>x. ((f x > ereal (M-C)) \<and> (g x > ereal C)) \<Longrightarrow> (f x + g x > ereal M)"
+      using ereal_add_strict_mono2 by fastforce
+    ultimately have "eventually (\<lambda>x. f x + g x > ereal M) F" using eventually_mono by force
+  }
+  then show ?thesis by (simp add: tendsto_PInfty)
+qed
+
+text\<open>One would like to deduce the next lemma from the previous one, but the fact
+that $-(x+y)$ is in general different from $(-x) + (-y)$ in ereal creates difficulties,
+so it is more efficient to copy the previous proof.\<close>
+
+lemma tendsto_add_ereal_MInf:
+  fixes y :: ereal
+  assumes y: "y \<noteq> \<infinity>"
+  assumes f: "(f \<longlongrightarrow> -\<infinity>) F" and g: "(g \<longlongrightarrow> y) F"
+  shows "((\<lambda>x. f x + g x) \<longlongrightarrow> -\<infinity>) F"
+proof -
+  have "\<exists>C. eventually (\<lambda>x. g x < ereal C) F"
+  proof (cases y)
+    case (real r)
+    have "y < y+1" using y real by (simp add: ereal_between(1))
+    then have "eventually (\<lambda>x. g x < y + 1) F" using g y order_tendsto_iff by force
+    moreover have "y+1 = ereal(real_of_ereal (y+1))" by (simp add: real)
+    ultimately have "eventually (\<lambda>x. g x < ereal(real_of_ereal(y + 1))) F" by simp
+    then show ?thesis by auto
+  next
+    case (MInf)
+    have "eventually (\<lambda>x. g x < ereal 0) F" using g MInf by (simp add: tendsto_MInfty)
+    then show ?thesis by auto
+  qed (simp add: y)
+  then obtain C::real where ge: "eventually (\<lambda>x. g x < ereal C) F" by auto
+
+  {
+    fix M::real
+    have "eventually (\<lambda>x. f x < ereal(M - C)) F" using f by (simp add: tendsto_MInfty)
+    then have "eventually (\<lambda>x. (f x < ereal (M- C)) \<and> (g x < ereal C)) F"
+      by (auto simp add: ge eventually_conj_iff)
+    moreover have "\<And>x. ((f x < ereal (M-C)) \<and> (g x < ereal C)) \<Longrightarrow> (f x + g x < ereal M)"
+      using ereal_add_strict_mono2 by fastforce
+    ultimately have "eventually (\<lambda>x. f x + g x < ereal M) F" using eventually_mono by force
+  }
+  then show ?thesis by (simp add: tendsto_MInfty)
+qed
+
+lemma tendsto_add_ereal_general1:
+  fixes x y :: ereal
+  assumes y: "\<bar>y\<bar> \<noteq> \<infinity>"
+  assumes f: "(f \<longlongrightarrow> x) F" and g: "(g \<longlongrightarrow> y) F"
+  shows "((\<lambda>x. f x + g x) \<longlongrightarrow> x + y) F"
+proof (cases x)
+  case (real r)
+  have a: "\<bar>x\<bar> \<noteq> \<infinity>" by (simp add: real)
+  show ?thesis by (rule tendsto_add_ereal[OF a, OF y, OF f, OF g])
+next
+  case PInf
+  then show ?thesis using tendsto_add_ereal_PInf assms by force
+next
+  case MInf
+  then show ?thesis using tendsto_add_ereal_MInf assms
+    by (metis abs_ereal.simps(3) ereal_MInfty_eq_plus)
+qed
+
+lemma tendsto_add_ereal_general2:
+  fixes x y :: ereal
+  assumes x: "\<bar>x\<bar> \<noteq> \<infinity>"
+      and f: "(f \<longlongrightarrow> x) F" and g: "(g \<longlongrightarrow> y) F"
+  shows "((\<lambda>x. f x + g x) \<longlongrightarrow> x + y) F"
+proof -
+  have "((\<lambda>x. g x + f x) \<longlongrightarrow> x + y) F"
+    using tendsto_add_ereal_general1[OF x, OF g, OF f] add.commute[of "y", of "x"] by simp
+  moreover have "\<And>x. g x + f x = f x + g x" using add.commute by auto
+  ultimately show ?thesis by simp
+qed
+
+text \<open>The next lemma says that the addition is continuous on ereal, except at
+the pairs $(-\infty, \infty)$ and $(\infty, -\infty)$.\<close>
+
+lemma tendsto_add_ereal_general [tendsto_intros]:
+  fixes x y :: ereal
+  assumes "\<not>((x=\<infinity> \<and> y=-\<infinity>) \<or> (x=-\<infinity> \<and> y=\<infinity>))"
+      and f: "(f \<longlongrightarrow> x) F" and g: "(g \<longlongrightarrow> y) F"
+  shows "((\<lambda>x. f x + g x) \<longlongrightarrow> x + y) F"
+proof (cases x)
+  case (real r)
+  show ?thesis
+    apply (rule tendsto_add_ereal_general2) using real assms by auto
+next
+  case (PInf)
+  then have "y \<noteq> -\<infinity>" using assms by simp
+  then show ?thesis using tendsto_add_ereal_PInf PInf assms by auto
+next
+  case (MInf)
+  then have "y \<noteq> \<infinity>" using assms by simp
+  then show ?thesis using tendsto_add_ereal_MInf MInf f g by (metis ereal_MInfty_eq_plus)
+qed
+
+subsubsection \<open>Continuity of multiplication\<close>
+
+text \<open>In the same way as for addition, we prove that the multiplication is continuous on
+ereal times ereal, except at $(\infty, 0)$ and $(-\infty, 0)$ and $(0, \infty)$ and $(0, -\infty)$,
+starting with specific situations.\<close>
+
+lemma tendsto_mult_real_ereal:
+  assumes "(u \<longlongrightarrow> ereal l) F" "(v \<longlongrightarrow> ereal m) F"
+  shows "((\<lambda>n. u n * v n) \<longlongrightarrow> ereal l * ereal m) F"
+proof -
+  have ureal: "eventually (\<lambda>n. u n = ereal(real_of_ereal(u n))) F" by (rule real_lim_then_eventually_real[OF assms(1)])
+  then have "((\<lambda>n. ereal(real_of_ereal(u n))) \<longlongrightarrow> ereal l) F" using assms by auto
+  then have limu: "((\<lambda>n. real_of_ereal(u n)) \<longlongrightarrow> l) F" by auto
+  have vreal: "eventually (\<lambda>n. v n = ereal(real_of_ereal(v n))) F" by (rule real_lim_then_eventually_real[OF assms(2)])
+  then have "((\<lambda>n. ereal(real_of_ereal(v n))) \<longlongrightarrow> ereal m) F" using assms by auto
+  then have limv: "((\<lambda>n. real_of_ereal(v n)) \<longlongrightarrow> m) F" by auto
+
+  {
+    fix n assume "u n = ereal(real_of_ereal(u n))" "v n = ereal(real_of_ereal(v n))"
+    then have "ereal(real_of_ereal(u n) * real_of_ereal(v n)) = u n * v n" by (metis times_ereal.simps(1))
+  }
+  then have *: "eventually (\<lambda>n. ereal(real_of_ereal(u n) * real_of_ereal(v n)) = u n * v n) F"
+    using eventually_elim2[OF ureal vreal] by auto
+
+  have "((\<lambda>n. real_of_ereal(u n) * real_of_ereal(v n)) \<longlongrightarrow> l * m) F" using tendsto_mult[OF limu limv] by auto
+  then have "((\<lambda>n. ereal(real_of_ereal(u n)) * real_of_ereal(v n)) \<longlongrightarrow> ereal(l * m)) F" by auto
+  then show ?thesis using * filterlim_cong by fastforce
+qed
+
+lemma tendsto_mult_ereal_PInf:
+  fixes f g::"_ \<Rightarrow> ereal"
+  assumes "(f \<longlongrightarrow> l) F" "l>0" "(g \<longlongrightarrow> \<infinity>) F"
+  shows "((\<lambda>x. f x * g x) \<longlongrightarrow> \<infinity>) F"
+proof -
+  obtain a::real where "0 < ereal a" "a < l" using assms(2) using ereal_dense2 by blast
+  have *: "eventually (\<lambda>x. f x > a) F" using \<open>a < l\<close> assms(1) by (simp add: order_tendsto_iff)
+  {
+    fix K::real
+    define M where "M = max K 1"
+    then have "M > 0" by simp
+    then have "ereal(M/a) > 0" using \<open>ereal a > 0\<close> by simp
+    then have "\<And>x. ((f x > a) \<and> (g x > M/a)) \<Longrightarrow> (f x * g x > ereal a * ereal(M/a))"
+      using ereal_mult_mono_strict'[where ?c = "M/a", OF \<open>0 < ereal a\<close>] by auto
+    moreover have "ereal a * ereal(M/a) = M" using \<open>ereal a > 0\<close> by simp
+    ultimately have "\<And>x. ((f x > a) \<and> (g x > M/a)) \<Longrightarrow> (f x * g x > M)" by simp
+    moreover have "M \<ge> K" unfolding M_def by simp
+    ultimately have Imp: "\<And>x. ((f x > a) \<and> (g x > M/a)) \<Longrightarrow> (f x * g x > K)"
+      using ereal_less_eq(3) le_less_trans by blast
+
+    have "eventually (\<lambda>x. g x > M/a) F" using assms(3) by (simp add: tendsto_PInfty)
+    then have "eventually (\<lambda>x. (f x > a) \<and> (g x > M/a)) F"
+      using * by (auto simp add: eventually_conj_iff)
+    then have "eventually (\<lambda>x. f x * g x > K) F" using eventually_mono Imp by force
+  }
+  then show ?thesis by (auto simp add: tendsto_PInfty)
+qed
+
+lemma tendsto_mult_ereal_pos:
+  fixes f g::"_ \<Rightarrow> ereal"
+  assumes "(f \<longlongrightarrow> l) F" "(g \<longlongrightarrow> m) F" "l>0" "m>0"
+  shows "((\<lambda>x. f x * g x) \<longlongrightarrow> l * m) F"
+proof (cases)
+  assume *: "l = \<infinity> \<or> m = \<infinity>"
+  then show ?thesis
+  proof (cases)
+    assume "m = \<infinity>"
+    then show ?thesis using tendsto_mult_ereal_PInf assms by auto
+  next
+    assume "\<not>(m = \<infinity>)"
+    then have "l = \<infinity>" using * by simp
+    then have "((\<lambda>x. g x * f x) \<longlongrightarrow> l * m) F" using tendsto_mult_ereal_PInf assms by auto
+    moreover have "\<And>x. g x * f x = f x * g x" using mult.commute by auto
+    ultimately show ?thesis by simp
+  qed
+next
+  assume "\<not>(l = \<infinity> \<or> m = \<infinity>)"
+  then have "l < \<infinity>" "m < \<infinity>" by auto
+  then obtain lr mr where "l = ereal lr" "m = ereal mr"
+    using \<open>l>0\<close> \<open>m>0\<close> by (metis ereal_cases ereal_less(6) not_less_iff_gr_or_eq)
+  then show ?thesis using tendsto_mult_real_ereal assms by auto
+qed
+
+text \<open>We reduce the general situation to the positive case by multiplying by suitable signs.
+Unfortunately, as ereal is not a ring, all the neat sign lemmas are not available there. We
+give the bare minimum we need.\<close>
+
+lemma ereal_sgn_abs:
+  fixes l::ereal
+  shows "sgn(l) * l = abs(l)"
+apply (cases l) by (auto simp add: sgn_if ereal_less_uminus_reorder)
+
+lemma sgn_squared_ereal:
+  assumes "l \<noteq> (0::ereal)"
+  shows "sgn(l) * sgn(l) = 1"
+apply (cases l) using assms by (auto simp add: one_ereal_def sgn_if)
+
+lemma tendsto_mult_ereal [tendsto_intros]:
+  fixes f g::"_ \<Rightarrow> ereal"
+  assumes "(f \<longlongrightarrow> l) F" "(g \<longlongrightarrow> m) F" "\<not>((l=0 \<and> abs(m) = \<infinity>) \<or> (m=0 \<and> abs(l) = \<infinity>))"
+  shows "((\<lambda>x. f x * g x) \<longlongrightarrow> l * m) F"
+proof (cases)
+  assume "l=0 \<or> m=0"
+  then have "abs(l) \<noteq> \<infinity>" "abs(m) \<noteq> \<infinity>" using assms(3) by auto
+  then obtain lr mr where "l = ereal lr" "m = ereal mr" by auto
+  then show ?thesis using tendsto_mult_real_ereal assms by auto
+next
+  have sgn_finite: "\<And>a::ereal. abs(sgn a) \<noteq> \<infinity>"
+    by (metis MInfty_neq_ereal(2) PInfty_neq_ereal(2) abs_eq_infinity_cases ereal_times(1) ereal_times(3) ereal_uminus_eq_reorder sgn_ereal.elims)
+  then have sgn_finite2: "\<And>a b::ereal. abs(sgn a * sgn b) \<noteq> \<infinity>"
+    by (metis abs_eq_infinity_cases abs_ereal.simps(2) abs_ereal.simps(3) ereal_mult_eq_MInfty ereal_mult_eq_PInfty)
+  assume "\<not>(l=0 \<or> m=0)"
+  then have "l \<noteq> 0" "m \<noteq> 0" by auto
+  then have "abs(l) > 0" "abs(m) > 0"
+    by (metis abs_ereal_ge0 abs_ereal_less0 abs_ereal_pos ereal_uminus_uminus ereal_uminus_zero less_le not_less)+
+  then have "sgn(l) * l > 0" "sgn(m) * m > 0" using ereal_sgn_abs by auto
+  moreover have "((\<lambda>x. sgn(l) * f x) \<longlongrightarrow> (sgn(l) * l)) F"
+    by (rule tendsto_cmult_ereal, auto simp add: sgn_finite assms(1))
+  moreover have "((\<lambda>x. sgn(m) * g x) \<longlongrightarrow> (sgn(m) * m)) F"
+    by (rule tendsto_cmult_ereal, auto simp add: sgn_finite assms(2))
+  ultimately have *: "((\<lambda>x. (sgn(l) * f x) * (sgn(m) * g x)) \<longlongrightarrow> (sgn(l) * l) * (sgn(m) * m)) F"
+    using tendsto_mult_ereal_pos by force
+  have "((\<lambda>x. (sgn(l) * sgn(m)) * ((sgn(l) * f x) * (sgn(m) * g x))) \<longlongrightarrow> (sgn(l) * sgn(m)) * ((sgn(l) * l) * (sgn(m) * m))) F"
+    by (rule tendsto_cmult_ereal, auto simp add: sgn_finite2 *)
+  moreover have "\<And>x. (sgn(l) * sgn(m)) * ((sgn(l) * f x) * (sgn(m) * g x)) = f x * g x"
+    by (metis mult.left_neutral sgn_squared_ereal[OF \<open>l \<noteq> 0\<close>] sgn_squared_ereal[OF \<open>m \<noteq> 0\<close>] mult.assoc mult.commute)
+  moreover have "(sgn(l) * sgn(m)) * ((sgn(l) * l) * (sgn(m) * m)) = l * m"
+    by (metis mult.left_neutral sgn_squared_ereal[OF \<open>l \<noteq> 0\<close>] sgn_squared_ereal[OF \<open>m \<noteq> 0\<close>] mult.assoc mult.commute)
+  ultimately show ?thesis by auto
+qed
+
+lemma tendsto_cmult_ereal_general [tendsto_intros]:
+  fixes f::"_ \<Rightarrow> ereal" and c::ereal
+  assumes "(f \<longlongrightarrow> l) F" "\<not> (l=0 \<and> abs(c) = \<infinity>)"
+  shows "((\<lambda>x. c * f x) \<longlongrightarrow> c * l) F"
+by (cases "c = 0", auto simp add: assms tendsto_mult_ereal)
+
+
+subsubsection \<open>Continuity of division\<close>
+
+lemma tendsto_inverse_ereal_PInf:
+  fixes u::"_ \<Rightarrow> ereal"
+  assumes "(u \<longlongrightarrow> \<infinity>) F"
+  shows "((\<lambda>x. 1/ u x) \<longlongrightarrow> 0) F"
+proof -
+  {
+    fix e::real assume "e>0"
+    have "1/e < \<infinity>" by auto
+    then have "eventually (\<lambda>n. u n > 1/e) F" using assms(1) by (simp add: tendsto_PInfty)
+    moreover
+    {
+      fix z::ereal assume "z>1/e"
+      then have "z>0" using \<open>e>0\<close> using less_le_trans not_le by fastforce
+      then have "1/z \<ge> 0" by auto
+      moreover have "1/z < e" using \<open>e>0\<close> \<open>z>1/e\<close>
+        apply (cases z) apply auto
+        by (metis (mono_tags, hide_lams) less_ereal.simps(2) less_ereal.simps(4) divide_less_eq ereal_divide_less_pos ereal_less(4)
+            ereal_less_eq(4) less_le_trans mult_eq_0_iff not_le not_one_less_zero times_ereal.simps(1))
+      ultimately have "1/z \<ge> 0" "1/z < e" by auto
+    }
+    ultimately have "eventually (\<lambda>n. 1/u n<e) F" "eventually (\<lambda>n. 1/u n\<ge>0) F" by (auto simp add: eventually_mono)
+  } note * = this
+  show ?thesis
+  proof (subst order_tendsto_iff, auto)
+    fix a::ereal assume "a<0"
+    then show "eventually (\<lambda>n. 1/u n > a) F" using *(2) eventually_mono less_le_trans linordered_field_no_ub by fastforce
+  next
+    fix a::ereal assume "a>0"
+    then obtain e::real where "e>0" "a>e" using ereal_dense2 ereal_less(2) by blast
+    then have "eventually (\<lambda>n. 1/u n < e) F" using *(1) by auto
+    then show "eventually (\<lambda>n. 1/u n < a) F" using \<open>a>e\<close> by (metis (mono_tags, lifting) eventually_mono less_trans)
+  qed
+qed
+
+text \<open>The next lemma deserves to exist by itself, as it is so common and useful.\<close>
+
+lemma tendsto_inverse_real [tendsto_intros]:
+  fixes u::"_ \<Rightarrow> real"
+  shows "(u \<longlongrightarrow> l) F \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> ((\<lambda>x. 1/ u x) \<longlongrightarrow> 1/l) F"
+  using tendsto_inverse unfolding inverse_eq_divide .
+
+lemma tendsto_inverse_ereal [tendsto_intros]:
+  fixes u::"_ \<Rightarrow> ereal"
+  assumes "(u \<longlongrightarrow> l) F" "l \<noteq> 0"
+  shows "((\<lambda>x. 1/ u x) \<longlongrightarrow> 1/l) F"
+proof (cases l)
+  case (real r)
+  then have "r \<noteq> 0" using assms(2) by auto
+  then have "1/l = ereal(1/r)" using real by (simp add: one_ereal_def)
+  define v where "v = (\<lambda>n. real_of_ereal(u n))"
+  have ureal: "eventually (\<lambda>n. u n = ereal(v n)) F" unfolding v_def using real_lim_then_eventually_real assms(1) real by auto
+  then have "((\<lambda>n. ereal(v n)) \<longlongrightarrow> ereal r) F" using assms real v_def by auto
+  then have *: "((\<lambda>n. v n) \<longlongrightarrow> r) F" by auto
+  then have "((\<lambda>n. 1/v n) \<longlongrightarrow> 1/r) F" using \<open>r \<noteq> 0\<close> tendsto_inverse_real by auto
+  then have lim: "((\<lambda>n. ereal(1/v n)) \<longlongrightarrow> 1/l) F" using \<open>1/l = ereal(1/r)\<close> by auto
+
+  have "r \<in> -{0}" "open (-{(0::real)})" using \<open>r \<noteq> 0\<close> by auto
+  then have "eventually (\<lambda>n. v n \<in> -{0}) F" using * using topological_tendstoD by blast
+  then have "eventually (\<lambda>n. v n \<noteq> 0) F" by auto
+  moreover
+  {
+    fix n assume H: "v n \<noteq> 0" "u n = ereal(v n)"
+    then have "ereal(1/v n) = 1/ereal(v n)" by (simp add: one_ereal_def)
+    then have "ereal(1/v n) = 1/u n" using H(2) by simp
+  }
+  ultimately have "eventually (\<lambda>n. ereal(1/v n) = 1/u n) F" using ureal eventually_elim2 by force
+  with Lim_transform_eventually[OF this lim] show ?thesis by simp
+next
+  case (PInf)
+  then have "1/l = 0" by auto
+  then show ?thesis using tendsto_inverse_ereal_PInf assms PInf by auto
+next
+  case (MInf)
+  then have "1/l = 0" by auto
+  have "1/z = -1/ -z" if "z < 0" for z::ereal
+    apply (cases z) using divide_ereal_def \<open> z < 0 \<close> by auto
+  moreover have "eventually (\<lambda>n. u n < 0) F" by (metis (no_types) MInf assms(1) tendsto_MInfty zero_ereal_def)
+  ultimately have *: "eventually (\<lambda>n. -1/-u n = 1/u n) F" by (simp add: eventually_mono)
+
+  define v where "v = (\<lambda>n. - u n)"
+  have "(v \<longlongrightarrow> \<infinity>) F" unfolding v_def using MInf assms(1) tendsto_uminus_ereal by fastforce
+  then have "((\<lambda>n. 1/v n) \<longlongrightarrow> 0) F" using tendsto_inverse_ereal_PInf by auto
+  then have "((\<lambda>n. -1/v n) \<longlongrightarrow> 0) F" using tendsto_uminus_ereal by fastforce
+  then show ?thesis unfolding v_def using Lim_transform_eventually[OF *] \<open> 1/l = 0 \<close> by auto
+qed
+
+lemma tendsto_divide_ereal [tendsto_intros]:
+  fixes f g::"_ \<Rightarrow> ereal"
+  assumes "(f \<longlongrightarrow> l) F" "(g \<longlongrightarrow> m) F" "m \<noteq> 0" "\<not>(abs(l) = \<infinity> \<and> abs(m) = \<infinity>)"
+  shows "((\<lambda>x. f x / g x) \<longlongrightarrow> l / m) F"
+proof -
+  define h where "h = (\<lambda>x. 1/ g x)"
+  have *: "(h \<longlongrightarrow> 1/m) F" unfolding h_def using assms(2) assms(3) tendsto_inverse_ereal by auto
+  have "((\<lambda>x. f x * h x) \<longlongrightarrow> l * (1/m)) F"
+    apply (rule tendsto_mult_ereal[OF assms(1) *]) using assms(3) assms(4) by (auto simp add: divide_ereal_def)
+  moreover have "f x * h x = f x / g x" for x unfolding h_def by (simp add: divide_ereal_def)
+  moreover have "l * (1/m) = l/m" by (simp add: divide_ereal_def)
+  ultimately show ?thesis unfolding h_def using Lim_transform_eventually by auto
+qed
+
+
+subsubsection \<open>Further limits\<close>
+
+lemma id_nat_ereal_tendsto_PInf [tendsto_intros]:
+  "(\<lambda> n::nat. real n) \<longlonglongrightarrow> \<infinity>"
+by (simp add: filterlim_real_sequentially tendsto_PInfty_eq_at_top)
+
+lemma tendsto_at_top_pseudo_inverse [tendsto_intros]:
+  fixes u::"nat \<Rightarrow> nat"
+  assumes "LIM n sequentially. u n :> at_top"
+  shows "LIM n sequentially. Inf {N. u N \<ge> n} :> at_top"
+proof -
+  {
+    fix C::nat
+    define M where "M = Max {u n| n. n \<le> C}+1"
+    {
+      fix n assume "n \<ge> M"
+      have "eventually (\<lambda>N. u N \<ge> n) sequentially" using assms
+        by (simp add: filterlim_at_top)
+      then have *: "{N. u N \<ge> n} \<noteq> {}" by force
+
+      have "N > C" if "u N \<ge> n" for N
+      proof (rule ccontr)
+        assume "\<not>(N > C)"
+        have "u N \<le> Max {u n| n. n \<le> C}"
+          apply (rule Max_ge) using \<open>\<not>(N > C)\<close> by auto
+        then show False using \<open>u N \<ge> n\<close> \<open>n \<ge> M\<close> unfolding M_def by auto
+      qed
+      then have **: "{N. u N \<ge> n} \<subseteq> {C..}" by fastforce
+      have "Inf {N. u N \<ge> n} \<ge> C"
+        by (metis "*" "**" Inf_nat_def1 atLeast_iff subset_eq)
+    }
+    then have "eventually (\<lambda>n. Inf {N. u N \<ge> n} \<ge> C) sequentially"
+      using eventually_sequentially by auto
+  }
+  then show ?thesis using filterlim_at_top by auto
+qed
+
+lemma pseudo_inverse_finite_set:
+  fixes u::"nat \<Rightarrow> nat"
+  assumes "LIM n sequentially. u n :> at_top"
+  shows "finite {N. u N \<le> n}"
+proof -
+  fix n
+  have "eventually (\<lambda>N. u N \<ge> n+1) sequentially" using assms
+    by (simp add: filterlim_at_top)
+  then obtain N1 where N1: "\<And>N. N \<ge> N1 \<Longrightarrow> u N \<ge> n + 1"
+    using eventually_sequentially by auto
+  have "{N. u N \<le> n} \<subseteq> {..<N1}"
+    apply auto using N1 by (metis Suc_eq_plus1 not_less not_less_eq_eq)
+  then show "finite {N. u N \<le> n}" by (simp add: finite_subset)
+qed
+
+lemma tendsto_at_top_pseudo_inverse2 [tendsto_intros]:
+  fixes u::"nat \<Rightarrow> nat"
+  assumes "LIM n sequentially. u n :> at_top"
+  shows "LIM n sequentially. Max {N. u N \<le> n} :> at_top"
+proof -
+  {
+    fix N0::nat
+    have "N0 \<le> Max {N. u N \<le> n}" if "n \<ge> u N0" for n
+      apply (rule Max.coboundedI) using pseudo_inverse_finite_set[OF assms] that by auto
+    then have "eventually (\<lambda>n. N0 \<le> Max {N. u N \<le> n}) sequentially"
+      using eventually_sequentially by blast
+  }
+  then show ?thesis using filterlim_at_top by auto
+qed
+
+lemma ereal_truncation_top [tendsto_intros]:
+  fixes x::ereal
+  shows "(\<lambda>n::nat. min x n) \<longlonglongrightarrow> x"
+proof (cases x)
+  case (real r)
+  then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
+  then have "min x n = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
+  then have "eventually (\<lambda>n. min x n = x) sequentially" using eventually_at_top_linorder by blast
+  then show ?thesis by (simp add: Lim_eventually)
+next
+  case (PInf)
+  then have "min x n = n" for n::nat by (auto simp add: min_def)
+  then show ?thesis using id_nat_ereal_tendsto_PInf PInf by auto
+next
+  case (MInf)
+  then have "min x n = x" for n::nat by (auto simp add: min_def)
+  then show ?thesis by auto
+qed
+
+lemma ereal_truncation_real_top [tendsto_intros]:
+  fixes x::ereal
+  assumes "x \<noteq> - \<infinity>"
+  shows "(\<lambda>n::nat. real_of_ereal(min x n)) \<longlonglongrightarrow> x"
+proof (cases x)
+  case (real r)
+  then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
+  then have "min x n = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
+  then have "real_of_ereal(min x n) = r" if "n \<ge> K" for n using real that by auto
+  then have "eventually (\<lambda>n. real_of_ereal(min x n) = r) sequentially" using eventually_at_top_linorder by blast
+  then have "(\<lambda>n. real_of_ereal(min x n)) \<longlonglongrightarrow> r" by (simp add: Lim_eventually)
+  then show ?thesis using real by auto
+next
+  case (PInf)
+  then have "real_of_ereal(min x n) = n" for n::nat by (auto simp add: min_def)
+  then show ?thesis using id_nat_ereal_tendsto_PInf PInf by auto
+qed (simp add: assms)
+
+lemma ereal_truncation_bottom [tendsto_intros]:
+  fixes x::ereal
+  shows "(\<lambda>n::nat. max x (- real n)) \<longlonglongrightarrow> x"
+proof (cases x)
+  case (real r)
+  then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
+  then have "max x (-real n) = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
+  then have "eventually (\<lambda>n. max x (-real n) = x) sequentially" using eventually_at_top_linorder by blast
+  then show ?thesis by (simp add: Lim_eventually)
+next
+  case (MInf)
+  then have "max x (-real n) = (-1)* ereal(real n)" for n::nat by (auto simp add: max_def)
+  moreover have "(\<lambda>n. (-1)* ereal(real n)) \<longlonglongrightarrow> -\<infinity>"
+    using tendsto_cmult_ereal[of "-1", OF _ id_nat_ereal_tendsto_PInf] by (simp add: one_ereal_def)
+  ultimately show ?thesis using MInf by auto
+next
+  case (PInf)
+  then have "max x (-real n) = x" for n::nat by (auto simp add: max_def)
+  then show ?thesis by auto
+qed
+
+lemma ereal_truncation_real_bottom [tendsto_intros]:
+  fixes x::ereal
+  assumes "x \<noteq> \<infinity>"
+  shows "(\<lambda>n::nat. real_of_ereal(max x (- real n))) \<longlonglongrightarrow> x"
+proof (cases x)
+  case (real r)
+  then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
+  then have "max x (-real n) = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
+  then have "real_of_ereal(max x (-real n)) = r" if "n \<ge> K" for n using real that by auto
+  then have "eventually (\<lambda>n. real_of_ereal(max x (-real n)) = r) sequentially" using eventually_at_top_linorder by blast
+  then have "(\<lambda>n. real_of_ereal(max x (-real n))) \<longlonglongrightarrow> r" by (simp add: Lim_eventually)
+  then show ?thesis using real by auto
+next
+  case (MInf)
+  then have "real_of_ereal(max x (-real n)) = (-1)* ereal(real n)" for n::nat by (auto simp add: max_def)
+  moreover have "(\<lambda>n. (-1)* ereal(real n)) \<longlonglongrightarrow> -\<infinity>"
+    using tendsto_cmult_ereal[of "-1", OF _ id_nat_ereal_tendsto_PInf] by (simp add: one_ereal_def)
+  ultimately show ?thesis using MInf by auto
+qed (simp add: assms)
+
+text \<open>the next one is copied from \verb+tendsto_sum+.\<close>
+lemma tendsto_sum_ereal [tendsto_intros]:
+  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> ereal"
+  assumes "\<And>i. i \<in> S \<Longrightarrow> (f i \<longlongrightarrow> a i) F"
+          "\<And>i. abs(a i) \<noteq> \<infinity>"
+  shows "((\<lambda>x. \<Sum>i\<in>S. f i x) \<longlongrightarrow> (\<Sum>i\<in>S. a i)) F"
+proof (cases "finite S")
+  assume "finite S" then show ?thesis using assms
+    by (induct, simp, simp add: tendsto_add_ereal_general2 assms)
+qed(simp)
+
+
 subsection \<open>monoset\<close>
 
 definition (in order) mono_set:
@@ -530,6 +1263,606 @@
     by auto
 qed
 
+lemma limsup_finite_then_bounded:
+  fixes u::"nat \<Rightarrow> real"
+  assumes "limsup u < \<infinity>"
+  shows "\<exists>C. \<forall>n. u n \<le> C"
+proof -
+  obtain C where C: "limsup u < C" "C < \<infinity>" using assms ereal_dense2 by blast
+  then have "C = ereal(real_of_ereal C)" using ereal_real by force
+  have "eventually (\<lambda>n. u n < C) sequentially" using C(1) unfolding Limsup_def
+    apply (auto simp add: INF_less_iff)
+    using SUP_lessD eventually_mono by fastforce
+  then obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> u n < C" using eventually_sequentially by auto
+  define D where "D = max (real_of_ereal C) (Max {u n |n. n \<le> N})"
+  have "\<And>n. u n \<le> D"
+  proof -
+    fix n show "u n \<le> D"
+    proof (cases)
+      assume *: "n \<le> N"
+      have "u n \<le> Max {u n |n. n \<le> N}" by (rule Max_ge, auto simp add: *)
+      then show "u n \<le> D" unfolding D_def by linarith
+    next
+      assume "\<not>(n \<le> N)"
+      then have "n \<ge> N" by simp
+      then have "u n < C" using N by auto
+      then have "u n < real_of_ereal C" using \<open>C = ereal(real_of_ereal C)\<close> less_ereal.simps(1) by fastforce
+      then show "u n \<le> D" unfolding D_def by linarith
+    qed
+  qed
+  then show ?thesis by blast
+qed
+
+lemma liminf_finite_then_bounded_below:
+  fixes u::"nat \<Rightarrow> real"
+  assumes "liminf u > -\<infinity>"
+  shows "\<exists>C. \<forall>n. u n \<ge> C"
+proof -
+  obtain C where C: "liminf u > C" "C > -\<infinity>" using assms using ereal_dense2 by blast
+  then have "C = ereal(real_of_ereal C)" using ereal_real by force
+  have "eventually (\<lambda>n. u n > C) sequentially" using C(1) unfolding Liminf_def
+    apply (auto simp add: less_SUP_iff)
+    using eventually_elim2 less_INF_D by fastforce
+  then obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> u n > C" using eventually_sequentially by auto
+  define D where "D = min (real_of_ereal C) (Min {u n |n. n \<le> N})"
+  have "\<And>n. u n \<ge> D"
+  proof -
+    fix n show "u n \<ge> D"
+    proof (cases)
+      assume *: "n \<le> N"
+      have "u n \<ge> Min {u n |n. n \<le> N}" by (rule Min_le, auto simp add: *)
+      then show "u n \<ge> D" unfolding D_def by linarith
+    next
+      assume "\<not>(n \<le> N)"
+      then have "n \<ge> N" by simp
+      then have "u n > C" using N by auto
+      then have "u n > real_of_ereal C" using \<open>C = ereal(real_of_ereal C)\<close> less_ereal.simps(1) by fastforce
+      then show "u n \<ge> D" unfolding D_def by linarith
+    qed
+  qed
+  then show ?thesis by blast
+qed
+
+lemma liminf_upper_bound:
+  fixes u:: "nat \<Rightarrow> ereal"
+  assumes "liminf u < l"
+  shows "\<exists>N>k. u N < l"
+by (metis assms gt_ex less_le_trans liminf_bounded_iff not_less)
+
+lemma limsup_shift:
+  "limsup (\<lambda>n. u (n+1)) = limsup u"
+proof -
+  have "(SUP m:{n+1..}. u m) = (SUP m:{n..}. u (m + 1))" for n
+    apply (rule SUP_eq) using Suc_le_D by auto
+  then have a: "(INF n. SUP m:{n..}. u (m + 1)) = (INF n. (SUP m:{n+1..}. u m))" by auto
+  have b: "(INF n. (SUP m:{n+1..}. u m)) = (INF n:{1..}. (SUP m:{n..}. u m))"
+    apply (rule INF_eq) using Suc_le_D by auto
+  have "(INF n:{1..}. v n) = (INF n. v n)" if "decseq v" for v::"nat \<Rightarrow> 'a"
+    apply (rule INF_eq) using \<open>decseq v\<close> decseq_Suc_iff by auto
+  moreover have "decseq (\<lambda>n. (SUP m:{n..}. u m))" by (simp add: SUP_subset_mono decseq_def)
+  ultimately have c: "(INF n:{1..}. (SUP m:{n..}. u m)) = (INF n. (SUP m:{n..}. u m))" by simp
+  have "(INF n. SUPREMUM {n..} u) = (INF n. SUP m:{n..}. u (m + 1))" using a b c by simp
+  then show ?thesis by (auto cong: limsup_INF_SUP)
+qed
+
+lemma limsup_shift_k:
+  "limsup (\<lambda>n. u (n+k)) = limsup u"
+proof (induction k)
+  case (Suc k)
+  have "limsup (\<lambda>n. u (n+k+1)) = limsup (\<lambda>n. u (n+k))" using limsup_shift[where ?u="\<lambda>n. u(n+k)"] by simp
+  then show ?case using Suc.IH by simp
+qed (auto)
+
+lemma liminf_shift:
+  "liminf (\<lambda>n. u (n+1)) = liminf u"
+proof -
+  have "(INF m:{n+1..}. u m) = (INF m:{n..}. u (m + 1))" for n
+    apply (rule INF_eq) using Suc_le_D by (auto)
+  then have a: "(SUP n. INF m:{n..}. u (m + 1)) = (SUP n. (INF m:{n+1..}. u m))" by auto
+  have b: "(SUP n. (INF m:{n+1..}. u m)) = (SUP n:{1..}. (INF m:{n..}. u m))"
+    apply (rule SUP_eq) using Suc_le_D by (auto)
+  have "(SUP n:{1..}. v n) = (SUP n. v n)" if "incseq v" for v::"nat \<Rightarrow> 'a"
+    apply (rule SUP_eq) using \<open>incseq v\<close> incseq_Suc_iff by auto
+  moreover have "incseq (\<lambda>n. (INF m:{n..}. u m))" by (simp add: INF_superset_mono mono_def)
+  ultimately have c: "(SUP n:{1..}. (INF m:{n..}. u m)) = (SUP n. (INF m:{n..}. u m))" by simp
+  have "(SUP n. INFIMUM {n..} u) = (SUP n. INF m:{n..}. u (m + 1))" using a b c by simp
+  then show ?thesis by (auto cong: liminf_SUP_INF)
+qed
+
+lemma liminf_shift_k:
+  "liminf (\<lambda>n. u (n+k)) = liminf u"
+proof (induction k)
+  case (Suc k)
+  have "liminf (\<lambda>n. u (n+k+1)) = liminf (\<lambda>n. u (n+k))" using liminf_shift[where ?u="\<lambda>n. u(n+k)"] by simp
+  then show ?case using Suc.IH by simp
+qed (auto)
+
+lemma Limsup_obtain:
+  fixes u::"_ \<Rightarrow> 'a :: complete_linorder"
+  assumes "Limsup F u > c"
+  shows "\<exists>i. u i > c"
+proof -
+  have "(INF P:{P. eventually P F}. SUP x:{x. P x}. u x) > c" using assms by (simp add: Limsup_def)
+  then show ?thesis by (metis eventually_True mem_Collect_eq less_INF_D less_SUP_iff)
+qed
+
+text \<open>The next lemma is extremely useful, as it often makes it possible to reduce statements
+about limsups to statements about limits.\<close>
+
+lemma limsup_subseq_lim:
+  fixes u::"nat \<Rightarrow> 'a :: {complete_linorder, linorder_topology}"
+  shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (u o r) \<longlonglongrightarrow> limsup u"
+proof (cases)
+  assume "\<forall>n. \<exists>p>n. \<forall>m\<ge>p. u m \<le> u p"
+  then have "\<exists>r. \<forall>n. (\<forall>m\<ge>r n. u m \<le> u (r n)) \<and> r n < r (Suc n)"
+    by (intro dependent_nat_choice) (auto simp: conj_commute)
+  then obtain r :: "nat \<Rightarrow> nat" where "strict_mono r" and mono: "\<And>n m. r n \<le> m \<Longrightarrow> u m \<le> u (r n)"
+    by (auto simp: strict_mono_Suc_iff)
+  define umax where "umax = (\<lambda>n. (SUP m:{n..}. u m))"
+  have "decseq umax" unfolding umax_def by (simp add: SUP_subset_mono antimono_def)
+  then have "umax \<longlonglongrightarrow> limsup u" unfolding umax_def by (metis LIMSEQ_INF limsup_INF_SUP)
+  then have *: "(umax o r) \<longlonglongrightarrow> limsup u" by (simp add: LIMSEQ_subseq_LIMSEQ \<open>strict_mono r\<close>)
+  have "\<And>n. umax(r n) = u(r n)" unfolding umax_def using mono
+    by (metis SUP_le_iff antisym atLeast_def mem_Collect_eq order_refl)
+  then have "umax o r = u o r" unfolding o_def by simp
+  then have "(u o r) \<longlonglongrightarrow> limsup u" using * by simp
+  then show ?thesis using \<open>strict_mono r\<close> by blast
+next
+  assume "\<not> (\<forall>n. \<exists>p>n. (\<forall>m\<ge>p. u m \<le> u p))"
+  then obtain N where N: "\<And>p. p > N \<Longrightarrow> \<exists>m>p. u p < u m" by (force simp: not_le le_less)
+  have "\<exists>r. \<forall>n. N < r n \<and> r n < r (Suc n) \<and> (\<forall>i\<in> {N<..r (Suc n)}. u i \<le> u (r (Suc n)))"
+  proof (rule dependent_nat_choice)
+    fix x assume "N < x"
+    then have a: "finite {N<..x}" "{N<..x} \<noteq> {}" by simp_all
+    have "Max {u i |i. i \<in> {N<..x}} \<in> {u i |i. i \<in> {N<..x}}" apply (rule Max_in) using a by (auto)
+    then obtain p where "p \<in> {N<..x}" and upmax: "u p = Max{u i |i. i \<in> {N<..x}}" by auto
+    define U where "U = {m. m > p \<and> u p < u m}"
+    have "U \<noteq> {}" unfolding U_def using N[of p] \<open>p \<in> {N<..x}\<close> by auto
+    define y where "y = Inf U"
+    then have "y \<in> U" using \<open>U \<noteq> {}\<close> by (simp add: Inf_nat_def1)
+    have a: "\<And>i. i \<in> {N<..x} \<Longrightarrow> u i \<le> u p"
+    proof -
+      fix i assume "i \<in> {N<..x}"
+      then have "u i \<in> {u i |i. i \<in> {N<..x}}" by blast
+      then show "u i \<le> u p" using upmax by simp
+    qed
+    moreover have "u p < u y" using \<open>y \<in> U\<close> U_def by auto
+    ultimately have "y \<notin> {N<..x}" using not_le by blast
+    moreover have "y > N" using \<open>y \<in> U\<close> U_def \<open>p \<in> {N<..x}\<close> by auto
+    ultimately have "y > x" by auto
+
+    have "\<And>i. i \<in> {N<..y} \<Longrightarrow> u i \<le> u y"
+    proof -
+      fix i assume "i \<in> {N<..y}" show "u i \<le> u y"
+      proof (cases)
+        assume "i = y"
+        then show ?thesis by simp
+      next
+        assume "\<not>(i=y)"
+        then have i:"i \<in> {N<..<y}" using \<open>i \<in> {N<..y}\<close> by simp
+        have "u i \<le> u p"
+        proof (cases)
+          assume "i \<le> x"
+          then have "i \<in> {N<..x}" using i by simp
+          then show ?thesis using a by simp
+        next
+          assume "\<not>(i \<le> x)"
+          then have "i > x" by simp
+          then have *: "i > p" using \<open>p \<in> {N<..x}\<close> by simp
+          have "i < Inf U" using i y_def by simp
+          then have "i \<notin> U" using Inf_nat_def not_less_Least by auto
+          then show ?thesis using U_def * by auto
+        qed
+        then show "u i \<le> u y" using \<open>u p < u y\<close> by auto
+      qed
+    qed
+    then have "N < y \<and> x < y \<and> (\<forall>i\<in>{N<..y}. u i \<le> u y)" using \<open>y > x\<close> \<open>y > N\<close> by auto
+    then show "\<exists>y>N. x < y \<and> (\<forall>i\<in>{N<..y}. u i \<le> u y)" by auto
+  qed (auto)
+  then obtain r where r: "\<forall>n. N < r n \<and> r n < r (Suc n) \<and> (\<forall>i\<in> {N<..r (Suc n)}. u i \<le> u (r (Suc n)))" by auto
+  have "strict_mono r" using r by (auto simp: strict_mono_Suc_iff)
+  have "incseq (u o r)" unfolding o_def using r by (simp add: incseq_SucI order.strict_implies_order)
+  then have "(u o r) \<longlonglongrightarrow> (SUP n. (u o r) n)" using LIMSEQ_SUP by blast
+  then have "limsup (u o r) = (SUP n. (u o r) n)" by (simp add: lim_imp_Limsup)
+  moreover have "limsup (u o r) \<le> limsup u" using \<open>strict_mono r\<close> by (simp add: limsup_subseq_mono)
+  ultimately have "(SUP n. (u o r) n) \<le> limsup u" by simp
+
+  {
+    fix i assume i: "i \<in> {N<..}"
+    obtain n where "i < r (Suc n)" using \<open>strict_mono r\<close> using Suc_le_eq seq_suble by blast
+    then have "i \<in> {N<..r(Suc n)}" using i by simp
+    then have "u i \<le> u (r(Suc n))" using r by simp
+    then have "u i \<le> (SUP n. (u o r) n)" unfolding o_def by (meson SUP_upper2 UNIV_I)
+  }
+  then have "(SUP i:{N<..}. u i) \<le> (SUP n. (u o r) n)" using SUP_least by blast
+  then have "limsup u \<le> (SUP n. (u o r) n)" unfolding Limsup_def
+    by (metis (mono_tags, lifting) INF_lower2 atLeast_Suc_greaterThan atLeast_def eventually_ge_at_top mem_Collect_eq)
+  then have "limsup u = (SUP n. (u o r) n)" using \<open>(SUP n. (u o r) n) \<le> limsup u\<close> by simp
+  then have "(u o r) \<longlonglongrightarrow> limsup u" using \<open>(u o r) \<longlonglongrightarrow> (SUP n. (u o r) n)\<close> by simp
+  then show ?thesis using \<open>strict_mono r\<close> by auto
+qed
+
+lemma liminf_subseq_lim:
+  fixes u::"nat \<Rightarrow> 'a :: {complete_linorder, linorder_topology}"
+  shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (u o r) \<longlonglongrightarrow> liminf u"
+proof (cases)
+  assume "\<forall>n. \<exists>p>n. \<forall>m\<ge>p. u m \<ge> u p"
+  then have "\<exists>r. \<forall>n. (\<forall>m\<ge>r n. u m \<ge> u (r n)) \<and> r n < r (Suc n)"
+    by (intro dependent_nat_choice) (auto simp: conj_commute)
+  then obtain r :: "nat \<Rightarrow> nat" where "strict_mono r" and mono: "\<And>n m. r n \<le> m \<Longrightarrow> u m \<ge> u (r n)"
+    by (auto simp: strict_mono_Suc_iff)
+  define umin where "umin = (\<lambda>n. (INF m:{n..}. u m))"
+  have "incseq umin" unfolding umin_def by (simp add: INF_superset_mono incseq_def)
+  then have "umin \<longlonglongrightarrow> liminf u" unfolding umin_def by (metis LIMSEQ_SUP liminf_SUP_INF)
+  then have *: "(umin o r) \<longlonglongrightarrow> liminf u" by (simp add: LIMSEQ_subseq_LIMSEQ \<open>strict_mono r\<close>)
+  have "\<And>n. umin(r n) = u(r n)" unfolding umin_def using mono
+    by (metis le_INF_iff antisym atLeast_def mem_Collect_eq order_refl)
+  then have "umin o r = u o r" unfolding o_def by simp
+  then have "(u o r) \<longlonglongrightarrow> liminf u" using * by simp
+  then show ?thesis using \<open>strict_mono r\<close> by blast
+next
+  assume "\<not> (\<forall>n. \<exists>p>n. (\<forall>m\<ge>p. u m \<ge> u p))"
+  then obtain N where N: "\<And>p. p > N \<Longrightarrow> \<exists>m>p. u p > u m" by (force simp: not_le le_less)
+  have "\<exists>r. \<forall>n. N < r n \<and> r n < r (Suc n) \<and> (\<forall>i\<in> {N<..r (Suc n)}. u i \<ge> u (r (Suc n)))"
+  proof (rule dependent_nat_choice)
+    fix x assume "N < x"
+    then have a: "finite {N<..x}" "{N<..x} \<noteq> {}" by simp_all
+    have "Min {u i |i. i \<in> {N<..x}} \<in> {u i |i. i \<in> {N<..x}}" apply (rule Min_in) using a by (auto)
+    then obtain p where "p \<in> {N<..x}" and upmin: "u p = Min{u i |i. i \<in> {N<..x}}" by auto
+    define U where "U = {m. m > p \<and> u p > u m}"
+    have "U \<noteq> {}" unfolding U_def using N[of p] \<open>p \<in> {N<..x}\<close> by auto
+    define y where "y = Inf U"
+    then have "y \<in> U" using \<open>U \<noteq> {}\<close> by (simp add: Inf_nat_def1)
+    have a: "\<And>i. i \<in> {N<..x} \<Longrightarrow> u i \<ge> u p"
+    proof -
+      fix i assume "i \<in> {N<..x}"
+      then have "u i \<in> {u i |i. i \<in> {N<..x}}" by blast
+      then show "u i \<ge> u p" using upmin by simp
+    qed
+    moreover have "u p > u y" using \<open>y \<in> U\<close> U_def by auto
+    ultimately have "y \<notin> {N<..x}" using not_le by blast
+    moreover have "y > N" using \<open>y \<in> U\<close> U_def \<open>p \<in> {N<..x}\<close> by auto
+    ultimately have "y > x" by auto
+
+    have "\<And>i. i \<in> {N<..y} \<Longrightarrow> u i \<ge> u y"
+    proof -
+      fix i assume "i \<in> {N<..y}" show "u i \<ge> u y"
+      proof (cases)
+        assume "i = y"
+        then show ?thesis by simp
+      next
+        assume "\<not>(i=y)"
+        then have i:"i \<in> {N<..<y}" using \<open>i \<in> {N<..y}\<close> by simp
+        have "u i \<ge> u p"
+        proof (cases)
+          assume "i \<le> x"
+          then have "i \<in> {N<..x}" using i by simp
+          then show ?thesis using a by simp
+        next
+          assume "\<not>(i \<le> x)"
+          then have "i > x" by simp
+          then have *: "i > p" using \<open>p \<in> {N<..x}\<close> by simp
+          have "i < Inf U" using i y_def by simp
+          then have "i \<notin> U" using Inf_nat_def not_less_Least by auto
+          then show ?thesis using U_def * by auto
+        qed
+        then show "u i \<ge> u y" using \<open>u p > u y\<close> by auto
+      qed
+    qed
+    then have "N < y \<and> x < y \<and> (\<forall>i\<in>{N<..y}. u i \<ge> u y)" using \<open>y > x\<close> \<open>y > N\<close> by auto
+    then show "\<exists>y>N. x < y \<and> (\<forall>i\<in>{N<..y}. u i \<ge> u y)" by auto
+  qed (auto)
+  then obtain r :: "nat \<Rightarrow> nat" 
+    where r: "\<forall>n. N < r n \<and> r n < r (Suc n) \<and> (\<forall>i\<in> {N<..r (Suc n)}. u i \<ge> u (r (Suc n)))" by auto
+  have "strict_mono r" using r by (auto simp: strict_mono_Suc_iff)
+  have "decseq (u o r)" unfolding o_def using r by (simp add: decseq_SucI order.strict_implies_order)
+  then have "(u o r) \<longlonglongrightarrow> (INF n. (u o r) n)" using LIMSEQ_INF by blast
+  then have "liminf (u o r) = (INF n. (u o r) n)" by (simp add: lim_imp_Liminf)
+  moreover have "liminf (u o r) \<ge> liminf u" using \<open>strict_mono r\<close> by (simp add: liminf_subseq_mono)
+  ultimately have "(INF n. (u o r) n) \<ge> liminf u" by simp
+
+  {
+    fix i assume i: "i \<in> {N<..}"
+    obtain n where "i < r (Suc n)" using \<open>strict_mono r\<close> using Suc_le_eq seq_suble by blast
+    then have "i \<in> {N<..r(Suc n)}" using i by simp
+    then have "u i \<ge> u (r(Suc n))" using r by simp
+    then have "u i \<ge> (INF n. (u o r) n)" unfolding o_def by (meson INF_lower2 UNIV_I)
+  }
+  then have "(INF i:{N<..}. u i) \<ge> (INF n. (u o r) n)" using INF_greatest by blast
+  then have "liminf u \<ge> (INF n. (u o r) n)" unfolding Liminf_def
+    by (metis (mono_tags, lifting) SUP_upper2 atLeast_Suc_greaterThan atLeast_def eventually_ge_at_top mem_Collect_eq)
+  then have "liminf u = (INF n. (u o r) n)" using \<open>(INF n. (u o r) n) \<ge> liminf u\<close> by simp
+  then have "(u o r) \<longlonglongrightarrow> liminf u" using \<open>(u o r) \<longlonglongrightarrow> (INF n. (u o r) n)\<close> by simp
+  then show ?thesis using \<open>strict_mono r\<close> by auto
+qed
+
+text \<open>The following statement about limsups is reduced to a statement about limits using
+subsequences thanks to \verb+limsup_subseq_lim+. The statement for limits follows for instance from
+\verb+tendsto_add_ereal_general+.\<close>
+
+lemma ereal_limsup_add_mono:
+  fixes u v::"nat \<Rightarrow> ereal"
+  shows "limsup (\<lambda>n. u n + v n) \<le> limsup u + limsup v"
+proof (cases)
+  assume "(limsup u = \<infinity>) \<or> (limsup v = \<infinity>)"
+  then have "limsup u + limsup v = \<infinity>" by simp
+  then show ?thesis by auto
+next
+  assume "\<not>((limsup u = \<infinity>) \<or> (limsup v = \<infinity>))"
+  then have "limsup u < \<infinity>" "limsup v < \<infinity>" by auto
+
+  define w where "w = (\<lambda>n. u n + v n)"
+  obtain r where r: "strict_mono r" "(w o r) \<longlonglongrightarrow> limsup w" using limsup_subseq_lim by auto
+  obtain s where s: "strict_mono s" "(u o r o s) \<longlonglongrightarrow> limsup (u o r)" using limsup_subseq_lim by auto
+  obtain t where t: "strict_mono t" "(v o r o s o t) \<longlonglongrightarrow> limsup (v o r o s)" using limsup_subseq_lim by auto
+
+  define a where "a = r o s o t"
+  have "strict_mono a" using r s t by (simp add: a_def strict_mono_o)
+  have l:"(w o a) \<longlonglongrightarrow> limsup w"
+         "(u o a) \<longlonglongrightarrow> limsup (u o r)"
+         "(v o a) \<longlonglongrightarrow> limsup (v o r o s)"
+  apply (metis (no_types, lifting) r(2) s(1) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
+  apply (metis (no_types, lifting) s(2) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
+  apply (metis (no_types, lifting) t(2) a_def comp_assoc)
+  done
+
+  have "limsup (u o r) \<le> limsup u" by (simp add: limsup_subseq_mono r(1))
+  then have a: "limsup (u o r) \<noteq> \<infinity>" using \<open>limsup u < \<infinity>\<close> by auto
+  have "limsup (v o r o s) \<le> limsup v" 
+    by (simp add: comp_assoc limsup_subseq_mono r(1) s(1) strict_mono_o)
+  then have b: "limsup (v o r o s) \<noteq> \<infinity>" using \<open>limsup v < \<infinity>\<close> by auto
+
+  have "(\<lambda>n. (u o a) n + (v o a) n) \<longlonglongrightarrow> limsup (u o r) + limsup (v o r o s)"
+    using l tendsto_add_ereal_general a b by fastforce
+  moreover have "(\<lambda>n. (u o a) n + (v o a) n) = (w o a)" unfolding w_def by auto
+  ultimately have "(w o a) \<longlonglongrightarrow> limsup (u o r) + limsup (v o r o s)" by simp
+  then have "limsup w = limsup (u o r) + limsup (v o r o s)" using l(1) LIMSEQ_unique by blast
+  then have "limsup w \<le> limsup u + limsup v"
+    using \<open>limsup (u o r) \<le> limsup u\<close> \<open>limsup (v o r o s) \<le> limsup v\<close> ereal_add_mono by simp
+  then show ?thesis unfolding w_def by simp
+qed
+
+text \<open>There is an asymmetry between liminfs and limsups in ereal, as $\infty + (-\infty) = \infty$.
+This explains why there are more assumptions in the next lemma dealing with liminfs that in the
+previous one about limsups.\<close>
+
+lemma ereal_liminf_add_mono:
+  fixes u v::"nat \<Rightarrow> ereal"
+  assumes "\<not>((liminf u = \<infinity> \<and> liminf v = -\<infinity>) \<or> (liminf u = -\<infinity> \<and> liminf v = \<infinity>))"
+  shows "liminf (\<lambda>n. u n + v n) \<ge> liminf u + liminf v"
+proof (cases)
+  assume "(liminf u = -\<infinity>) \<or> (liminf v = -\<infinity>)"
+  then have *: "liminf u + liminf v = -\<infinity>" using assms by auto
+  show ?thesis by (simp add: *)
+next
+  assume "\<not>((liminf u = -\<infinity>) \<or> (liminf v = -\<infinity>))"
+  then have "liminf u > -\<infinity>" "liminf v > -\<infinity>" by auto
+
+  define w where "w = (\<lambda>n. u n + v n)"
+  obtain r where r: "strict_mono r" "(w o r) \<longlonglongrightarrow> liminf w" using liminf_subseq_lim by auto
+  obtain s where s: "strict_mono s" "(u o r o s) \<longlonglongrightarrow> liminf (u o r)" using liminf_subseq_lim by auto
+  obtain t where t: "strict_mono t" "(v o r o s o t) \<longlonglongrightarrow> liminf (v o r o s)" using liminf_subseq_lim by auto
+
+  define a where "a = r o s o t"
+  have "strict_mono a" using r s t by (simp add: a_def strict_mono_o)
+  have l:"(w o a) \<longlonglongrightarrow> liminf w"
+         "(u o a) \<longlonglongrightarrow> liminf (u o r)"
+         "(v o a) \<longlonglongrightarrow> liminf (v o r o s)"
+  apply (metis (no_types, lifting) r(2) s(1) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
+  apply (metis (no_types, lifting) s(2) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
+  apply (metis (no_types, lifting) t(2) a_def comp_assoc)
+  done
+
+  have "liminf (u o r) \<ge> liminf u" by (simp add: liminf_subseq_mono r(1))
+  then have a: "liminf (u o r) \<noteq> -\<infinity>" using \<open>liminf u > -\<infinity>\<close> by auto
+  have "liminf (v o r o s) \<ge> liminf v" 
+    by (simp add: comp_assoc liminf_subseq_mono r(1) s(1) strict_mono_o)
+  then have b: "liminf (v o r o s) \<noteq> -\<infinity>" using \<open>liminf v > -\<infinity>\<close> by auto
+
+  have "(\<lambda>n. (u o a) n + (v o a) n) \<longlonglongrightarrow> liminf (u o r) + liminf (v o r o s)"
+    using l tendsto_add_ereal_general a b by fastforce
+  moreover have "(\<lambda>n. (u o a) n + (v o a) n) = (w o a)" unfolding w_def by auto
+  ultimately have "(w o a) \<longlonglongrightarrow> liminf (u o r) + liminf (v o r o s)" by simp
+  then have "liminf w = liminf (u o r) + liminf (v o r o s)" using l(1) LIMSEQ_unique by blast
+  then have "liminf w \<ge> liminf u + liminf v"
+    using \<open>liminf (u o r) \<ge> liminf u\<close> \<open>liminf (v o r o s) \<ge> liminf v\<close> ereal_add_mono by simp
+  then show ?thesis unfolding w_def by simp
+qed
+
+lemma ereal_limsup_lim_add:
+  fixes u v::"nat \<Rightarrow> ereal"
+  assumes "u \<longlonglongrightarrow> a" "abs(a) \<noteq> \<infinity>"
+  shows "limsup (\<lambda>n. u n + v n) = a + limsup v"
+proof -
+  have "limsup u = a" using assms(1) using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
+  have "(\<lambda>n. -u n) \<longlonglongrightarrow> -a" using assms(1) by auto
+  then have "limsup (\<lambda>n. -u n) = -a" using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
+
+  have "limsup (\<lambda>n. u n + v n) \<le> limsup u + limsup v"
+    by (rule ereal_limsup_add_mono)
+  then have up: "limsup (\<lambda>n. u n + v n) \<le> a + limsup v" using \<open>limsup u = a\<close> by simp
+
+  have a: "limsup (\<lambda>n. (u n + v n) + (-u n)) \<le> limsup (\<lambda>n. u n + v n) + limsup (\<lambda>n. -u n)"
+    by (rule ereal_limsup_add_mono)
+  have "eventually (\<lambda>n. u n = ereal(real_of_ereal(u n))) sequentially" using assms
+    real_lim_then_eventually_real by auto
+  moreover have "\<And>x. x = ereal(real_of_ereal(x)) \<Longrightarrow> x + (-x) = 0"
+    by (metis plus_ereal.simps(1) right_minus uminus_ereal.simps(1) zero_ereal_def)
+  ultimately have "eventually (\<lambda>n. u n + (-u n) = 0) sequentially"
+    by (metis (mono_tags, lifting) eventually_mono)
+  moreover have "\<And>n. u n + (-u n) = 0 \<Longrightarrow> u n + v n + (-u n) = v n"
+    by (metis add.commute add.left_commute add.left_neutral)
+  ultimately have "eventually (\<lambda>n. u n + v n + (-u n) = v n) sequentially"
+    using eventually_mono by force
+  then have "limsup v = limsup (\<lambda>n. u n + v n + (-u n))" using Limsup_eq by force
+  then have "limsup v \<le> limsup (\<lambda>n. u n + v n) -a" using a \<open>limsup (\<lambda>n. -u n) = -a\<close> by (simp add: minus_ereal_def)
+  then have "limsup (\<lambda>n. u n + v n) \<ge> a + limsup v" using assms(2) by (metis add.commute ereal_le_minus)
+  then show ?thesis using up by simp
+qed
+
+lemma ereal_limsup_lim_mult:
+  fixes u v::"nat \<Rightarrow> ereal"
+  assumes "u \<longlonglongrightarrow> a" "a>0" "a \<noteq> \<infinity>"
+  shows "limsup (\<lambda>n. u n * v n) = a * limsup v"
+proof -
+  define w where "w = (\<lambda>n. u n * v n)"
+  obtain r where r: "strict_mono r" "(v o r) \<longlonglongrightarrow> limsup v" using limsup_subseq_lim by auto
+  have "(u o r) \<longlonglongrightarrow> a" using assms(1) LIMSEQ_subseq_LIMSEQ r by auto
+  with tendsto_mult_ereal[OF this r(2)] have "(\<lambda>n. (u o r) n * (v o r) n) \<longlonglongrightarrow> a * limsup v" using assms(2) assms(3) by auto
+  moreover have "\<And>n. (w o r) n = (u o r) n * (v o r) n" unfolding w_def by auto
+  ultimately have "(w o r) \<longlonglongrightarrow> a * limsup v" unfolding w_def by presburger
+  then have "limsup (w o r) = a * limsup v" by (simp add: tendsto_iff_Liminf_eq_Limsup)
+  then have I: "limsup w \<ge> a * limsup v" by (metis limsup_subseq_mono r(1))
+
+  obtain s where s: "strict_mono s" "(w o s) \<longlonglongrightarrow> limsup w" using limsup_subseq_lim by auto
+  have *: "(u o s) \<longlonglongrightarrow> a" using assms(1) LIMSEQ_subseq_LIMSEQ s by auto
+  have "eventually (\<lambda>n. (u o s) n > 0) sequentially" using assms(2) * order_tendsto_iff by blast
+  moreover have "eventually (\<lambda>n. (u o s) n < \<infinity>) sequentially" using assms(3) * order_tendsto_iff by blast
+  moreover have "(w o s) n / (u o s) n = (v o s) n" if "(u o s) n > 0" "(u o s) n < \<infinity>" for n
+    unfolding w_def using that by (auto simp add: ereal_divide_eq)
+  ultimately have "eventually (\<lambda>n. (w o s) n / (u o s) n = (v o s) n) sequentially" using eventually_elim2 by force
+  moreover have "(\<lambda>n. (w o s) n / (u o s) n) \<longlonglongrightarrow> (limsup w) / a"
+    apply (rule tendsto_divide_ereal[OF s(2) *]) using assms(2) assms(3) by auto
+  ultimately have "(v o s) \<longlonglongrightarrow> (limsup w) / a" using Lim_transform_eventually by fastforce
+  then have "limsup (v o s) = (limsup w) / a" by (simp add: tendsto_iff_Liminf_eq_Limsup)
+  then have "limsup v \<ge> (limsup w) / a" by (metis limsup_subseq_mono s(1))
+  then have "a * limsup v \<ge> limsup w" using assms(2) assms(3) by (simp add: ereal_divide_le_pos)
+  then show ?thesis using I unfolding w_def by auto
+qed
+
+lemma ereal_liminf_lim_mult:
+  fixes u v::"nat \<Rightarrow> ereal"
+  assumes "u \<longlonglongrightarrow> a" "a>0" "a \<noteq> \<infinity>"
+  shows "liminf (\<lambda>n. u n * v n) = a * liminf v"
+proof -
+  define w where "w = (\<lambda>n. u n * v n)"
+  obtain r where r: "strict_mono r" "(v o r) \<longlonglongrightarrow> liminf v" using liminf_subseq_lim by auto
+  have "(u o r) \<longlonglongrightarrow> a" using assms(1) LIMSEQ_subseq_LIMSEQ r by auto
+  with tendsto_mult_ereal[OF this r(2)] have "(\<lambda>n. (u o r) n * (v o r) n) \<longlonglongrightarrow> a * liminf v" using assms(2) assms(3) by auto
+  moreover have "\<And>n. (w o r) n = (u o r) n * (v o r) n" unfolding w_def by auto
+  ultimately have "(w o r) \<longlonglongrightarrow> a * liminf v" unfolding w_def by presburger
+  then have "liminf (w o r) = a * liminf v" by (simp add: tendsto_iff_Liminf_eq_Limsup)
+  then have I: "liminf w \<le> a * liminf v" by (metis liminf_subseq_mono r(1))
+
+  obtain s where s: "strict_mono s" "(w o s) \<longlonglongrightarrow> liminf w" using liminf_subseq_lim by auto
+  have *: "(u o s) \<longlonglongrightarrow> a" using assms(1) LIMSEQ_subseq_LIMSEQ s by auto
+  have "eventually (\<lambda>n. (u o s) n > 0) sequentially" using assms(2) * order_tendsto_iff by blast
+  moreover have "eventually (\<lambda>n. (u o s) n < \<infinity>) sequentially" using assms(3) * order_tendsto_iff by blast
+  moreover have "(w o s) n / (u o s) n = (v o s) n" if "(u o s) n > 0" "(u o s) n < \<infinity>" for n
+    unfolding w_def using that by (auto simp add: ereal_divide_eq)
+  ultimately have "eventually (\<lambda>n. (w o s) n / (u o s) n = (v o s) n) sequentially" using eventually_elim2 by force
+  moreover have "(\<lambda>n. (w o s) n / (u o s) n) \<longlonglongrightarrow> (liminf w) / a"
+    apply (rule tendsto_divide_ereal[OF s(2) *]) using assms(2) assms(3) by auto
+  ultimately have "(v o s) \<longlonglongrightarrow> (liminf w) / a" using Lim_transform_eventually by fastforce
+  then have "liminf (v o s) = (liminf w) / a" by (simp add: tendsto_iff_Liminf_eq_Limsup)
+  then have "liminf v \<le> (liminf w) / a" by (metis liminf_subseq_mono s(1))
+  then have "a * liminf v \<le> liminf w" using assms(2) assms(3) by (simp add: ereal_le_divide_pos)
+  then show ?thesis using I unfolding w_def by auto
+qed
+
+lemma ereal_liminf_lim_add:
+  fixes u v::"nat \<Rightarrow> ereal"
+  assumes "u \<longlonglongrightarrow> a" "abs(a) \<noteq> \<infinity>"
+  shows "liminf (\<lambda>n. u n + v n) = a + liminf v"
+proof -
+  have "liminf u = a" using assms(1) tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
+  then have *: "abs(liminf u) \<noteq> \<infinity>" using assms(2) by auto
+  have "(\<lambda>n. -u n) \<longlonglongrightarrow> -a" using assms(1) by auto
+  then have "liminf (\<lambda>n. -u n) = -a" using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
+  then have **: "abs(liminf (\<lambda>n. -u n)) \<noteq> \<infinity>" using assms(2) by auto
+
+  have "liminf (\<lambda>n. u n + v n) \<ge> liminf u + liminf v"
+    apply (rule ereal_liminf_add_mono) using * by auto
+  then have up: "liminf (\<lambda>n. u n + v n) \<ge> a + liminf v" using \<open>liminf u = a\<close> by simp
+
+  have a: "liminf (\<lambda>n. (u n + v n) + (-u n)) \<ge> liminf (\<lambda>n. u n + v n) + liminf (\<lambda>n. -u n)"
+    apply (rule ereal_liminf_add_mono) using ** by auto
+  have "eventually (\<lambda>n. u n = ereal(real_of_ereal(u n))) sequentially" using assms
+    real_lim_then_eventually_real by auto
+  moreover have "\<And>x. x = ereal(real_of_ereal(x)) \<Longrightarrow> x + (-x) = 0"
+    by (metis plus_ereal.simps(1) right_minus uminus_ereal.simps(1) zero_ereal_def)
+  ultimately have "eventually (\<lambda>n. u n + (-u n) = 0) sequentially"
+    by (metis (mono_tags, lifting) eventually_mono)
+  moreover have "\<And>n. u n + (-u n) = 0 \<Longrightarrow> u n + v n + (-u n) = v n"
+    by (metis add.commute add.left_commute add.left_neutral)
+  ultimately have "eventually (\<lambda>n. u n + v n + (-u n) = v n) sequentially"
+    using eventually_mono by force
+  then have "liminf v = liminf (\<lambda>n. u n + v n + (-u n))" using Liminf_eq by force
+  then have "liminf v \<ge> liminf (\<lambda>n. u n + v n) -a" using a \<open>liminf (\<lambda>n. -u n) = -a\<close> by (simp add: minus_ereal_def)
+  then have "liminf (\<lambda>n. u n + v n) \<le> a + liminf v" using assms(2) by (metis add.commute ereal_minus_le)
+  then show ?thesis using up by simp
+qed
+
+lemma ereal_liminf_limsup_add:
+  fixes u v::"nat \<Rightarrow> ereal"
+  shows "liminf (\<lambda>n. u n + v n) \<le> liminf u + limsup v"
+proof (cases)
+  assume "limsup v = \<infinity> \<or> liminf u = \<infinity>"
+  then show ?thesis by auto
+next
+  assume "\<not>(limsup v = \<infinity> \<or> liminf u = \<infinity>)"
+  then have "limsup v < \<infinity>" "liminf u < \<infinity>" by auto
+
+  define w where "w = (\<lambda>n. u n + v n)"
+  obtain r where r: "strict_mono r" "(u o r) \<longlonglongrightarrow> liminf u" using liminf_subseq_lim by auto
+  obtain s where s: "strict_mono s" "(w o r o s) \<longlonglongrightarrow> liminf (w o r)" using liminf_subseq_lim by auto
+  obtain t where t: "strict_mono t" "(v o r o s o t) \<longlonglongrightarrow> limsup (v o r o s)" using limsup_subseq_lim by auto
+
+  define a where "a = r o s o t"
+  have "strict_mono a" using r s t by (simp add: a_def strict_mono_o)
+  have l:"(u o a) \<longlonglongrightarrow> liminf u"
+         "(w o a) \<longlonglongrightarrow> liminf (w o r)"
+         "(v o a) \<longlonglongrightarrow> limsup (v o r o s)"
+  apply (metis (no_types, lifting) r(2) s(1) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
+  apply (metis (no_types, lifting) s(2) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
+  apply (metis (no_types, lifting) t(2) a_def comp_assoc)
+  done
+
+  have "liminf (w o r) \<ge> liminf w" by (simp add: liminf_subseq_mono r(1))
+  have "limsup (v o r o s) \<le> limsup v" 
+    by (simp add: comp_assoc limsup_subseq_mono r(1) s(1) strict_mono_o)
+  then have b: "limsup (v o r o s) < \<infinity>" using \<open>limsup v < \<infinity>\<close> by auto
+
+  have "(\<lambda>n. (u o a) n + (v o a) n) \<longlonglongrightarrow> liminf u + limsup (v o r o s)"
+    apply (rule tendsto_add_ereal_general) using b \<open>liminf u < \<infinity>\<close> l(1) l(3) by force+
+  moreover have "(\<lambda>n. (u o a) n + (v o a) n) = (w o a)" unfolding w_def by auto
+  ultimately have "(w o a) \<longlonglongrightarrow> liminf u + limsup (v o r o s)" by simp
+  then have "liminf (w o r) = liminf u + limsup (v o r o s)" using l(2) using LIMSEQ_unique by blast
+  then have "liminf w \<le> liminf u + limsup v"
+    using \<open>liminf (w o r) \<ge> liminf w\<close> \<open>limsup (v o r o s) \<le> limsup v\<close>
+    by (metis add_mono_thms_linordered_semiring(2) le_less_trans not_less)
+  then show ?thesis unfolding w_def by simp
+qed
+
+lemma ereal_liminf_limsup_minus:
+  fixes u v::"nat \<Rightarrow> ereal"
+  shows "liminf (\<lambda>n. u n - v n) \<le> limsup u - limsup v"
+  unfolding minus_ereal_def
+  apply (subst add.commute)
+  apply (rule order_trans[OF ereal_liminf_limsup_add])
+  using ereal_Limsup_uminus[of sequentially "\<lambda>n. - v n"]
+  apply (simp add: add.commute)
+  done
+
+
+lemma liminf_minus_ennreal:
+  fixes u v::"nat \<Rightarrow> ennreal"
+  shows "(\<And>n. v n \<le> u n) \<Longrightarrow> liminf (\<lambda>n. u n - v n) \<le> limsup u - limsup v"
+  unfolding liminf_SUP_INF limsup_INF_SUP
+  including ennreal.lifting
+proof (transfer, clarsimp)
+  fix v u :: "nat \<Rightarrow> ereal" assume *: "\<forall>x. 0 \<le> v x" "\<forall>x. 0 \<le> u x" "\<And>n. v n \<le> u n"
+  moreover have "0 \<le> limsup u - limsup v"
+    using * by (intro ereal_diff_positive Limsup_mono always_eventually) simp
+  moreover have "0 \<le> (SUPREMUM {x..} v)" for x
+    using * by (intro SUP_upper2[of x]) auto
+  moreover have "0 \<le> (SUPREMUM {x..} u)" for x
+    using * by (intro SUP_upper2[of x]) auto
+  ultimately show "(SUP n. INF n:{n..}. max 0 (u n - v n))
+            \<le> max 0 ((INF x. max 0 (SUPREMUM {x..} u)) - (INF x. max 0 (SUPREMUM {x..} v)))"
+    by (auto simp: * ereal_diff_positive max.absorb2 liminf_SUP_INF[symmetric] limsup_INF_SUP[symmetric] ereal_liminf_limsup_minus)
+qed
+
 subsection "Relate extended reals and the indicator function"
 
 lemma ereal_indicator_le_0: "(indicator S x::ereal) \<le> 0 \<longleftrightarrow> x \<notin> S"
--- a/src/HOL/Analysis/Path_Connected.thy	Fri Aug 18 22:55:54 2017 +0200
+++ b/src/HOL/Analysis/Path_Connected.thy	Sun Aug 20 03:35:20 2017 +0200
@@ -6747,6 +6747,16 @@
     by (metis countable_subset)
 qed
 
+lemma ball_minus_countable_nonempty:
+  assumes "countable (A :: 'a :: euclidean_space set)" "r > 0"
+  shows   "ball z r - A \<noteq> {}"
+proof
+  assume *: "ball z r - A = {}"
+  have "uncountable (ball z r - A)"
+    by (intro uncountable_minus_countable assms uncountable_ball)
+  thus False by (subst (asm) *) auto
+qed
+
 lemma uncountable_cball:
   fixes a :: "'a::euclidean_space"
   assumes "r > 0"
--- a/src/HOL/Analysis/Set_Integral.thy	Fri Aug 18 22:55:54 2017 +0200
+++ b/src/HOL/Analysis/Set_Integral.thy	Sun Aug 20 03:35:20 2017 +0200
@@ -11,1343 +11,6 @@
   imports Radon_Nikodym
 begin
 
-lemma bij_inv_eq_iff: "bij p \<Longrightarrow> x = inv p y \<longleftrightarrow> p x = y" (* COPIED FROM Permutations *)
-  using surj_f_inv_f[of p] by (auto simp add: bij_def)
-
-subsection \<open>Fun.thy\<close>
-
-lemma inj_fn:
-  fixes f::"'a \<Rightarrow> 'a"
-  assumes "inj f"
-  shows "inj (f^^n)"
-proof (induction n)
-  case (Suc n)
-  have "inj (f o (f^^n))"
-    using inj_comp[OF assms Suc.IH] by simp
-  then show "inj (f^^(Suc n))"
-    by auto
-qed (auto)
-
-lemma surj_fn:
-  fixes f::"'a \<Rightarrow> 'a"
-  assumes "surj f"
-  shows "surj (f^^n)"
-proof (induction n)
-  case (Suc n)
-  have "surj (f o (f^^n))"
-    using assms Suc.IH by (simp add: comp_surj)
-  then show "surj (f^^(Suc n))"
-    by auto
-qed (auto)
-
-lemma bij_fn:
-  fixes f::"'a \<Rightarrow> 'a"
-  assumes "bij f"
-  shows "bij (f^^n)"
-by (rule bijI[OF inj_fn[OF bij_is_inj[OF assms]] surj_fn[OF bij_is_surj[OF assms]]])
-
-lemma inv_fn_o_fn_is_id:
-  fixes f::"'a \<Rightarrow> 'a"
-  assumes "bij f"
-  shows "((inv f)^^n) o (f^^n) = (\<lambda>x. x)"
-proof -
-  have "((inv f)^^n)((f^^n) x) = x" for x n
-  proof (induction n)
-    case (Suc n)
-    have *: "(inv f) (f y) = y" for y
-      by (simp add: assms bij_is_inj)
-    have "(inv f ^^ Suc n) ((f ^^ Suc n) x) = (inv f^^n) (inv f (f ((f^^n) x)))"
-      by (simp add: funpow_swap1)
-    also have "... = (inv f^^n) ((f^^n) x)"
-      using * by auto
-    also have "... = x" using Suc.IH by auto
-    finally show ?case by simp
-  qed (auto)
-  then show ?thesis unfolding o_def by blast
-qed
-
-lemma fn_o_inv_fn_is_id:
-  fixes f::"'a \<Rightarrow> 'a"
-  assumes "bij f"
-  shows "(f^^n) o ((inv f)^^n) = (\<lambda>x. x)"
-proof -
-  have "(f^^n) (((inv f)^^n) x) = x" for x n
-  proof (induction n)
-    case (Suc n)
-    have *: "f(inv f y) = y" for y
-      using assms by (meson bij_inv_eq_iff)
-    have "(f ^^ Suc n) ((inv f ^^ Suc n) x) = (f^^n) (f (inv f ((inv f^^n) x)))"
-      by (simp add: funpow_swap1)
-    also have "... = (f^^n) ((inv f^^n) x)"
-      using * by auto
-    also have "... = x" using Suc.IH by auto
-    finally show ?case by simp
-  qed (auto)
-  then show ?thesis unfolding o_def by blast
-qed
-
-lemma inv_fn:
-  fixes f::"'a \<Rightarrow> 'a"
-  assumes "bij f"
-  shows "inv (f^^n) = ((inv f)^^n)"
-proof -
-  have "inv (f^^n) x = ((inv f)^^n) x" for x
-  apply (rule inv_into_f_eq, auto simp add: inj_fn[OF bij_is_inj[OF assms]])
-  using fn_o_inv_fn_is_id[OF assms, of n] by (metis comp_apply)
-  then show ?thesis by auto
-qed
-
-
-lemma mono_inv:
-  fixes f::"'a::linorder \<Rightarrow> 'b::linorder"
-  assumes "mono f" "bij f"
-  shows "mono (inv f)"
-proof
-  fix x y::'b assume "x \<le> y"
-  then show "inv f x \<le> inv f y"
-    by (metis (no_types, lifting) assms bij_is_surj eq_iff le_cases mono_def surj_f_inv_f)
-qed
-
-lemma mono_bij_Inf:
-  fixes f :: "'a::complete_linorder \<Rightarrow> 'b::complete_linorder"
-  assumes "mono f" "bij f"
-  shows "f (Inf A) = Inf (f`A)"
-proof -
-  have "(inv f) (Inf (f`A)) \<le> Inf ((inv f)`(f`A))"
-    using mono_Inf[OF mono_inv[OF assms], of "f`A"] by simp
-  then have "Inf (f`A) \<le> f (Inf ((inv f)`(f`A)))"
-    by (metis (no_types, lifting) assms mono_def bij_inv_eq_iff)
-  also have "... = f(Inf A)"
-    using assms by (simp add: bij_is_inj)
-  finally show ?thesis using mono_Inf[OF assms(1), of A] by auto
-qed
-
-
-lemma Inf_nat_def1:
-  fixes K::"nat set"
-  assumes "K \<noteq> {}"
-  shows "Inf K \<in> K"
-by (auto simp add: Min_def Inf_nat_def) (meson LeastI assms bot.extremum_unique subsetI)
-
-subsection \<open>Liminf-Limsup.thy\<close>
-
-lemma limsup_shift:
-  "limsup (\<lambda>n. u (n+1)) = limsup u"
-proof -
-  have "(SUP m:{n+1..}. u m) = (SUP m:{n..}. u (m + 1))" for n
-    apply (rule SUP_eq) using Suc_le_D by auto
-  then have a: "(INF n. SUP m:{n..}. u (m + 1)) = (INF n. (SUP m:{n+1..}. u m))" by auto
-  have b: "(INF n. (SUP m:{n+1..}. u m)) = (INF n:{1..}. (SUP m:{n..}. u m))"
-    apply (rule INF_eq) using Suc_le_D by auto
-  have "(INF n:{1..}. v n) = (INF n. v n)" if "decseq v" for v::"nat \<Rightarrow> 'a"
-    apply (rule INF_eq) using \<open>decseq v\<close> decseq_Suc_iff by auto
-  moreover have "decseq (\<lambda>n. (SUP m:{n..}. u m))" by (simp add: SUP_subset_mono decseq_def)
-  ultimately have c: "(INF n:{1..}. (SUP m:{n..}. u m)) = (INF n. (SUP m:{n..}. u m))" by simp
-  have "(INF n. SUPREMUM {n..} u) = (INF n. SUP m:{n..}. u (m + 1))" using a b c by simp
-  then show ?thesis by (auto cong: limsup_INF_SUP)
-qed
-
-lemma limsup_shift_k:
-  "limsup (\<lambda>n. u (n+k)) = limsup u"
-proof (induction k)
-  case (Suc k)
-  have "limsup (\<lambda>n. u (n+k+1)) = limsup (\<lambda>n. u (n+k))" using limsup_shift[where ?u="\<lambda>n. u(n+k)"] by simp
-  then show ?case using Suc.IH by simp
-qed (auto)
-
-lemma liminf_shift:
-  "liminf (\<lambda>n. u (n+1)) = liminf u"
-proof -
-  have "(INF m:{n+1..}. u m) = (INF m:{n..}. u (m + 1))" for n
-    apply (rule INF_eq) using Suc_le_D by (auto)
-  then have a: "(SUP n. INF m:{n..}. u (m + 1)) = (SUP n. (INF m:{n+1..}. u m))" by auto
-  have b: "(SUP n. (INF m:{n+1..}. u m)) = (SUP n:{1..}. (INF m:{n..}. u m))"
-    apply (rule SUP_eq) using Suc_le_D by (auto)
-  have "(SUP n:{1..}. v n) = (SUP n. v n)" if "incseq v" for v::"nat \<Rightarrow> 'a"
-    apply (rule SUP_eq) using \<open>incseq v\<close> incseq_Suc_iff by auto
-  moreover have "incseq (\<lambda>n. (INF m:{n..}. u m))" by (simp add: INF_superset_mono mono_def)
-  ultimately have c: "(SUP n:{1..}. (INF m:{n..}. u m)) = (SUP n. (INF m:{n..}. u m))" by simp
-  have "(SUP n. INFIMUM {n..} u) = (SUP n. INF m:{n..}. u (m + 1))" using a b c by simp
-  then show ?thesis by (auto cong: liminf_SUP_INF)
-qed
-
-lemma liminf_shift_k:
-  "liminf (\<lambda>n. u (n+k)) = liminf u"
-proof (induction k)
-  case (Suc k)
-  have "liminf (\<lambda>n. u (n+k+1)) = liminf (\<lambda>n. u (n+k))" using liminf_shift[where ?u="\<lambda>n. u(n+k)"] by simp
-  then show ?case using Suc.IH by simp
-qed (auto)
-
-lemma Limsup_obtain:
-  fixes u::"_ \<Rightarrow> 'a :: complete_linorder"
-  assumes "Limsup F u > c"
-  shows "\<exists>i. u i > c"
-proof -
-  have "(INF P:{P. eventually P F}. SUP x:{x. P x}. u x) > c" using assms by (simp add: Limsup_def)
-  then show ?thesis by (metis eventually_True mem_Collect_eq less_INF_D less_SUP_iff)
-qed
-
-text \<open>The next lemma is extremely useful, as it often makes it possible to reduce statements
-about limsups to statements about limits.\<close>
-
-lemma limsup_subseq_lim:
-  fixes u::"nat \<Rightarrow> 'a :: {complete_linorder, linorder_topology}"
-  shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (u o r) \<longlonglongrightarrow> limsup u"
-proof (cases)
-  assume "\<forall>n. \<exists>p>n. \<forall>m\<ge>p. u m \<le> u p"
-  then have "\<exists>r. \<forall>n. (\<forall>m\<ge>r n. u m \<le> u (r n)) \<and> r n < r (Suc n)"
-    by (intro dependent_nat_choice) (auto simp: conj_commute)
-  then obtain r :: "nat \<Rightarrow> nat" where "strict_mono r" and mono: "\<And>n m. r n \<le> m \<Longrightarrow> u m \<le> u (r n)"
-    by (auto simp: strict_mono_Suc_iff)
-  define umax where "umax = (\<lambda>n. (SUP m:{n..}. u m))"
-  have "decseq umax" unfolding umax_def by (simp add: SUP_subset_mono antimono_def)
-  then have "umax \<longlonglongrightarrow> limsup u" unfolding umax_def by (metis LIMSEQ_INF limsup_INF_SUP)
-  then have *: "(umax o r) \<longlonglongrightarrow> limsup u" by (simp add: LIMSEQ_subseq_LIMSEQ \<open>strict_mono r\<close>)
-  have "\<And>n. umax(r n) = u(r n)" unfolding umax_def using mono
-    by (metis SUP_le_iff antisym atLeast_def mem_Collect_eq order_refl)
-  then have "umax o r = u o r" unfolding o_def by simp
-  then have "(u o r) \<longlonglongrightarrow> limsup u" using * by simp
-  then show ?thesis using \<open>strict_mono r\<close> by blast
-next
-  assume "\<not> (\<forall>n. \<exists>p>n. (\<forall>m\<ge>p. u m \<le> u p))"
-  then obtain N where N: "\<And>p. p > N \<Longrightarrow> \<exists>m>p. u p < u m" by (force simp: not_le le_less)
-  have "\<exists>r. \<forall>n. N < r n \<and> r n < r (Suc n) \<and> (\<forall>i\<in> {N<..r (Suc n)}. u i \<le> u (r (Suc n)))"
-  proof (rule dependent_nat_choice)
-    fix x assume "N < x"
-    then have a: "finite {N<..x}" "{N<..x} \<noteq> {}" by simp_all
-    have "Max {u i |i. i \<in> {N<..x}} \<in> {u i |i. i \<in> {N<..x}}" apply (rule Max_in) using a by (auto)
-    then obtain p where "p \<in> {N<..x}" and upmax: "u p = Max{u i |i. i \<in> {N<..x}}" by auto
-    define U where "U = {m. m > p \<and> u p < u m}"
-    have "U \<noteq> {}" unfolding U_def using N[of p] \<open>p \<in> {N<..x}\<close> by auto
-    define y where "y = Inf U"
-    then have "y \<in> U" using \<open>U \<noteq> {}\<close> by (simp add: Inf_nat_def1)
-    have a: "\<And>i. i \<in> {N<..x} \<Longrightarrow> u i \<le> u p"
-    proof -
-      fix i assume "i \<in> {N<..x}"
-      then have "u i \<in> {u i |i. i \<in> {N<..x}}" by blast
-      then show "u i \<le> u p" using upmax by simp
-    qed
-    moreover have "u p < u y" using \<open>y \<in> U\<close> U_def by auto
-    ultimately have "y \<notin> {N<..x}" using not_le by blast
-    moreover have "y > N" using \<open>y \<in> U\<close> U_def \<open>p \<in> {N<..x}\<close> by auto
-    ultimately have "y > x" by auto
-
-    have "\<And>i. i \<in> {N<..y} \<Longrightarrow> u i \<le> u y"
-    proof -
-      fix i assume "i \<in> {N<..y}" show "u i \<le> u y"
-      proof (cases)
-        assume "i = y"
-        then show ?thesis by simp
-      next
-        assume "\<not>(i=y)"
-        then have i:"i \<in> {N<..<y}" using \<open>i \<in> {N<..y}\<close> by simp
-        have "u i \<le> u p"
-        proof (cases)
-          assume "i \<le> x"
-          then have "i \<in> {N<..x}" using i by simp
-          then show ?thesis using a by simp
-        next
-          assume "\<not>(i \<le> x)"
-          then have "i > x" by simp
-          then have *: "i > p" using \<open>p \<in> {N<..x}\<close> by simp
-          have "i < Inf U" using i y_def by simp
-          then have "i \<notin> U" using Inf_nat_def not_less_Least by auto
-          then show ?thesis using U_def * by auto
-        qed
-        then show "u i \<le> u y" using \<open>u p < u y\<close> by auto
-      qed
-    qed
-    then have "N < y \<and> x < y \<and> (\<forall>i\<in>{N<..y}. u i \<le> u y)" using \<open>y > x\<close> \<open>y > N\<close> by auto
-    then show "\<exists>y>N. x < y \<and> (\<forall>i\<in>{N<..y}. u i \<le> u y)" by auto
-  qed (auto)
-  then obtain r where r: "\<forall>n. N < r n \<and> r n < r (Suc n) \<and> (\<forall>i\<in> {N<..r (Suc n)}. u i \<le> u (r (Suc n)))" by auto
-  have "strict_mono r" using r by (auto simp: strict_mono_Suc_iff)
-  have "incseq (u o r)" unfolding o_def using r by (simp add: incseq_SucI order.strict_implies_order)
-  then have "(u o r) \<longlonglongrightarrow> (SUP n. (u o r) n)" using LIMSEQ_SUP by blast
-  then have "limsup (u o r) = (SUP n. (u o r) n)" by (simp add: lim_imp_Limsup)
-  moreover have "limsup (u o r) \<le> limsup u" using \<open>strict_mono r\<close> by (simp add: limsup_subseq_mono)
-  ultimately have "(SUP n. (u o r) n) \<le> limsup u" by simp
-
-  {
-    fix i assume i: "i \<in> {N<..}"
-    obtain n where "i < r (Suc n)" using \<open>strict_mono r\<close> using Suc_le_eq seq_suble by blast
-    then have "i \<in> {N<..r(Suc n)}" using i by simp
-    then have "u i \<le> u (r(Suc n))" using r by simp
-    then have "u i \<le> (SUP n. (u o r) n)" unfolding o_def by (meson SUP_upper2 UNIV_I)
-  }
-  then have "(SUP i:{N<..}. u i) \<le> (SUP n. (u o r) n)" using SUP_least by blast
-  then have "limsup u \<le> (SUP n. (u o r) n)" unfolding Limsup_def
-    by (metis (mono_tags, lifting) INF_lower2 atLeast_Suc_greaterThan atLeast_def eventually_ge_at_top mem_Collect_eq)
-  then have "limsup u = (SUP n. (u o r) n)" using \<open>(SUP n. (u o r) n) \<le> limsup u\<close> by simp
-  then have "(u o r) \<longlonglongrightarrow> limsup u" using \<open>(u o r) \<longlonglongrightarrow> (SUP n. (u o r) n)\<close> by simp
-  then show ?thesis using \<open>strict_mono r\<close> by auto
-qed
-
-lemma liminf_subseq_lim:
-  fixes u::"nat \<Rightarrow> 'a :: {complete_linorder, linorder_topology}"
-  shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (u o r) \<longlonglongrightarrow> liminf u"
-proof (cases)
-  assume "\<forall>n. \<exists>p>n. \<forall>m\<ge>p. u m \<ge> u p"
-  then have "\<exists>r. \<forall>n. (\<forall>m\<ge>r n. u m \<ge> u (r n)) \<and> r n < r (Suc n)"
-    by (intro dependent_nat_choice) (auto simp: conj_commute)
-  then obtain r :: "nat \<Rightarrow> nat" where "strict_mono r" and mono: "\<And>n m. r n \<le> m \<Longrightarrow> u m \<ge> u (r n)"
-    by (auto simp: strict_mono_Suc_iff)
-  define umin where "umin = (\<lambda>n. (INF m:{n..}. u m))"
-  have "incseq umin" unfolding umin_def by (simp add: INF_superset_mono incseq_def)
-  then have "umin \<longlonglongrightarrow> liminf u" unfolding umin_def by (metis LIMSEQ_SUP liminf_SUP_INF)
-  then have *: "(umin o r) \<longlonglongrightarrow> liminf u" by (simp add: LIMSEQ_subseq_LIMSEQ \<open>strict_mono r\<close>)
-  have "\<And>n. umin(r n) = u(r n)" unfolding umin_def using mono
-    by (metis le_INF_iff antisym atLeast_def mem_Collect_eq order_refl)
-  then have "umin o r = u o r" unfolding o_def by simp
-  then have "(u o r) \<longlonglongrightarrow> liminf u" using * by simp
-  then show ?thesis using \<open>strict_mono r\<close> by blast
-next
-  assume "\<not> (\<forall>n. \<exists>p>n. (\<forall>m\<ge>p. u m \<ge> u p))"
-  then obtain N where N: "\<And>p. p > N \<Longrightarrow> \<exists>m>p. u p > u m" by (force simp: not_le le_less)
-  have "\<exists>r. \<forall>n. N < r n \<and> r n < r (Suc n) \<and> (\<forall>i\<in> {N<..r (Suc n)}. u i \<ge> u (r (Suc n)))"
-  proof (rule dependent_nat_choice)
-    fix x assume "N < x"
-    then have a: "finite {N<..x}" "{N<..x} \<noteq> {}" by simp_all
-    have "Min {u i |i. i \<in> {N<..x}} \<in> {u i |i. i \<in> {N<..x}}" apply (rule Min_in) using a by (auto)
-    then obtain p where "p \<in> {N<..x}" and upmin: "u p = Min{u i |i. i \<in> {N<..x}}" by auto
-    define U where "U = {m. m > p \<and> u p > u m}"
-    have "U \<noteq> {}" unfolding U_def using N[of p] \<open>p \<in> {N<..x}\<close> by auto
-    define y where "y = Inf U"
-    then have "y \<in> U" using \<open>U \<noteq> {}\<close> by (simp add: Inf_nat_def1)
-    have a: "\<And>i. i \<in> {N<..x} \<Longrightarrow> u i \<ge> u p"
-    proof -
-      fix i assume "i \<in> {N<..x}"
-      then have "u i \<in> {u i |i. i \<in> {N<..x}}" by blast
-      then show "u i \<ge> u p" using upmin by simp
-    qed
-    moreover have "u p > u y" using \<open>y \<in> U\<close> U_def by auto
-    ultimately have "y \<notin> {N<..x}" using not_le by blast
-    moreover have "y > N" using \<open>y \<in> U\<close> U_def \<open>p \<in> {N<..x}\<close> by auto
-    ultimately have "y > x" by auto
-
-    have "\<And>i. i \<in> {N<..y} \<Longrightarrow> u i \<ge> u y"
-    proof -
-      fix i assume "i \<in> {N<..y}" show "u i \<ge> u y"
-      proof (cases)
-        assume "i = y"
-        then show ?thesis by simp
-      next
-        assume "\<not>(i=y)"
-        then have i:"i \<in> {N<..<y}" using \<open>i \<in> {N<..y}\<close> by simp
-        have "u i \<ge> u p"
-        proof (cases)
-          assume "i \<le> x"
-          then have "i \<in> {N<..x}" using i by simp
-          then show ?thesis using a by simp
-        next
-          assume "\<not>(i \<le> x)"
-          then have "i > x" by simp
-          then have *: "i > p" using \<open>p \<in> {N<..x}\<close> by simp
-          have "i < Inf U" using i y_def by simp
-          then have "i \<notin> U" using Inf_nat_def not_less_Least by auto
-          then show ?thesis using U_def * by auto
-        qed
-        then show "u i \<ge> u y" using \<open>u p > u y\<close> by auto
-      qed
-    qed
-    then have "N < y \<and> x < y \<and> (\<forall>i\<in>{N<..y}. u i \<ge> u y)" using \<open>y > x\<close> \<open>y > N\<close> by auto
-    then show "\<exists>y>N. x < y \<and> (\<forall>i\<in>{N<..y}. u i \<ge> u y)" by auto
-  qed (auto)
-  then obtain r :: "nat \<Rightarrow> nat" 
-    where r: "\<forall>n. N < r n \<and> r n < r (Suc n) \<and> (\<forall>i\<in> {N<..r (Suc n)}. u i \<ge> u (r (Suc n)))" by auto
-  have "strict_mono r" using r by (auto simp: strict_mono_Suc_iff)
-  have "decseq (u o r)" unfolding o_def using r by (simp add: decseq_SucI order.strict_implies_order)
-  then have "(u o r) \<longlonglongrightarrow> (INF n. (u o r) n)" using LIMSEQ_INF by blast
-  then have "liminf (u o r) = (INF n. (u o r) n)" by (simp add: lim_imp_Liminf)
-  moreover have "liminf (u o r) \<ge> liminf u" using \<open>strict_mono r\<close> by (simp add: liminf_subseq_mono)
-  ultimately have "(INF n. (u o r) n) \<ge> liminf u" by simp
-
-  {
-    fix i assume i: "i \<in> {N<..}"
-    obtain n where "i < r (Suc n)" using \<open>strict_mono r\<close> using Suc_le_eq seq_suble by blast
-    then have "i \<in> {N<..r(Suc n)}" using i by simp
-    then have "u i \<ge> u (r(Suc n))" using r by simp
-    then have "u i \<ge> (INF n. (u o r) n)" unfolding o_def by (meson INF_lower2 UNIV_I)
-  }
-  then have "(INF i:{N<..}. u i) \<ge> (INF n. (u o r) n)" using INF_greatest by blast
-  then have "liminf u \<ge> (INF n. (u o r) n)" unfolding Liminf_def
-    by (metis (mono_tags, lifting) SUP_upper2 atLeast_Suc_greaterThan atLeast_def eventually_ge_at_top mem_Collect_eq)
-  then have "liminf u = (INF n. (u o r) n)" using \<open>(INF n. (u o r) n) \<ge> liminf u\<close> by simp
-  then have "(u o r) \<longlonglongrightarrow> liminf u" using \<open>(u o r) \<longlonglongrightarrow> (INF n. (u o r) n)\<close> by simp
-  then show ?thesis using \<open>strict_mono r\<close> by auto
-qed
-
-
-subsection \<open>Extended-Real.thy\<close>
-
-text\<open>The proof of this one is copied from \verb+ereal_add_mono+.\<close>
-lemma ereal_add_strict_mono2:
-  fixes a b c d :: ereal
-  assumes "a < b"
-    and "c < d"
-  shows "a + c < b + d"
-using assms
-apply (cases a)
-apply (cases rule: ereal3_cases[of b c d], auto)
-apply (cases rule: ereal3_cases[of b c d], auto)
-done
-
-text \<open>The next ones are analogues of \verb+mult_mono+ and \verb+mult_mono'+ in ereal.\<close>
-
-lemma ereal_mult_mono:
-  fixes a b c d::ereal
-  assumes "b \<ge> 0" "c \<ge> 0" "a \<le> b" "c \<le> d"
-  shows "a * c \<le> b * d"
-by (metis ereal_mult_right_mono mult.commute order_trans assms)
-
-lemma ereal_mult_mono':
-  fixes a b c d::ereal
-  assumes "a \<ge> 0" "c \<ge> 0" "a \<le> b" "c \<le> d"
-  shows "a * c \<le> b * d"
-by (metis ereal_mult_right_mono mult.commute order_trans assms)
-
-lemma ereal_mult_mono_strict:
-  fixes a b c d::ereal
-  assumes "b > 0" "c > 0" "a < b" "c < d"
-  shows "a * c < b * d"
-proof -
-  have "c < \<infinity>" using \<open>c < d\<close> by auto
-  then have "a * c < b * c" by (metis ereal_mult_strict_left_mono[OF assms(3) assms(2)] mult.commute)
-  moreover have "b * c \<le> b * d" using assms(2) assms(4) by (simp add: assms(1) ereal_mult_left_mono less_imp_le)
-  ultimately show ?thesis by simp
-qed
-
-lemma ereal_mult_mono_strict':
-  fixes a b c d::ereal
-  assumes "a > 0" "c > 0" "a < b" "c < d"
-  shows "a * c < b * d"
-apply (rule ereal_mult_mono_strict, auto simp add: assms) using assms by auto
-
-lemma ereal_abs_add:
-  fixes a b::ereal
-  shows "abs(a+b) \<le> abs a + abs b"
-by (cases rule: ereal2_cases[of a b]) (auto)
-
-lemma ereal_abs_diff:
-  fixes a b::ereal
-  shows "abs(a-b) \<le> abs a + abs b"
-by (cases rule: ereal2_cases[of a b]) (auto)
-
-lemma sum_constant_ereal:
-  fixes a::ereal
-  shows "(\<Sum>i\<in>I. a) = a * card I"
-apply (cases "finite I", induct set: finite, simp_all)
-apply (cases a, auto, metis (no_types, hide_lams) add.commute mult.commute semiring_normalization_rules(3))
-done
-
-lemma real_lim_then_eventually_real:
-  assumes "(u \<longlongrightarrow> ereal l) F"
-  shows "eventually (\<lambda>n. u n = ereal(real_of_ereal(u n))) F"
-proof -
-  have "ereal l \<in> {-\<infinity><..<(\<infinity>::ereal)}" by simp
-  moreover have "open {-\<infinity><..<(\<infinity>::ereal)}" by simp
-  ultimately have "eventually (\<lambda>n. u n \<in> {-\<infinity><..<(\<infinity>::ereal)}) F" using assms tendsto_def by blast
-  moreover have "\<And>x. x \<in> {-\<infinity><..<(\<infinity>::ereal)} \<Longrightarrow> x = ereal(real_of_ereal x)" using ereal_real by auto
-  ultimately show ?thesis by (metis (mono_tags, lifting) eventually_mono)
-qed
-
-lemma ereal_Inf_cmult:
-  assumes "c>(0::real)"
-  shows "Inf {ereal c * x |x. P x} = ereal c * Inf {x. P x}"
-proof -
-  have "(\<lambda>x::ereal. c * x) (Inf {x::ereal. P x}) = Inf ((\<lambda>x::ereal. c * x)`{x::ereal. P x})"
-    apply (rule mono_bij_Inf)
-    apply (simp add: assms ereal_mult_left_mono less_imp_le mono_def)
-    apply (rule bij_betw_byWitness[of _ "\<lambda>x. (x::ereal) / c"], auto simp add: assms ereal_mult_divide)
-    using assms ereal_divide_eq apply auto
-    done
-  then show ?thesis by (simp only: setcompr_eq_image[symmetric])
-qed
-
-
-subsubsection \<open>Continuity of addition\<close>
-
-text \<open>The next few lemmas remove an unnecessary assumption in \verb+tendsto_add_ereal+, culminating
-in \verb+tendsto_add_ereal_general+ which essentially says that the addition
-is continuous on ereal times ereal, except at $(-\infty, \infty)$ and $(\infty, -\infty)$.
-It is much more convenient in many situations, see for instance the proof of
-\verb+tendsto_sum_ereal+ below.\<close>
-
-lemma tendsto_add_ereal_PInf:
-  fixes y :: ereal
-  assumes y: "y \<noteq> -\<infinity>"
-  assumes f: "(f \<longlongrightarrow> \<infinity>) F" and g: "(g \<longlongrightarrow> y) F"
-  shows "((\<lambda>x. f x + g x) \<longlongrightarrow> \<infinity>) F"
-proof -
-  have "\<exists>C. eventually (\<lambda>x. g x > ereal C) F"
-  proof (cases y)
-    case (real r)
-    have "y > y-1" using y real by (simp add: ereal_between(1))
-    then have "eventually (\<lambda>x. g x > y - 1) F" using g y order_tendsto_iff by auto
-    moreover have "y-1 = ereal(real_of_ereal(y-1))"
-      by (metis real ereal_eq_1(1) ereal_minus(1) real_of_ereal.simps(1))
-    ultimately have "eventually (\<lambda>x. g x > ereal(real_of_ereal(y - 1))) F" by simp
-    then show ?thesis by auto
-  next
-    case (PInf)
-    have "eventually (\<lambda>x. g x > ereal 0) F" using g PInf by (simp add: tendsto_PInfty)
-    then show ?thesis by auto
-  qed (simp add: y)
-  then obtain C::real where ge: "eventually (\<lambda>x. g x > ereal C) F" by auto
-
-  {
-    fix M::real
-    have "eventually (\<lambda>x. f x > ereal(M - C)) F" using f by (simp add: tendsto_PInfty)
-    then have "eventually (\<lambda>x. (f x > ereal (M-C)) \<and> (g x > ereal C)) F"
-      by (auto simp add: ge eventually_conj_iff)
-    moreover have "\<And>x. ((f x > ereal (M-C)) \<and> (g x > ereal C)) \<Longrightarrow> (f x + g x > ereal M)"
-      using ereal_add_strict_mono2 by fastforce
-    ultimately have "eventually (\<lambda>x. f x + g x > ereal M) F" using eventually_mono by force
-  }
-  then show ?thesis by (simp add: tendsto_PInfty)
-qed
-
-text\<open>One would like to deduce the next lemma from the previous one, but the fact
-that $-(x+y)$ is in general different from $(-x) + (-y)$ in ereal creates difficulties,
-so it is more efficient to copy the previous proof.\<close>
-
-lemma tendsto_add_ereal_MInf:
-  fixes y :: ereal
-  assumes y: "y \<noteq> \<infinity>"
-  assumes f: "(f \<longlongrightarrow> -\<infinity>) F" and g: "(g \<longlongrightarrow> y) F"
-  shows "((\<lambda>x. f x + g x) \<longlongrightarrow> -\<infinity>) F"
-proof -
-  have "\<exists>C. eventually (\<lambda>x. g x < ereal C) F"
-  proof (cases y)
-    case (real r)
-    have "y < y+1" using y real by (simp add: ereal_between(1))
-    then have "eventually (\<lambda>x. g x < y + 1) F" using g y order_tendsto_iff by force
-    moreover have "y+1 = ereal(real_of_ereal (y+1))" by (simp add: real)
-    ultimately have "eventually (\<lambda>x. g x < ereal(real_of_ereal(y + 1))) F" by simp
-    then show ?thesis by auto
-  next
-    case (MInf)
-    have "eventually (\<lambda>x. g x < ereal 0) F" using g MInf by (simp add: tendsto_MInfty)
-    then show ?thesis by auto
-  qed (simp add: y)
-  then obtain C::real where ge: "eventually (\<lambda>x. g x < ereal C) F" by auto
-
-  {
-    fix M::real
-    have "eventually (\<lambda>x. f x < ereal(M - C)) F" using f by (simp add: tendsto_MInfty)
-    then have "eventually (\<lambda>x. (f x < ereal (M- C)) \<and> (g x < ereal C)) F"
-      by (auto simp add: ge eventually_conj_iff)
-    moreover have "\<And>x. ((f x < ereal (M-C)) \<and> (g x < ereal C)) \<Longrightarrow> (f x + g x < ereal M)"
-      using ereal_add_strict_mono2 by fastforce
-    ultimately have "eventually (\<lambda>x. f x + g x < ereal M) F" using eventually_mono by force
-  }
-  then show ?thesis by (simp add: tendsto_MInfty)
-qed
-
-lemma tendsto_add_ereal_general1:
-  fixes x y :: ereal
-  assumes y: "\<bar>y\<bar> \<noteq> \<infinity>"
-  assumes f: "(f \<longlongrightarrow> x) F" and g: "(g \<longlongrightarrow> y) F"
-  shows "((\<lambda>x. f x + g x) \<longlongrightarrow> x + y) F"
-proof (cases x)
-  case (real r)
-  have a: "\<bar>x\<bar> \<noteq> \<infinity>" by (simp add: real)
-  show ?thesis by (rule tendsto_add_ereal[OF a, OF y, OF f, OF g])
-next
-  case PInf
-  then show ?thesis using tendsto_add_ereal_PInf assms by force
-next
-  case MInf
-  then show ?thesis using tendsto_add_ereal_MInf assms
-    by (metis abs_ereal.simps(3) ereal_MInfty_eq_plus)
-qed
-
-lemma tendsto_add_ereal_general2:
-  fixes x y :: ereal
-  assumes x: "\<bar>x\<bar> \<noteq> \<infinity>"
-      and f: "(f \<longlongrightarrow> x) F" and g: "(g \<longlongrightarrow> y) F"
-  shows "((\<lambda>x. f x + g x) \<longlongrightarrow> x + y) F"
-proof -
-  have "((\<lambda>x. g x + f x) \<longlongrightarrow> x + y) F"
-    using tendsto_add_ereal_general1[OF x, OF g, OF f] add.commute[of "y", of "x"] by simp
-  moreover have "\<And>x. g x + f x = f x + g x" using add.commute by auto
-  ultimately show ?thesis by simp
-qed
-
-text \<open>The next lemma says that the addition is continuous on ereal, except at
-the pairs $(-\infty, \infty)$ and $(\infty, -\infty)$.\<close>
-
-lemma tendsto_add_ereal_general [tendsto_intros]:
-  fixes x y :: ereal
-  assumes "\<not>((x=\<infinity> \<and> y=-\<infinity>) \<or> (x=-\<infinity> \<and> y=\<infinity>))"
-      and f: "(f \<longlongrightarrow> x) F" and g: "(g \<longlongrightarrow> y) F"
-  shows "((\<lambda>x. f x + g x) \<longlongrightarrow> x + y) F"
-proof (cases x)
-  case (real r)
-  show ?thesis
-    apply (rule tendsto_add_ereal_general2) using real assms by auto
-next
-  case (PInf)
-  then have "y \<noteq> -\<infinity>" using assms by simp
-  then show ?thesis using tendsto_add_ereal_PInf PInf assms by auto
-next
-  case (MInf)
-  then have "y \<noteq> \<infinity>" using assms by simp
-  then show ?thesis using tendsto_add_ereal_MInf MInf f g by (metis ereal_MInfty_eq_plus)
-qed
-
-subsubsection \<open>Continuity of multiplication\<close>
-
-text \<open>In the same way as for addition, we prove that the multiplication is continuous on
-ereal times ereal, except at $(\infty, 0)$ and $(-\infty, 0)$ and $(0, \infty)$ and $(0, -\infty)$,
-starting with specific situations.\<close>
-
-lemma tendsto_mult_real_ereal:
-  assumes "(u \<longlongrightarrow> ereal l) F" "(v \<longlongrightarrow> ereal m) F"
-  shows "((\<lambda>n. u n * v n) \<longlongrightarrow> ereal l * ereal m) F"
-proof -
-  have ureal: "eventually (\<lambda>n. u n = ereal(real_of_ereal(u n))) F" by (rule real_lim_then_eventually_real[OF assms(1)])
-  then have "((\<lambda>n. ereal(real_of_ereal(u n))) \<longlongrightarrow> ereal l) F" using assms by auto
-  then have limu: "((\<lambda>n. real_of_ereal(u n)) \<longlongrightarrow> l) F" by auto
-  have vreal: "eventually (\<lambda>n. v n = ereal(real_of_ereal(v n))) F" by (rule real_lim_then_eventually_real[OF assms(2)])
-  then have "((\<lambda>n. ereal(real_of_ereal(v n))) \<longlongrightarrow> ereal m) F" using assms by auto
-  then have limv: "((\<lambda>n. real_of_ereal(v n)) \<longlongrightarrow> m) F" by auto
-
-  {
-    fix n assume "u n = ereal(real_of_ereal(u n))" "v n = ereal(real_of_ereal(v n))"
-    then have "ereal(real_of_ereal(u n) * real_of_ereal(v n)) = u n * v n" by (metis times_ereal.simps(1))
-  }
-  then have *: "eventually (\<lambda>n. ereal(real_of_ereal(u n) * real_of_ereal(v n)) = u n * v n) F"
-    using eventually_elim2[OF ureal vreal] by auto
-
-  have "((\<lambda>n. real_of_ereal(u n) * real_of_ereal(v n)) \<longlongrightarrow> l * m) F" using tendsto_mult[OF limu limv] by auto
-  then have "((\<lambda>n. ereal(real_of_ereal(u n)) * real_of_ereal(v n)) \<longlongrightarrow> ereal(l * m)) F" by auto
-  then show ?thesis using * filterlim_cong by fastforce
-qed
-
-lemma tendsto_mult_ereal_PInf:
-  fixes f g::"_ \<Rightarrow> ereal"
-  assumes "(f \<longlongrightarrow> l) F" "l>0" "(g \<longlongrightarrow> \<infinity>) F"
-  shows "((\<lambda>x. f x * g x) \<longlongrightarrow> \<infinity>) F"
-proof -
-  obtain a::real where "0 < ereal a" "a < l" using assms(2) using ereal_dense2 by blast
-  have *: "eventually (\<lambda>x. f x > a) F" using \<open>a < l\<close> assms(1) by (simp add: order_tendsto_iff)
-  {
-    fix K::real
-    define M where "M = max K 1"
-    then have "M > 0" by simp
-    then have "ereal(M/a) > 0" using \<open>ereal a > 0\<close> by simp
-    then have "\<And>x. ((f x > a) \<and> (g x > M/a)) \<Longrightarrow> (f x * g x > ereal a * ereal(M/a))"
-      using ereal_mult_mono_strict'[where ?c = "M/a", OF \<open>0 < ereal a\<close>] by auto
-    moreover have "ereal a * ereal(M/a) = M" using \<open>ereal a > 0\<close> by simp
-    ultimately have "\<And>x. ((f x > a) \<and> (g x > M/a)) \<Longrightarrow> (f x * g x > M)" by simp
-    moreover have "M \<ge> K" unfolding M_def by simp
-    ultimately have Imp: "\<And>x. ((f x > a) \<and> (g x > M/a)) \<Longrightarrow> (f x * g x > K)"
-      using ereal_less_eq(3) le_less_trans by blast
-
-    have "eventually (\<lambda>x. g x > M/a) F" using assms(3) by (simp add: tendsto_PInfty)
-    then have "eventually (\<lambda>x. (f x > a) \<and> (g x > M/a)) F"
-      using * by (auto simp add: eventually_conj_iff)
-    then have "eventually (\<lambda>x. f x * g x > K) F" using eventually_mono Imp by force
-  }
-  then show ?thesis by (auto simp add: tendsto_PInfty)
-qed
-
-lemma tendsto_mult_ereal_pos:
-  fixes f g::"_ \<Rightarrow> ereal"
-  assumes "(f \<longlongrightarrow> l) F" "(g \<longlongrightarrow> m) F" "l>0" "m>0"
-  shows "((\<lambda>x. f x * g x) \<longlongrightarrow> l * m) F"
-proof (cases)
-  assume *: "l = \<infinity> \<or> m = \<infinity>"
-  then show ?thesis
-  proof (cases)
-    assume "m = \<infinity>"
-    then show ?thesis using tendsto_mult_ereal_PInf assms by auto
-  next
-    assume "\<not>(m = \<infinity>)"
-    then have "l = \<infinity>" using * by simp
-    then have "((\<lambda>x. g x * f x) \<longlongrightarrow> l * m) F" using tendsto_mult_ereal_PInf assms by auto
-    moreover have "\<And>x. g x * f x = f x * g x" using mult.commute by auto
-    ultimately show ?thesis by simp
-  qed
-next
-  assume "\<not>(l = \<infinity> \<or> m = \<infinity>)"
-  then have "l < \<infinity>" "m < \<infinity>" by auto
-  then obtain lr mr where "l = ereal lr" "m = ereal mr"
-    using \<open>l>0\<close> \<open>m>0\<close> by (metis ereal_cases ereal_less(6) not_less_iff_gr_or_eq)
-  then show ?thesis using tendsto_mult_real_ereal assms by auto
-qed
-
-text \<open>We reduce the general situation to the positive case by multiplying by suitable signs.
-Unfortunately, as ereal is not a ring, all the neat sign lemmas are not available there. We
-give the bare minimum we need.\<close>
-
-lemma ereal_sgn_abs:
-  fixes l::ereal
-  shows "sgn(l) * l = abs(l)"
-apply (cases l) by (auto simp add: sgn_if ereal_less_uminus_reorder)
-
-lemma sgn_squared_ereal:
-  assumes "l \<noteq> (0::ereal)"
-  shows "sgn(l) * sgn(l) = 1"
-apply (cases l) using assms by (auto simp add: one_ereal_def sgn_if)
-
-lemma tendsto_mult_ereal [tendsto_intros]:
-  fixes f g::"_ \<Rightarrow> ereal"
-  assumes "(f \<longlongrightarrow> l) F" "(g \<longlongrightarrow> m) F" "\<not>((l=0 \<and> abs(m) = \<infinity>) \<or> (m=0 \<and> abs(l) = \<infinity>))"
-  shows "((\<lambda>x. f x * g x) \<longlongrightarrow> l * m) F"
-proof (cases)
-  assume "l=0 \<or> m=0"
-  then have "abs(l) \<noteq> \<infinity>" "abs(m) \<noteq> \<infinity>" using assms(3) by auto
-  then obtain lr mr where "l = ereal lr" "m = ereal mr" by auto
-  then show ?thesis using tendsto_mult_real_ereal assms by auto
-next
-  have sgn_finite: "\<And>a::ereal. abs(sgn a) \<noteq> \<infinity>"
-    by (metis MInfty_neq_ereal(2) PInfty_neq_ereal(2) abs_eq_infinity_cases ereal_times(1) ereal_times(3) ereal_uminus_eq_reorder sgn_ereal.elims)
-  then have sgn_finite2: "\<And>a b::ereal. abs(sgn a * sgn b) \<noteq> \<infinity>"
-    by (metis abs_eq_infinity_cases abs_ereal.simps(2) abs_ereal.simps(3) ereal_mult_eq_MInfty ereal_mult_eq_PInfty)
-  assume "\<not>(l=0 \<or> m=0)"
-  then have "l \<noteq> 0" "m \<noteq> 0" by auto
-  then have "abs(l) > 0" "abs(m) > 0"
-    by (metis abs_ereal_ge0 abs_ereal_less0 abs_ereal_pos ereal_uminus_uminus ereal_uminus_zero less_le not_less)+
-  then have "sgn(l) * l > 0" "sgn(m) * m > 0" using ereal_sgn_abs by auto
-  moreover have "((\<lambda>x. sgn(l) * f x) \<longlongrightarrow> (sgn(l) * l)) F"
-    by (rule tendsto_cmult_ereal, auto simp add: sgn_finite assms(1))
-  moreover have "((\<lambda>x. sgn(m) * g x) \<longlongrightarrow> (sgn(m) * m)) F"
-    by (rule tendsto_cmult_ereal, auto simp add: sgn_finite assms(2))
-  ultimately have *: "((\<lambda>x. (sgn(l) * f x) * (sgn(m) * g x)) \<longlongrightarrow> (sgn(l) * l) * (sgn(m) * m)) F"
-    using tendsto_mult_ereal_pos by force
-  have "((\<lambda>x. (sgn(l) * sgn(m)) * ((sgn(l) * f x) * (sgn(m) * g x))) \<longlongrightarrow> (sgn(l) * sgn(m)) * ((sgn(l) * l) * (sgn(m) * m))) F"
-    by (rule tendsto_cmult_ereal, auto simp add: sgn_finite2 *)
-  moreover have "\<And>x. (sgn(l) * sgn(m)) * ((sgn(l) * f x) * (sgn(m) * g x)) = f x * g x"
-    by (metis mult.left_neutral sgn_squared_ereal[OF \<open>l \<noteq> 0\<close>] sgn_squared_ereal[OF \<open>m \<noteq> 0\<close>] mult.assoc mult.commute)
-  moreover have "(sgn(l) * sgn(m)) * ((sgn(l) * l) * (sgn(m) * m)) = l * m"
-    by (metis mult.left_neutral sgn_squared_ereal[OF \<open>l \<noteq> 0\<close>] sgn_squared_ereal[OF \<open>m \<noteq> 0\<close>] mult.assoc mult.commute)
-  ultimately show ?thesis by auto
-qed
-
-lemma tendsto_cmult_ereal_general [tendsto_intros]:
-  fixes f::"_ \<Rightarrow> ereal" and c::ereal
-  assumes "(f \<longlongrightarrow> l) F" "\<not> (l=0 \<and> abs(c) = \<infinity>)"
-  shows "((\<lambda>x. c * f x) \<longlongrightarrow> c * l) F"
-by (cases "c = 0", auto simp add: assms tendsto_mult_ereal)
-
-
-subsubsection \<open>Continuity of division\<close>
-
-lemma tendsto_inverse_ereal_PInf:
-  fixes u::"_ \<Rightarrow> ereal"
-  assumes "(u \<longlongrightarrow> \<infinity>) F"
-  shows "((\<lambda>x. 1/ u x) \<longlongrightarrow> 0) F"
-proof -
-  {
-    fix e::real assume "e>0"
-    have "1/e < \<infinity>" by auto
-    then have "eventually (\<lambda>n. u n > 1/e) F" using assms(1) by (simp add: tendsto_PInfty)
-    moreover
-    {
-      fix z::ereal assume "z>1/e"
-      then have "z>0" using \<open>e>0\<close> using less_le_trans not_le by fastforce
-      then have "1/z \<ge> 0" by auto
-      moreover have "1/z < e" using \<open>e>0\<close> \<open>z>1/e\<close>
-        apply (cases z) apply auto
-        by (metis (mono_tags, hide_lams) less_ereal.simps(2) less_ereal.simps(4) divide_less_eq ereal_divide_less_pos ereal_less(4)
-            ereal_less_eq(4) less_le_trans mult_eq_0_iff not_le not_one_less_zero times_ereal.simps(1))
-      ultimately have "1/z \<ge> 0" "1/z < e" by auto
-    }
-    ultimately have "eventually (\<lambda>n. 1/u n<e) F" "eventually (\<lambda>n. 1/u n\<ge>0) F" by (auto simp add: eventually_mono)
-  } note * = this
-  show ?thesis
-  proof (subst order_tendsto_iff, auto)
-    fix a::ereal assume "a<0"
-    then show "eventually (\<lambda>n. 1/u n > a) F" using *(2) eventually_mono less_le_trans linordered_field_no_ub by fastforce
-  next
-    fix a::ereal assume "a>0"
-    then obtain e::real where "e>0" "a>e" using ereal_dense2 ereal_less(2) by blast
-    then have "eventually (\<lambda>n. 1/u n < e) F" using *(1) by auto
-    then show "eventually (\<lambda>n. 1/u n < a) F" using \<open>a>e\<close> by (metis (mono_tags, lifting) eventually_mono less_trans)
-  qed
-qed
-
-text \<open>The next lemma deserves to exist by itself, as it is so common and useful.\<close>
-
-lemma tendsto_inverse_real [tendsto_intros]:
-  fixes u::"_ \<Rightarrow> real"
-  shows "(u \<longlongrightarrow> l) F \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> ((\<lambda>x. 1/ u x) \<longlongrightarrow> 1/l) F"
-  using tendsto_inverse unfolding inverse_eq_divide .
-
-lemma tendsto_inverse_ereal [tendsto_intros]:
-  fixes u::"_ \<Rightarrow> ereal"
-  assumes "(u \<longlongrightarrow> l) F" "l \<noteq> 0"
-  shows "((\<lambda>x. 1/ u x) \<longlongrightarrow> 1/l) F"
-proof (cases l)
-  case (real r)
-  then have "r \<noteq> 0" using assms(2) by auto
-  then have "1/l = ereal(1/r)" using real by (simp add: one_ereal_def)
-  define v where "v = (\<lambda>n. real_of_ereal(u n))"
-  have ureal: "eventually (\<lambda>n. u n = ereal(v n)) F" unfolding v_def using real_lim_then_eventually_real assms(1) real by auto
-  then have "((\<lambda>n. ereal(v n)) \<longlongrightarrow> ereal r) F" using assms real v_def by auto
-  then have *: "((\<lambda>n. v n) \<longlongrightarrow> r) F" by auto
-  then have "((\<lambda>n. 1/v n) \<longlongrightarrow> 1/r) F" using \<open>r \<noteq> 0\<close> tendsto_inverse_real by auto
-  then have lim: "((\<lambda>n. ereal(1/v n)) \<longlongrightarrow> 1/l) F" using \<open>1/l = ereal(1/r)\<close> by auto
-
-  have "r \<in> -{0}" "open (-{(0::real)})" using \<open>r \<noteq> 0\<close> by auto
-  then have "eventually (\<lambda>n. v n \<in> -{0}) F" using * using topological_tendstoD by blast
-  then have "eventually (\<lambda>n. v n \<noteq> 0) F" by auto
-  moreover
-  {
-    fix n assume H: "v n \<noteq> 0" "u n = ereal(v n)"
-    then have "ereal(1/v n) = 1/ereal(v n)" by (simp add: one_ereal_def)
-    then have "ereal(1/v n) = 1/u n" using H(2) by simp
-  }
-  ultimately have "eventually (\<lambda>n. ereal(1/v n) = 1/u n) F" using ureal eventually_elim2 by force
-  with Lim_transform_eventually[OF this lim] show ?thesis by simp
-next
-  case (PInf)
-  then have "1/l = 0" by auto
-  then show ?thesis using tendsto_inverse_ereal_PInf assms PInf by auto
-next
-  case (MInf)
-  then have "1/l = 0" by auto
-  have "1/z = -1/ -z" if "z < 0" for z::ereal
-    apply (cases z) using divide_ereal_def \<open> z < 0 \<close> by auto
-  moreover have "eventually (\<lambda>n. u n < 0) F" by (metis (no_types) MInf assms(1) tendsto_MInfty zero_ereal_def)
-  ultimately have *: "eventually (\<lambda>n. -1/-u n = 1/u n) F" by (simp add: eventually_mono)
-
-  define v where "v = (\<lambda>n. - u n)"
-  have "(v \<longlongrightarrow> \<infinity>) F" unfolding v_def using MInf assms(1) tendsto_uminus_ereal by fastforce
-  then have "((\<lambda>n. 1/v n) \<longlongrightarrow> 0) F" using tendsto_inverse_ereal_PInf by auto
-  then have "((\<lambda>n. -1/v n) \<longlongrightarrow> 0) F" using tendsto_uminus_ereal by fastforce
-  then show ?thesis unfolding v_def using Lim_transform_eventually[OF *] \<open> 1/l = 0 \<close> by auto
-qed
-
-lemma tendsto_divide_ereal [tendsto_intros]:
-  fixes f g::"_ \<Rightarrow> ereal"
-  assumes "(f \<longlongrightarrow> l) F" "(g \<longlongrightarrow> m) F" "m \<noteq> 0" "\<not>(abs(l) = \<infinity> \<and> abs(m) = \<infinity>)"
-  shows "((\<lambda>x. f x / g x) \<longlongrightarrow> l / m) F"
-proof -
-  define h where "h = (\<lambda>x. 1/ g x)"
-  have *: "(h \<longlongrightarrow> 1/m) F" unfolding h_def using assms(2) assms(3) tendsto_inverse_ereal by auto
-  have "((\<lambda>x. f x * h x) \<longlongrightarrow> l * (1/m)) F"
-    apply (rule tendsto_mult_ereal[OF assms(1) *]) using assms(3) assms(4) by (auto simp add: divide_ereal_def)
-  moreover have "f x * h x = f x / g x" for x unfolding h_def by (simp add: divide_ereal_def)
-  moreover have "l * (1/m) = l/m" by (simp add: divide_ereal_def)
-  ultimately show ?thesis unfolding h_def using Lim_transform_eventually by auto
-qed
-
-
-subsubsection \<open>Further limits\<close>
-
-lemma id_nat_ereal_tendsto_PInf [tendsto_intros]:
-  "(\<lambda> n::nat. real n) \<longlonglongrightarrow> \<infinity>"
-by (simp add: filterlim_real_sequentially tendsto_PInfty_eq_at_top)
-
-lemma tendsto_at_top_pseudo_inverse [tendsto_intros]:
-  fixes u::"nat \<Rightarrow> nat"
-  assumes "LIM n sequentially. u n :> at_top"
-  shows "LIM n sequentially. Inf {N. u N \<ge> n} :> at_top"
-proof -
-  {
-    fix C::nat
-    define M where "M = Max {u n| n. n \<le> C}+1"
-    {
-      fix n assume "n \<ge> M"
-      have "eventually (\<lambda>N. u N \<ge> n) sequentially" using assms
-        by (simp add: filterlim_at_top)
-      then have *: "{N. u N \<ge> n} \<noteq> {}" by force
-
-      have "N > C" if "u N \<ge> n" for N
-      proof (rule ccontr)
-        assume "\<not>(N > C)"
-        have "u N \<le> Max {u n| n. n \<le> C}"
-          apply (rule Max_ge) using \<open>\<not>(N > C)\<close> by auto
-        then show False using \<open>u N \<ge> n\<close> \<open>n \<ge> M\<close> unfolding M_def by auto
-      qed
-      then have **: "{N. u N \<ge> n} \<subseteq> {C..}" by fastforce
-      have "Inf {N. u N \<ge> n} \<ge> C"
-        by (metis "*" "**" Inf_nat_def1 atLeast_iff subset_eq)
-    }
-    then have "eventually (\<lambda>n. Inf {N. u N \<ge> n} \<ge> C) sequentially"
-      using eventually_sequentially by auto
-  }
-  then show ?thesis using filterlim_at_top by auto
-qed
-
-lemma pseudo_inverse_finite_set:
-  fixes u::"nat \<Rightarrow> nat"
-  assumes "LIM n sequentially. u n :> at_top"
-  shows "finite {N. u N \<le> n}"
-proof -
-  fix n
-  have "eventually (\<lambda>N. u N \<ge> n+1) sequentially" using assms
-    by (simp add: filterlim_at_top)
-  then obtain N1 where N1: "\<And>N. N \<ge> N1 \<Longrightarrow> u N \<ge> n + 1"
-    using eventually_sequentially by auto
-  have "{N. u N \<le> n} \<subseteq> {..<N1}"
-    apply auto using N1 by (metis Suc_eq_plus1 not_less not_less_eq_eq)
-  then show "finite {N. u N \<le> n}" by (simp add: finite_subset)
-qed
-
-lemma tendsto_at_top_pseudo_inverse2 [tendsto_intros]:
-  fixes u::"nat \<Rightarrow> nat"
-  assumes "LIM n sequentially. u n :> at_top"
-  shows "LIM n sequentially. Max {N. u N \<le> n} :> at_top"
-proof -
-  {
-    fix N0::nat
-    have "N0 \<le> Max {N. u N \<le> n}" if "n \<ge> u N0" for n
-      apply (rule Max.coboundedI) using pseudo_inverse_finite_set[OF assms] that by auto
-    then have "eventually (\<lambda>n. N0 \<le> Max {N. u N \<le> n}) sequentially"
-      using eventually_sequentially by blast
-  }
-  then show ?thesis using filterlim_at_top by auto
-qed
-
-lemma ereal_truncation_top [tendsto_intros]:
-  fixes x::ereal
-  shows "(\<lambda>n::nat. min x n) \<longlonglongrightarrow> x"
-proof (cases x)
-  case (real r)
-  then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
-  then have "min x n = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
-  then have "eventually (\<lambda>n. min x n = x) sequentially" using eventually_at_top_linorder by blast
-  then show ?thesis by (simp add: Lim_eventually)
-next
-  case (PInf)
-  then have "min x n = n" for n::nat by (auto simp add: min_def)
-  then show ?thesis using id_nat_ereal_tendsto_PInf PInf by auto
-next
-  case (MInf)
-  then have "min x n = x" for n::nat by (auto simp add: min_def)
-  then show ?thesis by auto
-qed
-
-lemma ereal_truncation_real_top [tendsto_intros]:
-  fixes x::ereal
-  assumes "x \<noteq> - \<infinity>"
-  shows "(\<lambda>n::nat. real_of_ereal(min x n)) \<longlonglongrightarrow> x"
-proof (cases x)
-  case (real r)
-  then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
-  then have "min x n = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
-  then have "real_of_ereal(min x n) = r" if "n \<ge> K" for n using real that by auto
-  then have "eventually (\<lambda>n. real_of_ereal(min x n) = r) sequentially" using eventually_at_top_linorder by blast
-  then have "(\<lambda>n. real_of_ereal(min x n)) \<longlonglongrightarrow> r" by (simp add: Lim_eventually)
-  then show ?thesis using real by auto
-next
-  case (PInf)
-  then have "real_of_ereal(min x n) = n" for n::nat by (auto simp add: min_def)
-  then show ?thesis using id_nat_ereal_tendsto_PInf PInf by auto
-qed (simp add: assms)
-
-lemma ereal_truncation_bottom [tendsto_intros]:
-  fixes x::ereal
-  shows "(\<lambda>n::nat. max x (- real n)) \<longlonglongrightarrow> x"
-proof (cases x)
-  case (real r)
-  then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
-  then have "max x (-real n) = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
-  then have "eventually (\<lambda>n. max x (-real n) = x) sequentially" using eventually_at_top_linorder by blast
-  then show ?thesis by (simp add: Lim_eventually)
-next
-  case (MInf)
-  then have "max x (-real n) = (-1)* ereal(real n)" for n::nat by (auto simp add: max_def)
-  moreover have "(\<lambda>n. (-1)* ereal(real n)) \<longlonglongrightarrow> -\<infinity>"
-    using tendsto_cmult_ereal[of "-1", OF _ id_nat_ereal_tendsto_PInf] by (simp add: one_ereal_def)
-  ultimately show ?thesis using MInf by auto
-next
-  case (PInf)
-  then have "max x (-real n) = x" for n::nat by (auto simp add: max_def)
-  then show ?thesis by auto
-qed
-
-lemma ereal_truncation_real_bottom [tendsto_intros]:
-  fixes x::ereal
-  assumes "x \<noteq> \<infinity>"
-  shows "(\<lambda>n::nat. real_of_ereal(max x (- real n))) \<longlonglongrightarrow> x"
-proof (cases x)
-  case (real r)
-  then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
-  then have "max x (-real n) = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
-  then have "real_of_ereal(max x (-real n)) = r" if "n \<ge> K" for n using real that by auto
-  then have "eventually (\<lambda>n. real_of_ereal(max x (-real n)) = r) sequentially" using eventually_at_top_linorder by blast
-  then have "(\<lambda>n. real_of_ereal(max x (-real n))) \<longlonglongrightarrow> r" by (simp add: Lim_eventually)
-  then show ?thesis using real by auto
-next
-  case (MInf)
-  then have "real_of_ereal(max x (-real n)) = (-1)* ereal(real n)" for n::nat by (auto simp add: max_def)
-  moreover have "(\<lambda>n. (-1)* ereal(real n)) \<longlonglongrightarrow> -\<infinity>"
-    using tendsto_cmult_ereal[of "-1", OF _ id_nat_ereal_tendsto_PInf] by (simp add: one_ereal_def)
-  ultimately show ?thesis using MInf by auto
-qed (simp add: assms)
-
-text \<open>the next one is copied from \verb+tendsto_sum+.\<close>
-lemma tendsto_sum_ereal [tendsto_intros]:
-  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> ereal"
-  assumes "\<And>i. i \<in> S \<Longrightarrow> (f i \<longlongrightarrow> a i) F"
-          "\<And>i. abs(a i) \<noteq> \<infinity>"
-  shows "((\<lambda>x. \<Sum>i\<in>S. f i x) \<longlongrightarrow> (\<Sum>i\<in>S. a i)) F"
-proof (cases "finite S")
-  assume "finite S" then show ?thesis using assms
-    by (induct, simp, simp add: tendsto_add_ereal_general2 assms)
-qed(simp)
-
-subsubsection \<open>Limsups and liminfs\<close>
-
-lemma limsup_finite_then_bounded:
-  fixes u::"nat \<Rightarrow> real"
-  assumes "limsup u < \<infinity>"
-  shows "\<exists>C. \<forall>n. u n \<le> C"
-proof -
-  obtain C where C: "limsup u < C" "C < \<infinity>" using assms ereal_dense2 by blast
-  then have "C = ereal(real_of_ereal C)" using ereal_real by force
-  have "eventually (\<lambda>n. u n < C) sequentially" using C(1) unfolding Limsup_def
-    apply (auto simp add: INF_less_iff)
-    using SUP_lessD eventually_mono by fastforce
-  then obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> u n < C" using eventually_sequentially by auto
-  define D where "D = max (real_of_ereal C) (Max {u n |n. n \<le> N})"
-  have "\<And>n. u n \<le> D"
-  proof -
-    fix n show "u n \<le> D"
-    proof (cases)
-      assume *: "n \<le> N"
-      have "u n \<le> Max {u n |n. n \<le> N}" by (rule Max_ge, auto simp add: *)
-      then show "u n \<le> D" unfolding D_def by linarith
-    next
-      assume "\<not>(n \<le> N)"
-      then have "n \<ge> N" by simp
-      then have "u n < C" using N by auto
-      then have "u n < real_of_ereal C" using \<open>C = ereal(real_of_ereal C)\<close> less_ereal.simps(1) by fastforce
-      then show "u n \<le> D" unfolding D_def by linarith
-    qed
-  qed
-  then show ?thesis by blast
-qed
-
-lemma liminf_finite_then_bounded_below:
-  fixes u::"nat \<Rightarrow> real"
-  assumes "liminf u > -\<infinity>"
-  shows "\<exists>C. \<forall>n. u n \<ge> C"
-proof -
-  obtain C where C: "liminf u > C" "C > -\<infinity>" using assms using ereal_dense2 by blast
-  then have "C = ereal(real_of_ereal C)" using ereal_real by force
-  have "eventually (\<lambda>n. u n > C) sequentially" using C(1) unfolding Liminf_def
-    apply (auto simp add: less_SUP_iff)
-    using eventually_elim2 less_INF_D by fastforce
-  then obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> u n > C" using eventually_sequentially by auto
-  define D where "D = min (real_of_ereal C) (Min {u n |n. n \<le> N})"
-  have "\<And>n. u n \<ge> D"
-  proof -
-    fix n show "u n \<ge> D"
-    proof (cases)
-      assume *: "n \<le> N"
-      have "u n \<ge> Min {u n |n. n \<le> N}" by (rule Min_le, auto simp add: *)
-      then show "u n \<ge> D" unfolding D_def by linarith
-    next
-      assume "\<not>(n \<le> N)"
-      then have "n \<ge> N" by simp
-      then have "u n > C" using N by auto
-      then have "u n > real_of_ereal C" using \<open>C = ereal(real_of_ereal C)\<close> less_ereal.simps(1) by fastforce
-      then show "u n \<ge> D" unfolding D_def by linarith
-    qed
-  qed
-  then show ?thesis by blast
-qed
-
-lemma liminf_upper_bound:
-  fixes u:: "nat \<Rightarrow> ereal"
-  assumes "liminf u < l"
-  shows "\<exists>N>k. u N < l"
-by (metis assms gt_ex less_le_trans liminf_bounded_iff not_less)
-
-text \<open>The following statement about limsups is reduced to a statement about limits using
-subsequences thanks to \verb+limsup_subseq_lim+. The statement for limits follows for instance from
-\verb+tendsto_add_ereal_general+.\<close>
-
-lemma ereal_limsup_add_mono:
-  fixes u v::"nat \<Rightarrow> ereal"
-  shows "limsup (\<lambda>n. u n + v n) \<le> limsup u + limsup v"
-proof (cases)
-  assume "(limsup u = \<infinity>) \<or> (limsup v = \<infinity>)"
-  then have "limsup u + limsup v = \<infinity>" by simp
-  then show ?thesis by auto
-next
-  assume "\<not>((limsup u = \<infinity>) \<or> (limsup v = \<infinity>))"
-  then have "limsup u < \<infinity>" "limsup v < \<infinity>" by auto
-
-  define w where "w = (\<lambda>n. u n + v n)"
-  obtain r where r: "strict_mono r" "(w o r) \<longlonglongrightarrow> limsup w" using limsup_subseq_lim by auto
-  obtain s where s: "strict_mono s" "(u o r o s) \<longlonglongrightarrow> limsup (u o r)" using limsup_subseq_lim by auto
-  obtain t where t: "strict_mono t" "(v o r o s o t) \<longlonglongrightarrow> limsup (v o r o s)" using limsup_subseq_lim by auto
-
-  define a where "a = r o s o t"
-  have "strict_mono a" using r s t by (simp add: a_def strict_mono_o)
-  have l:"(w o a) \<longlonglongrightarrow> limsup w"
-         "(u o a) \<longlonglongrightarrow> limsup (u o r)"
-         "(v o a) \<longlonglongrightarrow> limsup (v o r o s)"
-  apply (metis (no_types, lifting) r(2) s(1) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
-  apply (metis (no_types, lifting) s(2) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
-  apply (metis (no_types, lifting) t(2) a_def comp_assoc)
-  done
-
-  have "limsup (u o r) \<le> limsup u" by (simp add: limsup_subseq_mono r(1))
-  then have a: "limsup (u o r) \<noteq> \<infinity>" using \<open>limsup u < \<infinity>\<close> by auto
-  have "limsup (v o r o s) \<le> limsup v" 
-    by (simp add: comp_assoc limsup_subseq_mono r(1) s(1) strict_mono_o)
-  then have b: "limsup (v o r o s) \<noteq> \<infinity>" using \<open>limsup v < \<infinity>\<close> by auto
-
-  have "(\<lambda>n. (u o a) n + (v o a) n) \<longlonglongrightarrow> limsup (u o r) + limsup (v o r o s)"
-    using l tendsto_add_ereal_general a b by fastforce
-  moreover have "(\<lambda>n. (u o a) n + (v o a) n) = (w o a)" unfolding w_def by auto
-  ultimately have "(w o a) \<longlonglongrightarrow> limsup (u o r) + limsup (v o r o s)" by simp
-  then have "limsup w = limsup (u o r) + limsup (v o r o s)" using l(1) LIMSEQ_unique by blast
-  then have "limsup w \<le> limsup u + limsup v"
-    using \<open>limsup (u o r) \<le> limsup u\<close> \<open>limsup (v o r o s) \<le> limsup v\<close> ereal_add_mono by simp
-  then show ?thesis unfolding w_def by simp
-qed
-
-text \<open>There is an asymmetry between liminfs and limsups in ereal, as $\infty + (-\infty) = \infty$.
-This explains why there are more assumptions in the next lemma dealing with liminfs that in the
-previous one about limsups.\<close>
-
-lemma ereal_liminf_add_mono:
-  fixes u v::"nat \<Rightarrow> ereal"
-  assumes "\<not>((liminf u = \<infinity> \<and> liminf v = -\<infinity>) \<or> (liminf u = -\<infinity> \<and> liminf v = \<infinity>))"
-  shows "liminf (\<lambda>n. u n + v n) \<ge> liminf u + liminf v"
-proof (cases)
-  assume "(liminf u = -\<infinity>) \<or> (liminf v = -\<infinity>)"
-  then have *: "liminf u + liminf v = -\<infinity>" using assms by auto
-  show ?thesis by (simp add: *)
-next
-  assume "\<not>((liminf u = -\<infinity>) \<or> (liminf v = -\<infinity>))"
-  then have "liminf u > -\<infinity>" "liminf v > -\<infinity>" by auto
-
-  define w where "w = (\<lambda>n. u n + v n)"
-  obtain r where r: "strict_mono r" "(w o r) \<longlonglongrightarrow> liminf w" using liminf_subseq_lim by auto
-  obtain s where s: "strict_mono s" "(u o r o s) \<longlonglongrightarrow> liminf (u o r)" using liminf_subseq_lim by auto
-  obtain t where t: "strict_mono t" "(v o r o s o t) \<longlonglongrightarrow> liminf (v o r o s)" using liminf_subseq_lim by auto
-
-  define a where "a = r o s o t"
-  have "strict_mono a" using r s t by (simp add: a_def strict_mono_o)
-  have l:"(w o a) \<longlonglongrightarrow> liminf w"
-         "(u o a) \<longlonglongrightarrow> liminf (u o r)"
-         "(v o a) \<longlonglongrightarrow> liminf (v o r o s)"
-  apply (metis (no_types, lifting) r(2) s(1) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
-  apply (metis (no_types, lifting) s(2) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
-  apply (metis (no_types, lifting) t(2) a_def comp_assoc)
-  done
-
-  have "liminf (u o r) \<ge> liminf u" by (simp add: liminf_subseq_mono r(1))
-  then have a: "liminf (u o r) \<noteq> -\<infinity>" using \<open>liminf u > -\<infinity>\<close> by auto
-  have "liminf (v o r o s) \<ge> liminf v" 
-    by (simp add: comp_assoc liminf_subseq_mono r(1) s(1) strict_mono_o)
-  then have b: "liminf (v o r o s) \<noteq> -\<infinity>" using \<open>liminf v > -\<infinity>\<close> by auto
-
-  have "(\<lambda>n. (u o a) n + (v o a) n) \<longlonglongrightarrow> liminf (u o r) + liminf (v o r o s)"
-    using l tendsto_add_ereal_general a b by fastforce
-  moreover have "(\<lambda>n. (u o a) n + (v o a) n) = (w o a)" unfolding w_def by auto
-  ultimately have "(w o a) \<longlonglongrightarrow> liminf (u o r) + liminf (v o r o s)" by simp
-  then have "liminf w = liminf (u o r) + liminf (v o r o s)" using l(1) LIMSEQ_unique by blast
-  then have "liminf w \<ge> liminf u + liminf v"
-    using \<open>liminf (u o r) \<ge> liminf u\<close> \<open>liminf (v o r o s) \<ge> liminf v\<close> ereal_add_mono by simp
-  then show ?thesis unfolding w_def by simp
-qed
-
-lemma ereal_limsup_lim_add:
-  fixes u v::"nat \<Rightarrow> ereal"
-  assumes "u \<longlonglongrightarrow> a" "abs(a) \<noteq> \<infinity>"
-  shows "limsup (\<lambda>n. u n + v n) = a + limsup v"
-proof -
-  have "limsup u = a" using assms(1) using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
-  have "(\<lambda>n. -u n) \<longlonglongrightarrow> -a" using assms(1) by auto
-  then have "limsup (\<lambda>n. -u n) = -a" using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
-
-  have "limsup (\<lambda>n. u n + v n) \<le> limsup u + limsup v"
-    by (rule ereal_limsup_add_mono)
-  then have up: "limsup (\<lambda>n. u n + v n) \<le> a + limsup v" using \<open>limsup u = a\<close> by simp
-
-  have a: "limsup (\<lambda>n. (u n + v n) + (-u n)) \<le> limsup (\<lambda>n. u n + v n) + limsup (\<lambda>n. -u n)"
-    by (rule ereal_limsup_add_mono)
-  have "eventually (\<lambda>n. u n = ereal(real_of_ereal(u n))) sequentially" using assms
-    real_lim_then_eventually_real by auto
-  moreover have "\<And>x. x = ereal(real_of_ereal(x)) \<Longrightarrow> x + (-x) = 0"
-    by (metis plus_ereal.simps(1) right_minus uminus_ereal.simps(1) zero_ereal_def)
-  ultimately have "eventually (\<lambda>n. u n + (-u n) = 0) sequentially"
-    by (metis (mono_tags, lifting) eventually_mono)
-  moreover have "\<And>n. u n + (-u n) = 0 \<Longrightarrow> u n + v n + (-u n) = v n"
-    by (metis add.commute add.left_commute add.left_neutral)
-  ultimately have "eventually (\<lambda>n. u n + v n + (-u n) = v n) sequentially"
-    using eventually_mono by force
-  then have "limsup v = limsup (\<lambda>n. u n + v n + (-u n))" using Limsup_eq by force
-  then have "limsup v \<le> limsup (\<lambda>n. u n + v n) -a" using a \<open>limsup (\<lambda>n. -u n) = -a\<close> by (simp add: minus_ereal_def)
-  then have "limsup (\<lambda>n. u n + v n) \<ge> a + limsup v" using assms(2) by (metis add.commute ereal_le_minus)
-  then show ?thesis using up by simp
-qed
-
-lemma ereal_limsup_lim_mult:
-  fixes u v::"nat \<Rightarrow> ereal"
-  assumes "u \<longlonglongrightarrow> a" "a>0" "a \<noteq> \<infinity>"
-  shows "limsup (\<lambda>n. u n * v n) = a * limsup v"
-proof -
-  define w where "w = (\<lambda>n. u n * v n)"
-  obtain r where r: "strict_mono r" "(v o r) \<longlonglongrightarrow> limsup v" using limsup_subseq_lim by auto
-  have "(u o r) \<longlonglongrightarrow> a" using assms(1) LIMSEQ_subseq_LIMSEQ r by auto
-  with tendsto_mult_ereal[OF this r(2)] have "(\<lambda>n. (u o r) n * (v o r) n) \<longlonglongrightarrow> a * limsup v" using assms(2) assms(3) by auto
-  moreover have "\<And>n. (w o r) n = (u o r) n * (v o r) n" unfolding w_def by auto
-  ultimately have "(w o r) \<longlonglongrightarrow> a * limsup v" unfolding w_def by presburger
-  then have "limsup (w o r) = a * limsup v" by (simp add: tendsto_iff_Liminf_eq_Limsup)
-  then have I: "limsup w \<ge> a * limsup v" by (metis limsup_subseq_mono r(1))
-
-  obtain s where s: "strict_mono s" "(w o s) \<longlonglongrightarrow> limsup w" using limsup_subseq_lim by auto
-  have *: "(u o s) \<longlonglongrightarrow> a" using assms(1) LIMSEQ_subseq_LIMSEQ s by auto
-  have "eventually (\<lambda>n. (u o s) n > 0) sequentially" using assms(2) * order_tendsto_iff by blast
-  moreover have "eventually (\<lambda>n. (u o s) n < \<infinity>) sequentially" using assms(3) * order_tendsto_iff by blast
-  moreover have "(w o s) n / (u o s) n = (v o s) n" if "(u o s) n > 0" "(u o s) n < \<infinity>" for n
-    unfolding w_def using that by (auto simp add: ereal_divide_eq)
-  ultimately have "eventually (\<lambda>n. (w o s) n / (u o s) n = (v o s) n) sequentially" using eventually_elim2 by force
-  moreover have "(\<lambda>n. (w o s) n / (u o s) n) \<longlonglongrightarrow> (limsup w) / a"
-    apply (rule tendsto_divide_ereal[OF s(2) *]) using assms(2) assms(3) by auto
-  ultimately have "(v o s) \<longlonglongrightarrow> (limsup w) / a" using Lim_transform_eventually by fastforce
-  then have "limsup (v o s) = (limsup w) / a" by (simp add: tendsto_iff_Liminf_eq_Limsup)
-  then have "limsup v \<ge> (limsup w) / a" by (metis limsup_subseq_mono s(1))
-  then have "a * limsup v \<ge> limsup w" using assms(2) assms(3) by (simp add: ereal_divide_le_pos)
-  then show ?thesis using I unfolding w_def by auto
-qed
-
-lemma ereal_liminf_lim_mult:
-  fixes u v::"nat \<Rightarrow> ereal"
-  assumes "u \<longlonglongrightarrow> a" "a>0" "a \<noteq> \<infinity>"
-  shows "liminf (\<lambda>n. u n * v n) = a * liminf v"
-proof -
-  define w where "w = (\<lambda>n. u n * v n)"
-  obtain r where r: "strict_mono r" "(v o r) \<longlonglongrightarrow> liminf v" using liminf_subseq_lim by auto
-  have "(u o r) \<longlonglongrightarrow> a" using assms(1) LIMSEQ_subseq_LIMSEQ r by auto
-  with tendsto_mult_ereal[OF this r(2)] have "(\<lambda>n. (u o r) n * (v o r) n) \<longlonglongrightarrow> a * liminf v" using assms(2) assms(3) by auto
-  moreover have "\<And>n. (w o r) n = (u o r) n * (v o r) n" unfolding w_def by auto
-  ultimately have "(w o r) \<longlonglongrightarrow> a * liminf v" unfolding w_def by presburger
-  then have "liminf (w o r) = a * liminf v" by (simp add: tendsto_iff_Liminf_eq_Limsup)
-  then have I: "liminf w \<le> a * liminf v" by (metis liminf_subseq_mono r(1))
-
-  obtain s where s: "strict_mono s" "(w o s) \<longlonglongrightarrow> liminf w" using liminf_subseq_lim by auto
-  have *: "(u o s) \<longlonglongrightarrow> a" using assms(1) LIMSEQ_subseq_LIMSEQ s by auto
-  have "eventually (\<lambda>n. (u o s) n > 0) sequentially" using assms(2) * order_tendsto_iff by blast
-  moreover have "eventually (\<lambda>n. (u o s) n < \<infinity>) sequentially" using assms(3) * order_tendsto_iff by blast
-  moreover have "(w o s) n / (u o s) n = (v o s) n" if "(u o s) n > 0" "(u o s) n < \<infinity>" for n
-    unfolding w_def using that by (auto simp add: ereal_divide_eq)
-  ultimately have "eventually (\<lambda>n. (w o s) n / (u o s) n = (v o s) n) sequentially" using eventually_elim2 by force
-  moreover have "(\<lambda>n. (w o s) n / (u o s) n) \<longlonglongrightarrow> (liminf w) / a"
-    apply (rule tendsto_divide_ereal[OF s(2) *]) using assms(2) assms(3) by auto
-  ultimately have "(v o s) \<longlonglongrightarrow> (liminf w) / a" using Lim_transform_eventually by fastforce
-  then have "liminf (v o s) = (liminf w) / a" by (simp add: tendsto_iff_Liminf_eq_Limsup)
-  then have "liminf v \<le> (liminf w) / a" by (metis liminf_subseq_mono s(1))
-  then have "a * liminf v \<le> liminf w" using assms(2) assms(3) by (simp add: ereal_le_divide_pos)
-  then show ?thesis using I unfolding w_def by auto
-qed
-
-lemma ereal_liminf_lim_add:
-  fixes u v::"nat \<Rightarrow> ereal"
-  assumes "u \<longlonglongrightarrow> a" "abs(a) \<noteq> \<infinity>"
-  shows "liminf (\<lambda>n. u n + v n) = a + liminf v"
-proof -
-  have "liminf u = a" using assms(1) tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
-  then have *: "abs(liminf u) \<noteq> \<infinity>" using assms(2) by auto
-  have "(\<lambda>n. -u n) \<longlonglongrightarrow> -a" using assms(1) by auto
-  then have "liminf (\<lambda>n. -u n) = -a" using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
-  then have **: "abs(liminf (\<lambda>n. -u n)) \<noteq> \<infinity>" using assms(2) by auto
-
-  have "liminf (\<lambda>n. u n + v n) \<ge> liminf u + liminf v"
-    apply (rule ereal_liminf_add_mono) using * by auto
-  then have up: "liminf (\<lambda>n. u n + v n) \<ge> a + liminf v" using \<open>liminf u = a\<close> by simp
-
-  have a: "liminf (\<lambda>n. (u n + v n) + (-u n)) \<ge> liminf (\<lambda>n. u n + v n) + liminf (\<lambda>n. -u n)"
-    apply (rule ereal_liminf_add_mono) using ** by auto
-  have "eventually (\<lambda>n. u n = ereal(real_of_ereal(u n))) sequentially" using assms
-    real_lim_then_eventually_real by auto
-  moreover have "\<And>x. x = ereal(real_of_ereal(x)) \<Longrightarrow> x + (-x) = 0"
-    by (metis plus_ereal.simps(1) right_minus uminus_ereal.simps(1) zero_ereal_def)
-  ultimately have "eventually (\<lambda>n. u n + (-u n) = 0) sequentially"
-    by (metis (mono_tags, lifting) eventually_mono)
-  moreover have "\<And>n. u n + (-u n) = 0 \<Longrightarrow> u n + v n + (-u n) = v n"
-    by (metis add.commute add.left_commute add.left_neutral)
-  ultimately have "eventually (\<lambda>n. u n + v n + (-u n) = v n) sequentially"
-    using eventually_mono by force
-  then have "liminf v = liminf (\<lambda>n. u n + v n + (-u n))" using Liminf_eq by force
-  then have "liminf v \<ge> liminf (\<lambda>n. u n + v n) -a" using a \<open>liminf (\<lambda>n. -u n) = -a\<close> by (simp add: minus_ereal_def)
-  then have "liminf (\<lambda>n. u n + v n) \<le> a + liminf v" using assms(2) by (metis add.commute ereal_minus_le)
-  then show ?thesis using up by simp
-qed
-
-lemma ereal_liminf_limsup_add:
-  fixes u v::"nat \<Rightarrow> ereal"
-  shows "liminf (\<lambda>n. u n + v n) \<le> liminf u + limsup v"
-proof (cases)
-  assume "limsup v = \<infinity> \<or> liminf u = \<infinity>"
-  then show ?thesis by auto
-next
-  assume "\<not>(limsup v = \<infinity> \<or> liminf u = \<infinity>)"
-  then have "limsup v < \<infinity>" "liminf u < \<infinity>" by auto
-
-  define w where "w = (\<lambda>n. u n + v n)"
-  obtain r where r: "strict_mono r" "(u o r) \<longlonglongrightarrow> liminf u" using liminf_subseq_lim by auto
-  obtain s where s: "strict_mono s" "(w o r o s) \<longlonglongrightarrow> liminf (w o r)" using liminf_subseq_lim by auto
-  obtain t where t: "strict_mono t" "(v o r o s o t) \<longlonglongrightarrow> limsup (v o r o s)" using limsup_subseq_lim by auto
-
-  define a where "a = r o s o t"
-  have "strict_mono a" using r s t by (simp add: a_def strict_mono_o)
-  have l:"(u o a) \<longlonglongrightarrow> liminf u"
-         "(w o a) \<longlonglongrightarrow> liminf (w o r)"
-         "(v o a) \<longlonglongrightarrow> limsup (v o r o s)"
-  apply (metis (no_types, lifting) r(2) s(1) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
-  apply (metis (no_types, lifting) s(2) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
-  apply (metis (no_types, lifting) t(2) a_def comp_assoc)
-  done
-
-  have "liminf (w o r) \<ge> liminf w" by (simp add: liminf_subseq_mono r(1))
-  have "limsup (v o r o s) \<le> limsup v" 
-    by (simp add: comp_assoc limsup_subseq_mono r(1) s(1) strict_mono_o)
-  then have b: "limsup (v o r o s) < \<infinity>" using \<open>limsup v < \<infinity>\<close> by auto
-
-  have "(\<lambda>n. (u o a) n + (v o a) n) \<longlonglongrightarrow> liminf u + limsup (v o r o s)"
-    apply (rule tendsto_add_ereal_general) using b \<open>liminf u < \<infinity>\<close> l(1) l(3) by force+
-  moreover have "(\<lambda>n. (u o a) n + (v o a) n) = (w o a)" unfolding w_def by auto
-  ultimately have "(w o a) \<longlonglongrightarrow> liminf u + limsup (v o r o s)" by simp
-  then have "liminf (w o r) = liminf u + limsup (v o r o s)" using l(2) using LIMSEQ_unique by blast
-  then have "liminf w \<le> liminf u + limsup v"
-    using \<open>liminf (w o r) \<ge> liminf w\<close> \<open>limsup (v o r o s) \<le> limsup v\<close>
-    by (metis add_mono_thms_linordered_semiring(2) le_less_trans not_less)
-  then show ?thesis unfolding w_def by simp
-qed
-
-lemma ereal_liminf_limsup_minus:
-  fixes u v::"nat \<Rightarrow> ereal"
-  shows "liminf (\<lambda>n. u n - v n) \<le> limsup u - limsup v"
-  unfolding minus_ereal_def
-  apply (subst add.commute)
-  apply (rule order_trans[OF ereal_liminf_limsup_add])
-  using ereal_Limsup_uminus[of sequentially "\<lambda>n. - v n"]
-  apply (simp add: add.commute)
-  done
-
-
-lemma liminf_minus_ennreal:
-  fixes u v::"nat \<Rightarrow> ennreal"
-  shows "(\<And>n. v n \<le> u n) \<Longrightarrow> liminf (\<lambda>n. u n - v n) \<le> limsup u - limsup v"
-  unfolding liminf_SUP_INF limsup_INF_SUP
-  including ennreal.lifting
-proof (transfer, clarsimp)
-  fix v u :: "nat \<Rightarrow> ereal" assume *: "\<forall>x. 0 \<le> v x" "\<forall>x. 0 \<le> u x" "\<And>n. v n \<le> u n"
-  moreover have "0 \<le> limsup u - limsup v"
-    using * by (intro ereal_diff_positive Limsup_mono always_eventually) simp
-  moreover have "0 \<le> (SUPREMUM {x..} v)" for x
-    using * by (intro SUP_upper2[of x]) auto
-  moreover have "0 \<le> (SUPREMUM {x..} u)" for x
-    using * by (intro SUP_upper2[of x]) auto
-  ultimately show "(SUP n. INF n:{n..}. max 0 (u n - v n))
-            \<le> max 0 ((INF x. max 0 (SUPREMUM {x..} u)) - (INF x. max 0 (SUPREMUM {x..} v)))"
-    by (auto simp: * ereal_diff_positive max.absorb2 liminf_SUP_INF[symmetric] limsup_INF_SUP[symmetric] ereal_liminf_limsup_minus)
-qed
-
 (*
     Notation
 *)
--- a/src/HOL/Analysis/Summation_Tests.thy	Fri Aug 18 22:55:54 2017 +0200
+++ b/src/HOL/Analysis/Summation_Tests.thy	Sun Aug 20 03:35:20 2017 +0200
@@ -10,6 +10,7 @@
   "HOL-Library.Discrete"
   "HOL-Library.Extended_Real"
   "HOL-Library.Liminf_Limsup"
+  "Extended_Real_Limits"
 begin
 
 text \<open>
@@ -707,6 +708,41 @@
     by (intro exI[of _ "of_real r"]) simp
 qed
 
+lemma conv_radius_conv_Sup:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  shows "conv_radius f = Sup {r. \<forall>z. ereal (norm z) < r \<longrightarrow> summable (\<lambda>n. f n * z ^ n)}"
+proof (rule Sup_eqI [symmetric], goal_cases)
+  case (1 r)
+  thus ?case
+    by (intro conv_radius_geI_ex') auto
+next
+  case (2 r)
+  from 2[of 0] have r: "r \<ge> 0" by auto
+  show ?case
+  proof (intro conv_radius_leI_ex' r)
+    fix R assume R: "R > 0" "ereal R > r"
+    with r obtain r' where [simp]: "r = ereal r'" by (cases r) auto
+    show "\<not>summable (\<lambda>n. f n * of_real R ^ n)"
+    proof
+      assume *: "summable (\<lambda>n. f n * of_real R ^ n)"
+      define R' where "R' = (R + r') / 2"
+      from R have R': "R' > r'" "R' < R" by (simp_all add: R'_def)
+      hence "\<forall>z. norm z < R' \<longrightarrow> summable (\<lambda>n. f n * z ^ n)"
+        using powser_inside[OF *] by auto
+      from 2[of R'] and this have "R' \<le> r'" by auto
+      with \<open>R' > r'\<close> show False by simp
+    qed
+  qed
+qed
+
+lemma conv_radius_shift:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  shows   "conv_radius (\<lambda>n. f (n + m)) = conv_radius f"
+  unfolding conv_radius_conv_Sup summable_powser_ignore_initial_segment ..
+
+lemma conv_radius_norm [simp]: "conv_radius (\<lambda>x. norm (f x)) = conv_radius f"
+  by (simp add: conv_radius_def)
+
 lemma conv_radius_ratio_limit_ereal:
   fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
   assumes nz:  "eventually (\<lambda>n. f n \<noteq> 0) sequentially"
@@ -773,6 +809,31 @@
   shows   "conv_radius f = c'"
   using assms by (intro conv_radius_ratio_limit_ereal_nonzero) simp_all
 
+lemma conv_radius_cmult_left:
+  assumes "c \<noteq> (0 :: 'a :: {banach, real_normed_div_algebra})"
+  shows   "conv_radius (\<lambda>n. c * f n) = conv_radius f"
+proof -
+  have "conv_radius (\<lambda>n. c * f n) = 
+          inverse (limsup (\<lambda>n. ereal (root n (norm (c * f n)))))"
+    unfolding conv_radius_def ..
+  also have "(\<lambda>n. ereal (root n (norm (c * f n)))) = 
+               (\<lambda>n. ereal (root n (norm c)) * ereal (root n (norm (f n))))"
+    by (rule ext) (auto simp: norm_mult real_root_mult)
+  also have "limsup \<dots> = ereal 1 * limsup (\<lambda>n. ereal (root n (norm (f n))))"
+    using assms by (intro ereal_limsup_lim_mult tendsto_ereal LIMSEQ_root_const) auto
+  also have "inverse \<dots> = conv_radius f" by (simp add: conv_radius_def)
+  finally show ?thesis .
+qed
+
+lemma conv_radius_cmult_right:
+  assumes "c \<noteq> (0 :: 'a :: {banach, real_normed_div_algebra})"
+  shows   "conv_radius (\<lambda>n. f n * c) = conv_radius f"
+proof -
+  have "conv_radius (\<lambda>n. f n * c) = conv_radius (\<lambda>n. c * f n)"
+    by (simp add: conv_radius_def norm_mult mult.commute)
+  with conv_radius_cmult_left[OF assms, of f] show ?thesis by simp
+qed
+
 lemma conv_radius_mult_power:
   assumes "c \<noteq> (0 :: 'a :: {real_normed_div_algebra,banach})"
   shows   "conv_radius (\<lambda>n. c ^ n * f n) = conv_radius f / ereal (norm c)"
--- a/src/HOL/Limits.thy	Fri Aug 18 22:55:54 2017 +0200
+++ b/src/HOL/Limits.thy	Sun Aug 20 03:35:20 2017 +0200
@@ -1617,6 +1617,17 @@
 qed simp
 
 
+lemma filterlim_divide_at_infinity:
+  fixes f g :: "'a \<Rightarrow> 'a :: real_normed_field"
+  assumes "filterlim f (nhds c) F" "filterlim g (at 0) F" "c \<noteq> 0"
+  shows   "filterlim (\<lambda>x. f x / g x) at_infinity F"
+proof -
+  have "filterlim (\<lambda>x. f x * inverse (g x)) at_infinity F"
+    by (intro tendsto_mult_filterlim_at_infinity[OF assms(1,3)]
+          filterlim_compose [OF filterlim_inverse_at_infinity assms(2)])
+  thus ?thesis by (simp add: field_simps)
+qed
+
 subsection \<open>Floor and Ceiling\<close>
 
 lemma eventually_floor_less:
--- a/src/HOL/Series.thy	Fri Aug 18 22:55:54 2017 +0200
+++ b/src/HOL/Series.thy	Sun Aug 20 03:35:20 2017 +0200
@@ -983,6 +983,20 @@
   finally show ?thesis .
 qed
 
+lemma summable_powser_ignore_initial_segment:
+  fixes f :: "nat \<Rightarrow> 'a :: real_normed_div_algebra"
+  shows "summable (\<lambda>n. f (n + m) * z ^ n) \<longleftrightarrow> summable (\<lambda>n. f n * z ^ n)"
+proof (induction m)
+  case (Suc m)
+  have "summable (\<lambda>n. f (n + Suc m) * z ^ n) = summable (\<lambda>n. f (Suc n + m) * z ^ n)"
+    by simp
+  also have "\<dots> = summable (\<lambda>n. f (n + m) * z ^ n)"
+    by (rule summable_powser_split_head)
+  also have "\<dots> = summable (\<lambda>n. f n * z ^ n)"
+    by (rule Suc.IH)
+  finally show ?case .
+qed simp_all
+
 lemma powser_split_head:
   fixes f :: "nat \<Rightarrow> 'a::{real_normed_div_algebra,banach}"
   assumes "summable (\<lambda>n. f n * z ^ n)"