--- a/src/HOL/Analysis/Borel_Space.thy Thu Oct 13 18:36:06 2016 +0200
+++ b/src/HOL/Analysis/Borel_Space.thy Tue Oct 18 12:01:54 2016 +0200
@@ -1974,4 +1974,189 @@
no_notation
eucl_less (infix "<e" 50)
+lemma borel_measurable_Max2[measurable (raw)]:
+ fixes f::"_ \<Rightarrow> _ \<Rightarrow> 'a::{second_countable_topology, dense_linorder, linorder_topology}"
+ assumes "finite I"
+ and [measurable]: "\<And>i. f i \<in> borel_measurable M"
+ shows "(\<lambda>x. Max{f i x |i. i \<in> I}) \<in> borel_measurable M"
+by (simp add: borel_measurable_Max[OF assms(1), where ?f=f and ?M=M] Setcompr_eq_image)
+
+lemma measurable_compose_n [measurable (raw)]:
+ assumes "T \<in> measurable M M"
+ shows "(T^^n) \<in> measurable M M"
+by (induction n, auto simp add: measurable_compose[OF _ assms])
+
+lemma measurable_real_imp_nat:
+ fixes f::"'a \<Rightarrow> nat"
+ assumes [measurable]: "(\<lambda>x. real(f x)) \<in> borel_measurable M"
+ shows "f \<in> measurable M (count_space UNIV)"
+proof -
+ let ?g = "(\<lambda>x. real(f x))"
+ have "\<And>(n::nat). ?g-`({real n}) \<inter> space M = f-`{n} \<inter> space M" by auto
+ moreover have "\<And>(n::nat). ?g-`({real n}) \<inter> space M \<in> sets M" using assms by measurable
+ ultimately have "\<And>(n::nat). f-`{n} \<inter> space M \<in> sets M" by simp
+ then show ?thesis using measurable_count_space_eq2_countable by blast
+qed
+
+lemma measurable_equality_set [measurable]:
+ fixes f g::"_\<Rightarrow> 'a::{second_countable_topology, t2_space}"
+ assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+ shows "{x \<in> space M. f x = g x} \<in> sets M"
+
+proof -
+ define A where "A = {x \<in> space M. f x = g x}"
+ define B where "B = {y. \<exists>x::'a. y = (x,x)}"
+ have "A = (\<lambda>x. (f x, g x))-`B \<inter> space M" unfolding A_def B_def by auto
+ moreover have "(\<lambda>x. (f x, g x)) \<in> borel_measurable M" by simp
+ moreover have "B \<in> sets borel" unfolding B_def by (simp add: closed_diagonal)
+ ultimately have "A \<in> sets M" by simp
+ then show ?thesis unfolding A_def by simp
+qed
+
+lemma measurable_inequality_set [measurable]:
+ fixes f g::"_ \<Rightarrow> 'a::{second_countable_topology, linorder_topology}"
+ assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+ shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
+ "{x \<in> space M. f x < g x} \<in> sets M"
+ "{x \<in> space M. f x \<ge> g x} \<in> sets M"
+ "{x \<in> space M. f x > g x} \<in> sets M"
+proof -
+ define F where "F = (\<lambda>x. (f x, g x))"
+ have * [measurable]: "F \<in> borel_measurable M" unfolding F_def by simp
+
+ have "{x \<in> space M. f x \<le> g x} = F-`{(x, y) | x y. x \<le> y} \<inter> space M" unfolding F_def by auto
+ moreover have "{(x, y) | x y. x \<le> (y::'a)} \<in> sets borel" using closed_subdiagonal borel_closed by blast
+ ultimately show "{x \<in> space M. f x \<le> g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
+
+ have "{x \<in> space M. f x < g x} = F-`{(x, y) | x y. x < y} \<inter> space M" unfolding F_def by auto
+ moreover have "{(x, y) | x y. x < (y::'a)} \<in> sets borel" using open_subdiagonal borel_open by blast
+ ultimately show "{x \<in> space M. f x < g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
+
+ have "{x \<in> space M. f x \<ge> g x} = F-`{(x, y) | x y. x \<ge> y} \<inter> space M" unfolding F_def by auto
+ moreover have "{(x, y) | x y. x \<ge> (y::'a)} \<in> sets borel" using closed_superdiagonal borel_closed by blast
+ ultimately show "{x \<in> space M. f x \<ge> g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
+
+ have "{x \<in> space M. f x > g x} = F-`{(x, y) | x y. x > y} \<inter> space M" unfolding F_def by auto
+ moreover have "{(x, y) | x y. x > (y::'a)} \<in> sets borel" using open_superdiagonal borel_open by blast
+ ultimately show "{x \<in> space M. f x > g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
+qed
+
+lemma measurable_limit [measurable]:
+ fixes f::"nat \<Rightarrow> 'a \<Rightarrow> 'b::first_countable_topology"
+ assumes [measurable]: "\<And>n::nat. f n \<in> borel_measurable M"
+ shows "Measurable.pred M (\<lambda>x. (\<lambda>n. f n x) \<longlonglongrightarrow> c)"
+proof -
+ obtain A :: "nat \<Rightarrow> 'b set" where A:
+ "\<And>i. open (A i)"
+ "\<And>i. c \<in> A i"
+ "\<And>S. open S \<Longrightarrow> c \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially"
+ by (rule countable_basis_at_decseq) blast
+
+ have [measurable]: "\<And>N i. (f N)-`(A i) \<inter> space M \<in> sets M" using A(1) by auto
+ then have mes: "(\<Inter>i. \<Union>n. \<Inter>N\<in>{n..}. (f N)-`(A i) \<inter> space M) \<in> sets M" by blast
+
+ have "(u \<longlonglongrightarrow> c) \<longleftrightarrow> (\<forall>i. eventually (\<lambda>n. u n \<in> A i) sequentially)" for u::"nat \<Rightarrow> 'b"
+ proof
+ assume "u \<longlonglongrightarrow> c"
+ then have "eventually (\<lambda>n. u n \<in> A i) sequentially" for i using A(1)[of i] A(2)[of i]
+ by (simp add: topological_tendstoD)
+ then show "(\<forall>i. eventually (\<lambda>n. u n \<in> A i) sequentially)" by auto
+ next
+ assume H: "(\<forall>i. eventually (\<lambda>n. u n \<in> A i) sequentially)"
+ show "(u \<longlonglongrightarrow> c)"
+ proof (rule topological_tendstoI)
+ fix S assume "open S" "c \<in> S"
+ with A(3)[OF this] obtain i where "A i \<subseteq> S"
+ using eventually_False_sequentially eventually_mono by blast
+ moreover have "eventually (\<lambda>n. u n \<in> A i) sequentially" using H by simp
+ ultimately show "\<forall>\<^sub>F n in sequentially. u n \<in> S"
+ by (simp add: eventually_mono subset_eq)
+ qed
+ qed
+ then have "{x. (\<lambda>n. f n x) \<longlonglongrightarrow> c} = (\<Inter>i. \<Union>n. \<Inter>N\<in>{n..}. (f N)-`(A i))"
+ by (auto simp add: atLeast_def eventually_at_top_linorder)
+ then have "{x \<in> space M. (\<lambda>n. f n x) \<longlonglongrightarrow> c} = (\<Inter>i. \<Union>n. \<Inter>N\<in>{n..}. (f N)-`(A i) \<inter> space M)"
+ by auto
+ then have "{x \<in> space M. (\<lambda>n. f n x) \<longlonglongrightarrow> c} \<in> sets M" using mes by simp
+ then show ?thesis by auto
+qed
+
+lemma measurable_limit2 [measurable]:
+ fixes u::"nat \<Rightarrow> 'a \<Rightarrow> real"
+ assumes [measurable]: "\<And>n. u n \<in> borel_measurable M" "v \<in> borel_measurable M"
+ shows "Measurable.pred M (\<lambda>x. (\<lambda>n. u n x) \<longlonglongrightarrow> v x)"
+proof -
+ define w where "w = (\<lambda>n x. u n x - v x)"
+ have [measurable]: "w n \<in> borel_measurable M" for n unfolding w_def by auto
+ have "((\<lambda>n. u n x) \<longlonglongrightarrow> v x) \<longleftrightarrow> ((\<lambda>n. w n x) \<longlonglongrightarrow> 0)" for x
+ unfolding w_def using Lim_null by auto
+ then show ?thesis using measurable_limit by auto
+qed
+
+lemma measurable_P_restriction [measurable (raw)]:
+ assumes [measurable]: "Measurable.pred M P" "A \<in> sets M"
+ shows "{x \<in> A. P x} \<in> sets M"
+proof -
+ have "A \<subseteq> space M" using sets.sets_into_space[OF assms(2)].
+ then have "{x \<in> A. P x} = A \<inter> {x \<in> space M. P x}" by blast
+ then show ?thesis by auto
+qed
+
+lemma measurable_sum_nat [measurable (raw)]:
+ fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> nat"
+ assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> measurable M (count_space UNIV)"
+ shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> measurable M (count_space UNIV)"
+proof cases
+ assume "finite S"
+ then show ?thesis using assms by induct auto
+qed simp
+
+
+lemma measurable_abs_powr [measurable]:
+ fixes p::real
+ assumes [measurable]: "f \<in> borel_measurable M"
+ shows "(\<lambda>x. \<bar>f x\<bar> powr p) \<in> borel_measurable M"
+unfolding powr_def by auto
+
+text {* The next one is a variation around \verb+measurable_restrict_space+.*}
+
+lemma measurable_restrict_space3:
+ assumes "f \<in> measurable M N" and
+ "f \<in> A \<rightarrow> B"
+ shows "f \<in> measurable (restrict_space M A) (restrict_space N B)"
+proof -
+ have "f \<in> measurable (restrict_space M A) N" using assms(1) measurable_restrict_space1 by auto
+ then show ?thesis by (metis Int_iff funcsetI funcset_mem
+ measurable_restrict_space2[of f, of "restrict_space M A", of B, of N] assms(2) space_restrict_space)
+qed
+
+text {* The next one is a variation around \verb+measurable_piecewise_restrict+.*}
+
+lemma measurable_piecewise_restrict2:
+ assumes [measurable]: "\<And>n. A n \<in> sets M"
+ and "space M = (\<Union>(n::nat). A n)"
+ "\<And>n. \<exists>h \<in> measurable M N. (\<forall>x \<in> A n. f x = h x)"
+ shows "f \<in> measurable M N"
+proof (rule measurableI)
+ fix B assume [measurable]: "B \<in> sets N"
+ {
+ fix n::nat
+ obtain h where [measurable]: "h \<in> measurable M N" and "\<forall>x \<in> A n. f x = h x" using assms(3) by blast
+ then have *: "f-`B \<inter> A n = h-`B \<inter> A n" by auto
+ have "h-`B \<inter> A n = h-`B \<inter> space M \<inter> A n" using assms(2) sets.sets_into_space by auto
+ then have "h-`B \<inter> A n \<in> sets M" by simp
+ then have "f-`B \<inter> A n \<in> sets M" using * by simp
+ }
+ then have "(\<Union>n. f-`B \<inter> A n) \<in> sets M" by measurable
+ moreover have "f-`B \<inter> space M = (\<Union>n. f-`B \<inter> A n)" using assms(2) by blast
+ ultimately show "f-`B \<inter> space M \<in> sets M" by simp
+next
+ fix x assume "x \<in> space M"
+ then obtain n where "x \<in> A n" using assms(2) by blast
+ obtain h where [measurable]: "h \<in> measurable M N" and "\<forall>x \<in> A n. f x = h x" using assms(3) by blast
+ then have "f x = h x" using `x \<in> A n` by blast
+ moreover have "h x \<in> space N" by (metis measurable_space `x \<in> space M` `h \<in> measurable M N`)
+ ultimately show "f x \<in> space N" by simp
+qed
+
end
--- a/src/HOL/Analysis/Complete_Measure.thy Thu Oct 13 18:36:06 2016 +0200
+++ b/src/HOL/Analysis/Complete_Measure.thy Tue Oct 18 12:01:54 2016 +0200
@@ -791,11 +791,11 @@
by (auto simp add: emeasure_density measurable_completion nn_integral_completion intro!: nn_integral_cong_AE)
qed
-lemma null_sets_subset: "A \<subseteq> B \<Longrightarrow> A \<in> sets M \<Longrightarrow> B \<in> null_sets M \<Longrightarrow> A \<in> null_sets M"
+lemma null_sets_subset: "B \<in> null_sets M \<Longrightarrow> A \<in> sets M \<Longrightarrow> A \<subseteq> B \<Longrightarrow> A \<in> null_sets M"
using emeasure_mono[of A B M] by (simp add: null_sets_def)
lemma (in complete_measure) complete2: "A \<subseteq> B \<Longrightarrow> B \<in> null_sets M \<Longrightarrow> A \<in> null_sets M"
- using complete[of A B] null_sets_subset[of A B M] by simp
+ using complete[of A B] null_sets_subset[of B M A] by simp
lemma (in complete_measure) AE_iff_null_sets: "(AE x in M. P x) \<longleftrightarrow> {x\<in>space M. \<not> P x} \<in> null_sets M"
unfolding eventually_ae_filter by (auto intro: complete2)
--- a/src/HOL/Analysis/Set_Integral.thy Thu Oct 13 18:36:06 2016 +0200
+++ b/src/HOL/Analysis/Set_Integral.thy Tue Oct 18 12:01:54 2016 +0200
@@ -11,6 +11,1339 @@
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 {*Fun.thy*}
+
+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 {*Liminf-Limsup.thy*}
+
+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 `decseq v` 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 `incseq v` 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 {* The next lemma is extremely useful, as it often makes it possible to reduce statements
+about limsups to statements about limits.*}
+
+lemma limsup_subseq_lim:
+ fixes u::"nat \<Rightarrow> 'a :: {complete_linorder, linorder_topology}"
+ shows "\<exists>r. subseq 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 where "subseq r" and mono: "\<And>n m. r n \<le> m \<Longrightarrow> u m \<le> u (r n)"
+ by (auto simp: subseq_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 `subseq r`)
+ 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 `subseq r` 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] `p \<in> {N<..x}` by auto
+ define y where "y = Inf U"
+ then have "y \<in> U" using `U \<noteq> {}` 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 `y \<in> U` U_def by auto
+ ultimately have "y \<notin> {N<..x}" using not_le by blast
+ moreover have "y > N" using `y \<in> U` U_def `p \<in> {N<..x}` 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 `i \<in> {N<..y}` 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 `p \<in> {N<..x}` 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 `u p < u y` by auto
+ qed
+ qed
+ then have "N < y \<and> x < y \<and> (\<forall>i\<in>{N<..y}. u i \<le> u y)" using `y > x` `y > N` 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 "subseq r" using r by (auto simp: subseq_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 `subseq r` 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 `subseq r` 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 `(SUP n. (u o r) n) \<le> limsup u` by simp
+ then have "(u o r) \<longlonglongrightarrow> limsup u" using `(u o r) \<longlonglongrightarrow> (SUP n. (u o r) n)` by simp
+ then show ?thesis using `subseq r` by auto
+qed
+
+lemma liminf_subseq_lim:
+ fixes u::"nat \<Rightarrow> 'a :: {complete_linorder, linorder_topology}"
+ shows "\<exists>r. subseq 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 where "subseq r" and mono: "\<And>n m. r n \<le> m \<Longrightarrow> u m \<ge> u (r n)"
+ by (auto simp: subseq_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 `subseq r`)
+ 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 `subseq r` 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] `p \<in> {N<..x}` by auto
+ define y where "y = Inf U"
+ then have "y \<in> U" using `U \<noteq> {}` 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 `y \<in> U` U_def by auto
+ ultimately have "y \<notin> {N<..x}" using not_le by blast
+ moreover have "y > N" using `y \<in> U` U_def `p \<in> {N<..x}` 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 `i \<in> {N<..y}` 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 `p \<in> {N<..x}` 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 `u p > u y` by auto
+ qed
+ qed
+ then have "N < y \<and> x < y \<and> (\<forall>i\<in>{N<..y}. u i \<ge> u y)" using `y > x` `y > N` 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 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 "subseq r" using r by (auto simp: subseq_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 `subseq r` 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 `subseq r` 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 `(INF n. (u o r) n) \<ge> liminf u` by simp
+ then have "(u o r) \<longlonglongrightarrow> liminf u" using `(u o r) \<longlonglongrightarrow> (INF n. (u o r) n)` by simp
+ then show ?thesis using `subseq r` by auto
+qed
+
+
+subsection {*Extended-Real.thy*}
+
+text{* The proof of this one is copied from \verb+ereal_add_mono+. *}
+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 {* The next ones are analogues of \verb+mult_mono+ and \verb+mult_mono'+ in ereal.*}
+
+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 `c < d` 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 {*Continuity of addition*}
+
+text {*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. *}
+
+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{* 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.*}
+
+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 {* The next lemma says that the addition is continuous on ereal, except at
+the pairs $(-\infty, \infty)$ and $(\infty, -\infty)$. *}
+
+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 {*Continuity of multiplication*}
+
+text {* 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.*}
+
+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 `a < l` 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 `ereal a > 0` 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 `0 < ereal a`] by auto
+ moreover have "ereal a * ereal(M/a) = M" using `ereal a > 0` 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 `l>0` `m>0` 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 {*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.*}
+
+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 `l \<noteq> 0`] sgn_squared_ereal[OF `m \<noteq> 0`] 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 `l \<noteq> 0`] sgn_squared_ereal[OF `m \<noteq> 0`] 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 {*Continuity of division*}
+
+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 `e>0` using less_le_trans not_le by fastforce
+ then have "1/z \<ge> 0" by auto
+ moreover have "1/z < e" using `e>0` `z>1/e`
+ 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 `a>e` by (metis (mono_tags, lifting) eventually_mono less_trans)
+ qed
+qed
+
+text {* The next lemma deserves to exist by itself, as it is so common and useful. *}
+
+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 {*Further limits*}
+
+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 `\<not>(N > C)` by auto
+ then show False using `u N \<ge> n` `n \<ge> M` 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 {* the next one is copied from \verb+tendsto_sum+. *}
+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 {*Limsups and liminfs*}
+
+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 `C = ereal(real_of_ereal C)` 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 `C = ereal(real_of_ereal C)` 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 {* 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+.*}
+
+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: "subseq r" "(w o r) \<longlonglongrightarrow> limsup w" using limsup_subseq_lim by auto
+ obtain s where s: "subseq s" "(u o r o s) \<longlonglongrightarrow> limsup (u o r)" using limsup_subseq_lim by auto
+ obtain t where t: "subseq 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 "subseq a" using r s t by (simp add: a_def subseq_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 `limsup u < \<infinity>` by auto
+ have "limsup (v o r o s) \<le> limsup v" by (simp add: comp_assoc limsup_subseq_mono r(1) s(1) subseq_o)
+ then have b: "limsup (v o r o s) \<noteq> \<infinity>" using `limsup v < \<infinity>` 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 `limsup (u o r) \<le> limsup u` `limsup (v o r o s) \<le> limsup v` ereal_add_mono by simp
+ then show ?thesis unfolding w_def by simp
+qed
+
+text {* 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.*}
+
+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: "subseq r" "(w o r) \<longlonglongrightarrow> liminf w" using liminf_subseq_lim by auto
+ obtain s where s: "subseq s" "(u o r o s) \<longlonglongrightarrow> liminf (u o r)" using liminf_subseq_lim by auto
+ obtain t where t: "subseq 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 "subseq a" using r s t by (simp add: a_def subseq_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 `liminf u > -\<infinity>` by auto
+ have "liminf (v o r o s) \<ge> liminf v" by (simp add: comp_assoc liminf_subseq_mono r(1) s(1) subseq_o)
+ then have b: "liminf (v o r o s) \<noteq> -\<infinity>" using `liminf v > -\<infinity>` 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 `liminf (u o r) \<ge> liminf u` `liminf (v o r o s) \<ge> liminf v` 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 `limsup u = a` 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 `limsup (\<lambda>n. -u n) = -a` 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: "subseq 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: "subseq 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: "subseq 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: "subseq 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 `liminf u = a` 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 `liminf (\<lambda>n. -u n) = -a` 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: "subseq r" "(u o r) \<longlonglongrightarrow> liminf u" using liminf_subseq_lim by auto
+ obtain s where s: "subseq s" "(w o r o s) \<longlonglongrightarrow> liminf (w o r)" using liminf_subseq_lim by auto
+ obtain t where t: "subseq 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 "subseq a" using r s t by (simp add: a_def subseq_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) subseq_o)
+ then have b: "limsup (v o r o s) < \<infinity>" using `limsup v < \<infinity>` 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 `liminf u < \<infinity>` 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 `liminf (w o r) \<ge> liminf w` `limsup (v o r o s) \<le> limsup v`
+ 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
*)
@@ -755,4 +2088,121 @@
then show ?thesis using * by auto
qed
+text {* The next lemma shows that $L^1$ convergence of a sequence of functions follows from almost
+everywhere convergence and the weaker condition of the convergence of the integrated norms (or even
+just the nontrivial inequality about them). Useful in a lot of contexts! This statement (or its
+variations) are known as Scheffe lemma.
+
+The formalization is more painful as one should jump back and forth between reals and ereals and justify
+all the time positivity or integrability (thankfully, measurability is handled more or less automatically).*}
+
+lemma Scheffe_lemma1:
+ assumes "\<And>n. integrable M (F n)" "integrable M f"
+ "AE x in M. (\<lambda>n. F n x) \<longlonglongrightarrow> f x"
+ "limsup (\<lambda>n. \<integral>\<^sup>+ x. norm(F n x) \<partial>M) \<le> (\<integral>\<^sup>+ x. norm(f x) \<partial>M)"
+ shows "(\<lambda>n. \<integral>\<^sup>+ x. norm(F n x - f x) \<partial>M) \<longlonglongrightarrow> 0"
+proof -
+ have [measurable]: "\<And>n. F n \<in> borel_measurable M" "f \<in> borel_measurable M"
+ using assms(1) assms(2) by simp_all
+ define G where "G = (\<lambda>n x. norm(f x) + norm(F n x) - norm(F n x - f x))"
+ have [measurable]: "\<And>n. G n \<in> borel_measurable M" unfolding G_def by simp
+ have G_pos[simp]: "\<And>n x. G n x \<ge> 0"
+ unfolding G_def by (metis ge_iff_diff_ge_0 norm_minus_commute norm_triangle_ineq4)
+ have finint: "(\<integral>\<^sup>+ x. norm(f x) \<partial>M) \<noteq> \<infinity>"
+ using has_bochner_integral_implies_finite_norm[OF has_bochner_integral_integrable[OF \<open>integrable M f\<close>]]
+ by simp
+ then have fin2: "2 * (\<integral>\<^sup>+ x. norm(f x) \<partial>M) \<noteq> \<infinity>"
+ by (auto simp: ennreal_mult_eq_top_iff)
+
+ {
+ fix x assume *: "(\<lambda>n. F n x) \<longlonglongrightarrow> f x"
+ then have "(\<lambda>n. norm(F n x)) \<longlonglongrightarrow> norm(f x)" using tendsto_norm by blast
+ moreover have "(\<lambda>n. norm(F n x - f x)) \<longlonglongrightarrow> 0" using * Lim_null tendsto_norm_zero_iff by fastforce
+ ultimately have a: "(\<lambda>n. norm(F n x) - norm(F n x - f x)) \<longlonglongrightarrow> norm(f x)" using tendsto_diff by fastforce
+ have "(\<lambda>n. norm(f x) + (norm(F n x) - norm(F n x - f x))) \<longlonglongrightarrow> norm(f x) + norm(f x)"
+ by (rule tendsto_add) (auto simp add: a)
+ moreover have "\<And>n. G n x = norm(f x) + (norm(F n x) - norm(F n x - f x))" unfolding G_def by simp
+ ultimately have "(\<lambda>n. G n x) \<longlonglongrightarrow> 2 * norm(f x)" by simp
+ then have "(\<lambda>n. ennreal(G n x)) \<longlonglongrightarrow> ennreal(2 * norm(f x))" by simp
+ then have "liminf (\<lambda>n. ennreal(G n x)) = ennreal(2 * norm(f x))"
+ using sequentially_bot tendsto_iff_Liminf_eq_Limsup by blast
+ }
+ then have "AE x in M. liminf (\<lambda>n. ennreal(G n x)) = ennreal(2 * norm(f x))" using assms(3) by auto
+ then have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. ennreal (G n x)) \<partial>M) = (\<integral>\<^sup>+ x. 2 * ennreal(norm(f x)) \<partial>M)"
+ by (simp add: nn_integral_cong_AE ennreal_mult)
+ also have "... = 2 * (\<integral>\<^sup>+ x. norm(f x) \<partial>M)" by (rule nn_integral_cmult) auto
+ finally have int_liminf: "(\<integral>\<^sup>+ x. liminf (\<lambda>n. ennreal (G n x)) \<partial>M) = 2 * (\<integral>\<^sup>+ x. norm(f x) \<partial>M)"
+ by simp
+
+ have "(\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M) = (\<integral>\<^sup>+x. norm(f x) \<partial>M) + (\<integral>\<^sup>+x. norm(F n x) \<partial>M)" for n
+ by (rule nn_integral_add) (auto simp add: assms)
+ then have "limsup (\<lambda>n. (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M)) =
+ limsup (\<lambda>n. (\<integral>\<^sup>+x. norm(f x) \<partial>M) + (\<integral>\<^sup>+x. norm(F n x) \<partial>M))"
+ by simp
+ also have "... = (\<integral>\<^sup>+x. norm(f x) \<partial>M) + limsup (\<lambda>n. (\<integral>\<^sup>+x. norm(F n x) \<partial>M))"
+ by (rule Limsup_const_add, auto simp add: finint)
+ also have "... \<le> (\<integral>\<^sup>+x. norm(f x) \<partial>M) + (\<integral>\<^sup>+x. norm(f x) \<partial>M)"
+ using assms(4) by (simp add: add_left_mono)
+ also have "... = 2 * (\<integral>\<^sup>+x. norm(f x) \<partial>M)"
+ unfolding one_add_one[symmetric] distrib_right by simp
+ ultimately have a: "limsup (\<lambda>n. (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M)) \<le>
+ 2 * (\<integral>\<^sup>+x. norm(f x) \<partial>M)" by simp
+
+ have le: "ennreal (norm (F n x - f x)) \<le> ennreal (norm (f x)) + ennreal (norm (F n x))" for n x
+ by (simp add: norm_minus_commute norm_triangle_ineq4 ennreal_plus[symmetric] ennreal_minus del: ennreal_plus)
+ then have le2: "(\<integral>\<^sup>+ x. ennreal (norm (F n x - f x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) + ennreal (norm (F n x)) \<partial>M)" for n
+ by (rule nn_integral_mono)
+
+ have "2 * (\<integral>\<^sup>+ x. norm(f x) \<partial>M) = (\<integral>\<^sup>+ x. liminf (\<lambda>n. ennreal (G n x)) \<partial>M)"
+ by (simp add: int_liminf)
+ also have "\<dots> \<le> liminf (\<lambda>n. (\<integral>\<^sup>+x. G n x \<partial>M))"
+ by (rule nn_integral_liminf) auto
+ also have "liminf (\<lambda>n. (\<integral>\<^sup>+x. G n x \<partial>M)) =
+ liminf (\<lambda>n. (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M) - (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M))"
+ proof (intro arg_cong[where f=liminf] ext)
+ fix n
+ have "\<And>x. ennreal(G n x) = ennreal(norm(f x)) + ennreal(norm(F n x)) - ennreal(norm(F n x - f x))"
+ unfolding G_def by (simp add: ennreal_plus[symmetric] ennreal_minus del: ennreal_plus)
+ moreover have "(\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) - ennreal(norm(F n x - f x)) \<partial>M)
+ = (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M) - (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M)"
+ proof (rule nn_integral_diff)
+ from le show "AE x in M. ennreal (norm (F n x - f x)) \<le> ennreal (norm (f x)) + ennreal (norm (F n x))"
+ by simp
+ from le2 have "(\<integral>\<^sup>+x. ennreal (norm (F n x - f x)) \<partial>M) < \<infinity>" using assms(1) assms(2)
+ by (metis has_bochner_integral_implies_finite_norm integrable.simps Bochner_Integration.integrable_diff)
+ then show "(\<integral>\<^sup>+x. ennreal (norm (F n x - f x)) \<partial>M) \<noteq> \<infinity>" by simp
+ qed (auto simp add: assms)
+ ultimately show "(\<integral>\<^sup>+x. G n x \<partial>M) = (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M) - (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M)"
+ by simp
+ qed
+ finally have "2 * (\<integral>\<^sup>+ x. norm(f x) \<partial>M) + limsup (\<lambda>n. (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M)) \<le>
+ liminf (\<lambda>n. (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M) - (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M)) +
+ limsup (\<lambda>n. (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M))"
+ by (intro add_mono) auto
+ also have "\<dots> \<le> (limsup (\<lambda>n. \<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M) - limsup (\<lambda>n. \<integral>\<^sup>+x. norm (F n x - f x) \<partial>M)) +
+ limsup (\<lambda>n. (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M))"
+ by (intro add_mono liminf_minus_ennreal le2) auto
+ also have "\<dots> = limsup (\<lambda>n. (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M))"
+ by (intro diff_add_cancel_ennreal Limsup_mono always_eventually allI le2)
+ also have "\<dots> \<le> 2 * (\<integral>\<^sup>+x. norm(f x) \<partial>M)"
+ by fact
+ finally have "limsup (\<lambda>n. (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M)) = 0"
+ using fin2 by simp
+ then show ?thesis
+ by (rule tendsto_0_if_Limsup_eq_0_ennreal)
+qed
+
+lemma Scheffe_lemma2:
+ fixes F::"nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes "\<And> n::nat. F n \<in> borel_measurable M" "integrable M f"
+ "AE x in M. (\<lambda>n. F n x) \<longlonglongrightarrow> f x"
+ "\<And>n. (\<integral>\<^sup>+ x. norm(F n x) \<partial>M) \<le> (\<integral>\<^sup>+ x. norm(f x) \<partial>M)"
+ shows "(\<lambda>n. \<integral>\<^sup>+ x. norm(F n x - f x) \<partial>M) \<longlonglongrightarrow> 0"
+proof (rule Scheffe_lemma1)
+ fix n::nat
+ have "(\<integral>\<^sup>+ x. norm(f x) \<partial>M) < \<infinity>" using assms(2) by (metis has_bochner_integral_implies_finite_norm integrable.cases)
+ then have "(\<integral>\<^sup>+ x. norm(F n x) \<partial>M) < \<infinity>" using assms(4)[of n] by auto
+ then show "integrable M (F n)" by (subst integrable_iff_bounded, simp add: assms(1)[of n])
+qed (auto simp add: assms Limsup_bounded)
+
end
--- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Thu Oct 13 18:36:06 2016 +0200
+++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Tue Oct 18 12:01:54 2016 +0200
@@ -459,6 +459,99 @@
qed
+lemma countable_separating_set_linorder1:
+ shows "\<exists>B::('a::{linorder_topology, second_countable_topology} set). countable B \<and> (\<forall>x y. x < y \<longrightarrow> (\<exists>b \<in> B. x < b \<and> b \<le> y))"
+proof -
+ obtain A::"'a set set" where "countable A" "topological_basis A" using ex_countable_basis by auto
+ define B1 where "B1 = {(LEAST x. x \<in> U)| U. U \<in> A}"
+ then have "countable B1" using `countable A` by (simp add: Setcompr_eq_image)
+ define B2 where "B2 = {(SOME x. x \<in> U)| U. U \<in> A}"
+ then have "countable B2" using `countable A` by (simp add: Setcompr_eq_image)
+ have "\<exists>b \<in> B1 \<union> B2. x < b \<and> b \<le> y" if "x < y" for x y
+ proof (cases)
+ assume "\<exists>z. x < z \<and> z < y"
+ then obtain z where z: "x < z \<and> z < y" by auto
+ define U where "U = {x<..<y}"
+ then have "open U" by simp
+ moreover have "z \<in> U" using z U_def by simp
+ ultimately obtain V where "V \<in> A" "z \<in> V" "V \<subseteq> U" using topological_basisE[OF `topological_basis A`] by auto
+ define w where "w = (SOME x. x \<in> V)"
+ then have "w \<in> V" using `z \<in> V` by (metis someI2)
+ then have "x < w \<and> w \<le> y" using `w \<in> V` `V \<subseteq> U` U_def by fastforce
+ moreover have "w \<in> B1 \<union> B2" using w_def B2_def `V \<in> A` by auto
+ ultimately show ?thesis by auto
+ next
+ assume "\<not>(\<exists>z. x < z \<and> z < y)"
+ then have *: "\<And>z. z > x \<Longrightarrow> z \<ge> y" by auto
+ define U where "U = {x<..}"
+ then have "open U" by simp
+ moreover have "y \<in> U" using `x < y` U_def by simp
+ ultimately obtain "V" where "V \<in> A" "y \<in> V" "V \<subseteq> U" using topological_basisE[OF `topological_basis A`] by auto
+ have "U = {y..}" unfolding U_def using * `x < y` by auto
+ then have "V \<subseteq> {y..}" using `V \<subseteq> U` by simp
+ then have "(LEAST w. w \<in> V) = y" using `y \<in> V` by (meson Least_equality atLeast_iff subsetCE)
+ then have "y \<in> B1 \<union> B2" using `V \<in> A` B1_def by auto
+ moreover have "x < y \<and> y \<le> y" using `x < y` by simp
+ ultimately show ?thesis by auto
+ qed
+ moreover have "countable (B1 \<union> B2)" using `countable B1` `countable B2` by simp
+ ultimately show ?thesis by auto
+qed
+
+lemma countable_separating_set_linorder2:
+ shows "\<exists>B::('a::{linorder_topology, second_countable_topology} set). countable B \<and> (\<forall>x y. x < y \<longrightarrow> (\<exists>b \<in> B. x \<le> b \<and> b < y))"
+proof -
+ obtain A::"'a set set" where "countable A" "topological_basis A" using ex_countable_basis by auto
+ define B1 where "B1 = {(GREATEST x. x \<in> U) | U. U \<in> A}"
+ then have "countable B1" using `countable A` by (simp add: Setcompr_eq_image)
+ define B2 where "B2 = {(SOME x. x \<in> U)| U. U \<in> A}"
+ then have "countable B2" using `countable A` by (simp add: Setcompr_eq_image)
+ have "\<exists>b \<in> B1 \<union> B2. x \<le> b \<and> b < y" if "x < y" for x y
+ proof (cases)
+ assume "\<exists>z. x < z \<and> z < y"
+ then obtain z where z: "x < z \<and> z < y" by auto
+ define U where "U = {x<..<y}"
+ then have "open U" by simp
+ moreover have "z \<in> U" using z U_def by simp
+ ultimately obtain "V" where "V \<in> A" "z \<in> V" "V \<subseteq> U" using topological_basisE[OF `topological_basis A`] by auto
+ define w where "w = (SOME x. x \<in> V)"
+ then have "w \<in> V" using `z \<in> V` by (metis someI2)
+ then have "x \<le> w \<and> w < y" using `w \<in> V` `V \<subseteq> U` U_def by fastforce
+ moreover have "w \<in> B1 \<union> B2" using w_def B2_def `V \<in> A` by auto
+ ultimately show ?thesis by auto
+ next
+ assume "\<not>(\<exists>z. x < z \<and> z < y)"
+ then have *: "\<And>z. z < y \<Longrightarrow> z \<le> x" using leI by blast
+ define U where "U = {..<y}"
+ then have "open U" by simp
+ moreover have "x \<in> U" using `x < y` U_def by simp
+ ultimately obtain "V" where "V \<in> A" "x \<in> V" "V \<subseteq> U" using topological_basisE[OF `topological_basis A`] by auto
+ have "U = {..x}" unfolding U_def using * `x < y` by auto
+ then have "V \<subseteq> {..x}" using `V \<subseteq> U` by simp
+ then have "(GREATEST x. x \<in> V) = x" using `x \<in> V` by (meson Greatest_equality atMost_iff subsetCE)
+ then have "x \<in> B1 \<union> B2" using `V \<in> A` B1_def by auto
+ moreover have "x \<le> x \<and> x < y" using `x < y` by simp
+ ultimately show ?thesis by auto
+ qed
+ moreover have "countable (B1 \<union> B2)" using `countable B1` `countable B2` by simp
+ ultimately show ?thesis by auto
+qed
+
+lemma countable_separating_set_dense_linorder:
+ shows "\<exists>B::('a::{linorder_topology, dense_linorder, second_countable_topology} set). countable B \<and> (\<forall>x y. x < y \<longrightarrow> (\<exists>b \<in> B. x < b \<and> b < y))"
+proof -
+ obtain B::"'a set" where B: "countable B" "\<And>x y. x < y \<Longrightarrow> (\<exists>b \<in> B. x < b \<and> b \<le> y)"
+ using countable_separating_set_linorder1 by auto
+ have "\<exists>b \<in> B. x < b \<and> b < y" if "x < y" for x y
+ proof -
+ obtain z where "x < z" "z < y" using `x < y` dense by blast
+ then obtain b where "b \<in> B" "x < b \<and> b \<le> z" using B(2) by auto
+ then have "x < b \<and> b < y" using `z < y` by auto
+ then show ?thesis using `b \<in> B` by auto
+ qed
+ then show ?thesis using B(1) by auto
+qed
+
subsection \<open>Polish spaces\<close>
text \<open>Textbooks define Polish spaces as completely metrizable.
@@ -8688,7 +8781,7 @@
unfolding homeomorphic_def homeomorphism_def
by (metis equalityI image_subset_iff subsetI)
qed
-
+
lemma homeomorphicI [intro?]:
"\<lbrakk>f ` S = T; g ` T = S;
continuous_on S f; continuous_on T g;
@@ -10037,7 +10130,7 @@
apply (rule openin_Union, clarify)
apply (metis (full_types) \<open>open U\<close> cont clo openin_trans continuous_openin_preimage_gen)
done
-qed
+qed
lemma pasting_lemma_exists:
fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space"
--- a/src/HOL/Library/Extended_Nonnegative_Real.thy Thu Oct 13 18:36:06 2016 +0200
+++ b/src/HOL/Library/Extended_Nonnegative_Real.thy Tue Oct 18 12:01:54 2016 +0200
@@ -949,6 +949,9 @@
by (cases "0 \<le> x")
(auto simp: ennreal_neg ennreal_1[symmetric] simp del: ennreal_1)
+lemma one_less_ennreal[simp]: "1 < ennreal x \<longleftrightarrow> 1 < x"
+ by transfer (auto simp: max.absorb2 less_max_iff_disj)
+
lemma ennreal_plus[simp]:
"0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> ennreal (a + b) = ennreal a + ennreal b"
by (transfer fixing: a b) (auto simp: max_absorb2)
--- a/src/HOL/Library/Permutations.thy Thu Oct 13 18:36:06 2016 +0200
+++ b/src/HOL/Library/Permutations.thy Tue Oct 18 12:01:54 2016 +0200
@@ -22,6 +22,23 @@
"Fun.swap a b id x = (if x = a then b else if x = b then a else x)"
by (simp add: Fun.swap_def)
+lemma bij_inv_eq_iff: "bij p \<Longrightarrow> x = inv p y \<longleftrightarrow> p x = y"
+ using surj_f_inv_f[of p] by (auto simp add: bij_def)
+
+lemma bij_swap_comp:
+ assumes bp: "bij p"
+ shows "Fun.swap a b id \<circ> p = Fun.swap (inv p a) (inv p b) p"
+ using surj_f_inv_f[OF bij_is_surj[OF bp]]
+ by (simp add: fun_eq_iff Fun.swap_def bij_inv_eq_iff[OF bp])
+
+lemma bij_swap_ompose_bij: "bij p \<Longrightarrow> bij (Fun.swap a b id \<circ> p)"
+proof -
+ assume H: "bij p"
+ show ?thesis
+ unfolding bij_swap_comp[OF H] bij_swap_iff
+ using H .
+qed
+
subsection \<open>Basic consequences of the definition\<close>
@@ -30,7 +47,7 @@
lemma permutes_in_image: "p permutes S \<Longrightarrow> p x \<in> S \<longleftrightarrow> x \<in> S"
unfolding permutes_def by metis
-
+
lemma permutes_not_in:
assumes "f permutes S" "x \<notin> S" shows "f x = x"
using assms by (auto simp: permutes_def)
@@ -107,7 +124,7 @@
(* Next three lemmas contributed by Lukas Bulwahn *)
lemma permutes_bij_inv_into:
- fixes A :: "'a set" and B :: "'b set"
+ fixes A :: "'a set" and B :: "'b set"
assumes "p permutes A"
assumes "bij_betw f A B"
shows "(\<lambda>x. if x \<in> B then f (p (inv_into A f x)) else x) permutes B"
@@ -167,9 +184,9 @@
unfolding fun_eq_iff permutes_inv_eq[OF pS] permutes_inv_eq[OF permutes_inv[OF pS]]
by blast
-lemma permutes_invI:
+lemma permutes_invI:
assumes perm: "p permutes S"
- and inv: "\<And>x. x \<in> S \<Longrightarrow> p' (p x) = x"
+ and inv: "\<And>x. x \<in> S \<Longrightarrow> p' (p x) = x"
and outside: "\<And>x. x \<notin> S \<Longrightarrow> p' x = x"
shows "inv p = p'"
proof
@@ -177,14 +194,14 @@
proof (cases "x \<in> S")
assume [simp]: "x \<in> S"
from assms have "p' x = p' (p (inv p x))" by (simp add: permutes_inverses)
- also from permutes_inv[OF perm]
+ also from permutes_inv[OF perm]
have "\<dots> = inv p x" by (subst inv) (simp_all add: permutes_in_image)
finally show "inv p x = p' x" ..
qed (insert permutes_inv[OF perm], simp_all add: outside permutes_not_in)
qed
lemma permutes_vimage: "f permutes A \<Longrightarrow> f -` A = A"
- by (simp add: bij_vimage_eq_inv_image permutes_bij permutes_image[OF permutes_inv])
+ by (simp add: bij_vimage_eq_inv_image permutes_bij permutes_image[OF permutes_inv])
subsection \<open>The number of permutations on a finite set\<close>
@@ -329,7 +346,7 @@
lemma finite_permutations:
assumes fS: "finite S"
shows "finite {p. p permutes S}"
- using card_permutations[OF refl fS]
+ using card_permutations[OF refl fS]
by (auto intro: card_ge_0_finite)
@@ -724,23 +741,6 @@
qed
qed
-lemma bij_inv_eq_iff: "bij p \<Longrightarrow> x = inv p y \<longleftrightarrow> p x = y"
- using surj_f_inv_f[of p] by (auto simp add: bij_def)
-
-lemma bij_swap_comp:
- assumes bp: "bij p"
- shows "Fun.swap a b id \<circ> p = Fun.swap (inv p a) (inv p b) p"
- using surj_f_inv_f[OF bij_is_surj[OF bp]]
- by (simp add: fun_eq_iff Fun.swap_def bij_inv_eq_iff[OF bp])
-
-lemma bij_swap_ompose_bij: "bij p \<Longrightarrow> bij (Fun.swap a b id \<circ> p)"
-proof -
- assume H: "bij p"
- show ?thesis
- unfolding bij_swap_comp[OF H] bij_swap_iff
- using H .
-qed
-
lemma permutation_lemma:
assumes fS: "finite S"
and p: "bij p"
@@ -881,7 +881,7 @@
lemma sign_idempotent: "sign p * sign p = 1"
by (simp add: sign_def)
-
+
subsection \<open>Permuting a list\<close>
text \<open>This function permutes a list by applying a permutation to the indices.\<close>
@@ -889,7 +889,7 @@
definition permute_list :: "(nat \<Rightarrow> nat) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
"permute_list f xs = map (\<lambda>i. xs ! (f i)) [0..<length xs]"
-lemma permute_list_map:
+lemma permute_list_map:
assumes "f permutes {..<length xs}"
shows "permute_list f (map g xs) = map g (permute_list f xs)"
using permutes_in_image[OF assms] by (auto simp: permute_list_def)
@@ -897,7 +897,7 @@
lemma permute_list_nth:
assumes "f permutes {..<length xs}" "i < length xs"
shows "permute_list f xs ! i = xs ! f i"
- using permutes_in_image[OF assms(1)] assms(2)
+ using permutes_in_image[OF assms(1)] assms(2)
by (simp add: permute_list_def)
lemma permute_list_Nil [simp]: "permute_list f [] = []"
@@ -906,7 +906,7 @@
lemma length_permute_list [simp]: "length (permute_list f xs) = length xs"
by (simp add: permute_list_def)
-lemma permute_list_compose:
+lemma permute_list_compose:
assumes "g permutes {..<length xs}"
shows "permute_list (f \<circ> g) xs = permute_list g (permute_list f xs)"
using assms[THEN permutes_in_image] by (auto simp add: permute_list_def)
@@ -924,7 +924,7 @@
fix y :: 'a
from assms have [simp]: "f x < length xs \<longleftrightarrow> x < length xs" for x
using permutes_in_image[OF assms] by auto
- have "count (mset (permute_list f xs)) y =
+ have "count (mset (permute_list f xs)) y =
card ((\<lambda>i. xs ! f i) -` {y} \<inter> {..<length xs})"
by (simp add: permute_list_def mset_map count_image_mset atLeast0LessThan)
also have "(\<lambda>i. xs ! f i) -` {y} \<inter> {..<length xs} = f -` {i. i < length xs \<and> y = xs ! i}"
@@ -935,7 +935,7 @@
finally show "count (mset (permute_list f xs)) y = count (mset xs) y" by simp
qed
-lemma set_permute_list [simp]:
+lemma set_permute_list [simp]:
assumes "f permutes {..<length xs}"
shows "set (permute_list f xs) = set xs"
by (rule mset_eq_setD[OF mset_permute_list]) fact
@@ -945,7 +945,7 @@
shows "distinct (permute_list f xs) = distinct xs"
by (simp add: distinct_count_atmost_1 assms)
-lemma permute_list_zip:
+lemma permute_list_zip:
assumes "f permutes A" "A = {..<length xs}"
assumes [simp]: "length xs = length ys"
shows "permute_list f (zip xs ys) = zip (permute_list f xs) (permute_list f ys)"
@@ -961,7 +961,7 @@
finally show ?thesis .
qed
-lemma map_of_permute:
+lemma map_of_permute:
assumes "\<sigma> permutes fst ` set xs"
shows "map_of xs \<circ> \<sigma> = map_of (map (\<lambda>(x,y). (inv \<sigma> x, y)) xs)" (is "_ = map_of (map ?f _)")
proof
@@ -993,7 +993,7 @@
from this have "count (image_mset f (mset_set (insert x F))) b = Suc (card {a \<in> F. f a = f x})"
using insert.hyps by auto
also have "\<dots> = card (insert x {a \<in> F. f a = f x})"
- using insert.hyps(1,2) by simp
+ using insert.hyps(1,2) by simp
also have "card (insert x {a \<in> F. f a = f x}) = card {a \<in> insert x F. f a = b}"
using \<open>f x = b\<close> by (auto intro: arg_cong[where f="card"])
finally show ?thesis using insert by auto
@@ -1003,7 +1003,7 @@
with insert A show ?thesis by simp
qed
qed
-
+
(* Prove image_mset_eq_implies_permutes *)
lemma image_mset_eq_implies_permutes:
fixes f :: "'a \<Rightarrow> 'b"
@@ -1317,7 +1317,7 @@
subsection \<open>Constructing permutations from association lists\<close>
definition list_permutes where
- "list_permutes xs A \<longleftrightarrow> set (map fst xs) \<subseteq> A \<and> set (map snd xs) = set (map fst xs) \<and>
+ "list_permutes xs A \<longleftrightarrow> set (map fst xs) \<subseteq> A \<and> set (map snd xs) = set (map fst xs) \<and>
distinct (map fst xs) \<and> distinct (map snd xs)"
lemma list_permutesI [simp]:
@@ -1349,8 +1349,8 @@
proof (rule inj_onI)
fix x y assume xy: "x \<in> set (map fst xs)" "y \<in> set (map fst xs)"
assume eq: "map_of xs x = map_of xs y"
- from xy obtain x' y'
- where x'y': "map_of xs x = Some x'" "map_of xs y = Some y'"
+ from xy obtain x' y'
+ where x'y': "map_of xs x = Some x'" "map_of xs y = Some y'"
by (cases "map_of xs x"; cases "map_of xs y")
(simp_all add: map_of_eq_None_iff)
moreover from x'y' have *: "(x,x') \<in> set xs" "(y,y') \<in> set xs"
@@ -1398,7 +1398,7 @@
also from assms have "?f ` set (map fst xs) = (the \<circ> map_of xs) ` set (map fst xs)"
by (intro image_cong refl)
(auto simp: permutation_of_list_def map_of_eq_None_iff split: option.splits)
- also from assms have "\<dots> = set (map fst xs)"
+ also from assms have "\<dots> = set (map fst xs)"
by (subst image_map_of') (simp_all add: list_permutes_def)
finally show "bij_betw ?f (set (map fst xs)) (set (map fst xs))" .
qed (force simp: permutation_of_list_def dest!: map_of_SomeD split: option.splits)+
--- a/src/HOL/Probability/Levy.thy Thu Oct 13 18:36:06 2016 +0200
+++ b/src/HOL/Probability/Levy.thy Tue Oct 18 12:01:54 2016 +0200
@@ -8,11 +8,6 @@
imports Characteristic_Functions Helly_Selection Sinc_Integral
begin
-lemma LIM_zero_cancel:
- fixes f :: "_ \<Rightarrow> 'b::real_normed_vector"
- shows "((\<lambda>x. f x - l) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> l) F"
-unfolding tendsto_iff dist_norm by simp
-
subsection \<open>The Levy inversion theorem\<close>
(* Actually, this is not needed for us -- but it is useful for other purposes. (See Billingsley.) *)
--- a/src/HOL/Topological_Spaces.thy Thu Oct 13 18:36:06 2016 +0200
+++ b/src/HOL/Topological_Spaces.thy Tue Oct 18 12:01:54 2016 +0200
@@ -3398,4 +3398,93 @@
lemma isCont_swap[continuous_intros]: "isCont prod.swap a"
using continuous_on_eq_continuous_within continuous_on_swap by blast
+lemma open_diagonal_complement:
+ "open {(x,y) | x y. x \<noteq> (y::('a::t2_space))}"
+proof (rule topological_space_class.openI)
+ fix t assume "t \<in> {(x, y) | x y. x \<noteq> (y::'a)}"
+ then obtain x y where "t = (x,y)" "x \<noteq> y" by blast
+ then obtain U V where *: "open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
+ by (auto simp add: separation_t2)
+ define T where "T = U \<times> V"
+ have "open T" using * open_Times T_def by auto
+ moreover have "t \<in> T" unfolding T_def using `t = (x,y)` * by auto
+ moreover have "T \<subseteq> {(x, y) | x y. x \<noteq> y}" unfolding T_def using * by auto
+ ultimately show "\<exists>T. open T \<and> t \<in> T \<and> T \<subseteq> {(x, y) | x y. x \<noteq> y}" by auto
+qed
+
+lemma closed_diagonal:
+ "closed {y. \<exists> x::('a::t2_space). y = (x,x)}"
+proof -
+ have "{y. \<exists> x::'a. y = (x,x)} = UNIV - {(x,y) | x y. x \<noteq> y}" by auto
+ then show ?thesis using open_diagonal_complement closed_Diff by auto
+qed
+
+lemma open_superdiagonal:
+ "open {(x,y) | x y. x > (y::'a::{linorder_topology})}"
+proof (rule topological_space_class.openI)
+ fix t assume "t \<in> {(x, y) | x y. y < (x::'a)}"
+ then obtain x y where "t = (x, y)" "x > y" by blast
+ show "\<exists>T. open T \<and> t \<in> T \<and> T \<subseteq> {(x, y) | x y. y < x}"
+ proof (cases)
+ assume "\<exists>z. y < z \<and> z < x"
+ then obtain z where z: "y < z \<and> z < x" by blast
+ define T where "T = {z<..} \<times> {..<z}"
+ have "open T" unfolding T_def by (simp add: open_Times)
+ moreover have "t \<in> T" using T_def z `t = (x,y)` by auto
+ moreover have "T \<subseteq> {(x, y) | x y. y < x}" unfolding T_def by auto
+ ultimately show ?thesis by auto
+ next
+ assume "\<not>(\<exists>z. y < z \<and> z < x)"
+ then have *: "{x ..} = {y<..}" "{..< x} = {..y}"
+ using `x > y` apply auto using leI by blast
+ define T where "T = {x ..} \<times> {.. y}"
+ then have "T = {y<..} \<times> {..< x}" using * by simp
+ then have "open T" unfolding T_def by (simp add: open_Times)
+ moreover have "t \<in> T" using T_def `t = (x,y)` by auto
+ moreover have "T \<subseteq> {(x, y) | x y. y < x}" unfolding T_def using `x > y` by auto
+ ultimately show ?thesis by auto
+ qed
+qed
+
+lemma closed_subdiagonal:
+ "closed {(x,y) | x y. x \<le> (y::'a::{linorder_topology})}"
+proof -
+ have "{(x,y) | x y. x \<le> (y::'a)} = UNIV - {(x,y) | x y. x > (y::'a)}" by auto
+ then show ?thesis using open_superdiagonal closed_Diff by auto
+qed
+
+lemma open_subdiagonal:
+ "open {(x,y) | x y. x < (y::'a::{linorder_topology})}"
+proof (rule topological_space_class.openI)
+ fix t assume "t \<in> {(x, y) | x y. y > (x::'a)}"
+ then obtain x y where "t = (x, y)" "x < y" by blast
+ show "\<exists>T. open T \<and> t \<in> T \<and> T \<subseteq> {(x, y) | x y. y > x}"
+ proof (cases)
+ assume "\<exists>z. y > z \<and> z > x"
+ then obtain z where z: "y > z \<and> z > x" by blast
+ define T where "T = {..<z} \<times> {z<..}"
+ have "open T" unfolding T_def by (simp add: open_Times)
+ moreover have "t \<in> T" using T_def z `t = (x,y)` by auto
+ moreover have "T \<subseteq> {(x, y) |x y. y > x}" unfolding T_def by auto
+ ultimately show ?thesis by auto
+ next
+ assume "\<not>(\<exists>z. y > z \<and> z > x)"
+ then have *: "{..x} = {..<y}" "{x<..} = {y..}"
+ using `x < y` apply auto using leI by blast
+ define T where "T = {..x} \<times> {y..}"
+ then have "T = {..<y} \<times> {x<..}" using * by simp
+ then have "open T" unfolding T_def by (simp add: open_Times)
+ moreover have "t \<in> T" using T_def `t = (x,y)` by auto
+ moreover have "T \<subseteq> {(x, y) |x y. y > x}" unfolding T_def using `x < y` by auto
+ ultimately show ?thesis by auto
+ qed
+qed
+
+lemma closed_superdiagonal:
+ "closed {(x,y) | x y. x \<ge> (y::('a::{linorder_topology}))}"
+proof -
+ have "{(x,y) | x y. x \<ge> (y::'a)} = UNIV - {(x,y) | x y. x < y}" by auto
+ then show ?thesis using open_subdiagonal closed_Diff by auto
+qed
+
end