Probability: move emeasure and nn_integral from ereal to ennreal
authorhoelzl
Thu, 14 Apr 2016 15:48:11 +0200
changeset 62975 1d066f6ab25d
parent 62974 f17602cbf76a
child 62976 38906f0e4633
Probability: move emeasure and nn_integral from ereal to ennreal
NEWS
src/HOL/Library/Extended_Nonnegative_Real.thy
src/HOL/Library/Extended_Real.thy
src/HOL/Library/Liminf_Limsup.thy
src/HOL/Probability/Binary_Product_Measure.thy
src/HOL/Probability/Bochner_Integration.thy
src/HOL/Probability/Borel_Space.thy
src/HOL/Probability/Caratheodory.thy
src/HOL/Probability/Complete_Measure.thy
src/HOL/Probability/Convolution.thy
src/HOL/Probability/Distribution_Functions.thy
src/HOL/Probability/Distributions.thy
src/HOL/Probability/Embed_Measure.thy
src/HOL/Probability/Finite_Product_Measure.thy
src/HOL/Probability/Giry_Monad.thy
src/HOL/Probability/Independent_Family.thy
src/HOL/Probability/Infinite_Product_Measure.thy
src/HOL/Probability/Information.thy
src/HOL/Probability/Lebesgue_Integral_Substitution.thy
src/HOL/Probability/Lebesgue_Measure.thy
src/HOL/Probability/Levy.thy
src/HOL/Probability/Measurable.thy
src/HOL/Probability/Measure_Space.thy
src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy
src/HOL/Probability/Probability_Mass_Function.thy
src/HOL/Probability/Probability_Measure.thy
src/HOL/Probability/Projective_Family.thy
src/HOL/Probability/Projective_Limit.thy
src/HOL/Probability/Radon_Nikodym.thy
src/HOL/Probability/Regularity.thy
src/HOL/Probability/Sigma_Algebra.thy
src/HOL/Probability/Sinc_Integral.thy
src/HOL/Probability/Weak_Convergence.thy
--- a/NEWS	Thu Apr 14 12:17:44 2016 +0200
+++ b/NEWS	Thu Apr 14 15:48:11 2016 +0200
@@ -251,6 +251,10 @@
 
 * Session HOL-NSA has been renamed to HOL-Nonstandard_Analysis.
 
+* In HOL-Probability the type of emeasure and nn_integral was changed
+from ereal to ennreal:
+  emeasure :: 'a measure => 'a set => ennreal
+  nn_integral :: 'a measure => ('a => ennreal) => ennreal
 
 *** ML ***
 
--- a/src/HOL/Library/Extended_Nonnegative_Real.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Library/Extended_Nonnegative_Real.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -8,6 +8,156 @@
   imports Extended_Real Indicator_Function
 begin
 
+lemma ereal_ineq_diff_add:
+  assumes "b \<noteq> (-\<infinity>::ereal)" "a \<ge> b"
+  shows "a = b + (a-b)"
+by (metis add.commute assms(1) assms(2) ereal_eq_minus_iff ereal_minus_le_iff ereal_plus_eq_PInfty)
+
+lemma Limsup_const_add:
+  fixes c :: "'a::{complete_linorder, linorder_topology, topological_monoid_add, ordered_ab_semigroup_add}"
+  shows "F \<noteq> bot \<Longrightarrow> Limsup F (\<lambda>x. c + f x) = c + Limsup F f"
+  by (rule Limsup_compose_continuous_mono)
+     (auto intro!: monoI add_mono continuous_on_add continuous_on_id continuous_on_const)
+
+lemma Liminf_const_add:
+  fixes c :: "'a::{complete_linorder, linorder_topology, topological_monoid_add, ordered_ab_semigroup_add}"
+  shows "F \<noteq> bot \<Longrightarrow> Liminf F (\<lambda>x. c + f x) = c + Liminf F f"
+  by (rule Liminf_compose_continuous_mono)
+     (auto intro!: monoI add_mono continuous_on_add continuous_on_id continuous_on_const)
+
+lemma Liminf_add_const:
+  fixes c :: "'a::{complete_linorder, linorder_topology, topological_monoid_add, ordered_ab_semigroup_add}"
+  shows "F \<noteq> bot \<Longrightarrow> Liminf F (\<lambda>x. f x + c) = Liminf F f + c"
+  by (rule Liminf_compose_continuous_mono)
+     (auto intro!: monoI add_mono continuous_on_add continuous_on_id continuous_on_const)
+
+lemma sums_offset:
+  fixes f g :: "nat \<Rightarrow> 'a :: {t2_space, topological_comm_monoid_add}"
+  assumes "(\<lambda>n. f (n + i)) sums l" shows "f sums (l + (\<Sum>j<i. f j))"
+proof  -
+  have "(\<lambda>k. (\<Sum>n<k. f (n + i)) + (\<Sum>j<i. f j)) \<longlonglongrightarrow> l + (\<Sum>j<i. f j)"
+    using assms by (auto intro!: tendsto_add simp: sums_def)
+  moreover
+  { fix k :: nat
+    have "(\<Sum>j<k + i. f j) = (\<Sum>j=i..<k + i. f j) + (\<Sum>j=0..<i. f j)"
+      by (subst setsum.union_disjoint[symmetric]) (auto intro!: setsum.cong)
+    also have "(\<Sum>j=i..<k + i. f j) = (\<Sum>j\<in>(\<lambda>n. n + i)`{0..<k}. f j)"
+      unfolding image_add_atLeastLessThan by simp
+    finally have "(\<Sum>j<k + i. f j) = (\<Sum>n<k. f (n + i)) + (\<Sum>j<i. f j)"
+      by (auto simp: inj_on_def atLeast0LessThan setsum.reindex) }
+  ultimately have "(\<lambda>k. (\<Sum>n<k + i. f n)) \<longlonglongrightarrow> l + (\<Sum>j<i. f j)"
+    by simp
+  then show ?thesis
+    unfolding sums_def by (rule LIMSEQ_offset)
+qed
+
+lemma suminf_offset:
+  fixes f g :: "nat \<Rightarrow> 'a :: {t2_space, topological_comm_monoid_add}"
+  shows "summable (\<lambda>j. f (j + i)) \<Longrightarrow> suminf f = (\<Sum>j. f (j + i)) + (\<Sum>j<i. f j)"
+  by (intro sums_unique[symmetric] sums_offset summable_sums)
+
+lemma eventually_at_left_1: "(\<And>z::real. 0 < z \<Longrightarrow> z < 1 \<Longrightarrow> P z) \<Longrightarrow> eventually P (at_left 1)"
+  by (subst eventually_at_left[of 0]) (auto intro: exI[of _ 0])
+
+lemma mult_eq_1:
+  fixes a b :: "'a :: {ordered_semiring, comm_monoid_mult}"
+  shows "0 \<le> a \<Longrightarrow> a \<le> 1 \<Longrightarrow> b \<le> 1 \<Longrightarrow> a * b = 1 \<longleftrightarrow> (a = 1 \<and> b = 1)"
+  by (metis mult.left_neutral eq_iff mult.commute mult_right_mono)
+
+lemma ereal_add_diff_cancel:
+  fixes a b :: ereal
+  shows "\<bar>b\<bar> \<noteq> \<infinity> \<Longrightarrow> (a + b) - b = a"
+  by (cases a b rule: ereal2_cases) auto
+
+lemma add_top:
+  fixes x :: "'a::{order_top, ordered_comm_monoid_add}"
+  shows "0 \<le> x \<Longrightarrow> x + top = top"
+  by (intro top_le add_increasing order_refl)
+
+lemma top_add:
+  fixes x :: "'a::{order_top, ordered_comm_monoid_add}"
+  shows "0 \<le> x \<Longrightarrow> top + x = top"
+  by (intro top_le add_increasing2 order_refl)
+
+lemma le_lfp: "mono f \<Longrightarrow> x \<le> lfp f \<Longrightarrow> f x \<le> lfp f"
+  by (subst lfp_unfold) (auto dest: monoD)
+
+lemma lfp_transfer:
+  assumes \<alpha>: "sup_continuous \<alpha>" and f: "sup_continuous f" and mg: "mono g"
+  assumes bot: "\<alpha> bot \<le> lfp g" and eq: "\<And>x. x \<le> lfp f \<Longrightarrow> \<alpha> (f x) = g (\<alpha> x)"
+  shows "\<alpha> (lfp f) = lfp g"
+proof (rule antisym)
+  note mf = sup_continuous_mono[OF f]
+  have f_le_lfp: "(f ^^ i) bot \<le> lfp f" for i
+    by (induction i) (auto intro: le_lfp mf)
+
+  have "\<alpha> ((f ^^ i) bot) \<le> lfp g" for i
+    by (induction i) (auto simp: bot eq f_le_lfp intro!: le_lfp mg)
+  then show "\<alpha> (lfp f) \<le> lfp g"
+    unfolding sup_continuous_lfp[OF f]
+    by (subst \<alpha>[THEN sup_continuousD])
+       (auto intro!: mono_funpow sup_continuous_mono[OF f] SUP_least)
+
+  show "lfp g \<le> \<alpha> (lfp f)"
+    by (rule lfp_lowerbound) (simp add: eq[symmetric] lfp_unfold[OF mf, symmetric])
+qed
+
+lemma sup_continuous_applyD: "sup_continuous f \<Longrightarrow> sup_continuous (\<lambda>x. f x h)"
+  using sup_continuous_apply[THEN sup_continuous_compose] .
+
+lemma sup_continuous_SUP[order_continuous_intros]:
+  fixes M :: "_ \<Rightarrow> _ \<Rightarrow> 'a::complete_lattice"
+  assumes M: "\<And>i. i \<in> I \<Longrightarrow> sup_continuous (M i)"
+  shows  "sup_continuous (SUP i:I. M i)"
+  unfolding sup_continuous_def by (auto simp add: sup_continuousD[OF M] intro: SUP_commute)
+
+lemma sup_continuous_apply_SUP[order_continuous_intros]:
+  fixes M :: "_ \<Rightarrow> _ \<Rightarrow> 'a::complete_lattice"
+  shows "(\<And>i. i \<in> I \<Longrightarrow> sup_continuous (M i)) \<Longrightarrow> sup_continuous (\<lambda>x. SUP i:I. M i x)"
+  unfolding SUP_apply[symmetric] by (rule sup_continuous_SUP)
+
+lemma sup_continuous_lfp'[order_continuous_intros]:
+  assumes 1: "sup_continuous f"
+  assumes 2: "\<And>g. sup_continuous g \<Longrightarrow> sup_continuous (f g)"
+  shows "sup_continuous (lfp f)"
+proof -
+  have "sup_continuous ((f ^^ i) bot)" for i
+  proof (induction i)
+    case (Suc i) then show ?case
+      by (auto intro!: 2)
+  qed (simp add: bot_fun_def sup_continuous_const)
+  then show ?thesis
+    unfolding sup_continuous_lfp[OF 1] by (intro order_continuous_intros)
+qed
+
+lemma sup_continuous_lfp''[order_continuous_intros]:
+  assumes 1: "\<And>s. sup_continuous (f s)"
+  assumes 2: "\<And>g. sup_continuous g \<Longrightarrow> sup_continuous (\<lambda>s. f s (g s))"
+  shows "sup_continuous (\<lambda>x. lfp (f x))"
+proof -
+  have "sup_continuous (\<lambda>x. (f x ^^ i) bot)" for i
+  proof (induction i)
+    case (Suc i) then show ?case
+      by (auto intro!: 2)
+  qed (simp add: bot_fun_def sup_continuous_const)
+  then show ?thesis
+    unfolding sup_continuous_lfp[OF 1] by (intro order_continuous_intros)
+qed
+
+lemma mono_INF_fun:
+    "(\<And>x y. mono (F x y)) \<Longrightarrow> mono (\<lambda>z x. INF y : X x. F x y z :: 'a :: complete_lattice)"
+  by (auto intro!: INF_mono[OF bexI] simp: le_fun_def mono_def)
+
+lemma continuous_on_max:
+  fixes f g :: "'a::topological_space \<Rightarrow> 'b::linorder_topology"
+  shows "continuous_on A f \<Longrightarrow> continuous_on A g \<Longrightarrow> continuous_on A (\<lambda>x. max (f x) (g x))"
+  by (auto simp: continuous_on_def intro!: tendsto_max)
+
+lemma continuous_on_cmult_ereal:
+  "\<bar>c::ereal\<bar> \<noteq> \<infinity> \<Longrightarrow> continuous_on A f \<Longrightarrow> continuous_on A (\<lambda>x. c * f x)"
+  using tendsto_cmult_ereal[of c f "f x" "at x within A" for x]
+  by (auto simp: continuous_on_def simp del: tendsto_cmult_ereal)
+
 context linordered_nonzero_semiring
 begin
 
@@ -19,6 +169,18 @@
 
 end
 
+lemma real_of_nat_Sup:
+  assumes "A \<noteq> {}" "bdd_above A"
+  shows "of_nat (Sup A) = (SUP a:A. of_nat a :: real)"
+proof (intro antisym)
+  show "(SUP a:A. of_nat a::real) \<le> of_nat (Sup A)"
+    using assms by (intro cSUP_least of_nat_mono) (auto intro: cSup_upper)
+  have "Sup A \<in> A"
+    unfolding Sup_nat_def using assms by (intro Max_in) (auto simp: bdd_above_nat)
+  then show "of_nat (Sup A) \<le> (SUP a:A. of_nat a::real)"
+    by (intro cSUP_upper bdd_above_image_mono assms) (auto simp: mono_def)
+qed
+
 lemma of_nat_less[simp]:
   "i < j \<Longrightarrow> of_nat i < (of_nat j::'a::{linordered_nonzero_semiring, semiring_char_0})"
   by (auto simp: less_le)
@@ -58,22 +220,30 @@
   shows "summable f \<Longrightarrow> finite I \<Longrightarrow> \<forall>m\<in>- I. 0 \<le> f m \<Longrightarrow> setsum f I \<le> suminf f"
   by (rule sums_le[OF _ sums_If_finite_set summable_sums]) auto
 
+subsection \<open>Defining the extended non-negative reals\<close>
+
+text \<open>Basic definitions and type class setup\<close>
+
 typedef ennreal = "{x :: ereal. 0 \<le> x}"
   morphisms enn2ereal e2ennreal'
   by auto
 
 definition "e2ennreal x = e2ennreal' (max 0 x)"
 
+lemma enn2ereal_range: "e2ennreal ` {0..} = UNIV"
+proof -
+  have "\<exists>y\<ge>0. x = e2ennreal y" for x
+    by (cases x) (auto simp: e2ennreal_def max_absorb2)
+  then show ?thesis
+    by (auto simp: image_iff Bex_def)
+qed
+
 lemma type_definition_ennreal': "type_definition enn2ereal e2ennreal {x. 0 \<le> x}"
   using type_definition_ennreal
   by (auto simp: type_definition_def e2ennreal_def max_absorb2)
 
 setup_lifting type_definition_ennreal'
 
-lift_definition ennreal :: "real \<Rightarrow> ennreal" is "sup 0 \<circ> ereal"
-  by simp
-
-declare [[coercion ennreal]]
 declare [[coercion e2ennreal]]
 
 instantiation ennreal :: complete_linorder
@@ -99,13 +269,11 @@
 
 end
 
-lemma ennreal_cases:
-  fixes x :: ennreal
-  obtains (real) r :: real where "0 \<le> r" "x = ennreal r" | (top) "x = top"
-  apply transfer
-  subgoal for x thesis
-    by (cases x) (auto simp: max.absorb2 top_ereal_def)
-  done
+lemma pcr_ennreal_enn2ereal[simp]: "pcr_ennreal (enn2ereal x) x"
+  by (simp add: ennreal.pcr_cr_eq cr_ennreal_def)
+
+lemma rel_fun_eq_pcr_ennreal: "rel_fun op = pcr_ennreal f g \<longleftrightarrow> f = enn2ereal \<circ> g"
+  by (auto simp: rel_fun_def ennreal.pcr_cr_eq cr_ennreal_def)
 
 instantiation ennreal :: infinity
 begin
@@ -156,7 +324,7 @@
 
 end
 
-lemma ennreal_zero_less_one: "0 < (1::ennreal)"
+lemma ennreal_zero_less_one: "0 < (1::ennreal)" -- \<open>TODO: remove \<close>
   by transfer auto
 
 instance ennreal :: dioid
@@ -173,6 +341,12 @@
 instance ennreal :: linordered_nonzero_semiring
   proof qed (transfer; simp)
 
+instance ennreal :: strict_ordered_ab_semigroup_add
+proof
+  fix a b c d :: ennreal show "a < b \<Longrightarrow> c < d \<Longrightarrow> a + c < b + d"
+    by transfer (auto intro!: ereal_add_strict_mono)
+qed
+
 declare [[coercion "of_nat :: nat \<Rightarrow> ennreal"]]
 
 lemma e2ennreal_neg: "x \<le> 0 \<Longrightarrow> e2ennreal x = 0"
@@ -182,6 +356,58 @@
   by (cases "0 \<le> x" "0 \<le> y" rule: bool.exhaust[case_product bool.exhaust])
      (auto simp: e2ennreal_neg less_eq_ennreal.abs_eq eq_onp_def)
 
+lemma enn2ereal_nonneg[simp]: "0 \<le> enn2ereal x"
+  using ennreal.enn2ereal[of x] by simp
+
+lemma ereal_ennreal_cases:
+  obtains b where "0 \<le> a" "a = enn2ereal b" | "a < 0"
+  using e2ennreal'_inverse[of a, symmetric] by (cases "0 \<le> a") (auto intro: enn2ereal_nonneg)
+
+lemma rel_fun_liminf[transfer_rule]: "rel_fun (rel_fun op = pcr_ennreal) pcr_ennreal liminf liminf"
+proof -
+  have "rel_fun (rel_fun op = pcr_ennreal) pcr_ennreal (\<lambda>x. sup 0 (liminf x)) liminf"
+    unfolding liminf_SUP_INF[abs_def] by (transfer_prover_start, transfer_step+; simp)
+  then show ?thesis
+    apply (subst (asm) (2) rel_fun_def)
+    apply (subst (2) rel_fun_def)
+    apply (auto simp: comp_def max.absorb2 Liminf_bounded rel_fun_eq_pcr_ennreal)
+    done
+qed
+
+lemma rel_fun_limsup[transfer_rule]: "rel_fun (rel_fun op = pcr_ennreal) pcr_ennreal limsup limsup"
+proof -
+  have "rel_fun (rel_fun op = pcr_ennreal) pcr_ennreal (\<lambda>x. INF n. sup 0 (SUP i:{n..}. x i)) limsup"
+    unfolding limsup_INF_SUP[abs_def] by (transfer_prover_start, transfer_step+; simp)
+  then show ?thesis
+    unfolding limsup_INF_SUP[abs_def]
+    apply (subst (asm) (2) rel_fun_def)
+    apply (subst (2) rel_fun_def)
+    apply (auto simp: comp_def max.absorb2 Sup_upper2 rel_fun_eq_pcr_ennreal)
+    apply (subst (asm) max.absorb2)
+    apply (rule SUP_upper2)
+    apply auto
+    done
+qed
+
+lemma setsum_enn2ereal[simp]: "(\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> (\<Sum>i\<in>I. enn2ereal (f i)) = enn2ereal (setsum f I)"
+  by (induction I rule: infinite_finite_induct) (auto simp: setsum_nonneg zero_ennreal.rep_eq plus_ennreal.rep_eq)
+
+lemma transfer_e2ennreal_setsum [transfer_rule]:
+  "rel_fun (rel_fun op = pcr_ennreal) (rel_fun op = pcr_ennreal) setsum setsum"
+  by (auto intro!: rel_funI simp: rel_fun_eq_pcr_ennreal comp_def)
+
+lemma enn2ereal_of_nat[simp]: "enn2ereal (of_nat n) = ereal n"
+  by (induction n) (auto simp: zero_ennreal.rep_eq one_ennreal.rep_eq plus_ennreal.rep_eq)
+
+lemma enn2ereal_numeral[simp]: "enn2ereal (numeral a) = numeral a"
+  apply (subst of_nat_numeral[of a, symmetric])
+  apply (subst enn2ereal_of_nat)
+  apply simp
+  done
+
+lemma transfer_numeral[transfer_rule]: "pcr_ennreal (numeral a) (numeral a)"
+  unfolding cr_ennreal_def pcr_ennreal_def by auto
+
 subsection \<open>Cancellation simprocs\<close>
 
 lemma ennreal_add_left_cancel: "a + b = a + c \<longleftrightarrow> a = (\<infinity>::ennreal) \<or> b = c"
@@ -260,130 +486,8 @@
   ("(l::ennreal) + m < n" | "(l::ennreal) < m + n") =
   \<open>fn phi => fn ctxt => fn ct => Less_Ennreal_Cancel.proc ctxt (Thm.term_of ct)\<close>
 
-instantiation ennreal :: linear_continuum_topology
-begin
 
-definition open_ennreal :: "ennreal set \<Rightarrow> bool"
-  where "(open :: ennreal set \<Rightarrow> bool) = generate_topology (range lessThan \<union> range greaterThan)"
-
-instance
-proof
-  show "\<exists>a b::ennreal. a \<noteq> b"
-    using zero_neq_one by (intro exI)
-  show "\<And>x y::ennreal. x < y \<Longrightarrow> \<exists>z>x. z < y"
-  proof transfer
-    fix x y :: ereal assume "0 \<le> x" "x < y"
-    moreover from dense[OF this(2)] guess z ..
-    ultimately show "\<exists>z\<in>Collect (op \<le> 0). x < z \<and> z < y"
-      by (intro bexI[of _ z]) auto
-  qed
-qed (rule open_ennreal_def)
-
-end
-
-lemma ennreal_rat_dense:
-  fixes x y :: ennreal
-  shows "x < y \<Longrightarrow> \<exists>r::rat. x < real_of_rat r \<and> real_of_rat r < y"
-proof transfer
-  fix x y :: ereal assume xy: "0 \<le> x" "0 \<le> y" "x < y"
-  moreover
-  from ereal_dense3[OF \<open>x < y\<close>]
-  obtain r where "x < ereal (real_of_rat r)" "ereal (real_of_rat r) < y"
-    by auto
-  moreover then have "0 \<le> r"
-    using le_less_trans[OF \<open>0 \<le> x\<close> \<open>x < ereal (real_of_rat r)\<close>] by auto
-  ultimately show "\<exists>r. x < (sup 0 \<circ> ereal) (real_of_rat r) \<and> (sup 0 \<circ> ereal) (real_of_rat r) < y"
-    by (intro exI[of _ r]) (auto simp: max_absorb2)
-qed
-
-lemma enn2ereal_range: "e2ennreal ` {0..} = UNIV"
-proof -
-  have "\<exists>y\<ge>0. x = e2ennreal y" for x
-    by (cases x) (auto simp: e2ennreal_def max_absorb2)
-  then show ?thesis
-    by (auto simp: image_iff Bex_def)
-qed
-
-lemma enn2ereal_nonneg: "0 \<le> enn2ereal x"
-  using ennreal.enn2ereal[of x] by simp
-
-lemma ereal_ennreal_cases:
-  obtains b where "0 \<le> a" "a = enn2ereal b" | "a < 0"
-  using e2ennreal'_inverse[of a, symmetric] by (cases "0 \<le> a") (auto intro: enn2ereal_nonneg)
-
-lemma enn2ereal_Iio: "enn2ereal -` {..<a} = (if 0 \<le> a then {..< e2ennreal a} else {})"
-  using enn2ereal_nonneg
-  by (cases a rule: ereal_ennreal_cases)
-     (auto simp add: vimage_def set_eq_iff ennreal.enn2ereal_inverse less_ennreal.rep_eq e2ennreal_def max_absorb2
-           intro: le_less_trans less_imp_le)
-
-lemma enn2ereal_Ioi: "enn2ereal -` {a <..} = (if 0 \<le> a then {e2ennreal a <..} else UNIV)"
-  using enn2ereal_nonneg
-  by (cases a rule: ereal_ennreal_cases)
-     (auto simp add: vimage_def set_eq_iff ennreal.enn2ereal_inverse less_ennreal.rep_eq e2ennreal_def max_absorb2
-           intro: less_le_trans)
-
-lemma continuous_on_e2ennreal: "continuous_on A e2ennreal"
-proof (rule continuous_on_subset)
-  show "continuous_on ({0..} \<union> {..0}) e2ennreal"
-  proof (rule continuous_on_closed_Un)
-    show "continuous_on {0 ..} e2ennreal"
-      by (rule continuous_onI_mono)
-         (auto simp add: less_eq_ennreal.abs_eq eq_onp_def enn2ereal_range)
-    show "continuous_on {.. 0} e2ennreal"
-      by (subst continuous_on_cong[OF refl, of _ _ "\<lambda>_. 0"])
-         (auto simp add: e2ennreal_neg continuous_on_const)
-  qed auto
-  show "A \<subseteq> {0..} \<union> {..0::ereal}"
-    by auto
-qed
-
-lemma continuous_at_e2ennreal: "continuous (at x within A) e2ennreal"
-  by (rule continuous_on_imp_continuous_within[OF continuous_on_e2ennreal, of _ UNIV]) auto
-
-lemma continuous_on_enn2ereal: "continuous_on UNIV enn2ereal"
-  by (rule continuous_on_generate_topology[OF open_generated_order])
-     (auto simp add: enn2ereal_Iio enn2ereal_Ioi)
-
-lemma continuous_at_enn2ereal: "continuous (at x within A) enn2ereal"
-  by (rule continuous_on_imp_continuous_within[OF continuous_on_enn2ereal]) auto
-
-lemma transfer_enn2ereal_continuous_on [transfer_rule]:
-  "rel_fun (op =) (rel_fun (rel_fun op = pcr_ennreal) op =) continuous_on continuous_on"
-proof -
-  have "continuous_on A f" if "continuous_on A (\<lambda>x. enn2ereal (f x))" for A and f :: "'a \<Rightarrow> ennreal"
-    using continuous_on_compose2[OF continuous_on_e2ennreal[of "{0..}"] that]
-    by (auto simp: ennreal.enn2ereal_inverse subset_eq enn2ereal_nonneg e2ennreal_def max_absorb2)
-  moreover
-  have "continuous_on A (\<lambda>x. enn2ereal (f x))" if "continuous_on A f" for A and f :: "'a \<Rightarrow> ennreal"
-    using continuous_on_compose2[OF continuous_on_enn2ereal that] by auto
-  ultimately
-  show ?thesis
-    by (auto simp add: rel_fun_def ennreal.pcr_cr_eq cr_ennreal_def)
-qed
-
-lemma continuous_on_add_ennreal[continuous_intros]:
-  fixes f g :: "'a::topological_space \<Rightarrow> ennreal"
-  shows "continuous_on A f \<Longrightarrow> continuous_on A g \<Longrightarrow> continuous_on A (\<lambda>x. f x + g x)"
-  by (transfer fixing: A) (auto intro!: tendsto_add_ereal_nonneg simp: continuous_on_def)
-
-lemma continuous_on_inverse_ennreal[continuous_intros]:
-  fixes f :: "'a::topological_space \<Rightarrow> ennreal"
-  shows "continuous_on A f \<Longrightarrow> continuous_on A (\<lambda>x. inverse (f x))"
-proof (transfer fixing: A)
-  show "pred_fun (\<lambda>_. True)  (op \<le> 0) f \<Longrightarrow> continuous_on A (\<lambda>x. inverse (f x))" if "continuous_on A f"
-    for f :: "'a \<Rightarrow> ereal"
-    using continuous_on_compose2[OF continuous_on_inverse_ereal that] by (auto simp: subset_eq)
-qed
-
-instance ennreal :: topological_comm_monoid_add
-proof
-  show "((\<lambda>x. fst x + snd x) \<longlongrightarrow> a + b) (nhds a \<times>\<^sub>F nhds b)" for a b :: ennreal
-    using continuous_on_add_ennreal[of UNIV fst snd]
-    using tendsto_at_iff_tendsto_nhds[symmetric, of "\<lambda>x::(ennreal \<times> ennreal). fst x + snd x"]
-    by (auto simp: continuous_on_eq_continuous_at)
-       (simp add: isCont_def nhds_prod[symmetric])
-qed
+subsection \<open>Order with top\<close>
 
 lemma ennreal_zero_less_top[simp]: "0 < (top::ennreal)"
   by transfer (simp add: top_ereal_def)
@@ -423,12 +527,90 @@
   shows "finite I \<Longrightarrow> (\<Sum>i\<in>I. f i) = top \<longleftrightarrow> (\<exists>i\<in>I. f i = top)"
   by (induction I rule: finite_induct) auto
 
+lemma ennreal_mult_eq_top_iff:
+  fixes a b :: ennreal
+  shows "a * b = top \<longleftrightarrow> (a = top \<and> b \<noteq> 0) \<or> (b = top \<and> a \<noteq> 0)"
+  by transfer (auto simp: top_ereal_def)
+
+lemma ennreal_top_eq_mult_iff:
+  fixes a b :: ennreal
+  shows "top = a * b \<longleftrightarrow> (a = top \<and> b \<noteq> 0) \<or> (b = top \<and> a \<noteq> 0)"
+  using ennreal_mult_eq_top_iff[of a b] by auto
+
+lemma ennreal_mult_less_top:
+  fixes a b :: ennreal
+  shows "a * b < top \<longleftrightarrow> (a = 0 \<or> b = 0 \<or> (a < top \<and> b < top))"
+  by transfer (auto simp add: top_ereal_def)
+
+lemma top_power_ennreal: "top ^ n = (if n = 0 then 1 else top :: ennreal)"
+  by (induction n) (simp_all add: ennreal_mult_eq_top_iff)
+
+lemma ennreal_setprod_eq_0[simp]:
+  fixes f :: "'a \<Rightarrow> ennreal"
+  shows "(setprod f A = 0) = (finite A \<and> (\<exists>i\<in>A. f i = 0))"
+  by (induction A rule: infinite_finite_induct) auto
+
+lemma ennreal_setprod_eq_top:
+  fixes f :: "'a \<Rightarrow> ennreal"
+  shows "(\<Prod>i\<in>I. f i) = top \<longleftrightarrow> (finite I \<and> ((\<forall>i\<in>I. f i \<noteq> 0) \<and> (\<exists>i\<in>I. f i = top)))"
+  by (induction I rule: infinite_finite_induct) (auto simp: ennreal_mult_eq_top_iff)
+
+lemma ennreal_top_mult: "top * a = (if a = 0 then 0 else top :: ennreal)"
+  by (simp add: ennreal_mult_eq_top_iff)
+
+lemma ennreal_mult_top: "a * top = (if a = 0 then 0 else top :: ennreal)"
+  by (simp add: ennreal_mult_eq_top_iff)
+
 lemma enn2ereal_eq_top_iff[simp]: "enn2ereal x = \<infinity> \<longleftrightarrow> x = top"
   by transfer (simp add: top_ereal_def)
 
-lemma ennreal_top_top: "top - top = (top::ennreal)"
+lemma enn2ereal_top: "enn2ereal top = \<infinity>"
+  by transfer (simp add: top_ereal_def)
+
+lemma e2ennreal_infty: "e2ennreal \<infinity> = top"
+  by (simp add: top_ennreal.abs_eq top_ereal_def)
+
+lemma ennreal_top_minus[simp]: "top - x = (top::ennreal)"
   by transfer (auto simp: top_ereal_def max_def)
 
+lemma minus_top_ennreal: "x - top = (if x = top then top else 0:: ennreal)"
+  apply transfer
+  subgoal for x
+    by (cases x) (auto simp: top_ereal_def max_def)
+  done
+
+lemma bot_ennreal: "bot = (0::ennreal)"
+  by transfer rule
+
+lemma ennreal_of_nat_neq_top[simp]: "of_nat i \<noteq> (top::ennreal)"
+  by (induction i) auto
+
+lemma numeral_eq_of_nat: "(numeral a::ennreal) = of_nat (numeral a)"
+  by simp
+
+lemma of_nat_less_top: "of_nat i < (top::ennreal)"
+  using less_le_trans[of "of_nat i" "of_nat (Suc i)" "top::ennreal"]
+  by simp
+
+lemma top_neq_numeral[simp]: "top \<noteq> (numeral i::ennreal)"
+  using of_nat_less_top[of "numeral i"] by simp
+
+lemma ennreal_numeral_less_top[simp]: "numeral i < (top::ennreal)"
+  using of_nat_less_top[of "numeral i"] by simp
+
+lemma ennreal_add_bot[simp]: "bot + x = (x::ennreal)"
+  by transfer simp
+
+instance ennreal :: semiring_char_0
+proof (standard, safe intro!: linorder_injI)
+  have *: "1 + of_nat k \<noteq> (0::ennreal)" for k
+    using add_pos_nonneg[OF zero_less_one, of "of_nat k :: ennreal"] by auto
+  fix x y :: nat assume "x < y" "of_nat x = (of_nat y::ennreal)" then show False
+    by (auto simp add: less_iff_Suc_add *)
+qed
+
+subsection \<open>Arithmetic\<close>
+
 lemma ennreal_minus_zero[simp]: "a - (0::ennreal) = a"
   by transfer (auto simp: max_def)
 
@@ -464,75 +646,6 @@
        (auto simp: top_ereal_def max_def split: if_splits)
   done
 
-lemma enn2ereal_ennreal[simp]: "0 \<le> x \<Longrightarrow> enn2ereal (ennreal x) = x"
-  by transfer (simp add: max_absorb2)
-
-lemma tendsto_ennrealD:
-  assumes lim: "((\<lambda>x. ennreal (f x)) \<longlongrightarrow> ennreal x) F"
-  assumes *: "\<forall>\<^sub>F x in F. 0 \<le> f x" and x: "0 \<le> x"
-  shows "(f \<longlongrightarrow> x) F"
-  using continuous_on_tendsto_compose[OF continuous_on_enn2ereal lim]
-  apply simp
-  apply (subst (asm) tendsto_cong)
-  using *
-  apply eventually_elim
-  apply (auto simp: max_absorb2 \<open>0 \<le> x\<close>)
-  done
-
-lemma tendsto_ennreal_iff[simp]:
-  "\<forall>\<^sub>F x in F. 0 \<le> f x \<Longrightarrow> 0 \<le> x \<Longrightarrow> ((\<lambda>x. ennreal (f x)) \<longlongrightarrow> ennreal x) F \<longleftrightarrow> (f \<longlongrightarrow> x) F"
-  by (auto dest: tendsto_ennrealD)
-     (auto simp: ennreal_def
-           intro!: continuous_on_tendsto_compose[OF continuous_on_e2ennreal[of UNIV]] tendsto_max)
-
-lemma ennreal_zero[simp]: "ennreal 0 = 0"
-  by (simp add: ennreal_def max.absorb1 zero_ennreal.abs_eq)
-
-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)
-
-lemma ennreal_inj[simp]:
-  "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> ennreal a = ennreal b \<longleftrightarrow> a = b"
-  by (transfer fixing: a b) (auto simp: max_absorb2)
-
-lemma ennreal_le_iff[simp]: "0 \<le> y \<Longrightarrow> ennreal x \<le> ennreal y \<longleftrightarrow> x \<le> y"
-  by (auto simp: ennreal_def zero_ereal_def less_eq_ennreal.abs_eq eq_onp_def split: split_max)
-
-lemma setsum_ennreal[simp]: "(\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> (\<Sum>i\<in>I. ennreal (f i)) = ennreal (setsum f I)"
-  by (induction I rule: infinite_finite_induct) (auto simp: setsum_nonneg)
-
-lemma sums_ennreal[simp]: "(\<And>i. 0 \<le> f i) \<Longrightarrow> 0 \<le> x \<Longrightarrow> (\<lambda>i. ennreal (f i)) sums ennreal x \<longleftrightarrow> f sums x"
-  unfolding sums_def by (simp add: always_eventually setsum_nonneg)
-
-lemma summable_suminf_not_top: "(\<And>i. 0 \<le> f i) \<Longrightarrow> (\<Sum>i. ennreal (f i)) \<noteq> top \<Longrightarrow> summable f"
-  using summable_sums[OF summableI, of "\<lambda>i. ennreal (f i)"]
-  by (cases "\<Sum>i. ennreal (f i)" rule: ennreal_cases)
-     (auto simp: summable_def)
-
-lemma suminf_ennreal[simp]:
-  "(\<And>i. 0 \<le> f i) \<Longrightarrow> (\<Sum>i. ennreal (f i)) \<noteq> top \<Longrightarrow> (\<Sum>i. ennreal (f i)) = ennreal (\<Sum>i. f i)"
-  by (rule sums_unique[symmetric]) (simp add: summable_suminf_not_top suminf_nonneg summable_sums)
-
-lemma suminf_less_top: "(\<Sum>i. f i :: ennreal) < top \<Longrightarrow> f i < top"
-  using le_less_trans[OF setsum_le_suminf[OF summableI, of "{i}" f]] by simp
-
-lemma add_top:
-  fixes x :: "'a::{order_top, ordered_comm_monoid_add}"
-  shows "0 \<le> x \<Longrightarrow> x + top = top"
-  by (intro top_le add_increasing order_refl)
-
-lemma top_add:
-  fixes x :: "'a::{order_top, ordered_comm_monoid_add}"
-  shows "0 \<le> x \<Longrightarrow> top + x = top"
-  by (intro top_le add_increasing2 order_refl)
-
-lemma enn2ereal_top: "enn2ereal top = \<infinity>"
-  by transfer (simp add: top_ereal_def)
-
-lemma e2ennreal_infty: "e2ennreal \<infinity> = top"
-  by (simp add: top_ennreal.abs_eq top_ereal_def)
-
 lemma sup_const_add_ennreal:
   fixes a b c :: "ennreal"
   shows "sup (c + a) (c + b) = c + sup a b"
@@ -546,337 +659,14 @@
     done
   done
 
-lemma bot_ennreal: "bot = (0::ennreal)"
-  by transfer rule
-
-lemma le_lfp: "mono f \<Longrightarrow> x \<le> lfp f \<Longrightarrow> f x \<le> lfp f"
-  by (subst lfp_unfold) (auto dest: monoD)
-
-lemma lfp_transfer:
-  assumes \<alpha>: "sup_continuous \<alpha>" and f: "sup_continuous f" and mg: "mono g"
-  assumes bot: "\<alpha> bot \<le> lfp g" and eq: "\<And>x. x \<le> lfp f \<Longrightarrow> \<alpha> (f x) = g (\<alpha> x)"
-  shows "\<alpha> (lfp f) = lfp g"
-proof (rule antisym)
-  note mf = sup_continuous_mono[OF f]
-  have f_le_lfp: "(f ^^ i) bot \<le> lfp f" for i
-    by (induction i) (auto intro: le_lfp mf)
-
-  have "\<alpha> ((f ^^ i) bot) \<le> lfp g" for i
-    by (induction i) (auto simp: bot eq f_le_lfp intro!: le_lfp mg)
-  then show "\<alpha> (lfp f) \<le> lfp g"
-    unfolding sup_continuous_lfp[OF f]
-    by (subst \<alpha>[THEN sup_continuousD])
-       (auto intro!: mono_funpow sup_continuous_mono[OF f] SUP_least)
-
-  show "lfp g \<le> \<alpha> (lfp f)"
-    by (rule lfp_lowerbound) (simp add: eq[symmetric] lfp_unfold[OF mf, symmetric])
-qed
-
-lemma sup_continuous_applyD: "sup_continuous f \<Longrightarrow> sup_continuous (\<lambda>x. f x h)"
-  using sup_continuous_apply[THEN sup_continuous_compose] .
-
-lemma INF_ennreal_add_const:
-  fixes f g :: "nat \<Rightarrow> ennreal"
-  shows "(INF i. f i + c) = (INF i. f i) + c"
-  using continuous_at_Inf_mono[of "\<lambda>x. x + c" "f`UNIV"]
-  using continuous_add[of "at_right (Inf (range f))", of "\<lambda>x. x" "\<lambda>x. c"]
-  by (auto simp: mono_def)
-
-lemma INF_ennreal_const_add:
-  fixes f g :: "nat \<Rightarrow> ennreal"
-  shows "(INF i. c + f i) = c + (INF i. f i)"
-  using INF_ennreal_add_const[of f c] by (simp add: ac_simps)
-
-lemma sup_continuous_e2ennreal[order_continuous_intros]:
-  assumes f: "sup_continuous f" shows "sup_continuous (\<lambda>x. e2ennreal (f x))"
-  apply (rule sup_continuous_compose[OF _ f])
-  apply (rule continuous_at_left_imp_sup_continuous)
-  apply (auto simp: mono_def e2ennreal_mono continuous_at_e2ennreal)
-  done
-
-lemma sup_continuous_enn2ereal[order_continuous_intros]:
-  assumes f: "sup_continuous f" shows "sup_continuous (\<lambda>x. enn2ereal (f x))"
-  apply (rule sup_continuous_compose[OF _ f])
-  apply (rule continuous_at_left_imp_sup_continuous)
-  apply (simp_all add: mono_def less_eq_ennreal.rep_eq continuous_at_enn2ereal)
-  done
-
-lemma ennreal_1[simp]: "ennreal 1 = 1"
-  by transfer (simp add: max_absorb2)
-
-lemma ennreal_of_nat_eq_real_of_nat: "of_nat i = ennreal (of_nat i)"
-  by (induction i) simp_all
-
-lemma ennreal_SUP_of_nat_eq_top: "(SUP x. of_nat x :: ennreal) = top"
-proof (intro antisym top_greatest le_SUP_iff[THEN iffD2] allI impI)
-  fix y :: ennreal assume "y < top"
-  then obtain r where "y = ennreal r"
-    by (cases y rule: ennreal_cases) auto
-  then show "\<exists>i\<in>UNIV. y < of_nat i"
-    using reals_Archimedean2[of "max 1 r"] zero_less_one
-    by (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_def less_ennreal.abs_eq eq_onp_def max.absorb2
-             dest!: one_less_of_natD intro: less_trans)
-qed
-
-lemma ennreal_SUP_eq_top:
-  fixes f :: "'a \<Rightarrow> ennreal"
-  assumes "\<And>n. \<exists>i\<in>I. of_nat n \<le> f i"
-  shows "(SUP i : I. f i) = top"
-proof -
-  have "(SUP x. of_nat x :: ennreal) \<le> (SUP i : I. f i)"
-    using assms by (auto intro!: SUP_least intro: SUP_upper2)
-  then show ?thesis
-    by (auto simp: ennreal_SUP_of_nat_eq_top top_unique)
-qed
-
-lemma sup_continuous_SUP[order_continuous_intros]:
-  fixes M :: "_ \<Rightarrow> _ \<Rightarrow> 'a::complete_lattice"
-  assumes M: "\<And>i. i \<in> I \<Longrightarrow> sup_continuous (M i)"
-  shows  "sup_continuous (SUP i:I. M i)"
-  unfolding sup_continuous_def by (auto simp add: sup_continuousD[OF M] intro: SUP_commute)
-
-lemma sup_continuous_apply_SUP[order_continuous_intros]:
-  fixes M :: "_ \<Rightarrow> _ \<Rightarrow> 'a::complete_lattice"
-  shows "(\<And>i. i \<in> I \<Longrightarrow> sup_continuous (M i)) \<Longrightarrow> sup_continuous (\<lambda>x. SUP i:I. M i x)"
-  unfolding SUP_apply[symmetric] by (rule sup_continuous_SUP)
-
-lemma sup_continuous_lfp'[order_continuous_intros]:
-  assumes 1: "sup_continuous f"
-  assumes 2: "\<And>g. sup_continuous g \<Longrightarrow> sup_continuous (f g)"
-  shows "sup_continuous (lfp f)"
-proof -
-  have "sup_continuous ((f ^^ i) bot)" for i
-  proof (induction i)
-    case (Suc i) then show ?case
-      by (auto intro!: 2)
-  qed (simp add: bot_fun_def sup_continuous_const)
-  then show ?thesis
-    unfolding sup_continuous_lfp[OF 1] by (intro order_continuous_intros)
-qed
-
-lemma sup_continuous_lfp''[order_continuous_intros]:
-  assumes 1: "\<And>s. sup_continuous (f s)"
-  assumes 2: "\<And>g. sup_continuous g \<Longrightarrow> sup_continuous (\<lambda>s. f s (g s))"
-  shows "sup_continuous (\<lambda>x. lfp (f x))"
-proof -
-  have "sup_continuous (\<lambda>x. (f x ^^ i) bot)" for i
-  proof (induction i)
-    case (Suc i) then show ?case
-      by (auto intro!: 2)
-  qed (simp add: bot_fun_def sup_continuous_const)
-  then show ?thesis
-    unfolding sup_continuous_lfp[OF 1] by (intro order_continuous_intros)
-qed
-
-lemma ennreal_INF_const_minus:
-  fixes f :: "'a \<Rightarrow> ennreal"
-  shows "I \<noteq> {} \<Longrightarrow> (SUP x:I. c - f x) = c - (INF x:I. f x)"
-  by (transfer fixing: I)
-     (simp add: sup_max[symmetric] SUP_sup_const1 SUP_ereal_minus_right del: sup_ereal_def)
-
 lemma ennreal_diff_add_assoc:
   fixes a b c :: ennreal
   shows "a \<le> b \<Longrightarrow> c + b - a = c + (b - a)"
   apply transfer
   subgoal for a b c
-    apply (cases a b c rule: ereal3_cases)
-    apply (auto simp: field_simps max_absorb2)
-    done
-  done
-
-lemma ennreal_top_minus[simp]:
-  fixes c :: ennreal
-  shows "top - c = top"
-  by transfer (auto simp: top_ereal_def max_absorb2)
-
-lemma le_ennreal_iff:
-  "0 \<le> r \<Longrightarrow> x \<le> ennreal r \<longleftrightarrow> (\<exists>q\<ge>0. x = ennreal q \<and> q \<le> r)"
-  apply (transfer fixing: r)
-  subgoal for x
-    by (cases x) (auto simp: max_absorb2 cong: conj_cong)
+    by (cases a b c rule: ereal3_cases) (auto simp: field_simps max_absorb2)
   done
 
-lemma ennreal_minus: "0 \<le> q \<Longrightarrow> q \<le> r \<Longrightarrow> ennreal r - ennreal q = ennreal (r - q)"
-  by transfer (simp add: max_absorb2)
-
-lemma ennreal_tendsto_const_minus:
-  fixes g :: "'a \<Rightarrow> ennreal"
-  assumes ae: "\<forall>\<^sub>F x in F. g x \<le> c"
-  assumes g: "((\<lambda>x. c - g x) \<longlongrightarrow> 0) F"
-  shows "(g \<longlongrightarrow> c) F"
-proof (cases c rule: ennreal_cases)
-  case top with tendsto_unique[OF _ g, of "top"] show ?thesis
-    by (cases "F = bot") auto
-next
-  case (real r)
-  then have "\<forall>x. \<exists>q\<ge>0. g x \<le> c \<longrightarrow> (g x = ennreal q \<and> q \<le> r)"
-    by (auto simp: le_ennreal_iff)
-  then obtain f where *: "\<And>x. g x \<le> c \<Longrightarrow> 0 \<le> f x" "\<And>x. g x \<le> c \<Longrightarrow> g x = ennreal (f x)" "\<And>x. g x \<le> c \<Longrightarrow> f x \<le> r"
-    by metis
-  from ae have ae2: "\<forall>\<^sub>F x in F. c - g x = ennreal (r - f x) \<and> f x \<le> r \<and> g x = ennreal (f x) \<and> 0 \<le> f x"
-  proof eventually_elim
-    fix x assume "g x \<le> c" with *[of x] \<open>0 \<le> r\<close> show "c - g x = ennreal (r - f x) \<and> f x \<le> r \<and> g x = ennreal (f x) \<and> 0 \<le> f x"
-      by (auto simp: real ennreal_minus)
-  qed
-  with g have "((\<lambda>x. ennreal (r - f x)) \<longlongrightarrow> ennreal 0) F"
-    by (auto simp add: tendsto_cong eventually_conj_iff)
-  with ae2 have "((\<lambda>x. r - f x) \<longlongrightarrow> 0) F"
-    by (subst (asm) tendsto_ennreal_iff) (auto elim: eventually_mono)
-  then have "(f \<longlongrightarrow> r) F"
-    by (rule Lim_transform2[OF tendsto_const])
-  with ae2 have "((\<lambda>x. ennreal (f x)) \<longlongrightarrow> ennreal r) F"
-    by (subst tendsto_ennreal_iff) (auto elim: eventually_mono simp: real)
-  with ae2 show ?thesis
-    by (auto simp: real tendsto_cong eventually_conj_iff)
-qed
-
-lemma ereal_add_diff_cancel:
-  fixes a b :: ereal
-  shows "\<bar>b\<bar> \<noteq> \<infinity> \<Longrightarrow> (a + b) - b = a"
-  by (cases a b rule: ereal2_cases) auto
-
-lemma ennreal_add_diff_cancel:
-  fixes a b :: ennreal
-  shows "b \<noteq> \<infinity> \<Longrightarrow> (a + b) - b = a"
-  unfolding infinity_ennreal_def
-  by transfer (simp add: max_absorb2 top_ereal_def ereal_add_diff_cancel)
-
-lemma ennreal_mult_eq_top_iff:
-  fixes a b :: ennreal
-  shows "a * b = top \<longleftrightarrow> (a = top \<and> b \<noteq> 0) \<or> (b = top \<and> a \<noteq> 0)"
-  by transfer (auto simp: top_ereal_def)
-
-lemma ennreal_top_eq_mult_iff:
-  fixes a b :: ennreal
-  shows "top = a * b \<longleftrightarrow> (a = top \<and> b \<noteq> 0) \<or> (b = top \<and> a \<noteq> 0)"
-  using ennreal_mult_eq_top_iff[of a b] by auto
-
-lemma ennreal_mult: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> ennreal (a * b) = ennreal a * ennreal b"
-  by transfer (simp add: max_absorb2)
-
-lemma setsum_enn2ereal[simp]: "(\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> (\<Sum>i\<in>I. enn2ereal (f i)) = enn2ereal (setsum f I)"
-  by (induction I rule: infinite_finite_induct) (auto simp: setsum_nonneg zero_ennreal.rep_eq plus_ennreal.rep_eq)
-
-lemma e2ennreal_enn2ereal[simp]: "e2ennreal (enn2ereal x) = x"
-  by (simp add: e2ennreal_def max_absorb2 enn2ereal_nonneg ennreal.enn2ereal_inverse)
-
-lemma tendsto_enn2ereal_iff[simp]: "((\<lambda>i. enn2ereal (f i)) \<longlongrightarrow> enn2ereal x) F \<longleftrightarrow> (f \<longlongrightarrow> x) F"
-  using continuous_on_enn2ereal[THEN continuous_on_tendsto_compose, of f x F]
-    continuous_on_e2ennreal[THEN continuous_on_tendsto_compose, of "\<lambda>x. enn2ereal (f x)" "enn2ereal x" F UNIV]
-  by auto
-
-lemma sums_enn2ereal[simp]: "(\<lambda>i. enn2ereal (f i)) sums enn2ereal x \<longleftrightarrow> f sums x"
-  unfolding sums_def by (simp add: always_eventually setsum_nonneg setsum_enn2ereal)
-
-lemma suminf_enn2real[simp]: "(\<Sum>i. enn2ereal (f i)) = enn2ereal (suminf f)"
-  by (rule sums_unique[symmetric]) (simp add: summable_sums)
-
-lemma pcr_ennreal_enn2ereal[simp]: "pcr_ennreal (enn2ereal x) x"
-  by (simp add: ennreal.pcr_cr_eq cr_ennreal_def)
-
-lemma rel_fun_eq_pcr_ennreal: "rel_fun op = pcr_ennreal f g \<longleftrightarrow> f = enn2ereal \<circ> g"
-  by (auto simp: rel_fun_def ennreal.pcr_cr_eq cr_ennreal_def)
-
-lemma transfer_e2ennreal_suminf [transfer_rule]: "rel_fun (rel_fun op = pcr_ennreal) pcr_ennreal suminf suminf"
-  by (auto simp: rel_funI rel_fun_eq_pcr_ennreal comp_def)
-
-lemma ennreal_suminf_cmult[simp]: "(\<Sum>i. r * f i) = r * (\<Sum>i. f i::ennreal)"
-  by transfer (auto intro!: suminf_cmult_ereal)
-
-lemma ennreal_suminf_SUP_eq_directed:
-  fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ennreal"
-  assumes *: "\<And>N i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> finite N \<Longrightarrow> \<exists>k\<in>I. \<forall>n\<in>N. f i n \<le> f k n \<and> f j n \<le> f k n"
-  shows "(\<Sum>n. SUP i:I. f i n) = (SUP i:I. \<Sum>n. f i n)"
-proof cases
-  assume "I \<noteq> {}"
-  then obtain i where "i \<in> I" by auto
-  from * show ?thesis
-    by (transfer fixing: I)
-       (auto simp: max_absorb2 SUP_upper2[OF \<open>i \<in> I\<close>] suminf_nonneg summable_ereal_pos \<open>I \<noteq> {}\<close>
-             intro!: suminf_SUP_eq_directed)
-qed (simp add: bot_ennreal)
-
-lemma ennreal_eq_zero_iff[simp]: "0 \<le> x \<Longrightarrow> ennreal x = 0 \<longleftrightarrow> x = 0"
-  by transfer (auto simp: max_absorb2)
-
-lemma ennreal_neq_top[simp]: "ennreal r \<noteq> top"
-  by transfer (simp add: top_ereal_def zero_ereal_def ereal_max[symmetric] del: ereal_max)
-
-lemma ennreal_of_nat_neq_top[simp]: "of_nat i \<noteq> (top::ennreal)"
-  by (induction i) auto
-
-lemma ennreal_suminf_neq_top: "summable f \<Longrightarrow> (\<And>i. 0 \<le> f i) \<Longrightarrow> (\<Sum>i. ennreal (f i)) \<noteq> top"
-  using sums_ennreal[of f "suminf f"]
-  by (simp add: suminf_nonneg sums_unique[symmetric] summable_sums_iff[symmetric] del: sums_ennreal)
-
-instance ennreal :: semiring_char_0
-proof (standard, safe intro!: linorder_injI)
-  have *: "1 + of_nat k \<noteq> (0::ennreal)" for k
-    using add_pos_nonneg[OF zero_less_one, of "of_nat k :: ennreal"] by auto
-  fix x y :: nat assume "x < y" "of_nat x = (of_nat y::ennreal)" then show False
-    by (auto simp add: less_iff_Suc_add *)
-qed
-
-lemma ennreal_suminf_lessD: "(\<Sum>i. f i :: ennreal) < x \<Longrightarrow> f i < x"
-  using le_less_trans[OF setsum_le_suminf[OF summableI, of "{i}" f]] by simp
-
-lemma ennreal_less_top[simp]: "ennreal x < top"
-  by transfer (simp add: top_ereal_def max_def)
-
-lemma ennreal_le_epsilon:
-  "(\<And>e::real. y < top \<Longrightarrow> 0 < e \<Longrightarrow> x \<le> y + ennreal e) \<Longrightarrow> x \<le> y"
-  apply (cases y rule: ennreal_cases)
-  apply (cases x rule: ennreal_cases)
-  apply (auto simp del: ennreal_plus simp add: top_unique ennreal_plus[symmetric]
-    intro: zero_less_one field_le_epsilon)
-  done
-
-lemma ennreal_less_zero_iff[simp]: "0 < ennreal x \<longleftrightarrow> 0 < x"
-  by transfer (auto simp: max_def)
-
-lemma suminf_ennreal_eq:
-  "(\<And>i. 0 \<le> f i) \<Longrightarrow> f sums x \<Longrightarrow> (\<Sum>i. ennreal (f i)) = ennreal x"
-  using suminf_nonneg[of f] sums_unique[of f x]
-  by (intro sums_unique[symmetric]) (auto simp: summable_sums_iff)
-
-lemma transfer_e2ennreal_sumset [transfer_rule]:
-  "rel_fun (rel_fun op = pcr_ennreal) (rel_fun op = pcr_ennreal) setsum setsum"
-  by (auto intro!: rel_funI simp: rel_fun_eq_pcr_ennreal comp_def)
-
-lemma ennreal_suminf_bound_add:
-  fixes f :: "nat \<Rightarrow> ennreal"
-  shows "(\<And>N. (\<Sum>n<N. f n) + y \<le> x) \<Longrightarrow> suminf f + y \<le> x"
-  by transfer (auto intro!: suminf_bound_add)
-
-lemma divide_right_mono_ennreal:
-  fixes a b c :: ennreal
-  shows "a \<le> b \<Longrightarrow> a / c \<le> b / c"
-  unfolding divide_ennreal_def by (intro mult_mono) auto
-
-lemma SUP_mult_left_ennreal: "c * (SUP i:I. f i) = (SUP i:I. c * f i ::ennreal)"
-proof cases
-  assume "I \<noteq> {}" then show ?thesis
-    by transfer (auto simp add: SUP_ereal_mult_left max_absorb2 SUP_upper2)
-qed (simp add: bot_ennreal)
-
-lemma SUP_mult_right_ennreal: "(SUP i:I. f i) * c = (SUP i:I. f i * c ::ennreal)"
-  using SUP_mult_left_ennreal by (simp add: mult.commute)
-
-lemma SUP_divide_ennreal: "(SUP i:I. f i) / c = (SUP i:I. f i / c ::ennreal)"
-  using SUP_mult_right_ennreal by (simp add: divide_ennreal_def)
-
-lemma of_nat_Sup_ennreal:
-  assumes "A \<noteq> {}" "bdd_above A"
-  shows "of_nat (Sup A) = (SUP a:A. of_nat a :: ennreal)"
-proof (intro antisym)
-  show "(SUP a:A. of_nat a::ennreal) \<le> of_nat (Sup A)"
-    by (intro SUP_least of_nat_mono) (auto intro: cSup_upper assms)
-  have "Sup A \<in> A"
-    unfolding Sup_nat_def using assms by (intro Max_in) (auto simp: bdd_above_nat)
-  then show "of_nat (Sup A) \<le> (SUP a:A. of_nat a::ennreal)"
-    by (intro SUP_upper)
-qed
-
 lemma mult_divide_eq_ennreal:
   fixes a b :: ennreal
   shows "b \<noteq> 0 \<Longrightarrow> b \<noteq> top \<Longrightarrow> (a * b) / b = a"
@@ -895,16 +685,6 @@
     done
   done
 
-lemma ennreal_power: "0 \<le> r \<Longrightarrow> ennreal r ^ n = ennreal (r ^ n)"
-  by (induction n) (auto simp: ennreal_mult)
-
-lemma top_power_ennreal: "top ^ n = (if n = 0 then 1 else top :: ennreal)"
-  by (induction n) (simp_all add: ennreal_mult_eq_top_iff)
-
-lemma power_eq_top_ennreal: "x ^ n = top \<longleftrightarrow> (n \<noteq> 0 \<and> (x::ennreal) = top)"
-  by (cases x rule: ennreal_cases)
-     (auto simp: ennreal_power top_power_ennreal)
-
 lemma ennreal_mult_divide_eq:
   fixes a b :: ennreal
   shows "b \<noteq> 0 \<Longrightarrow> b \<noteq> top \<Longrightarrow> (a * b) / b = a"
@@ -914,69 +694,11 @@
   apply (simp add: top_ereal_def divide_ereal_def[symmetric])
   done
 
-lemma enn2ereal_of_nat[simp]: "enn2ereal (of_nat n) = ereal n"
-  by (induction n) (auto simp: zero_ennreal.rep_eq one_ennreal.rep_eq plus_ennreal.rep_eq)
-
-lemma enn2ereal_numeral[simp]: "enn2ereal (numeral a) = numeral a"
-  apply (subst of_nat_numeral[of a, symmetric])
-  apply (subst enn2ereal_of_nat)
-  apply simp
-  done
-
-lemma transfer_numeral[transfer_rule]: "pcr_ennreal (numeral a) (numeral a)"
-  unfolding cr_ennreal_def pcr_ennreal_def by auto
-
-lemma ennreal_half[simp]: "ennreal (1/2) = inverse 2"
-  by transfer (simp add: max.absorb2)
-
-lemma numeral_eq_of_nat: "(numeral a::ennreal) = of_nat (numeral a)"
-  by simp
-
-lemma of_nat_less_top: "of_nat i < (top::ennreal)"
-  using less_le_trans[of "of_nat i" "of_nat (Suc i)" "top::ennreal"]
-  by simp
-
-lemma top_neq_numeral[simp]: "top \<noteq> ((numeral i)::ennreal)"
-  using of_nat_less_top[of "numeral i"] by simp
-
-lemma sup_continuous_mult_left_ennreal':
-  fixes c :: "ennreal"
-  shows "sup_continuous (\<lambda>x. c * x)"
-  unfolding sup_continuous_def
-  by transfer (auto simp: SUP_ereal_mult_left max.absorb2 SUP_upper2)
-
-lemma sup_continuous_mult_left_ennreal[order_continuous_intros]:
-  "sup_continuous f \<Longrightarrow> sup_continuous (\<lambda>x. c * f x :: ennreal)"
-  by (rule sup_continuous_compose[OF sup_continuous_mult_left_ennreal'])
-
-lemma sup_continuous_mult_right_ennreal[order_continuous_intros]:
-  "sup_continuous f \<Longrightarrow> sup_continuous (\<lambda>x. f x * c :: ennreal)"
-  using sup_continuous_mult_left_ennreal[of f c] by (simp add: mult.commute)
-
-lemma sup_continuous_divide_ennreal[order_continuous_intros]:
-  fixes f g :: "'a::complete_lattice \<Rightarrow> ennreal"
-  shows "sup_continuous f \<Longrightarrow> sup_continuous (\<lambda>x. f x / c)"
-  unfolding divide_ennreal_def by (rule sup_continuous_mult_right_ennreal)
-
-lemma ennreal_add_bot[simp]: "bot + x = (x::ennreal)"
-  by transfer simp
-
-lemma sup_continuous_transfer[transfer_rule]:
-  "(rel_fun (rel_fun (op =) pcr_ennreal) op =) sup_continuous sup_continuous"
-proof (safe intro!: rel_funI dest!: rel_fun_eq_pcr_ennreal[THEN iffD1])
-  show "sup_continuous (enn2ereal \<circ> f) \<Longrightarrow> sup_continuous f" for f :: "'a \<Rightarrow> _"
-    using sup_continuous_e2ennreal[of "enn2ereal \<circ> f"] by simp
-  show "sup_continuous f \<Longrightarrow> sup_continuous (enn2ereal \<circ> f)" for f :: "'a \<Rightarrow> _"
-    using sup_continuous_enn2ereal[of f] by (simp add: comp_def)
-qed
-
-lemma sup_continuous_add_ennreal[order_continuous_intros]:
-  fixes f g :: "'a::complete_lattice \<Rightarrow> ennreal"
-  shows "sup_continuous f \<Longrightarrow> sup_continuous g \<Longrightarrow> sup_continuous (\<lambda>x. f x + g x)"
-  by transfer (auto intro!: sup_continuous_add)
-
-lemmas ennreal2_cases = ennreal_cases[case_product ennreal_cases]
-lemmas ennreal3_cases = ennreal_cases[case_product ennreal2_cases]
+lemma ennreal_add_diff_cancel:
+  fixes a b :: ennreal
+  shows "b \<noteq> \<infinity> \<Longrightarrow> (a + b) - b = a"
+  unfolding infinity_ennreal_def
+  by transfer (simp add: max_absorb2 top_ereal_def ereal_add_diff_cancel)
 
 lemma ennreal_minus_eq_0:
   "a - b = 0 \<Longrightarrow> a \<le> (b::ennreal)"
@@ -1011,121 +733,6 @@
     by (cases a b rule: ereal2_cases) (auto simp: less_max_iff_disj)
   done
 
-lemma ennreal_SUP_add:
-  fixes f g :: "nat \<Rightarrow> ennreal"
-  shows "incseq f \<Longrightarrow> incseq g \<Longrightarrow> (SUP i. f i + g i) = SUPREMUM UNIV f + SUPREMUM UNIV g"
-  unfolding incseq_def le_fun_def
-  by transfer
-     (simp add: SUP_ereal_add incseq_def le_fun_def max_absorb2 SUP_upper2)
-
-lemma ennreal_top_mult: "top * a = (if a = 0 then 0 else top :: ennreal)"
-  by (simp add: ennreal_mult_eq_top_iff)
-
-lemma ennreal_mult_top: "a * top = (if a = 0 then 0 else top :: ennreal)"
-  by (simp add: ennreal_mult_eq_top_iff)
-
-lemma ennreal_less: "0 \<le> r \<Longrightarrow> ennreal r < ennreal q \<longleftrightarrow> r < q"
-  unfolding not_le[symmetric] by auto
-
-lemma ennreal_numeral_less_top[simp]: "numeral i < (top::ennreal)"
-  using of_nat_less_top[of "numeral i"] by simp
-
-lemma real_of_nat_Sup:
-  assumes "A \<noteq> {}" "bdd_above A"
-  shows "of_nat (Sup A) = (SUP a:A. of_nat a :: real)"
-proof (intro antisym)
-  show "(SUP a:A. of_nat a::real) \<le> of_nat (Sup A)"
-    using assms by (intro cSUP_least of_nat_mono) (auto intro: cSup_upper)
-  have "Sup A \<in> A"
-    unfolding Sup_nat_def using assms by (intro Max_in) (auto simp: bdd_above_nat)
-  then show "of_nat (Sup A) \<le> (SUP a:A. of_nat a::real)"
-    by (intro cSUP_upper bdd_above_image_mono assms) (auto simp: mono_def)
-qed
-
-definition "enn2real x = real_of_ereal (enn2ereal x)"
-
-lemma enn2real_nonneg: "0 \<le> enn2real x"
-  by (auto simp: enn2real_def intro!: real_of_ereal_pos enn2ereal_nonneg)
-
-lemma enn2real_mono: "a \<le> b \<Longrightarrow> b < top \<Longrightarrow> enn2real a \<le> enn2real b"
-  by (auto simp add: enn2real_def less_eq_ennreal.rep_eq intro!: real_of_ereal_positive_mono enn2ereal_nonneg)
-
-lemma enn2real_of_nat[simp]: "enn2real (of_nat n) = n"
-  by (auto simp: enn2real_def)
-
-lemma enn2real_ennreal[simp]: "0 \<le> r \<Longrightarrow> enn2real (ennreal r) = r"
-  by (simp add: enn2real_def)
-
-lemma of_nat_le_ennreal_iff[simp]: "0 \<le> r \<Longrightarrow> of_nat i \<le> ennreal r \<longleftrightarrow> of_nat i \<le> r"
-  by (simp add: ennreal_of_nat_eq_real_of_nat)
-
-lemma min_ennreal: "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> min (ennreal x) (ennreal y) = ennreal (min x y)"
-  by (auto split: split_min)
-
-lemma ennreal_approx_unit:
-  "(\<And>a::ennreal. 0 < a \<Longrightarrow> a < 1 \<Longrightarrow> a * z \<le> y) \<Longrightarrow> z \<le> y"
-  apply (subst SUP_mult_right_ennreal[of "\<lambda>x. x" "{0 <..< 1}" z, simplified])
-  apply (rule SUP_least)
-  apply auto
-  done
-
-lemma ennreal_mult_strict_right_mono: "(a::ennreal) < c \<Longrightarrow> 0 < b \<Longrightarrow> b < top \<Longrightarrow> a * b < c * b"
-  by transfer (auto intro!: ereal_mult_strict_right_mono)
-
-lemma ennreal_SUP_setsum:
-  fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ennreal"
-  shows "(\<And>i. i \<in> I \<Longrightarrow> incseq (f i)) \<Longrightarrow> (SUP n. \<Sum>i\<in>I. f i n) = (\<Sum>i\<in>I. SUP n. f i n)"
-  unfolding incseq_def
-  by transfer
-     (simp add: SUP_ereal_setsum incseq_def SUP_upper2 max_absorb2 setsum_nonneg)
-
-lemma ennreal_indicator_less[simp]:
-  "indicator A x \<le> (indicator B x::ennreal) \<longleftrightarrow> (x \<in> A \<longrightarrow> x \<in> B)"
-  by (simp add: indicator_def not_le)
-
-lemma rel_fun_liminf[transfer_rule]: "rel_fun (rel_fun op = pcr_ennreal) pcr_ennreal liminf liminf"
-proof -
-  have "rel_fun (rel_fun op = pcr_ennreal) pcr_ennreal (\<lambda>x. sup 0 (liminf x)) liminf"
-    unfolding liminf_SUP_INF[abs_def] by (transfer_prover_start, transfer_step+; simp)
-  then show ?thesis
-    apply (subst (asm) (2) rel_fun_def)
-    apply (subst (2) rel_fun_def)
-    apply (auto simp: comp_def max.absorb2 Liminf_bounded enn2ereal_nonneg rel_fun_eq_pcr_ennreal)
-    done
-qed
-
-lemma rel_fun_limsup[transfer_rule]: "rel_fun (rel_fun op = pcr_ennreal) pcr_ennreal limsup limsup"
-proof -
-  have "rel_fun (rel_fun op = pcr_ennreal) pcr_ennreal (\<lambda>x. INF n. sup 0 (SUP i:{n..}. x i)) limsup"
-    unfolding limsup_INF_SUP[abs_def] by (transfer_prover_start, transfer_step+; simp)
-  then show ?thesis
-    unfolding limsup_INF_SUP[abs_def]
-    apply (subst (asm) (2) rel_fun_def)
-    apply (subst (2) rel_fun_def)
-    apply (auto simp: comp_def max.absorb2 Sup_upper2 enn2ereal_nonneg rel_fun_eq_pcr_ennreal)
-    apply (subst (asm) max.absorb2)
-    apply (rule SUP_upper2)
-    apply (auto simp: enn2ereal_nonneg)
-    done
-qed
-
-lemma ennreal_liminf_minus:
-  fixes f :: "nat \<Rightarrow> ennreal"
-  shows "(\<And>n. f n \<le> c) \<Longrightarrow> liminf (\<lambda>n. c - f n) = c - limsup f"
-  apply transfer
-  apply (simp add: ereal_diff_positive max.absorb2 liminf_ereal_cminus)
-  apply (subst max.absorb2)
-  apply (rule ereal_diff_positive)
-  apply (rule Limsup_bounded)
-  apply auto
-  done
-
-lemma inverse_ennreal: "0 < r \<Longrightarrow> inverse (ennreal r) = ennreal (inverse r)"
-  by transfer (simp add: max.absorb2)
-
-lemma divide_ennreal: "0 \<le> r \<Longrightarrow> 0 < q \<Longrightarrow> ennreal r / ennreal q = ennreal (r / q)"
-  by (simp add: divide_ennreal_def inverse_ennreal ennreal_mult[symmetric] inverse_eq_divide)
-
 lemma ennreal_inverse_top[simp]: "inverse top = (0::ennreal)"
   by transfer (simp add: top_ereal_def ereal_inverse_eq_0)
 
@@ -1153,4 +760,986 @@
   unfolding divide_ennreal_def
   by transfer (auto simp: ereal_zero_less_0_iff top_ereal_def ereal_0_gt_inverse)
 
+lemma divide_right_mono_ennreal:
+  fixes a b c :: ennreal
+  shows "a \<le> b \<Longrightarrow> a / c \<le> b / c"
+  unfolding divide_ennreal_def by (intro mult_mono) auto
+
+lemma ennreal_mult_strict_right_mono: "(a::ennreal) < c \<Longrightarrow> 0 < b \<Longrightarrow> b < top \<Longrightarrow> a * b < c * b"
+  by transfer (auto intro!: ereal_mult_strict_right_mono)
+
+lemma ennreal_indicator_less[simp]:
+  "indicator A x \<le> (indicator B x::ennreal) \<longleftrightarrow> (x \<in> A \<longrightarrow> x \<in> B)"
+  by (simp add: indicator_def not_le)
+
+lemma ennreal_inverse_positive: "0 < inverse x \<longleftrightarrow> (x::ennreal) \<noteq> top"
+  by transfer (simp add: ereal_0_gt_inverse top_ereal_def)
+
+lemma ennreal_inverse_mult': "((0 < b \<or> a < top) \<and> (0 < a \<or> b < top)) \<Longrightarrow> inverse (a * b::ennreal) = inverse a * inverse b"
+  apply transfer
+  subgoal for a b
+    by (cases a b rule: ereal2_cases) (auto simp: top_ereal_def)
+  done
+
+lemma ennreal_inverse_mult: "a < top \<Longrightarrow> b < top \<Longrightarrow> inverse (a * b::ennreal) = inverse a * inverse b"
+  apply transfer
+  subgoal for a b
+    by (cases a b rule: ereal2_cases) (auto simp: top_ereal_def)
+  done
+
+lemma ennreal_inverse_1[simp]: "inverse (1::ennreal) = 1"
+  by transfer simp
+
+lemma ennreal_inverse_eq_0_iff[simp]: "inverse (a::ennreal) = 0 \<longleftrightarrow> a = top"
+  by transfer (simp add: ereal_inverse_eq_0 top_ereal_def)
+
+lemma ennreal_inverse_eq_top_iff[simp]: "inverse (a::ennreal) = top \<longleftrightarrow> a = 0"
+  by transfer (simp add: top_ereal_def)
+
+lemma ennreal_divide_eq_0_iff[simp]: "(a::ennreal) / b = 0 \<longleftrightarrow> (a = 0 \<or> b = top)"
+  by (simp add: divide_ennreal_def)
+
+lemma ennreal_divide_eq_top_iff: "(a::ennreal) / b = top \<longleftrightarrow> ((a \<noteq> 0 \<and> b = 0) \<or> (a = top \<and> b \<noteq> top))"
+  by (auto simp add: divide_ennreal_def ennreal_mult_eq_top_iff)
+
+lemma one_divide_one_divide_ennreal[simp]: "1 / (1 / c) = (c::ennreal)"
+  including ennreal.lifting
+  unfolding divide_ennreal_def
+  by transfer auto
+
+lemma ennreal_mult_left_cong:
+  "((a::ennreal) \<noteq> 0 \<Longrightarrow> b = c) \<Longrightarrow> a * b = a * c"
+  by (cases "a = 0") simp_all
+
+lemma ennreal_mult_right_cong:
+  "((a::ennreal) \<noteq> 0 \<Longrightarrow> b = c) \<Longrightarrow> b * a = c * a"
+  by (cases "a = 0") simp_all
+
+lemma ennreal_zero_less_mult_iff: "0 < a * b \<longleftrightarrow> 0 < a \<and> 0 < (b::ennreal)"
+  by transfer (auto simp add: ereal_zero_less_0_iff le_less)
+
+lemma less_diff_eq_ennreal:
+  fixes a b c :: ennreal
+  shows "b < top \<or> c < top \<Longrightarrow> a < b - c \<longleftrightarrow> a + c < b"
+  apply transfer
+  subgoal for a b c
+    by (cases a b c rule: ereal3_cases) (auto split: split_max)
+  done
+
+lemma diff_add_cancel_ennreal:
+  fixes a b :: ennreal shows "a \<le> b \<Longrightarrow> b - a + a = b"
+  unfolding infinity_ennreal_def
+  apply transfer
+  subgoal for a b
+    by (cases a b rule: ereal2_cases) (auto simp: max_absorb2)
+  done
+
+lemma ennreal_diff_self[simp]: "a \<noteq> top \<Longrightarrow> a - a = (0::ennreal)"
+  by transfer (simp add: top_ereal_def)
+
+lemma ennreal_minus_mono:
+  fixes a b c :: ennreal
+  shows "a \<le> c \<Longrightarrow> d \<le> b \<Longrightarrow> a - b \<le> c - d"
+  apply transfer
+  apply (rule max.mono)
+  apply simp
+  subgoal for a b c d
+    by (cases a b c d rule: ereal3_cases[case_product ereal_cases]) auto
+  done
+
+lemma ennreal_minus_eq_top[simp]: "a - (b::ennreal) = top \<longleftrightarrow> a = top"
+  by transfer (auto simp: top_ereal_def max.absorb2 ereal_minus_eq_PInfty_iff split: split_max)
+
+lemma ennreal_divide_self[simp]: "a \<noteq> 0 \<Longrightarrow> a < top \<Longrightarrow> a / a = (1::ennreal)"
+  unfolding divide_ennreal_def
+  apply transfer
+  subgoal for a
+    by (cases a) (auto simp: top_ereal_def)
+  done
+
+subsection \<open>Coercion from @{typ real} to @{typ ennreal}\<close>
+
+lift_definition ennreal :: "real \<Rightarrow> ennreal" is "sup 0 \<circ> ereal"
+  by simp
+
+declare [[coercion ennreal]]
+
+lemma ennreal_cases[cases type: ennreal]:
+  fixes x :: ennreal
+  obtains (real) r :: real where "0 \<le> r" "x = ennreal r" | (top) "x = top"
+  apply transfer
+  subgoal for x thesis
+    by (cases x) (auto simp: max.absorb2 top_ereal_def)
+  done
+
+lemmas ennreal2_cases = ennreal_cases[case_product ennreal_cases]
+lemmas ennreal3_cases = ennreal_cases[case_product ennreal2_cases]
+
+lemma ennreal_neq_top[simp]: "ennreal r \<noteq> top"
+  by transfer (simp add: top_ereal_def zero_ereal_def ereal_max[symmetric] del: ereal_max)
+
+lemma top_neq_ennreal[simp]: "top \<noteq> ennreal r"
+  using ennreal_neq_top[of r] by (auto simp del: ennreal_neq_top)
+
+lemma ennreal_less_top[simp]: "ennreal x < top"
+  by transfer (simp add: top_ereal_def max_def)
+
+lemma ennreal_neg: "x \<le> 0 \<Longrightarrow> ennreal x = 0"
+  by transfer (simp add: max.absorb1)
+
+lemma ennreal_inj[simp]:
+  "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> ennreal a = ennreal b \<longleftrightarrow> a = b"
+  by (transfer fixing: a b) (auto simp: max_absorb2)
+
+lemma ennreal_le_iff[simp]: "0 \<le> y \<Longrightarrow> ennreal x \<le> ennreal y \<longleftrightarrow> x \<le> y"
+  by (auto simp: ennreal_def zero_ereal_def less_eq_ennreal.abs_eq eq_onp_def split: split_max)
+
+lemma le_ennreal_iff: "0 \<le> r \<Longrightarrow> x \<le> ennreal r \<longleftrightarrow> (\<exists>q\<ge>0. x = ennreal q \<and> q \<le> r)"
+  by (cases x) (auto simp: top_unique)
+
+lemma ennreal_less_iff: "0 \<le> r \<Longrightarrow> ennreal r < ennreal q \<longleftrightarrow> r < q"
+  unfolding not_le[symmetric] by auto
+
+lemma ennreal_eq_zero_iff[simp]: "0 \<le> x \<Longrightarrow> ennreal x = 0 \<longleftrightarrow> x = 0"
+  by transfer (auto simp: max_absorb2)
+
+lemma ennreal_less_zero_iff[simp]: "0 < ennreal x \<longleftrightarrow> 0 < x"
+  by transfer (auto simp: max_def)
+
+lemma ennreal_lessI: "0 < q \<Longrightarrow> r < q \<Longrightarrow> ennreal r < ennreal q"
+  by (cases "0 \<le> r") (auto simp: ennreal_less_iff ennreal_neg)
+
+lemma ennreal_leI: "x \<le> y \<Longrightarrow> ennreal x \<le> ennreal y"
+  by (cases "0 \<le> y") (auto simp: ennreal_neg)
+
+lemma enn2ereal_ennreal[simp]: "0 \<le> x \<Longrightarrow> enn2ereal (ennreal x) = x"
+  by transfer (simp add: max_absorb2)
+
+lemma e2ennreal_enn2ereal[simp]: "e2ennreal (enn2ereal x) = x"
+  by (simp add: e2ennreal_def max_absorb2 ennreal.enn2ereal_inverse)
+
+lemma ennreal_0[simp]: "ennreal 0 = 0"
+  by (simp add: ennreal_def max.absorb1 zero_ennreal.abs_eq)
+
+lemma ennreal_1[simp]: "ennreal 1 = 1"
+  by transfer (simp add: max_absorb2)
+
+lemma ennreal_eq_0_iff: "ennreal x = 0 \<longleftrightarrow> x \<le> 0"
+  by (cases "0 \<le> x") (auto simp: ennreal_neg)
+
+lemma ennreal_le_iff2: "ennreal x \<le> ennreal y \<longleftrightarrow> ((0 \<le> y \<and> x \<le> y) \<or> (x \<le> 0 \<and> y \<le> 0))"
+  by (cases "0 \<le> y") (auto simp: ennreal_eq_0_iff ennreal_neg)
+
+lemma ennreal_eq_1[simp]: "ennreal x = 1 \<longleftrightarrow> x = 1"
+  by (cases "0 \<le> x")
+     (auto simp: ennreal_neg ennreal_1[symmetric] simp del: ennreal_1)
+
+lemma ennreal_le_1[simp]: "ennreal x \<le> 1 \<longleftrightarrow> x \<le> 1"
+  by (cases "0 \<le> x")
+     (auto simp: ennreal_neg ennreal_1[symmetric] simp del: ennreal_1)
+
+lemma ennreal_ge_1[simp]: "ennreal x \<ge> 1 \<longleftrightarrow> x \<ge> 1"
+  by (cases "0 \<le> x")
+     (auto simp: ennreal_neg ennreal_1[symmetric] simp del: ennreal_1)
+
+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)
+
+lemma setsum_ennreal[simp]: "(\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> (\<Sum>i\<in>I. ennreal (f i)) = ennreal (setsum f I)"
+  by (induction I rule: infinite_finite_induct) (auto simp: setsum_nonneg)
+
+lemma ennreal_of_nat_eq_real_of_nat: "of_nat i = ennreal (of_nat i)"
+  by (induction i) simp_all
+
+lemma of_nat_le_ennreal_iff[simp]: "0 \<le> r \<Longrightarrow> of_nat i \<le> ennreal r \<longleftrightarrow> of_nat i \<le> r"
+  by (simp add: ennreal_of_nat_eq_real_of_nat)
+
+lemma ennreal_le_of_nat_iff[simp]: "ennreal r \<le> of_nat i \<longleftrightarrow> r \<le> of_nat i"
+  by (simp add: ennreal_of_nat_eq_real_of_nat)
+
+lemma ennreal_indicator: "ennreal (indicator A x) = indicator A x"
+  by (auto split: split_indicator)
+
+lemma ennreal_numeral[simp]: "ennreal (numeral n) = numeral n"
+  using ennreal_of_nat_eq_real_of_nat[of "numeral n"] by simp
+
+lemma min_ennreal: "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> min (ennreal x) (ennreal y) = ennreal (min x y)"
+  by (auto split: split_min)
+
+lemma ennreal_half[simp]: "ennreal (1/2) = inverse 2"
+  by transfer (simp add: max.absorb2)
+
+lemma ennreal_minus: "0 \<le> q \<Longrightarrow> ennreal r - ennreal q = ennreal (r - q)"
+  by transfer
+     (simp add: max.absorb2 zero_ereal_def ereal_max[symmetric] del: ereal_max)
+
+lemma ennreal_minus_top[simp]: "ennreal a - top = 0"
+  by (simp add: minus_top_ennreal)
+
+lemma ennreal_mult: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> ennreal (a * b) = ennreal a * ennreal b"
+  by transfer (simp add: max_absorb2)
+
+lemma ennreal_mult': "0 \<le> a \<Longrightarrow> ennreal (a * b) = ennreal a * ennreal b"
+  by (cases "0 \<le> b") (auto simp: ennreal_mult ennreal_neg mult_nonneg_nonpos)
+
+lemma indicator_mult_ennreal: "indicator A x * ennreal r = ennreal (indicator A x * r)"
+  by (simp split: split_indicator)
+
+lemma ennreal_mult'': "0 \<le> b \<Longrightarrow> ennreal (a * b) = ennreal a * ennreal b"
+  by (cases "0 \<le> a") (auto simp: ennreal_mult ennreal_neg mult_nonpos_nonneg)
+
+lemma numeral_mult_ennreal: "0 \<le> x \<Longrightarrow> numeral b * ennreal x = ennreal (numeral b * x)"
+  by (simp add: ennreal_mult)
+
+lemma ennreal_power: "0 \<le> r \<Longrightarrow> ennreal r ^ n = ennreal (r ^ n)"
+  by (induction n) (auto simp: ennreal_mult)
+
+lemma power_eq_top_ennreal: "x ^ n = top \<longleftrightarrow> (n \<noteq> 0 \<and> (x::ennreal) = top)"
+  by (cases x rule: ennreal_cases)
+     (auto simp: ennreal_power top_power_ennreal)
+
+lemma inverse_ennreal: "0 < r \<Longrightarrow> inverse (ennreal r) = ennreal (inverse r)"
+  by transfer (simp add: max.absorb2)
+
+lemma divide_ennreal: "0 \<le> r \<Longrightarrow> 0 < q \<Longrightarrow> ennreal r / ennreal q = ennreal (r / q)"
+  by (simp add: divide_ennreal_def inverse_ennreal ennreal_mult[symmetric] inverse_eq_divide)
+
+lemma ennreal_inverse_power: "inverse (x ^ n :: ennreal) = inverse x ^ n"
+proof (cases x rule: ennreal_cases)
+  case top with power_eq_top_ennreal[of x n] show ?thesis
+    by (cases "n = 0") auto
+next
+  case (real r) then show ?thesis
+  proof cases
+    assume "x = 0" then show ?thesis
+      using power_eq_top_ennreal[of top "n - 1"]
+      by (cases n) (auto simp: ennreal_top_mult)
+  next
+    assume "x \<noteq> 0"
+    with real have "0 < r" by auto
+    with real show ?thesis
+      by (induction n)
+         (auto simp add: ennreal_power ennreal_mult[symmetric] inverse_ennreal)
+  qed
+qed
+
+lemma ennreal_divide_numeral: "0 \<le> x \<Longrightarrow> ennreal x / numeral b = ennreal (x / numeral b)"
+  by (subst divide_ennreal[symmetric]) auto
+
+lemma setprod_ennreal: "(\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> (\<Prod>i\<in>A. ennreal (f i)) = ennreal (setprod f A)"
+  by (induction A rule: infinite_finite_induct)
+     (auto simp: ennreal_mult setprod_nonneg)
+
+lemma mult_right_ennreal_cancel: "a * ennreal c = b * ennreal c \<longleftrightarrow> (a = b \<or> c \<le> 0)"
+  apply (cases "0 \<le> c")
+  apply (cases a b rule: ennreal2_cases)
+  apply (auto simp: ennreal_mult[symmetric] ennreal_neg ennreal_top_mult)
+  done
+
+lemma ennreal_le_epsilon:
+  "(\<And>e::real. y < top \<Longrightarrow> 0 < e \<Longrightarrow> x \<le> y + ennreal e) \<Longrightarrow> x \<le> y"
+  apply (cases y rule: ennreal_cases)
+  apply (cases x rule: ennreal_cases)
+  apply (auto simp del: ennreal_plus simp add: top_unique ennreal_plus[symmetric]
+    intro: zero_less_one field_le_epsilon)
+  done
+
+lemma ennreal_rat_dense:
+  fixes x y :: ennreal
+  shows "x < y \<Longrightarrow> \<exists>r::rat. x < real_of_rat r \<and> real_of_rat r < y"
+proof transfer
+  fix x y :: ereal assume xy: "0 \<le> x" "0 \<le> y" "x < y"
+  moreover
+  from ereal_dense3[OF \<open>x < y\<close>]
+  obtain r where "x < ereal (real_of_rat r)" "ereal (real_of_rat r) < y"
+    by auto
+  moreover then have "0 \<le> r"
+    using le_less_trans[OF \<open>0 \<le> x\<close> \<open>x < ereal (real_of_rat r)\<close>] by auto
+  ultimately show "\<exists>r. x < (sup 0 \<circ> ereal) (real_of_rat r) \<and> (sup 0 \<circ> ereal) (real_of_rat r) < y"
+    by (intro exI[of _ r]) (auto simp: max_absorb2)
+qed
+
+lemma ennreal_Ex_less_of_nat: "(x::ennreal) < top \<Longrightarrow> \<exists>n. x < of_nat n"
+  by (cases x rule: ennreal_cases)
+     (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_less_iff reals_Archimedean2)
+
+subsection \<open>Coercion from @{typ ennreal} to @{typ real}\<close>
+
+definition "enn2real x = real_of_ereal (enn2ereal x)"
+
+lemma enn2real_nonneg[simp]: "0 \<le> enn2real x"
+  by (auto simp: enn2real_def intro!: real_of_ereal_pos enn2ereal_nonneg)
+
+lemma enn2real_mono: "a \<le> b \<Longrightarrow> b < top \<Longrightarrow> enn2real a \<le> enn2real b"
+  by (auto simp add: enn2real_def less_eq_ennreal.rep_eq intro!: real_of_ereal_positive_mono enn2ereal_nonneg)
+
+lemma enn2real_of_nat[simp]: "enn2real (of_nat n) = n"
+  by (auto simp: enn2real_def)
+
+lemma enn2real_ennreal[simp]: "0 \<le> r \<Longrightarrow> enn2real (ennreal r) = r"
+  by (simp add: enn2real_def)
+
+lemma ennreal_enn2real[simp]: "r < top \<Longrightarrow> ennreal (enn2real r) = r"
+  by (cases r rule: ennreal_cases) auto
+
+lemma real_of_ereal_enn2ereal[simp]: "real_of_ereal (enn2ereal x) = enn2real x"
+  by (simp add: enn2real_def)
+
+lemma enn2real_top[simp]: "enn2real top = 0"
+  unfolding enn2real_def top_ennreal.rep_eq top_ereal_def by simp
+
+lemma enn2real_0[simp]: "enn2real 0 = 0"
+  unfolding enn2real_def zero_ennreal.rep_eq by simp
+
+lemma enn2real_1[simp]: "enn2real 1 = 1"
+  unfolding enn2real_def one_ennreal.rep_eq by simp
+
+lemma enn2real_numeral[simp]: "enn2real (numeral n) = (numeral n)"
+  unfolding enn2real_def by simp
+
+lemma enn2real_mult: "enn2real (a * b) = enn2real a * enn2real b"
+  unfolding enn2real_def
+  by (simp del: real_of_ereal_enn2ereal add: times_ennreal.rep_eq)
+
+lemma enn2real_leI: "0 \<le> B \<Longrightarrow> x \<le> ennreal B \<Longrightarrow> enn2real x \<le> B"
+  by (cases x rule: ennreal_cases) (auto simp: top_unique)
+
+lemma enn2real_positive_iff: "0 < enn2real x \<longleftrightarrow> (0 < x \<and> x < top)"
+  by (cases x rule: ennreal_cases) auto
+
+subsection \<open>Coercion from @{typ enat} to @{typ ennreal}\<close>
+
+
+definition ennreal_of_enat :: "enat \<Rightarrow> ennreal"
+where
+  "ennreal_of_enat n = (case n of \<infinity> \<Rightarrow> top | enat n \<Rightarrow> of_nat n)"
+
+declare [[coercion ennreal_of_enat]]
+declare [[coercion "of_nat :: nat \<Rightarrow> ennreal"]]
+
+lemma ennreal_of_enat_infty[simp]: "ennreal_of_enat \<infinity> = \<infinity>"
+  by (simp add: ennreal_of_enat_def)
+
+lemma ennreal_of_enat_enat[simp]: "ennreal_of_enat (enat n) = of_nat n"
+  by (simp add: ennreal_of_enat_def)
+
+lemma ennreal_of_enat_0[simp]: "ennreal_of_enat 0 = 0"
+  using ennreal_of_enat_enat[of 0] unfolding enat_0 by simp
+
+lemma ennreal_of_enat_1[simp]: "ennreal_of_enat 1 = 1"
+  using ennreal_of_enat_enat[of 1] unfolding enat_1 by simp
+
+lemma ennreal_top_neq_of_nat[simp]: "(top::ennreal) \<noteq> of_nat i"
+  using ennreal_of_nat_neq_top[of i] by metis
+
+lemma ennreal_of_enat_inj[simp]: "ennreal_of_enat i = ennreal_of_enat j \<longleftrightarrow> i = j"
+  by (cases i j rule: enat.exhaust[case_product enat.exhaust]) auto
+
+lemma ennreal_of_enat_le_iff[simp]: "ennreal_of_enat m \<le> ennreal_of_enat n \<longleftrightarrow> m \<le> n"
+  by (auto simp: ennreal_of_enat_def top_unique split: enat.split)
+
+lemma of_nat_less_ennreal_of_nat[simp]: "of_nat n \<le> ennreal_of_enat x \<longleftrightarrow> of_nat n \<le> x"
+  by (cases x) (auto simp: of_nat_eq_enat)
+
+lemma ennreal_of_enat_Sup: "ennreal_of_enat (Sup X) = (SUP x:X. ennreal_of_enat x)"
+proof -
+  have "ennreal_of_enat (Sup X) \<le> (SUP x : X. ennreal_of_enat x)"
+    unfolding Sup_enat_def
+  proof (clarsimp, intro conjI impI)
+    fix x assume "finite X" "X \<noteq> {}"
+    then show "ennreal_of_enat (Max X) \<le> (SUP x : X. ennreal_of_enat x)"
+      by (intro SUP_upper Max_in)
+  next
+    assume "infinite X" "X \<noteq> {}"
+    have "\<exists>y\<in>X. r < ennreal_of_enat y" if r: "r < top" for r
+    proof -
+      from ennreal_Ex_less_of_nat[OF r] guess n .. note n = this
+      have "\<not> (X \<subseteq> enat ` {.. n})"
+        using \<open>infinite X\<close> by (auto dest: finite_subset)
+      then obtain x where "x \<in> X" "x \<notin> enat ` {..n}"
+        by blast
+      moreover then have "of_nat n \<le> x"
+        by (cases x) (auto simp: of_nat_eq_enat)
+      ultimately show ?thesis
+        by (auto intro!: bexI[of _ x] less_le_trans[OF n])
+    qed
+    then have "(SUP x : X. ennreal_of_enat x) = top"
+      by simp
+    then show "top \<le> (SUP x : X. ennreal_of_enat x)"
+      unfolding top_unique by simp
+  qed
+  then show ?thesis
+    by (auto intro!: antisym Sup_least intro: Sup_upper)
+qed
+
+lemma ennreal_of_enat_eSuc[simp]: "ennreal_of_enat (eSuc x) = 1 + ennreal_of_enat x"
+  by (cases x) (auto simp: eSuc_enat)
+
+subsection \<open>Topology on @{typ ennreal}\<close>
+
+lemma enn2ereal_Iio: "enn2ereal -` {..<a} = (if 0 \<le> a then {..< e2ennreal a} else {})"
+  using enn2ereal_nonneg
+  by (cases a rule: ereal_ennreal_cases)
+     (auto simp add: vimage_def set_eq_iff ennreal.enn2ereal_inverse less_ennreal.rep_eq e2ennreal_def max_absorb2
+           simp del: enn2ereal_nonneg
+           intro: le_less_trans less_imp_le)
+
+lemma enn2ereal_Ioi: "enn2ereal -` {a <..} = (if 0 \<le> a then {e2ennreal a <..} else UNIV)"
+  by (cases a rule: ereal_ennreal_cases)
+     (auto simp add: vimage_def set_eq_iff ennreal.enn2ereal_inverse less_ennreal.rep_eq e2ennreal_def max_absorb2
+           intro: less_le_trans)
+
+instantiation ennreal :: linear_continuum_topology
+begin
+
+definition open_ennreal :: "ennreal set \<Rightarrow> bool"
+  where "(open :: ennreal set \<Rightarrow> bool) = generate_topology (range lessThan \<union> range greaterThan)"
+
+instance
+proof
+  show "\<exists>a b::ennreal. a \<noteq> b"
+    using zero_neq_one by (intro exI)
+  show "\<And>x y::ennreal. x < y \<Longrightarrow> \<exists>z>x. z < y"
+  proof transfer
+    fix x y :: ereal assume "0 \<le> x" "x < y"
+    moreover from dense[OF this(2)] guess z ..
+    ultimately show "\<exists>z\<in>Collect (op \<le> 0). x < z \<and> z < y"
+      by (intro bexI[of _ z]) auto
+  qed
+qed (rule open_ennreal_def)
+
 end
+
+lemma continuous_on_e2ennreal: "continuous_on A e2ennreal"
+proof (rule continuous_on_subset)
+  show "continuous_on ({0..} \<union> {..0}) e2ennreal"
+  proof (rule continuous_on_closed_Un)
+    show "continuous_on {0 ..} e2ennreal"
+      by (rule continuous_onI_mono)
+         (auto simp add: less_eq_ennreal.abs_eq eq_onp_def enn2ereal_range)
+    show "continuous_on {.. 0} e2ennreal"
+      by (subst continuous_on_cong[OF refl, of _ _ "\<lambda>_. 0"])
+         (auto simp add: e2ennreal_neg continuous_on_const)
+  qed auto
+  show "A \<subseteq> {0..} \<union> {..0::ereal}"
+    by auto
+qed
+
+lemma continuous_at_e2ennreal: "continuous (at x within A) e2ennreal"
+  by (rule continuous_on_imp_continuous_within[OF continuous_on_e2ennreal, of _ UNIV]) auto
+
+lemma continuous_on_enn2ereal: "continuous_on UNIV enn2ereal"
+  by (rule continuous_on_generate_topology[OF open_generated_order])
+     (auto simp add: enn2ereal_Iio enn2ereal_Ioi)
+
+lemma continuous_at_enn2ereal: "continuous (at x within A) enn2ereal"
+  by (rule continuous_on_imp_continuous_within[OF continuous_on_enn2ereal]) auto
+
+lemma sup_continuous_e2ennreal[order_continuous_intros]:
+  assumes f: "sup_continuous f" shows "sup_continuous (\<lambda>x. e2ennreal (f x))"
+  apply (rule sup_continuous_compose[OF _ f])
+  apply (rule continuous_at_left_imp_sup_continuous)
+  apply (auto simp: mono_def e2ennreal_mono continuous_at_e2ennreal)
+  done
+
+lemma sup_continuous_enn2ereal[order_continuous_intros]:
+  assumes f: "sup_continuous f" shows "sup_continuous (\<lambda>x. enn2ereal (f x))"
+  apply (rule sup_continuous_compose[OF _ f])
+  apply (rule continuous_at_left_imp_sup_continuous)
+  apply (simp_all add: mono_def less_eq_ennreal.rep_eq continuous_at_enn2ereal)
+  done
+
+lemma sup_continuous_mult_left_ennreal':
+  fixes c :: "ennreal"
+  shows "sup_continuous (\<lambda>x. c * x)"
+  unfolding sup_continuous_def
+  by transfer (auto simp: SUP_ereal_mult_left max.absorb2 SUP_upper2)
+
+lemma sup_continuous_mult_left_ennreal[order_continuous_intros]:
+  "sup_continuous f \<Longrightarrow> sup_continuous (\<lambda>x. c * f x :: ennreal)"
+  by (rule sup_continuous_compose[OF sup_continuous_mult_left_ennreal'])
+
+lemma sup_continuous_mult_right_ennreal[order_continuous_intros]:
+  "sup_continuous f \<Longrightarrow> sup_continuous (\<lambda>x. f x * c :: ennreal)"
+  using sup_continuous_mult_left_ennreal[of f c] by (simp add: mult.commute)
+
+lemma sup_continuous_divide_ennreal[order_continuous_intros]:
+  fixes f g :: "'a::complete_lattice \<Rightarrow> ennreal"
+  shows "sup_continuous f \<Longrightarrow> sup_continuous (\<lambda>x. f x / c)"
+  unfolding divide_ennreal_def by (rule sup_continuous_mult_right_ennreal)
+
+lemma transfer_enn2ereal_continuous_on [transfer_rule]:
+  "rel_fun (op =) (rel_fun (rel_fun op = pcr_ennreal) op =) continuous_on continuous_on"
+proof -
+  have "continuous_on A f" if "continuous_on A (\<lambda>x. enn2ereal (f x))" for A and f :: "'a \<Rightarrow> ennreal"
+    using continuous_on_compose2[OF continuous_on_e2ennreal[of "{0..}"] that]
+    by (auto simp: ennreal.enn2ereal_inverse subset_eq e2ennreal_def max_absorb2)
+  moreover
+  have "continuous_on A (\<lambda>x. enn2ereal (f x))" if "continuous_on A f" for A and f :: "'a \<Rightarrow> ennreal"
+    using continuous_on_compose2[OF continuous_on_enn2ereal that] by auto
+  ultimately
+  show ?thesis
+    by (auto simp add: rel_fun_def ennreal.pcr_cr_eq cr_ennreal_def)
+qed
+
+lemma transfer_sup_continuous[transfer_rule]:
+  "(rel_fun (rel_fun (op =) pcr_ennreal) op =) sup_continuous sup_continuous"
+proof (safe intro!: rel_funI dest!: rel_fun_eq_pcr_ennreal[THEN iffD1])
+  show "sup_continuous (enn2ereal \<circ> f) \<Longrightarrow> sup_continuous f" for f :: "'a \<Rightarrow> _"
+    using sup_continuous_e2ennreal[of "enn2ereal \<circ> f"] by simp
+  show "sup_continuous f \<Longrightarrow> sup_continuous (enn2ereal \<circ> f)" for f :: "'a \<Rightarrow> _"
+    using sup_continuous_enn2ereal[of f] by (simp add: comp_def)
+qed
+
+lemma continuous_on_ennreal[tendsto_intros]:
+  "continuous_on A f \<Longrightarrow> continuous_on A (\<lambda>x. ennreal (f x))"
+  by transfer (auto intro!: continuous_on_max continuous_on_const continuous_on_ereal)
+
+lemma tendsto_ennrealD:
+  assumes lim: "((\<lambda>x. ennreal (f x)) \<longlongrightarrow> ennreal x) F"
+  assumes *: "\<forall>\<^sub>F x in F. 0 \<le> f x" and x: "0 \<le> x"
+  shows "(f \<longlongrightarrow> x) F"
+  using continuous_on_tendsto_compose[OF continuous_on_enn2ereal lim]
+  apply simp
+  apply (subst (asm) tendsto_cong)
+  using *
+  apply eventually_elim
+  apply (auto simp: max_absorb2 \<open>0 \<le> x\<close>)
+  done
+
+lemma tendsto_ennreal_iff[simp]:
+  "\<forall>\<^sub>F x in F. 0 \<le> f x \<Longrightarrow> 0 \<le> x \<Longrightarrow> ((\<lambda>x. ennreal (f x)) \<longlongrightarrow> ennreal x) F \<longleftrightarrow> (f \<longlongrightarrow> x) F"
+  by (auto dest: tendsto_ennrealD)
+     (auto simp: ennreal_def
+           intro!: continuous_on_tendsto_compose[OF continuous_on_e2ennreal[of UNIV]] tendsto_max)
+
+lemma tendsto_enn2ereal_iff[simp]: "((\<lambda>i. enn2ereal (f i)) \<longlongrightarrow> enn2ereal x) F \<longleftrightarrow> (f \<longlongrightarrow> x) F"
+  using continuous_on_enn2ereal[THEN continuous_on_tendsto_compose, of f x F]
+    continuous_on_e2ennreal[THEN continuous_on_tendsto_compose, of "\<lambda>x. enn2ereal (f x)" "enn2ereal x" F UNIV]
+  by auto
+
+lemma continuous_on_add_ennreal:
+  fixes f g :: "'a::topological_space \<Rightarrow> ennreal"
+  shows "continuous_on A f \<Longrightarrow> continuous_on A g \<Longrightarrow> continuous_on A (\<lambda>x. f x + g x)"
+  by (transfer fixing: A) (auto intro!: tendsto_add_ereal_nonneg simp: continuous_on_def)
+
+lemma continuous_on_inverse_ennreal[continuous_intros]:
+  fixes f :: "'a::topological_space \<Rightarrow> ennreal"
+  shows "continuous_on A f \<Longrightarrow> continuous_on A (\<lambda>x. inverse (f x))"
+proof (transfer fixing: A)
+  show "pred_fun (\<lambda>_. True)  (op \<le> 0) f \<Longrightarrow> continuous_on A (\<lambda>x. inverse (f x))" if "continuous_on A f"
+    for f :: "'a \<Rightarrow> ereal"
+    using continuous_on_compose2[OF continuous_on_inverse_ereal that] by (auto simp: subset_eq)
+qed
+
+instance ennreal :: topological_comm_monoid_add
+proof
+  show "((\<lambda>x. fst x + snd x) \<longlongrightarrow> a + b) (nhds a \<times>\<^sub>F nhds b)" for a b :: ennreal
+    using continuous_on_add_ennreal[of UNIV fst snd]
+    using tendsto_at_iff_tendsto_nhds[symmetric, of "\<lambda>x::(ennreal \<times> ennreal). fst x + snd x"]
+    by (auto simp: continuous_on_eq_continuous_at)
+       (simp add: isCont_def nhds_prod[symmetric])
+qed
+
+lemma sup_continuous_add_ennreal[order_continuous_intros]:
+  fixes f g :: "'a::complete_lattice \<Rightarrow> ennreal"
+  shows "sup_continuous f \<Longrightarrow> sup_continuous g \<Longrightarrow> sup_continuous (\<lambda>x. f x + g x)"
+  by transfer (auto intro!: sup_continuous_add)
+
+lemma ennreal_suminf_lessD: "(\<Sum>i. f i :: ennreal) < x \<Longrightarrow> f i < x"
+  using le_less_trans[OF setsum_le_suminf[OF summableI, of "{i}" f]] by simp
+
+lemma sums_ennreal[simp]: "(\<And>i. 0 \<le> f i) \<Longrightarrow> 0 \<le> x \<Longrightarrow> (\<lambda>i. ennreal (f i)) sums ennreal x \<longleftrightarrow> f sums x"
+  unfolding sums_def by (simp add: always_eventually setsum_nonneg)
+
+lemma summable_suminf_not_top: "(\<And>i. 0 \<le> f i) \<Longrightarrow> (\<Sum>i. ennreal (f i)) \<noteq> top \<Longrightarrow> summable f"
+  using summable_sums[OF summableI, of "\<lambda>i. ennreal (f i)"]
+  by (cases "\<Sum>i. ennreal (f i)" rule: ennreal_cases)
+     (auto simp: summable_def)
+
+lemma suminf_ennreal[simp]:
+  "(\<And>i. 0 \<le> f i) \<Longrightarrow> (\<Sum>i. ennreal (f i)) \<noteq> top \<Longrightarrow> (\<Sum>i. ennreal (f i)) = ennreal (\<Sum>i. f i)"
+  by (rule sums_unique[symmetric]) (simp add: summable_suminf_not_top suminf_nonneg summable_sums)
+
+lemma sums_enn2ereal[simp]: "(\<lambda>i. enn2ereal (f i)) sums enn2ereal x \<longleftrightarrow> f sums x"
+  unfolding sums_def by (simp add: always_eventually setsum_nonneg)
+
+lemma suminf_enn2ereal[simp]: "(\<Sum>i. enn2ereal (f i)) = enn2ereal (suminf f)"
+  by (rule sums_unique[symmetric]) (simp add: summable_sums)
+
+lemma transfer_e2ennreal_suminf [transfer_rule]: "rel_fun (rel_fun op = pcr_ennreal) pcr_ennreal suminf suminf"
+  by (auto simp: rel_funI rel_fun_eq_pcr_ennreal comp_def)
+
+lemma ennreal_suminf_cmult[simp]: "(\<Sum>i. r * f i) = r * (\<Sum>i. f i::ennreal)"
+  by transfer (auto intro!: suminf_cmult_ereal)
+
+lemma ennreal_suminf_multc[simp]: "(\<Sum>i. f i * r) = (\<Sum>i. f i::ennreal) * r"
+  using ennreal_suminf_cmult[of r f] by (simp add: ac_simps)
+
+lemma ennreal_suminf_divide[simp]: "(\<Sum>i. f i / r) = (\<Sum>i. f i::ennreal) / r"
+  by (simp add: divide_ennreal_def)
+
+lemma ennreal_suminf_neq_top: "summable f \<Longrightarrow> (\<And>i. 0 \<le> f i) \<Longrightarrow> (\<Sum>i. ennreal (f i)) \<noteq> top"
+  using sums_ennreal[of f "suminf f"]
+  by (simp add: suminf_nonneg sums_unique[symmetric] summable_sums_iff[symmetric] del: sums_ennreal)
+
+lemma suminf_ennreal_eq:
+  "(\<And>i. 0 \<le> f i) \<Longrightarrow> f sums x \<Longrightarrow> (\<Sum>i. ennreal (f i)) = ennreal x"
+  using suminf_nonneg[of f] sums_unique[of f x]
+  by (intro sums_unique[symmetric]) (auto simp: summable_sums_iff)
+
+lemma ennreal_suminf_bound_add:
+  fixes f :: "nat \<Rightarrow> ennreal"
+  shows "(\<And>N. (\<Sum>n<N. f n) + y \<le> x) \<Longrightarrow> suminf f + y \<le> x"
+  by transfer (auto intro!: suminf_bound_add)
+
+lemma ennreal_suminf_SUP_eq_directed:
+  fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ennreal"
+  assumes *: "\<And>N i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> finite N \<Longrightarrow> \<exists>k\<in>I. \<forall>n\<in>N. f i n \<le> f k n \<and> f j n \<le> f k n"
+  shows "(\<Sum>n. SUP i:I. f i n) = (SUP i:I. \<Sum>n. f i n)"
+proof cases
+  assume "I \<noteq> {}"
+  then obtain i where "i \<in> I" by auto
+  from * show ?thesis
+    by (transfer fixing: I)
+       (auto simp: max_absorb2 SUP_upper2[OF \<open>i \<in> I\<close>] suminf_nonneg summable_ereal_pos \<open>I \<noteq> {}\<close>
+             intro!: suminf_SUP_eq_directed)
+qed (simp add: bot_ennreal)
+
+lemma INF_ennreal_add_const:
+  fixes f g :: "nat \<Rightarrow> ennreal"
+  shows "(INF i. f i + c) = (INF i. f i) + c"
+  using continuous_at_Inf_mono[of "\<lambda>x. x + c" "f`UNIV"]
+  using continuous_add[of "at_right (Inf (range f))", of "\<lambda>x. x" "\<lambda>x. c"]
+  by (auto simp: mono_def)
+
+lemma INF_ennreal_const_add:
+  fixes f g :: "nat \<Rightarrow> ennreal"
+  shows "(INF i. c + f i) = c + (INF i. f i)"
+  using INF_ennreal_add_const[of f c] by (simp add: ac_simps)
+
+lemma SUP_mult_left_ennreal: "c * (SUP i:I. f i) = (SUP i:I. c * f i ::ennreal)"
+proof cases
+  assume "I \<noteq> {}" then show ?thesis
+    by transfer (auto simp add: SUP_ereal_mult_left max_absorb2 SUP_upper2)
+qed (simp add: bot_ennreal)
+
+lemma SUP_mult_right_ennreal: "(SUP i:I. f i) * c = (SUP i:I. f i * c ::ennreal)"
+  using SUP_mult_left_ennreal by (simp add: mult.commute)
+
+lemma SUP_divide_ennreal: "(SUP i:I. f i) / c = (SUP i:I. f i / c ::ennreal)"
+  using SUP_mult_right_ennreal by (simp add: divide_ennreal_def)
+
+lemma ennreal_SUP_of_nat_eq_top: "(SUP x. of_nat x :: ennreal) = top"
+proof (intro antisym top_greatest le_SUP_iff[THEN iffD2] allI impI)
+  fix y :: ennreal assume "y < top"
+  then obtain r where "y = ennreal r"
+    by (cases y rule: ennreal_cases) auto
+  then show "\<exists>i\<in>UNIV. y < of_nat i"
+    using reals_Archimedean2[of "max 1 r"] zero_less_one
+    by (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_def less_ennreal.abs_eq eq_onp_def max.absorb2
+             dest!: one_less_of_natD intro: less_trans)
+qed
+
+lemma ennreal_SUP_eq_top:
+  fixes f :: "'a \<Rightarrow> ennreal"
+  assumes "\<And>n. \<exists>i\<in>I. of_nat n \<le> f i"
+  shows "(SUP i : I. f i) = top"
+proof -
+  have "(SUP x. of_nat x :: ennreal) \<le> (SUP i : I. f i)"
+    using assms by (auto intro!: SUP_least intro: SUP_upper2)
+  then show ?thesis
+    by (auto simp: ennreal_SUP_of_nat_eq_top top_unique)
+qed
+
+lemma ennreal_INF_const_minus:
+  fixes f :: "'a \<Rightarrow> ennreal"
+  shows "I \<noteq> {} \<Longrightarrow> (SUP x:I. c - f x) = c - (INF x:I. f x)"
+  by (transfer fixing: I)
+     (simp add: sup_max[symmetric] SUP_sup_const1 SUP_ereal_minus_right del: sup_ereal_def)
+
+lemma of_nat_Sup_ennreal:
+  assumes "A \<noteq> {}" "bdd_above A"
+  shows "of_nat (Sup A) = (SUP a:A. of_nat a :: ennreal)"
+proof (intro antisym)
+  show "(SUP a:A. of_nat a::ennreal) \<le> of_nat (Sup A)"
+    by (intro SUP_least of_nat_mono) (auto intro: cSup_upper assms)
+  have "Sup A \<in> A"
+    unfolding Sup_nat_def using assms by (intro Max_in) (auto simp: bdd_above_nat)
+  then show "of_nat (Sup A) \<le> (SUP a:A. of_nat a::ennreal)"
+    by (intro SUP_upper)
+qed
+
+lemma ennreal_tendsto_const_minus:
+  fixes g :: "'a \<Rightarrow> ennreal"
+  assumes ae: "\<forall>\<^sub>F x in F. g x \<le> c"
+  assumes g: "((\<lambda>x. c - g x) \<longlongrightarrow> 0) F"
+  shows "(g \<longlongrightarrow> c) F"
+proof (cases c rule: ennreal_cases)
+  case top with tendsto_unique[OF _ g, of "top"] show ?thesis
+    by (cases "F = bot") auto
+next
+  case (real r)
+  then have "\<forall>x. \<exists>q\<ge>0. g x \<le> c \<longrightarrow> (g x = ennreal q \<and> q \<le> r)"
+    by (auto simp: le_ennreal_iff)
+  then obtain f where *: "\<And>x. g x \<le> c \<Longrightarrow> 0 \<le> f x" "\<And>x. g x \<le> c \<Longrightarrow> g x = ennreal (f x)" "\<And>x. g x \<le> c \<Longrightarrow> f x \<le> r"
+    by metis
+  from ae have ae2: "\<forall>\<^sub>F x in F. c - g x = ennreal (r - f x) \<and> f x \<le> r \<and> g x = ennreal (f x) \<and> 0 \<le> f x"
+  proof eventually_elim
+    fix x assume "g x \<le> c" with *[of x] \<open>0 \<le> r\<close> show "c - g x = ennreal (r - f x) \<and> f x \<le> r \<and> g x = ennreal (f x) \<and> 0 \<le> f x"
+      by (auto simp: real ennreal_minus)
+  qed
+  with g have "((\<lambda>x. ennreal (r - f x)) \<longlongrightarrow> ennreal 0) F"
+    by (auto simp add: tendsto_cong eventually_conj_iff)
+  with ae2 have "((\<lambda>x. r - f x) \<longlongrightarrow> 0) F"
+    by (subst (asm) tendsto_ennreal_iff) (auto elim: eventually_mono)
+  then have "(f \<longlongrightarrow> r) F"
+    by (rule Lim_transform2[OF tendsto_const])
+  with ae2 have "((\<lambda>x. ennreal (f x)) \<longlongrightarrow> ennreal r) F"
+    by (subst tendsto_ennreal_iff) (auto elim: eventually_mono simp: real)
+  with ae2 show ?thesis
+    by (auto simp: real tendsto_cong eventually_conj_iff)
+qed
+
+lemma ennreal_SUP_add:
+  fixes f g :: "nat \<Rightarrow> ennreal"
+  shows "incseq f \<Longrightarrow> incseq g \<Longrightarrow> (SUP i. f i + g i) = SUPREMUM UNIV f + SUPREMUM UNIV g"
+  unfolding incseq_def le_fun_def
+  by transfer
+     (simp add: SUP_ereal_add incseq_def le_fun_def max_absorb2 SUP_upper2)
+
+lemma ennreal_SUP_setsum:
+  fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ennreal"
+  shows "(\<And>i. i \<in> I \<Longrightarrow> incseq (f i)) \<Longrightarrow> (SUP n. \<Sum>i\<in>I. f i n) = (\<Sum>i\<in>I. SUP n. f i n)"
+  unfolding incseq_def
+  by transfer
+     (simp add: SUP_ereal_setsum incseq_def SUP_upper2 max_absorb2 setsum_nonneg)
+
+lemma ennreal_liminf_minus:
+  fixes f :: "nat \<Rightarrow> ennreal"
+  shows "(\<And>n. f n \<le> c) \<Longrightarrow> liminf (\<lambda>n. c - f n) = c - limsup f"
+  apply transfer
+  apply (simp add: ereal_diff_positive max.absorb2 liminf_ereal_cminus)
+  apply (subst max.absorb2)
+  apply (rule ereal_diff_positive)
+  apply (rule Limsup_bounded)
+  apply auto
+  done
+
+lemma ennreal_continuous_on_cmult:
+  "(c::ennreal) < top \<Longrightarrow> continuous_on A f \<Longrightarrow> continuous_on A (\<lambda>x. c * f x)"
+  by (transfer fixing: A) (auto intro: continuous_on_cmult_ereal)
+
+lemma ennreal_tendsto_cmult:
+  "(c::ennreal) < top \<Longrightarrow> (f \<longlongrightarrow> x) F \<Longrightarrow> ((\<lambda>x. c * f x) \<longlongrightarrow> c * x) F"
+  by (rule continuous_on_tendsto_compose[where g=f, OF ennreal_continuous_on_cmult, where s=UNIV])
+     (auto simp: continuous_on_id)
+
+lemma tendsto_ennrealI[intro, simp]:
+  "(f \<longlongrightarrow> x) F \<Longrightarrow> ((\<lambda>x. ennreal (f x)) \<longlongrightarrow> ennreal x) F"
+  by (auto simp: ennreal_def
+           intro!: continuous_on_tendsto_compose[OF continuous_on_e2ennreal[of UNIV]] tendsto_max)
+
+lemma ennreal_suminf_minus:
+  fixes f g :: "nat \<Rightarrow> ennreal"
+  shows "(\<And>i. g i \<le> f i) \<Longrightarrow> suminf f \<noteq> top \<Longrightarrow> suminf g \<noteq> top \<Longrightarrow> (\<Sum>i. f i - g i) = suminf f - suminf g"
+  by transfer
+     (auto simp add: max.absorb2 ereal_diff_positive suminf_le_pos top_ereal_def intro!: suminf_ereal_minus)
+
+lemma ennreal_Sup_countable_SUP:
+  "A \<noteq> {} \<Longrightarrow> \<exists>f::nat \<Rightarrow> ennreal. incseq f \<and> range f \<subseteq> A \<and> Sup A = (SUP i. f i)"
+  unfolding incseq_def
+  apply transfer
+  subgoal for A
+    using Sup_countable_SUP[of A]
+    apply (clarsimp simp add: incseq_def[symmetric] SUP_upper2 max.absorb2 image_subset_iff Sup_upper2 cong: conj_cong)
+    subgoal for f
+      by (intro exI[of _ f]) auto
+    done
+  done
+
+lemma ennreal_SUP_countable_SUP:
+  "A \<noteq> {} \<Longrightarrow> \<exists>f::nat \<Rightarrow> ennreal. range f \<subseteq> g`A \<and> SUPREMUM A g = SUPREMUM UNIV f"
+  using ennreal_Sup_countable_SUP [of "g`A"] by auto
+
+lemma of_nat_tendsto_top_ennreal: "(\<lambda>n::nat. of_nat n :: ennreal) \<longlonglongrightarrow> top"
+  using LIMSEQ_SUP[of "of_nat :: nat \<Rightarrow> ennreal"]
+  by (simp add: ennreal_SUP_of_nat_eq_top incseq_def)
+
+lemma SUP_sup_continuous_ennreal:
+  fixes f :: "ennreal \<Rightarrow> 'a::complete_lattice"
+  assumes f: "sup_continuous f" and "I \<noteq> {}"
+  shows "(SUP i:I. f (g i)) = f (SUP i:I. g i)"
+proof (rule antisym)
+  show "(SUP i:I. f (g i)) \<le> f (SUP i:I. g i)"
+    by (rule mono_SUP[OF sup_continuous_mono[OF f]])
+  from ennreal_Sup_countable_SUP[of "g`I"] \<open>I \<noteq> {}\<close>
+  obtain M :: "nat \<Rightarrow> ennreal" where "incseq M" and M: "range M \<subseteq> g ` I" and eq: "(SUP i : I. g i) = (SUP i. M i)"
+    by auto
+  have "f (SUP i : I. g i) = (SUP i : range M. f i)"
+    unfolding eq sup_continuousD[OF f \<open>mono M\<close>] by simp
+  also have "\<dots> \<le> (SUP i : I. f (g i))"
+    by (insert M, drule SUP_subset_mono) auto
+  finally show "f (SUP i : I. g i) \<le> (SUP i : I. f (g i))" .
+qed
+
+lemma ennreal_suminf_SUP_eq:
+  fixes f :: "nat \<Rightarrow> nat \<Rightarrow> ennreal"
+  shows "(\<And>i. incseq (\<lambda>n. f n i)) \<Longrightarrow> (\<Sum>i. SUP n. f n i) = (SUP n. \<Sum>i. f n i)"
+  apply (rule ennreal_suminf_SUP_eq_directed)
+  subgoal for N n j
+    by (auto simp: incseq_def intro!:exI[of _ "max n j"])
+  done
+
+lemma ennreal_SUP_add_left:
+  fixes c :: ennreal
+  shows "I \<noteq> {} \<Longrightarrow> (SUP i:I. f i + c) = (SUP i:I. f i) + c"
+  apply transfer
+  apply (simp add: SUP_ereal_add_left)
+  apply (subst (1 2) max.absorb2)
+  apply (auto intro: SUP_upper2 ereal_add_nonneg_nonneg)
+  done
+
+lemma ennreal_SUP_const_minus:
+  fixes f :: "'a \<Rightarrow> ennreal"
+  shows "I \<noteq> {} \<Longrightarrow> c < top \<Longrightarrow> (INF x:I. c - f x) = c - (SUP x:I. f x)"
+  apply (transfer fixing: I)
+  unfolding ex_in_conv[symmetric]
+  apply (auto simp add: sup_max[symmetric] SUP_upper2 sup_absorb2
+              simp del: sup_ereal_def)
+  apply (subst INF_ereal_minus_right[symmetric])
+  apply (auto simp del: sup_ereal_def simp add: sup_INF)
+  done
+
+subsection \<open>Approximation lemmas\<close>
+
+lemma INF_approx_ennreal:
+  fixes x::ennreal and e::real
+  assumes "e > 0"
+  assumes INF: "x = (INF i : A. f i)"
+  assumes "x \<noteq> \<infinity>"
+  shows "\<exists>i \<in> A. f i < x + e"
+proof -
+  have "(INF i : A. f i) < x + e"
+    unfolding INF[symmetric] using \<open>0<e\<close> \<open>x \<noteq> \<infinity>\<close> by (cases x) auto
+  then show ?thesis
+    unfolding INF_less_iff .
+qed
+
+lemma SUP_approx_ennreal:
+  fixes x::ennreal and e::real
+  assumes "e > 0" "A \<noteq> {}"
+  assumes SUP: "x = (SUP i : A. f i)"
+  assumes "x \<noteq> \<infinity>"
+  shows "\<exists>i \<in> A. x < f i + e"
+proof -
+  have "x < x + e"
+    using \<open>0<e\<close> \<open>x \<noteq> \<infinity>\<close> by (cases x) auto
+  also have "x + e = (SUP i : A. f i + e)"
+    unfolding SUP ennreal_SUP_add_left[OF \<open>A \<noteq> {}\<close>] ..
+  finally show ?thesis
+    unfolding less_SUP_iff .
+qed
+
+lemma ennreal_approx_SUP:
+  fixes x::ennreal
+  assumes f_bound: "\<And>i. i \<in> A \<Longrightarrow> f i \<le> x"
+  assumes approx: "\<And>e. (e::real) > 0 \<Longrightarrow> \<exists>i \<in> A. x \<le> f i + e"
+  shows "x = (SUP i : A. f i)"
+proof (rule antisym)
+  show "x \<le> (SUP i:A. f i)"
+  proof (rule ennreal_le_epsilon)
+    fix e :: real assume "0 < e"
+    from approx[OF this] guess i ..
+    then have "x \<le> f i + e"
+      by simp
+    also have "\<dots> \<le> (SUP i:A. f i) + e"
+      by (intro add_mono \<open>i \<in> A\<close> SUP_upper order_refl)
+    finally show "x \<le> (SUP i:A. f i) + e" .
+  qed
+qed (intro SUP_least f_bound)
+
+lemma ennreal_approx_INF:
+  fixes x::ennreal
+  assumes f_bound: "\<And>i. i \<in> A \<Longrightarrow> x \<le> f i"
+  assumes approx: "\<And>e. (e::real) > 0 \<Longrightarrow> \<exists>i \<in> A. f i \<le> x + e"
+  shows "x = (INF i : A. f i)"
+proof (rule antisym)
+  show "(INF i:A. f i) \<le> x"
+  proof (rule ennreal_le_epsilon)
+    fix e :: real assume "0 < e"
+    from approx[OF this] guess i .. note i = this
+    then have "(INF i:A. f i) \<le> f i"
+      by (intro INF_lower)
+    also have "\<dots> \<le> x + e"
+      by fact
+    finally show "(INF i:A. f i) \<le> x + e" .
+  qed
+qed (intro INF_greatest f_bound)
+
+lemma ennreal_approx_unit:
+  "(\<And>a::ennreal. 0 < a \<Longrightarrow> a < 1 \<Longrightarrow> a * z \<le> y) \<Longrightarrow> z \<le> y"
+  apply (subst SUP_mult_right_ennreal[of "\<lambda>x. x" "{0 <..< 1}" z, simplified])
+  apply (rule SUP_least)
+  apply auto
+  done
+
+lemma suminf_ennreal2:
+  "(\<And>i. 0 \<le> f i) \<Longrightarrow> summable f \<Longrightarrow> (\<Sum>i. ennreal (f i)) = ennreal (\<Sum>i. f i)"
+  using suminf_ennreal_eq by blast
+
+lemma less_top_ennreal: "x < top \<longleftrightarrow> (\<exists>r\<ge>0. x = ennreal r)"
+  by (cases x) auto
+
+lemma tendsto_top_iff_ennreal:
+  fixes f :: "'a \<Rightarrow> ennreal"
+  shows "(f \<longlongrightarrow> top) F \<longleftrightarrow> (\<forall>l\<ge>0. eventually (\<lambda>x. ennreal l < f x) F)"
+  by (auto simp: less_top_ennreal order_tendsto_iff )
+
+lemma ennreal_tendsto_top_eq_at_top:
+  "((\<lambda>z. ennreal (f z)) \<longlongrightarrow> top) F \<longleftrightarrow> (LIM z F. f z :> at_top)"
+  unfolding filterlim_at_top_dense tendsto_top_iff_ennreal
+  apply (auto simp: ennreal_less_iff)
+  subgoal for y
+    by (auto elim!: eventually_mono allE[of _ "max 0 y"])
+  done
+
+lemma tendsto_0_if_Limsup_eq_0_ennreal:
+  fixes f :: "_ \<Rightarrow> ennreal"
+  shows "Limsup F f = 0 \<Longrightarrow> (f \<longlongrightarrow> 0) F"
+  using Liminf_le_Limsup[of F f] tendsto_iff_Liminf_eq_Limsup[of F f 0]
+  by (cases "F = bot") auto
+
+lemma diff_le_self_ennreal[simp]: "a - b \<le> (a::ennreal)"
+  by (cases a b rule: ennreal2_cases) (auto simp: ennreal_minus)
+
+lemma ennreal_ineq_diff_add: "b \<le> a \<Longrightarrow> a = b + (a - b::ennreal)"
+  by transfer (auto simp: ereal_diff_positive max.absorb2 ereal_ineq_diff_add)
+
+lemma ennreal_mult_strict_left_mono: "(a::ennreal) < c \<Longrightarrow> 0 < b \<Longrightarrow> b < top \<Longrightarrow> b * a < b * c"
+  by transfer (auto intro!: ereal_mult_strict_left_mono)
+
+lemma ennreal_between: "0 < e \<Longrightarrow> 0 < x \<Longrightarrow> x < top \<Longrightarrow> x - e < (x::ennreal)"
+  by transfer (auto intro!: ereal_between)
+
+lemma minus_less_iff_ennreal: "b < top \<Longrightarrow> b \<le> a \<Longrightarrow> a - b < c \<longleftrightarrow> a < c + (b::ennreal)"
+  by transfer
+     (auto simp: top_ereal_def ereal_minus_less le_less)
+
+lemma tendsto_zero_ennreal:
+  assumes ev: "\<And>r. 0 < r \<Longrightarrow> \<forall>\<^sub>F x in F. f x < ennreal r"
+  shows "(f \<longlongrightarrow> 0) F"
+proof (rule order_tendstoI)
+  fix e::ennreal assume "e > 0"
+  obtain e'::real where "e' > 0" "ennreal e' < e"
+    using `0 < e` dense[of 0 "if e = top then 1 else (enn2real e)"]
+    by (cases e) (auto simp: ennreal_less_iff)
+  from ev[OF \<open>e' > 0\<close>] show "\<forall>\<^sub>F x in F. f x < e"
+    by eventually_elim (insert \<open>ennreal e' < e\<close>, auto)
+qed simp
+
+lifting_update ennreal.lifting
+lifting_forget ennreal.lifting
+
+end
--- a/src/HOL/Library/Extended_Real.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Library/Extended_Real.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -3547,7 +3547,7 @@
   using Liminf_le_Limsup[OF assms, of f]
   by auto
 
-lemma convergent_ereal:
+lemma convergent_ereal: -- \<open>RENAME\<close>
   fixes X :: "nat \<Rightarrow> 'a :: {complete_linorder,linorder_topology}"
   shows "convergent X \<longleftrightarrow> limsup X = liminf X"
   using tendsto_iff_Liminf_eq_Limsup[of sequentially]
--- a/src/HOL/Library/Liminf_Limsup.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Library/Liminf_Limsup.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -555,4 +555,14 @@
     using \<open>subseq r\<close> by auto
 qed
 
+lemma tendsto_Limsup:
+  fixes f :: "_ \<Rightarrow> 'a :: {complete_linorder,linorder_topology}"
+  shows "F \<noteq> bot \<Longrightarrow> Limsup F f = Liminf F f \<Longrightarrow> (f \<longlongrightarrow> Limsup F f) F"
+  by (subst tendsto_iff_Liminf_eq_Limsup) auto
+
+lemma tendsto_Liminf:
+  fixes f :: "_ \<Rightarrow> 'a :: {complete_linorder,linorder_topology}"
+  shows "F \<noteq> bot \<Longrightarrow> Limsup F f = Liminf F f \<Longrightarrow> (f \<longlongrightarrow> Liminf F f) F"
+  by (subst tendsto_iff_Liminf_eq_Limsup) auto
+
 end
--- a/src/HOL/Probability/Binary_Product_Measure.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Binary_Product_Measure.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -110,8 +110,8 @@
   shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
   using measurable_Pair[OF assms] by simp
 
-lemma 
-  assumes f[measurable]: "f \<in> measurable M (N \<Otimes>\<^sub>M P)" 
+lemma
+  assumes f[measurable]: "f \<in> measurable M (N \<Otimes>\<^sub>M P)"
   shows measurable_fst': "(\<lambda>x. fst (f x)) \<in> measurable M N"
     and measurable_snd': "(\<lambda>x. snd (f x)) \<in> measurable M P"
   by simp_all
@@ -134,7 +134,7 @@
     finally have "a \<times> b \<in> sets (Sup_sigma {?fst, ?snd})" . }
   moreover have "sets ?fst \<subseteq> sets (A \<Otimes>\<^sub>M B)"
     by (rule sets_image_in_sets) (auto simp: space_pair_measure[symmetric])
-  moreover have "sets ?snd \<subseteq> sets (A \<Otimes>\<^sub>M B)"  
+  moreover have "sets ?snd \<subseteq> sets (A \<Otimes>\<^sub>M B)"
     by (rule sets_image_in_sets) (auto simp: space_pair_measure)
   ultimately show ?thesis
     by (intro antisym[of "sets A" for A] sets_Sup_in_sets sets_pair_in_sets )
@@ -143,7 +143,7 @@
 
 lemma measurable_pair_iff:
   "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2) \<longleftrightarrow> (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
-  by (auto intro: measurable_pair[of f M M1 M2]) 
+  by (auto intro: measurable_pair[of f M M1 M2])
 
 lemma measurable_split_conv:
   "(\<lambda>(x, y). f x y) \<in> measurable A B \<longleftrightarrow> (\<lambda>x. f (fst x) (snd x)) \<in> measurable A B"
@@ -190,7 +190,7 @@
   shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M"
   using measurable_comp[OF measurable_Pair1' f, OF x]
   by (simp add: comp_def)
-  
+
 lemma measurable_Pair1:
   assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and y: "y \<in> space M2"
   shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M"
@@ -279,7 +279,7 @@
   shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator X (x, y) \<partial>M \<partial>N)" (is "_ = ?\<mu> X")
 proof (rule emeasure_measure_of[OF pair_measure_def])
   show "positive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>"
-    by (auto simp: positive_def nn_integral_nonneg)
+    by (auto simp: positive_def)
   have eq[simp]: "\<And>A x y. indicator A (x, y) = indicator (Pair x -` A) y"
     by (auto simp: indicator_def)
   show "countably_additive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>"
@@ -291,7 +291,7 @@
     moreover have "\<And>x. range (\<lambda>i. Pair x -` F i) \<subseteq> sets M"
       using F by (auto simp: sets_Pair1)
     ultimately show "(\<Sum>n. ?\<mu> (F n)) = ?\<mu> (\<Union>i. F i)"
-      by (auto simp add: nn_integral_suminf[symmetric] vimage_UN suminf_emeasure emeasure_nonneg
+      by (auto simp add: nn_integral_suminf[symmetric] vimage_UN suminf_emeasure
                intro!: nn_integral_cong nn_integral_indicator[symmetric])
   qed
   show "{a \<times> b |a b. a \<in> sets N \<and> b \<in> sets M} \<subseteq> Pow (space N \<times> space M)"
@@ -315,7 +315,7 @@
   have "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = (\<integral>\<^sup>+x. emeasure M B * indicator A x \<partial>N)"
     using A B by (auto intro!: nn_integral_cong simp: emeasure_pair_measure_alt)
   also have "\<dots> = emeasure M B * emeasure N A"
-    using A by (simp add: emeasure_nonneg nn_integral_cmult_indicator)
+    using A by (simp add: nn_integral_cmult_indicator)
   finally show ?thesis
     by (simp add: ac_simps)
 qed
@@ -373,9 +373,8 @@
   next
     fix i
     from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto
-    with F1 F2 emeasure_nonneg[of M1 "F1 i"] emeasure_nonneg[of M2 "F2 i"]
-    show "emeasure (M1 \<Otimes>\<^sub>M M2) (F1 i \<times> F2 i) \<noteq> \<infinity>"
-      by (auto simp add: emeasure_pair_measure_Times)
+    with F1 F2 show "emeasure (M1 \<Otimes>\<^sub>M M2) (F1 i \<times> F2 i) \<noteq> \<infinity>"
+      by (auto simp add: emeasure_pair_measure_Times ennreal_mult_eq_top_iff)
   qed
 qed
 
@@ -386,8 +385,7 @@
   ultimately show
     "\<exists>A. countable A \<and> A \<subseteq> sets (M1 \<Otimes>\<^sub>M M2) \<and> \<Union>A = space (M1 \<Otimes>\<^sub>M M2) \<and> (\<forall>a\<in>A. emeasure (M1 \<Otimes>\<^sub>M M2) a \<noteq> \<infinity>)"
     by (intro exI[of _ "(\<lambda>(a, b). a \<times> b) ` (F1 \<times> F2)"] conjI)
-       (auto simp: M2.emeasure_pair_measure_Times space_pair_measure set_eq_iff subset_eq
-             dest: sets.sets_into_space)
+       (auto simp: M2.emeasure_pair_measure_Times space_pair_measure set_eq_iff subset_eq ennreal_mult_eq_top_iff)
 qed
 
 lemma sigma_finite_pair_measure:
@@ -455,9 +453,9 @@
   proof (rule AE_I)
     from N measurable_emeasure_Pair1[OF \<open>N \<in> sets (M1 \<Otimes>\<^sub>M M2)\<close>]
     show "emeasure M1 {x\<in>space M1. emeasure M2 (Pair x -` N) \<noteq> 0} = 0"
-      by (auto simp: M2.emeasure_pair_measure_alt nn_integral_0_iff emeasure_nonneg)
+      by (auto simp: M2.emeasure_pair_measure_alt nn_integral_0_iff)
     show "{x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0} \<in> sets M1"
-      by (intro borel_measurable_ereal_neq_const measurable_emeasure_Pair1 N)
+      by (intro borel_measurable_eq measurable_emeasure_Pair1 N sets.sets_Collect_neg N) simp
     { fix x assume "x \<in> space M1" "emeasure M2 (Pair x -` N) = 0"
       have "AE y in M2. Q (x, y)"
       proof (rule AE_I)
@@ -523,8 +521,8 @@
   "x \<in> space M1 \<Longrightarrow> g \<in> measurable (M1 \<Otimes>\<^sub>M M2) L \<Longrightarrow> (\<lambda>y. g (x, y)) \<in> measurable M2 L"
   by simp
 
-lemma (in sigma_finite_measure) borel_measurable_nn_integral_fst':
-  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)" "\<And>x. 0 \<le> f x"
+lemma (in sigma_finite_measure) borel_measurable_nn_integral_fst:
+  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)"
   shows "(\<lambda>x. \<integral>\<^sup>+ y. f (x, y) \<partial>M) \<in> borel_measurable M1"
 using f proof induct
   case (cong u v)
@@ -547,8 +545,8 @@
                    nn_integral_monotone_convergence_SUP incseq_def le_fun_def
               cong: measurable_cong)
 
-lemma (in sigma_finite_measure) nn_integral_fst':
-  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)" "\<And>x. 0 \<le> f x"
+lemma (in sigma_finite_measure) nn_integral_fst:
+  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)"
   shows "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>M \<partial>M1) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M) f" (is "?I f = _")
 using f proof induct
   case (cong u v)
@@ -557,26 +555,16 @@
   with cong show ?case
     by (simp cong: nn_integral_cong)
 qed (simp_all add: emeasure_pair_measure nn_integral_cmult nn_integral_add
-                   nn_integral_monotone_convergence_SUP
-                   measurable_compose_Pair1 nn_integral_nonneg
-                   borel_measurable_nn_integral_fst' nn_integral_mono incseq_def le_fun_def
+                   nn_integral_monotone_convergence_SUP measurable_compose_Pair1
+                   borel_measurable_nn_integral_fst nn_integral_mono incseq_def le_fun_def
               cong: nn_integral_cong)
 
-lemma (in sigma_finite_measure) nn_integral_fst:
-  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)"
-  shows "(\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x, y) \<partial>M) \<partial>M1) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M) f"
-  using f
-    borel_measurable_nn_integral_fst'[of "\<lambda>x. max 0 (f x)"]
-    nn_integral_fst'[of "\<lambda>x. max 0 (f x)"]
-  unfolding nn_integral_max_0 by auto
-
 lemma (in sigma_finite_measure) borel_measurable_nn_integral[measurable (raw)]:
   "case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M) \<Longrightarrow> (\<lambda>x. \<integral>\<^sup>+ y. f x y \<partial>M) \<in> borel_measurable N"
-  using borel_measurable_nn_integral_fst'[of "\<lambda>x. max 0 (case_prod f x)" N]
-  by (simp add: nn_integral_max_0)
+  using borel_measurable_nn_integral_fst[of "case_prod f" N] by simp
 
 lemma (in pair_sigma_finite) nn_integral_snd:
-  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
+  assumes f[measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
   shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M2) f"
 proof -
   note measurable_pair_swap[OF f]
@@ -584,8 +572,7 @@
   have "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1))"
     by simp
   also have "(\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M2) f"
-    by (subst distr_pair_swap)
-       (auto simp: nn_integral_distr[OF measurable_pair_swap' f] intro!: nn_integral_cong)
+    by (subst distr_pair_swap) (auto simp add: nn_integral_distr intro!: nn_integral_cong)
   finally show ?thesis .
 qed
 
@@ -610,13 +597,13 @@
   have [simp]: "snd \<in> X \<times> Y \<rightarrow> Y" "fst \<in> X \<times> Y \<rightarrow> X"
     by auto
   let ?XY = "{{fst -` a \<inter> X \<times> Y | a. a \<in> A}, {snd -` b \<inter> X \<times> Y | b. b \<in> B}}"
-  have "sets ?P = 
+  have "sets ?P =
     sets (\<Squnion>\<^sub>\<sigma> xy\<in>?XY. sigma (X \<times> Y) xy)"
     by (simp add: vimage_algebra_sigma sets_pair_eq_sets_fst_snd A B)
   also have "\<dots> = sets (sigma (X \<times> Y) (\<Union>?XY))"
     by (intro Sup_sigma_sigma arg_cong[where f=sets]) auto
   also have "\<dots> = sets ?S"
-  proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI) 
+  proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI)
     show "\<Union>?XY \<subseteq> Pow (X \<times> Y)" "{a \<times> b |a b. a \<in> A \<and> b \<in> B} \<subseteq> Pow (X \<times> Y)"
       using A B by auto
   next
@@ -710,20 +697,17 @@
   with A B have fin_Pair: "\<And>x. finite (Pair x -` X)"
     by (intro finite_subset[OF _ B]) auto
   have fin_X: "finite X" using X_subset by (rule finite_subset) (auto simp: A B)
+  have pos_card: "(0::ennreal) < of_nat (card (Pair x -` X)) \<longleftrightarrow> Pair x -` X \<noteq> {}" for x
+    by (auto simp: card_eq_0_iff fin_Pair) blast
+
   show "emeasure ?P X = emeasure ?C X"
+    using X_subset A fin_Pair fin_X
     apply (subst B.emeasure_pair_measure_alt[OF X])
     apply (subst emeasure_count_space)
-    using X_subset apply auto []
-    apply (simp add: fin_Pair emeasure_count_space X_subset fin_X)
-    apply (subst nn_integral_count_space)
-    using A apply simp
-    apply (simp del: of_nat_setsum add: of_nat_setsum[symmetric])
-    apply (subst card_gt_0_iff)
-    apply (simp add: fin_Pair)
-    apply (subst card_SigmaI[symmetric])
-    using A apply simp
-    using fin_Pair apply simp
-    using X_subset apply (auto intro!: arg_cong[where f=card])
+    apply (auto simp add: emeasure_count_space nn_integral_count_space
+                          pos_card of_nat_setsum[symmetric] card_SigmaI[symmetric]
+                simp del: of_nat_setsum card_SigmaI
+                intro!: arg_cong[where f=card])
     done
 qed
 
@@ -732,16 +716,15 @@
   assumes A: "A \<in> sets (count_space UNIV \<Otimes>\<^sub>M M)" (is "A \<in> sets (?A \<Otimes>\<^sub>M ?B)")
   shows "emeasure (?A \<Otimes>\<^sub>M ?B) A = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator A (x, y) \<partial>?B \<partial>?A)"
   by (rule emeasure_measure_of[OF pair_measure_def])
-     (auto simp: countably_additive_def positive_def suminf_indicator nn_integral_nonneg A
+     (auto simp: countably_additive_def positive_def suminf_indicator A
                  nn_integral_suminf[symmetric] dest: sets.sets_into_space)
 
 lemma emeasure_prod_count_space_single[simp]: "emeasure (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) {x} = 1"
 proof -
-  have [simp]: "\<And>a b x y. indicator {(a, b)} (x, y) = (indicator {a} x * indicator {b} y::ereal)"
+  have [simp]: "\<And>a b x y. indicator {(a, b)} (x, y) = (indicator {a} x * indicator {b} y::ennreal)"
     by (auto split: split_indicator)
   show ?thesis
-    by (cases x)
-       (auto simp: emeasure_prod_count_space nn_integral_cmult sets_Pair nn_integral_max_0 one_ereal_def[symmetric])
+    by (cases x) (auto simp: emeasure_prod_count_space nn_integral_cmult sets_Pair)
 qed
 
 lemma emeasure_count_space_prod_eq:
@@ -771,17 +754,16 @@
     also have "emeasure (?A \<Otimes>\<^sub>M ?B) C = emeasure (count_space UNIV) C"
       using \<open>countable C\<close> by (rule *)
     finally show ?thesis
-      using \<open>infinite C\<close> \<open>infinite A\<close> by simp
+      using \<open>infinite C\<close> \<open>infinite A\<close> by (simp add: top_unique)
   qed
 qed
 
-lemma nn_intergal_count_space_prod_eq':
-  assumes [simp]: "\<And>x. 0 \<le> f x"
-  shows "nn_integral (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) f = nn_integral (count_space UNIV) f"
+lemma nn_integral_count_space_prod_eq:
+  "nn_integral (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) f = nn_integral (count_space UNIV) f"
     (is "nn_integral ?P f = _")
 proof cases
   assume cntbl: "countable {x. f x \<noteq> 0}"
-  have [simp]: "\<And>x. ereal (real (card ({x} \<inter> {x. f x \<noteq> 0}))) = indicator {x. f x \<noteq> 0} x"
+  have [simp]: "\<And>x. card ({x} \<inter> {x. f x \<noteq> 0}) = (indicator {x. f x \<noteq> 0} x::ennreal)"
     by (auto split: split_indicator)
   have [measurable]: "\<And>y. (\<lambda>x. indicator {y} x) \<in> borel_measurable ?P"
     by (rule measurable_discrete_difference[of "\<lambda>x. 0" _ borel "{y}" "\<lambda>x. indicator {y} x" for y])
@@ -799,9 +781,9 @@
     by (auto simp add: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator)
 next
   { fix x assume "f x \<noteq> 0"
-    with \<open>0 \<le> f x\<close> have "(\<exists>r. 0 < r \<and> f x = ereal r) \<or> f x = \<infinity>"
-      by (cases "f x") (auto simp: less_le)
-    then have "\<exists>n. ereal (1 / real (Suc n)) \<le> f x"
+    then have "(\<exists>r\<ge>0. 0 < r \<and> f x = ennreal r) \<or> f x = \<infinity>"
+      by (cases "f x" rule: ennreal_cases) (auto simp: less_le)
+    then have "\<exists>n. ennreal (1 / real (Suc n)) \<le> f x"
       by (auto elim!: nat_approx_posE intro!: less_imp_le) }
   note * = this
 
@@ -816,25 +798,21 @@
   have [measurable]: "C \<in> sets ?P"
     using sets.countable[OF _ \<open>countable C\<close>, of ?P] by (auto simp: sets_Pair)
 
-  have "(\<integral>\<^sup>+x. ereal (1/Suc n) * indicator C x \<partial>?P) \<le> nn_integral ?P f"
+  have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>?P) \<le> nn_integral ?P f"
     using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric])
-  moreover have "(\<integral>\<^sup>+x. ereal (1/Suc n) * indicator C x \<partial>?P) = \<infinity>"
-    using \<open>infinite C\<close> by (simp add: nn_integral_cmult emeasure_count_space_prod_eq)
-  moreover have "(\<integral>\<^sup>+x. ereal (1/Suc n) * indicator C x \<partial>count_space UNIV) \<le> nn_integral (count_space UNIV) f"
+  moreover have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>?P) = \<infinity>"
+    using \<open>infinite C\<close> by (simp add: nn_integral_cmult emeasure_count_space_prod_eq ennreal_mult_top)
+  moreover have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>count_space UNIV) \<le> nn_integral (count_space UNIV) f"
     using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric])
-  moreover have "(\<integral>\<^sup>+x. ereal (1/Suc n) * indicator C x \<partial>count_space UNIV) = \<infinity>"
-    using \<open>infinite C\<close> by (simp add: nn_integral_cmult)
+  moreover have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>count_space UNIV) = \<infinity>"
+    using \<open>infinite C\<close> by (simp add: nn_integral_cmult ennreal_mult_top)
   ultimately show ?thesis
-    by simp
+    by (simp add: top_unique)
 qed
 
-lemma nn_intergal_count_space_prod_eq:
-  "nn_integral (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) f = nn_integral (count_space UNIV) f"
-  by (subst (1 2) nn_integral_max_0[symmetric]) (auto intro!: nn_intergal_count_space_prod_eq')
-
 lemma pair_measure_density:
-  assumes f: "f \<in> borel_measurable M1" "AE x in M1. 0 \<le> f x"
-  assumes g: "g \<in> borel_measurable M2" "AE x in M2. 0 \<le> g x"
+  assumes f: "f \<in> borel_measurable M1"
+  assumes g: "g \<in> borel_measurable M2"
   assumes "sigma_finite_measure M2" "sigma_finite_measure (density M2 g)"
   shows "density M1 f \<Otimes>\<^sub>M density M2 g = density (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x,y). f x * g y)" (is "?L = ?R")
 proof (rule measure_eqI)
@@ -916,7 +894,7 @@
       by simp
   qed
 qed
-  
+
 lemma sets_pair_countable:
   assumes "countable S1" "countable S2"
   assumes M: "sets M = Pow S1" and N: "sets N = Pow S2"
@@ -948,25 +926,22 @@
     by (subst sets_pair_countable[OF assms]) auto
 next
   fix A B assume "A \<in> sets (count_space S1)" "B \<in> sets (count_space S2)"
-  then show "emeasure (count_space S1) A * emeasure (count_space S2) B = 
+  then show "emeasure (count_space S1) A * emeasure (count_space S2) B =
     emeasure (count_space (S1 \<times> S2)) (A \<times> B)"
-    by (subst (1 2 3) emeasure_count_space) (auto simp: finite_cartesian_product_iff)
+    by (subst (1 2 3) emeasure_count_space) (auto simp: finite_cartesian_product_iff ennreal_mult_top ennreal_top_mult)
 qed
 
-lemma nn_integral_fst_count_space':
-  assumes nonneg: "\<And>xy. 0 \<le> f xy"
-  shows "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f"
+lemma nn_integral_fst_count_space:
+  "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f"
   (is "?lhs = ?rhs")
 proof(cases)
   assume *: "countable {xy. f xy \<noteq> 0}"
   let ?A = "fst ` {xy. f xy \<noteq> 0}"
   let ?B = "snd ` {xy. f xy \<noteq> 0}"
   from * have [simp]: "countable ?A" "countable ?B" by(rule countable_image)+
-  from nonneg have f_neq_0: "\<And>xy. f xy \<noteq> 0 \<longleftrightarrow> f xy > 0"
-    by(auto simp add: order.order_iff_strict)
   have "?lhs = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space UNIV \<partial>count_space ?A)"
     by(rule nn_integral_count_space_eq)
-      (auto simp add: f_neq_0 nn_integral_0_iff_AE AE_count_space not_le intro: rev_image_eqI)
+      (auto simp add: nn_integral_0_iff_AE AE_count_space not_le intro: rev_image_eqI)
   also have "\<dots> = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space ?B \<partial>count_space ?A)"
     by(intro nn_integral_count_space_eq nn_integral_cong)(auto intro: rev_image_eqI)
   also have "\<dots> = (\<integral>\<^sup>+ xy. f xy \<partial>count_space (?A \<times> ?B))"
@@ -977,9 +952,9 @@
   finally show ?thesis .
 next
   { fix xy assume "f xy \<noteq> 0"
-    with \<open>0 \<le> f xy\<close> have "(\<exists>r. 0 < r \<and> f xy = ereal r) \<or> f xy = \<infinity>"
-      by (cases "f xy") (auto simp: less_le)
-    then have "\<exists>n. ereal (1 / real (Suc n)) \<le> f xy"
+    then have "(\<exists>r\<ge>0. 0 < r \<and> f xy = ennreal r) \<or> f xy = \<infinity>"
+      by (cases "f xy" rule: ennreal_cases) (auto simp: less_le)
+    then have "\<exists>n. ennreal (1 / real (Suc n)) \<le> f xy"
       by (auto elim!: nat_approx_posE intro!: less_imp_le) }
   note * = this
 
@@ -991,15 +966,15 @@
   then obtain C where C: "C \<subseteq> {xy. 1/Suc n \<le> f xy}" and "countable C" "infinite C"
     by (metis infinite_countable_subset')
 
-  have "\<infinity> = (\<integral>\<^sup>+ xy. ereal (1 / Suc n) * indicator C xy \<partial>count_space UNIV)"
-    using \<open>infinite C\<close> by(simp add: nn_integral_cmult)
+  have "\<infinity> = (\<integral>\<^sup>+ xy. ennreal (1 / Suc n) * indicator C xy \<partial>count_space UNIV)"
+    using \<open>infinite C\<close> by(simp add: nn_integral_cmult ennreal_mult_top)
   also have "\<dots> \<le> ?rhs" using C
-    by(intro nn_integral_mono)(auto split: split_indicator simp add: nonneg)
-  finally have "?rhs = \<infinity>" by simp
+    by(intro nn_integral_mono)(auto split: split_indicator)
+  finally have "?rhs = \<infinity>" by (simp add: top_unique)
   moreover have "?lhs = \<infinity>"
   proof(cases "finite (fst ` C)")
     case True
-    then obtain x C' where x: "x \<in> fst ` C" 
+    then obtain x C' where x: "x \<in> fst ` C"
       and C': "C' = fst -` {x} \<inter> C"
       and "infinite C'"
       using \<open>infinite C\<close> by(auto elim!: inf_img_fin_domE')
@@ -1007,37 +982,33 @@
 
     from C' \<open>infinite C'\<close> have "infinite (snd ` C')"
       by(auto dest!: finite_imageD simp add: inj_on_def)
-    then have "\<infinity> = (\<integral>\<^sup>+ y. ereal (1 / Suc n) * indicator (snd ` C') y \<partial>count_space UNIV)"
-      by(simp add: nn_integral_cmult)
-    also have "\<dots> = (\<integral>\<^sup>+ y. ereal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV)"
+    then have "\<infinity> = (\<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator (snd ` C') y \<partial>count_space UNIV)"
+      by(simp add: nn_integral_cmult ennreal_mult_top)
+    also have "\<dots> = (\<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV)"
       by(rule nn_integral_cong)(force split: split_indicator intro: rev_image_eqI simp add: C')
-    also have "\<dots> = (\<integral>\<^sup>+ x'. (\<integral>\<^sup>+ y. ereal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV) * indicator {x} x' \<partial>count_space UNIV)"
-      by(simp add: one_ereal_def[symmetric] nn_integral_nonneg max_def)
-    also have "\<dots> \<le> (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ereal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV \<partial>count_space UNIV)"
-      by(rule nn_integral_mono)(simp split: split_indicator add: nn_integral_nonneg)
+    also have "\<dots> = (\<integral>\<^sup>+ x'. (\<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV) * indicator {x} x' \<partial>count_space UNIV)"
+      by(simp add: one_ereal_def[symmetric])
+    also have "\<dots> \<le> (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV \<partial>count_space UNIV)"
+      by(rule nn_integral_mono)(simp split: split_indicator)
     also have "\<dots> \<le> ?lhs" using **
-      by(intro nn_integral_mono)(auto split: split_indicator simp add: nonneg)
-    finally show ?thesis by simp
+      by(intro nn_integral_mono)(auto split: split_indicator)
+    finally show ?thesis by (simp add: top_unique)
   next
     case False
     def C' \<equiv> "fst ` C"
-    have "\<infinity> = \<integral>\<^sup>+ x. ereal (1 / Suc n) * indicator C' x \<partial>count_space UNIV"
-      using C'_def False by(simp add: nn_integral_cmult)
-    also have "\<dots> = \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ereal (1 / Suc n) * indicator C' x * indicator {SOME y. (x, y) \<in> C} y \<partial>count_space UNIV \<partial>count_space UNIV"
+    have "\<infinity> = \<integral>\<^sup>+ x. ennreal (1 / Suc n) * indicator C' x \<partial>count_space UNIV"
+      using C'_def False by(simp add: nn_integral_cmult ennreal_mult_top)
+    also have "\<dots> = \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' x * indicator {SOME y. (x, y) \<in> C} y \<partial>count_space UNIV \<partial>count_space UNIV"
       by(auto simp add: one_ereal_def[symmetric] max_def intro: nn_integral_cong)
-    also have "\<dots> \<le> \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ereal (1 / Suc n) * indicator C (x, y) \<partial>count_space UNIV \<partial>count_space UNIV"
+    also have "\<dots> \<le> \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C (x, y) \<partial>count_space UNIV \<partial>count_space UNIV"
       by(intro nn_integral_mono)(auto simp add: C'_def split: split_indicator intro: someI)
     also have "\<dots> \<le> ?lhs" using C
-      by(intro nn_integral_mono)(auto split: split_indicator simp add: nonneg)
-    finally show ?thesis by simp
+      by(intro nn_integral_mono)(auto split: split_indicator)
+    finally show ?thesis by (simp add: top_unique)
   qed
   ultimately show ?thesis by simp
 qed
 
-lemma nn_integral_fst_count_space:
-  "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f"
-by(subst (2 3) nn_integral_max_0[symmetric])(rule nn_integral_fst_count_space', simp)
-
 lemma nn_integral_snd_count_space:
   "(\<integral>\<^sup>+ y. \<integral>\<^sup>+ x. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f"
   (is "?lhs = ?rhs")
@@ -1118,7 +1089,7 @@
   interpret M: finite_measure M by fact
   interpret N: finite_measure N by fact
   show "emeasure (N  \<Otimes>\<^sub>M M) (space (N  \<Otimes>\<^sub>M M)) \<noteq> \<infinity>"
-    by (auto simp: space_pair_measure M.emeasure_pair_measure_Times)
+    by (auto simp: space_pair_measure M.emeasure_pair_measure_Times ennreal_mult_eq_top_iff)
 qed
 
 end
--- a/src/HOL/Probability/Bochner_Integration.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Bochner_Integration.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -36,9 +36,9 @@
 
   def A \<equiv> "\<lambda>m n. {x\<in>space M. dist (f x) (e n) < 1 / (Suc m) \<and> 1 / (Suc m) \<le> dist (f x) z}"
   def B \<equiv> "\<lambda>m. disjointed (A m)"
-  
+
   def m \<equiv> "\<lambda>N x. Max {m::nat. m \<le> N \<and> x \<in> (\<Union>n\<le>N. B m n)}"
-  def F \<equiv> "\<lambda>N::nat. \<lambda>x. if (\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)) \<and> (\<exists>n\<le>N. x \<in> B (m N x) n) 
+  def F \<equiv> "\<lambda>N::nat. \<lambda>x. if (\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)) \<and> (\<exists>n\<le>N. x \<in> B (m N x) n)
     then e (LEAST n. x \<in> B (m N x) n) else z"
 
   have B_imp_A[intro, simp]: "\<And>x m n. x \<in> B m n \<Longrightarrow> x \<in> A m n"
@@ -75,13 +75,13 @@
       { fix n x assume "x \<in> B (m i x) n"
         then have "(LEAST n. x \<in> B (m i x) n) \<le> n"
           by (intro Least_le)
-        also assume "n \<le> i" 
+        also assume "n \<le> i"
         finally have "(LEAST n. x \<in> B (m i x) n) \<le> i" . }
       then have "F i ` space M \<subseteq> {z} \<union> e ` {.. i}"
         by (auto simp: F_def)
       then show "finite (F i ` space M)"
         by (rule finite_subset) auto }
-    
+
     { fix N i n x assume "i \<le> N" "n \<le> N" "x \<in> B i n"
       then have 1: "\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)" by auto
       from m[OF this] obtain n where n: "m N x \<le> N" "n \<le> N" "x \<in> B (m N x) n" by auto
@@ -137,7 +137,7 @@
         qed
       qed
     qed
-    fix i 
+    fix i
     { fix n m assume "x \<in> A n m"
       then have "dist (e m) (f x) + dist (f x) z \<le> 2 * dist (f x) z"
         unfolding A_def by (auto simp: dist_commute)
@@ -153,13 +153,6 @@
   qed
 qed
 
-lemma real_indicator: "real_of_ereal (indicator A x :: ereal) = indicator A x"
-  unfolding indicator_def by auto
-
-lemma split_indicator_asm:
-  "P (indicator S x) \<longleftrightarrow> \<not> ((x \<in> S \<and> \<not> P 1) \<or> (x \<notin> S \<and> \<not> P 0))"
-  unfolding indicator_def by auto
-
 lemma
   fixes f :: "'a \<Rightarrow> 'b::semiring_1" assumes "finite A"
   shows setsum_mult_indicator[simp]: "(\<Sum>x \<in> A. f x * indicator (B x) (g x)) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)"
@@ -176,40 +169,39 @@
   assumes seq: "\<And>U. (\<And>i. U i \<in> borel_measurable M) \<Longrightarrow> (\<And>i x. 0 \<le> U i x) \<Longrightarrow> (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. U i x) \<longlonglongrightarrow> u x) \<Longrightarrow> P u"
   shows "P u"
 proof -
-  have "(\<lambda>x. ereal (u x)) \<in> borel_measurable M" using u by auto
+  have "(\<lambda>x. ennreal (u x)) \<in> borel_measurable M" using u by auto
   from borel_measurable_implies_simple_function_sequence'[OF this]
-  obtain U where U: "\<And>i. simple_function M (U i)" "incseq U" "\<And>i. \<infinity> \<notin> range (U i)" and
-    sup: "\<And>x. (SUP i. U i x) = max 0 (ereal (u x))" and nn: "\<And>i x. 0 \<le> U i x"
+  obtain U where U: "\<And>i. simple_function M (U i)" "incseq U" "\<And>i x. U i x < top" and
+    sup: "\<And>x. (SUP i. U i x) = ennreal (u x)"
     by blast
 
-  def U' \<equiv> "\<lambda>i x. indicator (space M) x * real_of_ereal (U i x)"
+  def U' \<equiv> "\<lambda>i x. indicator (space M) x * enn2real (U i x)"
   then have U'_sf[measurable]: "\<And>i. simple_function M (U' i)"
-    using U by (auto intro!: simple_function_compose1[where g=real_of_ereal])
+    using U by (auto intro!: simple_function_compose1[where g=enn2real])
 
   show "P u"
   proof (rule seq)
-    fix i show "U' i \<in> borel_measurable M" "\<And>x. 0 \<le> U' i x"
-      using U nn by (auto
-          intro: borel_measurable_simple_function 
-          intro!: borel_measurable_real_of_ereal real_of_ereal_pos borel_measurable_times
-          simp: U'_def zero_le_mult_iff)
+    show U': "U' i \<in> borel_measurable M" "\<And>x. 0 \<le> U' i x" for i
+      using U by (auto
+          intro: borel_measurable_simple_function
+          intro!: borel_measurable_enn2real borel_measurable_times
+          simp: U'_def zero_le_mult_iff enn2real_nonneg)
     show "incseq U'"
-      using U(2,3) nn
-      by (auto simp: incseq_def le_fun_def image_iff eq_commute U'_def indicator_def
-               intro!: real_of_ereal_positive_mono)
-  next
+      using U(2,3)
+      by (auto simp: incseq_def le_fun_def image_iff eq_commute U'_def indicator_def enn2real_mono)
+
     fix x assume x: "x \<in> space M"
     have "(\<lambda>i. U i x) \<longlonglongrightarrow> (SUP i. U i x)"
       using U(2) by (intro LIMSEQ_SUP) (auto simp: incseq_def le_fun_def)
-    moreover have "(\<lambda>i. U i x) = (\<lambda>i. ereal (U' i x))"
-      using x nn U(3) by (auto simp: fun_eq_iff U'_def ereal_real image_iff eq_commute)
-    moreover have "(SUP i. U i x) = ereal (u x)"
+    moreover have "(\<lambda>i. U i x) = (\<lambda>i. ennreal (U' i x))"
+      using x U(3) by (auto simp: fun_eq_iff U'_def image_iff eq_commute)
+    moreover have "(SUP i. U i x) = ennreal (u x)"
       using sup u(2) by (simp add: max_def)
-    ultimately show "(\<lambda>i. U' i x) \<longlonglongrightarrow> u x" 
-      by simp
+    ultimately show "(\<lambda>i. U' i x) \<longlonglongrightarrow> u x"
+      using u U' by simp
   next
     fix i
-    have "U' i ` space M \<subseteq> real_of_ereal ` (U i ` space M)" "finite (U i ` space M)"
+    have "U' i ` space M \<subseteq> enn2real ` (U i ` space M)" "finite (U i ` space M)"
       unfolding U'_def using U(1) by (auto dest: simple_functionD)
     then have fin: "finite (U' i ` space M)"
       by (metis finite_subset finite_imageI)
@@ -218,7 +210,7 @@
     ultimately have U': "(\<lambda>z. \<Sum>y\<in>U' i`space M. y * indicator {x\<in>space M. U' i x = y} z) = U' i"
       by (simp add: U'_def fun_eq_iff)
     have "\<And>x. x \<in> U' i ` space M \<Longrightarrow> 0 \<le> x"
-      using nn by (auto simp: U'_def real_of_ereal_pos)
+      by (auto simp: U'_def enn2real_nonneg)
     with fin have "P (\<lambda>z. \<Sum>y\<in>U' i`space M. y * indicator {x\<in>space M. U' i x = y} z)"
     proof induct
       case empty from set[of "{}"] show ?case
@@ -227,7 +219,7 @@
       case (insert x F)
       then show ?case
         by (auto intro!: add mult set setsum_nonneg split: split_indicator split_indicator_asm
-                 simp del: setsum_mult_indicator simp: setsum_nonneg_eq_0_iff )
+                 simp del: setsum_mult_indicator simp: setsum_nonneg_eq_0_iff)
     qed
     with U' show "P (U' i)" by simp
   qed
@@ -257,14 +249,14 @@
     by (auto intro: measurable_simple_function)
 
   assume fin: "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>" "emeasure M {y \<in> space M. g y \<noteq> 0} \<noteq> \<infinity>"
-   
+
   have "emeasure M {x\<in>space M. p (f x) (g x) \<noteq> 0} \<le>
       emeasure M ({x\<in>space M. f x \<noteq> 0} \<union> {x\<in>space M. g x \<noteq> 0})"
     by (intro emeasure_mono) (auto simp: p_0)
   also have "\<dots> \<le> emeasure M {x\<in>space M. f x \<noteq> 0} + emeasure M {x\<in>space M. g x \<noteq> 0}"
     by (intro emeasure_subadditive) auto
   finally show "emeasure M {y \<in> space M. p (f y) (g y) \<noteq> 0} \<noteq> \<infinity>"
-    using fin by auto
+    using fin by (auto simp: top_unique)
 qed
 
 lemma simple_function_finite_support:
@@ -282,7 +274,7 @@
   then have m: "0 < m"
     using nn by (auto simp: less_le)
 
-  from m have "m * emeasure M {x\<in>space M. 0 \<noteq> f x} = 
+  from m have "m * emeasure M {x\<in>space M. 0 \<noteq> f x} =
     (\<integral>\<^sup>+x. m * indicator {x\<in>space M. 0 \<noteq> f x} x \<partial>M)"
     using f by (intro nn_integral_cmult_indicator[symmetric]) auto
   also have "\<dots> \<le> (\<integral>\<^sup>+x. f x \<partial>M)"
@@ -294,7 +286,7 @@
   qed
   also note \<open>\<dots> < \<infinity>\<close>
   finally show ?thesis
-    using m by auto 
+    using m by (auto simp: ennreal_mult_less_top)
 next
   assume "\<not> (\<exists>x\<in>space M. f x \<noteq> 0)"
   with nn have *: "{x\<in>space M. f x \<noteq> 0} = {}"
@@ -306,11 +298,11 @@
   assumes f: "simple_function M f" and fin: "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
   shows "simple_bochner_integrable M f"
 proof
-  have "emeasure M {y \<in> space M. ereal (norm (f y)) \<noteq> 0} \<noteq> \<infinity>"
+  have "emeasure M {y \<in> space M. ennreal (norm (f y)) \<noteq> 0} \<noteq> \<infinity>"
   proof (rule simple_function_finite_support)
-    show "simple_function M (\<lambda>x. ereal (norm (f x)))"
+    show "simple_function M (\<lambda>x. ennreal (norm (f x)))"
       using f by (rule simple_function_compose1)
-    show "(\<integral>\<^sup>+ y. ereal (norm (f y)) \<partial>M) < \<infinity>" by fact
+    show "(\<integral>\<^sup>+ y. ennreal (norm (f y)) \<partial>M) < \<infinity>" by fact
   qed simp
   then show "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>" by simp
 qed fact
@@ -340,12 +332,12 @@
   note eq = this
 
   have "simple_bochner_integral M f =
-    (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M. 
+    (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M.
       if \<exists>x\<in>space M. y = f x \<and> z = g x then measure M {x\<in>space M. g x = z} else 0) *\<^sub>R y)"
     unfolding simple_bochner_integral_def
   proof (safe intro!: setsum.cong scaleR_cong_right)
     fix y assume y: "y \<in> space M" "f y \<noteq> 0"
-    have [simp]: "g ` space M \<inter> {z. \<exists>x\<in>space M. f y = f x \<and> z = g x} = 
+    have [simp]: "g ` space M \<inter> {z. \<exists>x\<in>space M. f y = f x \<and> z = g x} =
         {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
       by auto
     have eq:"{x \<in> space M. f x = f y} =
@@ -361,16 +353,16 @@
       then have "emeasure M {y \<in> space M. g y = g x} \<le> emeasure M {y \<in> space M. f y \<noteq> 0}"
         by (auto intro!: emeasure_mono cong: sub)
       then have "emeasure M {xa \<in> space M. g xa = g x} < \<infinity>"
-        using f by (auto simp: simple_bochner_integrable.simps) }
+        using f by (auto simp: simple_bochner_integrable.simps less_top) }
     ultimately
     show "measure M {x \<in> space M. f x = f y} =
       (\<Sum>z\<in>g ` space M. if \<exists>x\<in>space M. f y = f x \<and> z = g x then measure M {x \<in> space M. g x = z} else 0)"
       apply (simp add: setsum.If_cases eq)
       apply (subst measure_finite_Union[symmetric])
-      apply (auto simp: disjoint_family_on_def)
+      apply (auto simp: disjoint_family_on_def less_top)
       done
   qed
-  also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M. 
+  also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M.
       if \<exists>x\<in>space M. y = f x \<and> z = g x then measure M {x\<in>space M. g x = z} *\<^sub>R y else 0))"
     by (auto intro!: setsum.cong simp: scaleR_setsum_left)
   also have "\<dots> = ?r"
@@ -435,11 +427,11 @@
   assumes f: "simple_bochner_integrable M f"
   shows "norm (simple_bochner_integral M f) \<le> simple_bochner_integral M (\<lambda>x. norm (f x))"
 proof -
-  have "norm (simple_bochner_integral M f) \<le> 
+  have "norm (simple_bochner_integral M f) \<le>
     (\<Sum>y\<in>f ` space M. norm (measure M {x \<in> space M. f x = y} *\<^sub>R y))"
     unfolding simple_bochner_integral_def by (rule norm_setsum)
   also have "\<dots> = (\<Sum>y\<in>f ` space M. measure M {x \<in> space M. f x = y} *\<^sub>R norm y)"
-    by (simp add: measure_nonneg)
+    by simp
   also have "\<dots> = simple_bochner_integral M (\<lambda>x. norm (f x))"
     using f
     by (intro simple_bochner_integral_partition[symmetric])
@@ -447,36 +439,41 @@
   finally show ?thesis .
 qed
 
+lemma simple_bochner_integral_nonneg[simp]:
+  fixes f :: "'a \<Rightarrow> real"
+  shows "(\<And>x. 0 \<le> f x) \<Longrightarrow> 0 \<le> simple_bochner_integral M f"
+  by (simp add: setsum_nonneg simple_bochner_integral_def)
+
 lemma simple_bochner_integral_eq_nn_integral:
   assumes f: "simple_bochner_integrable M f" "\<And>x. 0 \<le> f x"
   shows "simple_bochner_integral M f = (\<integral>\<^sup>+x. f x \<partial>M)"
 proof -
-  { fix x y z have "(x \<noteq> 0 \<Longrightarrow> y = z) \<Longrightarrow> ereal x * y = ereal x * z"
-      by (cases "x = 0") (auto simp: zero_ereal_def[symmetric]) }
-  note ereal_cong_mult = this
+  { fix x y z have "(x \<noteq> 0 \<Longrightarrow> y = z) \<Longrightarrow> ennreal x * y = ennreal x * z"
+      by (cases "x = 0") (auto simp: zero_ennreal_def[symmetric]) }
+  note ennreal_cong_mult = this
 
   have [measurable]: "f \<in> borel_measurable M"
     using f(1) by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
 
   { fix y assume y: "y \<in> space M" "f y \<noteq> 0"
-    have "ereal (measure M {x \<in> space M. f x = f y}) = emeasure M {x \<in> space M. f x = f y}"
-    proof (rule emeasure_eq_ereal_measure[symmetric])
+    have "ennreal (measure M {x \<in> space M. f x = f y}) = emeasure M {x \<in> space M. f x = f y}"
+    proof (rule emeasure_eq_ennreal_measure[symmetric])
       have "emeasure M {x \<in> space M. f x = f y} \<le> emeasure M {x \<in> space M. f x \<noteq> 0}"
         using y by (intro emeasure_mono) auto
-      with f show "emeasure M {x \<in> space M. f x = f y} \<noteq> \<infinity>"
-        by (auto simp: simple_bochner_integrable.simps)
+      with f show "emeasure M {x \<in> space M. f x = f y} \<noteq> top"
+        by (auto simp: simple_bochner_integrable.simps top_unique)
     qed
-    moreover have "{x \<in> space M. f x = f y} = (\<lambda>x. ereal (f x)) -` {ereal (f y)} \<inter> space M"
-      by auto
-    ultimately have "ereal (measure M {x \<in> space M. f x = f y}) =
-          emeasure M ((\<lambda>x. ereal (f x)) -` {ereal (f y)} \<inter> space M)" by simp }
+    moreover have "{x \<in> space M. f x = f y} = (\<lambda>x. ennreal (f x)) -` {ennreal (f y)} \<inter> space M"
+      using f by auto
+    ultimately have "ennreal (measure M {x \<in> space M. f x = f y}) =
+          emeasure M ((\<lambda>x. ennreal (f x)) -` {ennreal (f y)} \<inter> space M)" by simp }
   with f have "simple_bochner_integral M f = (\<integral>\<^sup>Sx. f x \<partial>M)"
     unfolding simple_integral_def
-    by (subst simple_bochner_integral_partition[OF f(1), where g="\<lambda>x. ereal (f x)" and v=real_of_ereal])
+    by (subst simple_bochner_integral_partition[OF f(1), where g="\<lambda>x. ennreal (f x)" and v=enn2real])
        (auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases
-             intro!: setsum.cong ereal_cong_mult
-             simp: setsum_ereal[symmetric] times_ereal.simps(1)[symmetric] ac_simps
-             simp del: setsum_ereal times_ereal.simps(1))
+             intro!: setsum.cong ennreal_cong_mult
+             simp: setsum_ennreal[symmetric] ac_simps ennreal_mult
+             simp del: setsum_ennreal)
   also have "\<dots> = (\<integral>\<^sup>+x. f x \<partial>M)"
     using f
     by (intro nn_integral_eq_simple_integral[symmetric])
@@ -488,14 +485,14 @@
   fixes f :: "'a \<Rightarrow> 'b::{real_normed_vector, second_countable_topology}"
   assumes f[measurable]: "f \<in> borel_measurable M"
   assumes s: "simple_bochner_integrable M s" and t: "simple_bochner_integrable M t"
-  shows "ereal (norm (simple_bochner_integral M s - simple_bochner_integral M t)) \<le>
+  shows "ennreal (norm (simple_bochner_integral M s - simple_bochner_integral M t)) \<le>
     (\<integral>\<^sup>+ x. norm (f x - s x) \<partial>M) + (\<integral>\<^sup>+ x. norm (f x - t x) \<partial>M)"
-    (is "ereal (norm (?s - ?t)) \<le> ?S + ?T")
+    (is "ennreal (norm (?s - ?t)) \<le> ?S + ?T")
 proof -
   have [measurable]: "s \<in> borel_measurable M" "t \<in> borel_measurable M"
     using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
 
-  have "ereal (norm (?s - ?t)) = norm (simple_bochner_integral M (\<lambda>x. s x - t x))"
+  have "ennreal (norm (?s - ?t)) = norm (simple_bochner_integral M (\<lambda>x. s x - t x))"
     using s t by (subst simple_bochner_integral_diff) auto
   also have "\<dots> \<le> simple_bochner_integral M (\<lambda>x. norm (s x - t x))"
     using simple_bochner_integrable_compose2[of "op -" M "s" "t"] s t
@@ -503,8 +500,8 @@
   also have "\<dots> = (\<integral>\<^sup>+x. norm (s x - t x) \<partial>M)"
     using simple_bochner_integrable_compose2[of "\<lambda>x y. norm (x - y)" M "s" "t"] s t
     by (auto intro!: simple_bochner_integral_eq_nn_integral)
-  also have "\<dots> \<le> (\<integral>\<^sup>+x. ereal (norm (f x - s x)) + ereal (norm (f x - t x)) \<partial>M)"
-    by (auto intro!: nn_integral_mono)
+  also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s x)) + ennreal (norm (f x - t x)) \<partial>M)"
+    by (auto intro!: nn_integral_mono simp: ennreal_plus[symmetric] simp del: ennreal_plus)
        (metis (erased, hide_lams) add_diff_cancel_left add_diff_eq diff_add_eq order_trans
               norm_minus_commute norm_triangle_ineq4 order_refl)
   also have "\<dots> = ?S + ?T"
@@ -545,9 +542,9 @@
 lemma has_bochner_integral_simple_bochner_integrable:
   "simple_bochner_integrable M f \<Longrightarrow> has_bochner_integral M f (simple_bochner_integral M f)"
   by (rule has_bochner_integral.intros[where s="\<lambda>_. f"])
-     (auto intro: borel_measurable_simple_function 
+     (auto intro: borel_measurable_simple_function
            elim: simple_bochner_integrable.cases
-           simp: zero_ereal_def[symmetric])
+           simp: zero_ennreal_def[symmetric])
 
 lemma has_bochner_integral_real_indicator:
   assumes [measurable]: "A \<in> sets M" and A: "emeasure M A < \<infinity>"
@@ -562,7 +559,7 @@
   qed (rule simple_function_indicator assms)+
   moreover have "simple_bochner_integral M (indicator A) = measure M A"
     using simple_bochner_integral_eq_nn_integral[OF sbi] A
-    by (simp add: ereal_indicator emeasure_eq_ereal_measure)
+    by (simp add: ennreal_indicator emeasure_eq_ennreal_measure)
   ultimately show ?thesis
     by (metis has_bochner_integral_simple_bochner_integrable)
 qed
@@ -588,18 +585,19 @@
     (is "?f \<longlonglongrightarrow> 0")
   proof (rule tendsto_sandwich)
     show "eventually (\<lambda>n. 0 \<le> ?f n) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> 0"
-      by (auto simp: nn_integral_nonneg)
+      by auto
     show "eventually (\<lambda>i. ?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) \<partial>M) + \<integral>\<^sup>+ x. (norm (g x - sg i x)) \<partial>M) sequentially"
       (is "eventually (\<lambda>i. ?f i \<le> ?g i) sequentially")
     proof (intro always_eventually allI)
-      fix i have "?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) + ereal (norm (g x - sg i x)) \<partial>M)"
-        by (auto intro!: nn_integral_mono norm_diff_triangle_ineq)
+      fix i have "?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) + ennreal (norm (g x - sg i x)) \<partial>M)"
+        by (auto intro!: nn_integral_mono norm_diff_triangle_ineq
+                 simp del: ennreal_plus simp add: ennreal_plus[symmetric])
       also have "\<dots> = ?g i"
         by (intro nn_integral_add) auto
       finally show "?f i \<le> ?g i" .
     qed
     show "?g \<longlonglongrightarrow> 0"
-      using tendsto_add_ereal[OF _ _ f_sf g_sg] by simp
+      using tendsto_add[OF f_sf g_sg] by simp
   qed
 qed (auto simp: simple_bochner_integral_add tendsto_add)
 
@@ -629,19 +627,19 @@
     (is "?f \<longlonglongrightarrow> 0")
   proof (rule tendsto_sandwich)
     show "eventually (\<lambda>n. 0 \<le> ?f n) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> 0"
-      by (auto simp: nn_integral_nonneg)
+      by auto
 
     show "eventually (\<lambda>i. ?f i \<le> K * (\<integral>\<^sup>+ x. norm (f x - s i x) \<partial>M)) sequentially"
       (is "eventually (\<lambda>i. ?f i \<le> ?g i) sequentially")
     proof (intro always_eventually allI)
-      fix i have "?f i \<le> (\<integral>\<^sup>+ x. ereal K * norm (f x - s i x) \<partial>M)"
-        using K by (intro nn_integral_mono) (auto simp: ac_simps)
+      fix i have "?f i \<le> (\<integral>\<^sup>+ x. ennreal K * norm (f x - s i x) \<partial>M)"
+        using K by (intro nn_integral_mono) (auto simp: ac_simps ennreal_mult[symmetric])
       also have "\<dots> = ?g i"
         using K by (intro nn_integral_cmult) auto
       finally show "?f i \<le> ?g i" .
     qed
     show "?g \<longlonglongrightarrow> 0"
-      using tendsto_cmult_ereal[OF _ f_s, of "ereal K"] by simp
+      using ennreal_tendsto_cmult[OF _ f_s] by simp
   qed
 
   assume "(\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x"
@@ -651,7 +649,7 @@
 
 lemma has_bochner_integral_zero[intro]: "has_bochner_integral M (\<lambda>x. 0) 0"
   by (auto intro!: has_bochner_integral.intros[where s="\<lambda>_ _. 0"]
-           simp: zero_ereal_def[symmetric] simple_bochner_integrable.simps
+           simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps
                  simple_bochner_integral_def image_constant_conv)
 
 lemma has_bochner_integral_scaleR_left[intro]:
@@ -672,7 +670,7 @@
   shows "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. c * f x) (c * x)"
   by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_right])
 
-lemmas has_bochner_integral_divide = 
+lemmas has_bochner_integral_divide =
   has_bochner_integral_bounded_linear[OF bounded_linear_divide]
 
 lemma has_bochner_integral_divide_zero[intro]:
@@ -724,11 +722,10 @@
 proof (elim has_bochner_integral.cases)
   fix s v
   assume [measurable]: "f \<in> borel_measurable M" and s: "\<And>i. simple_bochner_integrable M (s i)" and
-    lim_0: "(\<lambda>i. \<integral>\<^sup>+ x. ereal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0"
+    lim_0: "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0"
   from order_tendstoD[OF lim_0, of "\<infinity>"]
-  obtain i where f_s_fin: "(\<integral>\<^sup>+ x. ereal (norm (f x - s i x)) \<partial>M) < \<infinity>"
-    by (metis (mono_tags, lifting) eventually_False_sequentially eventually_mono
-              less_ereal.simps(4) zero_ereal_def)
+  obtain i where f_s_fin: "(\<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) < \<infinity>"
+    by (auto simp: eventually_sequentially)
 
   have [measurable]: "\<And>i. s i \<in> borel_measurable M"
     using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
@@ -740,19 +737,20 @@
     by (rule finite_imageI)
   then have "\<And>x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> m" "0 \<le> m"
     by (auto simp: m_def image_comp comp_def Max_ge_iff)
-  then have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ereal m * indicator {x\<in>space M. s i x \<noteq> 0} x \<partial>M)"
+  then have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal m * indicator {x\<in>space M. s i x \<noteq> 0} x \<partial>M)"
     by (auto split: split_indicator intro!: Max_ge nn_integral_mono simp:)
   also have "\<dots> < \<infinity>"
-    using s by (subst nn_integral_cmult_indicator) (auto simp: \<open>0 \<le> m\<close> simple_bochner_integrable.simps)
+    using s by (subst nn_integral_cmult_indicator) (auto simp: \<open>0 \<le> m\<close> simple_bochner_integrable.simps ennreal_mult_less_top less_top)
   finally have s_fin: "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) < \<infinity>" .
 
-  have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ereal (norm (f x - s i x)) + ereal (norm (s i x)) \<partial>M)"
-    by (auto intro!: nn_integral_mono) (metis add.commute norm_triangle_sub)
+  have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) + ennreal (norm (s i x)) \<partial>M)"
+    by (auto intro!: nn_integral_mono simp del: ennreal_plus simp add: ennreal_plus[symmetric])
+       (metis add.commute norm_triangle_sub)
   also have "\<dots> = (\<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) + (\<integral>\<^sup>+x. norm (s i x) \<partial>M)"
     by (rule nn_integral_add) auto
   also have "\<dots> < \<infinity>"
     using s_fin f_s_fin by auto
-  finally show "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) < \<infinity>" .
+  finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" .
 qed
 
 lemma has_bochner_integral_norm_bound:
@@ -762,7 +760,7 @@
   fix s assume
     x: "(\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x" (is "?s \<longlonglongrightarrow> x") and
     s[simp]: "\<And>i. simple_bochner_integrable M (s i)" and
-    lim: "(\<lambda>i. \<integral>\<^sup>+ x. ereal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" and
+    lim: "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" and
     f[measurable]: "f \<in> borel_measurable M"
 
   have [measurable]: "\<And>i. s i \<in> borel_measurable M"
@@ -770,28 +768,28 @@
 
   show "norm x \<le> (\<integral>\<^sup>+x. norm (f x) \<partial>M)"
   proof (rule LIMSEQ_le)
-    show "(\<lambda>i. ereal (norm (?s i))) \<longlonglongrightarrow> norm x"
-      using x by (intro tendsto_intros lim_ereal[THEN iffD2])
+    show "(\<lambda>i. ennreal (norm (?s i))) \<longlonglongrightarrow> norm x"
+      using x by (auto simp: tendsto_ennreal_iff intro: tendsto_intros)
     show "\<exists>N. \<forall>n\<ge>N. norm (?s n) \<le> (\<integral>\<^sup>+x. norm (f x - s n x) \<partial>M) + (\<integral>\<^sup>+x. norm (f x) \<partial>M)"
       (is "\<exists>N. \<forall>n\<ge>N. _ \<le> ?t n")
     proof (intro exI allI impI)
       fix n
-      have "ereal (norm (?s n)) \<le> simple_bochner_integral M (\<lambda>x. norm (s n x))"
+      have "ennreal (norm (?s n)) \<le> simple_bochner_integral M (\<lambda>x. norm (s n x))"
         by (auto intro!: simple_bochner_integral_norm_bound)
       also have "\<dots> = (\<integral>\<^sup>+x. norm (s n x) \<partial>M)"
         by (intro simple_bochner_integral_eq_nn_integral)
            (auto intro: s simple_bochner_integrable_compose2)
-      also have "\<dots> \<le> (\<integral>\<^sup>+x. ereal (norm (f x - s n x)) + norm (f x) \<partial>M)"
-        by (auto intro!: nn_integral_mono)
+      also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s n x)) + norm (f x) \<partial>M)"
+        by (auto intro!: nn_integral_mono simp del: ennreal_plus simp add: ennreal_plus[symmetric])
            (metis add.commute norm_minus_commute norm_triangle_sub)
-      also have "\<dots> = ?t n" 
+      also have "\<dots> = ?t n"
         by (rule nn_integral_add) auto
       finally show "norm (?s n) \<le> ?t n" .
     qed
-    have "?t \<longlonglongrightarrow> 0 + (\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M)"
+    have "?t \<longlonglongrightarrow> 0 + (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
       using has_bochner_integral_implies_finite_norm[OF i]
-      by (intro tendsto_add_ereal tendsto_const lim) auto
-    then show "?t \<longlonglongrightarrow> \<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M"
+      by (intro tendsto_add tendsto_const lim)
+    then show "?t \<longlonglongrightarrow> \<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M"
       by simp
   qed
 qed
@@ -816,19 +814,18 @@
   then have "(\<lambda>i. norm (?s i - ?t i)) \<longlonglongrightarrow> norm (x - y)"
     by (intro tendsto_intros)
   moreover
-  have "(\<lambda>i. ereal (norm (?s i - ?t i))) \<longlonglongrightarrow> ereal 0"
+  have "(\<lambda>i. ennreal (norm (?s i - ?t i))) \<longlonglongrightarrow> ennreal 0"
   proof (rule tendsto_sandwich)
-    show "eventually (\<lambda>i. 0 \<le> ereal (norm (?s i - ?t i))) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> ereal 0"
-      by (auto simp: nn_integral_nonneg zero_ereal_def[symmetric])
+    show "eventually (\<lambda>i. 0 \<le> ennreal (norm (?s i - ?t i))) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> ennreal 0"
+      by auto
 
     show "eventually (\<lambda>i. norm (?s i - ?t i) \<le> ?S i + ?T i) sequentially"
       by (intro always_eventually allI simple_bochner_integral_bounded s t f)
-    show "(\<lambda>i. ?S i + ?T i) \<longlonglongrightarrow> ereal 0"
-      using tendsto_add_ereal[OF _ _ \<open>?S \<longlonglongrightarrow> 0\<close> \<open>?T \<longlonglongrightarrow> 0\<close>]
-      by (simp add: zero_ereal_def[symmetric])
+    show "(\<lambda>i. ?S i + ?T i) \<longlonglongrightarrow> ennreal 0"
+      using tendsto_add[OF \<open>?S \<longlonglongrightarrow> 0\<close> \<open>?T \<longlonglongrightarrow> 0\<close>] by simp
   qed
   then have "(\<lambda>i. norm (?s i - ?t i)) \<longlonglongrightarrow> 0"
-    by simp
+    by (simp add: ennreal_0[symmetric] del: ennreal_0)
   ultimately have "norm (x - y) = 0"
     by (rule LIMSEQ_unique)
   then show "x = y" by simp
@@ -841,11 +838,11 @@
   shows "has_bochner_integral M g x"
   using f
 proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases)
-  fix s assume "(\<lambda>i. \<integral>\<^sup>+ x. ereal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0"
-  also have "(\<lambda>i. \<integral>\<^sup>+ x. ereal (norm (f x - s i x)) \<partial>M) = (\<lambda>i. \<integral>\<^sup>+ x. ereal (norm (g x - s i x)) \<partial>M)"
+  fix s assume "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0"
+  also have "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) = (\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (g x - s i x)) \<partial>M)"
     using ae
     by (intro ext nn_integral_cong_AE, eventually_elim) simp
-  finally show "(\<lambda>i. \<integral>\<^sup>+ x. ereal (norm (g x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" .
+  finally show "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (g x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" .
 qed (auto intro: g)
 
 lemma has_bochner_integral_eq_AE:
@@ -966,14 +963,14 @@
   by (auto simp: integrable.simps)
 
 lemma integrable_zero[simp, intro]: "integrable M (\<lambda>x. 0)"
-  by (metis has_bochner_integral_zero integrable.simps) 
+  by (metis has_bochner_integral_zero integrable.simps)
 
 lemma integrable_setsum[simp, intro]: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow> integrable M (\<lambda>x. \<Sum>i\<in>I. f i x)"
-  by (metis has_bochner_integral_setsum integrable.simps) 
+  by (metis has_bochner_integral_setsum integrable.simps)
 
 lemma integrable_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
   integrable M (\<lambda>x. indicator A x *\<^sub>R c)"
-  by (metis has_bochner_integral_indicator integrable.simps) 
+  by (metis has_bochner_integral_indicator integrable.simps)
 
 lemma integrable_real_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
   integrable M (indicator A :: 'a \<Rightarrow> real)"
@@ -981,7 +978,7 @@
 
 lemma integrable_diff[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x - g x)"
   by (auto simp: integrable.simps intro: has_bochner_integral_diff)
-  
+
 lemma integrable_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. T (f x))"
   by (auto simp: integrable.simps intro: has_bochner_integral_bounded_linear)
 
@@ -1140,7 +1137,7 @@
 lemmas integral_snd[simp] =
   integral_bounded_linear[OF bounded_linear_snd]
 
-lemma integral_norm_bound_ereal:
+lemma integral_norm_bound_ennreal:
   "integrable M f \<Longrightarrow> norm (integral\<^sup>L M f) \<le> (\<integral>\<^sup>+x. norm (f x) \<partial>M)"
   by (metis has_bochner_integral_integrable has_bochner_integral_norm_bound)
 
@@ -1157,7 +1154,7 @@
     unfolding convergent_eq_cauchy
   proof (rule metric_CauchyI)
     fix e :: real assume "0 < e"
-    then have "0 < ereal (e / 2)" by auto
+    then have "0 < ennreal (e / 2)" by auto
     from order_tendstoD(2)[OF lim this]
     obtain M where M: "\<And>n. M \<le> n \<Longrightarrow> ?S n < e / 2"
       by (auto simp: eventually_sequentially)
@@ -1168,12 +1165,11 @@
         using M[OF n] by auto
       have "norm (?s n - ?s m) \<le> ?S n + ?S m"
         by (intro simple_bochner_integral_bounded s f)
-      also have "\<dots> < ereal (e / 2) + e / 2"
-        using ereal_add_strict_mono[OF less_imp_le[OF M[OF n]] _ \<open>?S n \<noteq> \<infinity>\<close> M[OF m]]
-        by (auto simp: nn_integral_nonneg)
-      also have "\<dots> = e" by simp
+      also have "\<dots> < ennreal (e / 2) + e / 2"
+        by (intro add_strict_mono M n m)
+      also have "\<dots> = e" using \<open>0<e\<close> by (simp del: ennreal_plus add: ennreal_plus[symmetric])
       finally show "dist (?s n) (?s m) < e"
-        by (simp add: dist_norm)
+        using \<open>0<e\<close> by (simp add: dist_norm ennreal_less_iff)
     qed
   qed
   then obtain x where "?s \<longlonglongrightarrow> x" ..
@@ -1203,20 +1199,22 @@
       by (rule norm_triangle_ineq4)
     finally (xtrans) show "norm (u' x - u i x) \<le> 2 * w x" .
   qed
-  
+  have w_nonneg: "AE x in M. 0 \<le> w x"
+    using bound[of 0] by (auto intro: order_trans[OF norm_ge_zero])
+
   have "(\<lambda>i. (\<integral>\<^sup>+x. norm (u' x - u i x) \<partial>M)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. 0 \<partial>M)"
-  proof (rule nn_integral_dominated_convergence)  
+  proof (rule nn_integral_dominated_convergence)
     show "(\<integral>\<^sup>+x. 2 * w x \<partial>M) < \<infinity>"
-      by (rule nn_integral_mult_bounded_inf[OF _ w, of 2]) auto
-    show "AE x in M. (\<lambda>i. ereal (norm (u' x - u i x))) \<longlonglongrightarrow> 0"
-      using u' 
+      by (rule nn_integral_mult_bounded_inf[OF _ w, of 2]) (insert w_nonneg, auto simp: ennreal_mult )
+    show "AE x in M. (\<lambda>i. ennreal (norm (u' x - u i x))) \<longlonglongrightarrow> 0"
+      using u'
     proof eventually_elim
       fix x assume "(\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
       from tendsto_diff[OF tendsto_const[of "u' x"] this]
-      show "(\<lambda>i. ereal (norm (u' x - u i x))) \<longlonglongrightarrow> 0"
-        by (simp add: zero_ereal_def tendsto_norm_zero_iff)
+      show "(\<lambda>i. ennreal (norm (u' x - u i x))) \<longlonglongrightarrow> 0"
+        by (simp add: tendsto_norm_zero_iff ennreal_0[symmetric] del: ennreal_0)
     qed
-  qed (insert bnd, auto)
+  qed (insert bnd w_nonneg, auto)
   then show ?thesis by simp
 qed
 
@@ -1230,29 +1228,31 @@
     pointwise: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x" and
     bound: "\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)"
     by simp metis
-  
+
   show ?thesis
   proof (rule integrableI_sequence)
     { fix i
-      have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. 2 * ereal (norm (f x)) \<partial>M)"
+      have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal (2 * norm (f x)) \<partial>M)"
         by (intro nn_integral_mono) (simp add: bound)
-      also have "\<dots> = 2 * (\<integral>\<^sup>+x. ereal (norm (f x)) \<partial>M)"
-        by (rule nn_integral_cmult) auto
+      also have "\<dots> = 2 * (\<integral>\<^sup>+x. ennreal (norm (f x)) \<partial>M)"
+        by (simp add: ennreal_mult nn_integral_cmult)
+      also have "\<dots> < top"
+        using fin by (simp add: ennreal_mult_less_top)
       finally have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) < \<infinity>"
-        using fin by auto }
+        by simp }
     note fin_s = this
 
     show "\<And>i. simple_bochner_integrable M (s i)"
       by (rule simple_bochner_integrableI_bounded) fact+
 
-    show "(\<lambda>i. \<integral>\<^sup>+ x. ereal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0"
+    show "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0"
     proof (rule nn_integral_dominated_convergence_norm)
       show "\<And>j. AE x in M. norm (s j x) \<le> 2 * norm (f x)"
         using bound by auto
       show "\<And>i. s i \<in> borel_measurable M" "(\<lambda>x. 2 * norm (f x)) \<in> borel_measurable M"
         using s by (auto intro: borel_measurable_simple_function)
-      show "(\<integral>\<^sup>+ x. ereal (2 * norm (f x)) \<partial>M) < \<infinity>"
-        using fin unfolding times_ereal.simps(1)[symmetric] by (subst nn_integral_cmult) auto
+      show "(\<integral>\<^sup>+ x. ennreal (2 * norm (f x)) \<partial>M) < \<infinity>"
+        using fin by (simp add: nn_integral_cmult ennreal_mult ennreal_mult_less_top)
       show "AE x in M. (\<lambda>i. s i x) \<longlonglongrightarrow> f x"
         using pointwise by auto
     qed fact
@@ -1269,11 +1269,11 @@
 proof (rule integrableI_bounded)
   { fix x :: 'b have "norm x \<le> B \<Longrightarrow> 0 \<le> B"
       using norm_ge_zero[of x] by arith }
-  with bnd null have "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ereal (max 0 B) * indicator A x \<partial>M)"
+  with bnd null have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (max 0 B) * indicator A x \<partial>M)"
     by (intro nn_integral_mono_AE) (auto split: split_indicator split_max)
   also have "\<dots> < \<infinity>"
-    using finite by (subst nn_integral_cmult_indicator) (auto simp: max_def)
-  finally show "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) < \<infinity>" .
+    using finite by (subst nn_integral_cmult_indicator) (auto simp: ennreal_mult_less_top)
+  finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" .
 qed simp
 
 lemma integrableI_bounded_set_indicator:
@@ -1309,11 +1309,11 @@
 proof safe
   assume "f \<in> borel_measurable M" "g \<in> borel_measurable M"
   assume "AE x in M. norm (g x) \<le> norm (f x)"
-  then have "(\<integral>\<^sup>+ x. ereal (norm (g x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M)"
+  then have "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
     by  (intro nn_integral_mono_AE) auto
-  also assume "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) < \<infinity>"
-  finally show "(\<integral>\<^sup>+ x. ereal (norm (g x)) \<partial>M) < \<infinity>" .
-qed 
+  also assume "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>"
+  finally show "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) < \<infinity>" .
+qed
 
 lemma integrable_mult_indicator:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
@@ -1334,7 +1334,7 @@
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
   assumes [measurable]: "integrable M f" shows "integrable M (\<lambda>x. norm (f x))"
   using assms by (rule integrable_bound) auto
-  
+
 lemma integrable_norm_cancel:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
   assumes [measurable]: "integrable M (\<lambda>x. norm (f x))" "f \<in> borel_measurable M" shows "integrable M f"
@@ -1376,7 +1376,7 @@
 
 lemma integrable_indicator_iff:
   "integrable M (indicator A::_ \<Rightarrow> real) \<longleftrightarrow> A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>"
-  by (simp add: integrable_iff_bounded borel_measurable_indicator_iff ereal_indicator nn_integral_indicator'
+  by (simp add: integrable_iff_bounded borel_measurable_indicator_iff ennreal_indicator nn_integral_indicator'
            cong: conj_cong)
 
 lemma integral_indicator[simp]: "integral\<^sup>L M (indicator A) = measure M (A \<inter> space M)"
@@ -1394,14 +1394,14 @@
   also have "\<dots> = 0"
     using * by (subst not_integrable_integral_eq) (auto simp: integrable_indicator_iff)
   also have "\<dots> = measure M (A \<inter> space M)"
-    using * by (auto simp: measure_def emeasure_notin_sets)
+    using * by (auto simp: measure_def emeasure_notin_sets not_less top_unique)
   finally show ?thesis .
 qed
 
 lemma integrable_discrete_difference:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
   assumes X: "countable X"
-  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" 
+  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
   assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
   assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
   shows "integrable M f \<longleftrightarrow> integrable M g"
@@ -1425,7 +1425,7 @@
 lemma integral_discrete_difference:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
   assumes X: "countable X"
-  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" 
+  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
   assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
   assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
   shows "integral\<^sup>L M f = integral\<^sup>L M g"
@@ -1449,7 +1449,7 @@
 lemma has_bochner_integral_discrete_difference:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
   assumes X: "countable X"
-  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" 
+  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
   assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
   assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
   shows "has_bochner_integral M f x \<longleftrightarrow> has_bochner_integral M g x"
@@ -1466,7 +1466,7 @@
     and integrable_dominated_convergence2: "\<And>i. integrable M (s i)"
     and integral_dominated_convergence: "(\<lambda>i. integral\<^sup>L M (s i)) \<longlonglongrightarrow> integral\<^sup>L M f"
 proof -
-  have "AE x in M. 0 \<le> w x"
+  have w_nonneg: "AE x in M. 0 \<le> w x"
     using bound[of 0] by eventually_elim (auto intro: norm_ge_zero order_trans)
   then have "(\<integral>\<^sup>+x. w x \<partial>M) = (\<integral>\<^sup>+x. norm (w x) \<partial>M)"
     by (intro nn_integral_cong_AE) auto
@@ -1476,10 +1476,10 @@
   show int_s: "\<And>i. integrable M (s i)"
     unfolding integrable_iff_bounded
   proof
-    fix i 
-    have "(\<integral>\<^sup>+ x. ereal (norm (s i x)) \<partial>M) \<le> (\<integral>\<^sup>+x. w x \<partial>M)"
-      using bound by (intro nn_integral_mono_AE) auto
-    with w show "(\<integral>\<^sup>+ x. ereal (norm (s i x)) \<partial>M) < \<infinity>" by auto
+    fix i
+    have "(\<integral>\<^sup>+ x. ennreal (norm (s i x)) \<partial>M) \<le> (\<integral>\<^sup>+x. w x \<partial>M)"
+      using bound[of i] w_nonneg by (intro nn_integral_mono_AE) auto
+    with w show "(\<integral>\<^sup>+ x. ennreal (norm (s i x)) \<partial>M) < \<infinity>" by auto
   qed fact
 
   have all_bound: "AE x in M. \<forall>i. norm (s i x) \<le> w x"
@@ -1488,30 +1488,30 @@
   show int_f: "integrable M f"
     unfolding integrable_iff_bounded
   proof
-    have "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+x. w x \<partial>M)"
-      using all_bound lim
+    have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+x. w x \<partial>M)"
+      using all_bound lim w_nonneg
     proof (intro nn_integral_mono_AE, eventually_elim)
-      fix x assume "\<forall>i. norm (s i x) \<le> w x" "(\<lambda>i. s i x) \<longlonglongrightarrow> f x"
-      then show "ereal (norm (f x)) \<le> ereal (w x)"
-        by (intro LIMSEQ_le_const2[where X="\<lambda>i. ereal (norm (s i x))"] tendsto_intros lim_ereal[THEN iffD2]) auto
+      fix x assume "\<forall>i. norm (s i x) \<le> w x" "(\<lambda>i. s i x) \<longlonglongrightarrow> f x" "0 \<le> w x"
+      then show "ennreal (norm (f x)) \<le> ennreal (w x)"
+        by (intro LIMSEQ_le_const2[where X="\<lambda>i. ennreal (norm (s i x))"]) (auto intro: tendsto_intros)
     qed
-    with w show "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) < \<infinity>" by auto
+    with w show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" by auto
   qed fact
 
-  have "(\<lambda>n. ereal (norm (integral\<^sup>L M (s n) - integral\<^sup>L M f))) \<longlonglongrightarrow> ereal 0" (is "?d \<longlonglongrightarrow> ereal 0")
+  have "(\<lambda>n. ennreal (norm (integral\<^sup>L M (s n) - integral\<^sup>L M f))) \<longlonglongrightarrow> ennreal 0" (is "?d \<longlonglongrightarrow> ennreal 0")
   proof (rule tendsto_sandwich)
-    show "eventually (\<lambda>n. ereal 0 \<le> ?d n) sequentially" "(\<lambda>_. ereal 0) \<longlonglongrightarrow> ereal 0" by auto
+    show "eventually (\<lambda>n. ennreal 0 \<le> ?d n) sequentially" "(\<lambda>_. ennreal 0) \<longlonglongrightarrow> ennreal 0" by auto
     show "eventually (\<lambda>n. ?d n \<le> (\<integral>\<^sup>+x. norm (s n x - f x) \<partial>M)) sequentially"
     proof (intro always_eventually allI)
       fix n
       have "?d n = norm (integral\<^sup>L M (\<lambda>x. s n x - f x))"
         using int_f int_s by simp
       also have "\<dots> \<le> (\<integral>\<^sup>+x. norm (s n x - f x) \<partial>M)"
-        by (intro int_f int_s integrable_diff integral_norm_bound_ereal)
+        by (intro int_f int_s integrable_diff integral_norm_bound_ennreal)
       finally show "?d n \<le> (\<integral>\<^sup>+x. norm (s n x - f x) \<partial>M)" .
     qed
-    show "(\<lambda>n. \<integral>\<^sup>+x. norm (s n x - f x) \<partial>M) \<longlonglongrightarrow> ereal 0"
-      unfolding zero_ereal_def[symmetric]
+    show "(\<lambda>n. \<integral>\<^sup>+x. norm (s n x - f x) \<partial>M) \<longlonglongrightarrow> ennreal 0"
+      unfolding ennreal_0
       apply (subst norm_minus_commute)
     proof (rule nn_integral_dominated_convergence_norm[where w=w])
       show "\<And>n. s n \<in> borel_measurable M"
@@ -1519,7 +1519,7 @@
     qed fact+
   qed
   then have "(\<lambda>n. integral\<^sup>L M (s n) - integral\<^sup>L M f) \<longlonglongrightarrow> 0"
-    unfolding lim_ereal tendsto_norm_zero_iff .
+    by (simp add: tendsto_norm_zero_iff del: ennreal_0)
   from tendsto_add[OF this tendsto_const[of "integral\<^sup>L M f"]]
   show "(\<lambda>i. integral\<^sup>L M (s i)) \<longlonglongrightarrow> integral\<^sup>L M f"  by simp
 qed
@@ -1580,59 +1580,101 @@
   using integrable_mult_left[of c M f] integrable_mult_left[of "1 / c" M "\<lambda>x. c * f x"]
   by (cases "c = 0") auto
 
+lemma integrableI_nn_integral_finite:
+  assumes [measurable]: "f \<in> borel_measurable M"
+    and nonneg: "AE x in M. 0 \<le> f x"
+    and finite: "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x"
+  shows "integrable M f"
+proof (rule integrableI_bounded)
+  have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M)"
+    using nonneg by (intro nn_integral_cong_AE) auto
+  with finite show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>"
+    by auto
+qed simp
+
+lemma integral_nonneg_AE:
+  fixes f :: "'a \<Rightarrow> real"
+  assumes nonneg: "AE x in M. 0 \<le> f x"
+  shows "0 \<le> integral\<^sup>L M f"
+proof cases
+  assume f: "integrable M f"
+  then have [measurable]: "f \<in> M \<rightarrow>\<^sub>M borel"
+    by auto
+  have "(\<lambda>x. max 0 (f x)) \<in> M \<rightarrow>\<^sub>M borel" "\<And>x. 0 \<le> max 0 (f x)" "integrable M (\<lambda>x. max 0 (f x))"
+    using f by auto
+  from this have "0 \<le> integral\<^sup>L M (\<lambda>x. max 0 (f x))"
+  proof (induction rule: borel_measurable_induct_real)
+    case (add f g)
+    then have "integrable M f" "integrable M g"
+      by (auto intro!: integrable_bound[OF add.prems])
+    with add show ?case
+      by (simp add: nn_integral_add)
+  next
+    case (seq U)
+    show ?case
+    proof (rule LIMSEQ_le_const)
+      have U_le: "x \<in> space M \<Longrightarrow> U i x \<le> max 0 (f x)" for x i
+        using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def)
+      with seq nonneg show "(\<lambda>i. integral\<^sup>L M (U i)) \<longlonglongrightarrow> LINT x|M. max 0 (f x)"
+        by (intro integral_dominated_convergence) auto
+      have "integrable M (U i)" for i
+        using seq.prems by (rule integrable_bound) (insert U_le seq, auto)
+      with seq show "\<exists>N. \<forall>n\<ge>N. 0 \<le> integral\<^sup>L M (U n)"
+        by auto
+    qed
+  qed (auto simp: measure_nonneg integrable_mult_left_iff)
+  also have "\<dots> = integral\<^sup>L M f"
+    using nonneg by (auto intro!: integral_cong_AE)
+  finally show ?thesis .
+qed (simp add: not_integrable_integral_eq)
+
+lemma integral_nonneg[simp]:
+  fixes f :: "'a \<Rightarrow> real"
+  shows "(\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> 0 \<le> integral\<^sup>L M f"
+  by (intro integral_nonneg_AE) auto
+
 lemma nn_integral_eq_integral:
   assumes f: "integrable M f"
-  assumes nonneg: "AE x in M. 0 \<le> f x" 
+  assumes nonneg: "AE x in M. 0 \<le> f x"
   shows "(\<integral>\<^sup>+ x. f x \<partial>M) = integral\<^sup>L M f"
 proof -
   { fix f :: "'a \<Rightarrow> real" assume f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" "integrable M f"
     then have "(\<integral>\<^sup>+ x. f x \<partial>M) = integral\<^sup>L M f"
     proof (induct rule: borel_measurable_induct_real)
       case (set A) then show ?case
-        by (simp add: integrable_indicator_iff ereal_indicator emeasure_eq_ereal_measure)
+        by (simp add: integrable_indicator_iff ennreal_indicator emeasure_eq_ennreal_measure)
     next
       case (mult f c) then show ?case
-        unfolding times_ereal.simps(1)[symmetric]
-        by (subst nn_integral_cmult)
-           (auto simp add: integrable_mult_left_iff zero_ereal_def[symmetric])
+        by (auto simp add: integrable_mult_left_iff nn_integral_cmult ennreal_mult integral_nonneg_AE)
     next
       case (add g f)
       then have "integrable M f" "integrable M g"
-        by (auto intro!: integrable_bound[OF add(8)])
+        by (auto intro!: integrable_bound[OF add.prems])
       with add show ?case
-        unfolding plus_ereal.simps(1)[symmetric]
-        by (subst nn_integral_add) auto
+        by (simp add: nn_integral_add integral_nonneg_AE)
     next
-      case (seq s)
-      { fix i x assume "x \<in> space M" with seq(4) have "s i x \<le> f x"
-          by (intro LIMSEQ_le_const[OF seq(5)] exI[of _ i]) (auto simp: incseq_def le_fun_def) }
-      note s_le_f = this
-
+      case (seq U)
       show ?case
       proof (rule LIMSEQ_unique)
-        show "(\<lambda>i. ereal (integral\<^sup>L M (s i))) \<longlonglongrightarrow> ereal (integral\<^sup>L M f)"
-          unfolding lim_ereal
-        proof (rule integral_dominated_convergence[where w=f])
-          show "integrable M f" by fact
-          from s_le_f seq show "\<And>i. AE x in M. norm (s i x) \<le> f x"
-            by auto
-        qed (insert seq, auto)
-        have int_s: "\<And>i. integrable M (s i)"
-          using seq f s_le_f by (intro integrable_bound[OF f(3)]) auto
-        have "(\<lambda>i. \<integral>\<^sup>+ x. s i x \<partial>M) \<longlonglongrightarrow> \<integral>\<^sup>+ x. f x \<partial>M"
-          using seq s_le_f f
-          by (intro nn_integral_dominated_convergence[where w=f])
-             (auto simp: integrable_iff_bounded)
-        also have "(\<lambda>i. \<integral>\<^sup>+x. s i x \<partial>M) = (\<lambda>i. \<integral>x. s i x \<partial>M)"
-          using seq int_s by simp
-        finally show "(\<lambda>i. \<integral>x. s i x \<partial>M) \<longlonglongrightarrow> \<integral>\<^sup>+x. f x \<partial>M"
-          by simp
+        have U_le_f: "x \<in> space M \<Longrightarrow> U i x \<le> f x" for x i
+          using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def)
+        have int_U: "\<And>i. integrable M (U i)"
+          using seq f U_le_f by (intro integrable_bound[OF f(3)]) auto
+        from U_le_f seq have "(\<lambda>i. integral\<^sup>L M (U i)) \<longlonglongrightarrow> integral\<^sup>L M f"
+          by (intro integral_dominated_convergence) auto
+        then show "(\<lambda>i. ennreal (integral\<^sup>L M (U i))) \<longlonglongrightarrow> ennreal (integral\<^sup>L M f)"
+          using seq f int_U by (simp add: f integral_nonneg_AE)
+        have "(\<lambda>i. \<integral>\<^sup>+ x. U i x \<partial>M) \<longlonglongrightarrow> \<integral>\<^sup>+ x. f x \<partial>M"
+          using seq U_le_f f
+          by (intro nn_integral_dominated_convergence[where w=f]) (auto simp: integrable_iff_bounded)
+        then show "(\<lambda>i. \<integral>x. U i x \<partial>M) \<longlonglongrightarrow> \<integral>\<^sup>+x. f x \<partial>M"
+          using seq int_U by simp
       qed
     qed }
   from this[of "\<lambda>x. max 0 (f x)"] assms have "(\<integral>\<^sup>+ x. max 0 (f x) \<partial>M) = integral\<^sup>L M (\<lambda>x. max 0 (f x))"
     by simp
   also have "\<dots> = integral\<^sup>L M f"
-    using assms by (auto intro!: integral_cong_AE)
+    using assms by (auto intro!: integral_cong_AE simp: integral_nonneg_AE)
   also have "(\<integral>\<^sup>+ x. max 0 (f x) \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M)"
     using assms by (auto intro!: nn_integral_cong_AE simp: max_def)
   finally show ?thesis .
@@ -1650,18 +1692,18 @@
 proof -
   have 1: "integrable M (\<lambda>x. \<Sum>i. norm (f i x))"
   proof (rule integrableI_bounded)
-    have "(\<integral>\<^sup>+ x. ereal (norm (\<Sum>i. norm (f i x))) \<partial>M) = (\<integral>\<^sup>+ x. (\<Sum>i. ereal (norm (f i x))) \<partial>M)"
-      apply (intro nn_integral_cong_AE) 
+    have "(\<integral>\<^sup>+ x. ennreal (norm (\<Sum>i. norm (f i x))) \<partial>M) = (\<integral>\<^sup>+ x. (\<Sum>i. ennreal (norm (f i x))) \<partial>M)"
+      apply (intro nn_integral_cong_AE)
       using summable
       apply eventually_elim
-      apply (simp add: suminf_ereal' suminf_nonneg)
+      apply (simp add: suminf_nonneg ennreal_suminf_neq_top)
       done
     also have "\<dots> = (\<Sum>i. \<integral>\<^sup>+ x. norm (f i x) \<partial>M)"
       by (intro nn_integral_suminf) auto
-    also have "\<dots> = (\<Sum>i. ereal (\<integral>x. norm (f i x) \<partial>M))"
+    also have "\<dots> = (\<Sum>i. ennreal (\<integral>x. norm (f i x) \<partial>M))"
       by (intro arg_cong[where f=suminf] ext nn_integral_eq_integral integrable_norm integrable) auto
-    finally show "(\<integral>\<^sup>+ x. ereal (norm (\<Sum>i. norm (f i x))) \<partial>M) < \<infinity>"
-      by (simp add: suminf_ereal' sums)
+    finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<Sum>i. norm (f i x))) \<partial>M) < \<infinity>"
+      by (simp add: sums ennreal_suminf_neq_top less_top[symmetric] integral_nonneg_AE)
   qed simp
 
   have 2: "AE x in M. (\<lambda>n. \<Sum>i<n. f i x) \<longlonglongrightarrow> (\<Sum>i. f i x)"
@@ -1693,43 +1735,57 @@
   fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
   shows "integrable M f \<Longrightarrow> norm (integral\<^sup>L M f) \<le> (\<integral>x. norm (f x) \<partial>M)"
   using nn_integral_eq_integral[of M "\<lambda>x. norm (f x)"]
-  using integral_norm_bound_ereal[of M f] by simp
-  
-lemma integrableI_nn_integral_finite:
-  assumes [measurable]: "f \<in> borel_measurable M"
-    and nonneg: "AE x in M. 0 \<le> f x"
-    and finite: "(\<integral>\<^sup>+x. f x \<partial>M) = ereal x"
-  shows "integrable M f"
-proof (rule integrableI_bounded)
-  have "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) = (\<integral>\<^sup>+ x. ereal (f x) \<partial>M)"
-    using nonneg by (intro nn_integral_cong_AE) auto
-  with finite show "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) < \<infinity>"
-    by auto
-qed simp
+  using integral_norm_bound_ennreal[of M f] by (simp add: integral_nonneg_AE)
 
 lemma integral_eq_nn_integral:
   assumes [measurable]: "f \<in> borel_measurable M"
   assumes nonneg: "AE x in M. 0 \<le> f x"
-  shows "integral\<^sup>L M f = real_of_ereal (\<integral>\<^sup>+ x. ereal (f x) \<partial>M)"
+  shows "integral\<^sup>L M f = enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M)"
 proof cases
-  assume *: "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) = \<infinity>"
-  also have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) = (\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M)"
+  assume *: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) = \<infinity>"
+  also have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
     using nonneg by (intro nn_integral_cong_AE) auto
   finally have "\<not> integrable M f"
     by (auto simp: integrable_iff_bounded)
   then show ?thesis
     by (simp add: * not_integrable_integral_eq)
 next
-  assume "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity>"
+  assume "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>"
   then have "integrable M f"
-    by (cases "\<integral>\<^sup>+ x. ereal (f x) \<partial>M") (auto intro!: integrableI_nn_integral_finite assms)
-  from nn_integral_eq_integral[OF this nonneg] show ?thesis
-    by simp
+    by (cases "\<integral>\<^sup>+ x. ennreal (f x) \<partial>M" rule: ennreal_cases)
+       (auto intro!: integrableI_nn_integral_finite assms)
+  from nn_integral_eq_integral[OF this] nonneg show ?thesis
+    by (simp add: integral_nonneg_AE)
 qed
-  
+
+lemma enn2real_nn_integral_eq_integral:
+  assumes eq: "AE x in M. f x = ennreal (g x)" and nn: "AE x in M. 0 \<le> g x"
+    and fin: "(\<integral>\<^sup>+x. f x \<partial>M) < top"
+    and [measurable]: "g \<in> M \<rightarrow>\<^sub>M borel"
+  shows "enn2real (\<integral>\<^sup>+x. f x \<partial>M) = (\<integral>x. g x \<partial>M)"
+proof -
+  have "ennreal (enn2real (\<integral>\<^sup>+x. f x \<partial>M)) = (\<integral>\<^sup>+x. f x \<partial>M)"
+    using fin by (intro ennreal_enn2real) auto
+  also have "\<dots> = (\<integral>\<^sup>+x. g x \<partial>M)"
+    using eq by (rule nn_integral_cong_AE)
+  also have "\<dots> = (\<integral>x. g x \<partial>M)"
+  proof (rule nn_integral_eq_integral)
+    show "integrable M g"
+    proof (rule integrableI_bounded)
+      have "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M)"
+        using eq nn by (auto intro!: nn_integral_cong_AE elim!: eventually_elim2)
+      also note fin
+      finally show "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) < \<infinity>"
+        by simp
+    qed simp
+  qed fact
+  finally show ?thesis
+    using nn by (simp add: integral_nonneg_AE)
+qed
+
 lemma has_bochner_integral_nn_integral:
-  assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
-  assumes "(\<integral>\<^sup>+x. f x \<partial>M) = ereal x"
+  assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "0 \<le> x"
+  assumes "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x"
   shows "has_bochner_integral M f x"
   unfolding has_bochner_integral_iff
   using assms by (auto simp: assms integral_eq_nn_integral intro: integrableI_nn_integral_finite)
@@ -1738,7 +1794,7 @@
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
   shows "simple_bochner_integrable M f \<Longrightarrow> integrable M f"
   by (intro integrableI_sequence[where s="\<lambda>_. f"] borel_measurable_simple_function)
-     (auto simp: zero_ereal_def[symmetric] simple_bochner_integrable.simps)
+     (auto simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps)
 
 lemma integrable_induct[consumes 1, case_names base add lim, induct pred: integrable]:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
@@ -1757,7 +1813,7 @@
     "\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)"
     unfolding norm_conv_dist by metis
 
-  { fix f A 
+  { fix f A
     have [simp]: "P (\<lambda>x. 0)"
       using base[of "{}" undefined] by simp
     have "(\<And>i::'b. i \<in> A \<Longrightarrow> integrable M (f i::'a \<Rightarrow> 'b)) \<Longrightarrow>
@@ -1773,7 +1829,7 @@
     unfolding s'_def using s(1)
     by (intro simple_function_compose2[where h="op *\<^sub>R"] simple_function_indicator) auto
 
-  { fix i 
+  { fix i
     have "\<And>z. {y. s' i z = y \<and> y \<in> s' i ` space M \<and> y \<noteq> 0 \<and> z \<in> space M} =
         (if z \<in> space M \<and> s' i z \<noteq> 0 then {s' i z} else {})"
       by (auto simp add: s'_def split: split_indicator)
@@ -1785,10 +1841,10 @@
   proof (rule lim)
     fix i
 
-    have "(\<integral>\<^sup>+x. norm (s' i x) \<partial>M) \<le> (\<integral>\<^sup>+x. 2 * ereal (norm (f x)) \<partial>M)"
+    have "(\<integral>\<^sup>+x. norm (s' i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal (2 * norm (f x)) \<partial>M)"
       using s by (intro nn_integral_mono) (auto simp: s'_eq_s)
     also have "\<dots> < \<infinity>"
-      using f by (subst nn_integral_cmult) auto
+      using f by (simp add: nn_integral_cmult ennreal_mult_less_top ennreal_mult)
     finally have sbi: "simple_bochner_integrable M (s' i)"
       using sf by (intro simple_bochner_integrableI_bounded) auto
     then show "integrable M (s' i)"
@@ -1798,10 +1854,10 @@
       then have "emeasure M {y \<in> space M. s' i y = s' i x} \<le> emeasure M {y \<in> space M. s' i y \<noteq> 0}"
         by (intro emeasure_mono) auto
       also have "\<dots> < \<infinity>"
-        using sbi by (auto elim: simple_bochner_integrable.cases)
+        using sbi by (auto elim: simple_bochner_integrable.cases simp: less_top)
       finally have "emeasure M {y \<in> space M. s' i y = s' i x} \<noteq> \<infinity>" by simp }
     then show "P (s' i)"
-      by (subst s'_eq) (auto intro!: setsum base)
+      by (subst s'_eq) (auto intro!: setsum base simp: less_top)
 
     fix x assume "x \<in> space M" with s show "(\<lambda>i. s' i x) \<longlonglongrightarrow> f x"
       by (simp add: s'_eq_s)
@@ -1810,20 +1866,6 @@
   qed fact
 qed
 
-lemma integral_nonneg_AE:
-  fixes f :: "'a \<Rightarrow> real"
-  assumes [measurable]: "AE x in M. 0 \<le> f x"
-  shows "0 \<le> integral\<^sup>L M f"
-proof cases
-  assume [measurable]: "integrable M f"
-  then have "0 \<le> ereal (integral\<^sup>L M (\<lambda>x. max 0 (f x)))"
-    by (subst integral_eq_nn_integral) (auto intro: real_of_ereal_pos nn_integral_nonneg assms)
-  also have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) = integral\<^sup>L M f"
-    using assms by (intro integral_cong_AE assms integrable_max) auto
-  finally show ?thesis
-    by simp
-qed (simp add: not_integrable_integral_eq)
-
 lemma integral_eq_zero_AE:
   "(AE x in M. f x = 0) \<Longrightarrow> integral\<^sup>L M f = 0"
   using integral_cong_AE[of f M "\<lambda>_. 0"]
@@ -1837,7 +1879,7 @@
   assume "integral\<^sup>L M f = 0"
   then have "integral\<^sup>N M f = 0"
     using nn_integral_eq_integral[OF f nonneg] by simp
-  then have "AE x in M. ereal (f x) \<le> 0"
+  then have "AE x in M. ennreal (f x) \<le> 0"
     by (simp add: nn_integral_0_iff_AE)
   with nonneg show "AE x in M. f x = 0"
     by auto
@@ -1857,11 +1899,11 @@
 
 lemma integral_mono:
   fixes f :: "'a \<Rightarrow> real"
-  shows "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x \<le> g x) \<Longrightarrow> 
+  shows "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x \<le> g x) \<Longrightarrow>
     integral\<^sup>L M f \<le> integral\<^sup>L M g"
   by (intro integral_mono_AE) auto
 
-lemma (in finite_measure) integrable_measure: 
+lemma (in finite_measure) integrable_measure:
   assumes I: "disjoint_family_on X I" "countable I"
   shows "integrable (count_space I) (\<lambda>i. measure M (X i))"
 proof -
@@ -1871,31 +1913,41 @@
     using I unfolding emeasure_eq_measure[symmetric]
     by (subst emeasure_UN_countable) (auto simp: disjoint_family_on_def)
   finally show ?thesis
-    by (auto intro!: integrableI_bounded simp: measure_nonneg)
+    by (auto intro!: integrableI_bounded)
 qed
 
 lemma integrableI_real_bounded:
   assumes f: "f \<in> borel_measurable M" and ae: "AE x in M. 0 \<le> f x" and fin: "integral\<^sup>N M f < \<infinity>"
   shows "integrable M f"
 proof (rule integrableI_bounded)
-  have "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) = \<integral>\<^sup>+ x. ereal (f x) \<partial>M"
+  have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) = \<integral>\<^sup>+ x. ennreal (f x) \<partial>M"
     using ae by (auto intro: nn_integral_cong_AE)
   also note fin
-  finally show "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) < \<infinity>" .
+  finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" .
 qed fact
 
 lemma integral_real_bounded:
-  assumes "AE x in M. 0 \<le> f x" "integral\<^sup>N M f \<le> ereal r"
+  assumes "0 \<le> r" "integral\<^sup>N M f \<le> ennreal r"
   shows "integral\<^sup>L M f \<le> r"
 proof cases
-  assume "integrable M f" from nn_integral_eq_integral[OF this] assms show ?thesis
+  assume [simp]: "integrable M f"
+
+  have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) = integral\<^sup>N M (\<lambda>x. max 0 (f x))"
+    by (intro nn_integral_eq_integral[symmetric]) auto
+  also have "\<dots> = integral\<^sup>N M f"
+    by (intro nn_integral_cong) (simp add: max_def ennreal_neg)
+  also have "\<dots> \<le> r"
+    by fact
+  finally have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) \<le> r"
+    using \<open>0 \<le> r\<close> by simp
+
+  moreover have "integral\<^sup>L M f \<le> integral\<^sup>L M (\<lambda>x. max 0 (f x))"
+    by (rule integral_mono_AE) auto
+  ultimately show ?thesis
     by simp
 next
-  assume "\<not> integrable M f"
-  moreover have "0 \<le> ereal r"
-    using nn_integral_nonneg assms(2) by (rule order_trans)
-  ultimately show ?thesis
-    by (simp add: not_integrable_integral_eq)
+  assume "\<not> integrable M f" then show ?thesis
+    using \<open>0 \<le> r\<close> by (simp add: not_integrable_integral_eq)
 qed
 
 subsection \<open>Restricted measure spaces\<close>
@@ -1907,7 +1959,7 @@
   unfolding integrable_iff_bounded
     borel_measurable_restrict_space_iff[OF \<Omega>]
     nn_integral_restrict_space[OF \<Omega>]
-  by (simp add: ac_simps ereal_indicator times_ereal.simps(1)[symmetric] del: times_ereal.simps)
+  by (simp add: ac_simps ennreal_indicator ennreal_mult)
 
 lemma integral_restrict_space:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
@@ -1929,13 +1981,12 @@
     proof (rule LIMSEQ_unique)
       show "(\<lambda>i. integral\<^sup>L (restrict_space M \<Omega>) (s i)) \<longlonglongrightarrow> integral\<^sup>L (restrict_space M \<Omega>) f"
         using lim by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) simp_all
-      
+
       show "(\<lambda>i. integral\<^sup>L (restrict_space M \<Omega>) (s i)) \<longlonglongrightarrow> (\<integral> x. indicator \<Omega> x *\<^sub>R f x \<partial>M)"
         unfolding lim
         using lim
         by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (indicator \<Omega> x *\<^sub>R f x)"])
-           (auto simp add: space_restrict_space integrable_restrict_space
-                 simp del: norm_scaleR
+           (auto simp add: space_restrict_space integrable_restrict_space simp del: norm_scaleR
                  split: split_indicator)
     qed
   qed
@@ -1958,9 +2009,9 @@
     and nn: "AE x in M. 0 \<le> g x"
   shows "integrable (density M g) f \<longleftrightarrow> integrable M (\<lambda>x. g x *\<^sub>R f x)"
   unfolding integrable_iff_bounded using nn
-  apply (simp add: nn_integral_density )
+  apply (simp add: nn_integral_density less_top[symmetric])
   apply (intro arg_cong2[where f="op ="] refl nn_integral_cong_AE)
-  apply auto
+  apply (auto simp: ennreal_mult)
   done
 
 lemma integral_density:
@@ -1974,21 +2025,25 @@
   proof induct
     case (base A c)
     then have [measurable]: "A \<in> sets M" by auto
-  
+
     have int: "integrable M (\<lambda>x. g x * indicator A x)"
       using g base integrable_density[of "indicator A :: 'a \<Rightarrow> real" M g] by simp
-    then have "integral\<^sup>L M (\<lambda>x. g x * indicator A x) = (\<integral>\<^sup>+ x. ereal (g x * indicator A x) \<partial>M)"
+    then have "integral\<^sup>L M (\<lambda>x. g x * indicator A x) = (\<integral>\<^sup>+ x. ennreal (g x * indicator A x) \<partial>M)"
       using g by (subst nn_integral_eq_integral) auto
-    also have "\<dots> = (\<integral>\<^sup>+ x. ereal (g x) * indicator A x \<partial>M)"
+    also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (g x) * indicator A x \<partial>M)"
       by (intro nn_integral_cong) (auto split: split_indicator)
     also have "\<dots> = emeasure (density M g) A"
       by (rule emeasure_density[symmetric]) auto
-    also have "\<dots> = ereal (measure (density M g) A)"
-      using base by (auto intro: emeasure_eq_ereal_measure)
+    also have "\<dots> = ennreal (measure (density M g) A)"
+      using base by (auto intro: emeasure_eq_ennreal_measure)
     also have "\<dots> = integral\<^sup>L (density M g) (indicator A)"
       using base by simp
     finally show ?case
-      using base by (simp add: int)
+      using base g
+      apply (simp add: int integral_nonneg_AE)
+      apply (subst (asm) ennreal_inj)
+      apply (auto intro!: integral_nonneg_AE)
+      done
   next
     case (add f h)
     then have [measurable]: "f \<in> borel_measurable M" "h \<in> borel_measurable M"
@@ -1999,7 +2054,7 @@
     case (lim f s)
     have [measurable]: "f \<in> borel_measurable M" "\<And>i. s i \<in> borel_measurable M"
       using lim(1,5)[THEN borel_measurable_integrable] by auto
-  
+
     show ?case
     proof (rule LIMSEQ_unique)
       show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. g x *\<^sub>R s i x)) \<longlonglongrightarrow> integral\<^sup>L M (\<lambda>x. g x *\<^sub>R f x)"
@@ -2058,7 +2113,7 @@
     then have [measurable]: "A \<in> sets N" by auto
     from base have int: "integrable (distr M N g) (\<lambda>a. indicator A a *\<^sub>R c)"
       by (intro integrable_indicator)
-  
+
     have "integral\<^sup>L (distr M N g) (\<lambda>a. indicator A a *\<^sub>R c) = measure (distr M N g) A *\<^sub>R c"
       using base by auto
     also have "\<dots> = measure M (g -` A \<inter> space M) *\<^sub>R c"
@@ -2078,13 +2133,13 @@
     case (lim f s)
     have [measurable]: "f \<in> borel_measurable N" "\<And>i. s i \<in> borel_measurable N"
       using lim(1,5)[THEN borel_measurable_integrable] by auto
-  
+
     show ?case
     proof (rule LIMSEQ_unique)
       show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. s i (g x))) \<longlonglongrightarrow> integral\<^sup>L M (\<lambda>x. f (g x))"
       proof (rule integral_dominated_convergence)
         show "integrable M (\<lambda>x. 2 * norm (f (g x)))"
-          using lim by (auto simp: integrable_distr_eq) 
+          using lim by (auto simp: integrable_distr_eq)
         show "AE x in M. (\<lambda>i. s i (g x)) \<longlonglongrightarrow> f (g x)"
           using lim(3) g[THEN measurable_space] by auto
         show "\<And>i. AE x in M. norm (s i (g x)) \<le> 2 * norm (f (g x))"
@@ -2121,7 +2176,7 @@
 proof -
   have eq: "\<And>x. x \<in> A \<Longrightarrow> (\<Sum>a | x = a \<and> a \<in> A \<and> f a \<noteq> 0. f a) = (\<Sum>x\<in>{x}. f x)"
     by (intro setsum.mono_neutral_cong_left) auto
-    
+
   have "(\<integral>x. f x \<partial>count_space A) = (\<integral>x. (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. indicator {a} x *\<^sub>R f a) \<partial>count_space A)"
     by (intro integral_cong refl) (simp add: f eq)
   also have "\<dots> = (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. measure (count_space A) {a} *\<^sub>R f a)"
@@ -2137,7 +2192,8 @@
 lemma integrable_count_space_nat_iff:
   fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
   shows "integrable (count_space UNIV) f \<longleftrightarrow> summable (\<lambda>x. norm (f x))"
-  by (auto simp add: integrable_iff_bounded nn_integral_count_space_nat summable_ereal suminf_ereal_finite)
+  by (auto simp add: integrable_iff_bounded nn_integral_count_space_nat ennreal_suminf_neq_top
+           intro:  summable_suminf_not_top)
 
 lemma sums_integral_count_space_nat:
   fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
@@ -2181,7 +2237,8 @@
   fixes g :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and f :: "'a \<Rightarrow> real"
   shows "finite A \<Longrightarrow> integrable (point_measure A f) g"
   unfolding point_measure_def
-  apply (subst density_ereal_max_0)
+  apply (subst density_cong[where f'="\<lambda>x. ennreal (max 0 (f x))"])
+  apply (auto split: split_max simp: ennreal_neg)
   apply (subst integrable_density)
   apply (auto simp: AE_count_space integrable_count_space)
   done
@@ -2204,67 +2261,60 @@
 
 lemma real_lebesgue_integral_def:
   assumes f[measurable]: "integrable M f"
-  shows "integral\<^sup>L M f = real_of_ereal (\<integral>\<^sup>+x. f x \<partial>M) - real_of_ereal (\<integral>\<^sup>+x. - f x \<partial>M)"
+  shows "integral\<^sup>L M f = enn2real (\<integral>\<^sup>+x. f x \<partial>M) - enn2real (\<integral>\<^sup>+x. ennreal (- f x) \<partial>M)"
 proof -
   have "integral\<^sup>L M f = integral\<^sup>L M (\<lambda>x. max 0 (f x) - max 0 (- f x))"
     by (auto intro!: arg_cong[where f="integral\<^sup>L M"])
   also have "\<dots> = integral\<^sup>L M (\<lambda>x. max 0 (f x)) - integral\<^sup>L M (\<lambda>x. max 0 (- f x))"
     by (intro integral_diff integrable_max integrable_minus integrable_zero f)
-  also have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) = real_of_ereal (\<integral>\<^sup>+x. max 0 (f x) \<partial>M)"
-    by (subst integral_eq_nn_integral[symmetric]) auto
-  also have "integral\<^sup>L M (\<lambda>x. max 0 (- f x)) = real_of_ereal (\<integral>\<^sup>+x. max 0 (- f x) \<partial>M)"
-    by (subst integral_eq_nn_integral[symmetric]) auto
-  also have "(\<lambda>x. ereal (max 0 (f x))) = (\<lambda>x. max 0 (ereal (f x)))"
-    by (auto simp: max_def)
-  also have "(\<lambda>x. ereal (max 0 (- f x))) = (\<lambda>x. max 0 (- ereal (f x)))"
-    by (auto simp: max_def)
-  finally show ?thesis
-    unfolding nn_integral_max_0 .
+  also have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) = enn2real (\<integral>\<^sup>+x. ennreal (f x) \<partial>M)"
+    by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg)
+  also have "integral\<^sup>L M (\<lambda>x. max 0 (- f x)) = enn2real (\<integral>\<^sup>+x. ennreal (- f x) \<partial>M)"
+    by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg)
+  finally show ?thesis .
 qed
 
 lemma real_integrable_def:
   "integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and>
-    (\<integral>\<^sup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
+    (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>"
   unfolding integrable_iff_bounded
 proof (safe del: notI)
-  assume *: "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) < \<infinity>"
-  have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M)"
+  assume *: "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>"
+  have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
     by (intro nn_integral_mono) auto
   also note *
-  finally show "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity>"
+  finally show "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>"
     by simp
-  have "(\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M)"
+  have "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
     by (intro nn_integral_mono) auto
   also note *
-  finally show "(\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
+  finally show "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>"
     by simp
 next
   assume [measurable]: "f \<in> borel_measurable M"
-  assume fin: "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
-  have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) = (\<integral>\<^sup>+ x. max 0 (ereal (f x)) + max 0 (ereal (- f x)) \<partial>M)"
-    by (intro nn_integral_cong) (auto simp: max_def)
-  also have"\<dots> = (\<integral>\<^sup>+ x. max 0 (ereal (f x)) \<partial>M) + (\<integral>\<^sup>+ x. max 0 (ereal (- f x)) \<partial>M)"
+  assume fin: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>"
+  have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (f x) + ennreal (- f x) \<partial>M)"
+    by (intro nn_integral_cong) (auto simp: abs_real_def ennreal_neg)
+  also have"\<dots> = (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) + (\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M)"
     by (intro nn_integral_add) auto
   also have "\<dots> < \<infinity>"
-    using fin by (auto simp: nn_integral_max_0)
+    using fin by (auto simp: less_top)
   finally show "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) < \<infinity>" .
 qed
 
 lemma integrableD[dest]:
   assumes "integrable M f"
-  shows "f \<in> borel_measurable M" "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
+  shows "f \<in> borel_measurable M" "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>"
   using assms unfolding real_integrable_def by auto
 
 lemma integrableE:
   assumes "integrable M f"
   obtains r q where
-    "(\<integral>\<^sup>+x. ereal (f x)\<partial>M) = ereal r"
-    "(\<integral>\<^sup>+x. ereal (-f x)\<partial>M) = ereal q"
+    "(\<integral>\<^sup>+x. ennreal (f x)\<partial>M) = ennreal r"
+    "(\<integral>\<^sup>+x. ennreal (-f x)\<partial>M) = ennreal q"
     "f \<in> borel_measurable M" "integral\<^sup>L M f = r - q"
   using assms unfolding real_integrable_def real_lebesgue_integral_def[OF assms]
-  using nn_integral_nonneg[of M "\<lambda>x. ereal (f x)"]
-  using nn_integral_nonneg[of M "\<lambda>x. ereal (-f x)"]
-  by (cases rule: ereal2_cases[of "(\<integral>\<^sup>+x. ereal (-f x)\<partial>M)" "(\<integral>\<^sup>+x. ereal (f x)\<partial>M)"]) auto
+  by (cases rule: ennreal2_cases[of "(\<integral>\<^sup>+x. ennreal (-f x)\<partial>M)" "(\<integral>\<^sup>+x. ennreal (f x)\<partial>M)"]) auto
 
 lemma integral_monotone_convergence_nonneg:
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> real"
@@ -2276,34 +2326,33 @@
   shows "integrable M u"
   and "integral\<^sup>L M u = x"
 proof -
-  have "(\<integral>\<^sup>+ x. ereal (u x) \<partial>M) = (SUP n. (\<integral>\<^sup>+ x. ereal (f n x) \<partial>M))"
+  have nn: "AE x in M. \<forall>i. 0 \<le> f i x"
+    using pos unfolding AE_all_countable by auto
+  with lim have u_nn: "AE x in M. 0 \<le> u x"
+    by eventually_elim (auto intro: LIMSEQ_le_const)
+  have [simp]: "0 \<le> x"
+    by (intro LIMSEQ_le_const[OF ilim] allI exI impI integral_nonneg_AE pos)
+  have "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = (SUP n. (\<integral>\<^sup>+ x. ennreal (f n x) \<partial>M))"
   proof (subst nn_integral_monotone_convergence_SUP_AE[symmetric])
     fix i
-    from mono pos show "AE x in M. ereal (f i x) \<le> ereal (f (Suc i) x) \<and> 0 \<le> ereal (f i x)"
+    from mono nn show "AE x in M. ennreal (f i x) \<le> ennreal (f (Suc i) x)"
       by eventually_elim (auto simp: mono_def)
-    show "(\<lambda>x. ereal (f i x)) \<in> borel_measurable M"
+    show "(\<lambda>x. ennreal (f i x)) \<in> borel_measurable M"
       using i by auto
   next
-    show "(\<integral>\<^sup>+ x. ereal (u x) \<partial>M) = \<integral>\<^sup>+ x. (SUP i. ereal (f i x)) \<partial>M"
+    show "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = \<integral>\<^sup>+ x. (SUP i. ennreal (f i x)) \<partial>M"
       apply (rule nn_integral_cong_AE)
-      using lim mono
-      by eventually_elim (simp add: SUP_eq_LIMSEQ[THEN iffD2])
+      using lim mono nn u_nn
+      apply eventually_elim
+      apply (simp add: LIMSEQ_unique[OF _ LIMSEQ_SUP] incseq_def)
+      done
   qed
-  also have "\<dots> = ereal x"
-    using mono i unfolding nn_integral_eq_integral[OF i pos]
-    by (subst SUP_eq_LIMSEQ) (auto simp: mono_def intro!: integral_mono_AE ilim)
-  finally have "(\<integral>\<^sup>+ x. ereal (u x) \<partial>M) = ereal x" .
-  moreover have "(\<integral>\<^sup>+ x. ereal (- u x) \<partial>M) = 0"
-  proof (subst nn_integral_0_iff_AE)
-    show "(\<lambda>x. ereal (- u x)) \<in> borel_measurable M"
-      using u by auto
-    from mono pos[of 0] lim show "AE x in M. ereal (- u x) \<le> 0"
-    proof eventually_elim
-      fix x assume "mono (\<lambda>n. f n x)" "0 \<le> f 0 x" "(\<lambda>i. f i x) \<longlonglongrightarrow> u x"
-      then show "ereal (- u x) \<le> 0"
-        using incseq_le[of "\<lambda>n. f n x" "u x" 0] by (simp add: mono_def incseq_def)
-    qed
-  qed
+  also have "\<dots> = ennreal x"
+    using mono i nn unfolding nn_integral_eq_integral[OF i pos]
+    by (subst LIMSEQ_unique[OF LIMSEQ_SUP]) (auto simp: mono_def integral_nonneg_AE pos intro!: integral_mono_AE ilim)
+  finally have "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = ennreal x" .
+  moreover have "(\<integral>\<^sup>+ x. ennreal (- u x) \<partial>M) = 0"
+    using u u_nn by (subst nn_integral_0_iff_AE) (auto simp add: ennreal_neg)
   ultimately show "integrable M u" "integral\<^sup>L M u = x"
     by (auto simp: real_integrable_def real_lebesgue_integral_def u)
 qed
@@ -2348,7 +2397,7 @@
     using f by (intro nn_integral_eq_integral integrable_norm) auto
   then have "(\<integral>x. norm (f x) \<partial>M) = 0 \<longleftrightarrow> (\<integral>\<^sup>+x. norm (f x) \<partial>M) = 0"
     by simp
-  also have "\<dots> \<longleftrightarrow> emeasure M {x\<in>space M. ereal (norm (f x)) \<noteq> 0} = 0"
+  also have "\<dots> \<longleftrightarrow> emeasure M {x\<in>space M. ennreal (norm (f x)) \<noteq> 0} = 0"
     by (intro nn_integral_0_iff) auto
   finally show ?thesis
     by simp
@@ -2360,15 +2409,15 @@
   using integral_norm_eq_0_iff[of M f] by simp
 
 lemma (in finite_measure) integrable_const[intro!, simp]: "integrable M (\<lambda>x. a)"
-  using integrable_indicator[of "space M" M a] by (simp cong: integrable_cong)
+  using integrable_indicator[of "space M" M a] by (simp cong: integrable_cong add: less_top[symmetric])
 
-lemma lebesgue_integral_const[simp]: 
+lemma lebesgue_integral_const[simp]:
   fixes a :: "'a :: {banach, second_countable_topology}"
   shows "(\<integral>x. a \<partial>M) = measure M (space M) *\<^sub>R a"
 proof -
   { assume "emeasure M (space M) = \<infinity>" "a \<noteq> 0"
     then have ?thesis
-      by (simp add: not_integrable_integral_eq measure_def integrable_iff_bounded) }
+      by (auto simp add: not_integrable_integral_eq ennreal_mult_less_top measure_def integrable_iff_bounded) }
   moreover
   { assume "a = 0" then have ?thesis by simp }
   moreover
@@ -2378,7 +2427,7 @@
     have "(\<integral>x. a \<partial>M) = (\<integral>x. indicator (space M) x *\<^sub>R a \<partial>M)"
       by (intro integral_cong) auto
     also have "\<dots> = measure M (space M) *\<^sub>R a"
-      by simp
+      by (simp add: less_top[symmetric])
     finally have ?thesis . }
   ultimately show ?thesis by blast
 qed
@@ -2409,18 +2458,16 @@
   finally show ?thesis .
 qed
 
-lemma (in finite_measure) ereal_integral_real:
-  assumes [measurable]: "f \<in> borel_measurable M" 
-  assumes ae: "AE x in M. 0 \<le> f x" "AE x in M. f x \<le> ereal B"
-  shows "ereal (\<integral>x. real_of_ereal (f x) \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)"
+lemma (in finite_measure) ennreal_integral_real:
+  assumes [measurable]: "f \<in> borel_measurable M"
+  assumes ae: "AE x in M. f x \<le> ennreal B" "0 \<le> B"
+  shows "ennreal (\<integral>x. enn2real (f x) \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)"
 proof (subst nn_integral_eq_integral[symmetric])
-  show "integrable M (\<lambda>x. real_of_ereal (f x))"
-    using ae by (intro integrable_const_bound[where B=B]) (auto simp: real_le_ereal_iff)
-  show "AE x in M. 0 \<le> real_of_ereal (f x)"
-    using ae by (auto simp: real_of_ereal_pos)
-  show "(\<integral>\<^sup>+ x. ereal (real_of_ereal (f x)) \<partial>M) = integral\<^sup>N M f"
-    using ae by (intro nn_integral_cong_AE) (auto simp: ereal_real)
-qed
+  show "integrable M (\<lambda>x. enn2real (f x))"
+    using ae by (intro integrable_const_bound[where B=B]) (auto simp: enn2real_leI enn2real_nonneg)
+  show "(\<integral>\<^sup>+ x. ennreal (enn2real (f x)) \<partial>M) = integral\<^sup>N M f"
+    using ae by (intro nn_integral_cong_AE) (auto simp: le_less_trans[OF _ ennreal_less_top])
+qed (auto simp: enn2real_nonneg)
 
 lemma (in finite_measure) integral_less_AE:
   fixes X Y :: "'a \<Rightarrow> real"
@@ -2447,7 +2494,7 @@
     then have "(emeasure M) A \<le> (emeasure M) {x \<in> space M. Y x - X x \<noteq> 0}"
       using int A by (simp add: integrable_def)
     ultimately have "emeasure M A = 0"
-      using emeasure_nonneg[of M A] by simp
+      by simp
     with \<open>(emeasure M) A \<noteq> 0\<close> show False by auto
   qed
   ultimately show ?thesis by auto
@@ -2473,7 +2520,7 @@
     show "AE x in M. (\<lambda>n. indicator {..X n} x *\<^sub>R f x) \<longlonglongrightarrow> f x"
     proof
       fix x
-      from \<open>filterlim X at_top sequentially\<close> 
+      from \<open>filterlim X at_top sequentially\<close>
       have "eventually (\<lambda>n. x \<le> X n) sequentially"
         unfolding filterlim_at_top_ge[where c=x] by auto
       then show "(\<lambda>n. indicator {..X n} x *\<^sub>R f x) \<longlonglongrightarrow> f x"
@@ -2528,13 +2575,13 @@
     by (simp cong: measurable_cong)
 qed
 
+lemma Collect_subset [simp]: "{x\<in>A. P x} \<subseteq> A" by auto
+
 lemma (in sigma_finite_measure) measurable_measure[measurable (raw)]:
   "(\<And>x. x \<in> space N \<Longrightarrow> A x \<subseteq> space M) \<Longrightarrow>
     {x \<in> space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} \<in> sets (N \<Otimes>\<^sub>M M) \<Longrightarrow>
     (\<lambda>x. measure M (A x)) \<in> borel_measurable N"
-  unfolding measure_def by (intro measurable_emeasure borel_measurable_real_of_ereal)
-
-lemma Collect_subset [simp]: "{x\<in>A. P x} \<subseteq> A" by auto
+  unfolding measure_def by (intro measurable_emeasure borel_measurable_enn2real) auto
 
 lemma (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]:
   fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}"
@@ -2593,17 +2640,17 @@
             using simple_functionD(1)[OF s(1), of i] x
             by (intro simple_function_borel_measurable)
                (auto simp: space_pair_measure dest: finite_subset)
-          have "(\<integral>\<^sup>+ y. ereal (norm (s i (x, y))) \<partial>M) \<le> (\<integral>\<^sup>+ y. 2 * norm (f x y) \<partial>M)"
+          have "(\<integral>\<^sup>+ y. ennreal (norm (s i (x, y))) \<partial>M) \<le> (\<integral>\<^sup>+ y. 2 * norm (f x y) \<partial>M)"
             using x s by (intro nn_integral_mono) auto
           also have "(\<integral>\<^sup>+ y. 2 * norm (f x y) \<partial>M) < \<infinity>"
             using int_2f by (simp add: integrable_iff_bounded)
-          finally show "(\<integral>\<^sup>+ xa. ereal (norm (s i (x, xa))) \<partial>M) < \<infinity>" .
+          finally show "(\<integral>\<^sup>+ xa. ennreal (norm (s i (x, xa))) \<partial>M) < \<infinity>" .
         qed
         then have "integral\<^sup>L M (\<lambda>y. s i (x, y)) = simple_bochner_integral M (\<lambda>y. s i (x, y))"
           by (rule simple_bochner_integrable_eq_integral[symmetric]) }
       ultimately have "(\<lambda>i. simple_bochner_integral M (\<lambda>y. s i (x, y))) \<longlonglongrightarrow> integral\<^sup>L M (f x)"
         by simp }
-    then 
+    then
     show "(\<lambda>i. f' i x) \<longlonglongrightarrow> integral\<^sup>L M (f x)"
       unfolding f'_def
       by (cases "integrable M (f x)") (simp_all add: not_integrable_integral_eq)
@@ -2657,11 +2704,11 @@
     fix x assume "integrable M2 (\<lambda>y. f (x, y))"
     then have f: "integrable M2 (\<lambda>y. norm (f (x, y)))"
       by simp
-    then have "(\<integral>\<^sup>+y. ereal (norm (f (x, y))) \<partial>M2) = ereal (LINT y|M2. norm (f (x, y)))"
+    then have "(\<integral>\<^sup>+y. ennreal (norm (f (x, y))) \<partial>M2) = ennreal (LINT y|M2. norm (f (x, y)))"
       by (rule nn_integral_eq_integral) simp
-    also have "\<dots> = ereal \<bar>LINT y|M2. norm (f (x, y))\<bar>"
-      using f by (auto intro!: abs_of_nonneg[symmetric] integral_nonneg_AE)
-    finally show "(\<integral>\<^sup>+y. ereal (norm (f (x, y))) \<partial>M2) = ereal \<bar>LINT y|M2. norm (f (x, y))\<bar>" .
+    also have "\<dots> = ennreal \<bar>LINT y|M2. norm (f (x, y))\<bar>"
+      using f by simp
+    finally show "(\<integral>\<^sup>+y. ennreal (norm (f (x, y))) \<partial>M2) = ennreal \<bar>LINT y|M2. norm (f (x, y))\<bar>" .
   qed
   also have "\<dots> < \<infinity>"
     using integ1 by (simp add: integrable_iff_bounded integral_nonneg_AE)
@@ -2679,7 +2726,7 @@
     by (rule nn_integral_PInf_AE[rotated]) (intro M2.measurable_emeasure_Pair A)
   moreover have "\<And>x. x \<in> space M1 \<Longrightarrow> Pair x -` A = {y\<in>space M2. (x, y) \<in> A}"
     using sets.sets_into_space[OF A] by (auto simp: space_pair_measure)
-  ultimately show ?thesis by auto
+  ultimately show ?thesis by (auto simp: less_top)
 qed
 
 lemma (in pair_sigma_finite) AE_integrable_fst':
@@ -2696,7 +2743,7 @@
        (auto simp: measurable_split_conv)
   with AE_space show ?thesis
     by eventually_elim
-       (auto simp: integrable_iff_bounded measurable_compose[OF _ borel_measurable_integrable[OF f]])
+       (auto simp: integrable_iff_bounded measurable_compose[OF _ borel_measurable_integrable[OF f]] less_top)
 qed
 
 lemma (in pair_sigma_finite) integrable_fst':
@@ -2707,13 +2754,13 @@
 proof
   show "(\<lambda>x. \<integral> y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
     by (rule M2.borel_measurable_lebesgue_integral) simp
-  have "(\<integral>\<^sup>+ x. ereal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) \<le> (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<partial>M1)"
-    using AE_integrable_fst'[OF f] by (auto intro!: nn_integral_mono_AE integral_norm_bound_ereal)
+  have "(\<integral>\<^sup>+ x. ennreal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) \<le> (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<partial>M1)"
+    using AE_integrable_fst'[OF f] by (auto intro!: nn_integral_mono_AE integral_norm_bound_ennreal)
   also have "(\<integral>\<^sup>+x. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<partial>M1) = (\<integral>\<^sup>+x. norm (f x) \<partial>(M1 \<Otimes>\<^sub>M M2))"
     by (rule M2.nn_integral_fst) simp
   also have "(\<integral>\<^sup>+x. norm (f x) \<partial>(M1 \<Otimes>\<^sub>M M2)) < \<infinity>"
     using f unfolding integrable_iff_bounded by simp
-  finally show "(\<integral>\<^sup>+ x. ereal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) < \<infinity>" .
+  finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) < \<infinity>" .
 qed
 
 lemma (in pair_sigma_finite) integral_fst':
@@ -2738,24 +2785,24 @@
   qed
   also have "\<dots> = measure (M1 \<Otimes>\<^sub>M M2) A *\<^sub>R c"
   proof (subst integral_scaleR_left)
-    have "(\<integral>\<^sup>+x. ereal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1) =
+    have "(\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1) =
       (\<integral>\<^sup>+x. emeasure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1)"
       using emeasure_pair_measure_finite[OF base]
-      by (intro nn_integral_cong_AE, eventually_elim) (simp add: emeasure_eq_ereal_measure)
+      by (intro nn_integral_cong_AE, eventually_elim) (simp add: emeasure_eq_ennreal_measure)
     also have "\<dots> = emeasure (M1 \<Otimes>\<^sub>M M2) A"
       using sets.sets_into_space[OF A]
       by (subst M2.emeasure_pair_measure_alt)
          (auto intro!: nn_integral_cong arg_cong[where f="emeasure M2"] simp: space_pair_measure)
-    finally have *: "(\<integral>\<^sup>+x. ereal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1) = emeasure (M1 \<Otimes>\<^sub>M M2) A" .
+    finally have *: "(\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1) = emeasure (M1 \<Otimes>\<^sub>M M2) A" .
 
     from base * show "integrable M1 (\<lambda>x. measure M2 {y \<in> space M2. (x, y) \<in> A})"
-      by (simp add: measure_nonneg integrable_iff_bounded)
-    then have "(\<integral>x. measure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1) = 
-      (\<integral>\<^sup>+x. ereal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1)"
-      by (rule nn_integral_eq_integral[symmetric]) (simp add: measure_nonneg)
+      by (simp add: integrable_iff_bounded)
+    then have "(\<integral>x. measure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1) =
+      (\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1)"
+      by (rule nn_integral_eq_integral[symmetric]) simp
     also note *
     finally show "(\<integral>x. measure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1) *\<^sub>R c = measure (M1 \<Otimes>\<^sub>M M2) A *\<^sub>R c"
-      using base by (simp add: emeasure_eq_ereal_measure)
+      using base by (simp add: emeasure_eq_ennreal_measure)
   qed
   also have "\<dots> = (\<integral> a. indicator A a *\<^sub>R c \<partial>(M1 \<Otimes>\<^sub>M M2))"
     using base by simp
@@ -2764,14 +2811,14 @@
   case (add f g)
   then have [measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" "g \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
     by auto
-  have "(\<integral> x. \<integral> y. f (x, y) + g (x, y) \<partial>M2 \<partial>M1) = 
+  have "(\<integral> x. \<integral> y. f (x, y) + g (x, y) \<partial>M2 \<partial>M1) =
     (\<integral> x. (\<integral> y. f (x, y) \<partial>M2) + (\<integral> y. g (x, y) \<partial>M2) \<partial>M1)"
     apply (rule integral_cong_AE)
     apply simp_all
     using AE_integrable_fst'[OF add(1)] AE_integrable_fst'[OF add(3)]
     apply eventually_elim
     apply simp
-    done 
+    done
   also have "\<dots> = (\<integral> x. f x \<partial>(M1 \<Otimes>\<^sub>M M2)) + (\<integral> x. g x \<partial>(M1 \<Otimes>\<^sub>M M2))"
     using integrable_fst'[OF add(1)] integrable_fst'[OF add(3)] add(2,4) by simp
   finally show ?case
@@ -2780,7 +2827,7 @@
   case (lim f s)
   then have [measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" "\<And>i. s i \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
     by auto
-  
+
   show ?case
   proof (rule LIMSEQ_unique)
     show "(\<lambda>i. integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (s i)) \<longlonglongrightarrow> integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f"
@@ -2811,11 +2858,11 @@
         by (intro integrable_mult_right integrable_norm integrable_fst' lim)
       fix i show "AE x in M1. norm (\<integral> y. s i (x, y) \<partial>M2) \<le> (\<integral> y. 2 * norm (f (x, y)) \<partial>M2)"
         using AE_space AE_integrable_fst'[OF lim(1), of i] AE_integrable_fst'[OF lim(5)]
-      proof eventually_elim 
+      proof eventually_elim
         fix x assume x: "x \<in> space M1"
           and s: "integrable M2 (\<lambda>y. s i (x, y))" and f: "integrable M2 (\<lambda>y. f (x, y))"
         from s have "norm (\<integral> y. s i (x, y) \<partial>M2) \<le> (\<integral>\<^sup>+y. norm (s i (x, y)) \<partial>M2)"
-          by (rule integral_norm_bound_ereal)
+          by (rule integral_norm_bound_ennreal)
         also have "\<dots> \<le> (\<integral>\<^sup>+y. 2 * norm (f (x, y)) \<partial>M2)"
           using x lim by (auto intro!: nn_integral_mono simp: space_pair_measure)
         also have "\<dots> = (\<integral>y. 2 * norm (f (x, y)) \<partial>M2)"
@@ -2927,14 +2974,14 @@
 
   show "?f \<in> borel_measurable (Pi\<^sub>M I M)"
     using assms by simp
-  have "(\<integral>\<^sup>+ x. ereal (norm (\<Prod>i\<in>I. f i (x i))) \<partial>Pi\<^sub>M I M) = 
-      (\<integral>\<^sup>+ x. (\<Prod>i\<in>I. ereal (norm (f i (x i)))) \<partial>Pi\<^sub>M I M)"
-    by (simp add: setprod_norm setprod_ereal)
-  also have "\<dots> = (\<Prod>i\<in>I. \<integral>\<^sup>+ x. ereal (norm (f i x)) \<partial>M i)"
+  have "(\<integral>\<^sup>+ x. ennreal (norm (\<Prod>i\<in>I. f i (x i))) \<partial>Pi\<^sub>M I M) =
+      (\<integral>\<^sup>+ x. (\<Prod>i\<in>I. ennreal (norm (f i (x i)))) \<partial>Pi\<^sub>M I M)"
+    by (simp add: setprod_norm setprod_ennreal)
+  also have "\<dots> = (\<Prod>i\<in>I. \<integral>\<^sup>+ x. ennreal (norm (f i x)) \<partial>M i)"
     using assms by (intro product_nn_integral_setprod) auto
   also have "\<dots> < \<infinity>"
-    using integrable by (simp add: setprod_PInf nn_integral_nonneg integrable_iff_bounded)
-  finally show "(\<integral>\<^sup>+ x. ereal (norm (\<Prod>i\<in>I. f i (x i))) \<partial>Pi\<^sub>M I M) < \<infinity>" .
+    using integrable by (simp add: less_top[symmetric] ennreal_setprod_eq_top integrable_iff_bounded)
+  finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<Prod>i\<in>I. f i (x i))) \<partial>Pi\<^sub>M I M) < \<infinity>" .
 qed
 
 lemma (in product_sigma_finite) product_integral_setprod:
@@ -3010,8 +3057,7 @@
   qed
 qed (simp add: not_integrable_integral_eq integrable_subalgebra[OF assms])
 
-hide_const simple_bochner_integral
-hide_const simple_bochner_integrable
+hide_const (open) simple_bochner_integral
+hide_const (open) simple_bochner_integrable
 
 end
-
--- a/src/HOL/Probability/Borel_Space.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Borel_Space.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -435,6 +435,14 @@
   by (subst measurable_restrict_space_iff)
      (auto simp: indicator_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_cong)
 
+lemma borel_measurable_restrict_space_iff_ennreal:
+  fixes f :: "'a \<Rightarrow> ennreal"
+  assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"
+  shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>
+    (\<lambda>x. f x * indicator \<Omega> x) \<in> borel_measurable M"
+  by (subst measurable_restrict_space_iff)
+     (auto simp: indicator_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_cong)
+
 lemma borel_measurable_restrict_space_iff:
   fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
   assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"
@@ -1639,12 +1647,6 @@
   from assms show ?thesis unfolding * by simp
 qed
 
-lemma
-  fixes f :: "'a \<Rightarrow> ereal" assumes f: "f \<in> borel_measurable M"
-  shows borel_measurable_ereal_eq_const: "{x\<in>space M. f x = c} \<in> sets M"
-    and borel_measurable_ereal_neq_const: "{x\<in>space M. f x \<noteq> c} \<in> sets M"
-  using f by auto
-
 lemma [measurable(raw)]:
   fixes f :: "'a \<Rightarrow> ereal"
   assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
@@ -1689,6 +1691,10 @@
 lemma measurable_e2ennreal[measurable]: "e2ennreal \<in> borel \<rightarrow>\<^sub>M borel"
   by (intro borel_measurable_continuous_on1 continuous_on_e2ennreal)
 
+lemma borel_measurable_enn2real[measurable (raw)]:
+  "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. enn2real (f x)) \<in> M \<rightarrow>\<^sub>M borel"
+  unfolding enn2real_def[abs_def] by measurable
+
 definition [simp]: "is_borel f M \<longleftrightarrow> f \<in> borel_measurable M"
 
 lemma is_borel_transfer[transfer_rule]: "rel_fun (rel_fun op = pcr_ennreal) op = is_borel is_borel"
@@ -1699,8 +1705,24 @@
     using measurable_compose[OF f measurable_e2ennreal] by simp
 qed simp
 
+context
+  includes ennreal.lifting
+begin
+
 lemma measurable_ennreal[measurable]: "ennreal \<in> borel \<rightarrow>\<^sub>M borel"
-  unfolding is_borel_def[symmetric] by transfer simp
+  unfolding is_borel_def[symmetric]
+  by transfer simp
+
+lemma borel_measurable_ennreal_iff[simp]:
+  assumes [simp]: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
+  shows "(\<lambda>x. ennreal (f x)) \<in> M \<rightarrow>\<^sub>M borel \<longleftrightarrow> f \<in> M \<rightarrow>\<^sub>M borel"
+proof safe
+  assume "(\<lambda>x. ennreal (f x)) \<in> M \<rightarrow>\<^sub>M borel"
+  then have "(\<lambda>x. enn2real (ennreal (f x))) \<in> M \<rightarrow>\<^sub>M borel"
+    by measurable
+  then show "f \<in> M \<rightarrow>\<^sub>M borel"
+    by (rule measurable_cong[THEN iffD1, rotated]) auto
+qed measurable
 
 lemma borel_measurable_times_ennreal[measurable (raw)]:
   fixes f g :: "'a \<Rightarrow> ennreal"
@@ -1728,6 +1750,8 @@
   shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
   using assms by (induction S rule: infinite_finite_induct) auto
 
+end
+
 hide_const (open) is_borel
 
 subsection \<open>LIMSEQ is borel measurable\<close>
--- a/src/HOL/Probability/Caratheodory.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Caratheodory.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -13,9 +13,8 @@
   Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson.
 \<close>
 
-lemma suminf_ereal_2dimen:
-  fixes f:: "nat \<times> nat \<Rightarrow> ereal"
-  assumes pos: "\<And>p. 0 \<le> f p"
+lemma suminf_ennreal_2dimen:
+  fixes f:: "nat \<times> nat \<Rightarrow> ennreal"
   assumes "\<And>m. g m = (\<Sum>n. f (m,n))"
   shows "(\<Sum>i. f (prod_decode i)) = suminf g"
 proof -
@@ -23,46 +22,42 @@
     using assms by (simp add: fun_eq_iff)
   have reindex: "\<And>B. (\<Sum>x\<in>B. f (prod_decode x)) = setsum f (prod_decode ` B)"
     by (simp add: setsum.reindex[OF inj_prod_decode] comp_def)
-  { fix n
+  have "(SUP n. \<Sum>i<n. f (prod_decode i)) = (SUP p : UNIV \<times> UNIV. \<Sum>i<fst p. \<Sum>n<snd p. f (i, n))"
+  proof (intro SUP_eq; clarsimp simp: setsum.cartesian_product reindex)
+    fix n
     let ?M = "\<lambda>f. Suc (Max (f ` prod_decode ` {..<n}))"
     { fix a b x assume "x < n" and [symmetric]: "(a, b) = prod_decode x"
       then have "a < ?M fst" "b < ?M snd"
         by (auto intro!: Max_ge le_imp_less_Suc image_eqI) }
     then have "setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<?M fst} \<times> {..<?M snd})"
-      by (auto intro!: setsum_mono3 simp: pos)
-    then have "\<exists>a b. setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<a} \<times> {..<b})" by auto }
-  moreover
-  { fix a b
+      by (auto intro!: setsum_mono3)
+    then show "\<exists>a b. setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<a} \<times> {..<b})" by auto
+  next
+    fix a b
     let ?M = "prod_decode ` {..<Suc (Max (prod_encode ` ({..<a} \<times> {..<b})))}"
     { fix a' b' assume "a' < a" "b' < b" then have "(a', b') \<in> ?M"
         by (auto intro!: Max_ge le_imp_less_Suc image_eqI[where x="prod_encode (a', b')"]) }
     then have "setsum f ({..<a} \<times> {..<b}) \<le> setsum f ?M"
-      by (auto intro!: setsum_mono3 simp: pos) }
-  ultimately
-  show ?thesis unfolding g_def using pos
-    by (auto intro!: SUP_eq  simp: setsum.cartesian_product reindex SUP_upper2
-                     suminf_ereal_eq_SUP SUP_pair
-                     SUP_ereal_setsum[symmetric] incseq_setsumI setsum_nonneg)
+      by (auto intro!: setsum_mono3)
+    then show "\<exists>n. setsum f ({..<a} \<times> {..<b}) \<le> setsum f (prod_decode ` {..<n})"
+      by auto
+  qed
+  also have "\<dots> = (SUP p. \<Sum>i<p. \<Sum>n. f (i, n))"
+    unfolding suminf_setsum[OF summableI, symmetric]
+    by (simp add: suminf_eq_SUP SUP_pair setsum.commute[of _ "{..< fst _}"])
+  finally show ?thesis unfolding g_def
+    by (simp add: suminf_eq_SUP)
 qed
 
 subsection \<open>Characterizations of Measures\<close>
 
-definition subadditive where
-  "subadditive M f \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
-
-definition countably_subadditive where
-  "countably_subadditive M f \<longleftrightarrow>
-    (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))))"
-
 definition outer_measure_space where
   "outer_measure_space M f \<longleftrightarrow> positive M f \<and> increasing M f \<and> countably_subadditive M f"
 
-lemma subadditiveD: "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) \<le> f x + f y"
-  by (auto simp add: subadditive_def)
-
 subsubsection \<open>Lambda Systems\<close>
 
-definition lambda_system where
+definition lambda_system :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set set"
+where
   "lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> M. f (l \<inter> x) + f ((\<Omega> - l) \<inter> x) = f x}"
 
 lemma (in algebra) lambda_system_eq:
@@ -81,7 +76,7 @@
   by (simp add: lambda_system_def)
 
 lemma (in algebra) lambda_system_Compl:
-  fixes f:: "'a set \<Rightarrow> ereal"
+  fixes f:: "'a set \<Rightarrow> ennreal"
   assumes x: "x \<in> lambda_system \<Omega> M f"
   shows "\<Omega> - x \<in> lambda_system \<Omega> M f"
 proof -
@@ -94,7 +89,7 @@
 qed
 
 lemma (in algebra) lambda_system_Int:
-  fixes f:: "'a set \<Rightarrow> ereal"
+  fixes f:: "'a set \<Rightarrow> ennreal"
   assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
   shows "x \<inter> y \<in> lambda_system \<Omega> M f"
 proof -
@@ -128,7 +123,7 @@
 qed
 
 lemma (in algebra) lambda_system_Un:
-  fixes f:: "'a set \<Rightarrow> ereal"
+  fixes f:: "'a set \<Rightarrow> ennreal"
   assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
   shows "x \<union> y \<in> lambda_system \<Omega> M f"
 proof -
@@ -176,25 +171,6 @@
     by (simp add: Un)
 qed
 
-lemma (in ring_of_sets) countably_subadditive_subadditive:
-  assumes f: "positive M f" and cs: "countably_subadditive M f"
-  shows  "subadditive M f"
-proof (auto simp add: subadditive_def)
-  fix x y
-  assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
-  hence "disjoint_family (binaryset x y)"
-    by (auto simp add: disjoint_family_on_def binaryset_def)
-  hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
-         (\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
-         f (\<Union>i. binaryset x y i) \<le> (\<Sum> n. f (binaryset x y n))"
-    using cs by (auto simp add: countably_subadditive_def)
-  hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
-         f (x \<union> y) \<le> (\<Sum> n. f (binaryset x y n))"
-    by (simp add: range_binaryset_eq UN_binaryset_eq)
-  thus "f (x \<union> y) \<le>  f x + f y" using f x y
-    by (auto simp add: Un o_def suminf_binaryset_eq positive_def)
-qed
-
 lemma lambda_system_increasing: "increasing M f \<Longrightarrow> increasing (lambda_system \<Omega> M f) f"
   by (simp add: increasing_def lambda_system_def)
 
@@ -202,7 +178,7 @@
   by (simp add: positive_def lambda_system_def)
 
 lemma (in algebra) lambda_system_strong_sum:
-  fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ereal"
+  fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ennreal"
   assumes f: "positive M f" and a: "a \<in> M"
       and A: "range A \<subseteq> lambda_system \<Omega> M f"
       and disj: "disjoint_family A"
@@ -247,11 +223,10 @@
   proof (rule antisym)
     show "f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))"
       using csa[unfolded countably_subadditive_def] A'' disj U_in by auto
-    have *: "\<And>i. 0 \<le> f (A i)" using pos A'' unfolding positive_def by auto
     have dis: "\<And>N. disjoint_family_on A {..<N}" by (intro disjoint_family_on_mono[OF _ disj]) auto
     show "(\<Sum>i. f (A i)) \<le> f (\<Union>i. A i)"
       using ls.additive_sum [OF lambda_system_positive[OF pos] lambda_system_additive _ A' dis] A''
-      by (intro suminf_bound[OF _ *]) (auto intro!: increasingD[OF inc] countable_UN)
+      by (intro suminf_le_const[OF summableI]) (auto intro!: increasingD[OF inc] countable_UN)
   qed
   have "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) = f a"
     if a [iff]: "a \<in> M" for a
@@ -271,7 +246,7 @@
     hence "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> (\<Sum>i. f (a \<inter> A i)) + f (a - (\<Union>i. A i))"
       by (rule add_right_mono)
     also have "\<dots> \<le> f a"
-    proof (intro suminf_bound_add allI)
+    proof (intro ennreal_suminf_bound_add)
       fix n
       have UNION_in: "(\<Union>i\<in>{0..<n}. A i) \<in> M"
         by (metis A'' UNION_in_sets)
@@ -285,14 +260,9 @@
         by (blast intro: increasingD [OF inc] UNION_in U_in)
       thus "(\<Sum>i<n. f (a \<inter> A i)) + f (a - (\<Union>i. A i)) \<le> f a"
         by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric])
-    next
-      have "\<And>i. a \<inter> A i \<in> M" using A'' by auto
-      then show "\<And>i. 0 \<le> f (a \<inter> A i)" using pos[unfolded positive_def] by auto
-      have "\<And>i. a - (\<Union>i. A i) \<in> M" using A'' by auto
-      then have "\<And>i. 0 \<le> f (a - (\<Union>i. A i))" using pos[unfolded positive_def] by auto
-      then show "f (a - (\<Union>i. A i)) \<noteq> -\<infinity>" by auto
     qed
-    finally show "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> f a" .
+    finally show "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> f a"
+      by simp
   next
     have "f a \<le> f (a \<inter> (\<Union>i. A i) \<union> (a - (\<Union>i. A i)))"
       by (blast intro:  increasingD [OF inc] U_in)
@@ -327,7 +297,7 @@
     using pos by (simp add: measure_space_def)
 qed
 
-definition outer_measure :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> 'a set \<Rightarrow> ereal" where
+definition outer_measure :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set \<Rightarrow> ennreal" where
    "outer_measure M f X =
      (INF A:{A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i)}. \<Sum>i. f (A i))"
 
@@ -346,7 +316,7 @@
     by (intro ca[unfolded countably_additive_def, rule_format, symmetric])
        (auto simp: Int_absorb1 disjoint_family_on_def)
   also have "... \<le> (\<Sum>i. f (A i))"
-    using A s by (intro suminf_le_pos increasingD[OF inc] positiveD2[OF posf]) auto
+    using A s by (auto intro!: suminf_le increasingD[OF inc])
   finally show "f s \<le> (\<Sum>i. f (A i))" .
 next
   have "(\<Sum>i. f (if i = 0 then s else {})) \<le> f s"
@@ -356,20 +326,14 @@
        (auto simp: disjoint_family_on_def)
 qed
 
-lemma outer_measure_nonneg: "positive M f \<Longrightarrow> 0 \<le> outer_measure M f X"
-  by (auto intro!: INF_greatest suminf_0_le intro: positiveD2 simp: outer_measure_def)
-
 lemma outer_measure_empty:
-  assumes posf: "positive M f" and "{} \<in> M"
-  shows "outer_measure M f {} = 0"
-proof (rule antisym)
-  show "outer_measure M f {} \<le> 0"
-    using assms by (auto intro!: INF_lower2[of "\<lambda>_. {}"] simp: outer_measure_def disjoint_family_on_def positive_def)
-qed (intro outer_measure_nonneg posf)
+  "positive M f \<Longrightarrow> {} \<in> M \<Longrightarrow> outer_measure M f {} = 0"
+  unfolding outer_measure_def
+  by (intro antisym INF_lower2[of  "\<lambda>_. {}"]) (auto simp: disjoint_family_on_def positive_def)
 
 lemma (in ring_of_sets) positive_outer_measure:
   assumes "positive M f" shows "positive (Pow \<Omega>) (outer_measure M f)"
-  unfolding positive_def by (auto simp: assms outer_measure_empty outer_measure_nonneg)
+  unfolding positive_def by (auto simp: assms outer_measure_empty)
 
 lemma (in ring_of_sets) increasing_outer_measure: "increasing (Pow \<Omega>) (outer_measure M f)"
   by (force simp: increasing_def outer_measure_def intro!: INF_greatest intro: INF_lower)
@@ -383,22 +347,13 @@
     by (auto intro!: A range_disjointed_sets)
   have "\<forall>n. f (disjointed A n) \<le> f (A n)"
     by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A)
-  moreover have "\<forall>i. 0 \<le> f (disjointed A i)"
-    using pos dA unfolding positive_def by auto
-  ultimately show "(\<Sum>i. f (disjointed A i)) \<le> (\<Sum>i. f (A i))"
-    by (blast intro!: suminf_le_pos)
+  then show "(\<Sum>i. f (disjointed A i)) \<le> (\<Sum>i. f (A i))"
+    by (blast intro!: suminf_le)
 qed (auto simp: X UN_disjointed_eq disjoint_family_disjointed)
 
 lemma (in ring_of_sets) outer_measure_close:
-  assumes posf: "positive M f" and "0 < e" and "outer_measure M f X \<noteq> \<infinity>"
-  shows "\<exists>A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) \<le> outer_measure M f X + e"
-proof -
-  from \<open>outer_measure M f X \<noteq> \<infinity>\<close> have fin: "\<bar>outer_measure M f X\<bar> \<noteq> \<infinity>"
-    using outer_measure_nonneg[OF posf, of X] by auto
-  show ?thesis
-    using Inf_ereal_close [OF fin [unfolded outer_measure_def], OF \<open>0 < e\<close>]
-    by (auto intro: less_imp_le simp add: outer_measure_def)
-qed
+  "outer_measure M f X < e \<Longrightarrow> \<exists>A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) < e"
+  unfolding outer_measure_def INF_less_iff by auto
 
 lemma (in ring_of_sets) countably_subadditive_outer_measure:
   assumes posf: "positive M f" and inc: "increasing M f"
@@ -406,16 +361,16 @@
 proof (simp add: countably_subadditive_def, safe)
   fix A :: "nat \<Rightarrow> _" assume A: "range A \<subseteq> Pow (\<Omega>)" and sb: "(\<Union>i. A i) \<subseteq> \<Omega>"
   let ?O = "outer_measure M f"
-
-  { fix e :: ereal assume e: "0 < e" and "\<forall>i. ?O (A i) \<noteq> \<infinity>"
-    hence "\<exists>B. \<forall>n. range (B n) \<subseteq> M \<and> disjoint_family (B n) \<and> A n \<subseteq> (\<Union>i. B n i) \<and>
-        (\<Sum>i. f (B n i)) \<le> ?O (A n) + e * (1/2)^(Suc n)"
-      using e sb by (auto intro!: choice outer_measure_close [of f, OF posf] simp: ereal_zero_less_0_iff one_ereal_def)
-    then obtain B
+  show "?O (\<Union>i. A i) \<le> (\<Sum>n. ?O (A n))"
+  proof (rule ennreal_le_epsilon)
+    fix b and e :: real assume "0 < e" "(\<Sum>n. outer_measure M f (A n)) < top"
+    then have *: "\<And>n. outer_measure M f (A n) < outer_measure M f (A n) + e * (1/2)^Suc n"
+      by (auto simp add: less_top dest!: ennreal_suminf_lessD)
+    obtain B
       where B: "\<And>n. range (B n) \<subseteq> M"
       and sbB: "\<And>n. A n \<subseteq> (\<Union>i. B n i)"
       and Ble: "\<And>n. (\<Sum>i. f (B n i)) \<le> ?O (A n) + e * (1/2)^(Suc n)"
-      by auto blast
+      by (metis less_imp_le outer_measure_close[OF *])
 
     def C \<equiv> "case_prod B o prod_decode"
     from B have B_in_M: "\<And>i j. B i j \<in> M"
@@ -425,26 +380,22 @@
     have A_C: "(\<Union>i. A i) \<subseteq> (\<Union>i. C i)"
       using sbB by (auto simp add: C_def subset_eq) (metis prod.case prod_encode_inverse)
 
-    have "?O (\<Union>i. A i) \<le> ?O (\<Union>i. C i)"  
+    have "?O (\<Union>i. A i) \<le> ?O (\<Union>i. C i)"
       using A_C A C by (intro increasing_outer_measure[THEN increasingD]) (auto dest!: sets_into_space)
     also have "\<dots> \<le> (\<Sum>i. f (C i))"
       using C by (intro outer_measure_le[OF posf inc]) auto
     also have "\<dots> = (\<Sum>n. \<Sum>i. f (B n i))"
-      using B_in_M unfolding C_def comp_def by (intro suminf_ereal_2dimen positiveD2[OF posf]) auto
-    also have "\<dots> \<le> (\<Sum>n. ?O (A n) + e*(1/2) ^ Suc n)"
-      using B_in_M by (intro suminf_le_pos[OF Ble] suminf_0_le posf[THEN positiveD2]) auto
-    also have "... = (\<Sum>n. ?O (A n)) + (\<Sum>n. e*(1/2) ^ Suc n)"
-      using e by (subst suminf_add_ereal) (auto simp add: ereal_zero_le_0_iff outer_measure_nonneg posf)
-    also have "(\<Sum>n. e*(1/2) ^ Suc n) = e"
-      using suminf_half_series_ereal e by (simp add: ereal_zero_le_0_iff suminf_cmult_ereal)
-    finally have "?O (\<Union>i. A i) \<le> (\<Sum>n. ?O (A n)) + e" . }
-  note * = this
-
-  show "?O (\<Union>i. A i) \<le> (\<Sum>n. ?O (A n))"
-  proof cases
-    assume "\<forall>i. ?O (A i) \<noteq> \<infinity>" with * show ?thesis
-      by (intro ereal_le_epsilon) auto
-  qed (metis suminf_PInfty[OF outer_measure_nonneg, OF posf] ereal_less_eq(1))
+      using B_in_M unfolding C_def comp_def by (intro suminf_ennreal_2dimen) auto
+    also have "\<dots> \<le> (\<Sum>n. ?O (A n) + e * (1/2) ^ Suc n)"
+      using B_in_M by (intro suminf_le suminf_nonneg allI Ble) auto
+    also have "... = (\<Sum>n. ?O (A n)) + (\<Sum>n. ennreal e * ennreal ((1/2) ^ Suc n))"
+      using \<open>0 < e\<close> by (subst suminf_add[symmetric])
+                       (auto simp del: ennreal_suminf_cmult simp add: ennreal_mult[symmetric])
+    also have "\<dots> = (\<Sum>n. ?O (A n)) + e"
+      unfolding ennreal_suminf_cmult
+      by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto
+    finally show "?O (\<Union>i. A i) \<le> (\<Sum>n. ?O (A n)) + e" .
+  qed
 qed
 
 lemma (in ring_of_sets) outer_measure_space_outer_measure:
@@ -481,7 +432,7 @@
     ultimately have "outer_measure M f (s \<inter> x) + outer_measure M f (s - x) \<le>
         (\<Sum>i. f (A i \<inter> x)) + (\<Sum>i. f (A i - x))" by (rule add_mono)
     also have "\<dots> = (\<Sum>i. f (A i \<inter> x) + f (A i - x))"
-      using A(2) x posf by (subst suminf_add_ereal) (auto simp: positive_def)
+      using A(2) x posf by (subst suminf_add) (auto simp: positive_def)
     also have "\<dots> = (\<Sum>i. f (A i))"
       using A x
       by (subst add[THEN additiveD, symmetric])
@@ -496,7 +447,7 @@
     also have "... \<le> outer_measure M f (s \<inter> x) + outer_measure M f (s - x)"
       apply (rule subadditiveD)
       apply (rule ring_of_sets.countably_subadditive_subadditive [OF ring_of_sets_Pow])
-      apply (simp add: positive_def outer_measure_empty[OF posf] outer_measure_nonneg[OF posf])
+      apply (simp add: positive_def outer_measure_empty[OF posf])
       apply (rule countably_subadditive_outer_measure)
       using s by (auto intro!: posf inc)
     finally show ?thesis .
@@ -513,7 +464,7 @@
 
 theorem (in ring_of_sets) caratheodory':
   assumes posf: "positive M f" and ca: "countably_additive M f"
-  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
+  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
 proof -
   have inc: "increasing M f"
     by (metis additive_increasing ca countably_additive_additive posf)
@@ -541,14 +492,14 @@
 lemma (in ring_of_sets) caratheodory_empty_continuous:
   assumes f: "positive M f" "additive M f" and fin: "\<And>A. A \<in> M \<Longrightarrow> f A \<noteq> \<infinity>"
   assumes cont: "\<And>A. range A \<subseteq> M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
-  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
+  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
 proof (intro caratheodory' empty_continuous_imp_countably_additive f)
   show "\<forall>A\<in>M. f A \<noteq> \<infinity>" using fin by auto
 qed (rule cont)
 
 subsection \<open>Volumes\<close>
 
-definition volume :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
+definition volume :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
   "volume M f \<longleftrightarrow>
   (f {} = 0) \<and> (\<forall>a\<in>M. 0 \<le> f a) \<and>
   (\<forall>C\<subseteq>M. disjoint C \<longrightarrow> finite C \<longrightarrow> \<Union>C \<in> M \<longrightarrow> f (\<Union>C) = (\<Sum>c\<in>C. f c))"
@@ -569,7 +520,7 @@
   by (auto simp: volume_def)
 
 lemma volume_finite_additive:
-  assumes "volume M f" 
+  assumes "volume M f"
   assumes A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M" "disjoint_family_on A I" "finite I" "UNION I A \<in> M"
   shows "f (UNION I A) = (\<Sum>i\<in>I. f (A i))"
 proof -
@@ -590,7 +541,7 @@
 qed
 
 lemma (in ring_of_sets) volume_additiveI:
-  assumes pos: "\<And>a. a \<in> M \<Longrightarrow> 0 \<le> \<mu> a" 
+  assumes pos: "\<And>a. a \<in> M \<Longrightarrow> 0 \<le> \<mu> a"
   assumes [simp]: "\<mu> {} = 0"
   assumes add: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b = {} \<Longrightarrow> \<mu> (a \<union> b) = \<mu> a + \<mu> b"
   shows "volume M \<mu>"
@@ -614,7 +565,7 @@
   have "\<forall>a\<in>?R. \<exists>m. \<exists>C\<subseteq>M. a = \<Union>C \<and> finite C \<and> disjoint C \<and> m = (\<Sum>c\<in>C. \<mu> c)"
     by (auto simp: generated_ring_def)
   from bchoice[OF this] guess \<mu>' .. note \<mu>'_spec = this
-  
+
   { fix C assume C: "C \<subseteq> M" "finite C" "disjoint C"
     fix D assume D: "D \<subseteq> M" "finite D" "disjoint D"
     assume "\<Union>C = \<Union>D"
@@ -688,7 +639,7 @@
 
 theorem (in semiring_of_sets) caratheodory:
   assumes pos: "positive M \<mu>" and ca: "countably_additive M \<mu>"
-  shows "\<exists>\<mu>' :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> M. \<mu>' s = \<mu> s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>'"
+  shows "\<exists>\<mu>' :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu>' s = \<mu> s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>'"
 proof -
   have "volume M \<mu>"
   proof (rule volumeI)
@@ -787,7 +738,7 @@
         then have "\<Union>range f = A i"
           using f C Ai unfolding bij_betw_def
             by (auto simp add: f_def cong del: strong_SUP_cong)
-        moreover 
+        moreover
         { have "(\<Sum>j. \<mu>_r (f j)) = (\<Sum>j. if j \<in> {..< card C} then \<mu>_r (f j) else 0)"
             using volume_empty[OF V(1)] by (auto intro!: arg_cong[where f=suminf] simp: f_def)
           also have "\<dots> = (\<Sum>j<card C. \<mu>_r (f j))"
@@ -827,7 +778,7 @@
         qed
       qed
       from f have "(\<Sum>n. \<mu>_r (A n)) = (\<Sum>n. \<mu>_r (case_prod f (prod_decode n)))"
-        by (intro suminf_ereal_2dimen[symmetric] positiveD2[OF pos] generated_ringI_Basic)
+        by (intro suminf_ennreal_2dimen[symmetric] generated_ringI_Basic)
          (auto split: prod.split)
       also have "\<dots> = (\<Sum>n. \<mu> (case_prod f (prod_decode n)))"
         using f V(2) by (auto intro!: arg_cong[where f=suminf] split: prod.split)
@@ -847,8 +798,7 @@
                intro: generated_ringI_Basic)
     also have "\<dots> = (\<Sum>c\<in>C'. \<Sum>n. \<mu>_r (A' n \<inter> c))"
       using C' A'
-      by (intro suminf_setsum_ereal positiveD2[OF pos] G.Int G.finite_Union)
-         (auto intro: generated_ringI_Basic)
+      by (intro suminf_setsum G.Int G.finite_Union) (auto intro: generated_ringI_Basic)
     also have "\<dots> = (\<Sum>c\<in>C'. \<mu>_r c)"
       using eq V C' by (auto intro!: setsum.cong)
     also have "\<dots> = \<mu>_r (\<Union>C')"
@@ -909,7 +859,7 @@
       using \<mu>' by (simp_all add: M sets_extend_measure measure_space_def)
   qed fact
 qed
-  
+
 lemma extend_measure_caratheodory_pair:
   fixes G :: "'i \<Rightarrow> 'j \<Rightarrow> 'a set"
   assumes M: "M = extend_measure \<Omega> {(a, b). P a b} (\<lambda>(a, b). G a b) (\<lambda>(a, b). \<mu> a b)"
--- a/src/HOL/Probability/Complete_Measure.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Complete_Measure.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -174,7 +174,7 @@
   let ?\<mu> = "emeasure M \<circ> main_part M"
   show "S \<in> sets (completion M)" "?\<mu> S = emeasure M (main_part M S) " using S by simp_all
   show "positive (sets (completion M)) ?\<mu>"
-    by (simp add: positive_def emeasure_nonneg)
+    by (simp add: positive_def)
   show "countably_additive (sets (completion M)) ?\<mu>"
   proof (intro countably_additiveI)
     fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets (completion M)" "disjoint_family A"
@@ -266,9 +266,9 @@
   qed
 qed
 
-lemma completion_ex_borel_measurable_pos:
-  fixes g :: "'a \<Rightarrow> ereal"
-  assumes g: "g \<in> borel_measurable (completion M)" and "\<And>x. 0 \<le> g x"
+lemma completion_ex_borel_measurable:
+  fixes g :: "'a \<Rightarrow> ennreal"
+  assumes g: "g \<in> borel_measurable (completion M)"
   shows "\<exists>g'\<in>borel_measurable M. (AE x in M. g x = g' x)"
 proof -
   from g[THEN borel_measurable_implies_simple_function_sequence'] guess f . note f = this
@@ -284,7 +284,7 @@
     proof (elim AE_mp, safe intro!: AE_I2)
       fix x assume eq: "\<forall>i. f i x = f' i x"
       moreover have "g x = (SUP i. f i x)"
-        unfolding f using \<open>0 \<le> g x\<close> by (auto split: split_max)
+        unfolding f by (auto split: split_max)
       ultimately show "g x = ?f x" by auto
     qed
     show "?f \<in> borel_measurable M"
@@ -292,28 +292,6 @@
   qed
 qed
 
-lemma completion_ex_borel_measurable:
-  fixes g :: "'a \<Rightarrow> ereal"
-  assumes g: "g \<in> borel_measurable (completion M)"
-  shows "\<exists>g'\<in>borel_measurable M. (AE x in M. g x = g' x)"
-proof -
-  have "(\<lambda>x. max 0 (g x)) \<in> borel_measurable (completion M)" "\<And>x. 0 \<le> max 0 (g x)" using g by auto
-  from completion_ex_borel_measurable_pos[OF this] guess g_pos ..
-  moreover
-  have "(\<lambda>x. max 0 (- g x)) \<in> borel_measurable (completion M)" "\<And>x. 0 \<le> max 0 (- g x)" using g by auto
-  from completion_ex_borel_measurable_pos[OF this] guess g_neg ..
-  ultimately
-  show ?thesis
-  proof (safe intro!: bexI[of _ "\<lambda>x. g_pos x - g_neg x"])
-    show "AE x in M. max 0 (- g x) = g_neg x \<longrightarrow> max 0 (g x) = g_pos x \<longrightarrow> g x = g_pos x - g_neg x"
-    proof (intro AE_I2 impI)
-      fix x assume g: "max 0 (- g x) = g_neg x" "max 0 (g x) = g_pos x"
-      show "g x = g_pos x - g_neg x" unfolding g[symmetric]
-        by (cases "g x") (auto split: split_max)
-    qed
-  qed auto
-qed
-
 lemma (in prob_space) prob_space_completion: "prob_space (completion M)"
   by (rule prob_spaceI) (simp add: emeasure_space_1)
 
--- a/src/HOL/Probability/Convolution.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Convolution.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -40,7 +40,7 @@
   assumes [simp]: "sets N = sets borel" "sets M = sets borel"
   assumes [simp]: "space M = space N" "space N = space borel"
   shows "emeasure (M \<star> N) A = \<integral>\<^sup>+x. (emeasure N {a. a + x \<in> A}) \<partial>M "
-  using assms by (auto intro!: nn_integral_cong simp del: nn_integral_indicator simp: nn_integral_convolution 
+  using assms by (auto intro!: nn_integral_cong simp del: nn_integral_indicator simp: nn_integral_convolution
     nn_integral_indicator [symmetric] ac_simps split:split_indicator)
 
 lemma convolution_emeasure':
@@ -64,7 +64,7 @@
   assumes [simp]: "sets N = sets borel" "sets M = sets borel" "sets L = sets borel"
   shows "emeasure (L \<star> (M \<star> N )) A = \<integral>\<^sup>+x. \<integral>\<^sup>+y. \<integral>\<^sup>+z. indicator A (x + y + z) \<partial>N \<partial>M \<partial>L"
   apply (subst nn_integral_indicator[symmetric], simp)
-  apply (subst nn_integral_convolution, 
+  apply (subst nn_integral_convolution,
         auto intro!: borel_measurable_indicator borel_measurable_indicator' convolution_finite)+
   by (rule nn_integral_cong)+ (auto simp: semigroup_add_class.add.assoc)
 
@@ -84,7 +84,7 @@
   assumes [simp]: "finite_measure M" "finite_measure N"
   assumes [measurable_cong, simp]: "sets N = sets borel" "sets M = sets borel"
   shows "(M \<star> N) = (N \<star> M)"
-proof (rule measure_eqI)  
+proof (rule measure_eqI)
   interpret M: finite_measure M by fact
   interpret N: finite_measure N by fact
   interpret pair_sigma_finite M N ..
@@ -117,16 +117,15 @@
 lemma (in prob_space) sum_indep_random_variable_lborel:
   assumes ind: "indep_var borel X borel Y"
   assumes [simp, measurable]: "random_variable lborel X"
-  assumes [simp, measurable]:"random_variable lborel Y" 
+  assumes [simp, measurable]:"random_variable lborel Y"
   shows "distr M lborel (\<lambda>x. X x + Y x) = convolution (distr M lborel X)  (distr M lborel Y)"
-  using ind unfolding indep_var_distribution_eq convolution_def 
+  using ind unfolding indep_var_distribution_eq convolution_def
   by (auto simp: distr_distr o_def intro!: arg_cong[where f = "distr M borel"] cong: distr_cong)
 
 lemma convolution_density:
-  fixes f g :: "real \<Rightarrow> ereal"
+  fixes f g :: "real \<Rightarrow> ennreal"
   assumes [measurable]: "f \<in> borel_measurable borel" "g \<in> borel_measurable borel"
   assumes [simp]:"finite_measure (density lborel f)" "finite_measure (density lborel g)"
-  assumes gt_0[simp]: "AE x in lborel. 0 \<le> f x" "AE x in lborel. 0 \<le> g x"
   shows "density lborel f \<star> density lborel g = density lborel (\<lambda>x. \<integral>\<^sup>+y. f (x - y) * g y \<partial>lborel)"
     (is "?l = ?r")
 proof (intro measure_eqI)
@@ -136,10 +135,9 @@
 
   have "(\<integral>\<^sup>+x. f x * (\<integral>\<^sup>+y. g y * indicator A (x + y) \<partial>lborel) \<partial>lborel) =
     (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. g y * (f x * indicator A (x + y)) \<partial>lborel) \<partial>lborel)"
-    using gt_0
   proof (intro nn_integral_cong_AE, eventually_elim)
-    fix x assume "0 \<le> f x"
-    then have "f x * (\<integral>\<^sup>+ y. g y * indicator A (x + y) \<partial>lborel) =
+    fix x
+    have "f x * (\<integral>\<^sup>+ y. g y * indicator A (x + y) \<partial>lborel) =
       (\<integral>\<^sup>+ y. f x * (g y * indicator A (x + y)) \<partial>lborel)"
       by (intro nn_integral_cmult[symmetric]) auto
     then show "f x * (\<integral>\<^sup>+ y. g y * indicator A (x + y) \<partial>lborel) =
@@ -149,18 +147,16 @@
   also have "\<dots> = (\<integral>\<^sup>+y. (\<integral>\<^sup>+x. g y * (f x * indicator A (x + y)) \<partial>lborel) \<partial>lborel)"
     by (intro lborel_pair.Fubini') simp
   also have "\<dots> = (\<integral>\<^sup>+y. (\<integral>\<^sup>+x. f (x - y) * g y * indicator A x \<partial>lborel) \<partial>lborel)"
-    using gt_0
   proof (intro nn_integral_cong_AE, eventually_elim)
-    fix y assume "0 \<le> g y"
-    then have "(\<integral>\<^sup>+x. g y * (f x * indicator A (x + y)) \<partial>lborel) =
+    fix y
+    have "(\<integral>\<^sup>+x. g y * (f x * indicator A (x + y)) \<partial>lborel) =
       g y * (\<integral>\<^sup>+x. f x * indicator A (x + y) \<partial>lborel)"
       by (intro nn_integral_cmult) auto
     also have "\<dots> = g y * (\<integral>\<^sup>+x. f (x - y) * indicator A x \<partial>lborel)"
-      using gt_0
       by (subst nn_integral_real_affine[where c=1 and t="-y"])
-         (auto simp del: gt_0 simp add: one_ereal_def[symmetric])
+         (auto simp add: one_ennreal_def[symmetric])
     also have "\<dots> = (\<integral>\<^sup>+x. g y * (f (x - y) * indicator A x) \<partial>lborel)"
-      using \<open>0 \<le> g y\<close> by (intro nn_integral_cmult[symmetric]) auto
+      by (intro nn_integral_cmult[symmetric]) auto
     finally show "(\<integral>\<^sup>+ x. g y * (f x * indicator A (x + y)) \<partial>lborel) =
       (\<integral>\<^sup>+ x. f (x - y) * g y * indicator A x \<partial>lborel)"
       by (simp add: ac_simps)
@@ -175,7 +171,7 @@
 lemma (in prob_space) distributed_finite_measure_density:
   "distributed M N X f \<Longrightarrow> finite_measure (density N f)"
   using finite_measure_distr[of X N] distributed_distr_eq_density[of M N X f] by simp
-  
+
 
 lemma (in prob_space) distributed_convolution:
   fixes f :: "real \<Rightarrow> _"
@@ -186,13 +182,10 @@
   shows "distributed M lborel (\<lambda>x. X x + Y x) (\<lambda>x. \<integral>\<^sup>+y. f (x - y) * g y \<partial>lborel)"
   unfolding distributed_def
 proof safe
-  show "AE x in lborel. 0 \<le> \<integral>\<^sup>+ xa. f (x - xa) * g xa \<partial>lborel"
-    by (auto simp: nn_integral_nonneg)
-
   have fg[measurable]: "f \<in> borel_measurable borel" "g \<in> borel_measurable borel"
     using distributed_borel_measurable[OF X] distributed_borel_measurable[OF Y] by simp_all
-  
-  show "(\<lambda>x. \<integral>\<^sup>+ xa. f (x - xa) * g xa \<partial>lborel) \<in> borel_measurable lborel" 
+
+  show "(\<lambda>x. \<integral>\<^sup>+ xa. f (x - xa) * g xa \<partial>lborel) \<in> borel_measurable lborel"
     by measurable
 
   have "distr M borel (\<lambda>x. X x + Y x) = (distr M borel X \<star> distr M borel Y)"
@@ -207,10 +200,6 @@
       using X by (rule distributed_finite_measure_density)
     show "finite_measure (density lborel g)"
       using Y by (rule distributed_finite_measure_density)
-    show "AE x in lborel. 0 \<le> f x"
-      using X by (rule distributed_AE)
-    show "AE x in lborel. 0 \<le> g x"
-      using Y by (rule distributed_AE)
   qed fact+
   finally show "distr M lborel (\<lambda>x. X x + Y x) = density lborel (\<lambda>x. \<integral>\<^sup>+ y. f (x - y) * g y \<partial>lborel)"
     by (simp cong: distr_cong)
--- a/src/HOL/Probability/Distribution_Functions.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Distribution_Functions.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -6,11 +6,11 @@
 section \<open>Distribution Functions\<close>
 
 text \<open>
-Shows that the cumulative distribution function (cdf) of a distribution (a measure on the reals) is 
+Shows that the cumulative distribution function (cdf) of a distribution (a measure on the reals) is
 nondecreasing and right continuous, which tends to 0 and 1 in either direction.
 
-Conversely, every such function is the cdf of a unique distribution. This direction defines the 
-measure in the obvious way on half-open intervals, and then applies the Caratheodory extension 
+Conversely, every such function is the cdf of a unique distribution. This direction defines the
+measure in the obvious way on half-open intervals, and then applies the Caratheodory extension
 theorem.
 \<close>
 
@@ -43,7 +43,7 @@
 lemma sets_M[intro]: "a \<in> sets borel \<Longrightarrow> a \<in> sets M"
   using M_super_borel by auto
 
-lemma cdf_diff_eq: 
+lemma cdf_diff_eq:
   assumes "x < y"
   shows "cdf M y - cdf M x = measure M {x<..y}"
 proof -
@@ -59,7 +59,7 @@
 
 lemma borel_UNIV: "space M = UNIV"
  by (metis in_mono sets.sets_into_space space_in_borel top_le M_super_borel)
- 
+
 lemma cdf_nonneg: "cdf M x \<ge> 0"
   unfolding cdf_def by (rule measure_nonneg)
 
@@ -76,11 +76,11 @@
   finally show ?thesis .
 qed
 
-lemma cdf_lim_at_top: "(cdf M \<longlongrightarrow> measure M (space M)) at_top" 
+lemma cdf_lim_at_top: "(cdf M \<longlongrightarrow> measure M (space M)) at_top"
   by (rule tendsto_at_topI_sequentially_real)
      (simp_all add: mono_def cdf_nondecreasing cdf_lim_infty)
 
-lemma cdf_lim_neg_infty: "((\<lambda>i. cdf M (- real i)) \<longlonglongrightarrow> 0)" 
+lemma cdf_lim_neg_infty: "((\<lambda>i. cdf M (- real i)) \<longlonglongrightarrow> 0)"
 proof -
   have "(\<lambda>i. cdf M (- real i)) \<longlonglongrightarrow> measure M (\<Inter> i::nat. {.. - real i })"
     unfolding cdf_def by (rule finite_Lim_measure_decseq) (auto simp: decseq_def)
@@ -91,7 +91,7 @@
 qed
 
 lemma cdf_lim_at_bot: "(cdf M \<longlongrightarrow> 0) at_bot"
-proof - 
+proof -
   have *: "((\<lambda>x :: real. - cdf M (- x)) \<longlongrightarrow> 0) at_top"
     by (intro tendsto_at_topI_sequentially_real monoI)
        (auto simp: cdf_nondecreasing cdf_lim_neg_infty tendsto_minus_cancel_left[symmetric])
@@ -105,7 +105,7 @@
 proof (rule tendsto_at_right_sequentially[where b="a + 1"])
   fix f :: "nat \<Rightarrow> real" and x assume f: "decseq f" "f \<longlonglongrightarrow> a"
   then have "(\<lambda>n. cdf M (f n)) \<longlonglongrightarrow> measure M (\<Inter>i. {.. f i})"
-    using `decseq f` unfolding cdf_def 
+    using `decseq f` unfolding cdf_def
     by (intro finite_Lim_measure_decseq) (auto simp: decseq_def)
   also have "(\<Inter>i. {.. f i}) = {.. a}"
     using decseq_le[OF f] by (auto intro: order_trans LIMSEQ_le_const[OF f(2)])
@@ -117,7 +117,7 @@
 proof (rule tendsto_at_left_sequentially[of "a - 1"])
   fix f :: "nat \<Rightarrow> real" and x assume f: "incseq f" "f \<longlonglongrightarrow> a" "\<And>x. f x < a" "\<And>x. a - 1 < f x"
   then have "(\<lambda>n. cdf M (f n)) \<longlonglongrightarrow> measure M (\<Union>i. {.. f i})"
-    using `incseq f` unfolding cdf_def 
+    using `incseq f` unfolding cdf_def
     by (intro finite_Lim_measure_incseq) (auto simp: incseq_def)
   also have "(\<Union>i. {.. f i}) = {..<a}"
     by (auto dest!: order_tendstoD(1)[OF f(2)] eventually_happens'[OF sequentially_bot]
@@ -139,7 +139,7 @@
 
 lemma countable_atoms: "countable {x. measure M {x} > 0}"
   using countable_support unfolding zero_less_measure_iff .
-    
+
 end
 
 locale real_distribution = prob_space M for M :: "real measure" +
@@ -155,7 +155,7 @@
 lemma cdf_lim_infty_prob: "(\<lambda>i. cdf M (real i)) \<longlonglongrightarrow> 1"
   by (subst prob_space [symmetric], rule cdf_lim_infty)
 
-lemma cdf_lim_at_top_prob: "(cdf M \<longlongrightarrow> 1) at_top" 
+lemma cdf_lim_at_top_prob: "(cdf M \<longlongrightarrow> 1) at_top"
   by (subst prob_space [symmetric], rule cdf_lim_at_top)
 
 lemma measurable_finite_borel [simp]:
@@ -200,7 +200,7 @@
 lemma real_distribution_interval_measure:
   fixes F :: "real \<Rightarrow> real"
   assumes nondecF : "\<And> x y. x \<le> y \<Longrightarrow> F x \<le> F y" and
-    right_cont_F : "\<And>a. continuous (at_right a) F" and 
+    right_cont_F : "\<And>a. continuous (at_right a) F" and
     lim_F_at_bot : "(F \<longlongrightarrow> 0) at_bot" and
     lim_F_at_top : "(F \<longlongrightarrow> 1) at_top"
   shows "real_distribution (interval_measure F)"
@@ -208,12 +208,12 @@
   let ?F = "interval_measure F"
   interpret prob_space ?F
   proof
-    have "ereal (1 - 0) = (SUP i::nat. ereal (F (real i) - F (- real i)))"
-      by (intro LIMSEQ_unique[OF _ LIMSEQ_SUP] lim_ereal[THEN iffD2] tendsto_intros
-         lim_F_at_bot[THEN filterlim_compose] lim_F_at_top[THEN filterlim_compose]
-         lim_F_at_bot[THEN filterlim_compose] filterlim_real_sequentially
-         filterlim_uminus_at_top[THEN iffD1])
-         (auto simp: incseq_def intro!: diff_mono nondecF)
+    have "ennreal (1 - 0) = (SUP i::nat. ennreal (F (real i) - F (- real i)))"
+      by (intro LIMSEQ_unique[OF _ LIMSEQ_SUP] tendsto_ennrealI tendsto_intros
+                lim_F_at_bot[THEN filterlim_compose] lim_F_at_top[THEN filterlim_compose]
+                lim_F_at_bot[THEN filterlim_compose] filterlim_real_sequentially
+                filterlim_uminus_at_top[THEN iffD1])
+         (auto simp: incseq_def nondecF intro!: diff_mono)
     also have "\<dots> = (SUP i::nat. emeasure ?F {- real i<..real i})"
       by (subst emeasure_interval_measure_Ioc) (simp_all add: nondecF right_cont_F)
     also have "\<dots> = emeasure ?F (\<Union>i::nat. {- real i<..real i})"
@@ -230,7 +230,7 @@
 lemma cdf_interval_measure:
   fixes F :: "real \<Rightarrow> real"
   assumes nondecF : "\<And> x y. x \<le> y \<Longrightarrow> F x \<le> F y" and
-    right_cont_F : "\<And>a. continuous (at_right a) F" and 
+    right_cont_F : "\<And>a. continuous (at_right a) F" and
     lim_F_at_bot : "(F \<longlongrightarrow> 0) at_bot" and
     lim_F_at_top : "(F \<longlongrightarrow> 1) at_top"
   shows "cdf (interval_measure F) = F"
--- a/src/HOL/Probability/Distributions.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Distributions.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -10,7 +10,7 @@
 begin
 
 lemma (in prob_space) distributed_affine:
-  fixes f :: "real \<Rightarrow> ereal"
+  fixes f :: "real \<Rightarrow> ennreal"
   assumes f: "distributed M lborel X f"
   assumes c: "c \<noteq> 0"
   shows "distributed M lborel (\<lambda>x. t + c * X x) (\<lambda>x. f ((x - t) / c) / \<bar>c\<bar>)"
@@ -26,30 +26,26 @@
   show "random_variable lborel (\<lambda>x. t + c * X x)"
     by simp
 
-  have "AE x in lborel. 0 \<le> f x"
-    using f by (simp add: distributed_def)
-  from AE_borel_affine[OF _ _ this, where c="1/c" and t="- t / c"] c
-  show "AE x in lborel. 0 \<le> f ((x - t) / c) / ereal \<bar>c\<bar>"
-    by (auto simp add: field_simps)
-
-  have eq: "\<And>x. ereal \<bar>c\<bar> * (f x / ereal \<bar>c\<bar>) = f x"
-    using c by (simp add: divide_ereal_def ac_simps one_ereal_def[symmetric])
+  have eq: "ennreal \<bar>c\<bar> * (f x / ennreal \<bar>c\<bar>) = f x" for x
+    using c
+    by (cases "f x")
+       (auto simp: divide_ennreal ennreal_mult[symmetric] ennreal_top_divide ennreal_mult_top)
 
   have "density lborel f = distr M lborel X"
     using f by (simp add: distributed_def)
-  with c show "distr M lborel (\<lambda>x. t + c * X x) = density lborel (\<lambda>x. f ((x - t) / c) / ereal \<bar>c\<bar>)"
+  with c show "distr M lborel (\<lambda>x. t + c * X x) = density lborel (\<lambda>x. f ((x - t) / c) / ennreal \<bar>c\<bar>)"
     by (subst (2) lborel_real_affine[where c="c" and t="t"])
        (simp_all add: density_density_eq density_distr distr_distr field_simps eq cong: distr_cong)
 qed
 
 lemma (in prob_space) distributed_affineI:
-  fixes f :: "real \<Rightarrow> ereal"
+  fixes f :: "real \<Rightarrow> ennreal" and c :: real
   assumes f: "distributed M lborel (\<lambda>x. (X x - t) / c) (\<lambda>x. \<bar>c\<bar> * f (x * c + t))"
   assumes c: "c \<noteq> 0"
   shows "distributed M lborel X f"
 proof -
-  have eq: "\<And>x. f x * ereal \<bar>c\<bar> / ereal \<bar>c\<bar> = f x"
-    using c by (simp add: divide_ereal_def ac_simps one_ereal_def[symmetric])
+  have eq: "f x * ennreal \<bar>c\<bar> / ennreal \<bar>c\<bar> = f x" for x
+    using c by (simp add: ennreal_times_divide[symmetric])
 
   show ?thesis
     using distributed_affine[OF f c, where t=t] c
@@ -73,19 +69,18 @@
 
 lemma nn_intergal_power_times_exp_Icc:
   assumes [arith]: "0 \<le> a"
-  shows "(\<integral>\<^sup>+x. ereal (x^k * exp (-x)) * indicator {0 .. a} x \<partial>lborel) =
+  shows "(\<integral>\<^sup>+x. ennreal (x^k * exp (-x)) * indicator {0 .. a} x \<partial>lborel) =
     (1 - (\<Sum>n\<le>k. (a^n * exp (-a)) / fact n)) * fact k" (is "?I = _")
 proof -
   let ?f = "\<lambda>k x. x^k * exp (-x) / fact k"
   let ?F = "\<lambda>k x. - (\<Sum>n\<le>k. (x^n * exp (-x)) / fact n)"
   have "?I * (inverse (real_of_nat (fact k))) =
-      (\<integral>\<^sup>+x. ereal (x^k * exp (-x)) * indicator {0 .. a} x * (inverse (real_of_nat (fact k))) \<partial>lborel)"
+      (\<integral>\<^sup>+x. ennreal (x^k * exp (-x)) * indicator {0 .. a} x * (inverse (real_of_nat (fact k))) \<partial>lborel)"
     by (intro nn_integral_multc[symmetric]) auto
-  also have "\<dots> = (\<integral>\<^sup>+x. ereal (?f k x) * indicator {0 .. a} x \<partial>lborel)"
-    apply (intro nn_integral_cong)
-    apply (simp add: field_simps)
-    by (metis (no_types) divide_inverse mult.commute mult.left_commute mult.left_neutral times_ereal.simps(1))
-  also have "\<dots> = ereal (?F k a - ?F k 0)"
+  also have "\<dots> = (\<integral>\<^sup>+x. ennreal (?f k x) * indicator {0 .. a} x \<partial>lborel)"
+    by (intro nn_integral_cong)
+       (simp add: field_simps ennreal_mult'[symmetric] indicator_mult_ennreal)
+  also have "\<dots> = ennreal (?F k a - ?F k 0)"
   proof (rule nn_integral_FTC_Icc)
     fix x assume "x \<in> {0..a}"
     show "DERIV (?F k) x :> ?f k x"
@@ -106,17 +101,19 @@
       finally show ?case .
     qed
   qed auto
-  also have "\<dots> = ereal (1 - (\<Sum>n\<le>k. (a^n * exp (-a)) / fact n))"
+  also have "\<dots> = ennreal (1 - (\<Sum>n\<le>k. (a^n * exp (-a)) / fact n))"
     by (auto simp: power_0_left if_distrib[where f="\<lambda>x. x / a" for a] setsum.If_cases)
+  also have "\<dots> = ennreal ((1 - (\<Sum>n\<le>k. (a^n * exp (-a)) / fact n)) * fact k) * ennreal (inverse (fact k))"
+    by (subst ennreal_mult''[symmetric]) (auto intro!: arg_cong[where f=ennreal])
   finally show ?thesis
-    by (cases "?I") (auto simp: field_simps)
+    by (auto simp add: mult_right_ennreal_cancel le_less)
 qed
 
 lemma nn_intergal_power_times_exp_Ici:
-  shows "(\<integral>\<^sup>+x. ereal (x^k * exp (-x)) * indicator {0 ..} x \<partial>lborel) = real_of_nat (fact k)"
+  shows "(\<integral>\<^sup>+x. ennreal (x^k * exp (-x)) * indicator {0 ..} x \<partial>lborel) = real_of_nat (fact k)"
 proof (rule LIMSEQ_unique)
-  let ?X = "\<lambda>n. \<integral>\<^sup>+ x. ereal (x^k * exp (-x)) * indicator {0 .. real n} x \<partial>lborel"
-  show "?X \<longlonglongrightarrow> (\<integral>\<^sup>+x. ereal (x^k * exp (-x)) * indicator {0 ..} x \<partial>lborel)"
+  let ?X = "\<lambda>n. \<integral>\<^sup>+ x. ennreal (x^k * exp (-x)) * indicator {0 .. real n} x \<partial>lborel"
+  show "?X \<longlonglongrightarrow> (\<integral>\<^sup>+x. ennreal (x^k * exp (-x)) * indicator {0 ..} x \<partial>lborel)"
     apply (intro nn_integral_LIMSEQ)
     apply (auto simp: incseq_def le_fun_def eventually_sequentially
                 split: split_indicator intro!: Lim_eventually)
@@ -137,7 +134,7 @@
 definition erlang_CDF ::  "nat \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
   "erlang_CDF k l x = (if x < 0 then 0 else 1 - (\<Sum>n\<le>k. ((l * x)^n * exp (- l * x) / fact n)))"
 
-lemma erlang_density_nonneg: "0 \<le> l \<Longrightarrow> 0 \<le> erlang_density k l x"
+lemma erlang_density_nonneg[simp]: "0 \<le> l \<Longrightarrow> 0 \<le> erlang_density k l x"
   by (simp add: erlang_density_def)
 
 lemma borel_measurable_erlang_density[measurable]: "erlang_density k l \<in> borel_measurable borel"
@@ -146,28 +143,47 @@
 lemma erlang_CDF_transform: "0 < l \<Longrightarrow> erlang_CDF k l a = erlang_CDF k 1 (l * a)"
   by (auto simp add: erlang_CDF_def mult_less_0_iff)
 
+lemma erlang_CDF_nonneg[simp]: assumes "0 < l" shows "0 \<le> erlang_CDF k l x"
+ unfolding erlang_CDF_def
+proof (clarsimp simp: not_less)
+  assume "0 \<le> x"
+  have "(\<Sum>n\<le>k. (l * x) ^ n * exp (- (l * x)) / fact n) =
+    exp (- (l * x)) * (\<Sum>n\<le>k. (l * x) ^ n / fact n)"
+    unfolding setsum_right_distrib by (intro setsum.cong) (auto simp: field_simps)
+  also have "\<dots> = (\<Sum>n\<le>k. (l * x) ^ n / fact n) / exp (l * x)"
+    by (simp add: exp_minus field_simps)
+  also have "\<dots> \<le> 1"
+  proof (subst divide_le_eq_1_pos)
+    show "(\<Sum>n\<le>k. (l * x) ^ n / fact n) \<le> exp (l * x)"
+      using \<open>0 < l\<close> \<open>0 \<le> x\<close> summable_exp_generic[of "l * x"]
+      by (auto simp: exp_def divide_inverse ac_simps intro!: setsum_le_suminf)
+  qed simp
+  finally show "(\<Sum>n\<le>k. (l * x) ^ n * exp (- (l * x)) / fact n) \<le> 1" .
+qed
+
 lemma nn_integral_erlang_density:
   assumes [arith]: "0 < l"
-  shows "(\<integral>\<^sup>+ x. ereal (erlang_density k l x) * indicator {.. a} x \<partial>lborel) = erlang_CDF k l a"
+  shows "(\<integral>\<^sup>+ x. ennreal (erlang_density k l x) * indicator {.. a} x \<partial>lborel) = erlang_CDF k l a"
 proof cases
   assume [arith]: "0 \<le> a"
   have eq: "\<And>x. indicator {0..a} (x / l) = indicator {0..a*l} x"
     by (simp add: field_simps split: split_indicator)
-  have "(\<integral>\<^sup>+x. ereal (erlang_density k l x) * indicator {.. a} x \<partial>lborel) =
-    (\<integral>\<^sup>+x. (l/fact k) * (ereal ((l*x)^k * exp (- (l*x))) * indicator {0 .. a} x) \<partial>lborel)"
-    by (intro nn_integral_cong) (auto simp: erlang_density_def power_mult_distrib split: split_indicator)
-  also have "\<dots> = (l/fact k) * (\<integral>\<^sup>+x. ereal ((l*x)^k * exp (- (l*x))) * indicator {0 .. a} x \<partial>lborel)"
+  have "(\<integral>\<^sup>+x. ennreal (erlang_density k l x) * indicator {.. a} x \<partial>lborel) =
+    (\<integral>\<^sup>+x. (l/fact k) * (ennreal ((l*x)^k * exp (- (l*x))) * indicator {0 .. a} x) \<partial>lborel)"
+    by (intro nn_integral_cong)
+       (auto simp: erlang_density_def power_mult_distrib ennreal_mult[symmetric] split: split_indicator)
+  also have "\<dots> = (l/fact k) * (\<integral>\<^sup>+x. ennreal ((l*x)^k * exp (- (l*x))) * indicator {0 .. a} x \<partial>lborel)"
     by (intro nn_integral_cmult) auto
-  also have "\<dots> = ereal (l/fact k) * ((1/l) * (\<integral>\<^sup>+x. ereal (x^k * exp (- x)) * indicator {0 .. l * a} x \<partial>lborel))"
+  also have "\<dots> = ennreal (l/fact k) * ((1/l) * (\<integral>\<^sup>+x. ennreal (x^k * exp (- x)) * indicator {0 .. l * a} x \<partial>lborel))"
     by (subst nn_integral_real_affine[where c="1 / l" and t=0]) (auto simp: field_simps eq)
   also have "\<dots> = (1 - (\<Sum>n\<le>k. ((l * a)^n * exp (-(l * a))) / fact n))"
-    by (subst nn_intergal_power_times_exp_Icc) auto
+    by (subst nn_intergal_power_times_exp_Icc) (auto simp: ennreal_mult'[symmetric])
   also have "\<dots> = erlang_CDF k l a"
     by (auto simp: erlang_CDF_def)
   finally show ?thesis .
 next
   assume "\<not> 0 \<le> a"
-  moreover then have "(\<integral>\<^sup>+ x. ereal (erlang_density k l x) * indicator {.. a} x \<partial>lborel) = (\<integral>\<^sup>+x. 0 \<partial>(lborel::real measure))"
+  moreover then have "(\<integral>\<^sup>+ x. ennreal (erlang_density k l x) * indicator {.. a} x \<partial>lborel) = (\<integral>\<^sup>+x. 0 \<partial>(lborel::real measure))"
     by (intro nn_integral_cong) (auto simp: erlang_density_def)
   ultimately show ?thesis
     by (simp add: erlang_CDF_def)
@@ -180,19 +196,20 @@
 lemma nn_integral_erlang_ith_moment:
   fixes k i :: nat and l :: real
   assumes [arith]: "0 < l"
-  shows "(\<integral>\<^sup>+ x. ereal (erlang_density k l x * x ^ i) \<partial>lborel) = fact (k + i) / (fact k * l ^ i)"
+  shows "(\<integral>\<^sup>+ x. ennreal (erlang_density k l x * x ^ i) \<partial>lborel) = fact (k + i) / (fact k * l ^ i)"
 proof -
   have eq: "\<And>x. indicator {0..} (x / l) = indicator {0..} x"
     by (simp add: field_simps split: split_indicator)
-  have "(\<integral>\<^sup>+ x. ereal (erlang_density k l x * x^i) \<partial>lborel) =
-    (\<integral>\<^sup>+x. (l/(fact k * l^i)) * (ereal ((l*x)^(k+i) * exp (- (l*x))) * indicator {0 ..} x) \<partial>lborel)"
-    by (intro nn_integral_cong) (auto simp: erlang_density_def power_mult_distrib power_add split: split_indicator)
-  also have "\<dots> = (l/(fact k * l^i)) * (\<integral>\<^sup>+x. ereal ((l*x)^(k+i) * exp (- (l*x))) * indicator {0 ..} x \<partial>lborel)"
+  have "(\<integral>\<^sup>+ x. ennreal (erlang_density k l x * x^i) \<partial>lborel) =
+    (\<integral>\<^sup>+x. (l/(fact k * l^i)) * (ennreal ((l*x)^(k+i) * exp (- (l*x))) * indicator {0 ..} x) \<partial>lborel)"
+    by (intro nn_integral_cong)
+       (auto simp: erlang_density_def power_mult_distrib power_add ennreal_mult'[symmetric] split: split_indicator)
+  also have "\<dots> = (l/(fact k * l^i)) * (\<integral>\<^sup>+x. ennreal ((l*x)^(k+i) * exp (- (l*x))) * indicator {0 ..} x \<partial>lborel)"
     by (intro nn_integral_cmult) auto
-  also have "\<dots> = ereal (l/(fact k * l^i)) * ((1/l) * (\<integral>\<^sup>+x. ereal (x^(k+i) * exp (- x)) * indicator {0 ..} x \<partial>lborel))"
+  also have "\<dots> = ennreal (l/(fact k * l^i)) * ((1/l) * (\<integral>\<^sup>+x. ennreal (x^(k+i) * exp (- x)) * indicator {0 ..} x \<partial>lborel))"
     by (subst nn_integral_real_affine[where c="1 / l" and t=0]) (auto simp: field_simps eq)
   also have "\<dots> = fact (k + i) / (fact k * l ^ i)"
-    by (subst nn_intergal_power_times_exp_Ici) auto
+    by (subst nn_intergal_power_times_exp_Ici) (auto simp: ennreal_mult'[symmetric])
   finally show ?thesis .
 qed
 
@@ -217,7 +234,7 @@
   also have "\<dots> = erlang_CDF k l a"
     by (auto intro!: emeasure_erlang_density)
   finally show ?thesis
-    by (auto simp: measure_def)
+    by (auto simp: emeasure_eq_measure measure_nonneg)
 qed
 
 lemma (in prob_space) erlang_distributed_gt:
@@ -246,17 +263,17 @@
       by (intro emeasure_mono) auto
     also have "... = 0"  by (auto intro!: erlang_CDF_at0 simp: X_distr[of 0])
     finally have "emeasure M {x\<in>space M. X x \<le> a} \<le> 0" by simp
-    then have "emeasure M {x\<in>space M. X x \<le> a} = 0" by (simp add:emeasure_le_0_iff)
+    then have "emeasure M {x\<in>space M. X x \<le> a} = 0" by simp
   }
   note eq_0 = this
 
-  show "(\<integral>\<^sup>+ x. erlang_density k l x * indicator {..a} x \<partial>lborel) = ereal (erlang_CDF k l a)"
+  show "(\<integral>\<^sup>+ x. erlang_density k l x * indicator {..a} x \<partial>lborel) = ennreal (erlang_CDF k l a)"
     using nn_integral_erlang_density[of l k a]
-    by (simp add: times_ereal.simps(1)[symmetric] ereal_indicator del: times_ereal.simps)
+    by (simp add: ennreal_indicator ennreal_mult)
 
-  show "emeasure M {x\<in>space M. X x \<le> a} = ereal (erlang_CDF k l a)"
-    using X_distr[of a] eq_0 by (auto simp: one_ereal_def erlang_CDF_def)
-qed (simp_all add: erlang_density_nonneg)
+  show "emeasure M {x\<in>space M. X x \<le> a} = ennreal (erlang_CDF k l a)"
+    using X_distr[of a] eq_0 by (auto simp: one_ennreal_def erlang_CDF_def)
+qed simp_all
 
 lemma (in prob_space) erlang_distributed_iff:
   assumes [arith]: "0<l"
@@ -266,7 +283,7 @@
     distributed_measurable[of M lborel X "erlang_density k l"]
     emeasure_erlang_density[of l]
     erlang_distributed_le[of X k l]
-  by (auto intro!: erlang_distributedI simp: one_ereal_def emeasure_eq_measure)
+  by (auto intro!: erlang_distributedI simp: one_ennreal_def emeasure_eq_measure)
 
 lemma (in prob_space) erlang_distributed_mult_const:
   assumes erlX: "distributed M lborel X (erlang_density k l)"
@@ -300,10 +317,10 @@
 proof (rule has_bochner_integral_nn_integral)
   show "AE x in M. 0 \<le> X x ^ i"
     by (subst distributed_AE2[OF D]) (auto simp: erlang_density_def)
-  show "(\<integral>\<^sup>+ x. ereal (X x ^ i) \<partial>M) = ereal (fact (k + i) / (fact k * l ^ i))"
+  show "(\<integral>\<^sup>+ x. ennreal (X x ^ i) \<partial>M) = ennreal (fact (k + i) / (fact k * l ^ i))"
     using nn_integral_erlang_ith_moment[of l k i]
-    by (subst distributed_nn_integral[symmetric, OF D]) auto
-qed (insert distributed_measurable[OF D], simp)
+    by (subst distributed_nn_integral[symmetric, OF D]) (auto simp: ennreal_mult')
+qed (insert distributed_measurable[OF D], auto)
 
 lemma (in prob_space) erlang_ith_moment_integrable:
   "0 < l \<Longrightarrow> distributed M lborel X (erlang_density k l) \<Longrightarrow> integrable M (\<lambda>x. X x ^ i)"
@@ -339,65 +356,31 @@
 lemma erlang_CDF_0: "erlang_CDF 0 l a = (if 0 \<le> a then 1 - exp (- l * a) else 0)"
   by (simp add: erlang_CDF_def)
 
-lemma (in prob_space) exponential_distributed_params:
-  assumes D: "distributed M lborel X (exponential_density l)"
-  shows "0 < l"
-proof (cases l "0 :: real" rule: linorder_cases)
-  assume "l < 0"
-  have "emeasure lborel {0 <.. 1::real} \<le>
-    emeasure lborel {x :: real \<in> space lborel. 0 < x}"
-    by (rule emeasure_mono) (auto simp: greaterThan_def[symmetric])
-  also have "emeasure lborel {x :: real \<in> space lborel. 0 < x} = 0"
-  proof -
-    have "AE x in lborel. 0 \<le> exponential_density l x"
-      using assms by (auto simp: distributed_real_AE)
-    then have "AE x in lborel. x \<le> (0::real)"
-      apply eventually_elim
-      using \<open>l < 0\<close>
-      apply (auto simp: exponential_density_def zero_le_mult_iff split: if_split_asm)
-      done
-    then show "emeasure lborel {x :: real \<in> space lborel. 0 < x} = 0"
-      by (subst (asm) AE_iff_measurable[OF _ refl]) (auto simp: not_le greaterThan_def[symmetric])
-  qed
-  finally show "0 < l" by simp
-next
-  assume "l = 0"
-  then have [simp]: "\<And>x. ereal (exponential_density l x) = 0"
-    by (simp add: exponential_density_def)
-  interpret X: prob_space "distr M lborel X"
-    using distributed_measurable[OF D] by (rule prob_space_distr)
-  from X.emeasure_space_1
-  show "0 < l"
-    by (simp add: emeasure_density distributed_distr_eq_density[OF D])
-qed assumption
-
 lemma prob_space_exponential_density: "0 < l \<Longrightarrow> prob_space (density lborel (exponential_density l))"
   by (rule prob_space_erlang_density)
 
 lemma (in prob_space) exponential_distributedD_le:
-  assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \<le> a"
+  assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \<le> a" and l: "0 < l"
   shows "\<P>(x in M. X x \<le> a) = 1 - exp (- a * l)"
-  using erlang_distributed_le[OF D exponential_distributed_params[OF D] a] a
-  by (simp add: erlang_CDF_def)
+  using erlang_distributed_le[OF D l a] a by (simp add: erlang_CDF_def)
 
 lemma (in prob_space) exponential_distributedD_gt:
-  assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \<le> a"
+  assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \<le> a" and l: "0 < l"
   shows "\<P>(x in M. a < X x ) = exp (- a * l)"
-  using erlang_distributed_gt[OF D exponential_distributed_params[OF D] a] a
-  by (simp add: erlang_CDF_def)
+  using erlang_distributed_gt[OF D l a] a by (simp add: erlang_CDF_def)
 
 lemma (in prob_space) exponential_distributed_memoryless:
-  assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \<le> a"and t: "0 \<le> t"
+  assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \<le> a" and l: "0 < l" and t: "0 \<le> t"
   shows "\<P>(x in M. a + t < X x \<bar> a < X x) = \<P>(x in M. t < X x)"
 proof -
   have "\<P>(x in M. a + t < X x \<bar> a < X x) = \<P>(x in M. a + t < X x) / \<P>(x in M. a < X x)"
     using \<open>0 \<le> t\<close> by (auto simp: cond_prob_def intro!: arg_cong[where f=prob] arg_cong2[where f="op /"])
   also have "\<dots> = exp (- (a + t) * l) / exp (- a * l)"
-    using a t by (simp add: exponential_distributedD_gt[OF D])
+    using a t by (simp add: exponential_distributedD_gt[OF D _ l])
   also have "\<dots> = exp (- t * l)"
-    using exponential_distributed_params[OF D] by (auto simp: field_simps exp_add[symmetric])
+    using l by (auto simp: field_simps exp_add[symmetric])
   finally show ?thesis
-    using t by (simp add: exponential_distributedD_gt[OF D])
+    using t by (simp add: exponential_distributedD_gt[OF D _ l])
 qed
 
 lemma exponential_distributedI:
@@ -405,40 +388,45 @@
     and X_distr: "\<And>a. 0 \<le> a \<Longrightarrow> emeasure M {x\<in>space M. X x \<le> a} = 1 - exp (- a * l)"
   shows "distributed M lborel X (exponential_density l)"
 proof (rule erlang_distributedI)
-  fix a :: real assume "0 \<le> a" then show "emeasure M {x \<in> space M. X x \<le> a} = ereal (erlang_CDF 0 l a)"
-    using X_distr[of a] by (simp add: erlang_CDF_def one_ereal_def)
+  fix a :: real assume "0 \<le> a" then show "emeasure M {x \<in> space M. X x \<le> a} = ennreal (erlang_CDF 0 l a)"
+    using X_distr[of a] by (simp add: erlang_CDF_def ennreal_minus ennreal_1[symmetric] del: ennreal_1)
 qed fact+
 
 lemma (in prob_space) exponential_distributed_iff:
-  "distributed M lborel X (exponential_density l) \<longleftrightarrow>
-    (X \<in> borel_measurable M \<and> 0 < l \<and> (\<forall>a\<ge>0. \<P>(x in M. X x \<le> a) = 1 - exp (- a * l)))"
-  using exponential_distributed_params[of X l] erlang_distributed_iff[of l X 0] by (auto simp: erlang_CDF_0)
+  assumes "0 < l"
+  shows "distributed M lborel X (exponential_density l) \<longleftrightarrow>
+    (X \<in> borel_measurable M \<and> (\<forall>a\<ge>0. \<P>(x in M. X x \<le> a) = 1 - exp (- a * l)))"
+  using assms erlang_distributed_iff[of l X 0] by (auto simp: erlang_CDF_0)
 
 
 lemma (in prob_space) exponential_distributed_expectation:
-  "distributed M lborel X (exponential_density l) \<Longrightarrow> expectation X = 1 / l"
-  using erlang_ith_moment[OF exponential_distributed_params, of X l X 0 1] by simp
+  "0 < l \<Longrightarrow> distributed M lborel X (exponential_density l) \<Longrightarrow> expectation X = 1 / l"
+  using erlang_ith_moment[of l X 0 1] by simp
 
 lemma exponential_density_nonneg: "0 < l \<Longrightarrow> 0 \<le> exponential_density l x"
   by (auto simp: exponential_density_def)
 
 lemma (in prob_space) exponential_distributed_min:
+  assumes "0 < l" "0 < u"
   assumes expX: "distributed M lborel X (exponential_density l)"
   assumes expY: "distributed M lborel Y (exponential_density u)"
   assumes ind: "indep_var borel X borel Y"
   shows "distributed M lborel (\<lambda>x. min (X x) (Y x)) (exponential_density (l + u))"
 proof (subst exponential_distributed_iff, safe)
-  have randX: "random_variable borel X" using expX by (simp add: exponential_distributed_iff)
-  moreover have randY: "random_variable borel Y" using expY by (simp add: exponential_distributed_iff)
+  have randX: "random_variable borel X"
+    using expX \<open>0 < l\<close> by (simp add: exponential_distributed_iff)
+  moreover have randY: "random_variable borel Y"
+    using expY \<open>0 < u\<close> by (simp add: exponential_distributed_iff)
   ultimately show "random_variable borel (\<lambda>x. min (X x) (Y x))" by auto
 
-  have "0 < l" by (rule exponential_distributed_params) fact
-  moreover have "0 < u" by (rule exponential_distributed_params) fact
-  ultimately  show " 0 < l + u" by force
+  show " 0 < l + u"
+    using \<open>0 < l\<close> \<open>0 < u\<close> by auto
 
   fix a::real assume a[arith]: "0 \<le> a"
-  have gt1[simp]: "\<P>(x in M. a < X x ) = exp (- a * l)" by (rule exponential_distributedD_gt[OF expX a])
-  have gt2[simp]: "\<P>(x in M. a < Y x ) = exp (- a * u)" by (rule exponential_distributedD_gt[OF expY a])
+  have gt1[simp]: "\<P>(x in M. a < X x ) = exp (- a * l)"
+    by (rule exponential_distributedD_gt[OF expX a]) fact
+  have gt2[simp]: "\<P>(x in M. a < Y x ) = exp (- a * u)"
+    by (rule exponential_distributedD_gt[OF expY a]) fact
 
   have "\<P>(x in M. a < (min (X x) (Y x)) ) =  \<P>(x in M. a < (X x) \<and> a < (Y x))" by (auto intro!:arg_cong[where f=prob])
 
@@ -458,6 +446,7 @@
 lemma (in prob_space) exponential_distributed_Min:
   assumes finI: "finite I"
   assumes A: "I \<noteq> {}"
+  assumes l: "\<And>i. i \<in> I \<Longrightarrow> 0 < l i"
   assumes expX: "\<And>i. i \<in> I \<Longrightarrow> distributed M lborel (X i) (exponential_density (l i))"
   assumes ind: "indep_vars (\<lambda>i. borel) X I"
   shows "distributed M lborel (\<lambda>x. Min ((\<lambda>i. X i x)`I)) (exponential_density (\<Sum>i\<in>I. l i))"
@@ -468,14 +457,14 @@
   case (insert i I)
   then have "distributed M lborel (\<lambda>x. min (X i x) (Min ((\<lambda>i. X i x)`I))) (exponential_density (l i + (\<Sum>i\<in>I. l i)))"
       by (intro exponential_distributed_min indep_vars_Min insert)
-         (auto intro: indep_vars_subset)
+         (auto intro: indep_vars_subset setsum_pos)
   then show ?case
     using insert by simp
 qed
 
 lemma (in prob_space) exponential_distributed_variance:
-  "distributed M lborel X (exponential_density l) \<Longrightarrow> variance X = 1 / l\<^sup>2"
-  using erlang_distributed_variance[OF exponential_distributed_params, of X l X 0] by simp
+  "0 < l \<Longrightarrow> distributed M lborel X (exponential_density l) \<Longrightarrow> variance X = 1 / l\<^sup>2"
+  using erlang_distributed_variance[of l X 0] by simp
 
 lemma nn_integral_zero': "AE x in M. f x = 0 \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>M) = 0"
   by (simp cong: nn_integral_cong_AE)
@@ -483,7 +472,7 @@
 lemma convolution_erlang_density:
   fixes k\<^sub>1 k\<^sub>2 :: nat
   assumes [simp, arith]: "0 < l"
-  shows "(\<lambda>x. \<integral>\<^sup>+y. ereal (erlang_density k\<^sub>1 l (x - y)) * ereal (erlang_density k\<^sub>2 l y) \<partial>lborel) =
+  shows "(\<lambda>x. \<integral>\<^sup>+y. ennreal (erlang_density k\<^sub>1 l (x - y)) * ennreal (erlang_density k\<^sub>2 l y) \<partial>lborel) =
     (erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l)"
       (is "?LHS = ?RHS")
 proof
@@ -503,22 +492,21 @@
     have I_eq1: "integral\<^sup>N lborel (erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l) = 1"
       using nn_integral_erlang_ith_moment[of l "Suc k\<^sub>1 + Suc k\<^sub>2 - 1" 0] by (simp del: fact_Suc)
 
-    have 1: "(\<integral>\<^sup>+ x. ereal (erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l x * indicator {0<..} x) \<partial>lborel) = 1"
+    have 1: "(\<integral>\<^sup>+ x. ennreal (erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l x * indicator {0<..} x) \<partial>lborel) = 1"
       apply (subst I_eq1[symmetric])
       unfolding erlang_density_def
       by (auto intro!: nn_integral_cong split:split_indicator)
 
     have "prob_space (density lborel ?LHS)"
-      unfolding times_ereal.simps[symmetric]
       by (intro prob_space_convolution_density)
          (auto intro!: prob_space_erlang_density erlang_density_nonneg)
     then have 2: "integral\<^sup>N lborel ?LHS = 1"
       by (auto dest!: prob_space.emeasure_space_1 simp: emeasure_density)
 
-    let ?I = "(integral\<^sup>N lborel (\<lambda>y. ereal ((1 - y)^ k\<^sub>1 * y^k\<^sub>2 * indicator {0..1} y)))"
+    let ?I = "(integral\<^sup>N lborel (\<lambda>y. ennreal ((1 - y)^ k\<^sub>1 * y^k\<^sub>2 * indicator {0..1} y)))"
     let ?C = "(fact (Suc (k\<^sub>1 + k\<^sub>2))) / ((fact k\<^sub>1) * (fact k\<^sub>2))"
     let ?s = "Suc k\<^sub>1 + Suc k\<^sub>2 - 1"
-    let ?L = "(\<lambda>x. \<integral>\<^sup>+y. ereal (erlang_density k\<^sub>1 l (x- y) * erlang_density k\<^sub>2 l y * indicator {0..x} y) \<partial>lborel)"
+    let ?L = "(\<lambda>x. \<integral>\<^sup>+y. ennreal (erlang_density k\<^sub>1 l (x- y) * erlang_density k\<^sub>2 l y * indicator {0..x} y) \<partial>lborel)"
 
     { fix x :: real assume [arith]: "0 < x"
       have *: "\<And>x y n. (x - y * x::real)^n = x^n * (1 - y)^n"
@@ -526,17 +514,18 @@
 
       have "?LHS x = ?L x"
         unfolding erlang_density_def
-        by (auto intro!: nn_integral_cong split:split_indicator)
-      also have "... = (\<lambda>x. ereal ?C * ?I * erlang_density ?s l x) x"
+        by (auto intro!: nn_integral_cong simp: ennreal_mult split:split_indicator)
+      also have "... = (\<lambda>x. ennreal ?C * ?I * erlang_density ?s l x) x"
         apply (subst nn_integral_real_affine[where c=x and t = 0])
-        apply (simp_all add: nn_integral_cmult[symmetric] nn_integral_multc[symmetric] erlang_density_nonneg del: fact_Suc)
+        apply (simp_all add: nn_integral_cmult[symmetric] nn_integral_multc[symmetric] del: fact_Suc)
         apply (intro nn_integral_cong)
         apply (auto simp add: erlang_density_def mult_less_0_iff exp_minus field_simps exp_diff power_add *
+                              ennreal_mult[symmetric]
                     simp del: fact_Suc split: split_indicator)
         done
-      finally have "(\<integral>\<^sup>+y. ereal (erlang_density k\<^sub>1 l (x - y) * erlang_density k\<^sub>2 l y) \<partial>lborel) =
-        (\<lambda>x. ereal ?C * ?I * erlang_density ?s l x) x"
-        by simp }
+      finally have "(\<integral>\<^sup>+y. ennreal (erlang_density k\<^sub>1 l (x - y) * erlang_density k\<^sub>2 l y) \<partial>lborel) =
+        (\<lambda>x. ennreal ?C * ?I * erlang_density ?s l x) x"
+        by (simp add: ennreal_mult) }
     note * = this
 
     assume [arith]: "0 < x"
@@ -544,15 +533,15 @@
       by (subst 2[symmetric])
          (auto intro!: nn_integral_cong_AE AE_I[where N="{0}"]
                simp: erlang_density_def  nn_integral_multc[symmetric] indicator_def split: if_split_asm)
-    also have "... = integral\<^sup>N lborel (\<lambda>x. (ereal (?C) * ?I) * ((erlang_density ?s l x) * indicator {0<..} x))"
-      by (auto intro!: nn_integral_cong simp: * split: split_indicator)
-    also have "... = ereal (?C) * ?I"
+    also have "... = integral\<^sup>N lborel (\<lambda>x. (ennreal (?C) * ?I) * ((erlang_density ?s l x) * indicator {0<..} x))"
+      by (auto intro!: nn_integral_cong simp: ennreal_mult[symmetric] * split: split_indicator)
+    also have "... = ennreal (?C) * ?I"
       using 1
-      by (auto simp: nn_integral_nonneg nn_integral_cmult)
-    finally have " ereal (?C) * ?I = 1" by presburger
+      by (auto simp: nn_integral_cmult)
+    finally have " ennreal (?C) * ?I = 1" by presburger
 
     then show ?thesis
-      using * by simp
+      using * by (simp add: ennreal_mult)
   qed
 qed
 
@@ -592,21 +581,17 @@
 lemma (in prob_space) exponential_distributed_setsum:
   assumes finI: "finite I"
   assumes A: "I \<noteq> {}"
+  assumes l: "0 < l"
   assumes expX: "\<And>i. i \<in> I \<Longrightarrow> distributed M lborel (X i) (exponential_density l)"
   assumes ind: "indep_vars (\<lambda>i. borel) X I"
   shows "distributed M lborel (\<lambda>x. \<Sum>i\<in>I. X i x) (erlang_density ((card I) - 1) l)"
-proof -
-  obtain i where "i \<in> I" using assms by auto
-  note exponential_distributed_params[OF expX[OF this]]
-  from erlang_distributed_setsum[OF assms(1,2) this assms(3,4)] show ?thesis by simp
-qed
+  using erlang_distributed_setsum[OF assms] by simp
 
 lemma (in information_space) entropy_exponential:
+  assumes l[simp, arith]: "0 < l"
   assumes D: "distributed M lborel X (exponential_density l)"
   shows "entropy b lborel X = log b (exp 1 / l)"
 proof -
-  have l[simp, arith]: "0 < l" by (rule exponential_distributed_params[OF D])
-
   have [simp]: "integrable lborel (exponential_density l)"
     using distributed_integrable[OF D, of "\<lambda>_. 1"] by simp
 
@@ -620,7 +605,7 @@
     using erlang_ith_moment[OF l D, of 1] distributed_integral[OF D, of "\<lambda>x. x"] by simp
 
   have "entropy b lborel X = - (\<integral> x. exponential_density l x * log b (exponential_density l x) \<partial>lborel)"
-    using D by (rule entropy_distr)
+    using D by (rule entropy_distr) simp
   also have "(\<integral> x. exponential_density l x * log b (exponential_density l x) \<partial>lborel) =
     (\<integral> x. (ln l * exponential_density l x - l * (exponential_density l x * x)) / ln b \<partial>lborel)"
     by (intro integral_cong) (auto simp: log_def ln_mult exponential_density_def field_simps)
@@ -644,10 +629,10 @@
   with X have "emeasure M (X -` B \<inter> space M) = emeasure M' (A \<inter> B) / emeasure M' A"
     by (simp add: distr[of B] measurable_sets)
   also have "\<dots> = (1 / emeasure M' A) * emeasure M' (A \<inter> B)"
-     by simp
+     by (simp add: divide_ennreal_def ac_simps)
   also have "\<dots> = (\<integral>\<^sup>+ x. (1 / emeasure M' A) * indicator (A \<inter> B) x \<partial>M')"
     using A B
-    by (intro nn_integral_cmult_indicator[symmetric]) (auto intro!: zero_le_divide_ereal)
+    by (intro nn_integral_cmult_indicator[symmetric]) (auto intro!: )
   also have "\<dots> = (\<integral>\<^sup>+ x. ?f x * indicator B x \<partial>M')"
     by (rule nn_integral_cong) (auto split: split_indicator)
   finally show "emeasure (distr M M' X) B = emeasure (density M' ?f) B"
@@ -656,7 +641,7 @@
 
 lemma uniform_distrI_borel:
   fixes A :: "real set"
-  assumes X[measurable]: "X \<in> borel_measurable M" and A: "emeasure lborel A = ereal r" "0 < r"
+  assumes X[measurable]: "X \<in> borel_measurable M" and A: "emeasure lborel A = ennreal r" "0 < r"
     and [measurable]: "A \<in> sets borel"
   assumes distr: "\<And>a. emeasure M {x\<in>space M. X x \<le> a} = emeasure lborel (A \<inter> {.. a}) / r"
   shows "distributed M lborel X (\<lambda>x. indicator A x / measure lborel A)"
@@ -667,24 +652,25 @@
     using A by (intro emeasure_mono) auto
   also have "\<dots> < \<infinity>"
     using A by simp
-  finally have fin: "emeasure lborel (A \<inter> {..a}) \<noteq> \<infinity>"
+  finally have fin: "emeasure lborel (A \<inter> {..a}) \<noteq> top"
     by simp
-  from emeasure_eq_ereal_measure[OF this]
-  have fin_eq: "emeasure lborel (A \<inter> {..a}) / r = ereal (measure lborel (A \<inter> {..a}) / r)"
-    using A by simp
-  then show "emeasure M {x\<in>space M. X x \<le> a} = ereal (measure lborel (A \<inter> {..a}) / r)"
+  from emeasure_eq_ennreal_measure[OF this]
+  have fin_eq: "emeasure lborel (A \<inter> {..a}) / r = ennreal (measure lborel (A \<inter> {..a}) / r)"
+    using A by (simp add: divide_ennreal measure_nonneg)
+  then show "emeasure M {x\<in>space M. X x \<le> a} = ennreal (measure lborel (A \<inter> {..a}) / r)"
     using distr by simp
 
-  have "(\<integral>\<^sup>+ x. ereal (indicator A x / measure lborel A * indicator {..a} x) \<partial>lborel) =
-    (\<integral>\<^sup>+ x. ereal (1 / measure lborel A) * indicator (A \<inter> {..a}) x \<partial>lborel)"
+  have "(\<integral>\<^sup>+ x. ennreal (indicator A x / measure lborel A * indicator {..a} x) \<partial>lborel) =
+    (\<integral>\<^sup>+ x. ennreal (1 / measure lborel A) * indicator (A \<inter> {..a}) x \<partial>lborel)"
     by (auto intro!: nn_integral_cong split: split_indicator)
-  also have "\<dots> = ereal (1 / measure lborel A) * emeasure lborel (A \<inter> {..a})"
+  also have "\<dots> = ennreal (1 / measure lborel A) * emeasure lborel (A \<inter> {..a})"
     using \<open>A \<in> sets borel\<close>
     by (intro nn_integral_cmult_indicator) (auto simp: measure_nonneg)
-  also have "\<dots> = ereal (measure lborel (A \<inter> {..a}) / r)"
-    unfolding emeasure_eq_ereal_measure[OF fin] using A by (simp add: measure_def)
-  finally show "(\<integral>\<^sup>+ x. ereal (indicator A x / measure lborel A * indicator {..a} x) \<partial>lborel) =
-    ereal (measure lborel (A \<inter> {..a}) / r)" .
+  also have "\<dots> = ennreal (measure lborel (A \<inter> {..a}) / r)"
+    unfolding emeasure_eq_ennreal_measure[OF fin] using A
+    by (simp add: measure_def ennreal_mult'[symmetric])
+  finally show "(\<integral>\<^sup>+ x. ennreal (indicator A x / measure lborel A * indicator {..a} x) \<partial>lborel) =
+    ennreal (measure lborel (A \<inter> {..a}) / r)" .
 qed (auto simp: measure_nonneg)
 
 lemma (in prob_space) uniform_distrI_borel_atLeastAtMost:
@@ -704,45 +690,45 @@
     also have "\<dots> = 0"
       using distr[of a] \<open>a < b\<close> by (simp add: emeasure_eq_measure)
     finally have "emeasure M {x\<in>space M. X x \<le> t} = 0"
-      by (simp add: antisym measure_nonneg emeasure_le_0_iff)
+      by (simp add: antisym measure_nonneg)
     with \<open>t < a\<close> show ?thesis by simp
   next
     assume bnds: "a \<le> t" "t \<le> b"
     have "{a..b} \<inter> {..t} = {a..t}"
       using bnds by auto
     then show ?thesis using \<open>a \<le> t\<close> \<open>a < b\<close>
-      using distr[OF bnds] by (simp add: emeasure_eq_measure)
+      using distr[OF bnds] by (simp add: emeasure_eq_measure divide_ennreal)
   next
     assume "b < t"
     have "1 = emeasure M {x\<in>space M. X x \<le> b}"
-      using distr[of b] \<open>a < b\<close> by (simp add: one_ereal_def emeasure_eq_measure)
+      using distr[of b] \<open>a < b\<close> by (simp add: one_ennreal_def emeasure_eq_measure)
     also have "\<dots> \<le> emeasure M {x\<in>space M. X x \<le> t}"
       using X \<open>b < t\<close> by (auto intro!: emeasure_mono measurable_sets)
     finally have "emeasure M {x\<in>space M. X x \<le> t} = 1"
-       by (simp add: antisym emeasure_eq_measure one_ereal_def)
-    with \<open>b < t\<close> \<open>a < b\<close> show ?thesis by (simp add: measure_def one_ereal_def)
+       by (simp add: antisym emeasure_eq_measure)
+    with \<open>b < t\<close> \<open>a < b\<close> show ?thesis by (simp add: measure_def divide_ennreal)
   qed
 qed (insert X \<open>a < b\<close>, auto)
 
 lemma (in prob_space) uniform_distributed_measure:
   fixes a b :: real
   assumes D: "distributed M lborel X (\<lambda>x. indicator {a .. b} x / measure lborel {a .. b})"
-  assumes " a \<le> t" "t \<le> b"
+  assumes t: "a \<le> t" "t \<le> b"
   shows "\<P>(x in M. X x \<le> t) = (t - a) / (b - a)"
 proof -
   have "emeasure M {x \<in> space M. X x \<le> t} = emeasure (distr M lborel X) {.. t}"
     using distributed_measurable[OF D]
     by (subst emeasure_distr) (auto intro!: arg_cong2[where f=emeasure])
-  also have "\<dots> = (\<integral>\<^sup>+x. ereal (1 / (b - a)) * indicator {a .. t} x \<partial>lborel)"
+  also have "\<dots> = (\<integral>\<^sup>+x. ennreal (1 / (b - a)) * indicator {a .. t} x \<partial>lborel)"
     using distributed_borel_measurable[OF D] \<open>a \<le> t\<close> \<open>t \<le> b\<close>
     unfolding distributed_distr_eq_density[OF D]
     by (subst emeasure_density)
        (auto intro!: nn_integral_cong simp: measure_def split: split_indicator)
-  also have "\<dots> = ereal (1 / (b - a)) * (t - a)"
+  also have "\<dots> = ennreal (1 / (b - a)) * (t - a)"
     using \<open>a \<le> t\<close> \<open>t \<le> b\<close>
     by (subst nn_integral_cmult_indicator) auto
   finally show ?thesis
-    by (simp add: measure_def)
+    using t by (simp add: emeasure_eq_measure ennreal_mult''[symmetric] measure_nonneg)
 qed
 
 lemma (in prob_space) uniform_distributed_bounds:
@@ -795,7 +781,7 @@
       unfolding * square_diff_square_factored by (auto simp: field_simps)
   qed (insert \<open>a < b\<close>, simp)
   finally show "(\<integral> x. indicator {a .. b} x / measure lborel {a .. b} * x \<partial>lborel) = (a + b) / 2" .
-qed auto
+qed (auto simp: measure_nonneg)
 
 lemma (in prob_space) uniform_distributed_variance:
   fixes a b :: real
@@ -810,7 +796,7 @@
     by (simp add: integral_power uniform_distributed_expectation[OF D])
        (simp add: eval_nat_numeral field_simps )
   finally show "(\<integral>x. x\<^sup>2 * ?D x \<partial>lborel) = (b - a)\<^sup>2 / 12" .
-qed fact
+qed (auto intro: D simp: measure_nonneg)
 
 subsection \<open>Normal distribution\<close>
 
@@ -824,7 +810,7 @@
 lemma std_normal_density_def: "std_normal_density x = (1 / sqrt (2 * pi)) * exp (- x\<^sup>2 / 2)"
   unfolding normal_density_def by simp
 
-lemma normal_density_nonneg: "0 \<le> normal_density \<mu> \<sigma> x"
+lemma normal_density_nonneg[simp]: "0 \<le> normal_density \<mu> \<sigma> x"
   by (auto simp: normal_density_def)
 
 lemma normal_density_pos: "0 < \<sigma> \<Longrightarrow> 0 < normal_density \<mu> \<sigma> x"
@@ -846,14 +832,14 @@
     by (intro nn_integral_cong_AE AE_I[where N="{0}"]) (auto split: split_indicator)
 
   have "?pI ?gauss * ?pI ?gauss = (\<integral>\<^sup>+x. \<integral>\<^sup>+s. ?gauss x * ?gauss s * ?I s * ?I x \<partial>lborel \<partial>lborel)"
-    by (auto simp: nn_integral_nonneg nn_integral_cmult[symmetric] nn_integral_multc[symmetric] *
+    by (auto simp: nn_integral_cmult[symmetric] nn_integral_multc[symmetric] * ennreal_mult[symmetric]
              intro!: nn_integral_cong split: split_indicator)
   also have "\<dots> = (\<integral>\<^sup>+x. \<integral>\<^sup>+s. ?ff x s * ?I s * ?I x \<partial>lborel \<partial>lborel)"
   proof (rule nn_integral_cong, cases)
     fix x :: real assume "x \<noteq> 0"
     then show "(\<integral>\<^sup>+s. ?gauss x * ?gauss s * ?I s * ?I x \<partial>lborel) = (\<integral>\<^sup>+s. ?ff x s * ?I s * ?I x \<partial>lborel)"
       by (subst nn_integral_real_affine[where t="0" and c="x"])
-         (auto simp: mult_exp_exp nn_integral_cmult[symmetric] field_simps zero_less_mult_iff
+         (auto simp: mult_exp_exp nn_integral_cmult[symmetric] field_simps zero_less_mult_iff ennreal_mult[symmetric]
                intro!: nn_integral_cong split: split_indicator)
   qed simp
   also have "... = \<integral>\<^sup>+s. \<integral>\<^sup>+x. ?ff x s * ?I s * ?I x \<partial>lborel \<partial>lborel"
@@ -861,9 +847,9 @@
   also have "... = ?pI (\<lambda>s. ?pI (\<lambda>x. ?ff x s))"
     by (rule nn_integral_cong_AE)
        (auto intro!: nn_integral_cong_AE AE_I[where N="{0}"] split: split_indicator_asm)
-  also have "\<dots> = ?pI (\<lambda>s. ereal (1 / (2 * (1 + s\<^sup>2))))"
-  proof (intro nn_integral_cong ereal_right_mult_cong)
-    fix s :: real show "?pI (\<lambda>x. ?ff x s) = ereal (1 / (2 * (1 + s\<^sup>2)))"
+  also have "\<dots> = ?pI (\<lambda>s. ennreal (1 / (2 * (1 + s\<^sup>2))))"
+  proof (intro nn_integral_cong ennreal_mult_right_cong)
+    fix s :: real show "?pI (\<lambda>x. ?ff x s) = ennreal (1 / (2 * (1 + s\<^sup>2)))"
     proof (subst nn_integral_FTC_atLeast)
       have "((\<lambda>a. - (exp (- (a\<^sup>2 * (1 + s\<^sup>2))) / (2 + 2 * s\<^sup>2))) \<longlongrightarrow> (- (0 / (2 + 2 * s\<^sup>2)))) at_top"
         apply (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_compose[OF filterlim_uminus_at_bot_at_top])
@@ -875,8 +861,8 @@
       then show "((\<lambda>a. - (exp (- a\<^sup>2 - s\<^sup>2 * a\<^sup>2) / (2 + 2 * s\<^sup>2))) \<longlongrightarrow> 0) at_top"
         by (simp add: field_simps)
     qed (auto intro!: derivative_eq_intros simp: field_simps add_nonneg_eq_0_iff)
-  qed rule
-  also have "... = ereal (pi / 4)"
+  qed
+  also have "... = ennreal (pi / 4)"
   proof (subst nn_integral_FTC_atLeast)
     show "((\<lambda>a. arctan a / 2) \<longlongrightarrow> (pi / 2) / 2 ) at_top"
       by (intro tendsto_intros) (simp_all add: tendsto_arctan_at_top)
@@ -884,25 +870,25 @@
   finally have "?pI ?gauss^2 = pi / 4"
     by (simp add: power2_eq_square)
   then have "?pI ?gauss = sqrt (pi / 4)"
-    using power_eq_iff_eq_base[of 2 "real_of_ereal (?pI ?gauss)" "sqrt (pi / 4)"]
-          nn_integral_nonneg[of lborel "\<lambda>x. ?gauss x * indicator {0..} x"]
-    by (cases "?pI ?gauss") auto
+    using power_eq_iff_eq_base[of 2 "enn2real (?pI ?gauss)" "sqrt (pi / 4)"]
+    by (cases "?pI ?gauss") (auto simp: power2_eq_square ennreal_mult[symmetric] ennreal_top_mult)
   also have "?pI ?gauss = (\<integral>\<^sup>+x. indicator {0..} x *\<^sub>R exp (- x\<^sup>2) \<partial>lborel)"
     by (intro nn_integral_cong) (simp split: split_indicator)
   also have "sqrt (pi / 4) = sqrt pi / 2"
     by (simp add: real_sqrt_divide)
   finally show ?thesis
-    by (rule has_bochner_integral_nn_integral[rotated 2]) auto
+    by (rule has_bochner_integral_nn_integral[rotated 3])
+       auto
 qed
 
 lemma gaussian_moment_1:
   "has_bochner_integral lborel (\<lambda>x::real. indicator {0..} x *\<^sub>R (exp (- x\<^sup>2) * x)) (1 / 2)"
 proof -
   have "(\<integral>\<^sup>+x. indicator {0..} x *\<^sub>R (exp (- x\<^sup>2) * x) \<partial>lborel) =
-    (\<integral>\<^sup>+x. ereal (x * exp (- x\<^sup>2)) * indicator {0..} x \<partial>lborel)"
+    (\<integral>\<^sup>+x. ennreal (x * exp (- x\<^sup>2)) * indicator {0..} x \<partial>lborel)"
     by (intro nn_integral_cong)
-       (auto simp: ac_simps times_ereal.simps(1)[symmetric] ereal_indicator simp del: times_ereal.simps)
-  also have "\<dots> = ereal (0 - (- exp (- 0\<^sup>2) / 2))"
+       (auto simp: ac_simps split: split_indicator)
+  also have "\<dots> = ennreal (0 - (- exp (- 0\<^sup>2) / 2))"
   proof (rule nn_integral_FTC_atLeast)
     have "((\<lambda>x::real. - exp (- x\<^sup>2) / 2) \<longlongrightarrow> - 0 / 2) at_top"
       by (intro tendsto_divide tendsto_minus filterlim_compose[OF exp_at_bot]
@@ -912,10 +898,11 @@
     then show "((\<lambda>a::real. - exp (- a\<^sup>2) / 2) \<longlongrightarrow> 0) at_top"
       by simp
   qed (auto intro!: derivative_eq_intros)
-  also have "\<dots> = ereal (1 / 2)"
+  also have "\<dots> = ennreal (1 / 2)"
     by simp
   finally show ?thesis
-    by (rule has_bochner_integral_nn_integral[rotated 2]) (auto split: split_indicator)
+    by (rule has_bochner_integral_nn_integral[rotated 3])
+        (auto split: split_indicator)
 qed
 
 lemma
@@ -1203,29 +1190,30 @@
     by (simp add: normal_density_def real_sqrt_mult field_simps)
        (simp add: power2_eq_square field_simps)
   show ?thesis
-    by (rule distributed_affineI[OF _ \<open>\<alpha> \<noteq> 0\<close>, where t=\<beta>]) (simp_all add: eq X)
+    by (rule distributed_affineI[OF _ \<open>\<alpha> \<noteq> 0\<close>, where t=\<beta>])
+       (simp_all add: eq X ennreal_mult'[symmetric])
 qed
 
 lemma (in prob_space) normal_standard_normal_convert:
   assumes pos_var[simp, arith]: "0 < \<sigma>"
   shows "distributed M lborel X (normal_density  \<mu> \<sigma>) = distributed M lborel (\<lambda>x. (X x - \<mu>) / \<sigma>) std_normal_density"
 proof auto
-  assume "distributed M lborel X (\<lambda>x. ereal (normal_density \<mu> \<sigma> x))"
-  then have "distributed M lborel (\<lambda>x. -\<mu> / \<sigma> + (1/\<sigma>) * X x) (\<lambda>x. ereal (normal_density (-\<mu> / \<sigma> + (1/\<sigma>)* \<mu>) (\<bar>1/\<sigma>\<bar> * \<sigma>) x))"
+  assume "distributed M lborel X (\<lambda>x. ennreal (normal_density \<mu> \<sigma> x))"
+  then have "distributed M lborel (\<lambda>x. -\<mu> / \<sigma> + (1/\<sigma>) * X x) (\<lambda>x. ennreal (normal_density (-\<mu> / \<sigma> + (1/\<sigma>)* \<mu>) (\<bar>1/\<sigma>\<bar> * \<sigma>) x))"
     by(rule normal_density_affine) auto
 
-  then show "distributed M lborel (\<lambda>x. (X x - \<mu>) / \<sigma>) (\<lambda>x. ereal (std_normal_density x))"
+  then show "distributed M lborel (\<lambda>x. (X x - \<mu>) / \<sigma>) (\<lambda>x. ennreal (std_normal_density x))"
     by (simp add: diff_divide_distrib[symmetric] field_simps)
 next
-  assume *: "distributed M lborel (\<lambda>x. (X x - \<mu>) / \<sigma>) (\<lambda>x. ereal (std_normal_density x))"
-  have "distributed M lborel (\<lambda>x. \<mu> + \<sigma> * ((X x - \<mu>) / \<sigma>)) (\<lambda>x. ereal (normal_density \<mu> \<bar>\<sigma>\<bar> x))"
+  assume *: "distributed M lborel (\<lambda>x. (X x - \<mu>) / \<sigma>) (\<lambda>x. ennreal (std_normal_density x))"
+  have "distributed M lborel (\<lambda>x. \<mu> + \<sigma> * ((X x - \<mu>) / \<sigma>)) (\<lambda>x. ennreal (normal_density \<mu> \<bar>\<sigma>\<bar> x))"
     using normal_density_affine[OF *, of \<sigma> \<mu>] by simp
-  then show "distributed M lborel X (\<lambda>x. ereal (normal_density \<mu> \<sigma> x))" by simp
+  then show "distributed M lborel X (\<lambda>x. ennreal (normal_density \<mu> \<sigma> x))" by simp
 qed
 
 lemma conv_normal_density_zero_mean:
   assumes [simp, arith]: "0 < \<sigma>" "0 < \<tau>"
-  shows "(\<lambda>x. \<integral>\<^sup>+y. ereal (normal_density 0 \<sigma> (x - y) * normal_density 0 \<tau> y) \<partial>lborel) =
+  shows "(\<lambda>x. \<integral>\<^sup>+y. ennreal (normal_density 0 \<sigma> (x - y) * normal_density 0 \<tau> y) \<partial>lborel) =
     normal_density 0 (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2))"  (is "?LHS = ?RHS")
 proof -
   def \<sigma>' \<equiv> "\<sigma>\<^sup>2" and \<tau>' \<equiv> "\<tau>\<^sup>2"
@@ -1237,7 +1225,7 @@
     by (subst power_eq_iff_eq_base[symmetric, where n=2])
        (simp_all add: real_sqrt_mult[symmetric] power2_eq_square)
   have "?LHS =
-    (\<lambda>x. \<integral>\<^sup>+y. ereal((normal_density 0 (sqrt (\<sigma>' + \<tau>')) x) * normal_density (\<tau>' * x / (\<sigma>' + \<tau>')) ?\<sigma> y) \<partial>lborel)"
+    (\<lambda>x. \<integral>\<^sup>+y. ennreal((normal_density 0 (sqrt (\<sigma>' + \<tau>')) x) * normal_density (\<tau>' * x / (\<sigma>' + \<tau>')) ?\<sigma> y) \<partial>lborel)"
     apply (intro ext nn_integral_cong)
     apply (simp add: normal_density_def \<sigma>'_def[symmetric] \<tau>'_def[symmetric] sqrt mult_exp_exp)
     apply (simp add: divide_simps power2_eq_square)
@@ -1245,8 +1233,9 @@
     done
 
   also have "... =
-    (\<lambda>x. (normal_density 0 (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)) x) * \<integral>\<^sup>+y. ereal( normal_density (\<tau>\<^sup>2* x / (\<sigma>\<^sup>2 + \<tau>\<^sup>2)) ?\<sigma> y) \<partial>lborel)"
-    by (subst nn_integral_cmult[symmetric]) (auto simp: \<sigma>'_def \<tau>'_def normal_density_def)
+    (\<lambda>x. (normal_density 0 (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)) x) * \<integral>\<^sup>+y. ennreal( normal_density (\<tau>\<^sup>2* x / (\<sigma>\<^sup>2 + \<tau>\<^sup>2)) ?\<sigma> y) \<partial>lborel)"
+    by (subst nn_integral_cmult[symmetric])
+       (auto simp: \<sigma>'_def \<tau>'_def normal_density_def ennreal_mult'[symmetric])
 
   also have "... = (\<lambda>x. (normal_density 0 (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)) x))"
     by (subst nn_integral_eq_integral) (auto simp: normal_density_nonneg)
@@ -1255,7 +1244,7 @@
 qed
 
 lemma conv_std_normal_density:
-  "(\<lambda>x. \<integral>\<^sup>+y. ereal (std_normal_density (x - y) * std_normal_density y) \<partial>lborel) =
+  "(\<lambda>x. \<integral>\<^sup>+y. ennreal (std_normal_density (x - y) * std_normal_density y) \<partial>lborel) =
   (normal_density 0 (sqrt 2))"
   by (subst conv_normal_density_zero_mean) simp_all
 
@@ -1281,11 +1270,12 @@
     by(rule normal_density_affine[OF normalY pos_var(2), of 1 "-\<nu>"]) simp
   then have 2[simp]: "distributed M lborel (\<lambda>x. - \<nu> +  Y x) (normal_density 0 \<tau>)" by simp
 
-  have *: "distributed M lborel (\<lambda>x. (- \<mu> + X x) + (- \<nu> + Y x)) (\<lambda>x. ereal (normal_density 0 (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)) x))"
-    using distributed_convolution[OF ind 1 2] conv_normal_density_zero_mean[OF pos_var] by simp
+  have *: "distributed M lborel (\<lambda>x. (- \<mu> + X x) + (- \<nu> + Y x)) (\<lambda>x. ennreal (normal_density 0 (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)) x))"
+    using distributed_convolution[OF ind 1 2] conv_normal_density_zero_mean[OF pos_var]
+    by (simp add: ennreal_mult'[symmetric] normal_density_nonneg)
 
   have "distributed M lborel (\<lambda>x. \<mu> + \<nu> + 1 * (- \<mu> + X x + (- \<nu> + Y x)))
-        (\<lambda>x. ereal (normal_density (\<mu> + \<nu> + 1 * 0) (\<bar>1\<bar> * sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)) x))"
+        (\<lambda>x. ennreal (normal_density (\<mu> + \<nu> + 1 * 0) (\<bar>1\<bar> * sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)) x))"
     by (rule normal_density_affine[OF *, of 1 "\<mu> + \<nu>"]) (auto simp: add_pos_pos)
 
   then show ?thesis by auto
@@ -1298,9 +1288,9 @@
   assumes normalY[simp]: "distributed M lborel Y (normal_density \<nu> \<tau>)"
   shows "distributed M lborel (\<lambda>x. X x - Y x) (normal_density (\<mu> - \<nu>) (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)))"
 proof -
-  have "distributed M lborel (\<lambda>x. 0 + - 1 * Y x) (\<lambda>x. ereal (normal_density (0 + - 1 * \<nu>) (\<bar>- 1\<bar> * \<tau>) x))"
+  have "distributed M lborel (\<lambda>x. 0 + - 1 * Y x) (\<lambda>x. ennreal (normal_density (0 + - 1 * \<nu>) (\<bar>- 1\<bar> * \<tau>) x))"
     by(rule normal_density_affine, auto)
-  then have [simp]:"distributed M lborel (\<lambda>x. - Y x) (\<lambda>x. ereal (normal_density (- \<nu>) \<tau> x))" by simp
+  then have [simp]:"distributed M lborel (\<lambda>x. - Y x) (\<lambda>x. ennreal (normal_density (- \<nu>) \<tau> x))" by simp
 
   have "distributed M lborel (\<lambda>x. X x + (- Y x)) (normal_density (\<mu> + - \<nu>) (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)))"
     apply (rule sum_indep_normal)
@@ -1340,7 +1330,8 @@
   assumes D: "distributed M lborel X std_normal_density"
   shows "expectation X = 0"
   using integral_std_normal_moment_odd[of 0]
-  by (auto simp: distributed_integral[OF D, of "\<lambda>x. x", symmetric])
+    distributed_integral[OF D, of "\<lambda>x. x", symmetric]
+  by (auto simp: )
 
 lemma (in prob_space) normal_distributed_expectation:
   assumes \<sigma>[arith]: "0 < \<sigma>"
@@ -1368,7 +1359,7 @@
   shows "entropy b lborel X = log b (2 * pi * exp 1 * \<sigma>\<^sup>2) / 2"
 proof -
   have "entropy b lborel X = - (\<integral> x. normal_density \<mu> \<sigma> x * log b (normal_density \<mu> \<sigma> x) \<partial>lborel)"
-    using D by (rule entropy_distr)
+    using D by (rule entropy_distr) simp
   also have "\<dots> = - (\<integral> x. normal_density \<mu> \<sigma> x * (- ln (2 * pi * \<sigma>\<^sup>2) - (x - \<mu>)\<^sup>2 / \<sigma>\<^sup>2) / (2 * ln b) \<partial>lborel)"
     by (intro arg_cong[where f="uminus"] integral_cong)
        (auto simp: normal_density_def field_simps ln_mult log_def ln_div ln_sqrt)
--- a/src/HOL/Probability/Embed_Measure.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Embed_Measure.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -1,8 +1,8 @@
 (*  Title:      HOL/Probability/Embed_Measure.thy
     Author:     Manuel Eberl, TU München
 
-    Defines measure embeddings with injective functions, i.e. lifting a measure on some values 
-    to a measure on "tagged" values (e.g. embed_measure M Inl lifts a measure on values 'a to a 
+    Defines measure embeddings with injective functions, i.e. lifting a measure on some values
+    to a measure on "tagged" values (e.g. embed_measure M Inl lifts a measure on values 'a to a
     measure on the left part of the sum type 'a + 'b)
 *)
 
@@ -13,15 +13,15 @@
 begin
 
 definition embed_measure :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
-  "embed_measure M f = measure_of (f ` space M) {f ` A |A. A \<in> sets M} 
+  "embed_measure M f = measure_of (f ` space M) {f ` A |A. A \<in> sets M}
                            (\<lambda>A. emeasure M (f -` A \<inter> space M))"
 
 lemma space_embed_measure: "space (embed_measure M f) = f ` space M"
-  unfolding embed_measure_def 
+  unfolding embed_measure_def
   by (subst space_measure_of) (auto dest: sets.sets_into_space)
 
 lemma sets_embed_measure':
-  assumes inj: "inj_on f (space M)" 
+  assumes inj: "inj_on f (space M)"
   shows "sets (embed_measure M f) = {f ` A |A. A \<in> sets M}"
   unfolding embed_measure_def
 proof (intro sigma_algebra.sets_measure_of_eq sigma_algebra_iff2[THEN iffD2] conjI allI ballI impI)
@@ -52,7 +52,7 @@
            intro!: image_cong the_inv_into_vimage[symmetric])
 
 lemma sets_embed_measure:
-  assumes inj: "inj f" 
+  assumes inj: "inj f"
   shows "sets (embed_measure M f) = {f ` A |A. A \<in> sets M}"
   using assms by (subst sets_embed_measure') (auto intro!: inj_onI dest: injD)
 
@@ -61,7 +61,7 @@
   by (intro in_measure_of) (auto dest: sets.sets_into_space)
 
 lemma measurable_embed_measure1:
-  assumes g: "(\<lambda>x. g (f x)) \<in> measurable M N" 
+  assumes g: "(\<lambda>x. g (f x)) \<in> measurable M N"
   shows "g \<in> measurable (embed_measure M f) N"
   unfolding measurable_def
 proof safe
@@ -76,7 +76,7 @@
 qed (insert measurable_space[OF assms], auto simp: space_embed_measure)
 
 lemma measurable_embed_measure2':
-  assumes "inj_on f (space M)" 
+  assumes "inj_on f (space M)"
   shows "f \<in> measurable M (embed_measure M f)"
 proof-
   {
@@ -99,7 +99,7 @@
   assumes "inj_on f (space M)"
   shows "embed_measure M f = distr M (embed_measure M f) f"
 proof-
-  have "distr M (embed_measure M f) f = 
+  have "distr M (embed_measure M f) f =
             measure_of (f ` space M) {f ` A |A. A \<in> sets M}
                        (\<lambda>A. emeasure M (f -` A \<inter> space M))" unfolding distr_def
       by (simp add: space_embed_measure sets_embed_measure'[OF assms])
@@ -135,19 +135,19 @@
     shows "emeasure (embed_measure M f) A = emeasure M (f -` A \<inter> space M)"
  using assms by (intro emeasure_embed_measure') (auto intro!: inj_onI dest: injD)
 
-lemma embed_measure_comp: 
+lemma embed_measure_comp:
   assumes [simp]: "inj f" "inj g"
   shows "embed_measure (embed_measure M f) g = embed_measure M (g \<circ> f)"
 proof-
   have [simp]: "inj (\<lambda>x. g (f x))" by (subst o_def[symmetric]) (auto intro: inj_comp)
   note measurable_embed_measure2[measurable]
-  have "embed_measure (embed_measure M f) g = 
+  have "embed_measure (embed_measure M f) g =
             distr M (embed_measure (embed_measure M f) g) (g \<circ> f)"
-      by (subst (1 2) embed_measure_eq_distr) 
+      by (subst (1 2) embed_measure_eq_distr)
          (simp_all add: distr_distr sets_embed_measure cong: distr_cong)
   also have "... = embed_measure M (g \<circ> f)"
       by (subst (3) embed_measure_eq_distr, simp add: o_def, rule distr_cong)
-         (auto simp: sets_embed_measure o_def image_image[symmetric] 
+         (auto simp: sets_embed_measure o_def image_image[symmetric]
                intro: inj_comp cong: distr_cong)
   finally show ?thesis .
 qed
@@ -161,7 +161,7 @@
       A_props: "countable A" "A \<subseteq> sets M" "\<Union>A = space M" "\<And>X. X\<in>A \<Longrightarrow> emeasure M X \<noteq> \<infinity>" by blast
   from A_props have "countable (op ` f`A)" by auto
   moreover
-  from inj and A_props have "op ` f`A \<subseteq> sets (embed_measure M f)" 
+  from inj and A_props have "op ` f`A \<subseteq> sets (embed_measure M f)"
     by (auto simp: sets_embed_measure)
   moreover
   from A_props and inj have "\<Union>(op ` f`A) = space (embed_measure M f)"
@@ -184,7 +184,7 @@
   apply (auto simp: space_embed_measure sets_embed_measure')
   done
 
-lemma embed_measure_count_space: 
+lemma embed_measure_count_space:
     "inj f \<Longrightarrow> embed_measure (count_space A) f = count_space (f`A)"
   by(rule embed_measure_count_space')(erule subset_inj_on, simp)
 
@@ -217,7 +217,7 @@
   moreover {
     fix X assume "X \<in> sets A"
     from asm have "emeasure ?M (f`X) = emeasure ?N (f`X)" by simp
-    with \<open>X \<in> sets A\<close> and \<open>sets A = sets B\<close> and assms 
+    with \<open>X \<in> sets A\<close> and \<open>sets A = sets B\<close> and assms
         have "emeasure A X = emeasure B X" by (simp add: emeasure_embed_measure_image)
   }
   ultimately show "A = B" by (rule measure_eqI)
@@ -240,7 +240,7 @@
 proof (rule pair_measure_eqI)
   have fg[simp]: "\<And>A. inj_on (map_prod f g) A" "\<And>A. inj_on f A" "\<And>A. inj_on g A"
     using f g by (auto simp: inj_on_def)
-  
+
   show sets: "sets ?L = sets (embed_measure (M \<Otimes>\<^sub>M N) (map_prod f g))"
     unfolding map_prod_def[symmetric]
     apply (simp add: sets_pair_eq_sets_fst_snd sets_embed_eq_vimage_algebra
@@ -275,18 +275,17 @@
   shows "density (embed_measure M f) g = embed_measure (density M (g \<circ> f)) f" (is "?M1 = ?M2")
 proof (rule measure_eqI)
   fix X assume X: "X \<in> sets ?M1"
-  from inj have Mf[measurable]: "f \<in> measurable M (embed_measure M f)" 
+  from inj have Mf[measurable]: "f \<in> measurable M (embed_measure M f)"
     by (rule measurable_embed_measure2)
-  from Mg and X have "emeasure ?M1 X = \<integral>\<^sup>+ x. g x * indicator X x \<partial>embed_measure M f" 
+  from Mg and X have "emeasure ?M1 X = \<integral>\<^sup>+ x. g x * indicator X x \<partial>embed_measure M f"
     by (subst emeasure_density) simp_all
   also from X have "... = \<integral>\<^sup>+ x. g (f x) * indicator X (f x) \<partial>M"
-    by (subst embed_measure_eq_distr[OF inj], subst nn_integral_distr)
-       (auto intro!: borel_measurable_ereal_times borel_measurable_indicator)
+    by (subst embed_measure_eq_distr[OF inj], subst nn_integral_distr) auto
   also have "... = \<integral>\<^sup>+ x. g (f x) * indicator (f -` X \<inter> space M) x \<partial>M"
     by (intro nn_integral_cong) (auto split: split_indicator)
   also from X have "... = emeasure (density M (g \<circ> f)) (f -` X \<inter> space M)"
     by (subst emeasure_density) (simp_all add: measurable_comp[OF Mf Mg] measurable_sets[OF Mf])
-  also from X and inj have "... = emeasure ?M2 X" 
+  also from X and inj have "... = emeasure ?M2 X"
     by (subst emeasure_embed_measure) (simp_all add: sets_embed_measure)
   finally show "emeasure ?M1 X = emeasure ?M2 X" .
 qed (simp_all add: sets_embed_measure inj)
@@ -297,7 +296,7 @@
 proof-
   have "density (embed_measure M f) (g \<circ> f') = embed_measure (density M (g \<circ> f' \<circ> f)) f"
     by (rule density_embed_measure[OF inj])
-       (rule measurable_comp, rule measurable_embed_measure1, subst measurable_cong, 
+       (rule measurable_comp, rule measurable_embed_measure1, subst measurable_cong,
         rule inv, rule measurable_ident_sets, simp, rule Mg)
   also have "density M (g \<circ> f' \<circ> f) = density M g"
     by (intro density_cong) (subst measurable_cong, simp add: o_def inv, simp_all add: Mg inv)
@@ -315,7 +314,7 @@
   with assms and B have "x = y" by (auto dest: inj_onD)
   with \<open>y \<in> B\<close> show "x \<in> B" by simp
 qed auto
-  
+
 
 lemma AE_embed_measure':
   assumes inj: "inj_on f (space M)"
@@ -326,7 +325,7 @@
   then obtain A where A_props: "A \<in> sets ?M" "emeasure ?M A = 0" "{x\<in>space ?M. \<not>P x} \<subseteq> A"
     by (force elim: AE_E)
   then obtain A' where A'_props: "A = f ` A'" "A' \<in> sets M" by (auto simp: sets_embed_measure' inj)
-  moreover have B: "{x\<in>space ?M. \<not>P x} = f ` {x\<in>space M. \<not>P (f x)}" 
+  moreover have B: "{x\<in>space ?M. \<not>P x} = f ` {x\<in>space M. \<not>P (f x)}"
     by (auto simp: inj space_embed_measure)
   from A_props(3) have "{x\<in>space M. \<not>P (f x)} \<subseteq> A'"
     by (subst (asm) B, subst (asm) A'_props, subst (asm) inj_on_image_subset_iff[OF inj])
@@ -350,7 +349,7 @@
   using assms by (intro AE_embed_measure') (auto intro!: inj_onI dest: injD)
 
 lemma nn_integral_monotone_convergence_SUP_countable:
-  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> ereal"
+  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> ennreal"
   assumes nonempty: "Y \<noteq> {}"
   and chain: "Complete_Partial_Order.chain op \<le> (f ` Y)"
   and countable: "countable B"
@@ -363,17 +362,17 @@
   also have "\<dots> = \<integral>\<^sup>+ x. (SUP i:Y. f i (from_nat_into B x)) \<partial>count_space (to_nat_on B ` B)"
     by(simp add: embed_measure_count_space'[symmetric] inj_on_to_nat_on countable nn_integral_embed_measure' measurable_embed_measure1)
   also have "\<dots> = \<integral>\<^sup>+ x. (SUP i:Y. ?f i x) \<partial>count_space UNIV"
-    by(simp add: nn_integral_count_space_indicator ereal_indicator[symmetric] Sup_ereal_mult_right' nonempty)
-  also have "\<dots> = (SUP i:Y. \<integral>\<^sup>+ x. ?f i x \<partial>count_space UNIV)" using nonempty
+    by(simp add: nn_integral_count_space_indicator ennreal_indicator[symmetric] SUP_mult_right_ennreal nonempty)
+  also have "\<dots> = (SUP i:Y. \<integral>\<^sup>+ x. ?f i x \<partial>count_space UNIV)"
   proof(rule nn_integral_monotone_convergence_SUP_nat)
     show "Complete_Partial_Order.chain op \<le> (?f ` Y)"
       by(rule chain_imageI[OF chain, unfolded image_image])(auto intro!: le_funI split: split_indicator dest: le_funD)
-  qed
+  qed fact
   also have "\<dots> = (SUP i:Y. \<integral>\<^sup>+ x. f i (from_nat_into B x) \<partial>count_space (to_nat_on B ` B))"
     by(simp add: nn_integral_count_space_indicator)
   also have "\<dots> = (SUP i:Y. \<integral>\<^sup>+ x. f i (from_nat_into B (to_nat_on B x)) \<partial>count_space B)"
     by(simp add: embed_measure_count_space'[symmetric] inj_on_to_nat_on countable nn_integral_embed_measure' measurable_embed_measure1)
-  also have "\<dots> = ?rhs" 
+  also have "\<dots> = ?rhs"
     by(intro arg_cong2[where f="SUPREMUM"] ext nn_integral_cong_AE)(simp_all add: AE_count_space countable)
   finally show ?thesis .
 qed
--- a/src/HOL/Probability/Finite_Product_Measure.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Finite_Product_Measure.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -116,7 +116,7 @@
 definition prod_emb where
   "prod_emb I M K X = (\<lambda>x. restrict x K) -` X \<inter> (PIE i:I. space (M i))"
 
-lemma prod_emb_iff: 
+lemma prod_emb_iff:
   "f \<in> prod_emb I M K X \<longleftrightarrow> f \<in> extensional I \<and> (restrict f K \<in> X) \<and> (\<forall>i\<in>I. f i \<in> space (M i))"
   unfolding prod_emb_def PiE_def by auto
 
@@ -180,7 +180,7 @@
 
 lemma Pi_cong_sets:
     "\<lbrakk>I = J; \<And>x. x \<in> I \<Longrightarrow> M x = N x\<rbrakk> \<Longrightarrow> Pi I M = Pi J N"
-  unfolding Pi_def by auto 
+  unfolding Pi_def by auto
 
 lemma PiM_cong:
   assumes "I = J" "\<And>x. x \<in> I \<Longrightarrow> M x = N x"
@@ -197,11 +197,11 @@
   thus ?case by (auto simp: assms)
 next
   case 3
-  show ?case using assms 
+  show ?case using assms
     by (intro ext) (auto simp: prod_emb_def dest: PiE_mem)
 next
   case (4 x)
-  thus ?case using assms 
+  thus ?case using assms
     by (auto intro!: setprod.cong split: if_split_asm)
 qed
 
@@ -251,7 +251,7 @@
 lemma prod_algebraE:
   assumes A: "A \<in> prod_algebra I M"
   obtains J E where "A = prod_emb I M J (PIE j:J. E j)"
-    "finite J" "J \<noteq> {} \<or> I = {}" "J \<subseteq> I" "\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i)" 
+    "finite J" "J \<noteq> {} \<or> I = {}" "J \<subseteq> I" "\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i)"
   using A by (auto simp: prod_algebra_def)
 
 lemma prod_algebraE_all:
@@ -276,7 +276,7 @@
   from prod_algebraE[OF this] guess J E . note A = this
   fix B assume "B \<in> prod_algebra I M"
   from prod_algebraE[OF this] guess K F . note B = this
-  have "A \<inter> B = prod_emb I M (J \<union> K) (\<Pi>\<^sub>E i\<in>J \<union> K. (if i \<in> J then E i else space (M i)) \<inter> 
+  have "A \<inter> B = prod_emb I M (J \<union> K) (\<Pi>\<^sub>E i\<in>J \<union> K. (if i \<in> J then E i else space (M i)) \<inter>
       (if i \<in> K then F i else space (M i)))"
     unfolding A B using A(2,3,4) A(5)[THEN sets.sets_into_space] B(2,3,4)
       B(5)[THEN sets.sets_into_space]
@@ -379,7 +379,7 @@
   qed
 next
   fix A assume "A \<in> ?R"
-  then obtain i B where A: "A = {f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> B}" "i \<in> I" "B \<in> sets (M i)" 
+  then obtain i B where A: "A = {f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> B}" "i \<in> I" "B \<in> sets (M i)"
     by auto
   then have "A = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. B)"
      by (auto simp: prod_emb_def)
@@ -409,7 +409,7 @@
   defines "P \<equiv> {{f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A i} | A j. j \<in> J \<and> A \<in> Pi j E}"
   shows "sets (\<Pi>\<^sub>M i\<in>I. sigma (\<Omega> i) (E i)) = sets (sigma (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P)"
 proof cases
-  assume "I = {}" 
+  assume "I = {}"
   with \<open>\<Union>J = I\<close> have "P = {{\<lambda>_. undefined}} \<or> P = {}"
     by (auto simp: P_def)
   with \<open>I = {}\<close> show ?thesis
@@ -417,7 +417,7 @@
 next
   let ?F = "\<lambda>i. {(\<lambda>x. x i) -` A \<inter> Pi\<^sub>E I \<Omega> |A. A \<in> E i}"
   assume "I \<noteq> {}"
-  then have "sets (Pi\<^sub>M I (\<lambda>i. sigma (\<Omega> i) (E i))) = 
+  then have "sets (Pi\<^sub>M I (\<lambda>i. sigma (\<Omega> i) (E i))) =
       sets (\<Squnion>\<^sub>\<sigma> i\<in>I. vimage_algebra (\<Pi>\<^sub>E i\<in>I. \<Omega> i) (\<lambda>x. x i) (sigma (\<Omega> i) (E i)))"
     by (subst sets_PiM_eq_proj) (auto simp: space_measure_of_conv)
   also have "\<dots> = sets (\<Squnion>\<^sub>\<sigma> i\<in>I. sigma (Pi\<^sub>E I \<Omega>) (?F i))"
@@ -520,7 +520,7 @@
 lemma measurable_PiM:
   assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
   assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
-    f -` prod_emb I M J (Pi\<^sub>E J X) \<inter> space N \<in> sets N" 
+    f -` prod_emb I M J (Pi\<^sub>E J X) \<inter> space N \<in> sets N"
   shows "f \<in> measurable N (PiM I M)"
   using sets_PiM prod_algebra_sets_into_space space
 proof (rule measurable_sigma_sets)
@@ -532,7 +532,7 @@
 lemma measurable_PiM_Collect:
   assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
   assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
-    {\<omega>\<in>space N. \<forall>i\<in>J. f \<omega> i \<in> X i} \<in> sets N" 
+    {\<omega>\<in>space N. \<forall>i\<in>J. f \<omega> i \<in> X i} \<in> sets N"
   shows "f \<in> measurable N (PiM I M)"
   using sets_PiM prod_algebra_sets_into_space space
 proof (rule measurable_sigma_sets)
@@ -546,7 +546,7 @@
 
 lemma measurable_PiM_single:
   assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
-  assumes sets: "\<And>A i. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {\<omega> \<in> space N. f \<omega> i \<in> A} \<in> sets N" 
+  assumes sets: "\<And>A i. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {\<omega> \<in> space N. f \<omega> i \<in> A} \<in> sets N"
   shows "f \<in> measurable N (PiM I M)"
   using sets_PiM_single
 proof (rule measurable_sigma_sets)
@@ -602,7 +602,7 @@
     "\<And>j. i = Suc j \<Longrightarrow> (\<lambda>x. g x j) \<in> measurable M N"
   shows "(\<lambda>x. case_nat (f x) (g x) i) \<in> measurable M N"
   by (cases i) simp_all
- 
+
 lemma measurable_case_nat'[measurable (raw)]:
   assumes fg[measurable]: "f \<in> measurable N M" "g \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
   shows "(\<lambda>x. case_nat (f x) (g x)) \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
@@ -666,7 +666,7 @@
     using A X by (auto intro!: measurable_sets)
 qed (insert X, auto simp add: PiE_def dest: measurable_space)
 
-lemma measurable_abs_UNIV: 
+lemma measurable_abs_UNIV:
   "(\<And>n. (\<lambda>\<omega>. f n \<omega>) \<in> measurable M (N n)) \<Longrightarrow> (\<lambda>\<omega> n. f n \<omega>) \<in> measurable M (PiM UNIV N)"
   by (intro measurable_PiM_single) (auto dest: measurable_space)
 
@@ -705,7 +705,7 @@
   fix x assume "x \<in> X"
   obtain \<omega> where "\<omega> \<in> space (PiM I M)"
     using ne by blast
-  from merge_in_prod_emb[OF this \<open>x\<in>X\<close> X J] * show "x \<in> {}" by auto 
+  from merge_in_prod_emb[OF this \<open>x\<in>X\<close> X J] * show "x \<in> {}" by auto
 qed
 
 lemma sets_in_Pi_aux:
@@ -775,7 +775,7 @@
   show ?case
   proof (intro exI[of _ "\<Union>i. J i"] bexI[of _ "\<Union>i. prod_emb (\<Union>i. J i) M (J i) (X i)"] conjI)
     show "(\<Union>i. J i) \<subseteq> I" "countable (\<Union>i. J i)" using J by auto
-    with J show "UNION UNIV K = prod_emb I M (\<Union>i. J i) (\<Union>i. prod_emb (\<Union>i. J i) M (J i) (X i))" 
+    with J show "UNION UNIV K = prod_emb I M (\<Union>i. J i) (\<Union>i. prod_emb (\<Union>i. J i) M (J i) (X i))"
       by (simp add: K[abs_def] SUP_upper)
   qed(auto intro: X)
 qed
@@ -867,7 +867,7 @@
 
 lemma emeasure_PiM_empty[simp]: "emeasure (PiM {} M) {\<lambda>_. undefined} = 1"
 proof -
-  let ?\<mu> = "\<lambda>A. if A = {} then 0 else (1::ereal)"
+  let ?\<mu> = "\<lambda>A. if A = {} then 0 else (1::ennreal)"
   have "emeasure (Pi\<^sub>M {} M) (prod_emb {} M {} (\<Pi>\<^sub>E i\<in>{}. {})) = 1"
   proof (subst emeasure_extend_measure_Pair[OF PiM_def])
     show "positive (PiM {} M) ?\<mu>"
@@ -883,7 +883,7 @@
 qed
 
 lemma PiM_empty: "PiM {} M = count_space {\<lambda>_. undefined}"
-  by (rule measure_eqI) (auto simp add: sets_PiM_empty one_ereal_def)
+  by (rule measure_eqI) (auto simp add: sets_PiM_empty)
 
 lemma (in product_sigma_finite) emeasure_PiM:
   "finite I \<Longrightarrow> (\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (PiM I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
@@ -956,11 +956,11 @@
       by (simp add: eq emeasure_PiM)
     def A \<equiv> "\<lambda>n. \<Pi>\<^sub>E i\<in>I. C i n"
     with C show "range A \<subseteq> prod_algebra I M" "\<And>i. emeasure (Pi\<^sub>M I M) (A i) \<noteq> \<infinity>" "(\<Union>i. A i) = space (PiM I M)"
-      by (auto intro!: prod_algebraI_finite simp: emeasure_PiM subset_eq setprod_PInf emeasure_nonneg)
+      by (auto intro!: prod_algebraI_finite simp: emeasure_PiM subset_eq ennreal_setprod_eq_top)
   qed
 qed
 
-lemma (in product_sigma_finite) sigma_finite: 
+lemma (in product_sigma_finite) sigma_finite:
   assumes "finite I"
   shows "sigma_finite_measure (PiM I M)"
 proof
@@ -975,7 +975,7 @@
   ultimately show "\<exists>A. countable A \<and> A \<subseteq> sets (Pi\<^sub>M I M) \<and> \<Union>A = space (Pi\<^sub>M I M) \<and> (\<forall>a\<in>A. emeasure (Pi\<^sub>M I M) a \<noteq> \<infinity>)"
     by (intro exI[of _ "PiE I ` PiE I F"])
        (auto intro!: countable_PiE sets_PiM_I_finite
-             simp: PiE_iff emeasure_PiM finite_index setprod_PInf emeasure_nonneg)
+             simp: PiE_iff emeasure_PiM finite_index ennreal_setprod_eq_top)
 qed
 
 sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure "Pi\<^sub>M I M"
@@ -1007,7 +1007,7 @@
 
 lemma (in product_sigma_finite) product_nn_integral_fold:
   assumes IJ: "I \<inter> J = {}" "finite I" "finite J"
-  and f: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)"
+  and f[measurable]: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)"
   shows "integral\<^sup>N (Pi\<^sub>M (I \<union> J) M) f =
     (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (merge I J (x, y)) \<partial>(Pi\<^sub>M J M)) \<partial>(Pi\<^sub>M I M))"
 proof -
@@ -1018,7 +1018,8 @@
     using measurable_comp[OF measurable_merge f] by (simp add: comp_def)
   show ?thesis
     apply (subst distr_merge[OF IJ, symmetric])
-    apply (subst nn_integral_distr[OF measurable_merge f])
+    apply (subst nn_integral_distr[OF measurable_merge])
+    apply measurable []
     apply (subst J.nn_integral_fst[symmetric, OF P_borel])
     apply simp
     done
@@ -1084,26 +1085,25 @@
   done
 
 lemma (in product_sigma_finite) product_nn_integral_setprod:
-  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ereal"
-  assumes "finite I" and borel: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
-  and pos: "\<And>i x. i \<in> I \<Longrightarrow> 0 \<le> f i x"
+  assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
   shows "(\<integral>\<^sup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>I. integral\<^sup>N (M i) (f i))"
-using assms proof induct
+using assms proof (induction I)
   case (insert i I)
+  note insert.prems[measurable]
   note \<open>finite I\<close>[intro, simp]
   interpret I: finite_product_sigma_finite M I by standard auto
   have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
     using insert by (auto intro!: setprod.cong)
   have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^sub>M J M)"
     using sets.sets_into_space insert
-    by (intro borel_measurable_ereal_setprod
+    by (intro borel_measurable_setprod_ennreal
               measurable_comp[OF measurable_component_singleton, unfolded comp_def])
        auto
   then show ?case
-    apply (simp add: product_nn_integral_insert[OF insert(1,2) prod])
-    apply (simp add: insert(2-) * pos borel setprod_ereal_pos nn_integral_multc)
+    apply (simp add: product_nn_integral_insert[OF insert(1,2)])
+    apply (simp add: insert(2-) * nn_integral_multc)
     apply (subst nn_integral_cmult)
-    apply (auto simp add: pos borel insert(2-) setprod_ereal_pos nn_integral_nonneg)
+    apply (auto simp add: insert(2-))
     done
 qed (simp add: space_PiM)
 
--- a/src/HOL/Probability/Giry_Monad.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Giry_Monad.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -7,7 +7,7 @@
 *)
 
 theory Giry_Monad
-  imports Probability_Measure Lebesgue_Integral_Substitution "~~/src/HOL/Library/Monad_Syntax" 
+  imports Probability_Measure Lebesgue_Integral_Substitution "~~/src/HOL/Library/Monad_Syntax"
 begin
 
 section \<open>Sub-probability spaces\<close>
@@ -23,7 +23,7 @@
 proof -
   interpret finite_measure M
   proof
-    show "emeasure M (space M) \<noteq> \<infinity>" using * by auto
+    show "emeasure M (space M) \<noteq> \<infinity>" using * by (auto simp: top_unique)
   qed
   show "subprob_space M" by standard fact+
 qed
@@ -66,8 +66,8 @@
   have "(\<integral>\<^sup>+ x. f x \<partial>M) \<le> (\<integral>\<^sup>+ x. c \<partial>M)"
     by(rule nn_integral_mono_AE) fact
   also have "\<dots> \<le> c * emeasure M (space M)"
-    using \<open>0 \<le> c\<close> by(simp add: nn_integral_const_If)
-  also have "\<dots> \<le> c * 1" using emeasure_space_le_1 \<open>0 \<le> c\<close> by(rule ereal_mult_left_mono)
+    using \<open>0 \<le> c\<close> by simp
+  also have "\<dots> \<le> c * 1" using emeasure_space_le_1 \<open>0 \<le> c\<close> by(rule mult_left_mono)
   finally show ?thesis by simp
 qed
 
@@ -84,7 +84,7 @@
   assumes contg': "continuous_on {a..b} g'"
   assumes mono: "strict_mono_on g {a..b}" and inv: "\<And>x. h x \<in> {a..b} \<Longrightarrow> g (h x) = x"
   assumes range: "{a..b} \<subseteq> range h"
-  shows "emeasure (distr (density lborel f) lborel h) {a..b} = 
+  shows "emeasure (distr (density lborel f) lborel h) {a..b} =
              emeasure (density lborel (\<lambda>x. f (g x) * g' x)) {a..b}"
 proof (cases "a < b")
   assume "a < b"
@@ -109,18 +109,18 @@
 
   have prob': "subprob_space (distr (density lborel f) lborel h)"
     by (rule subprob_space.subprob_space_distr[OF prob]) (simp_all add: Mh)
-  have B: "emeasure (distr (density lborel f) lborel h) {a..b} = 
+  have B: "emeasure (distr (density lborel f) lborel h) {a..b} =
             \<integral>\<^sup>+x. f x * indicator (h -` {a..b}) x \<partial>lborel"
     by (subst emeasure_distr) (simp_all add: emeasure_density Mf Mh measurable_sets_borel[OF Mh])
   also note A
   also have "emeasure (distr (density lborel f) lborel h) {a..b} \<le> 1"
     by (rule subprob_space.subprob_emeasure_le_1) (rule prob')
-  hence "emeasure (distr (density lborel f) lborel h) {a..b} \<noteq> \<infinity>" by auto
-  with assms have "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) = 
+  hence "emeasure (distr (density lborel f) lborel h) {a..b} \<noteq> \<infinity>" by (auto simp: top_unique)
+  with assms have "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) =
                       (\<integral>\<^sup>+x. f (g x) * g' x * indicator {a..b} x \<partial>lborel)"
     by (intro nn_integral_substitution_aux)
        (auto simp: derivg_nonneg A B emeasure_density mult.commute \<open>a < b\<close>)
-  also have "... = emeasure (density lborel (\<lambda>x. f (g x) * g' x)) {a..b}" 
+  also have "... = emeasure (density lborel (\<lambda>x. f (g x) * g' x)) {a..b}"
     by (simp add: emeasure_density)
   finally show ?thesis .
 next
@@ -130,16 +130,14 @@
   thus ?thesis by (simp_all add: emeasure_distr emeasure_density measurable_sets_borel[OF Mh])
 qed
 
-locale pair_subprob_space = 
+locale pair_subprob_space =
   pair_sigma_finite M1 M2 + M1: subprob_space M1 + M2: subprob_space M2 for M1 M2
 
 sublocale pair_subprob_space \<subseteq> P?: subprob_space "M1 \<Otimes>\<^sub>M M2"
 proof
-  have "\<And>a b. \<lbrakk>a \<ge> 0; b \<ge> 0; a \<le> 1; b \<le> 1\<rbrakk> \<Longrightarrow> a * b \<le> (1::ereal)"
-    by (metis monoid_mult_class.mult.left_neutral dual_order.trans ereal_mult_right_mono)
-  from this[OF _ _ M1.emeasure_space_le_1 M2.emeasure_space_le_1]
-    show "emeasure (M1 \<Otimes>\<^sub>M M2) (space (M1 \<Otimes>\<^sub>M M2)) \<le> 1"
-    by (simp add: M2.emeasure_pair_measure_Times space_pair_measure emeasure_nonneg)
+  from mult_le_one[OF M1.emeasure_space_le_1 _ M2.emeasure_space_le_1]
+  show "emeasure (M1 \<Otimes>\<^sub>M M2) (space (M1 \<Otimes>\<^sub>M M2)) \<le> 1"
+    by (simp add: M2.emeasure_pair_measure_Times space_pair_measure)
   from M1.subprob_not_empty and M2.subprob_not_empty show "space (M1 \<Otimes>\<^sub>M M2) \<noteq> {}"
     by (simp add: space_pair_measure)
 qed
@@ -172,10 +170,14 @@
 lemma subprob_algebra_cong: "sets M = sets N \<Longrightarrow> subprob_algebra M = subprob_algebra N"
   by (simp add: subprob_algebra_def)
 
-lemma measurable_emeasure_subprob_algebra[measurable]: 
+lemma measurable_emeasure_subprob_algebra[measurable]:
   "a \<in> sets A \<Longrightarrow> (\<lambda>M. emeasure M a) \<in> borel_measurable (subprob_algebra A)"
   by (auto intro!: measurable_Sup_sigma1 measurable_vimage_algebra1 simp: subprob_algebra_def)
 
+lemma measurable_measure_subprob_algebra[measurable]:
+  "a \<in> sets A \<Longrightarrow> (\<lambda>M. measure M a) \<in> borel_measurable (subprob_algebra A)"
+  unfolding measure_def by measurable
+
 lemma subprob_measurableD:
   assumes N: "N \<in> measurable M (subprob_algebra S)" and x: "x \<in> space M"
   shows "space (N x) = space S"
@@ -214,7 +216,7 @@
 lemma sets_kernel: "a \<in> space M \<Longrightarrow> sets (K a) = sets N"
   using measurable_space[OF K] by (simp add: space_subprob_algebra)
 
-lemma measurable_emeasure_kernel[measurable]: 
+lemma measurable_emeasure_kernel[measurable]:
     "A \<in> sets N \<Longrightarrow> (\<lambda>a. emeasure (K a) A) \<in> borel_measurable M"
   using measurable_compose[OF K measurable_emeasure_subprob_algebra] .
 
@@ -258,8 +260,8 @@
   ultimately show "space (subprob_algebra N) = {}" by (auto simp: space_subprob_algebra)
 qed
 
-lemma nn_integral_measurable_subprob_algebra':
-  assumes f: "f \<in> borel_measurable N" "\<And>x. 0 \<le> f x"
+lemma nn_integral_measurable_subprob_algebra[measurable]:
+  assumes f: "f \<in> borel_measurable N"
   shows "(\<lambda>M. integral\<^sup>N M f) \<in> borel_measurable (subprob_algebra N)" (is "_ \<in> ?B")
   using f
 proof induct
@@ -295,11 +297,6 @@
     by (simp add: ac_simps)
 qed
 
-lemma nn_integral_measurable_subprob_algebra:
-  "f \<in> borel_measurable N \<Longrightarrow> (\<lambda>M. integral\<^sup>N M f) \<in> borel_measurable (subprob_algebra N)"
-  by (subst nn_integral_max_0[symmetric])
-     (auto intro!: nn_integral_measurable_subprob_algebra')
-
 lemma measurable_distr:
   assumes [measurable]: "f \<in> measurable M N"
   shows "(\<lambda>M'. distr M' N f) \<in> measurable (subprob_algebra M) (subprob_algebra N)"
@@ -329,33 +326,56 @@
   finally show ?thesis .
 qed
 
-lemma integral_measurable_subprob_algebra:
-  fixes f :: "_ \<Rightarrow> real"
-  assumes f_measurable [measurable]: "f \<in> borel_measurable N"
-  and f_bounded: "\<And>x. x \<in> space N \<Longrightarrow> \<bar>f x\<bar> \<le> B"
-  shows "(\<lambda>M. integral\<^sup>L M f) \<in> borel_measurable (subprob_algebra N)"
+lemma integrable_measurable_subprob_algebra[measurable]:
+  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+  assumes [measurable]: "f \<in> borel_measurable N"
+  shows "Measurable.pred (subprob_algebra N) (\<lambda>M. integrable M f)"
+proof (rule measurable_cong[THEN iffD2])
+  show "M \<in> space (subprob_algebra N) \<Longrightarrow> integrable M f \<longleftrightarrow> (\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>" for M
+    by (auto simp: space_subprob_algebra integrable_iff_bounded)
+qed measurable
+
+lemma integral_measurable_subprob_algebra[measurable]:
+  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+  assumes f [measurable]: "f \<in> borel_measurable N"
+  shows "(\<lambda>M. integral\<^sup>L M f) \<in> subprob_algebra N \<rightarrow>\<^sub>M borel"
 proof -
-  note [measurable] = nn_integral_measurable_subprob_algebra
-  have "?thesis \<longleftrightarrow> (\<lambda>M. real_of_ereal (\<integral>\<^sup>+ x. f x \<partial>M) - real_of_ereal (\<integral>\<^sup>+ x. - f x \<partial>M)) \<in> borel_measurable (subprob_algebra N)"
-  proof(rule measurable_cong)
-    fix M
-    assume "M \<in> space (subprob_algebra N)"
-    hence "subprob_space M" and M [measurable_cong]: "sets M = sets N" 
-      by(simp_all add: space_subprob_algebra)
+  from borel_measurable_implies_sequence_metric[OF f, of 0]
+  obtain F where F: "\<And>i. simple_function N (F i)"
+    "\<And>x. x \<in> space N \<Longrightarrow> (\<lambda>i. F i x) \<longlonglongrightarrow> f x"
+    "\<And>i x. x \<in> space N \<Longrightarrow> norm (F i x) \<le> 2 * norm (f x)"
+    unfolding norm_conv_dist by blast
+
+  have [measurable]: "F i \<in> N \<rightarrow>\<^sub>M count_space UNIV" for i
+    using F(1) by (rule measurable_simple_function)
+
+  def F' \<equiv> "\<lambda>M i. if integrable M f then integral\<^sup>L M (F i) else 0"
+
+  have "(\<lambda>M. F' M i) \<in> subprob_algebra N \<rightarrow>\<^sub>M borel" for i
+  proof (rule measurable_cong[THEN iffD2])
+    fix M assume "M \<in> space (subprob_algebra N)"
+    then have [simp]: "sets M = sets N" "space M = space N" "subprob_space M"
+      by (auto simp: space_subprob_algebra intro!: sets_eq_imp_space_eq)
     interpret subprob_space M by fact
-    have "(\<integral>\<^sup>+ x. ereal \<bar>f x\<bar> \<partial>M) \<le> (\<integral>\<^sup>+ x. ereal B \<partial>M)"
-      by(rule nn_integral_mono)(simp add: sets_eq_imp_space_eq[OF M] f_bounded)
-    also have "\<dots> = max B 0 * emeasure M (space M)" by(simp add: nn_integral_const_If max_def)
-    also have "\<dots> \<le> ereal (max B 0) * 1"
-      by(rule ereal_mult_left_mono)(simp_all add: emeasure_space_le_1 zero_ereal_def)
-    finally have "(\<integral>\<^sup>+ x. ereal \<bar>f x\<bar> \<partial>M) \<noteq> \<infinity>" by(auto simp add: max_def)
-    then have "integrable M f" using f_measurable
-      by(auto intro: integrableI_bounded)
-    thus "(\<integral> x. f x \<partial>M) = real_of_ereal (\<integral>\<^sup>+ x. f x \<partial>M) - real_of_ereal (\<integral>\<^sup>+ x. - f x \<partial>M)"
-      by(simp add: real_lebesgue_integral_def)
-  qed
-  also have "\<dots>" by measurable
-  finally show ?thesis .
+    have "F' M i = (if integrable M f then Bochner_Integration.simple_bochner_integral M (F i) else 0)"
+      using F(1)
+      by (subst simple_bochner_integrable_eq_integral)
+         (auto simp: simple_bochner_integrable.simps simple_function_def F'_def)
+    then show "F' M i = (if integrable M f then \<Sum>y\<in>F i ` space N. measure M {x\<in>space N. F i x = y} *\<^sub>R y else 0)"
+      unfolding simple_bochner_integral_def by simp
+  qed measurable
+  moreover
+  have "F' M \<longlonglongrightarrow> integral\<^sup>L M f" if M: "M \<in> space (subprob_algebra N)" for M
+  proof cases
+    from M have [simp]: "sets M = sets N" "space M = space N"
+      by (auto simp: space_subprob_algebra intro!: sets_eq_imp_space_eq)
+    assume "integrable M f" then show ?thesis
+      unfolding F'_def using F(1)[THEN borel_measurable_simple_function] F
+      by (auto intro!: integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]
+               cong: measurable_cong_sets)
+  qed (auto simp: F'_def not_integrable_integral_eq)
+  ultimately show ?thesis
+    by (rule borel_measurable_LIMSEQ_metric)
 qed
 
 (* TODO: Rename. This name is too general -- Manuel *)
@@ -379,13 +399,13 @@
       using fx gx by (simp add: space_subprob_algebra)
 
     have 1: "\<And>A B. A \<in> sets N \<Longrightarrow> B \<in> sets L \<Longrightarrow> emeasure (f x \<Otimes>\<^sub>M g x) (A \<times> B) = emeasure (f x) A * emeasure (g x) B"
-      using fx gx by (intro G.emeasure_pair_measure_Times) (auto simp: space_subprob_algebra) 
-    have "emeasure (f x \<Otimes>\<^sub>M g x) (space (f x \<Otimes>\<^sub>M g x)) = 
+      using fx gx by (intro G.emeasure_pair_measure_Times) (auto simp: space_subprob_algebra)
+    have "emeasure (f x \<Otimes>\<^sub>M g x) (space (f x \<Otimes>\<^sub>M g x)) =
               emeasure (f x) (space (f x)) * emeasure (g x) (space (g x))"
       by (subst G.emeasure_pair_measure_Times[symmetric]) (simp_all add: space_pair_measure)
     hence 2: "\<And>A. A \<in> sets (N \<Otimes>\<^sub>M L) \<Longrightarrow> emeasure (f x \<Otimes>\<^sub>M g x) (space N \<times> space L - A) =
                                              ... - emeasure (f x \<Otimes>\<^sub>M g x) A"
-      using emeasure_compl[OF _ P.emeasure_finite]
+      using emeasure_compl[simplified, OF _ P.emeasure_finite]
       unfolding sets_eq
       unfolding sets_eq_imp_space_eq[OF sets_eq]
       by (simp add: space_pair_measure G.emeasure_pair_measure_Times)
@@ -398,13 +418,13 @@
     unfolding sets_pair_measure
   proof (induct A rule: sigma_sets_induct_disjoint)
     case (basic A) then show ?case
-      by (auto intro!: borel_measurable_ereal_times simp: Times cong: measurable_cong)
+      by (auto intro!: borel_measurable_times_ennreal simp: Times cong: measurable_cong)
          (auto intro!: measurable_emeasure_kernel f g)
   next
     case (compl A)
     then have A: "A \<in> sets (N \<Otimes>\<^sub>M L)"
       by (auto simp: sets_pair_measure)
-    have "(\<lambda>x. emeasure (f x) (space (f x)) * emeasure (g x) (space (g x)) - 
+    have "(\<lambda>x. emeasure (f x) (space (f x)) * emeasure (g x) (space (g x)) -
                    emeasure (f x \<Otimes>\<^sub>M g x) A) \<in> borel_measurable M" (is "?f \<in> ?M")
       using compl(2) f g by measurable
     thus ?case by (simp add: Compl A cong: measurable_cong)
@@ -472,10 +492,10 @@
   by (simp add: return_def)
 
 lemma measurable_return1[simp]: "measurable (return N x) L = measurable N L"
-  by (simp cong: measurable_cong_sets) 
+  by (simp cong: measurable_cong_sets)
 
 lemma measurable_return2[simp]: "measurable L (return N x) = measurable L N"
-  by (simp cong: measurable_cong_sets) 
+  by (simp cong: measurable_cong_sets)
 
 lemma return_sets_cong: "sets M = sets N \<Longrightarrow> return M = return N"
   by (auto dest: sets_eq_imp_space_eq simp: fun_eq_iff return_def)
@@ -491,7 +511,7 @@
   show "positive (sets (return M x)) (\<lambda>A. indicator A x)" by (simp add: positive_def)
   from assms show "A \<in> sets (return M x)" unfolding return_def by simp
   show "countably_additive (sets (return M x)) (\<lambda>A. indicator A x)"
-    by (auto intro: countably_additiveI simp: suminf_indicator)
+    by (auto intro!: countably_additiveI suminf_indicator)
 qed
 
 lemma prob_space_return: "x \<in> space M \<Longrightarrow> prob_space (return M x)"
@@ -500,7 +520,7 @@
 lemma subprob_space_return: "x \<in> space M \<Longrightarrow> subprob_space (return M x)"
   by (intro prob_space_return prob_space_imp_subprob_space)
 
-lemma subprob_space_return_ne: 
+lemma subprob_space_return_ne:
   assumes "space M \<noteq> {}" shows "subprob_space (return M x)"
 proof
   show "emeasure (return M x) (space (return M x)) \<le> 1"
@@ -520,16 +540,16 @@
     by (rule AE_cong) auto
   finally show ?thesis .
 qed
-  
+
 lemma nn_integral_return:
-  assumes "g x \<ge> 0" "x \<in> space M" "g \<in> borel_measurable M"
+  assumes "x \<in> space M" "g \<in> borel_measurable M"
   shows "(\<integral>\<^sup>+ a. g a \<partial>return M x) = g x"
 proof-
   interpret prob_space "return M x" by (rule prob_space_return[OF \<open>x \<in> space M\<close>])
   have "(\<integral>\<^sup>+ a. g a \<partial>return M x) = (\<integral>\<^sup>+ a. g x \<partial>return M x)" using assms
     by (intro nn_integral_cong_AE) (auto simp: AE_return)
   also have "... = g x"
-    using nn_integral_const[OF \<open>g x \<ge> 0\<close>, of "return M x"] emeasure_space_1 by simp
+    using nn_integral_const[of "return M x"] emeasure_space_1 by simp
   finally show ?thesis .
 qed
 
@@ -640,14 +660,14 @@
 
   have *: "\<And>x. fst x \<in> space M \<Longrightarrow> snd x \<in> A (fst x) \<longleftrightarrow> x \<in> (SIGMA x:space M. A x)"
     by (auto simp: fun_eq_iff)
-  have "(\<lambda>(x, y). indicator (A x) y::ereal) \<in> borel_measurable (M \<Otimes>\<^sub>M N)"
+  have "(\<lambda>(x, y). indicator (A x) y::ennreal) \<in> borel_measurable (M \<Otimes>\<^sub>M N)"
     apply measurable
     apply (subst measurable_cong)
     apply (rule *)
     apply (auto simp: space_pair_measure)
     done
   then have "(\<lambda>x. integral\<^sup>N (L x) (indicator (A x))) \<in> borel_measurable M"
-    by (intro nn_integral_measurable_subprob_algebra2[where N=N] ereal_indicator_nonneg L)
+    by (intro nn_integral_measurable_subprob_algebra2[where N=N] L)
   then show "(\<lambda>x. emeasure (L x) (A x)) \<in> borel_measurable M"
     apply (rule measurable_cong[THEN iffD1, rotated])
     apply (rule nn_integral_indicator)
@@ -660,7 +680,7 @@
   assumes L[measurable]: "L \<in> measurable M (subprob_algebra N)"
   shows "(\<lambda>x. measure (L x) (A x)) \<in> borel_measurable M"
   unfolding measure_def
-  by (intro borel_measurable_real_of_ereal emeasure_measurable_subprob_algebra2[OF assms])
+  by (intro borel_measurable_enn2real emeasure_measurable_subprob_algebra2[OF assms])
 
 definition "select_sets M = (SOME N. sets M = sets (subprob_algebra N))"
 
@@ -723,7 +743,7 @@
     qed
     finally show "(\<Sum>i. \<integral>\<^sup>+M'. emeasure M' (A i) \<partial>M) = (\<integral>\<^sup>+M'. emeasure M' (\<Union>i. A i) \<partial>M)" .
   qed
-qed (auto simp: A sets.space_closed positive_def nn_integral_nonneg)
+qed (auto simp: A sets.space_closed positive_def)
 
 lemma measurable_join:
   "join \<in> measurable (subprob_algebra (subprob_algebra N)) (subprob_algebra N)"
@@ -745,7 +765,7 @@
   fix M assume M: "M \<in> space (subprob_algebra (subprob_algebra N))"
   then have "(\<integral>\<^sup>+M'. emeasure M' (space N) \<partial>M) \<le> (\<integral>\<^sup>+M'. 1 \<partial>M)"
     apply (intro nn_integral_mono)
-    apply (auto simp: space_subprob_algebra 
+    apply (auto simp: space_subprob_algebra
                  dest!: sets_eq_imp_space_eq subprob_space.emeasure_space_le_1)
     done
   with M show "subprob_space (join M)"
@@ -756,8 +776,8 @@
   thus ?thesis by (simp add: measurable_empty_iff space_subprob_algebra_empty_iff)
 qed (auto simp: space_subprob_algebra)
 
-lemma nn_integral_join':
-  assumes f: "f \<in> borel_measurable N" "\<And>x. 0 \<le> f x"
+lemma nn_integral_join:
+  assumes f: "f \<in> borel_measurable N"
     and M[measurable_cong]: "sets M = sets (subprob_algebra N)"
   shows "(\<integral>\<^sup>+x. f x \<partial>join M) = (\<integral>\<^sup>+M'. \<integral>\<^sup>+x. f x \<partial>M' \<partial>M)"
   using f
@@ -772,7 +792,7 @@
     by simp
 next
   case (set A)
-  moreover with M have "(\<integral>\<^sup>+ M'. integral\<^sup>N M' (indicator A) \<partial>M) = (\<integral>\<^sup>+ M'. emeasure M' A \<partial>M)" 
+  moreover with M have "(\<integral>\<^sup>+ M'. integral\<^sup>N M' (indicator A) \<partial>M) = (\<integral>\<^sup>+ M'. emeasure M' A \<partial>M)"
     by (intro nn_integral_cong nn_integral_indicator)
        (auto simp: space_subprob_algebra dest!: sets_eq_imp_space_eq)
   ultimately show ?case
@@ -783,7 +803,7 @@
     using mult M M[THEN sets_eq_imp_space_eq]
     by (intro nn_integral_cong nn_integral_cmult) (auto simp add: space_subprob_algebra)
   also have "\<dots> = c * (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M)"
-    using nn_integral_measurable_subprob_algebra[OF mult(3)]
+    using nn_integral_measurable_subprob_algebra[OF mult(2)]
     by (intro nn_integral_cmult mult) (simp add: M)
   also have "\<dots> = c * (integral\<^sup>N (join M) f)"
     by (simp add: mult)
@@ -797,8 +817,8 @@
     by (intro nn_integral_cong nn_integral_add) (auto simp add: space_subprob_algebra)
   also have "\<dots> = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M) + (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. g x \<partial>M' \<partial>M)"
     using nn_integral_measurable_subprob_algebra[OF add(1)]
-    using nn_integral_measurable_subprob_algebra[OF add(5)]
-    by (intro nn_integral_add add) (simp_all add: M nn_integral_nonneg)
+    using nn_integral_measurable_subprob_algebra[OF add(4)]
+    by (intro nn_integral_add add) (simp_all add: M)
   also have "\<dots> = (integral\<^sup>N (join M) f) + (integral\<^sup>N (join M) g)"
     by (simp add: add)
   also have "\<dots> = (\<integral>\<^sup>+ x. f x + g x \<partial>join M)"
@@ -813,7 +833,7 @@
   also have "\<dots> = (SUP i. \<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. F i x \<partial>M' \<partial>M)"
     using nn_integral_measurable_subprob_algebra[OF seq(1)] seq
     by (intro nn_integral_monotone_convergence_SUP)
-       (simp_all add: M nn_integral_nonneg incseq_nn_integral incseq_def le_fun_def nn_integral_mono )
+       (simp_all add: M incseq_nn_integral incseq_def le_fun_def nn_integral_mono )
   also have "\<dots> = (SUP i. integral\<^sup>N (join M) (F i))"
     by (simp add: seq)
   also have "\<dots> = (\<integral>\<^sup>+ x. (SUP i. F i x) \<partial>join M)"
@@ -822,31 +842,23 @@
   finally show ?case by (simp add: ac_simps)
 qed
 
-lemma nn_integral_join:
-  assumes f[measurable]: "f \<in> borel_measurable N" "sets M = sets (subprob_algebra N)"
-  shows "(\<integral>\<^sup>+x. f x \<partial>join M) = (\<integral>\<^sup>+M'. \<integral>\<^sup>+x. f x \<partial>M' \<partial>M)"
-  apply (subst (1 3) nn_integral_max_0[symmetric])
-  apply (rule nn_integral_join')
-  apply (auto simp: f)
-  done
-
 lemma measurable_join1:
   "\<lbrakk> f \<in> measurable N K; sets M = sets (subprob_algebra N) \<rbrakk>
   \<Longrightarrow> f \<in> measurable (join M) K"
 by(simp add: measurable_def)
 
-lemma 
+lemma
   fixes f :: "_ \<Rightarrow> real"
   assumes f_measurable [measurable]: "f \<in> borel_measurable N"
-  and f_bounded: "\<And>x. x \<in> space N \<Longrightarrow> \<bar>f x\<bar> \<le> B" 
+  and f_bounded: "\<And>x. x \<in> space N \<Longrightarrow> \<bar>f x\<bar> \<le> B"
   and M [measurable_cong]: "sets M = sets (subprob_algebra N)"
   and fin: "finite_measure M"
-  and M_bounded: "AE M' in M. emeasure M' (space M') \<le> ereal B'"
+  and M_bounded: "AE M' in M. emeasure M' (space M') \<le> ennreal B'"
   shows integrable_join: "integrable (join M) f" (is ?integrable)
   and integral_join: "integral\<^sup>L (join M) f = \<integral> M'. integral\<^sup>L M' f \<partial>M" (is ?integral)
 proof(case_tac [!] "space N = {}")
   assume *: "space N = {}"
-  show ?integrable 
+  show ?integrable
     using M * by(simp add: real_integrable_def measurable_def nn_integral_empty)
   have "(\<integral> M'. integral\<^sup>L M' f \<partial>M) = (\<integral> M'. 0 \<partial>M)"
   proof(rule integral_cong)
@@ -869,120 +881,106 @@
   { fix f M'
     assume [measurable]: "f \<in> borel_measurable N"
       and f_bounded: "\<And>x. x \<in> space N \<Longrightarrow> f x \<le> B"
-      and "M' \<in> space M" "emeasure M' (space M') \<le> ereal B'"
-    have "AE x in M'. ereal (f x) \<le> ereal B"
+      and "M' \<in> space M" "emeasure M' (space M') \<le> ennreal B'"
+    have "AE x in M'. ennreal (f x) \<le> ennreal B"
     proof(rule AE_I2)
       fix x
       assume "x \<in> space M'"
       with \<open>M' \<in> space M\<close> sets_eq_imp_space_eq[OF M]
       have "x \<in> space N" by(auto simp add: space_subprob_algebra dest: sets_eq_imp_space_eq)
-      from f_bounded[OF this] show "ereal (f x) \<le> ereal B" by simp
+      from f_bounded[OF this] show "ennreal (f x) \<le> ennreal B" by simp
     qed
-    then have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M') \<le> (\<integral>\<^sup>+ x. ereal B \<partial>M')"
+    then have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M') \<le> (\<integral>\<^sup>+ x. ennreal B \<partial>M')"
       by(rule nn_integral_mono_AE)
-    also have "\<dots> = ereal B * emeasure M' (space M')" by(simp)
-    also have "\<dots> \<le> ereal B * ereal B'" by(rule ereal_mult_left_mono)(fact, simp)
-    also have "\<dots> \<le> ereal B * ereal \<bar>B'\<bar>" by(rule ereal_mult_left_mono)(simp_all)
-    finally have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M') \<le> ereal (B * \<bar>B'\<bar>)" by simp }
+    also have "\<dots> = ennreal B * emeasure M' (space M')" by(simp)
+    also have "\<dots> \<le> ennreal B * ennreal B'" by(rule mult_left_mono)(fact, simp)
+    also have "\<dots> \<le> ennreal B * ennreal \<bar>B'\<bar>" by(rule mult_left_mono)(simp_all)
+    finally have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M') \<le> ennreal (B * \<bar>B'\<bar>)" by (simp add: ennreal_mult) }
   note bounded1 = this
 
   have bounded:
     "\<And>f. \<lbrakk> f \<in> borel_measurable N; \<And>x. x \<in> space N \<Longrightarrow> f x \<le> B \<rbrakk>
-    \<Longrightarrow> (\<integral>\<^sup>+ x. ereal (f x) \<partial>join M) \<noteq> \<infinity>"
+    \<Longrightarrow> (\<integral>\<^sup>+ x. ennreal (f x) \<partial>join M) \<noteq> top"
   proof -
     fix f
     assume [measurable]: "f \<in> borel_measurable N"
       and f_bounded: "\<And>x. x \<in> space N \<Longrightarrow> f x \<le> B"
-    have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>join M) = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. ereal (f x) \<partial>M' \<partial>M)"
+    have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>join M) = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. ennreal (f x) \<partial>M' \<partial>M)"
       by(rule nn_integral_join[OF _ M]) simp
     also have "\<dots> \<le> \<integral>\<^sup>+ M'. B * \<bar>B'\<bar> \<partial>M"
       using bounded1[OF \<open>f \<in> borel_measurable N\<close> f_bounded]
       by(rule nn_integral_mono_AE[OF AE_mp[OF M_bounded AE_I2], rule_format])
     also have "\<dots> = B * \<bar>B'\<bar> * emeasure M (space M)" by simp
-    also have "\<dots> < \<infinity>" by(simp add: finite_measure.finite_emeasure_space[OF fin])
+    also have "\<dots> < \<infinity>"
+      using finite_measure.finite_emeasure_space[OF fin]
+      by(simp add: ennreal_mult_less_top less_top)
     finally show "?thesis f" by simp
   qed
-  have f_pos: "(\<integral>\<^sup>+ x. ereal (f x) \<partial>join M) \<noteq> \<infinity>"
-    and f_neg: "(\<integral>\<^sup>+ x. ereal (- f x) \<partial>join M) \<noteq> \<infinity>"
+  have f_pos: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>join M) \<noteq> \<infinity>"
+    and f_neg: "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>join M) \<noteq> \<infinity>"
     using f_bounded by(auto del: notI intro!: bounded simp add: abs_le_iff)
-  
+
   show ?integrable using f_pos f_neg by(simp add: real_integrable_def)
 
   note [measurable] = nn_integral_measurable_subprob_algebra
 
-  have "(\<integral>\<^sup>+ x. f x \<partial>join M) = (\<integral>\<^sup>+ x. max 0 (f x) \<partial>join M)"
-    by(subst nn_integral_max_0[symmetric])(simp add: zero_ereal_def)
-  also have "\<dots> = \<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. max 0 (f x) \<partial>M' \<partial>M"
+  have int_f: "(\<integral>\<^sup>+ x. f x \<partial>join M) = \<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M"
     by(simp add: nn_integral_join[OF _ M])
-  also have "\<dots> = \<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M"
-    by(subst nn_integral_max_0[symmetric])(simp add: zero_ereal_def)
-  finally have int_f: "(\<integral>\<^sup>+ x. f x \<partial>join M) = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M)" .
-
-  have "(\<integral>\<^sup>+ x. - f x \<partial>join M) = (\<integral>\<^sup>+ x. max 0 (- f x) \<partial>join M)"
-    by(subst nn_integral_max_0[symmetric])(simp add: zero_ereal_def)
-  also have "\<dots> = \<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. max 0 (- f x) \<partial>M' \<partial>M"
+  have int_mf: "(\<integral>\<^sup>+ x. - f x \<partial>join M) = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. - f x \<partial>M' \<partial>M)"
     by(simp add: nn_integral_join[OF _ M])
-  also have "\<dots> = \<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. - f x \<partial>M' \<partial>M"
-    by(subst nn_integral_max_0[symmetric])(simp add: zero_ereal_def)
-  finally have int_mf: "(\<integral>\<^sup>+ x. - f x \<partial>join M) = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. - f x \<partial>M' \<partial>M)" .
 
-  have f_pos1:
-    "\<And>M'. \<lbrakk> M' \<in> space M; emeasure M' (space M') \<le> ereal B' \<rbrakk>
-    \<Longrightarrow> (\<integral>\<^sup>+ x. ereal (f x) \<partial>M') \<le> ereal (B * \<bar>B'\<bar>)"
-    using f_measurable by(auto intro!: bounded1 dest: f_bounded)
-  have "AE M' in M. (\<integral>\<^sup>+ x. f x \<partial>M') \<noteq> \<infinity>"
-    using M_bounded by(rule AE_mp[OF _ AE_I2])(auto dest: f_pos1)
-  hence [simp]: "(\<integral>\<^sup>+ M'. ereal (real_of_ereal (\<integral>\<^sup>+ x. f x \<partial>M')) \<partial>M) = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M)"
-    by(rule nn_integral_cong_AE[OF AE_mp])(simp add: ereal_real nn_integral_nonneg)
-  from f_pos have [simp]: "integrable M (\<lambda>M'. real_of_ereal (\<integral>\<^sup>+ x. f x \<partial>M'))"
-    by(simp add: int_f real_integrable_def nn_integral_nonneg real_of_ereal[symmetric] nn_integral_0_iff_AE[THEN iffD2] del: real_of_ereal)
+  have pos_finite: "AE M' in M. (\<integral>\<^sup>+ x. f x \<partial>M') \<noteq> \<infinity>"
+    using AE_space M_bounded
+  proof eventually_elim
+    fix M' assume "M' \<in> space M" "emeasure M' (space M') \<le> ennreal B'"
+    then have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M') \<le> ennreal (B * \<bar>B'\<bar>)"
+      using f_measurable by(auto intro!: bounded1 dest: f_bounded)
+    then show "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M') \<noteq> \<infinity>"
+      by (auto simp: top_unique)
+  qed
+  hence [simp]: "(\<integral>\<^sup>+ M'. ennreal (enn2real (\<integral>\<^sup>+ x. f x \<partial>M')) \<partial>M) = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M)"
+    by (rule nn_integral_cong_AE[OF AE_mp]) (simp add: less_top)
+  from f_pos have [simp]: "integrable M (\<lambda>M'. enn2real (\<integral>\<^sup>+ x. f x \<partial>M'))"
+    by(simp add: int_f real_integrable_def nn_integral_0_iff_AE[THEN iffD2] ennreal_neg enn2real_nonneg)
 
-  have f_neg1:
-    "\<And>M'. \<lbrakk> M' \<in> space M; emeasure M' (space M') \<le> ereal B' \<rbrakk>
-    \<Longrightarrow> (\<integral>\<^sup>+ x. ereal (- f x) \<partial>M') \<le> ereal (B * \<bar>B'\<bar>)"
-    using f_measurable by(auto intro!: bounded1 dest: f_bounded)
-  have "AE M' in M. (\<integral>\<^sup>+ x. - f x \<partial>M') \<noteq> \<infinity>"
-    using M_bounded by(rule AE_mp[OF _ AE_I2])(auto dest: f_neg1)
-  hence [simp]: "(\<integral>\<^sup>+ M'. ereal (real_of_ereal (\<integral>\<^sup>+ x. - f x \<partial>M')) \<partial>M) = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. - f x \<partial>M' \<partial>M)"
-    by(rule nn_integral_cong_AE[OF AE_mp])(simp add: ereal_real nn_integral_nonneg)
-  from f_neg have [simp]: "integrable M (\<lambda>M'. real_of_ereal (\<integral>\<^sup>+ x. - f x \<partial>M'))"
-    by(simp add: int_mf real_integrable_def nn_integral_nonneg real_of_ereal[symmetric] nn_integral_0_iff_AE[THEN iffD2] del: real_of_ereal)
+  have neg_finite: "AE M' in M. (\<integral>\<^sup>+ x. - f x \<partial>M') \<noteq> \<infinity>"
+    using AE_space M_bounded
+  proof eventually_elim
+    fix M' assume "M' \<in> space M" "emeasure M' (space M') \<le> ennreal B'"
+    then have "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M') \<le> ennreal (B * \<bar>B'\<bar>)"
+      using f_measurable by(auto intro!: bounded1 dest: f_bounded)
+    then show "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M') \<noteq> \<infinity>"
+      by (auto simp: top_unique)
+  qed
+  hence [simp]: "(\<integral>\<^sup>+ M'. ennreal (enn2real (\<integral>\<^sup>+ x. - f x \<partial>M')) \<partial>M) = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. - f x \<partial>M' \<partial>M)"
+    by (rule nn_integral_cong_AE[OF AE_mp]) (simp add: less_top)
+  from f_neg have [simp]: "integrable M (\<lambda>M'. enn2real (\<integral>\<^sup>+ x. - f x \<partial>M'))"
+    by(simp add: int_mf real_integrable_def nn_integral_0_iff_AE[THEN iffD2] ennreal_neg enn2real_nonneg)
 
-  from \<open>?integrable\<close>
-  have "ereal (\<integral> x. f x \<partial>join M) = (\<integral>\<^sup>+ x. f x \<partial>join M) - (\<integral>\<^sup>+ x. - f x \<partial>join M)"
-    by(simp add: real_lebesgue_integral_def ereal_minus(1)[symmetric] ereal_real nn_integral_nonneg f_pos f_neg del: ereal_minus(1))
-  also note int_f
-  also note int_mf
-  also have "(\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M) - (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. - f x \<partial>M' \<partial>M) = 
-    ((\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M) - (\<integral>\<^sup>+ M'. - \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M)) - 
-    ((\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. - f x \<partial>M' \<partial>M) - (\<integral>\<^sup>+ M'. - \<integral>\<^sup>+ x. - f x \<partial>M' \<partial>M))"
-    by(subst (7 11) nn_integral_0_iff_AE[THEN iffD2])(simp_all add: nn_integral_nonneg)
-  also have "(\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M) - (\<integral>\<^sup>+ M'. - \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M) = \<integral> M'. real_of_ereal (\<integral>\<^sup>+ x. f x \<partial>M') \<partial>M"
-    using f_pos
-    by(simp add: real_lebesgue_integral_def)(simp add: ereal_minus(1)[symmetric] ereal_real int_f nn_integral_nonneg nn_integral_0_iff_AE[THEN iffD2] real_of_ereal_pos zero_ereal_def[symmetric])
-  also have "(\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. - f x \<partial>M' \<partial>M) - (\<integral>\<^sup>+ M'. - \<integral>\<^sup>+ x. - f x \<partial>M' \<partial>M) = \<integral> M'. real_of_ereal (\<integral>\<^sup>+ x. - f x \<partial>M') \<partial>M"
-    using f_neg
-    by(simp add: real_lebesgue_integral_def)(simp add: ereal_minus(1)[symmetric] ereal_real int_mf nn_integral_nonneg nn_integral_0_iff_AE[THEN iffD2] real_of_ereal_pos zero_ereal_def[symmetric])
-  also note ereal_minus(1)
-  also have "(\<integral> M'. real_of_ereal (\<integral>\<^sup>+ x. f x \<partial>M') \<partial>M) - (\<integral> M'. real_of_ereal (\<integral>\<^sup>+ x. - f x \<partial>M') \<partial>M) = 
-    \<integral>M'. real_of_ereal (\<integral>\<^sup>+ x. f x \<partial>M') - real_of_ereal (\<integral>\<^sup>+ x. - f x \<partial>M') \<partial>M"
+  have "(\<integral> x. f x \<partial>join M) = enn2real (\<integral>\<^sup>+ N. \<integral>\<^sup>+x. f x \<partial>N \<partial>M) - enn2real (\<integral>\<^sup>+ N. \<integral>\<^sup>+x. - f x \<partial>N \<partial>M)"
+    unfolding real_lebesgue_integral_def[OF \<open>?integrable\<close>] by (simp add: nn_integral_join[OF _ M])
+  also have "\<dots> = (\<integral>N. enn2real (\<integral>\<^sup>+x. f x \<partial>N) \<partial>M) - (\<integral>N. enn2real (\<integral>\<^sup>+x. - f x \<partial>N) \<partial>M)"
+    using pos_finite neg_finite by (subst (1 2) integral_eq_nn_integral) (auto simp: enn2real_nonneg)
+  also have "\<dots> = (\<integral>N. enn2real (\<integral>\<^sup>+x. f x \<partial>N) - enn2real (\<integral>\<^sup>+x. - f x \<partial>N) \<partial>M)"
     by simp
-  also have "\<dots> = \<integral>M'. \<integral> x. f x \<partial>M' \<partial>M" using _ _ M_bounded
-  proof(rule integral_cong_AE[OF _ _ AE_mp[OF _ AE_I2], rule_format])
-    show "(\<lambda>M'. integral\<^sup>L M' f) \<in> borel_measurable M"
-      by measurable(simp add: integral_measurable_subprob_algebra[OF _ f_bounded])
-      
-    fix M'
-    assume "M' \<in> space M" "emeasure M' (space M') \<le> B'"
-    then interpret finite_measure M' by(auto intro: finite_measureI)
-    
-    from \<open>M' \<in> space M\<close> sets_eq_imp_space_eq[OF M]
-    have [measurable_cong]: "sets M' = sets N" by(simp add: space_subprob_algebra)
-    hence [simp]: "space M' = space N" by(rule sets_eq_imp_space_eq)
-    have "integrable M' f"
-      by(rule integrable_const_bound[where B=B])(auto simp add: f_bounded)
-    then show "real_of_ereal (\<integral>\<^sup>+ x. f x \<partial>M') - real_of_ereal (\<integral>\<^sup>+ x. - f x \<partial>M') = \<integral> x. f x \<partial>M'"
-      by(simp add: real_lebesgue_integral_def)
+  also have "\<dots> = \<integral>M'. \<integral> x. f x \<partial>M' \<partial>M"
+  proof (rule integral_cong_AE)
+    show "AE x in M.
+        enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>x) - enn2real (\<integral>\<^sup>+ x. ennreal (- f x) \<partial>x) = integral\<^sup>L x f"
+      using AE_space M_bounded
+    proof eventually_elim
+      fix M' assume "M' \<in> space M" "emeasure M' (space M') \<le> B'"
+      then interpret subprob_space M'
+        by (auto simp: M[THEN sets_eq_imp_space_eq] space_subprob_algebra)
+
+      from \<open>M' \<in> space M\<close> sets_eq_imp_space_eq[OF M]
+      have [measurable_cong]: "sets M' = sets N" by(simp add: space_subprob_algebra)
+      hence [simp]: "space M' = space N" by(rule sets_eq_imp_space_eq)
+      have "integrable M' f"
+        by(rule integrable_const_bound[where B=B])(auto simp add: f_bounded)
+      then show "enn2real (\<integral>\<^sup>+ x. f x \<partial>M') - enn2real (\<integral>\<^sup>+ x. - f x \<partial>M') = \<integral> x. f x \<partial>M'"
+        by(simp add: real_lebesgue_integral_def)
+    qed
   qed simp_all
   finally show ?integral by simp
 qed
@@ -995,16 +993,16 @@
   then have A: "A \<in> sets N" by simp
   show "emeasure (join (distr M (subprob_algebra N) join)) A = emeasure (join (join M)) A"
     using measurable_join[of N]
-    by (auto simp: M A nn_integral_distr emeasure_join measurable_emeasure_subprob_algebra emeasure_nonneg
+    by (auto simp: M A nn_integral_distr emeasure_join measurable_emeasure_subprob_algebra
                    sets_eq_imp_space_eq[OF M] space_subprob_algebra nn_integral_join[OF _ M]
              intro!: nn_integral_cong emeasure_join)
 qed (simp add: M)
 
-lemma join_return: 
+lemma join_return:
   assumes "sets M = sets N" and "subprob_space M"
   shows "join (return (subprob_algebra N) M) = M"
   by (rule measure_eqI)
-     (simp_all add: emeasure_join emeasure_nonneg space_subprob_algebra  
+     (simp_all add: emeasure_join space_subprob_algebra
                     measurable_emeasure_subprob_algebra nn_integral_return assms)
 
 lemma join_return':
@@ -1025,13 +1023,13 @@
   fix A assume "A \<in> sets ?r"
   hence A_in_N: "A \<in> sets N" by simp
 
-  from assms have "f \<in> measurable (join M) N" 
+  from assms have "f \<in> measurable (join M) N"
       by (simp cong: measurable_cong_sets)
-  moreover from assms and A_in_N have "f-`A \<inter> space R \<in> sets R" 
+  moreover from assms and A_in_N have "f-`A \<inter> space R \<in> sets R"
       by (intro measurable_sets) simp_all
   ultimately have "emeasure (distr (join M) N f) A = \<integral>\<^sup>+M'. emeasure M' (f-`A \<inter> space R) \<partial>M"
       by (simp_all add: A_in_N emeasure_distr emeasure_join assms)
-  
+
   also have "... = \<integral>\<^sup>+ x. emeasure (distr x N f) A \<partial>M" using A_in_N
   proof (intro nn_integral_cong, subst emeasure_distr)
     fix M' assume "M' \<in> space M"
@@ -1045,7 +1043,7 @@
 
   also have "(\<lambda>M. distr M N f) \<in> measurable M (subprob_algebra N)"
       by (simp cong: measurable_cong_sets add: assms measurable_distr)
-  hence "(\<integral>\<^sup>+ x. emeasure (distr x N f) A \<partial>M) = 
+  hence "(\<integral>\<^sup>+ x. emeasure (distr x N f) A \<partial>M) =
              emeasure (join (distr M (subprob_algebra N) (\<lambda>M. distr M N f))) A"
       by (simp_all add: emeasure_join assms A_in_N nn_integral_distr measurable_emeasure_subprob_algebra)
   finally show "emeasure ?r A = emeasure ?l A" ..
@@ -1057,7 +1055,7 @@
 
 adhoc_overloading Monad_Syntax.bind bind
 
-lemma bind_empty: 
+lemma bind_empty:
   "space M = {} \<Longrightarrow> bind M f = count_space {}"
   by (simp add: bind_def)
 
@@ -1076,12 +1074,12 @@
   shows "sets (bind M f) = sets N"
   using f [of "SOME x. x \<in> space M"] by (simp add: bind_nonempty M some_in_eq)
 
-lemma space_bind[simp]: 
+lemma space_bind[simp]:
   assumes "\<And>x. x \<in> space M \<Longrightarrow> sets (f x) = sets N" and "space M \<noteq> {}"
   shows "space (bind M f) = space N"
   using assms by (intro sets_eq_imp_space_eq sets_bind)
 
-lemma bind_cong: 
+lemma bind_cong:
   assumes "\<forall>x \<in> space M. f x = g x"
   shows "bind M f = bind M g"
 proof (cases "space M = {}")
@@ -1117,7 +1115,8 @@
 proof cases
   assume M: "space M \<noteq> {}" show ?thesis
     unfolding bind_nonempty''[OF N M] nn_integral_join[OF f sets_distr]
-    by (rule nn_integral_distr[OF N nn_integral_measurable_subprob_algebra[OF f]])
+    by (rule nn_integral_distr[OF N])
+       (simp add: f nn_integral_measurable_subprob_algebra)
 qed (simp add: bind_empty space_empty[of M] nn_integral_count_space)
 
 lemma AE_bind:
@@ -1141,7 +1140,7 @@
     apply (rule measurable_compose[OF N nn_integral_measurable_subprob_algebra])
     apply measurable
     apply (intro eventually_subst AE_I2)
-    apply (auto simp add: emeasure_le_0_iff subprob_measurableD(1)[OF N]
+    apply (auto simp add: subprob_measurableD(1)[OF N]
                 intro!: AE_iff_measurable[symmetric])
     done
   finally show ?thesis .
@@ -1153,12 +1152,12 @@
   shows "(\<lambda>x. bind (f x) (g x)) \<in> measurable M (subprob_algebra R)"
 proof (subst measurable_cong)
   fix x assume x_in_M: "x \<in> space M"
-  with assms have "space (f x) \<noteq> {}" 
+  with assms have "space (f x) \<noteq> {}"
       by (blast dest: subprob_space_kernel subprob_space.subprob_not_empty)
   moreover from M2 x_in_M have "g x \<in> measurable (f x) (subprob_algebra R)"
       by (subst measurable_cong_sets[OF sets_kernel[OF M1 x_in_M] refl])
          (auto dest: measurable_Pair2)
-  ultimately show "bind (f x) (g x) = join (distr (f x) (subprob_algebra R) (g x))" 
+  ultimately show "bind (f x) (g x) = join (distr (f x) (subprob_algebra R) (g x))"
       by (simp_all add: bind_nonempty'')
 next
   show "(\<lambda>w. join (distr (f w) (subprob_algebra R) (g w))) \<in> measurable M (subprob_algebra R)"
@@ -1183,19 +1182,19 @@
   shows "subprob_space (M \<bind> f)"
 proof (rule subprob_space_kernel[of "\<lambda>x. x \<bind> f"])
   show "(\<lambda>x. x \<bind> f) \<in> measurable (subprob_algebra M) (subprob_algebra N)"
-    by (rule measurable_bind, rule measurable_ident_sets, rule refl, 
+    by (rule measurable_bind, rule measurable_ident_sets, rule refl,
         rule measurable_compose[OF measurable_snd assms(2)])
   from assms(1) show "M \<in> space (subprob_algebra M)"
     by (simp add: space_subprob_algebra)
 qed
 
-lemma 
+lemma
   fixes f :: "_ \<Rightarrow> real"
   assumes f_measurable [measurable]: "f \<in> borel_measurable K"
-  and f_bounded: "\<And>x. x \<in> space K \<Longrightarrow> \<bar>f x\<bar> \<le> B" 
+  and f_bounded: "\<And>x. x \<in> space K \<Longrightarrow> \<bar>f x\<bar> \<le> B"
   and N [measurable]: "N \<in> measurable M (subprob_algebra K)"
   and fin: "finite_measure M"
-  and M_bounded: "AE x in M. emeasure (N x) (space (N x)) \<le> ereal B'"
+  and M_bounded: "AE x in M. emeasure (N x) (space (N x)) \<le> ennreal B'"
   shows integrable_bind: "integrable (bind M N) f" (is ?integrable)
   and integral_bind: "integral\<^sup>L (bind M N) f = \<integral> x. integral\<^sup>L (N x) f \<partial>M" (is ?integral)
 proof(case_tac [!] "space M = {}")
@@ -1211,11 +1210,11 @@
     using f_measurable f_bounded
     by(rule integral_join[where B'=B'])(simp_all add: finite_measure_distr AE_distr_iff M_bounded)
   also have "\<dots> = \<integral> x. integral\<^sup>L (N x) f \<partial>M"
-    by(rule integral_distr)(simp_all add: integral_measurable_subprob_algebra[OF _ f_bounded])
+    by(rule integral_distr)(simp_all add: integral_measurable_subprob_algebra[OF _])
   finally show ?integral by(simp add: bind_nonempty''[where N=K])
 qed(simp_all add: bind_def integrable_count_space lebesgue_integral_count_space_finite integral_empty)
 
-lemma (in prob_space) prob_space_bind: 
+lemma (in prob_space) prob_space_bind:
   assumes ae: "AE x in M. prob_space (N x)"
     and N[measurable]: "N \<in> measurable M (subprob_algebra S)"
   shows "prob_space (M \<bind> N)"
@@ -1244,27 +1243,27 @@
   { fix x assume "x \<in> space M"
     from f[THEN measurable_space, OF this] interpret subprob_space "f x"
       by (simp add: space_subprob_algebra)
-    have "emeasure (f x) X = ereal (measure (f x) X)" "measure (f x) X \<le> 1"
+    have "emeasure (f x) X = ennreal (measure (f x) X)" "measure (f x) X \<le> 1"
       by (auto simp: emeasure_eq_measure subprob_measure_le_1) }
   note this[simp]
 
   have "emeasure (M \<bind> f) X = \<integral>\<^sup>+x. emeasure (f x) X \<partial>M"
     using subprob_not_empty f X by (rule emeasure_bind)
-  also have "\<dots> = \<integral>\<^sup>+x. ereal (measure (f x) X) \<partial>M"
+  also have "\<dots> = \<integral>\<^sup>+x. ennreal (measure (f x) X) \<partial>M"
     by (intro nn_integral_cong) simp
   also have "\<dots> = \<integral>x. measure (f x) X \<partial>M"
     by (intro nn_integral_eq_integral integrable_const_bound[where B=1]
               measure_measurable_subprob_algebra2[OF _ f] pair_measureI X)
        (auto simp: measure_nonneg)
   finally show ?thesis
-    by (simp add: Mf.emeasure_eq_measure)
+    by (simp add: Mf.emeasure_eq_measure measure_nonneg integral_nonneg)
 qed
 
-lemma emeasure_bind_const: 
-    "space M \<noteq> {} \<Longrightarrow> X \<in> sets N \<Longrightarrow> subprob_space N \<Longrightarrow> 
+lemma emeasure_bind_const:
+    "space M \<noteq> {} \<Longrightarrow> X \<in> sets N \<Longrightarrow> subprob_space N \<Longrightarrow>
          emeasure (M \<bind> (\<lambda>x. N)) X = emeasure N X * emeasure M (space M)"
-  by (simp add: bind_nonempty emeasure_join nn_integral_distr 
-                space_subprob_algebra measurable_emeasure_subprob_algebra emeasure_nonneg)
+  by (simp add: bind_nonempty emeasure_join nn_integral_distr
+                space_subprob_algebra measurable_emeasure_subprob_algebra)
 
 lemma emeasure_bind_const':
   assumes "subprob_space M" "subprob_space N"
@@ -1273,7 +1272,7 @@
 proof (case_tac "X \<in> sets N")
   fix X assume "X \<in> sets N"
   thus "emeasure (M \<bind> (\<lambda>x. N)) X = emeasure N X * emeasure M (space M)" using assms
-      by (subst emeasure_bind_const) 
+      by (subst emeasure_bind_const)
          (simp_all add: subprob_space.subprob_not_empty subprob_space.emeasure_space_le_1)
 next
   fix X assume "X \<notin> sets N"
@@ -1285,10 +1284,10 @@
 lemma emeasure_bind_const_prob_space:
   assumes "prob_space M" "subprob_space N"
   shows "emeasure (M \<bind> (\<lambda>x. N)) X = emeasure N X"
-  using assms by (simp add: emeasure_bind_const' prob_space_imp_subprob_space 
+  using assms by (simp add: emeasure_bind_const' prob_space_imp_subprob_space
                             prob_space.emeasure_space_1)
 
-lemma bind_return: 
+lemma bind_return:
   assumes "f \<in> measurable M (subprob_algebra N)" and "x \<in> space M"
   shows "bind (return M x) f = f x"
   using sets_kernel[OF assms] assms
@@ -1298,7 +1297,7 @@
 lemma bind_return':
   shows "bind M (return M) = M"
   by (cases "space M = {}")
-     (simp_all add: bind_empty space_empty[symmetric] bind_nonempty join_return' 
+     (simp_all add: bind_empty space_empty[symmetric] bind_nonempty join_return'
                cong: subprob_algebra_cong)
 
 lemma distr_bind:
@@ -1372,7 +1371,7 @@
   hence "?x \<in> space (restrict_space M A)" by(simp add: space_restrict_space)
   with f have "f ?x \<in> space (subprob_algebra N)" by(rule measurable_space)
   hence sps: "subprob_space (f ?x)"
-    and sets: "sets (f ?x) = sets N" 
+    and sets: "sets (f ?x) = sets N"
     by(simp_all add: space_subprob_algebra)
   have "space (f ?x) \<noteq> {}" using sps by(rule subprob_space.subprob_not_empty)
   moreover have "sets ?c = sets N" by(simp add: null_measure_def)(simp add: sets)
@@ -1392,10 +1391,10 @@
 qed
 
 lemma bind_const': "\<lbrakk>prob_space M; subprob_space N\<rbrakk> \<Longrightarrow> M \<bind> (\<lambda>x. N) = N"
-  by (intro measure_eqI) 
+  by (intro measure_eqI)
      (simp_all add: space_subprob_algebra prob_space.not_empty emeasure_bind_const_prob_space)
 
-lemma bind_return_distr: 
+lemma bind_return_distr:
     "space M \<noteq> {} \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> bind M (return N \<circ> f) = distr M N f"
   apply (simp add: bind_nonempty)
   apply (subst subprob_algebra_cong)
@@ -1423,16 +1422,16 @@
                          sets_kernel[OF M2 someI_ex[OF ex_in[OF \<open>space N \<noteq> {}\<close>]]]
   note space_some[simp] = sets_eq_imp_space_eq[OF this(1)] sets_eq_imp_space_eq[OF this(2)]
 
-  have "bind M (\<lambda>x. bind (f x) g) = 
+  have "bind M (\<lambda>x. bind (f x) g) =
         join (distr M (subprob_algebra R) (join \<circ> (\<lambda>x. (distr x (subprob_algebra R) g)) \<circ> f))"
     by (simp add: sets_eq_imp_space_eq[OF sets_fx] bind_nonempty o_def
              cong: subprob_algebra_cong distr_cong)
   also have "distr M (subprob_algebra R) (join \<circ> (\<lambda>x. (distr x (subprob_algebra R) g)) \<circ> f) =
              distr (distr (distr M (subprob_algebra N) f)
                           (subprob_algebra (subprob_algebra R))
-                          (\<lambda>x. distr x (subprob_algebra R) g)) 
+                          (\<lambda>x. distr x (subprob_algebra R) g))
                    (subprob_algebra R) join"
-      apply (subst distr_distr, 
+      apply (subst distr_distr,
              (blast intro: measurable_comp measurable_distr measurable_join M1 M2)+)+
       apply (simp add: o_assoc)
       done
@@ -1447,12 +1446,12 @@
   assumes Mh: "case_prod h \<in> measurable (M \<Otimes>\<^sub>M M') N"
   shows "do {x \<leftarrow> M; y \<leftarrow> f x; g (h x y)} = do {x \<leftarrow> M; y \<leftarrow> f x; return N (h x y)} \<bind> g"
 proof-
-  have "do {x \<leftarrow> M; y \<leftarrow> f x; return N (h x y)} \<bind> g = 
+  have "do {x \<leftarrow> M; y \<leftarrow> f x; return N (h x y)} \<bind> g =
             do {x \<leftarrow> M; do {y \<leftarrow> f x; return N (h x y)} \<bind> g}"
     using Mh by (auto intro!: bind_assoc measurable_bind'[OF Mf] Mf Mg
                       measurable_compose[OF _ return_measurable] simp: measurable_split_conv)
   also from Mh have "\<And>x. x \<in> space M \<Longrightarrow> h x \<in> measurable M' N" by measurable
-  hence "do {x \<leftarrow> M; do {y \<leftarrow> f x; return N (h x y)} \<bind> g} = 
+  hence "do {x \<leftarrow> M; do {y \<leftarrow> f x; return N (h x y)} \<bind> g} =
             do {x \<leftarrow> M; y \<leftarrow> f x; return N (h x y) \<bind> g}"
     apply (intro ballI bind_cong bind_assoc)
     apply (subst measurable_cong_sets[OF sets_kernel[OF Mf] refl], simp)
@@ -1460,7 +1459,7 @@
     done
   also have "\<And>x. x \<in> space M \<Longrightarrow> space (f x) = space M'"
     by (intro sets_eq_imp_space_eq sets_kernel[OF Mf])
-  with measurable_space[OF Mh] 
+  with measurable_space[OF Mh]
     have "do {x \<leftarrow> M; y \<leftarrow> f x; return N (h x y) \<bind> g} = do {x \<leftarrow> M; y \<leftarrow> f x; g (h x y)}"
     by (intro ballI bind_cong bind_return[OF Mg]) (auto simp: space_pair_measure)
   finally show ?thesis ..
@@ -1485,13 +1484,13 @@
 lemma (in pair_prob_space) bind_rotate:
   assumes C[measurable]: "(\<lambda>(x, y). C x y) \<in> measurable (M1 \<Otimes>\<^sub>M M2) (subprob_algebra N)"
   shows "(M1 \<bind> (\<lambda>x. M2 \<bind> (\<lambda>y. C x y))) = (M2 \<bind> (\<lambda>y. M1 \<bind> (\<lambda>x. C x y)))"
-proof - 
+proof -
   interpret swap: pair_prob_space M2 M1 by unfold_locales
   note measurable_bind[where N="M2", measurable]
   note measurable_bind[where N="M1", measurable]
   have [simp]: "M1 \<in> space (subprob_algebra M1)" "M2 \<in> space (subprob_algebra M2)"
     by (auto simp: space_subprob_algebra) unfold_locales
-  have "(M1 \<bind> (\<lambda>x. M2 \<bind> (\<lambda>y. C x y))) = 
+  have "(M1 \<bind> (\<lambda>x. M2 \<bind> (\<lambda>y. C x y))) =
     (M1 \<bind> (\<lambda>x. M2 \<bind> (\<lambda>y. return (M1 \<Otimes>\<^sub>M M2) (x, y)))) \<bind> (\<lambda>(x, y). C x y)"
     by (auto intro!: bind_cong simp: bind_return[where N=N] space_pair_measure bind_assoc[where N="M1 \<Otimes>\<^sub>M M2" and R=N])
   also have "\<dots> = (distr (M2 \<Otimes>\<^sub>M M1) (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x, y). (y, x))) \<bind> (\<lambda>(x, y). C x y)"
@@ -1539,7 +1538,8 @@
         using sets_SUP_measure[of M, OF const] by simp
       moreover assume "disjoint_family A"
       ultimately show "(\<Sum>i. SUP ia. emeasure (M ia) (A i)) = (SUP i. emeasure (M i) (\<Union>i. A i))"
-        using mono by (subst suminf_SUP_eq) (auto simp: mono_def le_fun_def intro!: SUP_cong suminf_emeasure)
+        using suminf_SUP_eq
+        using mono by (subst ennreal_suminf_SUP_eq) (auto simp: mono_def le_fun_def intro!: SUP_cong suminf_emeasure)
     qed
     show "positive (sets (SUP_measure M)) (\<lambda>A. SUP i. emeasure (M i) A)"
       by (auto simp: positive_def intro: SUP_upper2)
@@ -1563,8 +1563,8 @@
 
 lemma return_count_space_eq_density:
     "return (count_space M) x = density (count_space M) (indicator {x})"
-  by (rule measure_eqI) 
-     (auto simp: indicator_inter_arith_ereal emeasure_density split: split_indicator)
+  by (rule measure_eqI)
+     (auto simp: indicator_inter_arith[symmetric] emeasure_density split: split_indicator)
 
 lemma null_measure_in_space_subprob_algebra [simp]:
   "null_measure M \<in> space (subprob_algebra M) \<longleftrightarrow> space M \<noteq> {}"
--- a/src/HOL/Probability/Independent_Family.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Independent_Family.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -401,7 +401,7 @@
       using P[OF \<open>i \<in> I\<close>] by blast
     finally show "{{x \<in> space M. P i (X i x)}} \<subseteq> {X i -` A \<inter> space M |A. A \<in> sets (N i)}" .
   qed
-qed                              
+qed
 
 lemma (in prob_space) indep_sets_collect_sigma:
   fixes I :: "'j \<Rightarrow> 'i set" and J :: "'j set" and E :: "'i \<Rightarrow> 'a set set"
@@ -524,7 +524,7 @@
   show "indep_sets (\<lambda>i. sigma_sets (space M) (?proj i (sets (Pi\<^sub>M (K i) M')))) L"
   proof (rule indep_sets_mono_sets)
     fix j assume j: "j \<in> L"
-    have "sigma_sets (space M) (?proj j (sets (Pi\<^sub>M (K j) M'))) = 
+    have "sigma_sets (space M) (?proj j (sets (Pi\<^sub>M (K j) M'))) =
       sigma_sets (space M) (sigma_sets (space M) (?proj j (prod_algebra (K j) M')))"
       using j K X[THEN measurable_space] unfolding sets_PiM
       by (subst sigma_sets_vimage_commute) (auto simp add: Pi_iff)
@@ -541,7 +541,7 @@
         assume "K j \<noteq> {}" with J have "J \<noteq> {}"
           by auto
         { interpret sigma_algebra "space M" "?UN j"
-            by (rule sigma_algebra_sigma_sets) auto 
+            by (rule sigma_algebra_sigma_sets) auto
           have "\<And>A. (\<And>i. i \<in> J \<Longrightarrow> A i \<in> ?UN j) \<Longrightarrow> INTER J A \<in> ?UN j"
             using \<open>finite J\<close> \<open>J \<noteq> {}\<close> by (rule finite_INT) blast }
         note INT = this
@@ -996,7 +996,7 @@
   assumes I: "finite I" "i \<notin> I" and indep: "indep_vars (\<lambda>_. borel) X (insert i I)"
   shows "indep_var borel (X i) borel (\<lambda>\<omega>. \<Sum>i\<in>I. X i \<omega>)"
 proof -
-  have "indep_var 
+  have "indep_var
     borel ((\<lambda>f. f i) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) {i}))
     borel ((\<lambda>f. \<Sum>i\<in>I. f i) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) I))"
     using I by (intro indep_var_compose[OF indep_var_restrict[OF indep]] ) auto
@@ -1012,7 +1012,7 @@
   assumes I: "finite I" "i \<notin> I" and indep: "indep_vars (\<lambda>_. borel) X (insert i I)"
   shows "indep_var borel (X i) borel (\<lambda>\<omega>. \<Prod>i\<in>I. X i \<omega>)"
 proof -
-  have "indep_var 
+  have "indep_var
     borel ((\<lambda>f. f i) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) {i}))
     borel ((\<lambda>f. \<Prod>i\<in>I. f i) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) I))"
     using I by (intro indep_var_compose[OF indep_var_restrict[OF indep]] ) auto
@@ -1063,7 +1063,7 @@
   let ?P' = "\<Pi>\<^sub>M i\<in>I. distr M (M' i) (X i)"
   interpret D': prob_space "?D' i" for i by (intro prob_space_distr rv)
   interpret P: product_prob_space ?D' I ..
-    
+
   show ?thesis
   proof
     assume "indep_vars M' X I"
@@ -1092,7 +1092,7 @@
         using J \<open>I \<noteq> {}\<close> measurable_space[OF rv] by (auto simp: prod_emb_def PiE_iff split: if_split_asm)
       also have "emeasure M (\<Inter>i\<in>J. X i -` Y i \<inter> space M) = (\<Prod> i\<in>J. emeasure M (X i -` Y i \<inter> space M))"
         using \<open>indep_vars M' X I\<close> J \<open>I \<noteq> {}\<close> using indep_varsD[of M' X I J]
-        by (auto simp: emeasure_eq_measure setprod_ereal)
+        by (auto simp: emeasure_eq_measure setprod_ennreal measure_nonneg setprod_nonneg)
       also have "\<dots> = (\<Prod> i\<in>J. emeasure (?D' i) (Y i))"
         using rv J by (simp add: emeasure_distr)
       also have "\<dots> = emeasure ?P' E"
@@ -1133,7 +1133,7 @@
         using rv J Y by (simp add: emeasure_distr)
       finally have "emeasure M (\<Inter>j\<in>J. Y' j) = (\<Prod> i\<in>J. emeasure M (Y' i))" .
       then show "prob (\<Inter>j\<in>J. Y' j) = (\<Prod> i\<in>J. prob (Y' i))"
-        by (auto simp: emeasure_eq_measure setprod_ereal)
+        by (auto simp: emeasure_eq_measure setprod_ennreal measure_nonneg setprod_nonneg)
     qed
   qed
 qed
@@ -1160,10 +1160,10 @@
   assumes [simp]: "A \<in> sets N" "B \<in> sets N"
   shows "\<P>(x in M. X x \<in> A \<and> Y x \<in> B) = \<P>(x in M. X x \<in> A) * \<P>(x in M. Y x \<in> B)"
 proof-
-  have  " \<P>(x in M. (X x)\<in>A \<and>  (Y x)\<in> B ) = prob ((\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M)" 
+  have  " \<P>(x in M. (X x)\<in>A \<and>  (Y x)\<in> B ) = prob ((\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M)"
     by (auto intro!: arg_cong[where f= prob])
-  also have "...=  prob (X -` A \<inter> space M) * prob (Y -` B \<inter> space M)"  
-    by (auto intro!: indep_varD[where Ma=N and Mb=N])     
+  also have "...=  prob (X -` A \<inter> space M) * prob (Y -` B \<inter> space M)"
+    by (auto intro!: indep_varD[where Ma=N and Mb=N])
   also have "... = \<P>(x in M. X x \<in> A) * \<P>(x in M. Y x \<in> B)"
     by (auto intro!: arg_cong2[where f= "op *"] arg_cong[where f= prob])
   finally show ?thesis .
@@ -1202,7 +1202,8 @@
     have "emeasure ?J (A \<times> B) = emeasure M ((\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M)"
       using A B by (intro emeasure_distr[OF XY]) auto
     also have "\<dots> = emeasure M (X -` A \<inter> space M) * emeasure M (Y -` B \<inter> space M)"
-      using indep_varD[OF \<open>indep_var S X T Y\<close>, of A B] A B by (simp add: emeasure_eq_measure)
+      using indep_varD[OF \<open>indep_var S X T Y\<close>, of A B] A B
+      by (simp add: emeasure_eq_measure measure_nonneg ennreal_mult)
     also have "\<dots> = emeasure ?S A * emeasure ?T B"
       using rvs A B by (simp add: emeasure_distr)
     finally show "emeasure ?S A * emeasure ?T B = emeasure ?J (A \<times> B)" by simp
@@ -1238,15 +1239,15 @@
         using \<open>Y \<in> measurable M T\<close> by (auto intro: measurable_sets) }
     next
       fix A B assume ab: "A \<in> sets S" "B \<in> sets T"
-      then have "ereal (prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M))) = emeasure ?J (A \<times> B)"
-        using XY by (auto simp add: emeasure_distr emeasure_eq_measure intro!: arg_cong[where f="prob"])
+      then have "prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)) = emeasure ?J (A \<times> B)"
+        using XY by (auto simp add: emeasure_distr emeasure_eq_measure measure_nonneg intro!: arg_cong[where f="prob"])
       also have "\<dots> = emeasure (?S \<Otimes>\<^sub>M ?T) (A \<times> B)"
         unfolding \<open>?S \<Otimes>\<^sub>M ?T = ?J\<close> ..
       also have "\<dots> = emeasure ?S A * emeasure ?T B"
         using ab by (simp add: Y.emeasure_pair_measure_Times)
       finally show "prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)) =
         prob (X -` A \<inter> space M) * prob (Y -` B \<inter> space M)"
-        using rvs ab by (simp add: emeasure_eq_measure emeasure_distr)
+        using rvs ab by (simp add: emeasure_eq_measure emeasure_distr measure_nonneg ennreal_mult[symmetric])
     qed
   qed
 qed
@@ -1272,26 +1273,25 @@
   { fix i have "random_variable borel (Y i)"
     using I(2) by (cases "i\<in>I") (auto simp: Y_def rv_X) }
   note rv_Y = this[measurable]
-    
+
   interpret Y: prob_space "distr M borel (Y i)" for i
     using I(2) by (cases "i \<in> I") (auto intro!: prob_space_distr simp: indep_vars_def prob_space_return)
   interpret product_sigma_finite "\<lambda>i. distr M borel (Y i)"
     ..
-  
+
   have indep_Y: "indep_vars (\<lambda>i. borel) Y I"
     by (rule indep_vars_cong[THEN iffD1, OF _ _ _ I(2)]) (auto simp: Y_def)
 
-  have "(\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. X i \<omega>) \<partial>M) = (\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. max 0 (Y i \<omega>)) \<partial>M)"
+  have "(\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. X i \<omega>) \<partial>M) = (\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. Y i \<omega>) \<partial>M)"
     using I(3) by (auto intro!: nn_integral_cong setprod.cong simp add: Y_def max_def)
-  also have "\<dots> = (\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. max 0 (\<omega> i)) \<partial>distr M (Pi\<^sub>M I (\<lambda>i. borel)) (\<lambda>x. \<lambda>i\<in>I. Y i x))"
+  also have "\<dots> = (\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. \<omega> i) \<partial>distr M (Pi\<^sub>M I (\<lambda>i. borel)) (\<lambda>x. \<lambda>i\<in>I. Y i x))"
     by (subst nn_integral_distr) auto
-  also have "\<dots> = (\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. max 0 (\<omega> i)) \<partial>Pi\<^sub>M I (\<lambda>i. distr M borel (Y i)))"
+  also have "\<dots> = (\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. \<omega> i) \<partial>Pi\<^sub>M I (\<lambda>i. distr M borel (Y i)))"
     unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF \<open>I \<noteq> {}\<close> rv_Y indep_Y] ..
-  also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^sup>+\<omega>. max 0 \<omega> \<partial>distr M borel (Y i)))"
+  also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^sup>+\<omega>. \<omega> \<partial>distr M borel (Y i)))"
     by (rule product_nn_integral_setprod) (auto intro: \<open>finite I\<close>)
   also have "\<dots> = (\<Prod>i\<in>I. \<integral>\<^sup>+\<omega>. X i \<omega> \<partial>M)"
-    by (intro setprod.cong nn_integral_cong)
-       (auto simp: nn_integral_distr nn_integral_max_0 Y_def rv_X)
+    by (intro setprod.cong nn_integral_cong) (auto simp: nn_integral_distr Y_def rv_X)
   finally show ?thesis .
 qed (simp add: emeasure_space_1)
 
@@ -1314,12 +1314,12 @@
   { fix i have "integrable M (Y i)"
     using I(3) by (cases "i\<in>I") (auto simp: Y_def) }
   note int_Y = this
-    
+
   interpret Y: prob_space "distr M borel (Y i)" for i
     using I(2) by (cases "i \<in> I") (auto intro!: prob_space_distr simp: indep_vars_def prob_space_return)
   interpret product_sigma_finite "\<lambda>i. distr M borel (Y i)"
     ..
-  
+
   have indep_Y: "indep_vars (\<lambda>i. borel) Y I"
     by (rule indep_vars_cong[THEN iffD1, OF _ _ _ I(2)]) (auto simp: Y_def)
 
--- a/src/HOL/Probability/Infinite_Product_Measure.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Infinite_Product_Measure.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -47,7 +47,7 @@
      (auto intro!: emeasure_distr_restrict[symmetric])
 
 lemma (in product_prob_space) emeasure_PiM_emb:
-  "J \<subseteq> I \<Longrightarrow> finite J \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow> 
+  "J \<subseteq> I \<Longrightarrow> finite J \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
     emeasure (Pi\<^sub>M I M) (emb I J (Pi\<^sub>E J X)) = (\<Prod> i\<in>J. emeasure (M i) (X i))"
   by (subst emeasure_PiM_emb') (auto intro!: emeasure_PiM)
 
@@ -78,7 +78,8 @@
   assumes "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
   shows "measure (PiM I M) (emb I J (Pi\<^sub>E J X)) = (\<Prod> i\<in>J. measure (M i) (X i))"
   using emeasure_PiM_emb[OF assms]
-  unfolding emeasure_eq_measure M.emeasure_eq_measure by (simp add: setprod_ereal)
+  unfolding emeasure_eq_measure M.emeasure_eq_measure
+  by (simp add: setprod_ennreal measure_nonneg setprod_nonneg)
 
 lemma sets_Collect_single':
   "i \<in> I \<Longrightarrow> {x\<in>space (M i). P x} \<in> sets (M i) \<Longrightarrow> {x\<in>space (PiM I M). P (x i)} \<in> sets (PiM I M)"
@@ -203,7 +204,7 @@
     by (simp add: finite_Lim_measure_decseq)
 qed
 
-lemma nat_eq_diff_eq: 
+lemma nat_eq_diff_eq:
   fixes a b c :: nat
   shows "c \<le> b \<Longrightarrow> a = b - c \<longleftrightarrow> a + c = b"
   by auto
--- a/src/HOL/Probability/Information.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Information.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -73,7 +73,7 @@
 Kullback$-$Leibler distance.\<close>
 
 definition
-  "entropy_density b M N = log b \<circ> real_of_ereal \<circ> RN_deriv M N"
+  "entropy_density b M N = log b \<circ> enn2real \<circ> RN_deriv M N"
 
 definition
   "KL_divergence b M N = integral\<^sup>L N (entropy_density b M N)"
@@ -88,17 +88,17 @@
   shows "KL_divergence b M (density M f) = (\<integral>x. f x * log b (f x) \<partial>M)"
   unfolding KL_divergence_def
 proof (subst integral_real_density)
-  show [measurable]: "entropy_density b M (density M (\<lambda>x. ereal (f x))) \<in> borel_measurable M"
+  show [measurable]: "entropy_density b M (density M (\<lambda>x. ennreal (f x))) \<in> borel_measurable M"
     using f
     by (auto simp: comp_def entropy_density_def)
   have "density M (RN_deriv M (density M f)) = density M f"
     using f nn by (intro density_RN_deriv_density) auto
   then have eq: "AE x in M. RN_deriv M (density M f) x = f x"
-    using f nn by (intro density_unique) (auto simp: RN_deriv_nonneg)
-  show "(\<integral>x. f x * entropy_density b M (density M (\<lambda>x. ereal (f x))) x \<partial>M) = (\<integral>x. f x * log b (f x) \<partial>M)"
+    using f nn by (intro density_unique) auto
+  show "(\<integral>x. f x * entropy_density b M (density M (\<lambda>x. ennreal (f x))) x \<partial>M) = (\<integral>x. f x * log b (f x) \<partial>M)"
     apply (intro integral_cong_AE)
     apply measurable
-    using eq
+    using eq nn
     apply eventually_elim
     apply (auto simp: entropy_density_def)
     done
@@ -141,12 +141,12 @@
   have [simp, intro]: "?D_set \<in> sets M"
     using D by auto
 
-  have D_neg: "(\<integral>\<^sup>+ x. ereal (- D x) \<partial>M) = 0"
-    using D by (subst nn_integral_0_iff_AE) auto
+  have D_neg: "(\<integral>\<^sup>+ x. ennreal (- D x) \<partial>M) = 0"
+    using D by (subst nn_integral_0_iff_AE) (auto simp: ennreal_neg)
 
-  have "(\<integral>\<^sup>+ x. ereal (D x) \<partial>M) = emeasure (density M D) (space M)"
+  have "(\<integral>\<^sup>+ x. ennreal (D x) \<partial>M) = emeasure (density M D) (space M)"
     using D by (simp add: emeasure_density cong: nn_integral_cong)
-  then have D_pos: "(\<integral>\<^sup>+ x. ereal (D x) \<partial>M) = 1"
+  then have D_pos: "(\<integral>\<^sup>+ x. ennreal (D x) \<partial>M) = 1"
     using N.emeasure_space_1 by simp
 
   have "integrable M D"
@@ -162,10 +162,10 @@
   also have "\<dots> < (\<integral> x. D x * (ln b * log b (D x)) \<partial>M)"
   proof (rule integral_less_AE)
     show "integrable M (\<lambda>x. D x - indicator ?D_set x)"
-      using \<open>integrable M D\<close> by auto
+      using \<open>integrable M D\<close> by (auto simp: less_top[symmetric])
   next
     from integrable_mult_left(1)[OF int, of "ln b"]
-    show "integrable M (\<lambda>x. D x * (ln b * log b (D x)))" 
+    show "integrable M (\<lambda>x. D x * (ln b * log b (D x)))"
       by (simp add: ac_simps)
   next
     show "emeasure M {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<noteq> 0"
@@ -176,12 +176,12 @@
 
       have "emeasure M {x\<in>space M. D x = 1} = (\<integral>\<^sup>+ x. indicator {x\<in>space M. D x = 1} x \<partial>M)"
         using D(1) by auto
-      also have "\<dots> = (\<integral>\<^sup>+ x. ereal (D x) \<partial>M)"
-        using disj by (auto intro!: nn_integral_cong_AE simp: indicator_def one_ereal_def)
+      also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (D x) \<partial>M)"
+        using disj by (auto intro!: nn_integral_cong_AE simp: indicator_def one_ennreal_def)
       finally have "AE x in M. D x = 1"
         using D D_pos by (intro AE_I_eq_1) auto
-      then have "(\<integral>\<^sup>+x. indicator A x\<partial>M) = (\<integral>\<^sup>+x. ereal (D x) * indicator A x\<partial>M)"
-        by (intro nn_integral_cong_AE) (auto simp: one_ereal_def[symmetric])
+      then have "(\<integral>\<^sup>+x. indicator A x\<partial>M) = (\<integral>\<^sup>+x. ennreal (D x) * indicator A x\<partial>M)"
+        by (intro nn_integral_cong_AE) (auto simp: one_ennreal_def[symmetric])
       also have "\<dots> = density M D A"
         using \<open>A \<in> sets M\<close> D by (simp add: emeasure_density)
       finally show False using \<open>A \<in> sets M\<close> \<open>emeasure (density M D) A \<noteq> emeasure M A\<close> by simp
@@ -240,14 +240,13 @@
 proof -
   have "AE x in M. 1 = RN_deriv M M x"
   proof (rule RN_deriv_unique)
-    show "(\<lambda>x. 1) \<in> borel_measurable M" "AE x in M. 0 \<le> (1 :: ereal)" by auto
     show "density M (\<lambda>x. 1) = M"
       apply (auto intro!: measure_eqI emeasure_density)
       apply (subst emeasure_density)
       apply auto
       done
-  qed
-  then have "AE x in M. log b (real_of_ereal (RN_deriv M M x)) = 0"
+  qed auto
+  then have "AE x in M. log b (enn2real (RN_deriv M M x)) = 0"
     by (elim AE_mp) simp
   from integral_cong_AE[OF _ _ this]
   have "integral\<^sup>L M (entropy_density b M M) = 0"
@@ -276,7 +275,7 @@
   interpret N: prob_space N by fact
   have "finite_measure N" by unfold_locales
   from real_RN_deriv[OF this ac] guess D . note D = this
-  
+
   have "N = density M (RN_deriv M N)"
     using ac by (rule density_RN_deriv[symmetric])
   also have "\<dots> = density M D"
@@ -334,36 +333,28 @@
 
 subsection \<open>Finite Entropy\<close>
 
-definition (in information_space) 
-  "finite_entropy S X f \<longleftrightarrow> distributed M S X f \<and> integrable S (\<lambda>x. f x * log b (f x))"
+definition (in information_space) finite_entropy :: "'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> real) \<Rightarrow> bool"
+where
+  "finite_entropy S X f \<longleftrightarrow>
+    distributed M S X f \<and>
+    integrable S (\<lambda>x. f x * log b (f x)) \<and>
+    (\<forall>x\<in>space S. 0 \<le> f x)"
 
 lemma (in information_space) finite_entropy_simple_function:
   assumes X: "simple_function M X"
   shows "finite_entropy (count_space (X`space M)) X (\<lambda>a. measure M {x \<in> space M. X x = a})"
   unfolding finite_entropy_def
-proof
+proof safe
   have [simp]: "finite (X ` space M)"
     using X by (auto simp: simple_function_def)
   then show "integrable (count_space (X ` space M))
      (\<lambda>x. prob {xa \<in> space M. X xa = x} * log b (prob {xa \<in> space M. X xa = x}))"
     by (rule integrable_count_space)
-  have d: "distributed M (count_space (X ` space M)) X (\<lambda>x. ereal (if x \<in> X`space M then prob {xa \<in> space M. X xa = x} else 0))"
+  have d: "distributed M (count_space (X ` space M)) X (\<lambda>x. ennreal (if x \<in> X`space M then prob {xa \<in> space M. X xa = x} else 0))"
     by (rule distributed_simple_function_superset[OF X]) (auto intro!: arg_cong[where f=prob])
-  show "distributed M (count_space (X ` space M)) X (\<lambda>x. ereal (prob {xa \<in> space M. X xa = x}))"
+  show "distributed M (count_space (X ` space M)) X (\<lambda>x. ennreal (prob {xa \<in> space M. X xa = x}))"
     by (rule distributed_cong_density[THEN iffD1, OF _ _ _ d]) auto
-qed
-
-lemma distributed_transform_AE:
-  assumes T: "T \<in> measurable P Q" "absolutely_continuous Q (distr P Q T)"
-  assumes g: "distributed M Q Y g"
-  shows "AE x in P. 0 \<le> g (T x)"
-  using g
-  apply (subst AE_distr_iff[symmetric, OF T(1)])
-  apply simp
-  apply (rule absolutely_continuous_AE[OF _ T(2)])
-  apply simp
-  apply (simp add: distributed_AE)
-  done
+qed (rule measure_nonneg)
 
 lemma ac_fst:
   assumes "sigma_finite_measure T"
@@ -411,15 +402,34 @@
   "finite_entropy S X Px \<Longrightarrow> distributed M S X Px"
   unfolding finite_entropy_def by auto
 
+lemma (in information_space) finite_entropy_nn:
+  "finite_entropy S X Px \<Longrightarrow> x \<in> space S \<Longrightarrow> 0 \<le> Px x"
+  by (auto simp: finite_entropy_def)
+
+lemma (in information_space) finite_entropy_measurable:
+  "finite_entropy S X Px \<Longrightarrow> Px \<in> S \<rightarrow>\<^sub>M borel"
+  using distributed_real_measurable[of S Px M X]
+    finite_entropy_nn[of S X Px] finite_entropy_distributed[of S X Px] by auto
+
+lemma (in information_space) subdensity_finite_entropy:
+  fixes g :: "'b \<Rightarrow> real" and f :: "'c \<Rightarrow> real"
+  assumes T: "T \<in> measurable P Q"
+  assumes f: "finite_entropy P X f"
+  assumes g: "finite_entropy Q Y g"
+  assumes Y: "Y = T \<circ> X"
+  shows "AE x in P. g (T x) = 0 \<longrightarrow> f x = 0"
+  using subdensity[OF T, of M X "\<lambda>x. ennreal (f x)" Y "\<lambda>x. ennreal (g x)"]
+    finite_entropy_distributed[OF f] finite_entropy_distributed[OF g]
+    finite_entropy_nn[OF f] finite_entropy_nn[OF g]
+    assms
+  by auto
+
 lemma (in information_space) finite_entropy_integrable_transform:
-  assumes Fx: "finite_entropy S X Px"
-  assumes Fy: "distributed M T Y Py"
-    and "X = (\<lambda>x. f (Y x))"
-    and "f \<in> measurable T S"
-  shows "integrable T (\<lambda>x. Py x * log b (Px (f x)))"
-  using assms unfolding finite_entropy_def
+  "finite_entropy S X Px \<Longrightarrow> distributed M T Y Py \<Longrightarrow> (\<And>x. x \<in> space T \<Longrightarrow> 0 \<le> Py x) \<Longrightarrow>
+    X = (\<lambda>x. f (Y x)) \<Longrightarrow> f \<in> measurable T S \<Longrightarrow> integrable T (\<lambda>x. Py x * log b (Px (f x)))"
   using distributed_transform_integrable[of M T Y Py S X Px f "\<lambda>x. log b (Px x)"]
-  by auto
+  using distributed_real_measurable[of S Px M X]
+  by (auto simp: finite_entropy_def)
 
 subsection \<open>Mutual Information\<close>
 
@@ -503,16 +513,25 @@
   shows mutual_information_distr': "mutual_information b S T X Y = integral\<^sup>L (S \<Otimes>\<^sub>M T) f" (is "?M = ?R")
     and mutual_information_nonneg': "0 \<le> mutual_information b S T X Y"
 proof -
-  have Px: "distributed M S X Px"
+  have Px: "distributed M S X Px" and Px_nn: "\<And>x. x \<in> space S \<Longrightarrow> 0 \<le> Px x"
     using Fx by (auto simp: finite_entropy_def)
-  have Py: "distributed M T Y Py"
+  have Py: "distributed M T Y Py" and Py_nn: "\<And>x. x \<in> space T \<Longrightarrow> 0 \<le> Py x"
     using Fy by (auto simp: finite_entropy_def)
   have Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
-    using Fxy by (auto simp: finite_entropy_def)
+    and Pxy_nn: "\<And>x. x \<in> space (S \<Otimes>\<^sub>M T) \<Longrightarrow> 0 \<le> Pxy x"
+      "\<And>x y. x \<in> space S \<Longrightarrow> y \<in> space T \<Longrightarrow> 0 \<le> Pxy (x, y)"
+    using Fxy by (auto simp: finite_entropy_def space_pair_measure)
 
-  have X: "random_variable S X"
+  have [measurable]: "Px \<in> S \<rightarrow>\<^sub>M borel"
+    using Px Px_nn by (intro distributed_real_measurable)
+  have [measurable]: "Py \<in> T \<rightarrow>\<^sub>M borel"
+    using Py Py_nn by (intro distributed_real_measurable)
+  have measurable_Pxy[measurable]: "Pxy \<in> (S \<Otimes>\<^sub>M T) \<rightarrow>\<^sub>M borel"
+    using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure)
+
+  have X[measurable]: "random_variable S X"
     using Px by auto
-  have Y: "random_variable T Y"
+  have Y[measurable]: "random_variable T Y"
     using Py by auto
   interpret S: sigma_finite_measure S by fact
   interpret T: sigma_finite_measure T by fact
@@ -524,53 +543,46 @@
   let ?D = "distr M ?P (\<lambda>x. (X x, Y x))"
 
   { fix A assume "A \<in> sets S"
-    with X Y have "emeasure (distr M S X) A = emeasure ?D (A \<times> space T)"
-      by (auto simp: emeasure_distr measurable_Pair measurable_space
-               intro!: arg_cong[where f="emeasure M"]) }
+    with X[THEN measurable_space] Y[THEN measurable_space]
+    have "emeasure (distr M S X) A = emeasure ?D (A \<times> space T)"
+      by (auto simp: emeasure_distr intro!: arg_cong[where f="emeasure M"]) }
   note marginal_eq1 = this
   { fix A assume "A \<in> sets T"
-    with X Y have "emeasure (distr M T Y) A = emeasure ?D (space S \<times> A)"
-      by (auto simp: emeasure_distr measurable_Pair measurable_space
-               intro!: arg_cong[where f="emeasure M"]) }
+    with X[THEN measurable_space] Y[THEN measurable_space]
+    have "emeasure (distr M T Y) A = emeasure ?D (space S \<times> A)"
+      by (auto simp: emeasure_distr intro!: arg_cong[where f="emeasure M"]) }
   note marginal_eq2 = this
 
-  have eq: "(\<lambda>x. ereal (Px (fst x) * Py (snd x))) = (\<lambda>(x, y). ereal (Px x) * ereal (Py y))"
-    by auto
-
-  have distr_eq: "distr M S X \<Otimes>\<^sub>M distr M T Y = density ?P (\<lambda>x. ereal (Px (fst x) * Py (snd x)))"
-    unfolding Px(1)[THEN distributed_distr_eq_density] Py(1)[THEN distributed_distr_eq_density] eq
+  have distr_eq: "distr M S X \<Otimes>\<^sub>M distr M T Y = density ?P (\<lambda>x. ennreal (Px (fst x) * Py (snd x)))"
+    unfolding Px(1)[THEN distributed_distr_eq_density] Py(1)[THEN distributed_distr_eq_density]
   proof (subst pair_measure_density)
-    show "(\<lambda>x. ereal (Px x)) \<in> borel_measurable S" "(\<lambda>y. ereal (Py y)) \<in> borel_measurable T"
-      "AE x in S. 0 \<le> ereal (Px x)" "AE y in T. 0 \<le> ereal (Py y)"
+    show "(\<lambda>x. ennreal (Px x)) \<in> borel_measurable S" "(\<lambda>y. ennreal (Py y)) \<in> borel_measurable T"
       using Px Py by (auto simp: distributed_def)
     show "sigma_finite_measure (density T Py)" unfolding Py(1)[THEN distributed_distr_eq_density, symmetric] ..
-  qed (fact | simp)+
-  
-  have M: "?M = KL_divergence b (density ?P (\<lambda>x. ereal (Px (fst x) * Py (snd x)))) (density ?P (\<lambda>x. ereal (Pxy x)))"
+    show "density (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). ennreal (Px x) * ennreal (Py y)) =
+      density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Px (fst x) * Py (snd x)))"
+      using Px_nn Py_nn by (auto intro!: density_cong simp: distributed_def ennreal_mult space_pair_measure)
+  qed fact
+
+  have M: "?M = KL_divergence b (density ?P (\<lambda>x. ennreal (Px (fst x) * Py (snd x)))) (density ?P (\<lambda>x. ennreal (Pxy x)))"
     unfolding mutual_information_def distr_eq Pxy(1)[THEN distributed_distr_eq_density] ..
 
   from Px Py have f: "(\<lambda>x. Px (fst x) * Py (snd x)) \<in> borel_measurable ?P"
     by (intro borel_measurable_times) (auto intro: distributed_real_measurable measurable_fst'' measurable_snd'')
   have PxPy_nonneg: "AE x in ?P. 0 \<le> Px (fst x) * Py (snd x)"
-  proof (rule ST.AE_pair_measure)
-    show "{x \<in> space ?P. 0 \<le> Px (fst x) * Py (snd x)} \<in> sets ?P"
-      using f by auto
-    show "AE x in S. AE y in T. 0 \<le> Px (fst (x, y)) * Py (snd (x, y))"
-      using Px Py by (auto simp: zero_le_mult_iff dest!: distributed_real_AE)
-  qed
+    using Px_nn Py_nn by (auto simp: space_pair_measure)
 
-  have "(AE x in ?P. Px (fst x) = 0 \<longrightarrow> Pxy x = 0)"
-    by (rule subdensity_real[OF measurable_fst Pxy Px]) auto
+  have A: "(AE x in ?P. Px (fst x) = 0 \<longrightarrow> Pxy x = 0)"
+    by (rule subdensity_real[OF measurable_fst Pxy Px]) (insert Px_nn Pxy_nn, auto simp: space_pair_measure)
   moreover
-  have "(AE x in ?P. Py (snd x) = 0 \<longrightarrow> Pxy x = 0)"
-    by (rule subdensity_real[OF measurable_snd Pxy Py]) auto
+  have B: "(AE x in ?P. Py (snd x) = 0 \<longrightarrow> Pxy x = 0)"
+    by (rule subdensity_real[OF measurable_snd Pxy Py]) (insert Py_nn Pxy_nn, auto simp: space_pair_measure)
   ultimately have ac: "AE x in ?P. Px (fst x) * Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
     by eventually_elim auto
 
   show "?M = ?R"
-    unfolding M f_def
-    using b_gt_1 f PxPy_nonneg Pxy[THEN distributed_real_measurable] Pxy[THEN distributed_real_AE] ac
-    by (rule ST.KL_density_density)
+    unfolding M f_def using Pxy_nn Px_nn Py_nn
+    by (intro ST.KL_density_density b_gt_1 f PxPy_nonneg ac) (auto simp: space_pair_measure)
 
   have X: "X = fst \<circ> (\<lambda>x. (X x, Y x))" and Y: "Y = snd \<circ> (\<lambda>x. (X x, Y x))"
     by auto
@@ -579,47 +591,52 @@
     using finite_entropy_integrable[OF Fxy]
     using finite_entropy_integrable_transform[OF Fx Pxy, of fst]
     using finite_entropy_integrable_transform[OF Fy Pxy, of snd]
-    by simp
+    by (simp add: Pxy_nn)
   moreover have "f \<in> borel_measurable (S \<Otimes>\<^sub>M T)"
     unfolding f_def using Px Py Pxy
     by (auto intro: distributed_real_measurable measurable_fst'' measurable_snd''
       intro!: borel_measurable_times borel_measurable_log borel_measurable_divide)
   ultimately have int: "integrable (S \<Otimes>\<^sub>M T) f"
     apply (rule integrable_cong_AE_imp)
-    using
-      distributed_transform_AE[OF measurable_fst ac_fst, of T, OF T Px]
-      distributed_transform_AE[OF measurable_snd ac_snd, of _ _ _ _ S, OF T Py]
-      subdensity_real[OF measurable_fst Pxy Px X]
-      subdensity_real[OF measurable_snd Pxy Py Y]
-      distributed_real_AE[OF Pxy]
+    using A B AE_space
     by eventually_elim
-       (auto simp: f_def log_divide_eq log_mult_eq field_simps zero_less_mult_iff)
+       (auto simp: f_def log_divide_eq log_mult_eq field_simps space_pair_measure Px_nn Py_nn Pxy_nn
+                  less_le)
 
   show "0 \<le> ?M" unfolding M
-  proof (rule ST.KL_density_density_nonneg
-    [OF b_gt_1 f PxPy_nonneg _ Pxy[THEN distributed_real_measurable] Pxy[THEN distributed_real_AE] _ ac int[unfolded f_def]])
-    show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ereal (Pxy x))) "
+  proof (intro ST.KL_density_density_nonneg)
+    show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Pxy x))) "
       unfolding distributed_distr_eq_density[OF Pxy, symmetric]
       using distributed_measurable[OF Pxy] by (rule prob_space_distr)
-    show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ereal (Px (fst x) * Py (snd x))))"
+    show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Px (fst x) * Py (snd x))))"
       unfolding distr_eq[symmetric] by unfold_locales
-  qed
+    show "integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))))"
+      using int unfolding f_def .
+  qed (insert ac, auto simp: b_gt_1 Px_nn Py_nn Pxy_nn space_pair_measure)
 qed
 
-
 lemma (in information_space)
   fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
   assumes "sigma_finite_measure S" "sigma_finite_measure T"
-  assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py"
-  assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  assumes Px: "distributed M S X Px" and Px_nn: "\<And>x. x \<in> space S \<Longrightarrow> 0 \<le> Px x"
+    and Py: "distributed M T Y Py" and Py_nn: "\<And>y. y \<in> space T \<Longrightarrow> 0 \<le> Py y"
+    and Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+    and Pxy_nn: "\<And>x y. x \<in> space S \<Longrightarrow> y \<in> space T \<Longrightarrow> 0 \<le> Pxy (x, y)"
   defines "f \<equiv> \<lambda>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))"
   shows mutual_information_distr: "mutual_information b S T X Y = integral\<^sup>L (S \<Otimes>\<^sub>M T) f" (is "?M = ?R")
     and mutual_information_nonneg: "integrable (S \<Otimes>\<^sub>M T) f \<Longrightarrow> 0 \<le> mutual_information b S T X Y"
 proof -
-  have X: "random_variable S X"
+  have X[measurable]: "random_variable S X"
     using Px by (auto simp: distributed_def)
-  have Y: "random_variable T Y"
+  have Y[measurable]: "random_variable T Y"
     using Py by (auto simp: distributed_def)
+  have [measurable]: "Px \<in> S \<rightarrow>\<^sub>M borel"
+    using Px Px_nn by (intro distributed_real_measurable)
+  have [measurable]: "Py \<in> T \<rightarrow>\<^sub>M borel"
+    using Py Py_nn by (intro distributed_real_measurable)
+  have measurable_Pxy[measurable]: "Pxy \<in> (S \<Otimes>\<^sub>M T) \<rightarrow>\<^sub>M borel"
+    using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure)
+
   interpret S: sigma_finite_measure S by fact
   interpret T: sigma_finite_measure T by fact
   interpret ST: pair_sigma_finite S T ..
@@ -630,100 +647,100 @@
   let ?D = "distr M ?P (\<lambda>x. (X x, Y x))"
 
   { fix A assume "A \<in> sets S"
-    with X Y have "emeasure (distr M S X) A = emeasure ?D (A \<times> space T)"
-      by (auto simp: emeasure_distr measurable_Pair measurable_space
-               intro!: arg_cong[where f="emeasure M"]) }
+    with X[THEN measurable_space] Y[THEN measurable_space]
+    have "emeasure (distr M S X) A = emeasure ?D (A \<times> space T)"
+      by (auto simp: emeasure_distr intro!: arg_cong[where f="emeasure M"]) }
   note marginal_eq1 = this
   { fix A assume "A \<in> sets T"
-    with X Y have "emeasure (distr M T Y) A = emeasure ?D (space S \<times> A)"
-      by (auto simp: emeasure_distr measurable_Pair measurable_space
-               intro!: arg_cong[where f="emeasure M"]) }
+    with X[THEN measurable_space] Y[THEN measurable_space]
+    have "emeasure (distr M T Y) A = emeasure ?D (space S \<times> A)"
+      by (auto simp: emeasure_distr intro!: arg_cong[where f="emeasure M"]) }
   note marginal_eq2 = this
 
-  have eq: "(\<lambda>x. ereal (Px (fst x) * Py (snd x))) = (\<lambda>(x, y). ereal (Px x) * ereal (Py y))"
-    by auto
-
-  have distr_eq: "distr M S X \<Otimes>\<^sub>M distr M T Y = density ?P (\<lambda>x. ereal (Px (fst x) * Py (snd x)))"
-    unfolding Px(1)[THEN distributed_distr_eq_density] Py(1)[THEN distributed_distr_eq_density] eq
+  have distr_eq: "distr M S X \<Otimes>\<^sub>M distr M T Y = density ?P (\<lambda>x. ennreal (Px (fst x) * Py (snd x)))"
+    unfolding Px(1)[THEN distributed_distr_eq_density] Py(1)[THEN distributed_distr_eq_density]
   proof (subst pair_measure_density)
-    show "(\<lambda>x. ereal (Px x)) \<in> borel_measurable S" "(\<lambda>y. ereal (Py y)) \<in> borel_measurable T"
-      "AE x in S. 0 \<le> ereal (Px x)" "AE y in T. 0 \<le> ereal (Py y)"
+    show "(\<lambda>x. ennreal (Px x)) \<in> borel_measurable S" "(\<lambda>y. ennreal (Py y)) \<in> borel_measurable T"
       using Px Py by (auto simp: distributed_def)
     show "sigma_finite_measure (density T Py)" unfolding Py(1)[THEN distributed_distr_eq_density, symmetric] ..
-  qed (fact | simp)+
-  
-  have M: "?M = KL_divergence b (density ?P (\<lambda>x. ereal (Px (fst x) * Py (snd x)))) (density ?P (\<lambda>x. ereal (Pxy x)))"
+    show "density (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). ennreal (Px x) * ennreal (Py y)) =
+      density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Px (fst x) * Py (snd x)))"
+      using Px_nn Py_nn by (auto intro!: density_cong simp: distributed_def ennreal_mult space_pair_measure)
+  qed fact
+
+  have M: "?M = KL_divergence b (density ?P (\<lambda>x. ennreal (Px (fst x) * Py (snd x)))) (density ?P (\<lambda>x. ennreal (Pxy x)))"
     unfolding mutual_information_def distr_eq Pxy(1)[THEN distributed_distr_eq_density] ..
 
   from Px Py have f: "(\<lambda>x. Px (fst x) * Py (snd x)) \<in> borel_measurable ?P"
     by (intro borel_measurable_times) (auto intro: distributed_real_measurable measurable_fst'' measurable_snd'')
   have PxPy_nonneg: "AE x in ?P. 0 \<le> Px (fst x) * Py (snd x)"
-  proof (rule ST.AE_pair_measure)
-    show "{x \<in> space ?P. 0 \<le> Px (fst x) * Py (snd x)} \<in> sets ?P"
-      using f by auto
-    show "AE x in S. AE y in T. 0 \<le> Px (fst (x, y)) * Py (snd (x, y))"
-      using Px Py by (auto simp: zero_le_mult_iff dest!: distributed_real_AE)
-  qed
+    using Px_nn Py_nn by (auto simp: space_pair_measure)
 
   have "(AE x in ?P. Px (fst x) = 0 \<longrightarrow> Pxy x = 0)"
-    by (rule subdensity_real[OF measurable_fst Pxy Px]) auto
+    by (rule subdensity_real[OF measurable_fst Pxy Px]) (insert Px_nn Pxy_nn, auto simp: space_pair_measure)
   moreover
   have "(AE x in ?P. Py (snd x) = 0 \<longrightarrow> Pxy x = 0)"
-    by (rule subdensity_real[OF measurable_snd Pxy Py]) auto
+    by (rule subdensity_real[OF measurable_snd Pxy Py]) (insert Py_nn Pxy_nn, auto simp: space_pair_measure)
   ultimately have ac: "AE x in ?P. Px (fst x) * Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
     by eventually_elim auto
 
   show "?M = ?R"
     unfolding M f_def
-    using b_gt_1 f PxPy_nonneg Pxy[THEN distributed_real_measurable] Pxy[THEN distributed_real_AE] ac
-    by (rule ST.KL_density_density)
+    using b_gt_1 f PxPy_nonneg ac Pxy_nn
+    by (intro ST.KL_density_density) (auto simp: space_pair_measure)
 
   assume int: "integrable (S \<Otimes>\<^sub>M T) f"
   show "0 \<le> ?M" unfolding M
-  proof (rule ST.KL_density_density_nonneg
-    [OF b_gt_1 f PxPy_nonneg _ Pxy[THEN distributed_real_measurable] Pxy[THEN distributed_real_AE] _ ac int[unfolded f_def]])
-    show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ereal (Pxy x))) "
+  proof (intro ST.KL_density_density_nonneg)
+    show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Pxy x))) "
       unfolding distributed_distr_eq_density[OF Pxy, symmetric]
       using distributed_measurable[OF Pxy] by (rule prob_space_distr)
-    show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ereal (Px (fst x) * Py (snd x))))"
+    show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Px (fst x) * Py (snd x))))"
       unfolding distr_eq[symmetric] by unfold_locales
-  qed
+    show "integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))))"
+      using int unfolding f_def .
+  qed (insert ac, auto simp: b_gt_1 Px_nn Py_nn Pxy_nn space_pair_measure)
 qed
 
 lemma (in information_space)
   fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
   assumes "sigma_finite_measure S" "sigma_finite_measure T"
-  assumes Px[measurable]: "distributed M S X Px" and Py[measurable]: "distributed M T Y Py"
-  assumes Pxy[measurable]: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  assumes Px[measurable]: "distributed M S X Px" and Px_nn: "\<And>x. x \<in> space S \<Longrightarrow> 0 \<le> Px x"
+    and Py[measurable]: "distributed M T Y Py" and Py_nn:  "\<And>x. x \<in> space T \<Longrightarrow> 0 \<le> Py x"
+    and Pxy[measurable]: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+    and Pxy_nn: "\<And>x. x \<in> space (S \<Otimes>\<^sub>M T) \<Longrightarrow> 0 \<le> Pxy x"
   assumes ae: "AE x in S. AE y in T. Pxy (x, y) = Px x * Py y"
   shows mutual_information_eq_0: "mutual_information b S T X Y = 0"
 proof -
   interpret S: sigma_finite_measure S by fact
   interpret T: sigma_finite_measure T by fact
   interpret ST: pair_sigma_finite S T ..
+  note
+    distributed_real_measurable[OF Px_nn Px, measurable]
+    distributed_real_measurable[OF Py_nn Py, measurable]
+    distributed_real_measurable[OF Pxy_nn Pxy, measurable]
 
   have "AE x in S \<Otimes>\<^sub>M T. Px (fst x) = 0 \<longrightarrow> Pxy x = 0"
-    by (rule subdensity_real[OF measurable_fst Pxy Px]) auto
+    by (rule subdensity_real[OF measurable_fst Pxy Px]) (auto simp: Px_nn Pxy_nn space_pair_measure)
   moreover
   have "AE x in S \<Otimes>\<^sub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
-    by (rule subdensity_real[OF measurable_snd Pxy Py]) auto
-  moreover 
+    by (rule subdensity_real[OF measurable_snd Pxy Py]) (auto simp: Py_nn Pxy_nn space_pair_measure)
+  moreover
   have "AE x in S \<Otimes>\<^sub>M T. Pxy x = Px (fst x) * Py (snd x)"
-    using distributed_real_measurable[OF Px] distributed_real_measurable[OF Py] distributed_real_measurable[OF Pxy]
     by (intro ST.AE_pair_measure) (auto simp: ae intro!: measurable_snd'' measurable_fst'')
   ultimately have "AE x in S \<Otimes>\<^sub>M T. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) = 0"
     by eventually_elim simp
   then have "(\<integral>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) \<partial>(S \<Otimes>\<^sub>M T)) = (\<integral>x. 0 \<partial>(S \<Otimes>\<^sub>M T))"
     by (intro integral_cong_AE) auto
   then show ?thesis
-    by (subst mutual_information_distr[OF assms(1-5)]) simp
+    by (subst mutual_information_distr[OF assms(1-8)]) (auto simp add: space_pair_measure)
 qed
 
 lemma (in information_space) mutual_information_simple_distributed:
   assumes X: "simple_distributed M X Px" and Y: "simple_distributed M Y Py"
   assumes XY: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy"
   shows "\<I>(X ; Y) = (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x))`space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y)))"
-proof (subst mutual_information_distr[OF _ _ simple_distributed[OF X] simple_distributed[OF Y] simple_distributed_joint[OF XY]])
+proof (subst mutual_information_distr[OF _ _ simple_distributed[OF X] _ simple_distributed[OF Y] _ simple_distributed_joint[OF XY]])
   note fin = simple_distributed_joint_finite[OF XY, simp]
   show "sigma_finite_measure (count_space (X ` space M))"
     by (simp add: sigma_finite_measure_count_space_finite)
@@ -737,7 +754,7 @@
     (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y)))"
     by (auto simp add: pair_measure_count_space lebesgue_integral_count_space_finite setsum.If_cases split_beta'
              intro!: setsum.cong)
-qed
+qed (insert X Y XY, auto simp: simple_distributed_def)
 
 lemma (in information_space)
   fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
@@ -766,7 +783,7 @@
   assumes X: "distributed M S X Px"
   shows "AE x in S. RN_deriv S (density S Px) x = Px x"
 proof -
-  note D = distributed_measurable[OF X] distributed_borel_measurable[OF X] distributed_AE[OF X]
+  note D = distributed_measurable[OF X] distributed_borel_measurable[OF X]
   interpret X: prob_space "distr M S X"
     using D(1) by (rule prob_space_distr)
 
@@ -775,19 +792,20 @@
     using D
     apply (subst eq_commute)
     apply (intro RN_deriv_unique_sigma_finite)
-    apply (auto simp: distributed_distr_eq_density[symmetric, OF X] sf measure_nonneg)
+    apply (auto simp: distributed_distr_eq_density[symmetric, OF X] sf)
     done
 qed
 
 lemma (in information_space)
   fixes X :: "'a \<Rightarrow> 'b"
-  assumes X[measurable]: "distributed M MX X f"
+  assumes X[measurable]: "distributed M MX X f" and nn: "\<And>x. x \<in> space MX \<Longrightarrow> 0 \<le> f x"
   shows entropy_distr: "entropy b MX X = - (\<integral>x. f x * log b (f x) \<partial>MX)" (is ?eq)
 proof -
-  note D = distributed_measurable[OF X] distributed_borel_measurable[OF X] distributed_AE[OF X]
+  note D = distributed_measurable[OF X] distributed_borel_measurable[OF X]
   note ae = distributed_RN_deriv[OF X]
+  note distributed_real_measurable[OF nn X, measurable]
 
-  have ae_eq: "AE x in distr M MX X. log b (real_of_ereal (RN_deriv MX (distr M MX X) x)) =
+  have ae_eq: "AE x in distr M MX X. log b (enn2real (RN_deriv MX (distr M MX X) x)) =
     log b (f x)"
     unfolding distributed_distr_eq_density[OF X]
     apply (subst AE_density)
@@ -799,103 +817,79 @@
   have int_eq: "(\<integral> x. f x * log b (f x) \<partial>MX) = (\<integral> x. log b (f x) \<partial>distr M MX X)"
     unfolding distributed_distr_eq_density[OF X]
     using D
-    by (subst integral_density)
-       (auto simp: borel_measurable_ereal_iff)
+    by (subst integral_density) (auto simp: nn)
 
   show ?eq
     unfolding entropy_def KL_divergence_def entropy_density_def comp_def int_eq neg_equal_iff_equal
-    using ae_eq by (intro integral_cong_AE) auto
-qed
-
-lemma (in prob_space) distributed_imp_emeasure_nonzero:
-  assumes X: "distributed M MX X Px"
-  shows "emeasure MX {x \<in> space MX. Px x \<noteq> 0} \<noteq> 0"
-proof
-  note Px = distributed_borel_measurable[OF X] distributed_AE[OF X]
-  interpret X: prob_space "distr M MX X"
-    using distributed_measurable[OF X] by (rule prob_space_distr)
-
-  assume "emeasure MX {x \<in> space MX. Px x \<noteq> 0} = 0"
-  with Px have "AE x in MX. Px x = 0"
-    by (intro AE_I[OF subset_refl]) (auto simp: borel_measurable_ereal_iff)
-  moreover
-  from X.emeasure_space_1 have "(\<integral>\<^sup>+x. Px x \<partial>MX) = 1"
-    unfolding distributed_distr_eq_density[OF X] using Px
-    by (subst (asm) emeasure_density)
-       (auto simp: borel_measurable_ereal_iff intro!: integral_cong cong: nn_integral_cong)
-  ultimately show False
-    by (simp add: nn_integral_cong_AE)
+    using ae_eq by (intro integral_cong_AE) (auto simp: nn)
 qed
 
 lemma (in information_space) entropy_le:
   fixes Px :: "'b \<Rightarrow> real" and MX :: "'b measure"
-  assumes X[measurable]: "distributed M MX X Px"
-  and fin: "emeasure MX {x \<in> space MX. Px x \<noteq> 0} \<noteq> \<infinity>"
+  assumes X[measurable]: "distributed M MX X Px" and Px_nn[simp]: "\<And>x. x \<in> space MX \<Longrightarrow> 0 \<le> Px x"
+  and fin: "emeasure MX {x \<in> space MX. Px x \<noteq> 0} \<noteq> top"
   and int: "integrable MX (\<lambda>x. - Px x * log b (Px x))"
   shows "entropy b MX X \<le> log b (measure MX {x \<in> space MX. Px x \<noteq> 0})"
 proof -
-  note Px = distributed_borel_measurable[OF X] distributed_AE[OF X]
+  note Px = distributed_borel_measurable[OF X]
   interpret X: prob_space "distr M MX X"
     using distributed_measurable[OF X] by (rule prob_space_distr)
 
-  have " - log b (measure MX {x \<in> space MX. Px x \<noteq> 0}) = 
+  have " - log b (measure MX {x \<in> space MX. Px x \<noteq> 0}) =
     - log b (\<integral> x. indicator {x \<in> space MX. Px x \<noteq> 0} x \<partial>MX)"
-    using Px fin
-    by (auto simp: measure_def borel_measurable_ereal_iff)
+    using Px Px_nn fin by (auto simp: measure_def)
   also have "- log b (\<integral> x. indicator {x \<in> space MX. Px x \<noteq> 0} x \<partial>MX) = - log b (\<integral> x. 1 / Px x \<partial>distr M MX X)"
-    unfolding distributed_distr_eq_density[OF X] using Px
+    unfolding distributed_distr_eq_density[OF X] using Px Px_nn
     apply (intro arg_cong[where f="log b"] arg_cong[where f=uminus])
-    by (subst integral_density) (auto simp: borel_measurable_ereal_iff simp del: integral_indicator intro!: integral_cong)
+    by (subst integral_density) (auto simp del: integral_indicator intro!: integral_cong)
   also have "\<dots> \<le> (\<integral> x. - log b (1 / Px x) \<partial>distr M MX X)"
   proof (rule X.jensens_inequality[of "\<lambda>x. 1 / Px x" "{0<..}" 0 1 "\<lambda>x. - log b x"])
     show "AE x in distr M MX X. 1 / Px x \<in> {0<..}"
       unfolding distributed_distr_eq_density[OF X]
       using Px by (auto simp: AE_density)
-    have [simp]: "\<And>x. x \<in> space MX \<Longrightarrow> ereal (if Px x = 0 then 0 else 1) = indicator {x \<in> space MX. Px x \<noteq> 0} x"
-      by (auto simp: one_ereal_def)
-    have "(\<integral>\<^sup>+ x. max 0 (ereal (- (if Px x = 0 then 0 else 1))) \<partial>MX) = (\<integral>\<^sup>+ x. 0 \<partial>MX)"
-      by (intro nn_integral_cong) (auto split: split_max)
+    have [simp]: "\<And>x. x \<in> space MX \<Longrightarrow> ennreal (if Px x = 0 then 0 else 1) = indicator {x \<in> space MX. Px x \<noteq> 0} x"
+      by (auto simp: one_ennreal_def)
+    have "(\<integral>\<^sup>+ x. ennreal (- (if Px x = 0 then 0 else 1)) \<partial>MX) = (\<integral>\<^sup>+ x. 0 \<partial>MX)"
+      by (intro nn_integral_cong) (auto simp: ennreal_neg)
     then show "integrable (distr M MX X) (\<lambda>x. 1 / Px x)"
       unfolding distributed_distr_eq_density[OF X] using Px
-      by (auto simp: nn_integral_density real_integrable_def borel_measurable_ereal_iff fin nn_integral_max_0
+      by (auto simp: nn_integral_density real_integrable_def fin ennreal_neg ennreal_mult[symmetric]
               cong: nn_integral_cong)
     have "integrable MX (\<lambda>x. Px x * log b (1 / Px x)) =
       integrable MX (\<lambda>x. - Px x * log b (Px x))"
       using Px
-      by (intro integrable_cong_AE)
-         (auto simp: borel_measurable_ereal_iff log_divide_eq
-                  intro!: measurable_If)
+      by (intro integrable_cong_AE) (auto simp: log_divide_eq less_le)
     then show "integrable (distr M MX X) (\<lambda>x. - log b (1 / Px x))"
       unfolding distributed_distr_eq_density[OF X]
       using Px int
-      by (subst integrable_real_density) (auto simp: borel_measurable_ereal_iff)
+      by (subst integrable_real_density) auto
   qed (auto simp: minus_log_convex[OF b_gt_1])
   also have "\<dots> = (\<integral> x. log b (Px x) \<partial>distr M MX X)"
     unfolding distributed_distr_eq_density[OF X] using Px
     by (intro integral_cong_AE) (auto simp: AE_density log_divide_eq)
   also have "\<dots> = - entropy b MX X"
     unfolding distributed_distr_eq_density[OF X] using Px
-    by (subst entropy_distr[OF X]) (auto simp: borel_measurable_ereal_iff integral_density)
+    by (subst entropy_distr[OF X]) (auto simp: integral_density)
   finally show ?thesis
     by simp
 qed
 
 lemma (in information_space) entropy_le_space:
   fixes Px :: "'b \<Rightarrow> real" and MX :: "'b measure"
-  assumes X: "distributed M MX X Px"
+  assumes X: "distributed M MX X Px" and Px_nn[simp]: "\<And>x. x \<in> space MX \<Longrightarrow> 0 \<le> Px x"
   and fin: "finite_measure MX"
   and int: "integrable MX (\<lambda>x. - Px x * log b (Px x))"
   shows "entropy b MX X \<le> log b (measure MX (space MX))"
 proof -
-  note Px = distributed_borel_measurable[OF X] distributed_AE[OF X]
+  note Px = distributed_borel_measurable[OF X]
   interpret finite_measure MX by fact
   have "entropy b MX X \<le> log b (measure MX {x \<in> space MX. Px x \<noteq> 0})"
     using int X by (intro entropy_le) auto
   also have "\<dots> \<le> log b (measure MX (space MX))"
     using Px distributed_imp_emeasure_nonzero[OF X]
     by (intro log_le)
-       (auto intro!: borel_measurable_ereal_iff finite_measure_mono b_gt_1
-                     less_le[THEN iffD2] measure_nonneg simp: emeasure_eq_measure)
+       (auto intro!: finite_measure_mono b_gt_1 less_le[THEN iffD2]
+             simp: emeasure_eq_measure cong: conj_cong)
   finally show ?thesis .
 qed
 
@@ -907,13 +901,13 @@
     using uniform_distributed_params[OF X] by (auto simp add: measure_def)
   have eq: "(\<integral> x. indicator A x / measure MX A * log b (indicator A x / measure MX A) \<partial>MX) =
     (\<integral> x. (- log b (measure MX A) / measure MX A) * indicator A x \<partial>MX)"
-    using measure_nonneg[of MX A] uniform_distributed_params[OF X]
-    by (intro integral_cong) (auto split: split_indicator simp: log_divide_eq)
+    using uniform_distributed_params[OF X]
+    by (intro integral_cong) (auto split: split_indicator simp: log_divide_eq zero_less_measure_iff)
   show "- (\<integral> x. indicator A x / measure MX A * log b (indicator A x / measure MX A) \<partial>MX) =
     log b (measure MX A)"
     unfolding eq using uniform_distributed_params[OF X]
-    by (subst integral_mult_right) (auto simp: measure_def)
-qed
+    by (subst integral_mult_right) (auto simp: measure_def less_top[symmetric] intro!: integrable_real_indicator)
+qed simp
 
 lemma (in information_space) entropy_simple_distributed:
   "simple_distributed M X f \<Longrightarrow> \<H>(X) = - (\<Sum>x\<in>X`space M. f x * log b (f x))"
@@ -927,7 +921,7 @@
   let ?X = "count_space (X`space M)"
   have "\<H>(X) \<le> log b (measure ?X {x \<in> space ?X. f x \<noteq> 0})"
     by (rule entropy_le[OF simple_distributed[OF X]])
-       (simp_all add: simple_distributed_finite[OF X] subset_eq integrable_count_space emeasure_count_space)
+       (insert X, auto simp: simple_distributed_finite[OF X] subset_eq integrable_count_space emeasure_count_space)
   also have "measure ?X {x \<in> space ?X. f x \<noteq> 0} = card (X ` space M \<inter> {x. f x \<noteq> 0})"
     by (simp_all add: simple_distributed_finite[OF X] subset_eq emeasure_count_space measure_def Int_def)
   finally show ?thesis .
@@ -940,7 +934,7 @@
   let ?X = "count_space (X`space M)"
   have "\<H>(X) \<le> log b (measure ?X (space ?X))"
     by (rule entropy_le_space[OF simple_distributed[OF X]])
-       (simp_all add: simple_distributed_finite[OF X] subset_eq integrable_count_space emeasure_count_space finite_measure_count_space)
+       (insert X, auto simp: simple_distributed_finite[OF X] subset_eq integrable_count_space emeasure_count_space finite_measure_count_space)
   also have "measure ?X (space ?X) = card (X ` space M)"
     by (simp_all add: simple_distributed_finite[OF X] subset_eq emeasure_count_space measure_def)
   finally show ?thesis .
@@ -961,16 +955,32 @@
 lemma (in information_space)
   assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" and P: "sigma_finite_measure P"
   assumes Px[measurable]: "distributed M S X Px"
+    and Px_nn[simp]: "\<And>x. x \<in> space S \<Longrightarrow> 0 \<le> Px x"
   assumes Pz[measurable]: "distributed M P Z Pz"
+    and Pz_nn[simp]: "\<And>z. z \<in> space P \<Longrightarrow> 0 \<le> Pz z"
   assumes Pyz[measurable]: "distributed M (T \<Otimes>\<^sub>M P) (\<lambda>x. (Y x, Z x)) Pyz"
+    and Pyz_nn[simp]: "\<And>y z. y \<in> space T \<Longrightarrow> z \<in> space P \<Longrightarrow> 0 \<le> Pyz (y, z)"
   assumes Pxz[measurable]: "distributed M (S \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Z x)) Pxz"
+    and Pxz_nn[simp]: "\<And>x z. x \<in> space S \<Longrightarrow> z \<in> space P \<Longrightarrow> 0 \<le> Pxz (x, z)"
   assumes Pxyz[measurable]: "distributed M (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Y x, Z x)) Pxyz"
+    and Pxyz_nn[simp]: "\<And>x y z. x \<in> space S \<Longrightarrow> y \<in> space T \<Longrightarrow> z \<in> space P \<Longrightarrow> 0 \<le> Pxyz (x, y, z)"
   assumes I1: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))"
   assumes I2: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))"
   shows conditional_mutual_information_generic_eq: "conditional_mutual_information b S T P X Y Z
     = (\<integral>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" (is "?eq")
     and conditional_mutual_information_generic_nonneg: "0 \<le> conditional_mutual_information b S T P X Y Z" (is "?nonneg")
 proof -
+  have [measurable]: "Px \<in> S \<rightarrow>\<^sub>M borel"
+    using Px Px_nn by (intro distributed_real_measurable)
+  have [measurable]: "Pz \<in> P \<rightarrow>\<^sub>M borel"
+    using Pz Pz_nn by (intro distributed_real_measurable)
+  have measurable_Pyz[measurable]: "Pyz \<in> (T \<Otimes>\<^sub>M P) \<rightarrow>\<^sub>M borel"
+    using Pyz Pyz_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure)
+  have measurable_Pxz[measurable]: "Pxz \<in> (S \<Otimes>\<^sub>M P) \<rightarrow>\<^sub>M borel"
+    using Pxz Pxz_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure)
+  have measurable_Pxyz[measurable]: "Pxyz \<in> (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) \<rightarrow>\<^sub>M borel"
+    using Pxyz Pxyz_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure)
+
   interpret S: sigma_finite_measure S by fact
   interpret T: sigma_finite_measure T by fact
   interpret P: sigma_finite_measure P by fact
@@ -984,43 +994,34 @@
   have SP: "sigma_finite_measure (S \<Otimes>\<^sub>M P)" ..
   have YZ: "random_variable (T \<Otimes>\<^sub>M P) (\<lambda>x. (Y x, Z x))"
     using Pyz by (simp add: distributed_measurable)
-  
+
   from Pxz Pxyz have distr_eq: "distr M (S \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Z x)) =
     distr (distr M (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Y x, Z x))) (S \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). (x, z))"
     by (simp add: comp_def distr_distr)
 
   have "mutual_information b S P X Z =
     (\<integral>x. Pxz x * log b (Pxz x / (Px (fst x) * Pz (snd x))) \<partial>(S \<Otimes>\<^sub>M P))"
-    by (rule mutual_information_distr[OF S P Px Pz Pxz])
+    by (rule mutual_information_distr[OF S P Px Px_nn Pz Pz_nn Pxz Pxz_nn])
   also have "\<dots> = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))"
     using b_gt_1 Pxz Px Pz
-    by (subst distributed_transform_integral[OF Pxyz Pxz, where T="\<lambda>(x, y, z). (x, z)"]) (auto simp: split_beta')
+    by (subst distributed_transform_integral[OF Pxyz _ Pxz _, where T="\<lambda>(x, y, z). (x, z)"])
+       (auto simp: split_beta' space_pair_measure)
   finally have mi_eq:
     "mutual_information b S P X Z = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" .
-  
+
   have ae1: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Px (fst x) = 0 \<longrightarrow> Pxyz x = 0"
-    by (intro subdensity_real[of fst, OF _ Pxyz Px]) auto
+    by (intro subdensity_real[of fst, OF _ Pxyz Px]) (auto simp: space_pair_measure)
   moreover have ae2: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
-    by (intro subdensity_real[of "\<lambda>x. snd (snd x)", OF _ Pxyz Pz]) auto
+    by (intro subdensity_real[of "\<lambda>x. snd (snd x)", OF _ Pxyz Pz]) (auto simp: space_pair_measure)
   moreover have ae3: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
-    by (intro subdensity_real[of "\<lambda>x. (fst x, snd (snd x))", OF _ Pxyz Pxz]) auto
+    by (intro subdensity_real[of "\<lambda>x. (fst x, snd (snd x))", OF _ Pxyz Pxz]) (auto simp: space_pair_measure)
   moreover have ae4: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0"
-    by (intro subdensity_real[of snd, OF _ Pxyz Pyz]) auto
-  moreover have ae5: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Px (fst x)"
-    using Px by (intro STP.AE_pair_measure) (auto simp: comp_def dest: distributed_real_AE)
-  moreover have ae6: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Pyz (snd x)"
-    using Pyz by (intro STP.AE_pair_measure) (auto simp: comp_def dest: distributed_real_AE)
-  moreover have ae7: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Pz (snd (snd x))"
-    using Pz Pz[THEN distributed_real_measurable]
-    by (auto intro!: TP.AE_pair_measure STP.AE_pair_measure AE_I2[of S] dest: distributed_real_AE)
-  moreover have ae8: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Pxz (fst x, snd (snd x))"
-    using Pxz[THEN distributed_real_AE, THEN SP.AE_pair]
-    by (auto intro!: TP.AE_pair_measure STP.AE_pair_measure)
-  moreover note Pxyz[THEN distributed_real_AE]
+    by (intro subdensity_real[of snd, OF _ Pxyz Pyz]) (auto simp: space_pair_measure)
   ultimately have ae: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P.
     Pxyz x * log b (Pxyz x / (Px (fst x) * Pyz (snd x))) -
     Pxyz x * log b (Pxz (fst x, snd (snd x)) / (Px (fst x) * Pz (snd (snd x)))) =
     Pxyz x * log b (Pxyz x * Pz (snd (snd x)) / (Pxz (fst x, snd (snd x)) * Pyz (snd x))) "
+    using AE_space
   proof eventually_elim
     case (elim x)
     show ?case
@@ -1028,7 +1029,7 @@
       assume "Pxyz x \<noteq> 0"
       with elim have "0 < Px (fst x)" "0 < Pz (snd (snd x))" "0 < Pxz (fst x, snd (snd x))"
         "0 < Pyz (snd x)" "0 < Pxyz x"
-        by auto
+        by (auto simp: space_pair_measure less_le)
       then show ?thesis
         using b_gt_1 by (simp add: log_simps less_imp_le field_simps)
     qed simp
@@ -1036,7 +1037,8 @@
   with I1 I2 show ?eq
     unfolding conditional_mutual_information_def
     apply (subst mi_eq)
-    apply (subst mutual_information_distr[OF S TP Px Pyz Pxyz])
+    apply (subst mutual_information_distr[OF S TP Px Px_nn Pyz _ Pxyz])
+    apply (auto simp: space_pair_measure)
     apply (subst integral_diff[symmetric])
     apply (auto intro!: integral_cong_AE simp: split_beta' simp del: integral_diff)
     done
@@ -1053,40 +1055,34 @@
 
   let ?f = "\<lambda>(x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) / Pxyz (x, y, z)"
 
-  from subdensity_real[of snd, OF _ Pyz Pz]
-  have aeX1: "AE x in T \<Otimes>\<^sub>M P. Pz (snd x) = 0 \<longrightarrow> Pyz x = 0" by (auto simp: comp_def)
+  from subdensity_real[of snd, OF _ Pyz Pz _ AE_I2 AE_I2]
+  have aeX1: "AE x in T \<Otimes>\<^sub>M P. Pz (snd x) = 0 \<longrightarrow> Pyz x = 0"
+    by (auto simp: comp_def space_pair_measure)
   have aeX2: "AE x in T \<Otimes>\<^sub>M P. 0 \<le> Pz (snd x)"
-    using Pz by (intro TP.AE_pair_measure) (auto simp: comp_def dest: distributed_real_AE)
+    using Pz by (intro TP.AE_pair_measure) (auto simp: comp_def)
 
-  have aeX3: "AE y in T \<Otimes>\<^sub>M P. (\<integral>\<^sup>+ x. ereal (Pxz (x, snd y)) \<partial>S) = ereal (Pz (snd y))"
+  have aeX3: "AE y in T \<Otimes>\<^sub>M P. (\<integral>\<^sup>+ x. ennreal (Pxz (x, snd y)) \<partial>S) = ennreal (Pz (snd y))"
     using Pz distributed_marginal_eq_joint2[OF P S Pz Pxz]
-    by (intro TP.AE_pair_measure) (auto dest: distributed_real_AE)
+    by (intro TP.AE_pair_measure) auto
 
   have "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<le> (\<integral>\<^sup>+ (x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))"
-    apply (subst nn_integral_density)
-    apply simp
-    apply (rule distributed_AE[OF Pxyz])
-    apply auto []
-    apply (rule nn_integral_mono_AE)
-    using ae5 ae6 ae7 ae8
+    by (subst nn_integral_density)
+       (auto intro!: nn_integral_mono simp: space_pair_measure ennreal_mult[symmetric])
+  also have "\<dots> = (\<integral>\<^sup>+(y, z). (\<integral>\<^sup>+ x. ennreal (Pxz (x, z)) * ennreal (Pyz (y, z) / Pz z) \<partial>S) \<partial>(T \<Otimes>\<^sub>M P))"
+    by (subst STP.nn_integral_snd[symmetric])
+       (auto simp add: split_beta' ennreal_mult[symmetric] space_pair_measure intro!: nn_integral_cong)
+  also have "\<dots> = (\<integral>\<^sup>+x. ennreal (Pyz x) * 1 \<partial>T \<Otimes>\<^sub>M P)"
+    apply (rule nn_integral_cong_AE)
+    using aeX1 aeX2 aeX3 AE_space
     apply eventually_elim
-    apply auto
-    done
-  also have "\<dots> = (\<integral>\<^sup>+(y, z). \<integral>\<^sup>+ x. ereal (Pxz (x, z)) * ereal (Pyz (y, z) / Pz z) \<partial>S \<partial>T \<Otimes>\<^sub>M P)"
-    by (subst STP.nn_integral_snd[symmetric]) (auto simp add: split_beta')
-  also have "\<dots> = (\<integral>\<^sup>+x. ereal (Pyz x) * 1 \<partial>T \<Otimes>\<^sub>M P)"
-    apply (rule nn_integral_cong_AE)
-    using aeX1 aeX2 aeX3 distributed_AE[OF Pyz] AE_space
-    apply eventually_elim
-  proof (case_tac x, simp del: times_ereal.simps add: space_pair_measure)
-    fix a b assume "Pz b = 0 \<longrightarrow> Pyz (a, b) = 0" "0 \<le> Pz b" "a \<in> space T \<and> b \<in> space P"
-      "(\<integral>\<^sup>+ x. ereal (Pxz (x, b)) \<partial>S) = ereal (Pz b)" "0 \<le> Pyz (a, b)" 
-    then show "(\<integral>\<^sup>+ x. ereal (Pxz (x, b)) * ereal (Pyz (a, b) / Pz b) \<partial>S) = ereal (Pyz (a, b))"
-      by (subst nn_integral_multc)
-         (auto split: prod.split)
+  proof (case_tac x, simp add: space_pair_measure)
+    fix a b assume "Pz b = 0 \<longrightarrow> Pyz (a, b) = 0" "a \<in> space T \<and> b \<in> space P"
+      "(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) \<partial>S) = ennreal (Pz b)"
+    then show "(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) * ennreal (Pyz (a, b) / Pz b) \<partial>S) = ennreal (Pyz (a, b))"
+      by (subst nn_integral_multc) (auto split: prod.split simp: ennreal_mult[symmetric])
   qed
   also have "\<dots> = 1"
-    using Q.emeasure_space_1 distributed_AE[OF Pyz] distributed_distr_eq_density[OF Pyz]
+    using Q.emeasure_space_1 distributed_distr_eq_density[OF Pyz]
     by (subst nn_integral_density[symmetric]) auto
   finally have le1: "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<le> 1" .
   also have "\<dots> < \<infinity>" by simp
@@ -1094,19 +1090,16 @@
 
   have pos: "(\<integral>\<^sup>+x. ?f x \<partial>?P) \<noteq> 0"
     apply (subst nn_integral_density)
-    apply simp
-    apply (rule distributed_AE[OF Pxyz])
-    apply auto []
-    apply (simp add: split_beta')
+    apply (simp_all add: split_beta')
   proof
-    let ?g = "\<lambda>x. ereal (if Pxyz x = 0 then 0 else Pxz (fst x, snd (snd x)) * Pyz (snd x) / Pz (snd (snd x)))"
+    let ?g = "\<lambda>x. ennreal (Pxyz x) * (Pxz (fst x, snd (snd x)) * Pyz (snd x) / (Pz (snd (snd x)) * Pxyz x))"
     assume "(\<integral>\<^sup>+x. ?g x \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)) = 0"
-    then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. ?g x \<le> 0"
+    then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. ?g x = 0"
       by (intro nn_integral_0_iff_AE[THEN iffD1]) auto
     then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pxyz x = 0"
-      using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
-      by eventually_elim (auto split: if_split_asm simp: mult_le_0_iff divide_le_0_iff)
-    then have "(\<integral>\<^sup>+ x. ereal (Pxyz x) \<partial>S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) = 0"
+      using ae1 ae2 ae3 ae4 AE_space
+      by eventually_elim (auto split: if_split_asm simp: mult_le_0_iff divide_le_0_iff space_pair_measure)
+    then have "(\<integral>\<^sup>+ x. ennreal (Pxyz x) \<partial>S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) = 0"
       by (subst nn_integral_cong_AE[of _ "\<lambda>x. 0"]) auto
     with P.emeasure_space_1 show False
       by (subst (asm) emeasure_density) (auto cong: nn_integral_cong)
@@ -1116,10 +1109,7 @@
     apply (rule nn_integral_0_iff_AE[THEN iffD2])
     apply simp
     apply (subst AE_density)
-    apply simp
-    using ae5 ae6 ae7 ae8
-    apply eventually_elim
-    apply auto
+    apply (auto simp: space_pair_measure ennreal_neg)
     done
 
   have I3: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))"
@@ -1142,45 +1132,45 @@
       apply (rule nn_integral_eq_integral)
       apply (subst AE_density)
       apply simp
-      using ae5 ae6 ae7 ae8
-      apply eventually_elim
-      apply auto
+      apply (auto simp: space_pair_measure ennreal_neg)
       done
-    with nn_integral_nonneg[of ?P ?f] pos le1
+    with pos le1
     show "0 < (\<integral>x. ?f x \<partial>?P)" "(\<integral>x. ?f x \<partial>?P) \<le> 1"
-      by (simp_all add: one_ereal_def)
+      by (simp_all add: one_ennreal.rep_eq zero_less_iff_neq_zero[symmetric])
   qed
   also have "- log b (integral\<^sup>L ?P ?f) \<le> (\<integral> x. - log b (?f x) \<partial>?P)"
   proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"])
     show "AE x in ?P. ?f x \<in> {0<..}"
       unfolding AE_density[OF distributed_borel_measurable[OF Pxyz]]
-      using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
-      by eventually_elim (auto)
+      using ae1 ae2 ae3 ae4 AE_space
+      by eventually_elim (auto simp: space_pair_measure less_le)
     show "integrable ?P ?f"
-      unfolding real_integrable_def 
+      unfolding real_integrable_def
       using fin neg by (auto simp: split_beta')
     show "integrable ?P (\<lambda>x. - log b (?f x))"
       apply (subst integrable_real_density)
       apply simp
-      apply (auto intro!: distributed_real_AE[OF Pxyz]) []
+      apply (auto simp: space_pair_measure) []
       apply simp
       apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ I3])
       apply simp
       apply simp
-      using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
+      using ae1 ae2 ae3 ae4 AE_space
       apply eventually_elim
-      apply (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff zero_less_divide_iff field_simps)
+      apply (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff zero_less_divide_iff field_simps
+        less_le space_pair_measure)
       done
   qed (auto simp: b_gt_1 minus_log_convex)
   also have "\<dots> = conditional_mutual_information b S T P X Y Z"
     unfolding \<open>?eq\<close>
     apply (subst integral_real_density)
     apply simp
-    apply (auto intro!: distributed_real_AE[OF Pxyz]) []
+    apply (auto simp: space_pair_measure) []
     apply simp
     apply (intro integral_cong_AE)
-    using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
-    apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff field_simps)
+    using ae1 ae2 ae3 ae4
+    apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff field_simps
+      space_pair_measure less_le)
     done
   finally show ?nonneg
     by simp
@@ -1204,6 +1194,18 @@
   note Pxz = Fxz[THEN finite_entropy_distributed, measurable]
   note Pxyz = Fxyz[THEN finite_entropy_distributed, measurable]
 
+  note Px_nn = Fx[THEN finite_entropy_nn]
+  note Pz_nn = Fz[THEN finite_entropy_nn]
+  note Pyz_nn = Fyz[THEN finite_entropy_nn]
+  note Pxz_nn = Fxz[THEN finite_entropy_nn]
+  note Pxyz_nn = Fxyz[THEN finite_entropy_nn]
+
+  note Px' = Fx[THEN finite_entropy_measurable, measurable]
+  note Pz' = Fz[THEN finite_entropy_measurable, measurable]
+  note Pyz' = Fyz[THEN finite_entropy_measurable, measurable]
+  note Pxz' = Fxz[THEN finite_entropy_measurable, measurable]
+  note Pxyz' = Fxyz[THEN finite_entropy_measurable, measurable]
+
   interpret S: sigma_finite_measure S by fact
   interpret T: sigma_finite_measure T by fact
   interpret P: sigma_finite_measure P by fact
@@ -1222,36 +1224,28 @@
 
   have "mutual_information b S P X Z =
     (\<integral>x. Pxz x * log b (Pxz x / (Px (fst x) * Pz (snd x))) \<partial>(S \<Otimes>\<^sub>M P))"
-    by (rule mutual_information_distr[OF S P Px Pz Pxz])
+    using Px Px_nn Pz Pz_nn Pxz Pxz_nn
+    by (rule mutual_information_distr[OF S P]) (auto simp: space_pair_measure)
   also have "\<dots> = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))"
-    using b_gt_1 Pxz Px Pz
-    by (subst distributed_transform_integral[OF Pxyz Pxz, where T="\<lambda>(x, y, z). (x, z)"])
+    using b_gt_1 Pxz Pxz_nn Pxyz Pxyz_nn
+    by (subst distributed_transform_integral[OF Pxyz _ Pxz, where T="\<lambda>(x, y, z). (x, z)"])
        (auto simp: split_beta')
   finally have mi_eq:
     "mutual_information b S P X Z = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" .
-  
+
   have ae1: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Px (fst x) = 0 \<longrightarrow> Pxyz x = 0"
-    by (intro subdensity_real[of fst, OF _ Pxyz Px]) auto
+    by (intro subdensity_finite_entropy[of fst, OF _ Fxyz Fx]) auto
   moreover have ae2: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
-    by (intro subdensity_real[of "\<lambda>x. snd (snd x)", OF _ Pxyz Pz]) auto
+    by (intro subdensity_finite_entropy[of "\<lambda>x. snd (snd x)", OF _ Fxyz Fz]) auto
   moreover have ae3: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
-    by (intro subdensity_real[of "\<lambda>x. (fst x, snd (snd x))", OF _ Pxyz Pxz]) auto
+    by (intro subdensity_finite_entropy[of "\<lambda>x. (fst x, snd (snd x))", OF _ Fxyz Fxz]) auto
   moreover have ae4: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0"
-    by (intro subdensity_real[of snd, OF _ Pxyz Pyz]) auto
-  moreover have ae5: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Px (fst x)"
-    using Px by (intro STP.AE_pair_measure) (auto dest: distributed_real_AE)
-  moreover have ae6: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Pyz (snd x)"
-    using Pyz by (intro STP.AE_pair_measure) (auto dest: distributed_real_AE)
-  moreover have ae7: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Pz (snd (snd x))"
-    using Pz Pz[THEN distributed_real_measurable] by (auto intro!: TP.AE_pair_measure STP.AE_pair_measure AE_I2[of S] dest: distributed_real_AE)
-  moreover have ae8: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Pxz (fst x, snd (snd x))"
-    using Pxz[THEN distributed_real_AE, THEN SP.AE_pair]
-    by (auto intro!: TP.AE_pair_measure STP.AE_pair_measure simp: comp_def)
-  moreover note ae9 = Pxyz[THEN distributed_real_AE]
+    by (intro subdensity_finite_entropy[of snd, OF _ Fxyz Fyz]) auto
   ultimately have ae: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P.
     Pxyz x * log b (Pxyz x / (Px (fst x) * Pyz (snd x))) -
     Pxyz x * log b (Pxz (fst x, snd (snd x)) / (Px (fst x) * Pz (snd (snd x)))) =
     Pxyz x * log b (Pxyz x * Pz (snd (snd x)) / (Pxz (fst x, snd (snd x)) * Pyz (snd x))) "
+    using AE_space
   proof eventually_elim
     case (elim x)
     show ?case
@@ -1259,7 +1253,8 @@
       assume "Pxyz x \<noteq> 0"
       with elim have "0 < Px (fst x)" "0 < Pz (snd (snd x))" "0 < Pxz (fst x, snd (snd x))"
         "0 < Pyz (snd x)" "0 < Pxyz x"
-        by auto
+        using Px_nn Pz_nn Pxz_nn Pyz_nn Pxyz_nn
+        by (auto simp: space_pair_measure less_le)
       then show ?thesis
         using b_gt_1 by (simp add: log_simps less_imp_le field_simps)
     qed simp
@@ -1268,36 +1263,41 @@
   have "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)
     (\<lambda>x. Pxyz x * log b (Pxyz x) - Pxyz x * log b (Px (fst x)) - Pxyz x * log b (Pyz (snd x)))"
     using finite_entropy_integrable[OF Fxyz]
-    using finite_entropy_integrable_transform[OF Fx Pxyz, of fst]
-    using finite_entropy_integrable_transform[OF Fyz Pxyz, of snd]
+    using finite_entropy_integrable_transform[OF Fx Pxyz Pxyz_nn, of fst]
+    using finite_entropy_integrable_transform[OF Fyz Pxyz Pxyz_nn, of snd]
     by simp
   moreover have "(\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z)))) \<in> borel_measurable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)"
     using Pxyz Px Pyz by simp
   ultimately have I1: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))"
     apply (rule integrable_cong_AE_imp)
-    using ae1 ae4 ae5 ae6 ae9
+    using ae1 ae4 AE_space
     by eventually_elim
-       (auto simp: log_divide_eq log_mult_eq field_simps zero_less_mult_iff)
+       (insert Px_nn Pyz_nn Pxyz_nn,
+        auto simp: log_divide_eq log_mult_eq field_simps zero_less_mult_iff space_pair_measure less_le)
 
   have "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)
     (\<lambda>x. Pxyz x * log b (Pxz (fst x, snd (snd x))) - Pxyz x * log b (Px (fst x)) - Pxyz x * log b (Pz (snd (snd x))))"
-    using finite_entropy_integrable_transform[OF Fxz Pxyz, of "\<lambda>x. (fst x, snd (snd x))"]
-    using finite_entropy_integrable_transform[OF Fx Pxyz, of fst]
-    using finite_entropy_integrable_transform[OF Fz Pxyz, of "snd \<circ> snd"]
+    using finite_entropy_integrable_transform[OF Fxz Pxyz Pxyz_nn, of "\<lambda>x. (fst x, snd (snd x))"]
+    using finite_entropy_integrable_transform[OF Fx Pxyz Pxyz_nn, of fst]
+    using finite_entropy_integrable_transform[OF Fz Pxyz Pxyz_nn, of "snd \<circ> snd"]
     by simp
   moreover have "(\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z))) \<in> borel_measurable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)"
     using Pxyz Px Pz
     by auto
   ultimately have I2: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))"
     apply (rule integrable_cong_AE_imp)
-    using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 ae9
+    using ae1 ae2 ae3 ae4 AE_space
     by eventually_elim
-       (auto simp: log_divide_eq log_mult_eq field_simps zero_less_mult_iff)
+       (insert Px_nn Pz_nn Pxz_nn Pyz_nn Pxyz_nn,
+         auto simp: log_divide_eq log_mult_eq field_simps zero_less_mult_iff less_le space_pair_measure)
 
   from ae I1 I2 show ?eq
     unfolding conditional_mutual_information_def
     apply (subst mi_eq)
-    apply (subst mutual_information_distr[OF S TP Px Pyz Pxyz])
+    apply (subst mutual_information_distr[OF S TP Px Px_nn Pyz Pyz_nn Pxyz Pxyz_nn])
+    apply simp
+    apply simp
+    apply (simp add: space_pair_measure)
     apply (subst integral_diff[symmetric])
     apply (auto intro!: integral_cong_AE simp: split_beta' simp del: integral_diff)
     done
@@ -1312,72 +1312,66 @@
 
   let ?f = "\<lambda>(x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) / Pxyz (x, y, z)"
 
-  from subdensity_real[of snd, OF _ Pyz Pz]
+  from subdensity_finite_entropy[of snd, OF _ Fyz Fz]
   have aeX1: "AE x in T \<Otimes>\<^sub>M P. Pz (snd x) = 0 \<longrightarrow> Pyz x = 0" by (auto simp: comp_def)
   have aeX2: "AE x in T \<Otimes>\<^sub>M P. 0 \<le> Pz (snd x)"
-    using Pz by (intro TP.AE_pair_measure) (auto dest: distributed_real_AE)
+    using Pz by (intro TP.AE_pair_measure) (auto intro: Pz_nn)
 
-  have aeX3: "AE y in T \<Otimes>\<^sub>M P. (\<integral>\<^sup>+ x. ereal (Pxz (x, snd y)) \<partial>S) = ereal (Pz (snd y))"
+  have aeX3: "AE y in T \<Otimes>\<^sub>M P. (\<integral>\<^sup>+ x. ennreal (Pxz (x, snd y)) \<partial>S) = ennreal (Pz (snd y))"
     using Pz distributed_marginal_eq_joint2[OF P S Pz Pxz]
-    by (intro TP.AE_pair_measure) (auto dest: distributed_real_AE)
+    by (intro TP.AE_pair_measure) (auto )
   have "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<le> (\<integral>\<^sup>+ (x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))"
-    apply (subst nn_integral_density)
-    apply (rule distributed_borel_measurable[OF Pxyz])
-    apply (rule distributed_AE[OF Pxyz])
-    apply simp
-    apply (rule nn_integral_mono_AE)
-    using ae5 ae6 ae7 ae8
+    using Px_nn Pz_nn Pxz_nn Pyz_nn Pxyz_nn
+    by (subst nn_integral_density)
+       (auto intro!: nn_integral_mono simp: space_pair_measure ennreal_mult[symmetric])
+  also have "\<dots> = (\<integral>\<^sup>+(y, z). \<integral>\<^sup>+ x. ennreal (Pxz (x, z)) * ennreal (Pyz (y, z) / Pz z) \<partial>S \<partial>T \<Otimes>\<^sub>M P)"
+    using Px_nn Pz_nn Pxz_nn Pyz_nn Pxyz_nn
+    by (subst STP.nn_integral_snd[symmetric])
+       (auto simp add: split_beta' ennreal_mult[symmetric] space_pair_measure intro!: nn_integral_cong)
+  also have "\<dots> = (\<integral>\<^sup>+x. ennreal (Pyz x) * 1 \<partial>T \<Otimes>\<^sub>M P)"
+    apply (rule nn_integral_cong_AE)
+    using aeX1 aeX2 aeX3 AE_space
     apply eventually_elim
-    apply auto
-    done
-  also have "\<dots> = (\<integral>\<^sup>+(y, z). \<integral>\<^sup>+ x. ereal (Pxz (x, z)) * ereal (Pyz (y, z) / Pz z) \<partial>S \<partial>T \<Otimes>\<^sub>M P)"
-    by (subst STP.nn_integral_snd[symmetric]) (auto simp add: split_beta')
-  also have "\<dots> = (\<integral>\<^sup>+x. ereal (Pyz x) * 1 \<partial>T \<Otimes>\<^sub>M P)"
-    apply (rule nn_integral_cong_AE)
-    using aeX1 aeX2 aeX3 distributed_AE[OF Pyz] AE_space
-    apply eventually_elim
-  proof (case_tac x, simp del: times_ereal.simps add: space_pair_measure)
+  proof (case_tac x, simp add: space_pair_measure)
     fix a b assume "Pz b = 0 \<longrightarrow> Pyz (a, b) = 0" "0 \<le> Pz b" "a \<in> space T \<and> b \<in> space P"
-      "(\<integral>\<^sup>+ x. ereal (Pxz (x, b)) \<partial>S) = ereal (Pz b)" "0 \<le> Pyz (a, b)" 
-    then show "(\<integral>\<^sup>+ x. ereal (Pxz (x, b)) * ereal (Pyz (a, b) / Pz b) \<partial>S) = ereal (Pyz (a, b))"
-      by (subst nn_integral_multc) auto
+      "(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) \<partial>S) = ennreal (Pz b)"
+    then show "(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) * ennreal (Pyz (a, b) / Pz b) \<partial>S) = ennreal (Pyz (a, b))"
+      using Pyz_nn[of "(a,b)"]
+      by (subst nn_integral_multc) (auto simp: space_pair_measure ennreal_mult[symmetric])
   qed
   also have "\<dots> = 1"
-    using Q.emeasure_space_1 distributed_AE[OF Pyz] distributed_distr_eq_density[OF Pyz]
+    using Q.emeasure_space_1 Pyz_nn distributed_distr_eq_density[OF Pyz]
     by (subst nn_integral_density[symmetric]) auto
   finally have le1: "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<le> 1" .
   also have "\<dots> < \<infinity>" by simp
   finally have fin: "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<noteq> \<infinity>" by simp
 
-  have pos: "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<noteq> 0"
+  have "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<noteq> 0"
+    using Pxyz_nn
     apply (subst nn_integral_density)
-    apply simp
-    apply (rule distributed_AE[OF Pxyz])
-    apply simp
-    apply (simp add: split_beta')
+    apply (simp_all add: split_beta'  ennreal_mult'[symmetric] cong: nn_integral_cong)
   proof
-    let ?g = "\<lambda>x. ereal (if Pxyz x = 0 then 0 else Pxz (fst x, snd (snd x)) * Pyz (snd x) / Pz (snd (snd x)))"
+    let ?g = "\<lambda>x. ennreal (if Pxyz x = 0 then 0 else Pxz (fst x, snd (snd x)) * Pyz (snd x) / Pz (snd (snd x)))"
     assume "(\<integral>\<^sup>+ x. ?g x \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)) = 0"
-    then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. ?g x \<le> 0"
-      by (intro nn_integral_0_iff_AE[THEN iffD1]) (auto intro!: borel_measurable_ereal measurable_If)
+    then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. ?g x = 0"
+      by (intro nn_integral_0_iff_AE[THEN iffD1]) auto
     then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pxyz x = 0"
-      using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
-      by eventually_elim (auto split: if_split_asm simp: mult_le_0_iff divide_le_0_iff)
-    then have "(\<integral>\<^sup>+ x. ereal (Pxyz x) \<partial>S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) = 0"
+      using ae1 ae2 ae3 ae4 AE_space
+      by eventually_elim
+         (insert Px_nn Pz_nn Pxz_nn Pyz_nn,
+           auto split: if_split_asm simp: mult_le_0_iff divide_le_0_iff space_pair_measure)
+    then have "(\<integral>\<^sup>+ x. ennreal (Pxyz x) \<partial>S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) = 0"
       by (subst nn_integral_cong_AE[of _ "\<lambda>x. 0"]) auto
     with P.emeasure_space_1 show False
       by (subst (asm) emeasure_density) (auto cong: nn_integral_cong)
   qed
+  then have pos: "0 < (\<integral>\<^sup>+ x. ?f x \<partial>?P)"
+    by (simp add: zero_less_iff_neq_zero)
 
   have neg: "(\<integral>\<^sup>+ x. - ?f x \<partial>?P) = 0"
-    apply (rule nn_integral_0_iff_AE[THEN iffD2])
-    apply (auto simp: split_beta') []
-    apply (subst AE_density)
-    apply (auto simp: split_beta') []
-    using ae5 ae6 ae7 ae8
-    apply eventually_elim
-    apply auto
-    done
+    using Pz_nn Pxz_nn Pyz_nn Pxyz_nn
+    by (intro nn_integral_0_iff_AE[THEN iffD2])
+       (auto simp: split_beta' AE_density space_pair_measure intro!: AE_I2 ennreal_neg)
 
   have I3: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))"
     apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ integrable_diff[OF I1 I2]])
@@ -1396,48 +1390,48 @@
         by simp
     qed simp
     then have "(\<integral>\<^sup>+ x. ?f x \<partial>?P) = (\<integral>x. ?f x \<partial>?P)"
-      apply (rule nn_integral_eq_integral)
-      apply (subst AE_density)
-      apply simp
-      using ae5 ae6 ae7 ae8
-      apply eventually_elim
-      apply auto
-      done
-    with nn_integral_nonneg[of ?P ?f] pos le1
+      using Pz_nn Pxz_nn Pyz_nn Pxyz_nn
+      by (intro nn_integral_eq_integral)
+         (auto simp: AE_density space_pair_measure)
+    with pos le1
     show "0 < (\<integral>x. ?f x \<partial>?P)" "(\<integral>x. ?f x \<partial>?P) \<le> 1"
-      by (simp_all add: one_ereal_def)
+      by (simp_all add: )
   qed
   also have "- log b (integral\<^sup>L ?P ?f) \<le> (\<integral> x. - log b (?f x) \<partial>?P)"
   proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"])
     show "AE x in ?P. ?f x \<in> {0<..}"
       unfolding AE_density[OF distributed_borel_measurable[OF Pxyz]]
-      using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
-      by eventually_elim (auto)
+      using ae1 ae2 ae3 ae4 AE_space
+      by eventually_elim (insert Pxyz_nn Pyz_nn Pz_nn Pxz_nn, auto simp: space_pair_measure less_le)
     show "integrable ?P ?f"
       unfolding real_integrable_def
       using fin neg by (auto simp: split_beta')
     show "integrable ?P (\<lambda>x. - log b (?f x))"
+      using Pz_nn Pxz_nn Pyz_nn Pxyz_nn
       apply (subst integrable_real_density)
       apply simp
-      apply (auto intro!: distributed_real_AE[OF Pxyz]) []
+      apply simp
       apply simp
       apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ I3])
       apply simp
       apply simp
-      using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
+      using ae1 ae2 ae3 ae4 AE_space
       apply eventually_elim
-      apply (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff zero_less_divide_iff field_simps)
+      apply (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff
+                        zero_less_divide_iff field_simps space_pair_measure less_le)
       done
   qed (auto simp: b_gt_1 minus_log_convex)
   also have "\<dots> = conditional_mutual_information b S T P X Y Z"
     unfolding \<open>?eq\<close>
+    using Pz_nn Pxz_nn Pyz_nn Pxyz_nn
     apply (subst integral_real_density)
     apply simp
-    apply (auto intro!: distributed_real_AE[OF Pxyz]) []
+    apply simp
     apply simp
     apply (intro integral_cong_AE)
-    using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
-    apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff field_simps)
+    using ae1 ae2 ae3 ae4 AE_space
+    apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff
+                      field_simps space_pair_measure less_le)
     done
   finally show ?nonneg
     by simp
@@ -1450,8 +1444,8 @@
   assumes Pxyz: "simple_distributed M (\<lambda>x. (X x, Y x, Z x)) Pxyz"
   shows "\<I>(X ; Y | Z) =
    (\<Sum>(x, y, z)\<in>(\<lambda>x. (X x, Y x, Z x))`space M. Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))"
-proof (subst conditional_mutual_information_generic_eq[OF _ _ _ _
-    simple_distributed[OF Pz] simple_distributed_joint[OF Pyz] simple_distributed_joint[OF Pxz]
+proof (subst conditional_mutual_information_generic_eq[OF _ _ _ _ _
+    simple_distributed[OF Pz] _ simple_distributed_joint[OF Pyz] _ simple_distributed_joint[OF Pxz] _
     simple_distributed_joint2[OF Pxyz]])
   note simple_distributed_joint2_finite[OF Pxyz, simp]
   show "sigma_finite_measure (count_space (X ` space M))"
@@ -1471,10 +1465,10 @@
   from measurable_comp[OF this measurable_fst]
   have "random_variable (count_space (X ` space M)) X"
     by (simp add: comp_def)
-  then have "simple_function M X"    
+  then have "simple_function M X"
     unfolding simple_function_def by (auto simp: measurable_count_space_eq2)
   then have "simple_distributed M X ?Px"
-    by (rule simple_distributedI) auto
+    by (rule simple_distributedI) (auto simp: measure_nonneg)
   then show "distributed M (count_space (X ` space M)) X ?Px"
     by (rule simple_distributed)
 
@@ -1491,7 +1485,7 @@
     by (auto intro!: ext)
   then show "(\<integral> (x, y, z). ?i x y z \<partial>?P) = (\<Sum>(x, y, z)\<in>(\<lambda>x. (X x, Y x, Z x)) ` space M. ?j x y z)"
     by (auto intro!: setsum.cong simp add: \<open>?P = ?C\<close> lebesgue_integral_count_space_finite simple_distributed_finite setsum.If_cases split_beta')
-qed
+qed (insert Pz Pyz Pxz Pxyz, auto intro: measure_nonneg)
 
 lemma (in information_space) conditional_mutual_information_nonneg:
   assumes X: "simple_function M X" and Y: "simple_function M Y" and Z: "simple_function M Z"
@@ -1501,15 +1495,25 @@
       count_space (X`space M \<times> Y`space M \<times> Z`space M)"
     by (simp add: pair_measure_count_space X Y Z simple_functionD)
   note sf = sigma_finite_measure_count_space_finite[OF simple_functionD(1)]
-  note sd = simple_distributedI[OF _ refl]
+  note sd = simple_distributedI[OF _ _ refl]
   note sp = simple_function_Pair
   show ?thesis
    apply (rule conditional_mutual_information_generic_nonneg[OF sf[OF X] sf[OF Y] sf[OF Z]])
    apply (rule simple_distributed[OF sd[OF X]])
+   apply simp
+   apply simp
    apply (rule simple_distributed[OF sd[OF Z]])
+   apply simp
+   apply simp
    apply (rule simple_distributed_joint[OF sd[OF sp[OF Y Z]]])
+   apply simp
+   apply simp
    apply (rule simple_distributed_joint[OF sd[OF sp[OF X Z]]])
+   apply simp
+   apply simp
    apply (rule simple_distributed_joint2[OF sd[OF sp[OF X sp[OF Y Z]]]])
+   apply simp
+   apply simp
    apply (auto intro!: integrable_count_space simp: X Y Z simple_functionD)
    done
 qed
@@ -1517,8 +1521,8 @@
 subsection \<open>Conditional Entropy\<close>
 
 definition (in prob_space)
-  "conditional_entropy b S T X Y = - (\<integral>(x, y). log b (real_of_ereal (RN_deriv (S \<Otimes>\<^sub>M T) (distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))) (x, y)) / 
-    real_of_ereal (RN_deriv T (distr M T Y) y)) \<partial>distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)))"
+  "conditional_entropy b S T X Y = - (\<integral>(x, y). log b (enn2real (RN_deriv (S \<Otimes>\<^sub>M T) (distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))) (x, y)) /
+    enn2real (RN_deriv T (distr M T Y) y)) \<partial>distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)))"
 
 abbreviation (in information_space)
   conditional_entropy_Pow ("\<H>'(_ | _')") where
@@ -1527,33 +1531,39 @@
 lemma (in information_space) conditional_entropy_generic_eq:
   fixes Pxy :: "_ \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
   assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
-  assumes Py[measurable]: "distributed M T Y Py"
+  assumes Py[measurable]: "distributed M T Y Py" and Py_nn[simp]: "\<And>x. x \<in> space T \<Longrightarrow> 0 \<le> Py x"
   assumes Pxy[measurable]: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+    and Pxy_nn[simp]: "\<And>x y. x \<in> space S \<Longrightarrow> y \<in> space T \<Longrightarrow> 0 \<le> Pxy (x, y)"
   shows "conditional_entropy b S T X Y = - (\<integral>(x, y). Pxy (x, y) * log b (Pxy (x, y) / Py y) \<partial>(S \<Otimes>\<^sub>M T))"
 proof -
   interpret S: sigma_finite_measure S by fact
   interpret T: sigma_finite_measure T by fact
   interpret ST: pair_sigma_finite S T ..
 
-  have "AE x in density (S \<Otimes>\<^sub>M T) (\<lambda>x. ereal (Pxy x)). Pxy x = real_of_ereal (RN_deriv (S \<Otimes>\<^sub>M T) (distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))) x)"
+  have [measurable]: "Py \<in> T \<rightarrow>\<^sub>M borel"
+    using Py Py_nn by (intro distributed_real_measurable)
+  have measurable_Pxy[measurable]: "Pxy \<in> (S \<Otimes>\<^sub>M T) \<rightarrow>\<^sub>M borel"
+    using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure)
+
+  have "AE x in density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Pxy x)). Pxy x = enn2real (RN_deriv (S \<Otimes>\<^sub>M T) (distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))) x)"
     unfolding AE_density[OF distributed_borel_measurable, OF Pxy]
     unfolding distributed_distr_eq_density[OF Pxy]
     using distributed_RN_deriv[OF Pxy]
     by auto
   moreover
-  have "AE x in density (S \<Otimes>\<^sub>M T) (\<lambda>x. ereal (Pxy x)). Py (snd x) = real_of_ereal (RN_deriv T (distr M T Y) (snd x))"
+  have "AE x in density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Pxy x)). Py (snd x) = enn2real (RN_deriv T (distr M T Y) (snd x))"
     unfolding AE_density[OF distributed_borel_measurable, OF Pxy]
     unfolding distributed_distr_eq_density[OF Py]
     apply (rule ST.AE_pair_measure)
     apply auto
     using distributed_RN_deriv[OF Py]
     apply auto
-    done    
+    done
   ultimately
   have "conditional_entropy b S T X Y = - (\<integral>x. Pxy x * log b (Pxy x / Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))"
     unfolding conditional_entropy_def neg_equal_iff_equal
     apply (subst integral_real_density[symmetric])
-    apply (auto simp: distributed_real_AE[OF Pxy] distributed_distr_eq_density[OF Pxy]
+    apply (auto simp: distributed_distr_eq_density[OF Pxy] space_pair_measure
                 intro!: integral_cong_AE)
     done
   then show ?thesis by (simp add: split_beta')
@@ -1563,7 +1573,9 @@
   fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
   assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
   assumes Py[measurable]: "distributed M T Y Py"
+    and Py_nn[simp]: "\<And>x. x \<in> space T \<Longrightarrow> 0 \<le> Py x"
   assumes Pxy[measurable]: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+    and Pxy_nn[simp]: "\<And>x y. x \<in> space S \<Longrightarrow> y \<in> space T \<Longrightarrow> 0 \<le> Pxy (x, y)"
   assumes I1: "integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x))"
   assumes I2: "integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))"
   shows "conditional_entropy b S T X Y = entropy b (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) - entropy b T Y"
@@ -1572,35 +1584,45 @@
   interpret T: sigma_finite_measure T by fact
   interpret ST: pair_sigma_finite S T ..
 
+  have [measurable]: "Py \<in> T \<rightarrow>\<^sub>M borel"
+    using Py Py_nn by (intro distributed_real_measurable)
+  have measurable_Pxy[measurable]: "Pxy \<in> (S \<Otimes>\<^sub>M T) \<rightarrow>\<^sub>M borel"
+    using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure)
+
   have "entropy b T Y = - (\<integral>y. Py y * log b (Py y) \<partial>T)"
-    by (rule entropy_distr[OF Py])
+    by (rule entropy_distr[OF Py Py_nn])
   also have "\<dots> = - (\<integral>(x,y). Pxy (x,y) * log b (Py y) \<partial>(S \<Otimes>\<^sub>M T))"
-    using b_gt_1 Py[THEN distributed_real_measurable]
-    by (subst distributed_transform_integral[OF Pxy Py, where T=snd]) (auto intro!: integral_cong)
+    using b_gt_1
+    by (subst distributed_transform_integral[OF Pxy _ Py, where T=snd])
+       (auto intro!: integral_cong simp: space_pair_measure)
   finally have e_eq: "entropy b T Y = - (\<integral>(x,y). Pxy (x,y) * log b (Py y) \<partial>(S \<Otimes>\<^sub>M T))" .
 
+  have **: "\<And>x. x \<in> space (S \<Otimes>\<^sub>M T) \<Longrightarrow> 0 \<le> Pxy x"
+    by (auto simp: space_pair_measure)
+
   have ae2: "AE x in S \<Otimes>\<^sub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
-    by (intro subdensity_real[of snd, OF _ Pxy Py]) (auto intro: measurable_Pair)
+    by (intro subdensity_real[of snd, OF _ Pxy Py])
+       (auto intro: measurable_Pair simp: space_pair_measure)
   moreover have ae4: "AE x in S \<Otimes>\<^sub>M T. 0 \<le> Py (snd x)"
-    using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable)
-  moreover note ae5 = Pxy[THEN distributed_real_AE]
+    using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'')
   ultimately have "AE x in S \<Otimes>\<^sub>M T. 0 \<le> Pxy x \<and> 0 \<le> Py (snd x) \<and>
     (Pxy x = 0 \<or> (Pxy x \<noteq> 0 \<longrightarrow> 0 < Pxy x \<and> 0 < Py (snd x)))"
-    by eventually_elim auto
+    using AE_space by eventually_elim (auto simp: space_pair_measure less_le)
   then have ae: "AE x in S \<Otimes>\<^sub>M T.
      Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) = Pxy x * log b (Pxy x / Py (snd x))"
     by eventually_elim (auto simp: log_simps field_simps b_gt_1)
-  have "conditional_entropy b S T X Y = 
+  have "conditional_entropy b S T X Y =
     - (\<integral>x. Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))"
-    unfolding conditional_entropy_generic_eq[OF S T Py Pxy] neg_equal_iff_equal
+    unfolding conditional_entropy_generic_eq[OF S T Py Py_nn Pxy Pxy_nn, simplified] neg_equal_iff_equal
     apply (intro integral_cong_AE)
     using ae
     apply auto
     done
   also have "\<dots> = - (\<integral>x. Pxy x * log b (Pxy x) \<partial>(S \<Otimes>\<^sub>M T)) - - (\<integral>x.  Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))"
     by (simp add: integral_diff[OF I1 I2])
-  finally show ?thesis 
-    unfolding conditional_entropy_generic_eq[OF S T Py Pxy] entropy_distr[OF Pxy] e_eq
+  finally show ?thesis
+    using conditional_entropy_generic_eq[OF S T Py Py_nn Pxy Pxy_nn, simplified]
+      entropy_distr[OF Pxy **, simplified] e_eq
     by (simp add: split_beta')
 qed
 
@@ -1612,9 +1634,9 @@
     (is "?P = ?C") using X Y by (simp add: simple_functionD pair_measure_count_space)
   show ?thesis
     by (rule conditional_entropy_eq_entropy sigma_finite_measure_count_space_finite
-                 simple_functionD  X Y simple_distributed simple_distributedI[OF _ refl]
-                 simple_distributed_joint simple_function_Pair integrable_count_space)+
-       (auto simp: \<open>?P = ?C\<close> intro!: integrable_count_space simple_functionD  X Y)
+             simple_functionD  X Y simple_distributed simple_distributedI[OF _ _ refl]
+             simple_distributed_joint simple_function_Pair integrable_count_space measure_nonneg)+
+       (auto simp: \<open>?P = ?C\<close> measure_nonneg intro!: integrable_count_space simple_functionD  X Y)
 qed
 
 lemma (in information_space) conditional_entropy_eq:
@@ -1622,7 +1644,7 @@
   assumes XY: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy"
     shows "\<H>(X | Y) = - (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))"
 proof (subst conditional_entropy_generic_eq[OF _ _
-  simple_distributed[OF Y] simple_distributed_joint[OF XY]])
+  simple_distributed[OF Y] _ simple_distributed_joint[OF XY]])
   have "finite ((\<lambda>x. (X x, Y x))`space M)"
     using XY unfolding simple_distributed_def by auto
   from finite_imageI[OF this, of fst]
@@ -1643,7 +1665,7 @@
   from Y show "- (\<integral> (x, y). ?f (x, y) * log b (?f (x, y) / Py y) \<partial>?P) =
     - (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))"
     by (auto intro!: setsum.cong simp add: \<open>?P = ?C\<close> lebesgue_integral_count_space_finite simple_distributed_finite eq setsum.If_cases split_beta')
-qed
+qed (insert Y XY, auto)
 
 lemma (in information_space) conditional_mutual_information_eq_conditional_entropy:
   assumes X: "simple_function M X" and Y: "simple_function M Y"
@@ -1657,11 +1679,11 @@
   note XY = simple_function_Pair[OF X Y]
   note XXY = simple_function_Pair[OF X XY]
   have Py: "simple_distributed M Y Py"
-    using Y by (rule simple_distributedI) (auto simp: Py_def)
+    using Y by (rule simple_distributedI) (auto simp: Py_def measure_nonneg)
   have Pxy: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy"
-    using XY by (rule simple_distributedI) (auto simp: Pxy_def)
+    using XY by (rule simple_distributedI) (auto simp: Pxy_def measure_nonneg)
   have Pxxy: "simple_distributed M (\<lambda>x. (X x, X x, Y x)) Pxxy"
-    using XXY by (rule simple_distributedI) (auto simp: Pxxy_def)
+    using XXY by (rule simple_distributedI) (auto simp: Pxxy_def measure_nonneg)
   have eq: "(\<lambda>x. (X x, X x, Y x)) ` space M = (\<lambda>(x, y). (x, x, y)) ` (\<lambda>x. (X x, Y x)) ` space M"
     by auto
   have inj: "\<And>A. inj_on (\<lambda>(x, y). (x, x, y)) A"
@@ -1669,14 +1691,17 @@
   have Pxxy_eq: "\<And>x y. Pxxy (x, x, y) = Pxy (x, y)"
     by (auto simp: Pxxy_def Pxy_def intro!: arg_cong[where f=prob])
   have "AE x in count_space ((\<lambda>x. (X x, Y x))`space M). Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
-    by (intro subdensity_real[of snd, OF _ Pxy[THEN simple_distributed] Py[THEN simple_distributed]]) (auto intro: measurable_Pair)
+    using Py Pxy
+    by (intro subdensity_real[of snd, OF _ Pxy[THEN simple_distributed] Py[THEN simple_distributed]])
+       (auto intro: measurable_Pair simp: AE_count_space)
   then show ?thesis
     apply (subst conditional_mutual_information_eq[OF Py Pxy Pxy Pxxy])
     apply (subst conditional_entropy_eq[OF Py Pxy])
     apply (auto intro!: setsum.cong simp: Pxxy_eq setsum_negf[symmetric] eq setsum.reindex[OF inj]
                 log_simps zero_less_mult_iff zero_le_mult_iff field_simps mult_less_0_iff AE_count_space)
-    using Py[THEN simple_distributed, THEN distributed_real_AE] Pxy[THEN simple_distributed, THEN distributed_real_AE]
-  apply (auto simp add: not_le[symmetric] AE_count_space)
+    using Py[THEN simple_distributed] Pxy[THEN simple_distributed]
+    apply (auto simp add: not_le AE_count_space less_le antisym
+      simple_distributed_nonneg[OF Py] simple_distributed_nonneg[OF Pxy])
     done
 qed
 
@@ -1690,25 +1715,36 @@
 lemma (in information_space) mutual_information_eq_entropy_conditional_entropy_distr:
   fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" and Pxy :: "('b \<times> 'c) \<Rightarrow> real"
   assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
-  assumes Px[measurable]: "distributed M S X Px" and Py[measurable]: "distributed M T Y Py"
-  assumes Pxy[measurable]: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  assumes Px[measurable]: "distributed M S X Px"
+    and Px_nn[simp]: "\<And>x. x \<in> space S \<Longrightarrow> 0 \<le> Px x"
+    and Py[measurable]: "distributed M T Y Py"
+    and Py_nn[simp]: "\<And>x. x \<in> space T \<Longrightarrow> 0 \<le> Py x"
+    and Pxy[measurable]: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+    and Pxy_nn[simp]: "\<And>x y. x \<in> space S \<Longrightarrow> y \<in> space T \<Longrightarrow> 0 \<le> Pxy (x, y)"
   assumes Ix: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Px (fst x)))"
   assumes Iy: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))"
   assumes Ixy: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x))"
   shows  "mutual_information b S T X Y = entropy b S X + entropy b T Y - entropy b (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))"
 proof -
+  have [measurable]: "Px \<in> S \<rightarrow>\<^sub>M borel"
+    using Px Px_nn by (intro distributed_real_measurable)
+  have [measurable]: "Py \<in> T \<rightarrow>\<^sub>M borel"
+    using Py Py_nn by (intro distributed_real_measurable)
+  have measurable_Pxy[measurable]: "Pxy \<in> (S \<Otimes>\<^sub>M T) \<rightarrow>\<^sub>M borel"
+    using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure)
+
   have X: "entropy b S X = - (\<integral>x. Pxy x * log b (Px (fst x)) \<partial>(S \<Otimes>\<^sub>M T))"
-    using b_gt_1 Px[THEN distributed_real_measurable]
-    apply (subst entropy_distr[OF Px])
-    apply (subst distributed_transform_integral[OF Pxy Px, where T=fst])
-    apply (auto intro!: integral_cong)
+    using b_gt_1
+    apply (subst entropy_distr[OF Px Px_nn], simp)
+    apply (subst distributed_transform_integral[OF Pxy _ Px, where T=fst])
+    apply (auto intro!: integral_cong simp: space_pair_measure)
     done
 
   have Y: "entropy b T Y = - (\<integral>x. Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))"
-    using b_gt_1 Py[THEN distributed_real_measurable]
-    apply (subst entropy_distr[OF Py])
-    apply (subst distributed_transform_integral[OF Pxy Py, where T=snd])
-    apply (auto intro!: integral_cong)
+    using b_gt_1
+    apply (subst entropy_distr[OF Py Py_nn], simp)
+    apply (subst distributed_transform_integral[OF Pxy _ Py, where T=snd])
+    apply (auto intro!: integral_cong simp: space_pair_measure)
     done
 
   interpret S: sigma_finite_measure S by fact
@@ -1717,27 +1753,27 @@
   have ST: "sigma_finite_measure (S \<Otimes>\<^sub>M T)" ..
 
   have XY: "entropy b (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) = - (\<integral>x. Pxy x * log b (Pxy x) \<partial>(S \<Otimes>\<^sub>M T))"
-    by (subst entropy_distr[OF Pxy]) (auto intro!: integral_cong)
-  
+    by (subst entropy_distr[OF Pxy]) (auto intro!: integral_cong simp: space_pair_measure)
+
   have "AE x in S \<Otimes>\<^sub>M T. Px (fst x) = 0 \<longrightarrow> Pxy x = 0"
-    by (intro subdensity_real[of fst, OF _ Pxy Px]) (auto intro: measurable_Pair)
+    by (intro subdensity_real[of fst, OF _ Pxy Px]) (auto intro: measurable_Pair simp: space_pair_measure)
   moreover have "AE x in S \<Otimes>\<^sub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
-    by (intro subdensity_real[of snd, OF _ Pxy Py]) (auto intro: measurable_Pair)
+    by (intro subdensity_real[of snd, OF _ Pxy Py]) (auto intro: measurable_Pair simp: space_pair_measure)
   moreover have "AE x in S \<Otimes>\<^sub>M T. 0 \<le> Px (fst x)"
-    using Px by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_fst'' dest: distributed_real_AE distributed_real_measurable)
+    using Px by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_fst'')
   moreover have "AE x in S \<Otimes>\<^sub>M T. 0 \<le> Py (snd x)"
-    using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable)
-  moreover note Pxy[THEN distributed_real_AE]
-  ultimately have "AE x in S \<Otimes>\<^sub>M T. Pxy x * log b (Pxy x) - Pxy x * log b (Px (fst x)) - Pxy x * log b (Py (snd x)) = 
+    using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'')
+  ultimately have "AE x in S \<Otimes>\<^sub>M T. Pxy x * log b (Pxy x) - Pxy x * log b (Px (fst x)) - Pxy x * log b (Py (snd x)) =
     Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))"
     (is "AE x in _. ?f x = ?g x")
+    using AE_space
   proof eventually_elim
     case (elim x)
     show ?case
     proof cases
       assume "Pxy x \<noteq> 0"
       with elim have "0 < Px (fst x)" "0 < Py (snd x)" "0 < Pxy x"
-        by auto
+        by (auto simp: space_pair_measure less_le)
       then show ?thesis
         using b_gt_1 by (simp add: log_simps less_imp_le field_simps)
     qed simp
@@ -1754,15 +1790,18 @@
   also have "\<dots> = integral\<^sup>L (S \<Otimes>\<^sub>M T) ?g"
     using \<open>AE x in _. ?f x = ?g x\<close> by (intro integral_cong_AE) auto
   also have "\<dots> = mutual_information b S T X Y"
-    by (rule mutual_information_distr[OF S T Px Py Pxy, symmetric])
+    by (rule mutual_information_distr[OF S T Px _ Py _ Pxy _ , symmetric])
+       (auto simp: space_pair_measure)
   finally show ?thesis ..
 qed
 
 lemma (in information_space) mutual_information_eq_entropy_conditional_entropy':
   fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" and Pxy :: "('b \<times> 'c) \<Rightarrow> real"
   assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
-  assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py"
+  assumes Px: "distributed M S X Px" "\<And>x. x \<in> space S \<Longrightarrow> 0 \<le> Px x"
+    and Py: "distributed M T Y Py" "\<And>x. x \<in> space T \<Longrightarrow> 0 \<le> Py x"
   assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+    "\<And>x. x \<in> space (S \<Otimes>\<^sub>M T) \<Longrightarrow> 0 \<le> Pxy x"
   assumes Ix: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Px (fst x)))"
   assumes Iy: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))"
   assumes Ixy: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x))"
@@ -1770,27 +1809,30 @@
   using
     mutual_information_eq_entropy_conditional_entropy_distr[OF S T Px Py Pxy Ix Iy Ixy]
     conditional_entropy_eq_entropy[OF S T Py Pxy Ixy Iy]
-  by simp
+  by (simp add: space_pair_measure)
 
 lemma (in information_space) mutual_information_eq_entropy_conditional_entropy:
   assumes sf_X: "simple_function M X" and sf_Y: "simple_function M Y"
   shows  "\<I>(X ; Y) = \<H>(X) - \<H>(X | Y)"
 proof -
   have X: "simple_distributed M X (\<lambda>x. measure M (X -` {x} \<inter> space M))"
-    using sf_X by (rule simple_distributedI) auto
+    using sf_X by (rule simple_distributedI) (auto simp: measure_nonneg)
   have Y: "simple_distributed M Y (\<lambda>x. measure M (Y -` {x} \<inter> space M))"
-    using sf_Y by (rule simple_distributedI) auto
+    using sf_Y by (rule simple_distributedI) (auto simp: measure_nonneg)
   have sf_XY: "simple_function M (\<lambda>x. (X x, Y x))"
     using sf_X sf_Y by (rule simple_function_Pair)
   then have XY: "simple_distributed M (\<lambda>x. (X x, Y x)) (\<lambda>x. measure M ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M))"
-    by (rule simple_distributedI) auto
+    by (rule simple_distributedI) (auto simp: measure_nonneg)
   from simple_distributed_joint_finite[OF this, simp]
   have eq: "count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M) = count_space (X ` space M \<times> Y ` space M)"
     by (simp add: pair_measure_count_space)
 
   have "\<I>(X ; Y) = \<H>(X) + \<H>(Y) - entropy b (count_space (X`space M) \<Otimes>\<^sub>M count_space (Y`space M)) (\<lambda>x. (X x, Y x))"
-    using sigma_finite_measure_count_space_finite sigma_finite_measure_count_space_finite simple_distributed[OF X] simple_distributed[OF Y] simple_distributed_joint[OF XY]
-    by (rule mutual_information_eq_entropy_conditional_entropy_distr) (auto simp: eq integrable_count_space)
+    using sigma_finite_measure_count_space_finite
+      sigma_finite_measure_count_space_finite
+      simple_distributed[OF X] measure_nonneg simple_distributed[OF Y] measure_nonneg simple_distributed_joint[OF XY]
+    by (rule mutual_information_eq_entropy_conditional_entropy_distr)
+       (auto simp: eq integrable_count_space measure_nonneg)
   then show ?thesis
     unfolding conditional_entropy_eq_entropy_simple[OF sf_X sf_Y] by simp
 qed
@@ -1800,22 +1842,22 @@
   shows  "0 \<le> \<I>(X ; Y)"
 proof -
   have X: "simple_distributed M X (\<lambda>x. measure M (X -` {x} \<inter> space M))"
-    using sf_X by (rule simple_distributedI) auto
+    using sf_X by (rule simple_distributedI) (auto simp: measure_nonneg)
   have Y: "simple_distributed M Y (\<lambda>x. measure M (Y -` {x} \<inter> space M))"
-    using sf_Y by (rule simple_distributedI) auto
+    using sf_Y by (rule simple_distributedI) (auto simp: measure_nonneg)
 
   have sf_XY: "simple_function M (\<lambda>x. (X x, Y x))"
     using sf_X sf_Y by (rule simple_function_Pair)
   then have XY: "simple_distributed M (\<lambda>x. (X x, Y x)) (\<lambda>x. measure M ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M))"
-    by (rule simple_distributedI) auto
+    by (rule simple_distributedI) (auto simp: measure_nonneg)
 
   from simple_distributed_joint_finite[OF this, simp]
   have eq: "count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M) = count_space (X ` space M \<times> Y ` space M)"
     by (simp add: pair_measure_count_space)
 
   show ?thesis
-    by (rule mutual_information_nonneg[OF _ _ simple_distributed[OF X] simple_distributed[OF Y] simple_distributed_joint[OF XY]])
-       (simp_all add: eq integrable_count_space sigma_finite_measure_count_space_finite)
+    by (rule mutual_information_nonneg[OF _ _ simple_distributed[OF X] _ simple_distributed[OF Y] _ simple_distributed_joint[OF XY]])
+       (simp_all add: eq integrable_count_space sigma_finite_measure_count_space_finite measure_nonneg)
 qed
 
 lemma (in information_space) conditional_entropy_less_eq_entropy:
@@ -1827,7 +1869,7 @@
   finally show ?thesis by auto
 qed
 
-lemma (in information_space) 
+lemma (in information_space)
   fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" and Pxy :: "('b \<times> 'c) \<Rightarrow> real"
   assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
   assumes Px: "finite_entropy S X Px" and Py: "finite_entropy T Y Py"
@@ -1835,14 +1877,15 @@
   shows "conditional_entropy b S T X Y \<le> entropy b S X"
 proof -
 
-  have "0 \<le> mutual_information b S T X Y" 
+  have "0 \<le> mutual_information b S T X Y"
     by (rule mutual_information_nonneg') fact+
   also have "\<dots> = entropy b S X - conditional_entropy b S T X Y"
     apply (rule mutual_information_eq_entropy_conditional_entropy')
     using assms
     by (auto intro!: finite_entropy_integrable finite_entropy_distributed
       finite_entropy_integrable_transform[OF Px]
-      finite_entropy_integrable_transform[OF Py])
+      finite_entropy_integrable_transform[OF Py]
+      intro: finite_entropy_nn)
   finally show ?thesis by auto
 qed
 
@@ -1850,8 +1893,8 @@
   assumes X: "simple_function M X" and Y: "simple_function M Y"
   shows  "\<H>(\<lambda>x. (X x, Y x)) = \<H>(X) + \<H>(Y|X)"
 proof -
-  note XY = simple_distributedI[OF simple_function_Pair[OF X Y] refl]
-  note YX = simple_distributedI[OF simple_function_Pair[OF Y X] refl]
+  note XY = simple_distributedI[OF simple_function_Pair[OF X Y] measure_nonneg refl]
+  note YX = simple_distributedI[OF simple_function_Pair[OF Y X] measure_nonneg refl]
   note simple_distributed_joint_finite[OF this, simp]
   let ?f = "\<lambda>x. prob ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M)"
   let ?g = "\<lambda>x. prob ((\<lambda>x. (Y x, X x)) -` {x} \<inter> space M)"
@@ -1865,7 +1908,7 @@
   also have "\<dots> = entropy b (count_space (Y ` space M) \<Otimes>\<^sub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))"
     by (subst entropy_distr[OF simple_distributed_joint[OF YX]])
        (auto simp: pair_measure_count_space sigma_finite_measure_count_space_finite lebesgue_integral_count_space_finite
-             cong del: setsum.cong  intro!: setsum.mono_neutral_left)
+             cong del: setsum.cong  intro!: setsum.mono_neutral_left measure_nonneg)
   finally have "\<H>(\<lambda>x. (X x, Y x)) = entropy b (count_space (Y ` space M) \<Otimes>\<^sub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))" .
   then show ?thesis
     unfolding conditional_entropy_eq_entropy_simple[OF Y X] by simp
@@ -1875,14 +1918,14 @@
   assumes X: "simple_function M X"
   shows "\<H>(X) = \<H>(f \<circ> X) + \<H>(X|f \<circ> X)"
 proof -
-  note fX = simple_function_compose[OF X, of f]  
+  note fX = simple_function_compose[OF X, of f]
   have eq: "(\<lambda>x. ((f \<circ> X) x, X x)) ` space M = (\<lambda>x. (f x, x)) ` X ` space M" by auto
   have inj: "\<And>A. inj_on (\<lambda>x. (f x, x)) A"
     by (auto simp: inj_on_def)
   show ?thesis
     apply (subst entropy_chain_rule[symmetric, OF fX X])
-    apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_Pair[OF fX X] refl]])
-    apply (subst entropy_simple_distributed[OF simple_distributedI[OF X refl]])
+    apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_Pair[OF fX X] measure_nonneg refl]])
+    apply (subst entropy_simple_distributed[OF simple_distributedI[OF X measure_nonneg refl]])
     unfolding eq
     apply (subst setsum.reindex[OF inj])
     apply (auto intro!: setsum.cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"])
@@ -1908,9 +1951,9 @@
     using X by auto
   have "\<H>(X) = \<H>(the_inv_into (X`space M) f \<circ> (f \<circ> X))"
     unfolding o_assoc
-    apply (subst entropy_simple_distributed[OF simple_distributedI[OF X refl]])
+    apply (subst entropy_simple_distributed[OF simple_distributedI[OF X measure_nonneg refl]])
     apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_compose[OF X]], where f="\<lambda>x. prob (X -` {x} \<inter> space M)"])
-    apply (auto intro!: setsum.cong arg_cong[where f=prob] image_eqI simp: the_inv_into_f_f[OF inj] comp_def)
+    apply (auto intro!: setsum.cong arg_cong[where f=prob] image_eqI simp: the_inv_into_f_f[OF inj] comp_def measure_nonneg)
     done
   also have "... \<le> \<H>(f \<circ> X)"
     using entropy_data_processing[OF sf] .
--- a/src/HOL/Probability/Lebesgue_Integral_Substitution.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Lebesgue_Integral_Substitution.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -13,29 +13,29 @@
 begin
 
 lemma nn_integral_substitution_aux:
-  fixes f :: "real \<Rightarrow> ereal"
+  fixes f :: "real \<Rightarrow> ennreal"
   assumes Mf: "f \<in> borel_measurable borel"
   assumes nonnegf: "\<And>x. f x \<ge> 0"
   assumes derivg: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
-  assumes contg': "continuous_on {a..b} g'" 
+  assumes contg': "continuous_on {a..b} g'"
   assumes derivg_nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
   assumes "a < b"
-  shows "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) = 
+  shows "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) =
              (\<integral>\<^sup>+x. f (g x) * g' x * indicator {a..b} x \<partial>lborel)"
 proof-
   from \<open>a < b\<close> have [simp]: "a \<le> b" by simp
   from derivg have contg: "continuous_on {a..b} g" by (rule has_real_derivative_imp_continuous_on)
-  from this and contg' have Mg: "set_borel_measurable borel {a..b} g" and 
-                             Mg': "set_borel_measurable borel {a..b} g'" 
+  from this and contg' have Mg: "set_borel_measurable borel {a..b} g" and
+                             Mg': "set_borel_measurable borel {a..b} g'"
       by (simp_all only: set_measurable_continuous_on_ivl)
   from derivg have derivg': "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
     by (simp only: has_field_derivative_iff_has_vector_derivative)
 
-  have real_ind[simp]: "\<And>A x. real_of_ereal (indicator A x :: ereal) = indicator A x" 
+  have real_ind[simp]: "\<And>A x. enn2real (indicator A x) = indicator A x"
       by (auto split: split_indicator)
-  have ereal_ind[simp]: "\<And>A x. ereal (indicator A x) = indicator A x" 
+  have ennreal_ind[simp]: "\<And>A x. ennreal (indicator A x) = indicator A x"
       by (auto split: split_indicator)
-  have [simp]: "\<And>x A. indicator A (g x) = indicator (g -` A) x" 
+  have [simp]: "\<And>x A. indicator A (g x) = indicator (g -` A) x"
       by (auto split: split_indicator)
 
   from derivg derivg_nonneg have monog: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
@@ -43,7 +43,7 @@
   with monog have [simp]: "g a \<le> g b" by (auto intro: mono_onD)
 
   show ?thesis
-  proof (induction rule: borel_measurable_induct[OF Mf nonnegf, case_names cong set mult add sup])
+  proof (induction rule: borel_measurable_induct[OF Mf, case_names cong set mult add sup])
     case (cong f1 f2)
     from cong.hyps(3) have "f1 = f2" by auto
     with cong show ?case by simp
@@ -57,42 +57,42 @@
       case (interval c d)
       {
         fix u v :: real assume asm: "{u..v} \<subseteq> {g a..g b}" "u \<le> v"
-        
+
         obtain u' v' where u'v': "{a..b} \<inter> g-`{u..v} = {u'..v'}" "u' \<le> v'" "g u' = u" "g v' = v"
              using asm by (rule_tac continuous_interval_vimage_Int[OF contg monog, of u v]) simp_all
         hence "{u'..v'} \<subseteq> {a..b}" "{u'..v'} \<subseteq> g -` {u..v}" by blast+
         with u'v'(2) have "u' \<in> g -` {u..v}" "v' \<in> g -` {u..v}" by auto
         from u'v'(1) have [simp]: "{a..b} \<inter> g -` {u..v} \<in> sets borel" by simp
-        
+
         have A: "continuous_on {min u' v'..max u' v'} g'"
-            by (simp only: u'v' max_absorb2 min_absorb1) 
+            by (simp only: u'v' max_absorb2 min_absorb1)
                (intro continuous_on_subset[OF contg'], insert u'v', auto)
         have "\<And>x. x \<in> {u'..v'} \<Longrightarrow> (g has_real_derivative g' x) (at x within {u'..v'})"
            using asm by (intro has_field_derivative_subset[OF derivg] set_mp[OF \<open>{u'..v'} \<subseteq> {a..b}\<close>]) auto
-        hence B: "\<And>x. min u' v' \<le> x \<Longrightarrow> x \<le> max u' v' \<Longrightarrow> 
-                      (g has_vector_derivative g' x) (at x within {min u' v'..max u' v'})" 
-            by (simp only: u'v' max_absorb2 min_absorb1) 
+        hence B: "\<And>x. min u' v' \<le> x \<Longrightarrow> x \<le> max u' v' \<Longrightarrow>
+                      (g has_vector_derivative g' x) (at x within {min u' v'..max u' v'})"
+            by (simp only: u'v' max_absorb2 min_absorb1)
                (auto simp: has_field_derivative_iff_has_vector_derivative)
         have "integrable lborel (\<lambda>x. indicator ({a..b} \<inter> g -` {u..v}) x *\<^sub>R g' x)"
           by (rule set_integrable_subset[OF borel_integrable_atLeastAtMost'[OF contg']]) simp_all
-        hence "(\<integral>\<^sup>+x. ereal (g' x) * indicator ({a..b} \<inter> g-` {u..v}) x \<partial>lborel) = 
-                   LBINT x:{a..b} \<inter> g-`{u..v}. g' x" 
-          by (subst ereal_ind[symmetric], subst times_ereal.simps, subst nn_integral_eq_integral)
-             (auto intro: measurable_sets Mg simp: derivg_nonneg mult.commute split: split_indicator)
+        hence "(\<integral>\<^sup>+x. ennreal (g' x) * indicator ({a..b} \<inter> g-` {u..v}) x \<partial>lborel) =
+                   LBINT x:{a..b} \<inter> g-`{u..v}. g' x"
+          by (subst nn_integral_eq_integral[symmetric])
+             (auto intro!: derivg_nonneg nn_integral_cong split: split_indicator)
         also from interval_integral_FTC_finite[OF A B]
             have "LBINT x:{a..b} \<inter> g-`{u..v}. g' x = v - u"
                 by (simp add: u'v' interval_integral_Icc \<open>u \<le> v\<close>)
-        finally have "(\<integral>\<^sup>+ x. ereal (g' x) * indicator ({a..b} \<inter> g -` {u..v}) x \<partial>lborel) =
-                           ereal (v - u)" .
+        finally have "(\<integral>\<^sup>+ x. ennreal (g' x) * indicator ({a..b} \<inter> g -` {u..v}) x \<partial>lborel) =
+                           ennreal (v - u)" .
       } note A = this
-  
-      have "(\<integral>\<^sup>+x. indicator {c..d} (g x) * ereal (g' x) * indicator {a..b} x \<partial>lborel) =
-               (\<integral>\<^sup>+ x. ereal (g' x) * indicator ({a..b} \<inter> g -` {c..d}) x \<partial>lborel)" 
+
+      have "(\<integral>\<^sup>+x. indicator {c..d} (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel) =
+               (\<integral>\<^sup>+ x. ennreal (g' x) * indicator ({a..b} \<inter> g -` {c..d}) x \<partial>lborel)"
         by (intro nn_integral_cong) (simp split: split_indicator)
-      also have "{a..b} \<inter> g-`{c..d} = {a..b} \<inter> g-`{max (g a) c..min (g b) d}" 
+      also have "{a..b} \<inter> g-`{c..d} = {a..b} \<inter> g-`{max (g a) c..min (g b) d}"
         using \<open>a \<le> b\<close> \<open>c \<le> d\<close>
         by (auto intro!: monog intro: order.trans)
-      also have "(\<integral>\<^sup>+ x. ereal (g' x) * indicator ... x \<partial>lborel) =
+      also have "(\<integral>\<^sup>+ x. ennreal (g' x) * indicator ... x \<partial>lborel) =
         (if max (g a) c \<le> min (g b) d then min (g b) d - max (g a) c else 0)"
          using \<open>c \<le> d\<close> by (simp add: A)
       also have "... = (\<integral>\<^sup>+ x. indicator ({g a..g b} \<inter> {c..d}) x \<partial>lborel)"
@@ -106,48 +106,48 @@
       case (compl A)
       note \<open>A \<in> sets borel\<close>[measurable]
       from emeasure_mono[of "A \<inter> {g a..g b}" "{g a..g b}" lborel]
-          have [simp]: "emeasure lborel (A \<inter> {g a..g b}) \<noteq> \<infinity>" by auto
+      have [simp]: "emeasure lborel (A \<inter> {g a..g b}) \<noteq> top" by (auto simp: top_unique)
       have [simp]: "g -` A \<inter> {a..b} \<in> sets borel"
         by (rule set_borel_measurable_sets[OF Mg]) auto
       have [simp]: "g -` (-A) \<inter> {a..b} \<in> sets borel"
         by (rule set_borel_measurable_sets[OF Mg]) auto
 
-      have "(\<integral>\<^sup>+x. indicator (-A) x * indicator {g a..g b} x \<partial>lborel) = 
-                (\<integral>\<^sup>+x. indicator (-A \<inter> {g a..g b}) x \<partial>lborel)" 
+      have "(\<integral>\<^sup>+x. indicator (-A) x * indicator {g a..g b} x \<partial>lborel) =
+                (\<integral>\<^sup>+x. indicator (-A \<inter> {g a..g b}) x \<partial>lborel)"
         by (rule nn_integral_cong) (simp split: split_indicator)
       also from compl have "... = emeasure lborel ({g a..g b} - A)" using derivg_nonneg
         by (simp add: vimage_Compl diff_eq Int_commute[of "-A"])
       also have "{g a..g b} - A = {g a..g b} - A \<inter> {g a..g b}" by blast
       also have "emeasure lborel ... = g b - g a - emeasure lborel (A \<inter> {g a..g b})"
-             using \<open>A \<in> sets borel\<close> by (subst emeasure_Diff) (auto simp: real_of_ereal_minus)
-     also have "emeasure lborel (A \<inter> {g a..g b}) = 
-                    \<integral>\<^sup>+x. indicator A x * indicator {g a..g b} x \<partial>lborel" 
+             using \<open>A \<in> sets borel\<close> by (subst emeasure_Diff) (auto simp: )
+     also have "emeasure lborel (A \<inter> {g a..g b}) =
+                    \<integral>\<^sup>+x. indicator A x * indicator {g a..g b} x \<partial>lborel"
        using \<open>A \<in> sets borel\<close>
        by (subst nn_integral_indicator[symmetric], simp, intro nn_integral_cong)
           (simp split: split_indicator)
-      also have "... = \<integral>\<^sup>+ x. indicator (g-`A \<inter> {a..b}) x * ereal (g' x * indicator {a..b} x) \<partial>lborel" (is "_ = ?I")
+      also have "... = \<integral>\<^sup>+ x. indicator (g-`A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x) \<partial>lborel" (is "_ = ?I")
         by (subst compl.IH, intro nn_integral_cong) (simp split: split_indicator)
       also have "g b - g a = LBINT x:{a..b}. g' x" using derivg'
         by (intro integral_FTC_atLeastAtMost[symmetric])
            (auto intro: continuous_on_subset[OF contg'] has_field_derivative_subset[OF derivg]
                  has_vector_derivative_at_within)
-      also have "ereal ... = \<integral>\<^sup>+ x. g' x * indicator {a..b} x \<partial>lborel"
+      also have "ennreal ... = \<integral>\<^sup>+ x. g' x * indicator {a..b} x \<partial>lborel"
         using borel_integrable_atLeastAtMost'[OF contg']
         by (subst nn_integral_eq_integral)
            (simp_all add: mult.commute derivg_nonneg split: split_indicator)
-      also have Mg'': "(\<lambda>x. indicator (g -` A \<inter> {a..b}) x * ereal (g' x * indicator {a..b} x))
+      also have Mg'': "(\<lambda>x. indicator (g -` A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x))
                             \<in> borel_measurable borel" using Mg'
-        by (intro borel_measurable_ereal_times borel_measurable_indicator)
+        by (intro borel_measurable_times_ennreal borel_measurable_indicator)
            (simp_all add: mult.commute)
-      have le: "(\<integral>\<^sup>+x. indicator (g-`A \<inter> {a..b}) x * ereal (g' x * indicator {a..b} x) \<partial>lborel) \<le>
-                        (\<integral>\<^sup>+x. ereal (g' x) * indicator {a..b} x \<partial>lborel)"
+      have le: "(\<integral>\<^sup>+x. indicator (g-`A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x) \<partial>lborel) \<le>
+                        (\<integral>\<^sup>+x. ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
          by (intro nn_integral_mono) (simp split: split_indicator add: derivg_nonneg)
       note integrable = borel_integrable_atLeastAtMost'[OF contg']
-      with le have notinf: "(\<integral>\<^sup>+x. indicator (g-`A \<inter> {a..b}) x * ereal (g' x * indicator {a..b} x) \<partial>lborel) \<noteq> \<infinity>"
-          by (auto simp: real_integrable_def nn_integral_set_ereal mult.commute)
-      have "(\<integral>\<^sup>+ x. g' x * indicator {a..b} x \<partial>lborel) - ?I = 
-                  \<integral>\<^sup>+ x. ereal (g' x * indicator {a..b} x) - 
-                        indicator (g -` A \<inter> {a..b}) x * ereal (g' x * indicator {a..b} x) \<partial>lborel"
+      with le have notinf: "(\<integral>\<^sup>+x. indicator (g-`A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x) \<partial>lborel) \<noteq> top"
+          by (auto simp: real_integrable_def nn_integral_set_ennreal mult.commute top_unique)
+      have "(\<integral>\<^sup>+ x. g' x * indicator {a..b} x \<partial>lborel) - ?I =
+                  \<integral>\<^sup>+ x. ennreal (g' x * indicator {a..b} x) -
+                        indicator (g -` A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x) \<partial>lborel"
         apply (intro nn_integral_diff[symmetric])
         apply (insert Mg', simp add: mult.commute) []
         apply (insert Mg'', simp) []
@@ -155,7 +155,7 @@
         apply (rule notinf)
         apply (simp split: split_indicator add: derivg_nonneg)
         done
-      also have "... = \<integral>\<^sup>+ x. indicator (-A) (g x) * ereal (g' x) * indicator {a..b} x \<partial>lborel"
+      also have "... = \<integral>\<^sup>+ x. indicator (-A) (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel"
         by (intro nn_integral_cong) (simp split: split_indicator)
       finally show ?case .
 
@@ -166,38 +166,38 @@
       have "g -` (\<Union>i. f i) \<inter> {a..b} = (\<Union>i. {a..b} \<inter> g -` f i)" by auto
       hence "g -` (\<Union>i. f i) \<inter> {a..b} \<in> sets borel" by (auto simp del: UN_simps)
 
-      have "(\<integral>\<^sup>+x. indicator (\<Union>i. f i) x * indicator {g a..g b} x \<partial>lborel) = 
+      have "(\<integral>\<^sup>+x. indicator (\<Union>i. f i) x * indicator {g a..g b} x \<partial>lborel) =
                 \<integral>\<^sup>+x. indicator (\<Union>i. {g a..g b} \<inter> f i) x \<partial>lborel"
           by (intro nn_integral_cong) (simp split: split_indicator)
       also from union have "... = emeasure lborel (\<Union>i. {g a..g b} \<inter> f i)" by simp
       also from union have "... = (\<Sum>i. emeasure lborel ({g a..g b} \<inter> f i))"
         by (intro suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def)
       also from union have "... = (\<Sum>i. \<integral>\<^sup>+x. indicator ({g a..g b} \<inter> f i) x \<partial>lborel)" by simp
-      also have "(\<lambda>i. \<integral>\<^sup>+x. indicator ({g a..g b} \<inter> f i) x \<partial>lborel) = 
+      also have "(\<lambda>i. \<integral>\<^sup>+x. indicator ({g a..g b} \<inter> f i) x \<partial>lborel) =
                            (\<lambda>i. \<integral>\<^sup>+x. indicator (f i) x * indicator {g a..g b} x \<partial>lborel)"
         by (intro ext nn_integral_cong) (simp split: split_indicator)
-      also from union.IH have "(\<Sum>i. \<integral>\<^sup>+x. indicator (f i) x * indicator {g a..g b} x \<partial>lborel) = 
-          (\<Sum>i. \<integral>\<^sup>+ x. indicator (f i) (g x) * ereal (g' x) * indicator {a..b} x \<partial>lborel)" by simp
-      also have "(\<lambda>i. \<integral>\<^sup>+ x. indicator (f i) (g x) * ereal (g' x) * indicator {a..b} x \<partial>lborel) =
-                         (\<lambda>i. \<integral>\<^sup>+ x. ereal (g' x * indicator {a..b} x) * indicator ({a..b} \<inter> g -` f i) x \<partial>lborel)"
+      also from union.IH have "(\<Sum>i. \<integral>\<^sup>+x. indicator (f i) x * indicator {g a..g b} x \<partial>lborel) =
+          (\<Sum>i. \<integral>\<^sup>+ x. indicator (f i) (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)" by simp
+      also have "(\<lambda>i. \<integral>\<^sup>+ x. indicator (f i) (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel) =
+                         (\<lambda>i. \<integral>\<^sup>+ x. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \<inter> g -` f i) x \<partial>lborel)"
         by (intro ext nn_integral_cong) (simp split: split_indicator)
-      also have "(\<Sum>i. ... i) = \<integral>\<^sup>+ x. (\<Sum>i. ereal (g' x * indicator {a..b} x) * indicator ({a..b} \<inter> g -` f i) x) \<partial>lborel"
+      also have "(\<Sum>i. ... i) = \<integral>\<^sup>+ x. (\<Sum>i. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \<inter> g -` f i) x) \<partial>lborel"
         using Mg'
         apply (intro nn_integral_suminf[symmetric])
-        apply (rule borel_measurable_ereal_times, simp add: borel_measurable_ereal mult.commute)
+        apply (rule borel_measurable_times_ennreal, simp add: mult.commute)
         apply (rule borel_measurable_indicator, subst sets_lborel)
         apply (simp_all split: split_indicator add: derivg_nonneg)
         done
-      also have "(\<lambda>x i. ereal (g' x * indicator {a..b} x) * indicator ({a..b} \<inter> g -` f i) x) =
-                      (\<lambda>x i. ereal (g' x * indicator {a..b} x) * indicator (g -` f i) x)"
+      also have "(\<lambda>x i. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \<inter> g -` f i) x) =
+                      (\<lambda>x i. ennreal (g' x * indicator {a..b} x) * indicator (g -` f i) x)"
         by (intro ext) (simp split: split_indicator)
-      also have "(\<integral>\<^sup>+ x. (\<Sum>i. ereal (g' x * indicator {a..b} x) * indicator (g -` f i) x) \<partial>lborel) =
-                     \<integral>\<^sup>+ x. ereal (g' x * indicator {a..b} x) * (\<Sum>i. indicator (g -` f i) x) \<partial>lborel"
-        by (intro nn_integral_cong suminf_cmult_ereal) (auto split: split_indicator simp: derivg_nonneg)
-      also from union have "(\<lambda>x. \<Sum>i. indicator (g -` f i) x :: ereal) = (\<lambda>x. indicator (\<Union>i. g -` f i) x)"
+      also have "(\<integral>\<^sup>+ x. (\<Sum>i. ennreal (g' x * indicator {a..b} x) * indicator (g -` f i) x) \<partial>lborel) =
+                     \<integral>\<^sup>+ x. ennreal (g' x * indicator {a..b} x) * (\<Sum>i. indicator (g -` f i) x) \<partial>lborel"
+        by (intro nn_integral_cong) (auto split: split_indicator simp: derivg_nonneg)
+      also from union have "(\<lambda>x. \<Sum>i. indicator (g -` f i) x :: ennreal) = (\<lambda>x. indicator (\<Union>i. g -` f i) x)"
         by (intro ext suminf_indicator) (auto simp: disjoint_family_on_def)
-      also have "(\<integral>\<^sup>+x. ereal (g' x * indicator {a..b} x) * ... x \<partial>lborel) =
-                    (\<integral>\<^sup>+x. indicator (\<Union>i. f i) (g x) * ereal (g' x) * indicator {a..b} x \<partial>lborel)"
+      also have "(\<integral>\<^sup>+x. ennreal (g' x * indicator {a..b} x) * ... x \<partial>lborel) =
+                    (\<integral>\<^sup>+x. indicator (\<Union>i. f i) (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
        by (intro nn_integral_cong) (simp split: split_indicator)
       finally show ?case .
   qed
@@ -206,42 +206,40 @@
   case (mult f c)
     note Mf[measurable] = \<open>f \<in> borel_measurable borel\<close>
     let ?I = "indicator {a..b}"
-    have "(\<lambda>x. f (g x * ?I x) * ereal (g' x * ?I x)) \<in> borel_measurable borel" using Mg Mg'
-      by (intro borel_measurable_ereal_times measurable_compose[OF _ Mf])
-         (simp_all add: borel_measurable_ereal mult.commute)
-    also have "(\<lambda>x. f (g x * ?I x) * ereal (g' x * ?I x)) = (\<lambda>x. f (g x) * ereal (g' x) * ?I x)"
+    have "(\<lambda>x. f (g x * ?I x) * ennreal (g' x * ?I x)) \<in> borel_measurable borel" using Mg Mg'
+      by (intro borel_measurable_times_ennreal measurable_compose[OF _ Mf])
+         (simp_all add: mult.commute)
+    also have "(\<lambda>x. f (g x * ?I x) * ennreal (g' x * ?I x)) = (\<lambda>x. f (g x) * ennreal (g' x) * ?I x)"
       by (intro ext) (simp split: split_indicator)
-    finally have Mf': "(\<lambda>x. f (g x) * ereal (g' x) * ?I x) \<in> borel_measurable borel" .
+    finally have Mf': "(\<lambda>x. f (g x) * ennreal (g' x) * ?I x) \<in> borel_measurable borel" .
     with mult show ?case
       by (subst (1 2 3) mult_ac, subst (1 2) nn_integral_cmult) (simp_all add: mult_ac)
- 
+
 next
   case (add f2 f1)
     let ?I = "indicator {a..b}"
     {
-      fix f :: "real \<Rightarrow> ereal" assume Mf: "f \<in> borel_measurable borel"
-      have "(\<lambda>x. f (g x * ?I x) * ereal (g' x * ?I x)) \<in> borel_measurable borel" using Mg Mg'
-        by (intro borel_measurable_ereal_times measurable_compose[OF _ Mf])
-           (simp_all add: borel_measurable_ereal mult.commute)
-      also have "(\<lambda>x. f (g x * ?I x) * ereal (g' x * ?I x)) = (\<lambda>x. f (g x) * ereal (g' x) * ?I x)"
+      fix f :: "real \<Rightarrow> ennreal" assume Mf: "f \<in> borel_measurable borel"
+      have "(\<lambda>x. f (g x * ?I x) * ennreal (g' x * ?I x)) \<in> borel_measurable borel" using Mg Mg'
+        by (intro borel_measurable_times_ennreal measurable_compose[OF _ Mf])
+           (simp_all add:  mult.commute)
+      also have "(\<lambda>x. f (g x * ?I x) * ennreal (g' x * ?I x)) = (\<lambda>x. f (g x) * ennreal (g' x) * ?I x)"
         by (intro ext) (simp split: split_indicator)
-      finally have "(\<lambda>x. f (g x) * ereal (g' x) * ?I x) \<in> borel_measurable borel" .
+      finally have "(\<lambda>x. f (g x) * ennreal (g' x) * ?I x) \<in> borel_measurable borel" .
     } note Mf' = this[OF \<open>f1 \<in> borel_measurable borel\<close>] this[OF \<open>f2 \<in> borel_measurable borel\<close>]
-    from add have not_neginf: "\<And>x. f1 x \<noteq> -\<infinity>" "\<And>x. f2 x \<noteq> -\<infinity>" 
-      by (metis Infty_neq_0(1) ereal_0_le_uminus_iff ereal_infty_less_eq(1))+
 
     have "(\<integral>\<^sup>+ x. (f1 x + f2 x) * indicator {g a..g b} x \<partial>lborel) =
              (\<integral>\<^sup>+ x. f1 x * indicator {g a..g b} x + f2 x * indicator {g a..g b} x \<partial>lborel)"
       by (intro nn_integral_cong) (simp split: split_indicator)
-    also from add have "... = (\<integral>\<^sup>+ x. f1 (g x) * ereal (g' x) * indicator {a..b} x \<partial>lborel) +
-                                (\<integral>\<^sup>+ x. f2 (g x) * ereal (g' x) * indicator {a..b} x \<partial>lborel)"
+    also from add have "... = (\<integral>\<^sup>+ x. f1 (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel) +
+                                (\<integral>\<^sup>+ x. f2 (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
       by (simp_all add: nn_integral_add)
-    also from add have "... = (\<integral>\<^sup>+ x. f1 (g x) * ereal (g' x) * indicator {a..b} x + 
-                                      f2 (g x) * ereal (g' x) * indicator {a..b} x \<partial>lborel)"
+    also from add have "... = (\<integral>\<^sup>+ x. f1 (g x) * ennreal (g' x) * indicator {a..b} x +
+                                      f2 (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
       by (intro nn_integral_add[symmetric])
          (auto simp add: Mf' derivg_nonneg split: split_indicator)
-    also from not_neginf have "... = \<integral>\<^sup>+ x. (f1 (g x) + f2 (g x)) * ereal (g' x) * indicator {a..b} x \<partial>lborel"
-      by (intro nn_integral_cong) (simp split: split_indicator add: ereal_distrib)
+    also have "... = \<integral>\<^sup>+ x. (f1 (g x) + f2 (g x)) * ennreal (g' x) * indicator {a..b} x \<partial>lborel"
+      by (intro nn_integral_cong) (simp split: split_indicator add: distrib_right)
     finally show ?case .
 
 next
@@ -249,28 +247,28 @@
   {
     fix i
     let ?I = "indicator {a..b}"
-    have "(\<lambda>x. F i (g x * ?I x) * ereal (g' x * ?I x)) \<in> borel_measurable borel" using Mg Mg'
-      by (rule_tac borel_measurable_ereal_times, rule_tac measurable_compose[OF _ sup.hyps(1)])
+    have "(\<lambda>x. F i (g x * ?I x) * ennreal (g' x * ?I x)) \<in> borel_measurable borel" using Mg Mg'
+      by (rule_tac borel_measurable_times_ennreal, rule_tac measurable_compose[OF _ sup.hyps(1)])
          (simp_all add: mult.commute)
-    also have "(\<lambda>x. F i (g x * ?I x) * ereal (g' x * ?I x)) = (\<lambda>x. F i (g x) * ereal (g' x) * ?I x)"
+    also have "(\<lambda>x. F i (g x * ?I x) * ennreal (g' x * ?I x)) = (\<lambda>x. F i (g x) * ennreal (g' x) * ?I x)"
       by (intro ext) (simp split: split_indicator)
      finally have "... \<in> borel_measurable borel" .
   } note Mf' = this
 
-    have "(\<integral>\<^sup>+x. (SUP i. F i x) * indicator {g a..g b} x \<partial>lborel) = 
+    have "(\<integral>\<^sup>+x. (SUP i. F i x) * indicator {g a..g b} x \<partial>lborel) =
                \<integral>\<^sup>+x. (SUP i. F i x* indicator {g a..g b} x) \<partial>lborel"
       by (intro nn_integral_cong) (simp split: split_indicator)
     also from sup have "... = (SUP i. \<integral>\<^sup>+x. F i x* indicator {g a..g b} x \<partial>lborel)"
       by (intro nn_integral_monotone_convergence_SUP)
          (auto simp: incseq_def le_fun_def split: split_indicator)
-    also from sup have "... = (SUP i. \<integral>\<^sup>+x. F i (g x) * ereal (g' x) * indicator {a..b} x \<partial>lborel)"
+    also from sup have "... = (SUP i. \<integral>\<^sup>+x. F i (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
       by simp
-    also from sup have "... =  \<integral>\<^sup>+x. (SUP i. F i (g x) * ereal (g' x) * indicator {a..b} x) \<partial>lborel"
+    also from sup have "... =  \<integral>\<^sup>+x. (SUP i. F i (g x) * ennreal (g' x) * indicator {a..b} x) \<partial>lborel"
       by (intro nn_integral_monotone_convergence_SUP[symmetric])
          (auto simp: incseq_def le_fun_def derivg_nonneg Mf' split: split_indicator
-               intro!: ereal_mult_right_mono)
-    also from sup have "... = \<integral>\<^sup>+x. (SUP i. F i (g x)) * ereal (g' x) * indicator {a..b} x \<partial>lborel"
-      by (subst mult.assoc, subst mult.commute, subst SUP_ereal_mult_left)
+               intro!: mult_right_mono)
+    also from sup have "... = \<integral>\<^sup>+x. (SUP i. F i (g x)) * ennreal (g' x) * indicator {a..b} x \<partial>lborel"
+      by (subst mult.assoc, subst mult.commute, subst SUP_mult_left_ennreal)
          (auto split: split_indicator simp: derivg_nonneg mult_ac)
     finally show ?case by simp
   qed
@@ -280,119 +278,110 @@
   fixes f :: "real \<Rightarrow> real"
   assumes Mf[measurable]: "set_borel_measurable borel {g a..g b} f"
   assumes derivg: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
-  assumes contg': "continuous_on {a..b} g'" 
+  assumes contg': "continuous_on {a..b} g'"
   assumes derivg_nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
   assumes "a \<le> b"
-  shows "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) = 
+  shows "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) =
              (\<integral>\<^sup>+x. f (g x) * g' x * indicator {a..b} x \<partial>lborel)"
 proof (cases "a = b")
   assume "a \<noteq> b"
   with \<open>a \<le> b\<close> have "a < b" by auto
-  let ?f' = "\<lambda>x. max 0 (f x * indicator {g a..g b} x)"
+  let ?f' = "\<lambda>x. f x * indicator {g a..g b} x"
 
   from derivg derivg_nonneg have monog: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
     by (rule deriv_nonneg_imp_mono) simp_all
   have bounds: "\<And>x. x \<ge> a \<Longrightarrow> x \<le> b \<Longrightarrow> g x \<ge> g a" "\<And>x. x \<ge> a \<Longrightarrow> x \<le> b \<Longrightarrow> g x \<le> g b"
     by (auto intro: monog)
 
-  from derivg_nonneg have nonneg: 
-    "\<And>f x. x \<ge> a \<Longrightarrow> x \<le> b \<Longrightarrow> g' x \<noteq> 0 \<Longrightarrow> f x * ereal (g' x) \<ge> 0 \<Longrightarrow> f x \<ge> 0"
-    by (force simp: ereal_zero_le_0_iff field_simps)
+  from derivg_nonneg have nonneg:
+    "\<And>f x. x \<ge> a \<Longrightarrow> x \<le> b \<Longrightarrow> g' x \<noteq> 0 \<Longrightarrow> f x * ennreal (g' x) \<ge> 0 \<Longrightarrow> f x \<ge> 0"
+    by (force simp: field_simps)
   have nonneg': "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<not> 0 \<le> f (g x) \<Longrightarrow> 0 \<le> f (g x) * g' x \<Longrightarrow> g' x = 0"
     by (metis atLeastAtMost_iff derivg_nonneg eq_iff mult_eq_0_iff mult_le_0_iff)
 
-  have "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) = 
-            (\<integral>\<^sup>+x. ereal (?f' x) * indicator {g a..g b} x \<partial>lborel)"
-    by (subst nn_integral_max_0[symmetric], intro nn_integral_cong) 
-       (auto split: split_indicator simp: zero_ereal_def)
-  also have "... = \<integral>\<^sup>+ x. ?f' (g x) * ereal (g' x) * indicator {a..b} x \<partial>lborel" using Mf
-    by (subst nn_integral_substitution_aux[OF _ _ derivg contg' derivg_nonneg \<open>a < b\<close>]) 
-       (auto simp add: zero_ereal_def mult.commute)
-  also have "... = \<integral>\<^sup>+ x. max 0 (f (g x)) * ereal (g' x) * indicator {a..b} x \<partial>lborel"
-    by (intro nn_integral_cong) 
-       (auto split: split_indicator simp: max_def dest: bounds)
-  also have "... = \<integral>\<^sup>+ x. max 0 (f (g x) * ereal (g' x) * indicator {a..b} x) \<partial>lborel"
+  have "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) =
+            (\<integral>\<^sup>+x. ennreal (?f' x) * indicator {g a..g b} x \<partial>lborel)"
     by (intro nn_integral_cong)
-       (auto simp: max_def derivg_nonneg split: split_indicator intro!: nonneg')
-  also have "... = \<integral>\<^sup>+ x. f (g x) * ereal (g' x) * indicator {a..b} x \<partial>lborel"
-    by (rule nn_integral_max_0)
-  also have "... = \<integral>\<^sup>+x. ereal (f (g x) * g' x * indicator {a..b} x) \<partial>lborel"
-    by (intro nn_integral_cong) (simp split: split_indicator)
+       (auto split: split_indicator split_max simp: zero_ennreal.rep_eq ennreal_neg)
+  also have "... = \<integral>\<^sup>+ x. ?f' (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel" using Mf
+    by (subst nn_integral_substitution_aux[OF _ _ derivg contg' derivg_nonneg \<open>a < b\<close>])
+       (auto simp add: mult.commute)
+  also have "... = \<integral>\<^sup>+ x. f (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel"
+    by (intro nn_integral_cong) (auto split: split_indicator simp: max_def dest: bounds)
+  also have "... = \<integral>\<^sup>+x. ennreal (f (g x) * g' x * indicator {a..b} x) \<partial>lborel"
+    by (intro nn_integral_cong) (auto simp: mult.commute derivg_nonneg ennreal_mult' split: split_indicator)
   finally show ?thesis .
 qed auto
 
 lemma integral_substitution:
   assumes integrable: "set_integrable lborel {g a..g b} f"
   assumes derivg: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
-  assumes contg': "continuous_on {a..b} g'" 
+  assumes contg': "continuous_on {a..b} g'"
   assumes derivg_nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
   assumes "a \<le> b"
   shows "set_integrable lborel {a..b} (\<lambda>x. f (g x) * g' x)"
     and "(LBINT x. f x * indicator {g a..g b} x) = (LBINT x. f (g x) * g' x * indicator {a..b} x)"
 proof-
   from derivg have contg: "continuous_on {a..b} g" by (rule has_real_derivative_imp_continuous_on)
-  from this and contg' have Mg: "set_borel_measurable borel {a..b} g" and 
-                             Mg': "set_borel_measurable borel {a..b} g'" 
+  from this and contg' have Mg: "set_borel_measurable borel {a..b} g" and
+                             Mg': "set_borel_measurable borel {a..b} g'"
       by (simp_all only: set_measurable_continuous_on_ivl)
   from derivg derivg_nonneg have monog: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
     by (rule deriv_nonneg_imp_mono) simp_all
 
-  have "(\<lambda>x. ereal (f x) * indicator {g a..g b} x) = 
-           (\<lambda>x. ereal (f x * indicator {g a..g b} x))" 
+  have "(\<lambda>x. ennreal (f x) * indicator {g a..g b} x) =
+           (\<lambda>x. ennreal (f x * indicator {g a..g b} x))"
     by (intro ext) (simp split: split_indicator)
   with integrable have M1: "(\<lambda>x. f x * indicator {g a..g b} x) \<in> borel_measurable borel"
     unfolding real_integrable_def by (force simp: mult.commute)
-  have "(\<lambda>x. ereal (-f x) * indicator {g a..g b} x) = 
-           (\<lambda>x. -ereal (f x * indicator {g a..g b} x))" 
-    by (intro ext) (simp split: split_indicator)
-  with integrable have M2: "(\<lambda>x. -f x * indicator {g a..g b} x) \<in> borel_measurable borel"
+  from integrable have M2: "(\<lambda>x. -f x * indicator {g a..g b} x) \<in> borel_measurable borel"
     unfolding real_integrable_def by (force simp: mult.commute)
 
-  have "LBINT x. (f x :: real) * indicator {g a..g b} x = 
-          real_of_ereal (\<integral>\<^sup>+ x. ereal (f x) * indicator {g a..g b} x \<partial>lborel) -
-          real_of_ereal (\<integral>\<^sup>+ x. ereal (- (f x)) * indicator {g a..g b} x \<partial>lborel)" using integrable
-    by (subst real_lebesgue_integral_def) (simp_all add: nn_integral_set_ereal mult.commute)
-  also have "(\<integral>\<^sup>+x. ereal (f x) * indicator {g a..g b} x \<partial>lborel) =
-               (\<integral>\<^sup>+x. ereal (f x * indicator {g a..g b} x) \<partial>lborel)"
+  have "LBINT x. (f x :: real) * indicator {g a..g b} x =
+          enn2real (\<integral>\<^sup>+ x. ennreal (f x) * indicator {g a..g b} x \<partial>lborel) -
+          enn2real (\<integral>\<^sup>+ x. ennreal (- (f x)) * indicator {g a..g b} x \<partial>lborel)" using integrable
+    by (subst real_lebesgue_integral_def) (simp_all add: nn_integral_set_ennreal mult.commute)
+  also have "(\<integral>\<^sup>+x. ennreal (f x) * indicator {g a..g b} x \<partial>lborel) =
+               (\<integral>\<^sup>+x. ennreal (f x * indicator {g a..g b} x) \<partial>lborel)"
     by (intro nn_integral_cong) (simp split: split_indicator)
-  also with M1 have A: "(\<integral>\<^sup>+ x. ereal (f x * indicator {g a..g b} x) \<partial>lborel) =
-                            (\<integral>\<^sup>+ x. ereal (f (g x) * g' x * indicator {a..b} x) \<partial>lborel)"
-    by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \<open>a \<le> b\<close>]) 
-       (auto simp: nn_integral_set_ereal mult.commute)
-  also have "(\<integral>\<^sup>+ x. ereal (- (f x)) * indicator {g a..g b} x \<partial>lborel) =
-               (\<integral>\<^sup>+ x. ereal (- (f x) * indicator {g a..g b} x) \<partial>lborel)"
+  also with M1 have A: "(\<integral>\<^sup>+ x. ennreal (f x * indicator {g a..g b} x) \<partial>lborel) =
+                            (\<integral>\<^sup>+ x. ennreal (f (g x) * g' x * indicator {a..b} x) \<partial>lborel)"
+    by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \<open>a \<le> b\<close>])
+       (auto simp: nn_integral_set_ennreal mult.commute)
+  also have "(\<integral>\<^sup>+ x. ennreal (- (f x)) * indicator {g a..g b} x \<partial>lborel) =
+               (\<integral>\<^sup>+ x. ennreal (- (f x) * indicator {g a..g b} x) \<partial>lborel)"
     by (intro nn_integral_cong) (simp split: split_indicator)
-  also with M2 have B: "(\<integral>\<^sup>+ x. ereal (- (f x) * indicator {g a..g b} x) \<partial>lborel) =
-                            (\<integral>\<^sup>+ x. ereal (- (f (g x)) * g' x * indicator {a..b} x) \<partial>lborel)"
+  also with M2 have B: "(\<integral>\<^sup>+ x. ennreal (- (f x) * indicator {g a..g b} x) \<partial>lborel) =
+                            (\<integral>\<^sup>+ x. ennreal (- (f (g x)) * g' x * indicator {a..b} x) \<partial>lborel)"
     by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \<open>a \<le> b\<close>])
-       (auto simp: nn_integral_set_ereal mult.commute)
+       (auto simp: nn_integral_set_ennreal mult.commute)
 
   also {
-    from integrable have Mf: "set_borel_measurable borel {g a..g b} f" 
+    from integrable have Mf: "set_borel_measurable borel {g a..g b} f"
       unfolding real_integrable_def by simp
     from borel_measurable_times[OF measurable_compose[OF Mg Mf] Mg']
       have "(\<lambda>x. f (g x * indicator {a..b} x) * indicator {g a..g b} (g x * indicator {a..b} x) *
-                     (g' x * indicator {a..b} x)) \<in> borel_measurable borel"  (is "?f \<in> _") 
+                     (g' x * indicator {a..b} x)) \<in> borel_measurable borel"  (is "?f \<in> _")
       by (simp add: mult.commute)
     also have "?f = (\<lambda>x. f (g x) * g' x * indicator {a..b} x)"
       using monog by (intro ext) (auto split: split_indicator)
     finally show "set_integrable lborel {a..b} (\<lambda>x. f (g x) * g' x)"
-      using A B integrable unfolding real_integrable_def 
-      by (simp_all add: nn_integral_set_ereal mult.commute)
+      using A B integrable unfolding real_integrable_def
+      by (simp_all add: nn_integral_set_ennreal mult.commute)
   } note integrable' = this
 
-  have "real_of_ereal (\<integral>\<^sup>+ x. ereal (f (g x) * g' x * indicator {a..b} x) \<partial>lborel) -
-                  real_of_ereal (\<integral>\<^sup>+ x. ereal (-f (g x) * g' x * indicator {a..b} x) \<partial>lborel) =
+  have "enn2real (\<integral>\<^sup>+ x. ennreal (f (g x) * g' x * indicator {a..b} x) \<partial>lborel) -
+                  enn2real (\<integral>\<^sup>+ x. ennreal (-f (g x) * g' x * indicator {a..b} x) \<partial>lborel) =
                 (LBINT x. f (g x) * g' x * indicator {a..b} x)" using integrable'
     by (subst real_lebesgue_integral_def) (simp_all add: field_simps)
-  finally show "(LBINT x. f x * indicator {g a..g b} x) = 
+  finally show "(LBINT x. f x * indicator {g a..g b} x) =
                      (LBINT x. f (g x) * g' x * indicator {a..b} x)" .
 qed
 
 lemma interval_integral_substitution:
   assumes integrable: "set_integrable lborel {g a..g b} f"
   assumes derivg: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
-  assumes contg': "continuous_on {a..b} g'" 
+  assumes contg': "continuous_on {a..b} g'"
   assumes derivg_nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
   assumes "a \<le> b"
   shows "set_integrable lborel {a..b} (\<lambda>x. f (g x) * g' x)"
--- a/src/HOL/Probability/Lebesgue_Measure.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Lebesgue_Measure.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -14,7 +14,7 @@
 subsection \<open>Every right continuous and nondecreasing function gives rise to a measure\<close>
 
 definition interval_measure :: "(real \<Rightarrow> real) \<Rightarrow> real measure" where
-  "interval_measure F = extend_measure UNIV {(a, b). a \<le> b} (\<lambda>(a, b). {a <.. b}) (\<lambda>(a, b). ereal (F b - F a))"
+  "interval_measure F = extend_measure UNIV {(a, b). a \<le> b} (\<lambda>(a, b). {a <.. b}) (\<lambda>(a, b). ennreal (F b - F a))"
 
 lemma emeasure_interval_measure_Ioc:
   assumes "a \<le> b"
@@ -40,7 +40,6 @@
       thus ?thesis ..
     qed
   qed (auto simp: Ioc_inj, metis linear)
-
 next
   fix l r :: "nat \<Rightarrow> real" and a b :: real
   assume l_r[simp]: "\<And>n. l n \<le> r n" and "a \<le> b" and disj: "disjoint_family (\<lambda>n. {l n<..r n})"
@@ -168,13 +167,14 @@
   note claim2 = this
 
   (* now prove the inequality going the other way *)
-
-  { fix epsilon :: real assume egt0: "epsilon > 0"
-    have "\<forall>i. \<exists>d. d > 0 &  F (r i + d) < F (r i) + epsilon / 2^(i+2)"
+  have "ennreal (F b - F a) \<le> (\<Sum>i. ennreal (F (r i) - F (l i)))"
+  proof (rule ennreal_le_epsilon)
+    fix epsilon :: real assume egt0: "epsilon > 0"
+    have "\<forall>i. \<exists>d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)"
     proof
       fix i
       note right_cont_F [of "r i"]
-      thus "\<exists>d. d > 0 \<and> F (r i + d) < F (r i) + epsilon / 2^(i+2)"
+      thus "\<exists>d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)"
         apply -
         apply (subst (asm) continuous_at_right_real_increasing)
         apply (rule mono_F, assumption)
@@ -264,22 +264,15 @@
       by auto
     also have "... \<le> (\<Sum>i\<le>n. F (r i) - F (l i)) + epsilon"
       using finS Sbound Sprop by (auto intro!: add_right_mono setsum_mono3)
-    finally have "ereal (F b - F a) \<le> (\<Sum>i\<le>n. ereal (F (r i) - F (l i))) + epsilon"
-      by simp
-    then have "ereal (F b - F a) \<le> (\<Sum>i. ereal (F (r i) - F (l i))) + (epsilon :: real)"
-      apply (rule_tac order_trans)
-      prefer 2
-      apply (rule add_mono[where c="ereal epsilon"])
-      apply (rule suminf_upper[of _ "Suc n"])
-      apply (auto simp add: lessThan_Suc_atMost)
-      done }
-  hence "ereal (F b - F a) \<le> (\<Sum>i. ereal (F (r i) - F (l i)))"
-    by (auto intro: ereal_le_epsilon2)
-  moreover
-  have "(\<Sum>i. ereal (F (r i) - F (l i))) \<le> ereal (F b - F a)"
-    by (auto simp add: claim1 intro!: suminf_bound)
-  ultimately show "(\<Sum>n. ereal (F (r n) - F (l n))) = ereal (F b - F a)"
-    by simp
+    finally have "ennreal (F b - F a) \<le> (\<Sum>i\<le>n. ennreal (F (r i) - F (l i))) + epsilon"
+      using egt0 by (simp add: ennreal_plus[symmetric] setsum_nonneg del: ennreal_plus)
+    then show "ennreal (F b - F a) \<le> (\<Sum>i. ennreal (F (r i) - F (l i))) + (epsilon :: real)"
+      by (rule order_trans) (auto intro!: add_mono setsum_le_suminf simp del: setsum_ennreal)
+  qed
+  moreover have "(\<Sum>i. ennreal (F (r i) - F (l i))) \<le> ennreal (F b - F a)"
+    using \<open>a \<le> b\<close> by (auto intro!: suminf_le_const ennreal_le_iff[THEN iffD2] claim1)
+  ultimately show "(\<Sum>n. ennreal (F (r n) - F (l n))) = ennreal (F b - F a)"
+    by (rule antisym[rotated])
 qed (auto simp: Ioc_inj mono_F)
 
 lemma measure_interval_measure_Ioc:
@@ -290,7 +283,7 @@
   unfolding measure_def
   apply (subst emeasure_interval_measure_Ioc)
   apply fact+
-  apply simp
+  apply (simp add: assms)
   done
 
 lemma emeasure_interval_measure_Ioc_eq:
@@ -344,10 +337,9 @@
       using \<open>\<And>n. X n < a\<close> \<open>a \<le> b\<close> by (subst *) (auto intro: less_imp_le less_le_trans)
     finally show "(\<lambda>n. F b - F (X n)) \<longlonglongrightarrow> emeasure ?F {a..b}" .
   qed
-  show "((\<lambda>a. ereal (F b - F a)) \<longlongrightarrow> F b - F a) (at_left a)"
-    using cont_F
-    by (intro lim_ereal[THEN iffD2] tendsto_intros )
-       (auto simp: continuous_on_def intro: tendsto_within_subset)
+  show "((\<lambda>a. ennreal (F b - F a)) \<longlongrightarrow> F b - F a) (at_left a)"
+    by (rule continuous_on_tendsto_compose[where g="\<lambda>x. x" and s=UNIV])
+       (auto simp: continuous_on_ennreal continuous_on_diff cont_F continuous_on_const)
 qed (rule trivial_limit_at_left_real)
 
 lemma sigma_finite_interval_measure:
@@ -406,29 +398,29 @@
     by (simp add: lborel_def emeasure_distr emeasure_PiM emeasure_interval_measure_Icc continuous_on_id)
 qed
 
-lemma emeasure_lborel_Icc_eq: "emeasure lborel {l .. u} = ereal (if l \<le> u then u - l else 0)"
+lemma emeasure_lborel_Icc_eq: "emeasure lborel {l .. u} = ennreal (if l \<le> u then u - l else 0)"
   by simp
 
 lemma emeasure_lborel_cbox[simp]:
   assumes [simp]: "\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
   shows "emeasure lborel (cbox l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
 proof -
-  have "(\<lambda>x. \<Prod>b\<in>Basis. indicator {l\<bullet>b .. u\<bullet>b} (x \<bullet> b) :: ereal) = indicator (cbox l u)"
-    by (auto simp: fun_eq_iff cbox_def setprod_ereal_0 split: split_indicator)
+  have "(\<lambda>x. \<Prod>b\<in>Basis. indicator {l\<bullet>b .. u\<bullet>b} (x \<bullet> b) :: ennreal) = indicator (cbox l u)"
+    by (auto simp: fun_eq_iff cbox_def split: split_indicator)
   then have "emeasure lborel (cbox l u) = (\<integral>\<^sup>+x. (\<Prod>b\<in>Basis. indicator {l\<bullet>b .. u\<bullet>b} (x \<bullet> b)) \<partial>lborel)"
     by simp
   also have "\<dots> = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
-    by (subst nn_integral_lborel_setprod) (simp_all add: setprod_ereal inner_diff_left)
+    by (subst nn_integral_lborel_setprod) (simp_all add: setprod_ennreal inner_diff_left)
   finally show ?thesis .
 qed
 
 lemma AE_lborel_singleton: "AE x in lborel::'a::euclidean_space measure. x \<noteq> c"
-  using SOME_Basis AE_discrete_difference [of "{c}" lborel]
-    emeasure_lborel_cbox [of c c] by (auto simp add: cbox_sing)
+  using SOME_Basis AE_discrete_difference [of "{c}" lborel] emeasure_lborel_cbox [of c c]
+  by (auto simp add: cbox_sing setprod_constant power_0_left)
 
 lemma emeasure_lborel_Ioo[simp]:
   assumes [simp]: "l \<le> u"
-  shows "emeasure lborel {l <..< u} = ereal (u - l)"
+  shows "emeasure lborel {l <..< u} = ennreal (u - l)"
 proof -
   have "emeasure lborel {l <..< u} = emeasure lborel {l .. u}"
     using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
@@ -438,7 +430,7 @@
 
 lemma emeasure_lborel_Ioc[simp]:
   assumes [simp]: "l \<le> u"
-  shows "emeasure lborel {l <.. u} = ereal (u - l)"
+  shows "emeasure lborel {l <.. u} = ennreal (u - l)"
 proof -
   have "emeasure lborel {l <.. u} = emeasure lborel {l .. u}"
     using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
@@ -448,7 +440,7 @@
 
 lemma emeasure_lborel_Ico[simp]:
   assumes [simp]: "l \<le> u"
-  shows "emeasure lborel {l ..< u} = ereal (u - l)"
+  shows "emeasure lborel {l ..< u} = ennreal (u - l)"
 proof -
   have "emeasure lborel {l ..< u} = emeasure lborel {l .. u}"
     using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
@@ -460,12 +452,12 @@
   assumes [simp]: "\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
   shows "emeasure lborel (box l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
 proof -
-  have "(\<lambda>x. \<Prod>b\<in>Basis. indicator {l\<bullet>b <..< u\<bullet>b} (x \<bullet> b) :: ereal) = indicator (box l u)"
-    by (auto simp: fun_eq_iff box_def setprod_ereal_0 split: split_indicator)
+  have "(\<lambda>x. \<Prod>b\<in>Basis. indicator {l\<bullet>b <..< u\<bullet>b} (x \<bullet> b) :: ennreal) = indicator (box l u)"
+    by (auto simp: fun_eq_iff box_def split: split_indicator)
   then have "emeasure lborel (box l u) = (\<integral>\<^sup>+x. (\<Prod>b\<in>Basis. indicator {l\<bullet>b <..< u\<bullet>b} (x \<bullet> b)) \<partial>lborel)"
     by simp
   also have "\<dots> = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
-    by (subst nn_integral_lborel_setprod) (simp_all add: setprod_ereal inner_diff_left)
+    by (subst nn_integral_lborel_setprod) (simp_all add: setprod_ennreal inner_diff_left)
   finally show ?thesis .
 qed
 
@@ -490,7 +482,7 @@
   assumes [simp]: "\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
   shows measure_lborel_box[simp]: "measure lborel (box l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
     and measure_lborel_cbox[simp]: "measure lborel (cbox l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
-  by (simp_all add: measure_def)
+  by (simp_all add: measure_def inner_diff_left setprod_nonneg)
 
 lemma sigma_finite_lborel: "sigma_finite_measure lborel"
 proof
@@ -519,13 +511,14 @@
     unfolding UN_box_eq_UNIV[symmetric]
     apply (subst SUP_emeasure_incseq[symmetric])
     apply (auto simp: incseq_def subset_box inner_add_left setprod_constant
-               intro!: SUP_PInfty)
+      simp del: Sup_eq_top_iff SUP_eq_top_iff
+      intro!: ennreal_SUP_eq_top)
     done
 qed
 
 lemma emeasure_lborel_singleton[simp]: "emeasure lborel {x} = 0"
   using emeasure_lborel_cbox[of x x] nonempty_Basis
-  by (auto simp del: emeasure_lborel_cbox nonempty_Basis simp add: cbox_sing)
+  by (auto simp del: emeasure_lborel_cbox nonempty_Basis simp add: cbox_sing setprod_constant)
 
 lemma emeasure_lborel_countable:
   fixes A :: "'a::euclidean_space set"
@@ -537,7 +530,7 @@
     by (rule emeasure_UN_eq_0) auto
   ultimately have "emeasure lborel A \<le> 0" using emeasure_mono
     by (smt UN_E emeasure_empty equalityI from_nat_into order_refl singletonD subsetI)
-  thus ?thesis by (auto simp add: emeasure_le_0_iff)
+  thus ?thesis by (auto simp add: )
 qed
 
 lemma countable_imp_null_set_lborel: "countable A \<Longrightarrow> A \<in> null_sets lborel"
@@ -597,33 +590,29 @@
       by (auto simp: field_simps box_def inner_simps)
     with \<open>0 < c\<close> show ?thesis
       using le
-      by (auto simp: field_simps emeasure_density nn_integral_distr nn_integral_cmult
-                     emeasure_lborel_box_eq inner_simps setprod_dividef setprod_constant
-                     borel_measurable_indicator' emeasure_distr)
+      by (auto simp: field_simps inner_simps setprod_dividef setprod_constant setprod_nonneg
+                     ennreal_mult[symmetric] emeasure_density nn_integral_distr emeasure_distr
+                     nn_integral_cmult emeasure_lborel_box_eq borel_measurable_indicator')
   next
     assume "\<not> 0 < c" with \<open>c \<noteq> 0\<close> have "c < 0" by auto
     then have "box ((u - t) /\<^sub>R c) ((l - t) /\<^sub>R c) = (\<lambda>x. t + c *\<^sub>R x) -` box l u"
       by (auto simp: field_simps box_def inner_simps)
-    then have "\<And>x. indicator (box l u) (t + c *\<^sub>R x) = (indicator (box ((u - t) /\<^sub>R c) ((l - t) /\<^sub>R c)) x :: ereal)"
+    then have *: "\<And>x. indicator (box l u) (t + c *\<^sub>R x) = (indicator (box ((u - t) /\<^sub>R c) ((l - t) /\<^sub>R c)) x :: ennreal)"
       by (auto split: split_indicator)
-    moreover
-    { have "(- c) ^ card ?B * (\<Prod>x\<in>?B. l \<bullet> x - u \<bullet> x) =
-         (-1 * c) ^ card ?B * (\<Prod>x\<in>?B. -1 * (u \<bullet> x - l \<bullet> x))"
-         by simp
-      also have "\<dots> = (-1 * -1)^card ?B * c ^ card ?B * (\<Prod>x\<in>?B. u \<bullet> x - l \<bullet> x)"
-        unfolding setprod.distrib power_mult_distrib by (simp add: setprod_constant)
-      finally have "(- c) ^ card ?B * (\<Prod>x\<in>?B. l \<bullet> x - u \<bullet> x) = c ^ card ?B * (\<Prod>b\<in>?B. u \<bullet> b - l \<bullet> b)"
-        by simp }
-    ultimately show ?thesis
+    have **: "(\<Prod>x\<in>Basis. (l \<bullet> x - u \<bullet> x) / c) = (\<Prod>x\<in>Basis. u \<bullet> x - l \<bullet> x) / (-c) ^ card (Basis::'a set)"
+      using \<open>c < 0\<close>
+      by (auto simp add: field_simps setprod_dividef[symmetric] setprod_constant[symmetric]
+               intro!: setprod.cong)
+    show ?thesis
       using \<open>c < 0\<close> le
-      by (auto simp: field_simps emeasure_density nn_integral_distr nn_integral_cmult
-                     emeasure_lborel_box_eq inner_simps setprod_dividef setprod_constant
-                     borel_measurable_indicator' emeasure_distr)
+      by (auto simp: * ** field_simps emeasure_density nn_integral_distr nn_integral_cmult
+                     emeasure_lborel_box_eq inner_simps setprod_nonneg ennreal_mult[symmetric]
+                     borel_measurable_indicator')
   qed
 qed simp
 
 lemma lborel_real_affine:
-  "c \<noteq> 0 \<Longrightarrow> lborel = density (distr lborel borel (\<lambda>x. t + c * x)) (\<lambda>_. \<bar>c\<bar>)"
+  "c \<noteq> 0 \<Longrightarrow> lborel = density (distr lborel borel (\<lambda>x. t + c * x)) (\<lambda>_. ennreal (abs c))"
   using lborel_affine[of c t] by simp
 
 lemma AE_borel_affine:
@@ -643,7 +632,7 @@
   assumes f: "integrable lborel f"
   shows "c \<noteq> 0 \<Longrightarrow> integrable lborel (\<lambda>x. f (t + c * x))"
   using f f[THEN borel_measurable_integrable] unfolding integrable_iff_bounded
-  by (subst (asm) nn_integral_real_affine[where c=c and t=t]) auto
+  by (subst (asm) nn_integral_real_affine[where c=c and t=t]) (auto simp: ennreal_mult_less_top)
 
 lemma lborel_integrable_real_affine_iff:
   fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
@@ -681,7 +670,7 @@
 
 lemma lborel_distr_uminus: "distr lborel borel uminus = (lborel :: real measure)"
   by (subst lborel_real_affine[of "-1" 0])
-     (auto simp: density_1 one_ereal_def[symmetric])
+     (auto simp: density_1 one_ennreal_def[symmetric])
 
 lemma lborel_distr_mult:
   assumes "(c::real) \<noteq> 0"
@@ -698,16 +687,16 @@
   shows "lborel = density (distr lborel borel (op * c)) (\<lambda>_. \<bar>c\<bar>)"
 proof-
   have "lborel = density lborel (\<lambda>_. 1)" by (rule density_1[symmetric])
-  also from assms have "(\<lambda>_. 1 :: ereal) = (\<lambda>_. inverse \<bar>c\<bar> * \<bar>c\<bar>)" by (intro ext) simp
+  also from assms have "(\<lambda>_. 1 :: ennreal) = (\<lambda>_. inverse \<bar>c\<bar> * \<bar>c\<bar>)" by (intro ext) simp
   also have "density lborel ... = density (density lborel (\<lambda>_. inverse \<bar>c\<bar>)) (\<lambda>_. \<bar>c\<bar>)"
-    by (subst density_density_eq) auto
+    by (subst density_density_eq) (auto simp: ennreal_mult)
   also from assms have "density lborel (\<lambda>_. inverse \<bar>c\<bar>) = distr lborel borel (op * c)"
     by (rule lborel_distr_mult[symmetric])
   finally show ?thesis .
 qed
 
 lemma lborel_distr_plus: "distr lborel borel (op + c) = (lborel :: real measure)"
-  by (subst lborel_real_affine[of 1 c]) (auto simp: density_1 one_ereal_def[symmetric])
+  by (subst lborel_real_affine[of 1 c]) (auto simp: density_1 one_ennreal_def[symmetric])
 
 interpretation lborel: sigma_finite_measure lborel
   by (rule sigma_finite_lborel)
@@ -727,9 +716,9 @@
     "box (la, lb) (ua, ub) = box la ua \<times> box lb ub"
     using lu[of _ 0] lu[of 0] by (auto intro!: inj_onI simp add: Basis_prod_def ball_Un box_def)
   show "emeasure (lborel \<Otimes>\<^sub>M lborel) (box (la, lb) (ua, ub)) =
-      ereal (setprod (op \<bullet> ((ua, ub) - (la, lb))) Basis)"
+      ennreal (setprod (op \<bullet> ((ua, ub) - (la, lb))) Basis)"
     by (simp add: lborel.emeasure_pair_measure_Times Basis_prod_def setprod.union_disjoint
-                  setprod.reindex)
+                  setprod.reindex ennreal_mult inner_diff_left setprod_nonneg)
 qed (simp add: borel_prod[symmetric])
 
 (* FIXME: conversion in measurable prover *)
@@ -811,14 +800,14 @@
         show "\<And>x. (\<lambda>k. ?f k x) \<longlonglongrightarrow> (if x \<in> UNION UNIV F \<inter> box a b then 1 else 0)"
           apply (auto simp: * setsum.If_cases Iio_Int_singleton)
           apply (rule_tac k="Suc xa" in LIMSEQ_offset)
-          apply (simp add: tendsto_const)
+          apply simp
           done
         have *: "emeasure lborel ((\<Union>x. F x) \<inter> box a b) \<le> emeasure lborel (box a b)"
           by (intro emeasure_mono) auto
 
         with union(1) show "(\<lambda>k. \<Sum>i<k. ?M (F i)) \<longlonglongrightarrow> ?M (\<Union>i. F i)"
           unfolding sums_def[symmetric] UN_extend_simps
-          by (intro measure_UNION) (auto simp: disjoint_family_on_def emeasure_lborel_box_eq)
+          by (intro measure_UNION) (auto simp: disjoint_family_on_def emeasure_lborel_box_eq top_unique)
       qed
       then show ?case
         by (subst (asm) has_integral_restrict) auto
@@ -844,43 +833,46 @@
     have "(\<lambda>n. emeasure lborel (A \<inter> ?B n)) \<longlonglongrightarrow> emeasure lborel (\<Union>n::nat. A \<inter> ?B n)"
       by (intro Lim_emeasure_incseq) (auto simp: incseq_def box_def)
     also have "(\<lambda>n. emeasure lborel (A \<inter> ?B n)) = (\<lambda>n. measure lborel (A \<inter> ?B n))"
-    proof (intro ext emeasure_eq_ereal_measure)
+    proof (intro ext emeasure_eq_ennreal_measure)
       fix n have "emeasure lborel (A \<inter> ?B n) \<le> emeasure lborel (?B n)"
         by (intro emeasure_mono) auto
-      then show "emeasure lborel (A \<inter> ?B n) \<noteq> \<infinity>"
-        by auto
+      then show "emeasure lborel (A \<inter> ?B n) \<noteq> top"
+        by (auto simp: top_unique)
     qed
     finally show "(\<lambda>n. measure lborel (A \<inter> ?B n)) \<longlonglongrightarrow> measure lborel A"
-      using emeasure_eq_ereal_measure[of lborel A] finite
-      by (simp add: UN_box_eq_UNIV)
+      using emeasure_eq_ennreal_measure[of lborel A] finite
+      by (simp add: UN_box_eq_UNIV less_top)
   qed
 qed
 
 lemma nn_integral_has_integral:
   fixes f::"'a::euclidean_space \<Rightarrow> real"
-  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) = ereal r"
+  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
   shows "(f has_integral r) UNIV"
-using f proof (induct arbitrary: r rule: borel_measurable_induct_real)
+using f proof (induct f arbitrary: r rule: borel_measurable_induct_real)
   case (set A)
   moreover then have "((\<lambda>x. 1) has_integral measure lborel A) A"
-    by (intro has_integral_measure_lborel) (auto simp: ereal_indicator)
+    by (intro has_integral_measure_lborel) (auto simp: ennreal_indicator)
   ultimately show ?case
-    by (simp add: ereal_indicator measure_def) (simp add: indicator_def)
+    by (simp add: ennreal_indicator measure_def) (simp add: indicator_def)
 next
   case (mult g c)
-  then have "ereal c * (\<integral>\<^sup>+ x. g x \<partial>lborel) = ereal r"
-    by (subst nn_integral_cmult[symmetric]) auto
-  then obtain r' where "(c = 0 \<and> r = 0) \<or> ((\<integral>\<^sup>+ x. ereal (g x) \<partial>lborel) = ereal r' \<and> r = c * r')"
-    by (cases "\<integral>\<^sup>+ x. ereal (g x) \<partial>lborel") (auto split: if_split_asm)
+  then have "ennreal c * (\<integral>\<^sup>+ x. g x \<partial>lborel) = ennreal r"
+    by (subst nn_integral_cmult[symmetric]) (auto simp: ennreal_mult)
+  with \<open>0 \<le> r\<close> \<open>0 \<le> c\<close>
+  obtain r' where "(c = 0 \<and> r = 0) \<or> (0 \<le> r' \<and> (\<integral>\<^sup>+ x. ennreal (g x) \<partial>lborel) = ennreal r' \<and> r = c * r')"
+    by (cases "\<integral>\<^sup>+ x. ennreal (g x) \<partial>lborel" rule: ennreal_cases)
+       (auto split: if_split_asm simp: ennreal_mult_top ennreal_mult[symmetric])
   with mult show ?case
     by (auto intro!: has_integral_cmult_real)
 next
   case (add g h)
   moreover
   then have "(\<integral>\<^sup>+ x. h x + g x \<partial>lborel) = (\<integral>\<^sup>+ x. h x \<partial>lborel) + (\<integral>\<^sup>+ x. g x \<partial>lborel)"
-    unfolding plus_ereal.simps[symmetric] by (subst nn_integral_add) auto
-  with add obtain a b where "(\<integral>\<^sup>+ x. h x \<partial>lborel) = ereal a" "(\<integral>\<^sup>+ x. g x \<partial>lborel) = ereal b" "r = a + b"
-    by (cases "\<integral>\<^sup>+ x. h x \<partial>lborel" "\<integral>\<^sup>+ x. g x \<partial>lborel" rule: ereal2_cases) auto
+    by (simp add: nn_integral_add)
+  with add obtain a b where "0 \<le> a" "0 \<le> b" "(\<integral>\<^sup>+ x. h x \<partial>lborel) = ennreal a" "(\<integral>\<^sup>+ x. g x \<partial>lborel) = ennreal b" "r = a + b"
+    by (cases "\<integral>\<^sup>+ x. h x \<partial>lborel" "\<integral>\<^sup>+ x. g x \<partial>lborel" rule: ennreal2_cases)
+       (auto simp: add_top nn_integral_add top_add ennreal_plus[symmetric] simp del: ennreal_plus)
   ultimately show ?case
     by (auto intro!: has_integral_add)
 next
@@ -897,16 +889,14 @@
   note U_le_f = this
 
   { fix i
-    have "(\<integral>\<^sup>+x. ereal (U i x) \<partial>lborel) \<le> (\<integral>\<^sup>+x. ereal (f x) \<partial>lborel)"
-      using U_le_f by (intro nn_integral_mono) simp
-    then obtain p where "(\<integral>\<^sup>+x. U i x \<partial>lborel) = ereal p" "p \<le> r"
-      using seq(6) by (cases "\<integral>\<^sup>+x. U i x \<partial>lborel") auto
-    moreover then have "0 \<le> p"
-      by (metis ereal_less_eq(5) nn_integral_nonneg)
+    have "(\<integral>\<^sup>+x. U i x \<partial>lborel) \<le> (\<integral>\<^sup>+x. f x \<partial>lborel)"
+      using seq(2) f(2) U_le_f by (intro nn_integral_mono) simp
+    then obtain p where "(\<integral>\<^sup>+x. U i x \<partial>lborel) = ennreal p" "p \<le> r" "0 \<le> p"
+      using seq(6) \<open>0\<le>r\<close> by (cases "\<integral>\<^sup>+x. U i x \<partial>lborel" rule: ennreal_cases) (auto simp: top_unique)
     moreover note seq
-    ultimately have "\<exists>p. (\<integral>\<^sup>+x. U i x \<partial>lborel) = ereal p \<and> 0 \<le> p \<and> p \<le> r \<and> (U i has_integral p) UNIV"
+    ultimately have "\<exists>p. (\<integral>\<^sup>+x. U i x \<partial>lborel) = ennreal p \<and> 0 \<le> p \<and> p \<le> r \<and> (U i has_integral p) UNIV"
       by auto }
-  then obtain p where p: "\<And>i. (\<integral>\<^sup>+x. ereal (U i x) \<partial>lborel) = ereal (p i)"
+  then obtain p where p: "\<And>i. (\<integral>\<^sup>+x. ennreal (U i x) \<partial>lborel) = ennreal (p i)"
     and bnd: "\<And>i. p i \<le> r" "\<And>i. 0 \<le> p i"
     and U_int: "\<And>i.(U i has_integral (p i)) UNIV" by metis
 
@@ -922,9 +912,9 @@
       using seq by auto
   qed
   moreover have "(\<lambda>i. (\<integral>\<^sup>+x. U i x \<partial>lborel)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. f x \<partial>lborel)"
-    using seq U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto
+    using seq f(2) U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto
   ultimately have "integral UNIV f = r"
-    by (auto simp add: int_eq p seq intro: LIMSEQ_unique)
+    by (auto simp add: bnd int_eq p seq intro: LIMSEQ_unique)
   with * show ?case
     by (simp add: has_integral_integral)
 qed
@@ -934,8 +924,8 @@
   assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) < \<infinity>"
   shows "(\<integral>\<^sup>+x. f x \<partial>lborel) = integral UNIV f"
 proof -
-  from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ereal r"
-    by (cases "\<integral>\<^sup>+x. f x \<partial>lborel") auto
+  from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
+    by (cases "\<integral>\<^sup>+x. f x \<partial>lborel" rule: ennreal_cases) auto
   then show ?thesis
     using nn_integral_has_integral[OF f(1,2) r] by (simp add: integral_unique)
 qed
@@ -945,8 +935,8 @@
   assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) < \<infinity>"
   shows "f integrable_on UNIV"
 proof -
-  from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ereal r"
-    by (cases "\<integral>\<^sup>+x. f x \<partial>lborel") auto
+  from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
+    by (cases "\<integral>\<^sup>+x. f x \<partial>lborel" rule: ennreal_cases) auto
   then show ?thesis
     by (intro has_integral_integrable[where i=r] nn_integral_has_integral[where r=r] f)
 qed
@@ -957,25 +947,26 @@
   assumes I: "(f has_integral I) UNIV"
   shows "integral\<^sup>N lborel f = I"
 proof -
-  from f_borel have "(\<lambda>x. ereal (f x)) \<in> borel_measurable lborel" by auto
+  from f_borel have "(\<lambda>x. ennreal (f x)) \<in> borel_measurable lborel" by auto
   from borel_measurable_implies_simple_function_sequence'[OF this] guess F . note F = this
   let ?B = "\<lambda>i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One) :: 'a set"
 
   note F(1)[THEN borel_measurable_simple_function, measurable]
 
-  { fix i x have "real_of_ereal (F i x) \<le> f x"
-      using F(3,5) F(4)[of x, symmetric] nonneg
-      unfolding real_le_ereal_iff
-      by (auto simp: image_iff eq_commute[of \<infinity>] max_def intro: SUP_upper split: if_split_asm) }
-  note F_le_f = this
+  have "0 \<le> I"
+    using I by (rule has_integral_nonneg) (simp add: nonneg)
+
+  have F_le_f: "enn2real (F i x) \<le> f x" for i x
+    using F(3,4)[where x=x] nonneg SUP_upper[of i UNIV "\<lambda>i. F i x"]
+    by (cases "F i x" rule: ennreal_cases) auto
   let ?F = "\<lambda>i x. F i x * indicator (?B i) x"
-  have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>lborel) = (SUP i. integral\<^sup>N lborel (\<lambda>x. ?F i x))"
+  have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) = (SUP i. integral\<^sup>N lborel (\<lambda>x. ?F i x))"
   proof (subst nn_integral_monotone_convergence_SUP[symmetric])
     { fix x
       obtain j where j: "x \<in> ?B j"
         using UN_box_eq_UNIV by auto
 
-      have "ereal (f x) = (SUP i. F i x)"
+      have "ennreal (f x) = (SUP i. F i x)"
         using F(4)[of x] nonneg[of x] by (simp add: max_def)
       also have "\<dots> = (SUP i. ?F i x)"
       proof (rule SUP_eq)
@@ -984,50 +975,50 @@
           by (intro bexI[of _ "max i j"])
              (auto split: split_max split_indicator simp: incseq_def le_fun_def box_def)
       qed (auto intro!: F split: split_indicator)
-      finally have "ereal (f x) =  (SUP i. ?F i x)" . }
-    then show "(\<integral>\<^sup>+ x. ereal (f x) \<partial>lborel) = (\<integral>\<^sup>+ x. (SUP i. ?F i x) \<partial>lborel)"
+      finally have "ennreal (f x) =  (SUP i. ?F i x)" . }
+    then show "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) = (\<integral>\<^sup>+ x. (SUP i. ?F i x) \<partial>lborel)"
       by simp
   qed (insert F, auto simp: incseq_def le_fun_def box_def split: split_indicator)
-  also have "\<dots> \<le> ereal I"
+  also have "\<dots> \<le> ennreal I"
   proof (rule SUP_least)
     fix i :: nat
-    have finite_F: "(\<integral>\<^sup>+ x. ereal (real_of_ereal (F i x) * indicator (?B i) x) \<partial>lborel) < \<infinity>"
+    have finite_F: "(\<integral>\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \<partial>lborel) < \<infinity>"
     proof (rule nn_integral_bound_simple_function)
-      have "emeasure lborel {x \<in> space lborel. ereal (real_of_ereal (F i x) * indicator (?B i) x) \<noteq> 0} \<le>
+      have "emeasure lborel {x \<in> space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \<noteq> 0} \<le>
         emeasure lborel (?B i)"
         by (intro emeasure_mono)  (auto split: split_indicator)
-      then show "emeasure lborel {x \<in> space lborel. ereal (real_of_ereal (F i x) * indicator (?B i) x) \<noteq> 0} < \<infinity>"
-        by auto
+      then show "emeasure lborel {x \<in> space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \<noteq> 0} < \<infinity>"
+        by (auto simp: less_top[symmetric] top_unique)
     qed (auto split: split_indicator
-              intro!: real_of_ereal_pos F simple_function_compose1[where g="real_of_ereal"] simple_function_ereal)
+              intro!: F simple_function_compose1[where g="enn2real"] simple_function_ennreal)
 
-    have int_F: "(\<lambda>x. real_of_ereal (F i x) * indicator (?B i) x) integrable_on UNIV"
-      using F(5) finite_F
-      by (intro nn_integral_integrable_on) (auto split: split_indicator intro: real_of_ereal_pos)
+    have int_F: "(\<lambda>x. enn2real (F i x) * indicator (?B i) x) integrable_on UNIV"
+      using F(4) finite_F
+      by (intro nn_integral_integrable_on) (auto split: split_indicator simp: enn2real_nonneg)
 
     have "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) =
-      (\<integral>\<^sup>+ x. ereal (real_of_ereal (F i x) * indicator (?B i) x) \<partial>lborel)"
-      using F(3,5)
-      by (intro nn_integral_cong) (auto simp: image_iff ereal_real eq_commute split: split_indicator)
-    also have "\<dots> = ereal (integral UNIV (\<lambda>x. real_of_ereal (F i x) * indicator (?B i) x))"
+      (\<integral>\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \<partial>lborel)"
+      using F(3,4)
+      by (intro nn_integral_cong) (auto simp: image_iff eq_commute split: split_indicator)
+    also have "\<dots> = ennreal (integral UNIV (\<lambda>x. enn2real (F i x) * indicator (?B i) x))"
       using F
       by (intro nn_integral_lborel_eq_integral[OF _ _ finite_F])
-         (auto split: split_indicator intro: real_of_ereal_pos)
-    also have "\<dots> \<le> ereal I"
+         (auto split: split_indicator intro: enn2real_nonneg)
+    also have "\<dots> \<le> ennreal I"
       by (auto intro!: has_integral_le[OF integrable_integral[OF int_F] I] nonneg F_le_f
-          split: split_indicator )
-    finally show "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) \<le> ereal I" .
+               simp: \<open>0 \<le> I\<close> split: split_indicator )
+    finally show "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) \<le> ennreal I" .
   qed
-  finally have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>lborel) < \<infinity>"
-    by auto
+  finally have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) < \<infinity>"
+    by (auto simp: less_top[symmetric] top_unique)
   from nn_integral_lborel_eq_integral[OF assms(1,2) this] I show ?thesis
     by (simp add: integral_unique)
 qed
 
 lemma has_integral_iff_emeasure_lborel:
   fixes A :: "'a::euclidean_space set"
-  assumes A[measurable]: "A \<in> sets borel"
-  shows "((\<lambda>x. 1) has_integral r) A \<longleftrightarrow> emeasure lborel A = ereal r"
+  assumes A[measurable]: "A \<in> sets borel" and [simp]: "0 \<le> r"
+  shows "((\<lambda>x. 1) has_integral r) A \<longleftrightarrow> emeasure lborel A = ennreal r"
 proof cases
   assume emeasure_A: "emeasure lborel A = \<infinity>"
   have "\<not> (\<lambda>x. 1::real) integrable_on A"
@@ -1038,16 +1029,16 @@
     then obtain r where "((indicator A::'a\<Rightarrow>real) has_integral r) UNIV"
       by auto
     from nn_integral_has_integral_lborel[OF _ _ this] emeasure_A show False
-      by (simp add: ereal_indicator)
+      by (simp add: ennreal_indicator)
   qed
   with emeasure_A show ?thesis
     by auto
 next
   assume "emeasure lborel A \<noteq> \<infinity>"
   moreover then have "((\<lambda>x. 1) has_integral measure lborel A) A"
-    by (simp add: has_integral_measure_lborel)
+    by (simp add: has_integral_measure_lborel less_top)
   ultimately show ?thesis
-    by (auto simp: emeasure_eq_ereal_measure has_integral_unique)
+    by (auto simp: emeasure_eq_ennreal_measure has_integral_unique)
 qed
 
 lemma has_integral_integral_real:
@@ -1057,7 +1048,7 @@
 using f proof induct
   case (base A c) then show ?case
     by (auto intro!: has_integral_mult_left simp: )
-       (simp add: emeasure_eq_ereal_measure indicator_def has_integral_measure_lborel)
+       (simp add: emeasure_eq_ennreal_measure indicator_def has_integral_measure_lborel)
 next
   case (add f g) then show ?case
     by (auto intro!: has_integral_add)
@@ -1069,8 +1060,7 @@
     show "(\<lambda>x. norm (2 * f x)) integrable_on UNIV"
       using \<open>integrable lborel f\<close>
       by (intro nn_integral_integrable_on)
-         (auto simp: integrable_iff_bounded abs_mult times_ereal.simps(1)[symmetric] nn_integral_cmult
-               simp del: times_ereal.simps)
+         (auto simp: integrable_iff_bounded abs_mult  nn_integral_cmult ennreal_mult ennreal_mult_less_top)
     show "\<And>k. \<forall>x\<in>UNIV. norm (s k x) \<le> norm (2 * f x)"
       using lim by (auto simp add: abs_mult)
     show "\<forall>x\<in>UNIV. (\<lambda>k. s k x) \<longlonglongrightarrow> f x"
@@ -1099,7 +1089,7 @@
 qed
 
 lemma integrable_on_lborel: "integrable lborel f \<Longrightarrow> f integrable_on UNIV"
-  using has_integral_integral_lborel by (auto intro: has_integral_integrable)
+  using has_integral_integral_lborel by auto
 
 lemma integral_lborel: "integrable lborel f \<Longrightarrow> integral UNIV f = (\<integral>x. f x \<partial>lborel)"
   using has_integral_integral_lborel by auto
@@ -1116,7 +1106,7 @@
   then have "emeasure lborel A \<le> emeasure lborel (cbox a b)"
     by (intro emeasure_mono) auto
   then show ?thesis
-    by (auto simp: emeasure_lborel_cbox_eq)
+    by (auto simp: emeasure_lborel_cbox_eq setprod_nonneg less_top[symmetric] top_unique split: if_split_asm)
 qed
 
 lemma emeasure_compact_finite: "compact A \<Longrightarrow> emeasure lborel A < \<infinity>"
@@ -1169,7 +1159,7 @@
   fixes f :: "real \<Rightarrow> real"
   assumes [measurable]: "f \<in> borel_measurable borel"
   assumes f: "\<And>x. x \<in> {a..b} \<Longrightarrow> DERIV F x :> f x" "\<And>x. x \<in> {a..b} \<Longrightarrow> 0 \<le> f x" and "a \<le> b"
-  shows nn_integral_FTC_Icc: "(\<integral>\<^sup>+x. ereal (f x) * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?nn)
+  shows nn_integral_FTC_Icc: "(\<integral>\<^sup>+x. ennreal (f x) * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?nn)
     and has_bochner_integral_FTC_Icc_nonneg:
       "has_bochner_integral lborel (\<lambda>x. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
     and integral_FTC_Icc_nonneg: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq)
@@ -1178,6 +1168,9 @@
   have *: "(\<lambda>x. f x * indicator {a..b} x) \<in> borel_measurable borel" "\<And>x. 0 \<le> f x * indicator {a..b} x"
     using f(2) by (auto split: split_indicator)
 
+  have F_mono: "a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b\<Longrightarrow> F x \<le> F y" for x y
+    using f by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans)
+
   have "(f has_integral F b - F a) {a..b}"
     by (intro fundamental_theorem_of_calculus)
        (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric]
@@ -1188,11 +1181,12 @@
   then have nn: "(\<integral>\<^sup>+x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a"
     by (rule nn_integral_has_integral_lborel[OF *])
   then show ?has
-    by (rule has_bochner_integral_nn_integral[rotated 2]) (simp_all add: *)
+    by (rule has_bochner_integral_nn_integral[rotated 3]) (simp_all add: * F_mono \<open>a \<le> b\<close>)
   then show ?eq ?int
     unfolding has_bochner_integral_iff by auto
-  from nn show ?nn
-    by (simp add: ereal_mult_indicator)
+  show ?nn
+    by (subst nn[symmetric])
+       (auto intro!: nn_integral_cong simp add: ennreal_mult f split: split_indicator)
 qed
 
 lemma
@@ -1245,22 +1239,25 @@
   assumes f: "\<And>x. a \<le> x \<Longrightarrow> DERIV F x :> f x"
   assumes nonneg: "\<And>x. a \<le> x \<Longrightarrow> 0 \<le> f x"
   assumes lim: "(F \<longlongrightarrow> T) at_top"
-  shows "(\<integral>\<^sup>+x. ereal (f x) * indicator {a ..} x \<partial>lborel) = T - F a"
+  shows "(\<integral>\<^sup>+x. ennreal (f x) * indicator {a ..} x \<partial>lborel) = T - F a"
 proof -
-  let ?f = "\<lambda>(i::nat) (x::real). ereal (f x) * indicator {a..a + real i} x"
-  let ?fR = "\<lambda>x. ereal (f x) * indicator {a ..} x"
-  have "\<And>x. (SUP i::nat. ?f i x) = ?fR x"
-  proof (rule SUP_Lim_ereal)
-    show "\<And>x. incseq (\<lambda>i. ?f i x)"
-      using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator)
+  let ?f = "\<lambda>(i::nat) (x::real). ennreal (f x) * indicator {a..a + real i} x"
+  let ?fR = "\<lambda>x. ennreal (f x) * indicator {a ..} x"
 
-    fix x
+  have F_mono: "a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> F x \<le> F y" for x y
+    using f nonneg by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans)
+  then have F_le_T: "a \<le> x \<Longrightarrow> F x \<le> T" for x
+    by (intro tendsto_le_const[OF _ lim])
+       (auto simp: trivial_limit_at_top_linorder eventually_at_top_linorder)
+
+  have "(SUP i::nat. ?f i x) = ?fR x" for x
+  proof (rule LIMSEQ_unique[OF LIMSEQ_SUP])
     from reals_Archimedean2[of "x - a"] guess n ..
     then have "eventually (\<lambda>n. ?f n x = ?fR x) sequentially"
       by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator)
     then show "(\<lambda>n. ?f n x) \<longlonglongrightarrow> ?fR x"
       by (rule Lim_eventually)
-  qed
+  qed (auto simp: nonneg incseq_def le_fun_def split: split_indicator)
   then have "integral\<^sup>N lborel ?fR = (\<integral>\<^sup>+ x. (SUP i::nat. ?f i x) \<partial>lborel)"
     by simp
   also have "\<dots> = (SUP i::nat. (\<integral>\<^sup>+ x. ?f i x \<partial>lborel))"
@@ -1270,26 +1267,18 @@
     show "\<And>i. (?f i) \<in> borel_measurable lborel"
       using f_borel by auto
   qed
-  also have "\<dots> = (SUP i::nat. ereal (F (a + real i) - F a))"
+  also have "\<dots> = (SUP i::nat. ennreal (F (a + real i) - F a))"
     by (subst nn_integral_FTC_Icc[OF f_borel f nonneg]) auto
   also have "\<dots> = T - F a"
-  proof (rule SUP_Lim_ereal)
-    show "incseq (\<lambda>n. ereal (F (a + real n) - F a))"
-    proof (simp add: incseq_def, safe)
-      fix m n :: nat assume "m \<le> n"
-      with f nonneg show "F (a + real m) \<le> F (a + real n)"
-        by (intro DERIV_nonneg_imp_nondecreasing[where f=F])
-           (simp, metis add_increasing2 order_refl order_trans of_nat_0_le_iff)
-    qed
+  proof (rule LIMSEQ_unique[OF LIMSEQ_SUP])
     have "(\<lambda>x. F (a + real x)) \<longlonglongrightarrow> T"
       apply (rule filterlim_compose[OF lim filterlim_tendsto_add_at_top])
       apply (rule LIMSEQ_const_iff[THEN iffD2, OF refl])
       apply (rule filterlim_real_sequentially)
       done
-    then show "(\<lambda>n. ereal (F (a + real n) - F a)) \<longlonglongrightarrow> ereal (T - F a)"
-      unfolding lim_ereal
-      by (intro tendsto_diff) auto
-  qed
+    then show "(\<lambda>n. ennreal (F (a + real n) - F a)) \<longlonglongrightarrow> ennreal (T - F a)"
+      by (simp add: F_mono F_le_T tendsto_diff)
+  qed (auto simp: incseq_def intro!: ennreal_le_iff[THEN iffD2] F_mono)
   finally show ?thesis .
 qed
 
--- a/src/HOL/Probability/Levy.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Levy.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -25,17 +25,17 @@
   have 1: "cmod (?F t - (b - a)) \<le> a^2 / 2 * abs t + b^2 / 2 * abs t" if "t \<noteq> 0" for t
   proof -
     have "cmod (?F t - (b - a)) = cmod (
-        (iexp (-(t * a)) - (1 + ii * -(t * a))) / (ii * t) - 
-        (iexp (-(t * b)) - (1 + ii * -(t * b))) / (ii * t))"  
+        (iexp (-(t * a)) - (1 + ii * -(t * a))) / (ii * t) -
+        (iexp (-(t * b)) - (1 + ii * -(t * b))) / (ii * t))"
            (is "_ = cmod (?one / (ii * t) - ?two / (ii * t))")
       using `t \<noteq> 0` by (intro arg_cong[where f=norm]) (simp add: field_simps)
-    also have "\<dots> \<le> cmod (?one / (ii * t)) + cmod (?two / (ii * t))" 
+    also have "\<dots> \<le> cmod (?one / (ii * t)) + cmod (?two / (ii * t))"
       by (rule norm_triangle_ineq4)
     also have "cmod (?one / (ii * t)) = cmod ?one / abs t"
       by (simp add: norm_divide norm_mult)
     also have "cmod (?two / (ii * t)) = cmod ?two / abs t"
-      by (simp add: norm_divide norm_mult)      
-    also have "cmod ?one / abs t + cmod ?two / abs t \<le> 
+      by (simp add: norm_divide norm_mult)
+    also have "cmod ?one / abs t + cmod ?two / abs t \<le>
         ((- (a * t))^2 / 2) / abs t + ((- (b * t))^2 / 2) / abs t"
       apply (rule add_mono)
       apply (rule divide_right_mono)
@@ -71,7 +71,7 @@
     using iexp_approx1 [of "t * (b - a)" 0]
     by (intro divide_right_mono) (auto simp add: field_simps eval_nat_numeral)
   also have "\<dots> = b - a"
-    using assms by (auto simp add: abs_mult) 
+    using assms by (auto simp add: abs_mult)
   finally show ?thesis .
 qed
 
@@ -109,7 +109,7 @@
       also have "\<dots> + ?t = (CLBINT t=(0::real)..T. ?f (-t) x + ?f t x)"
         using 1
         by (intro interval_lebesgue_integral_add(2) [symmetric] interval_integrable_mirror[THEN iffD2]) simp_all
-      also have "\<dots> = (CLBINT t=(0::real)..T. ((iexp(t * (x - a)) - iexp (-(t * (x - a)))) -  
+      also have "\<dots> = (CLBINT t=(0::real)..T. ((iexp(t * (x - a)) - iexp (-(t * (x - a)))) -
           (iexp(t * (x - b)) - iexp (-(t * (x - b))))) / (ii * t))"
         using `T \<ge> 0` by (intro interval_integral_cong) (auto simp add: divide_simps)
       also have "\<dots> = (CLBINT t=(0::real)..T. complex_of_real(
@@ -120,7 +120,7 @@
         unfolding minus_diff_eq[symmetric, of "y * x" "y * a" for y a] sin_minus cos_minus
         apply (simp add: field_simps power2_eq_square)
         done
-      also have "\<dots> = complex_of_real (LBINT t=(0::real)..T. 
+      also have "\<dots> = complex_of_real (LBINT t=(0::real)..T.
           2 * (sin (t * (x - a)) / t) - 2 * (sin (t * (x - b)) / t))"
         by (rule interval_lebesgue_integral_of_real)
       also have "\<dots> = complex_of_real (2 * (sgn (x - a) * Si (T * abs (x - a)) -
@@ -133,7 +133,7 @@
       finally have "(CLBINT t. ?f' (t, x)) =
           2 * (sgn (x - a) * Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b)))" .
     } note main_eq = this
-    have "(CLBINT t=-T..T. ?F t * \<phi> t) = 
+    have "(CLBINT t=-T..T. ?F t * \<phi> t) =
       (CLBINT t. (CLINT x | M. ?F t * iexp (t * x) * indicator {-T<..<T} t))"
       using `T \<ge> 0` unfolding \<phi>_def char_def interval_lebesgue_integral_def
       by (auto split: split_indicator intro!: integral_cong)
@@ -142,24 +142,26 @@
     also have "\<dots> = (CLINT x | M. (CLBINT t. ?f' (t, x)))"
     proof (intro P.Fubini_integral [symmetric] integrableI_bounded_set [where B="b - a"])
       show "emeasure (lborel \<Otimes>\<^sub>M M) ({- T<..<T} \<times> space M) < \<infinity>"
-        using \<open>T \<ge> 0\<close> by (subst emeasure_pair_measure_Times) auto
+        using \<open>T \<ge> 0\<close>
+        by (subst emeasure_pair_measure_Times)
+           (auto simp: ennreal_mult_less_top not_less top_unique)
       show "AE x\<in>{- T<..<T} \<times> space M in lborel \<Otimes>\<^sub>M M. cmod (case x of (t, x) \<Rightarrow> ?f' (t, x)) \<le> b - a"
         using Levy_Inversion_aux2[of "x - b" "x - a" for x] `a \<le> b`
         by (intro AE_I [of _ _ "{0} \<times> UNIV"]) (force simp: emeasure_pair_measure_Times)+
     qed (auto split: split_indicator split_indicator_asm)
-    also have "\<dots> = (CLINT x | M. (complex_of_real (2 * (sgn (x - a) * 
+    also have "\<dots> = (CLINT x | M. (complex_of_real (2 * (sgn (x - a) *
          Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b))))))"
        using main_eq by (intro integral_cong, auto)
-    also have "\<dots> = complex_of_real (LINT x | M. (2 * (sgn (x - a) * 
+    also have "\<dots> = complex_of_real (LINT x | M. (2 * (sgn (x - a) *
          Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b)))))"
        by (rule integral_complex_of_real)
-    finally have "(CLBINT t=-T..T. ?F t * \<phi> t) = 
-        complex_of_real (LINT x | M. (2 * (sgn (x - a) * 
+    finally have "(CLBINT t=-T..T. ?F t * \<phi> t) =
+        complex_of_real (LINT x | M. (2 * (sgn (x - a) *
          Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b)))))" .
   } note main_eq2 = this
 
-  have "(\<lambda>T :: nat. LINT x | M. (2 * (sgn (x - a) * 
-         Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b))))) \<longlonglongrightarrow> 
+  have "(\<lambda>T :: nat. LINT x | M. (2 * (sgn (x - a) *
+         Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b))))) \<longlonglongrightarrow>
        (LINT x | M. 2 * pi * indicator {a<..b} x)"
   proof (rule integral_dominated_convergence [where w="\<lambda>x. 4 * B"])
     show "integrable M (\<lambda>x. 4 * B)"
@@ -191,11 +193,11 @@
         using x `a \<le> b` by (auto split: split_indicator)
       finally show "(\<lambda>n. 2 * ?S n x) \<longlonglongrightarrow> 2 * pi * indicator {a<..b} x" .
     qed
-  qed simp_all 
+  qed simp_all
   also have "(LINT x | M. 2 * pi * indicator {a<..b} x) = 2 * pi * \<mu> {a<..b}"
     by (simp add: \<mu>_def)
-  finally have "(\<lambda>T. LINT x | M. (2 * (sgn (x - a) * 
-         Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b))))) \<longlonglongrightarrow> 
+  finally have "(\<lambda>T. LINT x | M. (2 * (sgn (x - a) *
+         Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b))))) \<longlonglongrightarrow>
        2 * pi * \<mu> {a<..b}" .
   with main_eq2 show ?thesis
     by (auto intro!: tendsto_eq_intros)
@@ -233,7 +235,7 @@
         "measure M1 {a} = 0" "measure M2 {a} = 0" "a < M" "a \<le> x" "\<bar>cdf M1 a\<bar> < \<epsilon> / 4" "\<bar>cdf M2 a\<bar> < \<epsilon> / 4"
         by auto
 
-      from `\<epsilon> > 0` `(cdf M1 \<longlongrightarrow> cdf M1 x) (at_right x)` `(cdf M2 \<longlongrightarrow> cdf M2 x) (at_right x)` 
+      from `\<epsilon> > 0` `(cdf M1 \<longlongrightarrow> cdf M1 x) (at_right x)` `(cdf M2 \<longlongrightarrow> cdf M2 x) (at_right x)`
       have "eventually (\<lambda>y. \<bar>cdf M1 y - cdf M1 x\<bar> < \<epsilon> / 4 \<and> \<bar>cdf M2 y - cdf M2 x\<bar> < \<epsilon> / 4 \<and> x < y) (at_right x)"
         by (simp only: tendsto_iff dist_real_def eventually_conj eventually_at_right_less)
       then obtain N where "N > x" "\<And>y. x < y \<Longrightarrow> y < N \<Longrightarrow> \<bar>cdf M1 y - cdf M1 x\<bar> < \<epsilon> / 4"
@@ -269,9 +271,9 @@
         by (intro add_mono less_imp_le \<open>\<bar>cdf M2 x - cdf M2 b\<bar> < \<epsilon> / 4\<close> \<open>\<bar>cdf M2 a\<bar> < \<epsilon> / 4\<close>)
       finally have 2: "abs (cdf M2 x - (cdf M2 b - cdf M2 a)) \<le> \<epsilon> / 2" by simp
 
-      have "abs (cdf M1 x - cdf M2 x) = abs ((cdf M1 x - (cdf M1 b - cdf M1 a)) - 
+      have "abs (cdf M1 x - cdf M2 x) = abs ((cdf M1 x - (cdf M1 b - cdf M1 a)) -
           (cdf M2 x - (cdf M2 b - cdf M2 a)))" by (subst *, simp)
-      also have "\<dots> \<le> abs (cdf M1 x - (cdf M1 b - cdf M1 a)) + 
+      also have "\<dots> \<le> abs (cdf M1 x - (cdf M1 b - cdf M1 a)) +
           abs (cdf M2 x - (cdf M2 b - cdf M2 a))" by (rule abs_triangle_ineq4)
       also have "\<dots> \<le> \<epsilon> / 2 + \<epsilon> / 2" by (rule add_mono [OF 1 2])
       finally have "abs (cdf M1 x - cdf M2 x) \<le> \<epsilon>" by simp }
@@ -295,13 +297,13 @@
   fixes M :: "nat \<Rightarrow> real measure" and M' :: "real measure"
   assumes real_distr_M : "\<And>n. real_distribution (M n)"
     and real_distr_M': "real_distribution M'"
-    and char_conv: "\<And>t. (\<lambda>n. char (M n) t) \<longlonglongrightarrow> char M' t" 
+    and char_conv: "\<And>t. (\<lambda>n. char (M n) t) \<longlonglongrightarrow> char M' t"
   shows "weak_conv_m M M'"
 proof -
   interpret Mn: real_distribution "M n" for n by fact
   interpret M': real_distribution M' by fact
 
-  have *: "\<And>u x. u > 0 \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> (CLBINT t:{-u..u}. 1 - iexp (t * x)) = 
+  have *: "\<And>u x. u > 0 \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> (CLBINT t:{-u..u}. 1 - iexp (t * x)) =
       2 * (u  - sin (u * x) / x)"
   proof -
     fix u :: real and x :: real
@@ -331,7 +333,7 @@
     finally show "(CLBINT t:{-u..u}. 1 - iexp (t * x)) = 2 * (u  - sin (u * x) / x)"
       by (simp add: field_simps)
   qed
-  have main_bound: "\<And>u n. u > 0 \<Longrightarrow> Re (CLBINT t:{-u..u}. 1 - char (M n) t) \<ge> 
+  have main_bound: "\<And>u n. u > 0 \<Longrightarrow> Re (CLBINT t:{-u..u}. 1 - char (M n) t) \<ge>
     u * measure (M n) {x. abs x \<ge> 2 / u}"
   proof -
     fix u :: real and n
@@ -345,9 +347,10 @@
     have Mn3: "set_integrable (M n \<Otimes>\<^sub>M lborel) (UNIV \<times> {- u..u}) (\<lambda>a. 1 - exp (\<i> * complex_of_real (snd a * fst a)))"
       using `0 < u`
       by (intro integrableI_bounded_set_indicator [where B="2"])
-         (auto simp: lborel.emeasure_pair_measure_Times split: split_indicator
+         (auto simp: lborel.emeasure_pair_measure_Times ennreal_mult_less_top not_less top_unique
+               split: split_indicator
                intro!: order_trans [OF norm_triangle_ineq4])
-    have "(CLBINT t:{-u..u}. 1 - char (M n) t) = 
+    have "(CLBINT t:{-u..u}. 1 - char (M n) t) =
         (CLBINT t:{-u..u}. (CLINT x | M n. 1 - iexp (t * x)))"
       unfolding char_def by (rule set_lebesgue_integral_cong, auto simp del: of_real_mult)
     also have "\<dots> = (CLBINT t. (CLINT x | M n. indicator {-u..u} t *\<^sub>R (1 - iexp (t * x))))"
@@ -358,7 +361,7 @@
       using `u > 0` by (intro integral_cong, auto simp add: * simp del: of_real_mult)
     also have "\<dots> = (LINT x | M n. (if x = 0 then 0 else 2 * (u  - sin (u * x) / x)))"
       by (rule integral_complex_of_real)
-    finally have "Re (CLBINT t:{-u..u}. 1 - char (M n) t) = 
+    finally have "Re (CLBINT t:{-u..u}. 1 - char (M n) t) =
        (LINT x | M n. (if x = 0 then 0 else 2 * (u  - sin (u * x) / x)))" by simp
     also have "\<dots> \<ge> (LINT x : {x. abs x \<ge> 2 / u} | M n. u)"
     proof -
@@ -377,9 +380,10 @@
       ultimately show ?thesis
         using `u > 0`
         by (intro integral_mono [OF _ **])
-           (auto simp: divide_simps sin_x_le_x mult.commute[of u] mult_neg_pos split: split_indicator)
+           (auto simp: divide_simps sin_x_le_x mult.commute[of u] mult_neg_pos top_unique less_top[symmetric]
+                 split: split_indicator)
     qed
-    also (xtrans) have "(LINT x : {x. abs x \<ge> 2 / u} | M n. u) = 
+    also (xtrans) have "(LINT x : {x. abs x \<ge> 2 / u} | M n. u) =
         u * measure (M n) {x. abs x \<ge> 2 / u}"
       by (simp add: Mn.emeasure_eq_measure)
     finally show "Re (CLBINT t:{-u..u}. 1 - char (M n) t) \<ge> u * measure (M n) {x. abs x \<ge> 2 / u}" .
@@ -423,7 +427,7 @@
     have "(\<lambda>n. CLBINT t:{-d/2..d/2}. 1 - char (M n) t) \<longlonglongrightarrow> (CLBINT t:{-d/2..d/2}. 1 - char M' t)"
       using bd1
       apply (intro integral_dominated_convergence[where w="\<lambda>x. indicator {-d/2..d/2} x *\<^sub>R 2"])
-      apply (auto intro!: char_conv tendsto_intros 
+      apply (auto intro!: char_conv tendsto_intros
                   simp: emeasure_lborel_Icc_eq
                   split: split_indicator)
       done
@@ -438,20 +442,20 @@
         (CLBINT t:{-d/2..d/2}. 1 - char M' t)) < d * \<epsilon> / 4" by auto
     { fix n
       assume "n \<ge> N"
-      have "cmod (CLBINT t:{-d/2..d/2}. 1 - char (M n) t) = 
+      have "cmod (CLBINT t:{-d/2..d/2}. 1 - char (M n) t) =
         cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) - (CLBINT t:{-d/2..d/2}. 1 - char M' t)
           + (CLBINT t:{-d/2..d/2}. 1 - char M' t))" by simp
-      also have "\<dots> \<le> cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) - 
+      also have "\<dots> \<le> cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) -
           (CLBINT t:{-d/2..d/2}. 1 - char M' t)) + cmod(CLBINT t:{-d/2..d/2}. 1 - char M' t)"
         by (rule norm_triangle_ineq)
-      also have "\<dots> < d * \<epsilon> / 4 + d * \<epsilon> / 4" 
+      also have "\<dots> < d * \<epsilon> / 4 + d * \<epsilon> / 4"
         by (rule add_less_le_mono [OF N [OF `n \<ge> N`] bound])
       also have "\<dots> = d * \<epsilon> / 2" by auto
       finally have "cmod (CLBINT t:{-d/2..d/2}. 1 - char (M n) t) < d * \<epsilon> / 2" .
       hence "d * \<epsilon> / 2 > Re (CLBINT t:{-d/2..d/2}. 1 - char (M n) t)"
         by (rule order_le_less_trans [OF complex_Re_le_cmod])
       hence "d * \<epsilon> / 2 > Re (CLBINT t:{-(d/2)..d/2}. 1 - char (M n) t)" (is "_ > ?lhs") by simp
-      also have "?lhs \<ge> (d / 2) * measure (M n) {x. abs x \<ge> 2 / (d / 2)}" 
+      also have "?lhs \<ge> (d / 2) * measure (M n) {x. abs x \<ge> 2 / (d / 2)}"
         using d0 by (intro main_bound, simp)
       finally (xtrans) have "d * \<epsilon> / 2 > (d / 2) * measure (M n) {x. abs x \<ge> 2 / (d / 2)}" .
       with d0 `\<epsilon> > 0` have "\<epsilon> > measure (M n) {x. abs x \<ge> 2 / (d / 2)}" by (simp add: field_simps)
@@ -473,7 +477,7 @@
     { fix n :: nat
       have *: "(UN (k :: nat). {- real k<..real k}) = UNIV"
         by (auto, metis leI le_less_trans less_imp_le minus_less_iff reals_Archimedean2)
-      have "(\<lambda>k. measure (M n) {- real k<..real k}) \<longlonglongrightarrow> 
+      have "(\<lambda>k. measure (M n) {- real k<..real k}) \<longlonglongrightarrow>
           measure (M n) (UN (k :: nat). {- real k<..real k})"
         by (rule Mn.finite_Lim_measure_incseq, auto simp add: incseq_def)
       hence "(\<lambda>k. measure (M n) {- real k<..real k}) \<longlonglongrightarrow> 1"
@@ -490,11 +494,11 @@
           by eventually_elim (auto simp add: less_Suc_eq)
       qed simp
     } note 8 = this
-    from 8 [of N] have "\<exists>K :: nat. \<forall>k \<ge> K. \<forall>m<N. 1 - \<epsilon> < 
+    from 8 [of N] have "\<exists>K :: nat. \<forall>k \<ge> K. \<forall>m<N. 1 - \<epsilon> <
         Sigma_Algebra.measure (M m) {- real k<..real k}"
       by (auto simp add: eventually_sequentially)
     hence "\<exists>K :: nat. \<forall>m<N. 1 - \<epsilon> < Sigma_Algebra.measure (M m) {- real K<..real K}" by auto
-    then obtain K :: nat where 
+    then obtain K :: nat where
       "\<forall>m<N. 1 - \<epsilon> < Sigma_Algebra.measure (M m) {- real K<..real K}" ..
     hence K: "\<And>m. m < N \<Longrightarrow> 1 - \<epsilon> < Sigma_Algebra.measure (M m) {- real K<..real K}"
       by auto
@@ -504,7 +508,7 @@
       apply (rule max.strict_coboundedI2, auto)
     proof -
       fix n
-      show " 1 - \<epsilon> < measure (M n) {- max (real K) (4 / d)<..max (real K) (4 / d)}"      
+      show " 1 - \<epsilon> < measure (M n) {- max (real K) (4 / d)<..max (real K) (4 / d)}"
         apply (case_tac "n < N")
         apply (rule order_less_le_trans)
         apply (erule K)
@@ -512,7 +516,7 @@
         apply (rule order_less_le_trans)
         apply (rule 6, erule leI)
         by (rule Mn.finite_measure_mono, auto)
-    qed 
+    qed
     thus "\<exists>a b. a < b \<and> (\<forall>n. 1 - \<epsilon> < measure (M n) {a<..b})" by (intro exI)
   qed
   have tight: "tight M"
@@ -530,8 +534,8 @@
       by (subst 4, rule LIMSEQ_subseq_LIMSEQ [OF _ s], rule assms)
     hence "char \<nu> = char M'" by (intro ext, intro LIMSEQ_unique [OF 3 5])
     hence "\<nu> = M'" by (rule Levy_uniqueness [OF * `real_distribution M'`])
-    thus "weak_conv_m (M \<circ> s) M'" 
-      by (elim subst) (rule nu)  
+    thus "weak_conv_m (M \<circ> s) M'"
+      by (elim subst) (rule nu)
   qed
 qed
 
--- a/src/HOL/Probability/Measurable.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Measurable.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -72,7 +72,7 @@
 setup \<open>
   Global_Theory.add_thms_dynamic (@{binding measurable}, Measurable.get_all)
 \<close>
-  
+
 declare
   pred_sets1[measurable_dest]
   pred_sets2[measurable_dest]
@@ -97,7 +97,7 @@
 declare sets_restrict_space_cong[measurable_cong]
 declare sets_restrict_UNIV[measurable_cong]
 
-lemma predE[measurable (raw)]: 
+lemma predE[measurable (raw)]:
   "pred M P \<Longrightarrow> {x\<in>space M. P x} \<in> sets M"
   unfolding pred_def .
 
@@ -140,14 +140,14 @@
 
 lemma pred_intros_countable[measurable (raw)]:
   fixes P :: "'a \<Rightarrow> 'i :: countable \<Rightarrow> bool"
-  shows 
+  shows
     "(\<And>i. pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i. P x i)"
     "(\<And>i. pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<exists>i. P x i)"
   by (auto intro!: sets.sets_Collect_countable_All sets.sets_Collect_countable_Ex simp: pred_def)
 
 lemma pred_intros_countable_bounded[measurable (raw)]:
   fixes X :: "'i :: countable set"
-  shows 
+  shows
     "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Inter>i\<in>X. N x i))"
     "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Union>i\<in>X. N x i))"
     "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i\<in>X. P x i)"
@@ -270,7 +270,7 @@
   shows "(\<lambda>x. LEAST i. P i x) \<in> measurable M (count_space UNIV)"
   unfolding measurable_def by (safe intro!: sets_Least) simp_all
 
-lemma measurable_Max_nat[measurable (raw)]: 
+lemma measurable_Max_nat[measurable (raw)]:
   fixes P :: "nat \<Rightarrow> 'a \<Rightarrow> bool"
   assumes [measurable]: "\<And>i. Measurable.pred M (P i)"
   shows "(\<lambda>x. Max {i. P i x}) \<in> measurable M (count_space UNIV)"
@@ -294,8 +294,8 @@
 
   have "(\<lambda>x. Max {i. P i x}) -` {n} \<inter> space M = {x\<in>space M. Max {i. P i x} = n}"
     by auto
-  also have "\<dots> = 
-    {x\<in>space M. if (\<forall>i. \<exists>n\<ge>i. P n x) then the None = n else 
+  also have "\<dots> =
+    {x\<in>space M. if (\<forall>i. \<exists>n\<ge>i. P n x) then the None = n else
       if (\<exists>i. P i x) then P n x \<and> (\<forall>i>n. \<not> P i x)
       else Max {} = n}"
     by (intro arg_cong[where f=Collect] ext conj_cong)
@@ -305,7 +305,7 @@
   finally show "(\<lambda>x. Max {i. P i x}) -` {n} \<inter> space M \<in> sets M" .
 qed simp
 
-lemma measurable_Min_nat[measurable (raw)]: 
+lemma measurable_Min_nat[measurable (raw)]:
   fixes P :: "nat \<Rightarrow> 'a \<Rightarrow> bool"
   assumes [measurable]: "\<And>i. Measurable.pred M (P i)"
   shows "(\<lambda>x. Min {i. P i x}) \<in> measurable M (count_space UNIV)"
@@ -329,8 +329,8 @@
 
   have "(\<lambda>x. Min {i. P i x}) -` {n} \<inter> space M = {x\<in>space M. Min {i. P i x} = n}"
     by auto
-  also have "\<dots> = 
-    {x\<in>space M. if (\<forall>i. \<exists>n\<ge>i. P n x) then the None = n else 
+  also have "\<dots> =
+    {x\<in>space M. if (\<forall>i. \<exists>n\<ge>i. P n x) then the None = n else
       if (\<exists>i. P i x) then P n x \<and> (\<forall>i<n. \<not> P i x)
       else Min {} = n}"
     by (intro arg_cong[where f=Collect] ext conj_cong)
@@ -374,7 +374,7 @@
 
 lemma measurable_pred_countable[measurable (raw)]:
   assumes "countable X"
-  shows 
+  shows
     "(\<And>i. i \<in> X \<Longrightarrow> Measurable.pred M (\<lambda>x. P x i)) \<Longrightarrow> Measurable.pred M (\<lambda>x. \<forall>i\<in>X. P x i)"
     "(\<And>i. i \<in> X \<Longrightarrow> Measurable.pred M (\<lambda>x. P x i)) \<Longrightarrow> Measurable.pred M (\<lambda>x. \<exists>i\<in>X. P x i)"
   unfolding pred_def
@@ -395,7 +395,7 @@
   shows "(\<lambda>x. SUP i:I. F i x) \<in> measurable M (count_space UNIV)"
   unfolding measurable_count_space_eq2_countable
 proof (safe intro!: UNIV_I)
-  fix a 
+  fix a
   have "(\<lambda>x. SUP i:I. F i x) -` {a} \<inter> space M =
     {x\<in>space M. (\<forall>i\<in>I. F i x \<le> a) \<and> (\<forall>b. (\<forall>i\<in>I. F i x \<le> b) \<longrightarrow> a \<le> b)}"
     unfolding SUP_le_iff[symmetric] by auto
@@ -411,7 +411,7 @@
   shows "(\<lambda>x. INF i:I. F i x) \<in> measurable M (count_space UNIV)"
   unfolding measurable_count_space_eq2_countable
 proof (safe intro!: UNIV_I)
-  fix a 
+  fix a
   have "(\<lambda>x. INF i:I. F i x) -` {a} \<inter> space M =
     {x\<in>space M. (\<forall>i\<in>I. a \<le> F i x) \<and> (\<forall>b. (\<forall>i\<in>I. b \<le> F i x) \<longrightarrow> b \<le> a)}"
     unfolding le_INF_iff[symmetric] by auto
@@ -501,7 +501,7 @@
 lemma measurable_enat_coinduct:
   fixes f :: "'a \<Rightarrow> enat"
   assumes "R f"
-  assumes *: "\<And>f. R f \<Longrightarrow> \<exists>g h i P. R g \<and> f = (\<lambda>x. if P x then h x else eSuc (g (i x))) \<and> 
+  assumes *: "\<And>f. R f \<Longrightarrow> \<exists>g h i P. R g \<and> f = (\<lambda>x. if P x then h x else eSuc (g (i x))) \<and>
     Measurable.pred M P \<and>
     i \<in> measurable M M \<and>
     h \<in> measurable M (count_space UNIV)"
@@ -581,6 +581,26 @@
   shows "Measurable.pred M (\<lambda>x. \<exists>!i\<in>I. P i x)"
   unfolding bex1_def by measurable
 
+lemma measurable_Sup_nat[measurable (raw)]:
+  fixes F :: "'a \<Rightarrow> nat set"
+  assumes [measurable]: "\<And>i. Measurable.pred M (\<lambda>x. i \<in> F x)"
+  shows "(\<lambda>x. Sup (F x)) \<in> M \<rightarrow>\<^sub>M count_space UNIV"
+proof (clarsimp simp add: measurable_count_space_eq2_countable)
+  fix a
+  have F_empty_iff: "F x = {} \<longleftrightarrow> (\<forall>i. i \<notin> F x)" for x
+    by auto
+  have "Measurable.pred M (\<lambda>x. if finite (F x) then if F x = {} then a = Max {}
+    else a \<in> F x \<and> (\<forall>j. j \<in> F x \<longrightarrow> j \<le> a) else a = the None)"
+    unfolding finite_nat_set_iff_bounded Ball_def F_empty_iff by measurable
+  moreover have "(\<lambda>x. Sup (F x)) -` {a} \<inter> space M =
+    {x\<in>space M. if finite (F x) then if F x = {} then a = Max {}
+      else a \<in> F x \<and> (\<forall>j. j \<in> F x \<longrightarrow> j \<le> a) else a = the None}"
+    by (intro set_eqI)
+       (auto simp: Sup_nat_def Max.infinite intro!: Max_in Max_eqI)
+  ultimately show "(\<lambda>x. Sup (F x)) -` {a} \<inter> space M \<in> sets M"
+    by auto
+qed
+
 lemma measurable_if_split[measurable (raw)]:
   "(c \<Longrightarrow> Measurable.pred M f) \<Longrightarrow> (\<not> c \<Longrightarrow> Measurable.pred M g) \<Longrightarrow>
    Measurable.pred M (if c then f else g)"
--- a/src/HOL/Probability/Measure_Space.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Measure_Space.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -14,26 +14,26 @@
 subsection "Relate extended reals and the indicator function"
 
 lemma suminf_cmult_indicator:
-  fixes f :: "nat \<Rightarrow> ereal"
-  assumes "disjoint_family A" "x \<in> A i" "\<And>i. 0 \<le> f i"
+  fixes f :: "nat \<Rightarrow> ennreal"
+  assumes "disjoint_family A" "x \<in> A i"
   shows "(\<Sum>n. f n * indicator (A n) x) = f i"
 proof -
-  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ereal)"
+  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ennreal)"
     using \<open>x \<in> A i\<close> assms unfolding disjoint_family_on_def indicator_def by auto
-  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ereal)"
+  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ennreal)"
     by (auto simp: setsum.If_cases)
-  moreover have "(SUP n. if i < n then f i else 0) = (f i :: ereal)"
+  moreover have "(SUP n. if i < n then f i else 0) = (f i :: ennreal)"
   proof (rule SUP_eqI)
-    fix y :: ereal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
+    fix y :: ennreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
     from this[of "Suc i"] show "f i \<le> y" by auto
-  qed (insert assms, simp)
+  qed (insert assms, simp add: zero_le)
   ultimately show ?thesis using assms
-    by (subst suminf_ereal_eq_SUP) (auto simp: indicator_def)
+    by (subst suminf_eq_SUP) (auto simp: indicator_def)
 qed
 
 lemma suminf_indicator:
   assumes "disjoint_family A"
-  shows "(\<Sum>n. indicator (A n) x :: ereal) = indicator (\<Union>i. A i) x"
+  shows "(\<Sum>n. indicator (A n) x :: ennreal) = indicator (\<Union>i. A i) x"
 proof cases
   assume *: "x \<in> (\<Union>i. A i)"
   then obtain i where "x \<in> A i" by auto
@@ -98,6 +98,36 @@
   necessary to define @{typ "'a measure"} in @{theory Sigma_Algebra}.
 \<close>
 
+definition subadditive where
+  "subadditive M f \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
+
+lemma subadditiveD: "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) \<le> f x + f y"
+  by (auto simp add: subadditive_def)
+
+definition countably_subadditive where
+  "countably_subadditive M f \<longleftrightarrow>
+    (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))))"
+
+lemma (in ring_of_sets) countably_subadditive_subadditive:
+  fixes f :: "'a set \<Rightarrow> ennreal"
+  assumes f: "positive M f" and cs: "countably_subadditive M f"
+  shows  "subadditive M f"
+proof (auto simp add: subadditive_def)
+  fix x y
+  assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
+  hence "disjoint_family (binaryset x y)"
+    by (auto simp add: disjoint_family_on_def binaryset_def)
+  hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
+         (\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
+         f (\<Union>i. binaryset x y i) \<le> (\<Sum> n. f (binaryset x y n))"
+    using cs by (auto simp add: countably_subadditive_def)
+  hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
+         f (x \<union> y) \<le> (\<Sum> n. f (binaryset x y n))"
+    by (simp add: range_binaryset_eq UN_binaryset_eq)
+  thus "f (x \<union> y) \<le>  f x + f y" using f x y
+    by (auto simp add: Un o_def suminf_binaryset_eq positive_def)
+qed
+
 definition additive where
   "additive M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> \<mu> (x \<union> y) = \<mu> x + \<mu> y)"
 
@@ -105,7 +135,6 @@
   "increasing M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<subseteq> y \<longrightarrow> \<mu> x \<le> \<mu> y)"
 
 lemma positiveD1: "positive M f \<Longrightarrow> f {} = 0" by (auto simp: positive_def)
-lemma positiveD2: "positive M f \<Longrightarrow> A \<in> M \<Longrightarrow> 0 \<le> f A" by (auto simp: positive_def)
 
 lemma positiveD_empty:
   "positive M f \<Longrightarrow> f {} = 0"
@@ -162,14 +191,14 @@
 qed
 
 lemma (in ring_of_sets) additive_increasing:
+  fixes f :: "'a set \<Rightarrow> ennreal"
   assumes posf: "positive M f" and addf: "additive M f"
   shows "increasing M f"
 proof (auto simp add: increasing_def)
   fix x y
   assume xy: "x \<in> M" "y \<in> M" "x \<subseteq> y"
   then have "y - x \<in> M" by auto
-  then have "0 \<le> f (y-x)" using posf[unfolded positive_def] by auto
-  then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono) auto
+  then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono zero_le)
   also have "... = f (x \<union> (y-x))" using addf
     by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
   also have "... = f y"
@@ -178,9 +207,10 @@
 qed
 
 lemma (in ring_of_sets) subadditive:
-  assumes f: "positive M f" "additive M f" and A: "range A \<subseteq> M" and S: "finite S"
+  fixes f :: "'a set \<Rightarrow> ennreal"
+  assumes f: "positive M f" "additive M f" and A: "A`S \<subseteq> M" and S: "finite S"
   shows "f (\<Union>i\<in>S. A i) \<le> (\<Sum>i\<in>S. f (A i))"
-using S
+using S A
 proof (induct S)
   case empty thus ?case using f by (auto simp: positive_def)
 next
@@ -199,6 +229,7 @@
 qed
 
 lemma (in ring_of_sets) countably_additive_additive:
+  fixes f :: "'a set \<Rightarrow> ennreal"
   assumes posf: "positive M f" and ca: "countably_additive M f"
   shows "additive M f"
 proof (auto simp add: additive_def)
@@ -219,13 +250,13 @@
 qed
 
 lemma (in algebra) increasing_additive_bound:
-  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> ereal"
+  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> ennreal"
   assumes f: "positive M f" and ad: "additive M f"
       and inc: "increasing M f"
       and A: "range A \<subseteq> M"
       and disj: "disjoint_family A"
   shows  "(\<Sum>i. f (A i)) \<le> f \<Omega>"
-proof (safe intro!: suminf_bound)
+proof (safe intro!: suminf_le_const)
   fix N
   note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
   have "(\<Sum>i<N. f (A i)) = f (\<Union>i\<in>{..<N}. A i)"
@@ -236,6 +267,7 @@
 qed (insert f A, auto simp: positive_def)
 
 lemma (in ring_of_sets) countably_additiveI_finite:
+  fixes \<mu> :: "'a set \<Rightarrow> ennreal"
   assumes "finite \<Omega>" "positive M \<mu>" "additive M \<mu>"
   shows "countably_additive M \<mu>"
 proof (rule countably_additiveI)
@@ -279,6 +311,7 @@
 qed
 
 lemma (in ring_of_sets) countably_additive_iff_continuous_from_below:
+  fixes f :: "'a set \<Rightarrow> ennreal"
   assumes f: "positive M f" "additive M f"
   shows "countably_additive M f \<longleftrightarrow>
     (\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i))"
@@ -292,7 +325,7 @@
     by (auto simp: UN_disjointed_eq disjoint_family_disjointed)
   moreover have "(\<lambda>n. (\<Sum>i<n. f (disjointed A i))) \<longlonglongrightarrow> (\<Sum>i. f (disjointed A i))"
     using f(1)[unfolded positive_def] dA
-    by (auto intro!: summable_LIMSEQ summable_ereal_pos)
+    by (auto intro!: summable_LIMSEQ summableI)
   from LIMSEQ_Suc[OF this]
   have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) \<longlonglongrightarrow> (\<Sum>i. f (disjointed A i))"
     unfolding lessThan_Suc_atMost .
@@ -320,6 +353,7 @@
 qed
 
 lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous:
+  fixes f :: "'a set \<Rightarrow> ennreal"
   assumes f: "positive M f" "additive M f"
   shows "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Inter>i. A i))
      \<longleftrightarrow> (\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0)"
@@ -344,11 +378,11 @@
   have "f (\<Inter>x. A x) \<le> f (A 0)"
     using A by (auto intro!: f_mono)
   then have f_Int_fin: "f (\<Inter>x. A x) \<noteq> \<infinity>"
-    using A by auto
+    using A by (auto simp: top_unique)
   { fix i
     have "f (A i - (\<Inter>i. A i)) \<le> f (A i)" using A by (auto intro!: f_mono)
     then have "f (A i - (\<Inter>i. A i)) \<noteq> \<infinity>"
-      using A by auto }
+      using A by (auto simp: top_unique) }
   note f_fin = this
   have "(\<lambda>i. f (A i - (\<Inter>i. A i))) \<longlonglongrightarrow> 0"
   proof (intro cont[rule_format, OF _ decseq _ f_fin])
@@ -361,7 +395,7 @@
     by auto
   ultimately have "(INF n. f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i)) = 0 + f (\<Inter>i. A i)"
     using A(4) f_fin f_Int_fin
-    by (subst INF_ereal_add) (auto simp: decseq_f)
+    by (subst INF_ennreal_add_const) (auto simp: decseq_f)
   moreover {
     fix n
     have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f ((A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i))"
@@ -376,38 +410,31 @@
 qed
 
 lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below:
+  fixes f :: "'a set \<Rightarrow> ennreal"
   assumes f: "positive M f" "additive M f" "\<forall>A\<in>M. f A \<noteq> \<infinity>"
   assumes cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
   assumes A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M"
   shows "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
 proof -
-  have "\<forall>A\<in>M. \<exists>x. f A = ereal x"
-  proof
-    fix A assume "A \<in> M" with f show "\<exists>x. f A = ereal x"
-      unfolding positive_def by (cases "f A") auto
-  qed
-  from bchoice[OF this] guess f' .. note f' = this[rule_format]
   from A have "(\<lambda>i. f ((\<Union>i. A i) - A i)) \<longlonglongrightarrow> 0"
     by (intro cont[rule_format]) (auto simp: decseq_def incseq_def)
   moreover
   { fix i
-    have "f ((\<Union>i. A i) - A i) + f (A i) = f ((\<Union>i. A i) - A i \<union> A i)"
-      using A by (intro f(2)[THEN additiveD, symmetric]) auto
-    also have "(\<Union>i. A i) - A i \<union> A i = (\<Union>i. A i)"
+    have "f ((\<Union>i. A i) - A i \<union> A i) = f ((\<Union>i. A i) - A i) + f (A i)"
+      using A by (intro f(2)[THEN additiveD]) auto
+    also have "((\<Union>i. A i) - A i) \<union> A i = (\<Union>i. A i)"
       by auto
-    finally have "f' (\<Union>i. A i) - f' (A i) = f' ((\<Union>i. A i) - A i)"
-      using A by (subst (asm) (1 2 3) f') auto
-    then have "f ((\<Union>i. A i) - A i) = ereal (f' (\<Union>i. A i) - f' (A i))"
-      using A f' by auto }
-  ultimately have "(\<lambda>i. f' (\<Union>i. A i) - f' (A i)) \<longlonglongrightarrow> 0"
-    by (simp add: zero_ereal_def)
-  then have "(\<lambda>i. f' (A i)) \<longlonglongrightarrow> f' (\<Union>i. A i)"
-    by (rule Lim_transform2[OF tendsto_const])
-  then show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
-    using A by (subst (1 2) f') auto
+    finally have "f ((\<Union>i. A i) - A i) = f (\<Union>i. A i) - f (A i)"
+      using f(3)[rule_format, of "A i"] A by (auto simp: ennreal_add_diff_cancel subset_eq) }
+  moreover have "\<forall>\<^sub>F i in sequentially. f (A i) \<le> f (\<Union>i. A i)"
+    using increasingD[OF additive_increasing[OF f(1, 2)], of "A _" "\<Union>i. A i"] A
+    by (auto intro!: always_eventually simp: subset_eq)
+  ultimately show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
+    by (auto intro: ennreal_tendsto_const_minus)
 qed
 
 lemma (in ring_of_sets) empty_continuous_imp_countably_additive:
+  fixes f :: "'a set \<Rightarrow> ennreal"
   assumes f: "positive M f" "additive M f" and fin: "\<forall>A\<in>M. f A \<noteq> \<infinity>"
   assumes cont: "\<And>A. range A \<subseteq> M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
   shows "countably_additive M f"
@@ -423,21 +450,8 @@
 lemma emeasure_empty[simp, intro]: "emeasure M {} = 0"
   using emeasure_positive[of M] by (simp add: positive_def)
 
-lemma emeasure_nonneg[intro!]: "0 \<le> emeasure M A"
-  using emeasure_notin_sets[of A M] emeasure_positive[of M]
-  by (cases "A \<in> sets M") (auto simp: positive_def)
-
-lemma emeasure_not_MInf[simp]: "emeasure M A \<noteq> - \<infinity>"
-  using emeasure_nonneg[of M A] by auto
-
-lemma emeasure_le_0_iff: "emeasure M A \<le> 0 \<longleftrightarrow> emeasure M A = 0"
-  using emeasure_nonneg[of M A] by auto
-
-lemma emeasure_less_0_iff: "emeasure M A < 0 \<longleftrightarrow> False"
-  using emeasure_nonneg[of M A] by auto
-
 lemma emeasure_single_in_space: "emeasure M {x} \<noteq> 0 \<Longrightarrow> x \<in> space M"
-  using emeasure_notin_sets[of "{x}" M] by (auto dest: sets.sets_into_space)
+  using emeasure_notin_sets[of "{x}" M] by (auto dest: sets.sets_into_space zero_less_iff_neq_zero[THEN iffD2])
 
 lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)"
   by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
@@ -449,7 +463,7 @@
 
 lemma sums_emeasure:
   "disjoint_family F \<Longrightarrow> (\<And>i. F i \<in> sets M) \<Longrightarrow> (\<lambda>i. emeasure M (F i)) sums emeasure M (\<Union>i. F i)"
-  unfolding sums_iff by (intro conjI summable_ereal_pos emeasure_nonneg suminf_emeasure) auto
+  unfolding sums_iff by (intro conjI suminf_emeasure) auto
 
 lemma emeasure_additive: "additive (sets M) (emeasure M)"
   by (metis sets.countably_additive_additive emeasure_positive emeasure_countably_additive)
@@ -465,43 +479,29 @@
 
 lemma emeasure_mono:
   "a \<subseteq> b \<Longrightarrow> b \<in> sets M \<Longrightarrow> emeasure M a \<le> emeasure M b"
-  by (metis sets.additive_increasing emeasure_additive emeasure_nonneg emeasure_notin_sets
-            emeasure_positive increasingD)
+  by (metis zero_le sets.additive_increasing emeasure_additive emeasure_notin_sets emeasure_positive increasingD)
 
 lemma emeasure_space:
   "emeasure M A \<le> emeasure M (space M)"
-  by (metis emeasure_mono emeasure_nonneg emeasure_notin_sets sets.sets_into_space sets.top)
-
-lemma emeasure_compl:
-  assumes s: "s \<in> sets M" and fin: "emeasure M s \<noteq> \<infinity>"
-  shows "emeasure M (space M - s) = emeasure M (space M) - emeasure M s"
-proof -
-  from s have "0 \<le> emeasure M s" by auto
-  have "emeasure M (space M) = emeasure M (s \<union> (space M - s))" using s
-    by (metis Un_Diff_cancel Un_absorb1 s sets.sets_into_space)
-  also have "... = emeasure M s + emeasure M (space M - s)"
-    by (rule plus_emeasure[symmetric]) (auto simp add: s)
-  finally have "emeasure M (space M) = emeasure M s + emeasure M (space M - s)" .
-  then show ?thesis
-    using fin \<open>0 \<le> emeasure M s\<close>
-    unfolding ereal_eq_minus_iff by (auto simp: ac_simps)
-qed
+  by (metis emeasure_mono emeasure_notin_sets sets.sets_into_space sets.top zero_le)
 
 lemma emeasure_Diff:
   assumes finite: "emeasure M B \<noteq> \<infinity>"
   and [measurable]: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
   shows "emeasure M (A - B) = emeasure M A - emeasure M B"
 proof -
-  have "0 \<le> emeasure M B" using assms by auto
   have "(A - B) \<union> B = A" using \<open>B \<subseteq> A\<close> by auto
   then have "emeasure M A = emeasure M ((A - B) \<union> B)" by simp
   also have "\<dots> = emeasure M (A - B) + emeasure M B"
     by (subst plus_emeasure[symmetric]) auto
   finally show "emeasure M (A - B) = emeasure M A - emeasure M B"
-    unfolding ereal_eq_minus_iff
-    using finite \<open>0 \<le> emeasure M B\<close> by auto
+    using finite by simp
 qed
 
+lemma emeasure_compl:
+  "s \<in> sets M \<Longrightarrow> emeasure M s \<noteq> \<infinity> \<Longrightarrow> emeasure M (space M - s) = emeasure M (space M) - emeasure M s"
+  by (rule emeasure_Diff) (auto dest: sets.sets_into_space)
+
 lemma Lim_emeasure_incseq:
   "range A \<subseteq> sets M \<Longrightarrow> incseq A \<Longrightarrow> (\<lambda>i. (emeasure M (A i))) \<longlonglongrightarrow> emeasure M (\<Union>i. A i)"
   using emeasure_countably_additive
@@ -531,15 +531,10 @@
 proof -
   have le_MI: "emeasure M (\<Inter>i. A i) \<le> emeasure M (A 0)"
     using A by (auto intro!: emeasure_mono)
-  hence *: "emeasure M (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by auto
-
-  have A0: "0 \<le> emeasure M (A 0)" using A by auto
+  hence *: "emeasure M (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by (auto simp: top_unique)
 
-  have "emeasure M (A 0) - (INF n. emeasure M (A n)) = emeasure M (A 0) + (SUP n. - emeasure M (A n))"
-    by (simp add: ereal_SUP_uminus minus_ereal_def)
-  also have "\<dots> = (SUP n. emeasure M (A 0) - emeasure M (A n))"
-    unfolding minus_ereal_def using A0 assms
-    by (subst SUP_ereal_add) (auto simp add: decseq_emeasure)
+  have "emeasure M (A 0) - (INF n. emeasure M (A n)) = (SUP n. emeasure M (A 0) - emeasure M (A n))"
+    by (simp add: ennreal_INF_const_minus)
   also have "\<dots> = (SUP n. emeasure M (A 0 - A n))"
     using A finite \<open>decseq A\<close>[unfolded decseq_def] by (subst emeasure_Diff) auto
   also have "\<dots> = emeasure M (\<Union>i. A 0 - A i)"
@@ -552,7 +547,8 @@
   also have "\<dots> = emeasure M (A 0) - emeasure M (\<Inter>i. A i)"
     using A finite * by (simp, subst emeasure_Diff) auto
   finally show ?thesis
-    unfolding ereal_minus_eq_minus_iff using finite A0 by auto
+    by (rule ennreal_minus_cancel[rotated 3])
+       (insert finite A, auto intro: INF_lower emeasure_mono)
 qed
 
 lemma emeasure_INT_decseq_subset:
@@ -637,7 +633,6 @@
 lemma emeasure_lfp:
   assumes [simp]: "\<And>s. sets (M s) = sets N"
   assumes cont: "sup_continuous F" "sup_continuous f"
-  assumes nonneg: "\<And>P s. 0 \<le> f P s"
   assumes meas: "\<And>P. Measurable.pred N P \<Longrightarrow> Measurable.pred N (F P)"
   assumes iter: "\<And>P s. Measurable.pred N P \<Longrightarrow> P \<le> lfp F \<Longrightarrow> emeasure (M s) {x\<in>space N. F P x} = f (\<lambda>s. emeasure (M s) {x\<in>space N. P x}) s"
   shows "emeasure (M s) {x\<in>space N. lfp F x} = lfp f s"
@@ -646,34 +641,15 @@
   then show "(\<lambda>s. emeasure (M s) {x \<in> space N. (SUP i. C i) x}) = (SUP i. (\<lambda>s. emeasure (M s) {x \<in> space N. C i x}))"
     unfolding SUP_apply[abs_def]
     by (subst SUP_emeasure_incseq) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure])
-qed (auto simp add: iter nonneg le_fun_def SUP_apply[abs_def] intro!: meas cont)
-
-lemma emeasure_subadditive:
-  assumes [measurable]: "A \<in> sets M" "B \<in> sets M"
-  shows "emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
-proof -
-  from plus_emeasure[of A M "B - A"]
-  have "emeasure M (A \<union> B) = emeasure M A + emeasure M (B - A)" by simp
-  also have "\<dots> \<le> emeasure M A + emeasure M B"
-    using assms by (auto intro!: add_left_mono emeasure_mono)
-  finally show ?thesis .
-qed
+qed (auto simp add: iter le_fun_def SUP_apply[abs_def] intro!: meas cont)
 
 lemma emeasure_subadditive_finite:
-  assumes "finite I" "A ` I \<subseteq> sets M"
-  shows "emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
-using assms proof induct
-  case (insert i I)
-  then have "emeasure M (\<Union>i\<in>insert i I. A i) = emeasure M (A i \<union> (\<Union>i\<in>I. A i))"
-    by simp
-  also have "\<dots> \<le> emeasure M (A i) + emeasure M (\<Union>i\<in>I. A i)"
-    using insert by (intro emeasure_subadditive) auto
-  also have "\<dots> \<le> emeasure M (A i) + (\<Sum>i\<in>I. emeasure M (A i))"
-    using insert by (intro add_mono) auto
-  also have "\<dots> = (\<Sum>i\<in>insert i I. emeasure M (A i))"
-    using insert by auto
-  finally show ?case .
-qed simp
+  "finite I \<Longrightarrow> A ` I \<subseteq> sets M \<Longrightarrow> emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
+  by (rule sets.subadditive[OF emeasure_positive emeasure_additive]) auto
+
+lemma emeasure_subadditive:
+  "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
+  using emeasure_subadditive_finite[of "{True, False}" "\<lambda>True \<Rightarrow> A | False \<Rightarrow> B" M] by simp
 
 lemma emeasure_subadditive_countably:
   assumes "range f \<subseteq> sets M"
@@ -686,7 +662,7 @@
     by (simp add:  disjoint_family_disjointed comp_def)
   also have "\<dots> \<le> (\<Sum>i. emeasure M (f i))"
     using sets.range_disjointed_sets[OF assms] assms
-    by (auto intro!: suminf_le_pos emeasure_mono disjointed_subset)
+    by (auto intro!: suminf_le emeasure_mono disjointed_subset)
   finally show ?thesis .
 qed
 
@@ -727,16 +703,16 @@
 
 lemma emeasure_eq_0:
   "N \<in> sets M \<Longrightarrow> emeasure M N = 0 \<Longrightarrow> K \<subseteq> N \<Longrightarrow> emeasure M K = 0"
-  by (metis emeasure_mono emeasure_nonneg order_eq_iff)
+  by (metis emeasure_mono order_eq_iff zero_le)
 
 lemma emeasure_UN_eq_0:
   assumes "\<And>i::nat. emeasure M (N i) = 0" and "range N \<subseteq> sets M"
   shows "emeasure M (\<Union>i. N i) = 0"
 proof -
-  have "0 \<le> emeasure M (\<Union>i. N i)" using assms by auto
-  moreover have "emeasure M (\<Union>i. N i) \<le> 0"
+  have "emeasure M (\<Union>i. N i) \<le> 0"
     using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp
-  ultimately show ?thesis by simp
+  then show ?thesis
+    by (auto intro: antisym zero_le)
 qed
 
 lemma measure_eqI_finite:
@@ -788,9 +764,9 @@
         and [intro]: "F \<inter> A \<in> sigma_sets \<Omega> E"
         using \<open>F \<in> E\<close> S.sets_into_space by (auto simp: M)
       have "?\<nu> (F \<inter> A) \<le> ?\<nu> F" by (auto intro!: emeasure_mono simp: M N)
-      then have "?\<nu> (F \<inter> A) \<noteq> \<infinity>" using \<open>?\<nu> F \<noteq> \<infinity>\<close> by auto
+      then have "?\<nu> (F \<inter> A) \<noteq> \<infinity>" using \<open>?\<nu> F \<noteq> \<infinity>\<close> by (auto simp: top_unique)
       have "?\<mu> (F \<inter> A) \<le> ?\<mu> F" by (auto intro!: emeasure_mono simp: M N)
-      then have "?\<mu> (F \<inter> A) \<noteq> \<infinity>" using \<open>?\<mu> F \<noteq> \<infinity>\<close> by auto
+      then have "?\<mu> (F \<inter> A) \<noteq> \<infinity>" using \<open>?\<mu> F \<noteq> \<infinity>\<close> by (auto simp: top_unique)
       then have "?\<mu> (F \<inter> (\<Omega> - A)) = ?\<mu> F - ?\<mu> (F \<inter> A)" unfolding **
         using \<open>F \<inter> A \<in> sigma_sets \<Omega> E\<close> by (auto intro!: emeasure_Diff simp: M N)
       also have "\<dots> = ?\<nu> F - ?\<nu> (F \<inter> A)" using eq \<open>F \<in> E\<close> compl by simp
@@ -840,7 +816,7 @@
 proof (intro measure_eqI emeasure_measure_of_sigma)
   show "sigma_algebra (space M) (sets M)" ..
   show "positive (sets M) (emeasure M)"
-    by (simp add: positive_def emeasure_nonneg)
+    by (simp add: positive_def)
   show "countably_additive (sets M) (emeasure M)"
     by (simp add: emeasure_countably_additive)
 qed simp_all
@@ -874,7 +850,7 @@
   then have "emeasure M B = 0" "emeasure M A = 0"
     using null_sets by auto
   with sets * show "A - B \<in> null_sets M" "A \<union> B \<in> null_sets M"
-    by (auto intro!: antisym)
+    by (auto intro!: antisym zero_le)
 qed
 
 lemma UN_from_nat_into:
@@ -906,7 +882,7 @@
     also have "(\<lambda>n. emeasure M (N (from_nat_into I n))) = (\<lambda>_. 0)"
       using assms \<open>I \<noteq> {}\<close> by (auto intro: from_nat_into)
     finally show "emeasure M (\<Union>i\<in>I. N i) = 0"
-      by (intro antisym emeasure_nonneg) simp
+      by (intro antisym zero_le) simp
   qed
 qed
 
@@ -965,7 +941,7 @@
   "_almost_everywhere" :: "pttrn \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
 
 translations
-  "AE x in M. P" == "CONST almost_everywhere M (\<lambda>x. P)"
+  "AE x in M. P" \<rightleftharpoons> "CONST almost_everywhere M (\<lambda>x. P)"
 
 lemma eventually_ae_filter: "eventually P (ae_filter M) \<longleftrightarrow> (\<exists>N\<in>null_sets M. {x \<in> space M. \<not> P x} \<subseteq> N)"
   unfolding ae_filter_def by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq)
@@ -980,10 +956,10 @@
 proof
   assume "AE x in M. P x" then obtain N where N: "N \<in> sets M" "?P \<subseteq> N" "emeasure M N = 0"
     unfolding eventually_ae_filter by auto
-  have "0 \<le> emeasure M ?P" by auto
-  moreover have "emeasure M ?P \<le> emeasure M N"
+  have "emeasure M ?P \<le> emeasure M N"
     using assms N(1,2) by (auto intro: emeasure_mono)
-  ultimately have "emeasure M ?P = 0" unfolding \<open>emeasure M N = 0\<close> by auto
+  then have "emeasure M ?P = 0"
+    unfolding \<open>emeasure M N = 0\<close> by auto
   then show "?P \<in> null_sets M" using assms by auto
 next
   assume "?P \<in> null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I')
@@ -1034,10 +1010,10 @@
 
   show ?thesis
   proof (intro AE_I)
-    have "0 \<le> emeasure M (A \<union> B)" using A B by auto
-    moreover have "emeasure M (A \<union> B) \<le> 0"
+    have "emeasure M (A \<union> B) \<le> 0"
       using emeasure_subadditive[of A M B] A B by auto
-    ultimately show "A \<union> B \<in> sets M" "emeasure M (A \<union> B) = 0" using A B by auto
+    then show "A \<union> B \<in> sets M" "emeasure M (A \<union> B) = 0"
+      using A B by auto
     show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B"
       using P imp by auto
   qed
@@ -1155,7 +1131,7 @@
   also have "emeasure M (B - N) = emeasure M B"
     using N B by (subst emeasure_Diff_null_set) auto
   finally show ?thesis .
-qed (simp add: emeasure_nonneg emeasure_notin_sets)
+qed (simp add: emeasure_notin_sets)
 
 lemma emeasure_eq_AE:
   assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
@@ -1241,7 +1217,7 @@
     proof -
       have "emeasure M (disjointed A i) \<le> emeasure M (A i)"
         using range disjointed_subset[of A i] by (auto intro!: emeasure_mono)
-      then show ?thesis using measure[of i] by auto
+      then show ?thesis using measure[of i] by (auto simp: top_unique)
     qed
   qed
 qed
@@ -1267,7 +1243,7 @@
       have "emeasure M (\<Union>i\<le>n. F i) \<le> (\<Sum>i\<le>n. emeasure M (F i))"
         using F by (auto intro!: emeasure_subadditive_finite)
       also have "\<dots> < \<infinity>"
-        using F by (auto simp: setsum_Pinfty)
+        using F by (auto simp: setsum_Pinfty less_top)
       finally show ?thesis by simp
     qed
     show "incseq (\<lambda>n. \<Union>i\<le>n. F i)"
@@ -1407,41 +1383,58 @@
 
 subsection \<open>Real measure values\<close>
 
-lemma measure_nonneg: "0 \<le> measure M A"
-  using emeasure_nonneg[of M A] unfolding measure_def by (auto intro: real_of_ereal_pos)
+lemma ring_of_finite_sets: "ring_of_sets (space M) {A\<in>sets M. emeasure M A \<noteq> top}"
+proof (rule ring_of_setsI)
+  show "a \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow> b \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow>
+    a \<union> b \<in> {A \<in> sets M. emeasure M A \<noteq> top}" for a b
+    using emeasure_subadditive[of a M b] by (auto simp: top_unique)
+
+  show "a \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow> b \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow>
+    a - b \<in> {A \<in> sets M. emeasure M A \<noteq> top}" for a b
+    using emeasure_mono[of "a - b" a M] by (auto simp: Diff_subset top_unique)
+qed (auto dest: sets.sets_into_space)
+
+lemma measure_nonneg[simp]: "0 \<le> measure M A"
+  unfolding measure_def by (auto simp: enn2real_nonneg)
 
 lemma zero_less_measure_iff: "0 < measure M A \<longleftrightarrow> measure M A \<noteq> 0"
   using measure_nonneg[of M A] by (auto simp add: le_less)
 
 lemma measure_le_0_iff: "measure M X \<le> 0 \<longleftrightarrow> measure M X = 0"
-  using measure_nonneg[of M X] by auto
+  using measure_nonneg[of M X] by linarith
 
 lemma measure_empty[simp]: "measure M {} = 0"
-  unfolding measure_def by simp
+  unfolding measure_def by (simp add: zero_ennreal.rep_eq)
+
+lemma emeasure_eq_ennreal_measure:
+  "emeasure M A \<noteq> top \<Longrightarrow> emeasure M A = ennreal (measure M A)"
+  by (cases "emeasure M A" rule: ennreal_cases) (auto simp: measure_def)
 
-lemma emeasure_eq_ereal_measure:
-  "emeasure M A \<noteq> \<infinity> \<Longrightarrow> emeasure M A = ereal (measure M A)"
-  using emeasure_nonneg[of M A]
-  by (cases "emeasure M A") (auto simp: measure_def)
+lemma measure_zero_top: "emeasure M A = top \<Longrightarrow> measure M A = 0"
+  by (simp add: measure_def enn2ereal_top)
 
-lemma max_0_ereal_measure [simp]: "max 0 (ereal (measure M X)) = ereal (measure M X)"
-by(simp add: max_def measure_nonneg)
+lemma measure_eq_emeasure_eq_ennreal: "0 \<le> x \<Longrightarrow> emeasure M A = ennreal x \<Longrightarrow> measure M A = x"
+  using emeasure_eq_ennreal_measure[of M A]
+  by (cases "A \<in> M") (auto simp: measure_notin_sets emeasure_notin_sets)
+
+lemma enn2real_plus:"a < top \<Longrightarrow> b < top \<Longrightarrow> enn2real (a + b) = enn2real a + enn2real b"
+  by (simp add: enn2real_def plus_ennreal.rep_eq real_of_ereal_add enn2ereal_nonneg less_top
+           del: real_of_ereal_enn2ereal)
 
 lemma measure_Union:
-  assumes finite: "emeasure M A \<noteq> \<infinity>" "emeasure M B \<noteq> \<infinity>"
-  and measurable: "A \<in> sets M" "B \<in> sets M" "A \<inter> B = {}"
-  shows "measure M (A \<union> B) = measure M A + measure M B"
-  unfolding measure_def
-  using plus_emeasure[OF measurable, symmetric] finite
-  by (simp add: emeasure_eq_ereal_measure)
+  "emeasure M A \<noteq> \<infinity> \<Longrightarrow> emeasure M B \<noteq> \<infinity> \<Longrightarrow> A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow>
+    measure M (A \<union> B) = measure M A + measure M B"
+  by (simp add: measure_def enn2ereal_nonneg plus_emeasure[symmetric] enn2real_plus less_top)
+
+lemma disjoint_family_on_insert:
+  "i \<notin> I \<Longrightarrow> disjoint_family_on A (insert i I) \<longleftrightarrow> A i \<inter> (\<Union>i\<in>I. A i) = {} \<and> disjoint_family_on A I"
+  by (fastforce simp: disjoint_family_on_def)
 
 lemma measure_finite_Union:
-  assumes measurable: "A`S \<subseteq> sets M" "disjoint_family_on A S" "finite S"
-  assumes finite: "\<And>i. i \<in> S \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>"
-  shows "measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
-  unfolding measure_def
-  using setsum_emeasure[OF measurable, symmetric] finite
-  by (simp add: emeasure_eq_ereal_measure)
+  "finite S \<Longrightarrow> A`S \<subseteq> sets M \<Longrightarrow> disjoint_family_on A S \<Longrightarrow> (\<And>i. i \<in> S \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>) \<Longrightarrow>
+    measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
+  by (induction S rule: finite_induct)
+     (auto simp: disjoint_family_on_insert measure_Union setsum_emeasure[symmetric] sets.countable_UN'[OF countable_finite])
 
 lemma measure_Diff:
   assumes finite: "emeasure M A \<noteq> \<infinity>"
@@ -1451,7 +1444,7 @@
   have "emeasure M (A - B) \<le> emeasure M A" "emeasure M B \<le> emeasure M A"
     using measurable by (auto intro!: emeasure_mono)
   hence "measure M ((A - B) \<union> B) = measure M (A - B) + measure M B"
-    using measurable finite by (rule_tac measure_Union) auto
+    using measurable finite by (rule_tac measure_Union) (auto simp: top_unique)
   thus ?thesis using \<open>B \<subseteq> A\<close> by (auto simp: Un_absorb2)
 qed
 
@@ -1460,29 +1453,32 @@
   assumes finite: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
   shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
 proof -
-  from summable_sums[OF summable_ereal_pos, of "\<lambda>i. emeasure M (A i)"]
-       suminf_emeasure[OF measurable] emeasure_nonneg[of M]
-  have "(\<lambda>i. emeasure M (A i)) sums (emeasure M (\<Union>i. A i))" by simp
+  have "(\<lambda>i. emeasure M (A i)) sums (emeasure M (\<Union>i. A i))"
+    unfolding suminf_emeasure[OF measurable, symmetric] by (simp add: summable_sums)
   moreover
   { fix i
     have "emeasure M (A i) \<le> emeasure M (\<Union>i. A i)"
       using measurable by (auto intro!: emeasure_mono)
-    then have "emeasure M (A i) = ereal ((measure M (A i)))"
-      using finite by (intro emeasure_eq_ereal_measure) auto }
+    then have "emeasure M (A i) = ennreal ((measure M (A i)))"
+      using finite by (intro emeasure_eq_ennreal_measure) (auto simp: top_unique) }
   ultimately show ?thesis using finite
-    unfolding sums_ereal[symmetric] by (simp add: emeasure_eq_ereal_measure)
+    by (subst (asm) (2) emeasure_eq_ennreal_measure)
+       (simp_all add: measure_nonneg)
 qed
 
 lemma measure_subadditive:
   assumes measurable: "A \<in> sets M" "B \<in> sets M"
   and fin: "emeasure M A \<noteq> \<infinity>" "emeasure M B \<noteq> \<infinity>"
-  shows "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)"
+  shows "measure M (A \<union> B) \<le> measure M A + measure M B"
 proof -
   have "emeasure M (A \<union> B) \<noteq> \<infinity>"
-    using emeasure_subadditive[OF measurable] fin by auto
+    using emeasure_subadditive[OF measurable] fin by (auto simp: top_unique)
   then show "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)"
     using emeasure_subadditive[OF measurable] fin
-    by (auto simp: emeasure_eq_ereal_measure)
+    apply simp
+    apply (subst (asm) (2 3 4) emeasure_eq_ennreal_measure)
+    apply (auto simp add: ennreal_plus[symmetric] simp del: ennreal_plus)
+    done
 qed
 
 lemma measure_subadditive_finite:
@@ -1492,82 +1488,90 @@
   { have "emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
       using emeasure_subadditive_finite[OF A] .
     also have "\<dots> < \<infinity>"
-      using fin by (simp add: setsum_Pinfty)
-    finally have "emeasure M (\<Union>i\<in>I. A i) \<noteq> \<infinity>" by simp }
-  then show ?thesis
+      using fin by (simp add: less_top A)
+    finally have "emeasure M (\<Union>i\<in>I. A i) \<noteq> top" by simp }
+  note * = this
+  show ?thesis
     using emeasure_subadditive_finite[OF A] fin
-    unfolding measure_def by (simp add: emeasure_eq_ereal_measure suminf_ereal measure_nonneg)
+    unfolding emeasure_eq_ennreal_measure[OF *]
+    by (simp_all add: setsum_ennreal measure_nonneg setsum_nonneg emeasure_eq_ennreal_measure)
 qed
 
 lemma measure_subadditive_countably:
   assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. emeasure M (A i)) \<noteq> \<infinity>"
   shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
 proof -
-  from emeasure_nonneg fin have "\<And>i. emeasure M (A i) \<noteq> \<infinity>" by (rule suminf_PInfty)
-  moreover
+  from fin have **: "\<And>i. emeasure M (A i) \<noteq> top"
+    using ennreal_suminf_lessD[of "\<lambda>i. emeasure M (A i)"] by (simp add: less_top)
   { have "emeasure M (\<Union>i. A i) \<le> (\<Sum>i. emeasure M (A i))"
       using emeasure_subadditive_countably[OF A] .
     also have "\<dots> < \<infinity>"
-      using fin by simp
-    finally have "emeasure M (\<Union>i. A i) \<noteq> \<infinity>" by simp }
-  ultimately  show ?thesis
-    using emeasure_subadditive_countably[OF A] fin
-    unfolding measure_def by (simp add: emeasure_eq_ereal_measure suminf_ereal measure_nonneg)
+      using fin by (simp add: less_top)
+    finally have "emeasure M (\<Union>i. A i) \<noteq> top" by simp }
+  then have "ennreal (measure M (\<Union>i. A i)) = emeasure M (\<Union>i. A i)"
+    by (rule emeasure_eq_ennreal_measure[symmetric])
+  also have "\<dots> \<le> (\<Sum>i. emeasure M (A i))"
+    using emeasure_subadditive_countably[OF A] .
+  also have "\<dots> = ennreal (\<Sum>i. measure M (A i))"
+    using fin unfolding emeasure_eq_ennreal_measure[OF **]
+    by (subst suminf_ennreal) (auto simp: **)
+  finally show ?thesis
+    apply (rule ennreal_le_iff[THEN iffD1, rotated])
+    apply (intro suminf_nonneg allI measure_nonneg summable_suminf_not_top)
+    using fin
+    apply (simp add: emeasure_eq_ennreal_measure[OF **])
+    done
 qed
 
 lemma measure_eq_setsum_singleton:
-  assumes S: "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
-  and fin: "\<And>x. x \<in> S \<Longrightarrow> emeasure M {x} \<noteq> \<infinity>"
-  shows "(measure M S) = (\<Sum>x\<in>S. (measure M {x}))"
-  unfolding measure_def
-  using emeasure_eq_setsum_singleton[OF S] fin
-  by simp (simp add: emeasure_eq_ereal_measure)
+  "finite S \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M) \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> emeasure M {x} \<noteq> \<infinity>) \<Longrightarrow>
+    measure M S = (\<Sum>x\<in>S. measure M {x})"
+  using emeasure_eq_setsum_singleton[of S M]
+  by (intro measure_eq_emeasure_eq_ennreal) (auto simp: setsum_nonneg emeasure_eq_ennreal_measure)
 
 lemma Lim_measure_incseq:
   assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
-  shows "(\<lambda>i. (measure M (A i))) \<longlonglongrightarrow> (measure M (\<Union>i. A i))"
-proof -
-  have "ereal ((measure M (\<Union>i. A i))) = emeasure M (\<Union>i. A i)"
-    using fin by (auto simp: emeasure_eq_ereal_measure)
-  then show ?thesis
-    using Lim_emeasure_incseq[OF A]
-    unfolding measure_def
-    by (intro lim_real_of_ereal) simp
-qed
+  shows "(\<lambda>i. measure M (A i)) \<longlonglongrightarrow> measure M (\<Union>i. A i)"
+proof (rule tendsto_ennrealD)
+  have "ennreal (measure M (\<Union>i. A i)) = emeasure M (\<Union>i. A i)"
+    using fin by (auto simp: emeasure_eq_ennreal_measure)
+  moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i
+    using assms emeasure_mono[of "A _" "\<Union>i. A i" M]
+    by (intro emeasure_eq_ennreal_measure[symmetric]) (auto simp: less_top UN_upper intro: le_less_trans)
+  ultimately show "(\<lambda>x. ennreal (Sigma_Algebra.measure M (A x))) \<longlonglongrightarrow> ennreal (Sigma_Algebra.measure M (\<Union>i. A i))"
+    using A by (auto intro!: Lim_emeasure_incseq)
+qed auto
 
 lemma Lim_measure_decseq:
   assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
   shows "(\<lambda>n. measure M (A n)) \<longlonglongrightarrow> measure M (\<Inter>i. A i)"
-proof -
-  have "emeasure M (\<Inter>i. A i) \<le> emeasure M (A 0)"
-    using A by (auto intro!: emeasure_mono)
-  also have "\<dots> < \<infinity>"
-    using fin[of 0] by auto
-  finally have "ereal ((measure M (\<Inter>i. A i))) = emeasure M (\<Inter>i. A i)"
-    by (auto simp: emeasure_eq_ereal_measure)
-  then show ?thesis
-    unfolding measure_def
-    using Lim_emeasure_decseq[OF A fin]
-    by (intro lim_real_of_ereal) simp
-qed
+proof (rule tendsto_ennrealD)
+  have "ennreal (measure M (\<Inter>i. A i)) = emeasure M (\<Inter>i. A i)"
+    using fin[of 0] A emeasure_mono[of "\<Inter>i. A i" "A 0" M]
+    by (auto intro!: emeasure_eq_ennreal_measure[symmetric] simp: INT_lower less_top intro: le_less_trans)
+  moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i
+    using A fin[of i] by (intro emeasure_eq_ennreal_measure[symmetric]) auto
+  ultimately show "(\<lambda>x. ennreal (Sigma_Algebra.measure M (A x))) \<longlonglongrightarrow> ennreal (Sigma_Algebra.measure M (\<Inter>i. A i))"
+    using fin A by (auto intro!: Lim_emeasure_decseq)
+qed auto
 
 subsection \<open>Measure spaces with @{term "emeasure M (space M) < \<infinity>"}\<close>
 
 locale finite_measure = sigma_finite_measure M for M +
-  assumes finite_emeasure_space: "emeasure M (space M) \<noteq> \<infinity>"
+  assumes finite_emeasure_space: "emeasure M (space M) \<noteq> top"
 
 lemma finite_measureI[Pure.intro!]:
   "emeasure M (space M) \<noteq> \<infinity> \<Longrightarrow> finite_measure M"
   proof qed (auto intro!: exI[of _ "{space M}"])
 
-lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \<noteq> \<infinity>"
-  using finite_emeasure_space emeasure_space[of M A] by auto
+lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \<noteq> top"
+  using finite_emeasure_space emeasure_space[of M A] by (auto simp: top_unique)
 
-lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ereal (measure M A)"
-  unfolding measure_def by (simp add: emeasure_eq_ereal_measure)
+lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ennreal (measure M A)"
+  by (intro emeasure_eq_ennreal_measure) simp
 
-lemma (in finite_measure) emeasure_real: "\<exists>r. 0 \<le> r \<and> emeasure M A = ereal r"
-  using emeasure_finite[of A] emeasure_nonneg[of M A] by (cases "emeasure M A") auto
+lemma (in finite_measure) emeasure_real: "\<exists>r. 0 \<le> r \<and> emeasure M A = ennreal r"
+  using emeasure_finite[of A] by (cases "emeasure M A" rule: ennreal_cases) auto
 
 lemma (in finite_measure) bounded_measure: "measure M A \<le> measure M (space M)"
   using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def)
@@ -1583,7 +1587,7 @@
   using measure_Union[OF _ _ assms] by simp
 
 lemma (in finite_measure) finite_measure_finite_Union:
-  assumes measurable: "A`S \<subseteq> sets M" "disjoint_family_on A S" "finite S"
+  assumes measurable: "finite S" "A`S \<subseteq> sets M" "disjoint_family_on A S"
   shows "measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
   using measure_finite_Union[OF assms] by simp
 
@@ -1606,16 +1610,9 @@
   using measure_subadditive_finite[OF assms] by simp
 
 lemma (in finite_measure) finite_measure_subadditive_countably:
-  assumes A: "range A \<subseteq> sets M" and sum: "summable (\<lambda>i. measure M (A i))"
-  shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
-proof -
-  from \<open>summable (\<lambda>i. measure M (A i))\<close>
-  have "(\<lambda>i. ereal (measure M (A i))) sums ereal (\<Sum>i. measure M (A i))"
-    by (simp add: sums_ereal) (rule summable_sums)
-  from sums_unique[OF this, symmetric]
-       measure_subadditive_countably[OF A]
-  show ?thesis by (simp add: emeasure_eq_measure)
-qed
+  "range A \<subseteq> sets M \<Longrightarrow> summable (\<lambda>i. measure M (A i)) \<Longrightarrow> measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
+  by (rule measure_subadditive_countably)
+     (simp_all add: ennreal_suminf_neq_top emeasure_eq_measure)
 
 lemma (in finite_measure) finite_measure_eq_setsum_singleton:
   assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
@@ -1691,7 +1688,7 @@
 proof (rule antisym)
   show "measure M (\<Union>i :: nat. c i) \<le> 0"
     using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2))
-qed (simp add: measure_nonneg)
+qed simp
 
 lemma (in finite_measure) measure_space_inter:
   assumes events:"s \<in> sets M" "t \<in> sets M"
@@ -1790,7 +1787,7 @@
 qed
 
 lemma emeasure_count_space:
-  assumes "X \<subseteq> A" shows "emeasure (count_space A) X = (if finite X then ereal (card X) else \<infinity>)"
+  assumes "X \<subseteq> A" shows "emeasure (count_space A) X = (if finite X then of_nat (card X) else \<infinity>)"
     (is "_ = ?M X")
   unfolding count_space_def
 proof (rule emeasure_measure_of_sigma)
@@ -1828,9 +1825,9 @@
       then have "(\<lambda>i. ?M (F i)) \<longlonglongrightarrow> (SUP n. ?M (F n))"
         by (rule LIMSEQ_SUP)
 
-      moreover have "(SUP n. ?M (F n)) = \<infinity>"
-      proof (rule SUP_PInfty)
-        fix n :: nat show "\<exists>k::nat\<in>UNIV. ereal n \<le> ?M (F k)"
+      moreover have "(SUP n. ?M (F n)) = top"
+      proof (rule ennreal_SUP_eq_top)
+        fix n :: nat show "\<exists>k::nat\<in>UNIV. of_nat n \<le> ?M (F k)"
         proof (induct n)
           case (Suc n)
           then guess k .. note k = this
@@ -1839,7 +1836,7 @@
           moreover have "finite (F (f k)) \<Longrightarrow> finite (F k)"
             using \<open>k \<le> f k\<close> \<open>incseq F\<close> by (auto simp: incseq_def dest: finite_subset)
           ultimately show ?case
-            by (auto intro!: exI[of _ "f k"])
+            by (auto intro!: exI[of _ "f k"] simp del: of_nat_Suc)
         qed auto
       qed
 
@@ -1878,16 +1875,16 @@
 qed simp
 
 lemma emeasure_count_space_finite[simp]:
-  "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (count_space A) X = ereal (card X)"
+  "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (count_space A) X = of_nat (card X)"
   using emeasure_count_space[of X A] by simp
 
 lemma emeasure_count_space_infinite[simp]:
   "X \<subseteq> A \<Longrightarrow> infinite X \<Longrightarrow> emeasure (count_space A) X = \<infinity>"
   using emeasure_count_space[of X A] by simp
 
-lemma measure_count_space: "measure (count_space A) X = (if X \<subseteq> A then card X else 0)"
-  unfolding measure_def
-  by (cases "finite X") (simp_all add: emeasure_notin_sets)
+lemma measure_count_space: "measure (count_space A) X = (if X \<subseteq> A then of_nat (card X) else 0)"
+  by (cases "finite X") (auto simp: measure_notin_sets ennreal_of_nat_eq_real_of_nat
+                                    measure_zero_top measure_eq_emeasure_eq_ennreal)
 
 lemma emeasure_count_space_eq_0:
   "emeasure (count_space A) X = 0 \<longleftrightarrow> (X \<subseteq> A \<longrightarrow> X = {})"
@@ -1913,7 +1910,7 @@
 lemma sigma_finite_measure_count_space_countable:
   assumes A: "countable A"
   shows "sigma_finite_measure (count_space A)"
-  proof qed (auto intro!: exI[of _ "(\<lambda>a. {a}) ` A"] simp: A)
+  proof qed (insert A, auto intro!: exI[of _ "(\<lambda>a. {a}) ` A"])
 
 lemma sigma_finite_measure_count_space:
   fixes A :: "'a::countable set" shows "sigma_finite_measure (count_space A)"
@@ -1943,7 +1940,7 @@
     show "op \<inter> \<Omega> ` sets M \<subseteq> Pow (\<Omega> \<inter> space M)" "A \<in> sets (restrict_space M \<Omega>)"
       using \<open>A \<subseteq> \<Omega>\<close> \<open>A \<in> sets M\<close> sets.space_closed by (auto simp: sets_restrict_space)
     show "positive (sets (restrict_space M \<Omega>)) (emeasure M)"
-      by (auto simp: positive_def emeasure_nonneg)
+      by (auto simp: positive_def)
     show "countably_additive (sets (restrict_space M \<Omega>)) (emeasure M)"
     proof (rule countably_additiveI)
       fix A :: "nat \<Rightarrow> _" assume "range A \<subseteq> sets (restrict_space M \<Omega>)" "disjoint_family A"
@@ -2028,7 +2025,7 @@
     then obtain a' where "a = A \<inter> a'" "a' \<in> C" ..
     then have "emeasure (restrict_space M A) a \<le> emeasure M a'"
       using A C by(auto simp add: emeasure_restrict_space intro: emeasure_mono)
-    also have "\<dots> < \<infinity>" using C(4)[rule_format, of a'] \<open>a' \<in> C\<close> by simp
+    also have "\<dots> < \<infinity>" using C(4)[rule_format, of a'] \<open>a' \<in> C\<close> by (simp add: less_top)
     finally have "emeasure (restrict_space M A) a \<noteq> \<infinity>" by simp }
   ultimately show ?thesis
     by unfold_locales (rule exI conjI|assumption|blast)+
@@ -2106,60 +2103,63 @@
            dest: sets.sets_into_space)
 
 lemma measure_null_measure[simp]: "measure (null_measure M) X = 0"
-  by (simp add: measure_def)
+  by (intro measure_eq_emeasure_eq_ennreal) auto
 
 lemma null_measure_idem [simp]: "null_measure (null_measure M) = null_measure M"
-by(rule measure_eqI) simp_all
+  by(rule measure_eqI) simp_all
 
 subsection \<open>Scaling a measure\<close>
 
-definition scale_measure :: "real \<Rightarrow> 'a measure \<Rightarrow> 'a measure"
-where "scale_measure r M = measure_of (space M) (sets M) (\<lambda>A. (max 0 r) * emeasure M A)"
+definition scale_measure :: "ennreal \<Rightarrow> 'a measure \<Rightarrow> 'a measure"
+where
+  "scale_measure r M = measure_of (space M) (sets M) (\<lambda>A. r * emeasure M A)"
 
 lemma space_scale_measure: "space (scale_measure r M) = space M"
-by(simp add: scale_measure_def)
+  by (simp add: scale_measure_def)
 
 lemma sets_scale_measure [simp, measurable_cong]: "sets (scale_measure r M) = sets M"
-by(simp add: scale_measure_def)
+  by (simp add: scale_measure_def)
 
 lemma emeasure_scale_measure [simp]:
-  "emeasure (scale_measure r M) A = max 0 r * emeasure M A"
+  "emeasure (scale_measure r M) A = r * emeasure M A"
   (is "_ = ?\<mu> A")
 proof(cases "A \<in> sets M")
   case True
   show ?thesis unfolding scale_measure_def
   proof(rule emeasure_measure_of_sigma)
     show "sigma_algebra (space M) (sets M)" ..
-    show "positive (sets M) ?\<mu>" by(simp add: positive_def emeasure_nonneg)
+    show "positive (sets M) ?\<mu>" by (simp add: positive_def)
     show "countably_additive (sets M) ?\<mu>"
     proof (rule countably_additiveI)
       fix A :: "nat \<Rightarrow> _"  assume *: "range A \<subseteq> sets M" "disjoint_family A"
-      have "(\<Sum>i. ?\<mu> (A i)) = max 0 (ereal r) * (\<Sum>i. emeasure M (A i))"
-        by(rule suminf_cmult_ereal)(simp_all add: emeasure_nonneg)
+      have "(\<Sum>i. ?\<mu> (A i)) = r * (\<Sum>i. emeasure M (A i))"
+        by simp
       also have "\<dots> = ?\<mu> (\<Union>i. A i)" using * by(simp add: suminf_emeasure)
       finally show "(\<Sum>i. ?\<mu> (A i)) = ?\<mu> (\<Union>i. A i)" .
     qed
   qed(fact True)
 qed(simp add: emeasure_notin_sets)
 
-lemma measure_scale_measure [simp]: "measure (scale_measure r M) A = max 0 r * measure M A"
-by(simp add: measure_def max_def)
+lemma scale_measure_1 [simp]: "scale_measure 1 M = M"
+  by(rule measure_eqI) simp_all
 
-lemma scale_measure_1 [simp]: "scale_measure 1 M = M"
-by(rule measure_eqI)(simp_all add: max_def)
+lemma scale_measure_0[simp]: "scale_measure 0 M = null_measure M"
+  by(rule measure_eqI) simp_all
 
-lemma scale_measure_le_0: "r \<le> 0 \<Longrightarrow> scale_measure r M = null_measure M"
-by(rule measure_eqI)(simp_all add: max_def)
-
-lemma scale_measure_0 [simp]: "scale_measure 0 M = null_measure M"
-by(simp add: scale_measure_le_0)
+lemma measure_scale_measure [simp]: "0 \<le> r \<Longrightarrow> measure (scale_measure r M) A = r * measure M A"
+  using emeasure_scale_measure[of r M A]
+    emeasure_eq_ennreal_measure[of M A]
+    measure_eq_emeasure_eq_ennreal[of _ "scale_measure r M" A]
+  by (cases "emeasure (scale_measure r M) A = top")
+     (auto simp del: emeasure_scale_measure
+           simp: ennreal_top_eq_mult_iff ennreal_mult_eq_top_iff measure_zero_top ennreal_mult[symmetric])
 
 lemma scale_scale_measure [simp]:
-  "scale_measure r (scale_measure r' M) = scale_measure (max 0 r * max 0 r') M"
-by(rule measure_eqI)(simp_all add: max_def mult.assoc times_ereal.simps(1)[symmetric] del: times_ereal.simps(1))
+  "scale_measure r (scale_measure r' M) = scale_measure (r * r') M"
+  by (rule measure_eqI) (simp_all add: max_def mult.assoc)
 
 lemma scale_null_measure [simp]: "scale_measure r (null_measure M) = null_measure M"
-by(rule measure_eqI) simp_all
+  by (rule measure_eqI) simp_all
 
 subsection \<open>Measures form a chain-complete partial order\<close>
 
@@ -2212,7 +2212,6 @@
     apply simp
     apply (blast intro: antisym)
     apply (simp)
-    apply (blast intro: antisym)
     apply simp
     done
 qed (rule less_eq_measure.intros)
@@ -2262,7 +2261,7 @@
     fix F :: "nat \<Rightarrow> 'a set" assume F: "range F \<subseteq> sets (Sup A)" "disjoint_family F"
     show "(\<Sum>i. (SUP a:A. emeasure a) (F i)) = SUPREMUM A emeasure (UNION UNIV F)"
       unfolding SUP_apply
-    proof (subst suminf_SUP_eq_directed)
+    proof (subst ennreal_suminf_SUP_eq_directed)
       fix N i j assume "i \<in> A" "j \<in> A"
       with A(1)
       show "\<exists>k\<in>A. \<forall>n\<in>N. emeasure i (F n) \<le> emeasure k (F n) \<and> emeasure j (F n) \<le> emeasure k (F n)"
@@ -2273,7 +2272,7 @@
         fix a assume "a \<in> A" then show "(\<Sum>i. emeasure a (F i)) = emeasure a (UNION UNIV F)"
           using sets_Sup[OF A(1), of a] F by (cases "a = bot") (auto simp: suminf_emeasure)
       qed
-    qed (insert F \<open>A \<noteq> {}\<close>, auto simp: suminf_emeasure intro!: SUP_cong)
+    qed
   qed
 qed (insert \<open>A \<noteq> {}\<close> \<open>X \<in> sets (Sup A)\<close>, auto simp: positive_def dest: sets.sets_into_space intro: SUP_upper2)
 
--- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -9,18 +9,6 @@
   imports Measure_Space Borel_Space
 begin
 
-lemma infinite_countable_subset':
-  assumes X: "infinite X" shows "\<exists>C\<subseteq>X. countable C \<and> infinite C"
-proof -
-  from infinite_countable_subset[OF X] guess f ..
-  then show ?thesis
-    by (intro exI[of _ "range f"]) (auto simp: range_inj_infinite)
-qed
-
-lemma indicator_less_ereal[simp]:
-  "indicator A x \<le> (indicator B x::ereal) \<longleftrightarrow> (x \<in> A \<longrightarrow> x \<in> B)"
-  by (simp add: indicator_def not_le)
-
 subsection "Simple function"
 
 text \<open>
@@ -73,7 +61,7 @@
 qed
 
 lemma simple_function_indicator_representation:
-  fixes f ::"'a \<Rightarrow> ereal"
+  fixes f ::"'a \<Rightarrow> ennreal"
   assumes f: "simple_function M f" and x: "x \<in> space M"
   shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
   (is "?l = ?r")
@@ -88,7 +76,7 @@
 qed
 
 lemma simple_function_notspace:
-  "simple_function M (\<lambda>x. h x * indicator (- space M) x::ereal)" (is "simple_function M ?h")
+  "simple_function M (\<lambda>x. h x * indicator (- space M) x::ennreal)" (is "simple_function M ?h")
 proof -
   have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
   hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset)
@@ -100,10 +88,10 @@
   assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
   shows "simple_function M f \<longleftrightarrow> simple_function M g"
 proof -
-  have "f ` space M = g ` space M"
-    "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
-    using assms by (auto intro!: image_eqI)
-  thus ?thesis unfolding simple_function_def using assms by simp
+  have "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
+    using assms by auto
+  with assms show ?thesis
+    by (simp add: simple_function_def cong: image_cong)
 qed
 
 lemma simple_function_cong_algebra:
@@ -118,11 +106,14 @@
   using assms unfolding simple_function_def
   by (auto intro: borel_measurable_vimage)
 
+lemma simple_function_iff_borel_measurable:
+  fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
+  shows "simple_function M f \<longleftrightarrow> finite (f ` space M) \<and> f \<in> borel_measurable M"
+  by (metis borel_measurable_simple_function simple_functionD(1) simple_function_borel_measurable)
+
 lemma simple_function_eq_measurable:
-  fixes f :: "'a \<Rightarrow> ereal"
-  shows "simple_function M f \<longleftrightarrow> finite (f`space M) \<and> f \<in> measurable M (count_space UNIV)"
-  using simple_function_borel_measurable[of f] measurable_simple_function[of M f]
-  by (fastforce simp: simple_function_def)
+  "simple_function M f \<longleftrightarrow> finite (f`space M) \<and> f \<in> measurable M (count_space UNIV)"
+  using measurable_simple_function[of M f] by (fastforce simp: simple_function_def)
 
 lemma simple_function_const[intro, simp]:
   "simple_function M (\<lambda>x. c)"
@@ -204,9 +195,9 @@
   assume "finite P" from this assms show ?thesis by induct auto
 qed auto
 
-lemma simple_function_ereal[intro, simp]:
+lemma simple_function_ennreal[intro, simp]:
   fixes f g :: "'a \<Rightarrow> real" assumes sf: "simple_function M f"
-  shows "simple_function M (\<lambda>x. ereal (f x))"
+  shows "simple_function M (\<lambda>x. ennreal (f x))"
   by (rule simple_function_compose1[OF sf])
 
 lemma simple_function_real_of_nat[intro, simp]:
@@ -215,122 +206,101 @@
   by (rule simple_function_compose1[OF sf])
 
 lemma borel_measurable_implies_simple_function_sequence:
-  fixes u :: "'a \<Rightarrow> ereal"
-  assumes u: "u \<in> borel_measurable M"
-  shows "\<exists>f. incseq f \<and> (\<forall>i. \<infinity> \<notin> range (f i) \<and> simple_function M (f i)) \<and>
-             (\<forall>x. (SUP i. f i x) = max 0 (u x)) \<and> (\<forall>i x. 0 \<le> f i x)"
+  fixes u :: "'a \<Rightarrow> ennreal"
+  assumes u[measurable]: "u \<in> borel_measurable M"
+  shows "\<exists>f. incseq f \<and> (\<forall>i. (\<forall>x. f i x < top) \<and> simple_function M (f i)) \<and> u = (SUP i. f i)"
 proof -
-  def f \<equiv> "\<lambda>x i. if real i \<le> u x then i * 2 ^ i else nat \<lfloor>real_of_ereal (u x) * 2 ^ i\<rfloor>"
-  { fix x j have "f x j \<le> j * 2 ^ j" unfolding f_def
-    proof (split if_split, intro conjI impI)
-      assume "\<not> real j \<le> u x"
-      then have "nat \<lfloor>real_of_ereal (u x) * 2 ^ j\<rfloor> \<le> nat \<lfloor>j * 2 ^ j\<rfloor>"
-         by (cases "u x") (auto intro!: nat_mono floor_mono)
-      moreover have "real (nat \<lfloor>j * 2 ^ j\<rfloor>) \<le> j * 2^j"
-        by linarith
-      ultimately show "nat \<lfloor>real_of_ereal (u x) * 2 ^ j\<rfloor> \<le> j * 2 ^ j"
-        unfolding of_nat_le_iff by auto
-    qed auto }
-  note f_upper = this
+  def f \<equiv> "\<lambda>i x. real_of_int (floor (enn2real (min i (u x)) * 2^i)) / 2^i"
+
+  have [simp]: "0 \<le> f i x" for i x
+    by (auto simp: f_def intro!: divide_nonneg_nonneg mult_nonneg_nonneg enn2real_nonneg)
 
-  have real_f:
-    "\<And>i x. real (f x i) =
-      (if real i \<le> u x then i * 2 ^ i else real (nat \<lfloor>real_of_ereal (u x) * 2 ^ i\<rfloor>))"
-    unfolding f_def by auto
+  have *: "2^n * real_of_int x = real_of_int (2^n * x)" for n x
+    by simp
+
+  have "real_of_int \<lfloor>real i * 2 ^ i\<rfloor> = real_of_int \<lfloor>i * 2 ^ i\<rfloor>" for i
+    by (intro arg_cong[where f=real_of_int]) simp
+  then have [simp]: "real_of_int \<lfloor>real i * 2 ^ i\<rfloor> = i * 2 ^ i" for i
+    unfolding floor_of_nat by simp
 
-  let ?g = "\<lambda>j x. real (f x j) / 2^j :: ereal"
-  show ?thesis
-  proof (intro exI[of _ ?g] conjI allI ballI)
-    fix i
-    have "simple_function M (\<lambda>x. real (f x i))"
-    proof (intro simple_function_borel_measurable)
-      show "(\<lambda>x. real (f x i)) \<in> borel_measurable M"
-        using u by (auto simp: real_f)
-      have "(\<lambda>x. real (f x i))`space M \<subseteq> real`{..i*2^i}"
-        using f_upper[of _ i] by auto
-      then show "finite ((\<lambda>x. real (f x i))`space M)"
-        by (rule finite_subset) auto
-    qed
-    then show "simple_function M (?g i)"
-      by (auto)
-  next
-    show "incseq ?g"
-    proof (intro incseq_ereal incseq_SucI le_funI)
-      fix x and i :: nat
-      have "f x i * 2 \<le> f x (Suc i)" unfolding f_def
-      proof ((split if_split)+, intro conjI impI)
-        assume "ereal (real i) \<le> u x" "\<not> ereal (real (Suc i)) \<le> u x"
-        then show "i * 2 ^ i * 2 \<le> nat \<lfloor>real_of_ereal (u x) * 2 ^ Suc i\<rfloor>"
-          by (cases "u x") (auto intro!: le_nat_floor)
-      next
-        assume "\<not> ereal (real i) \<le> u x" "ereal (real (Suc i)) \<le> u x"
-        then show "nat \<lfloor>real_of_ereal (u x) * 2 ^ i\<rfloor> * 2 \<le> Suc i * 2 ^ Suc i"
-          by (cases "u x") auto
-      next
-        assume "\<not> ereal (real i) \<le> u x" "\<not> ereal (real (Suc i)) \<le> u x"
-        have "nat \<lfloor>real_of_ereal (u x) * 2 ^ i\<rfloor> * 2 = nat \<lfloor>real_of_ereal (u x) * 2 ^ i\<rfloor> * nat \<lfloor>2::real\<rfloor>"
-          by simp
-        also have "\<dots> \<le> nat \<lfloor>real_of_ereal (u x) * 2 ^ i * 2\<rfloor>"
-        proof cases
-          assume "0 \<le> u x" then show ?thesis
-            by (intro le_mult_nat_floor)
-        next
-          assume "\<not> 0 \<le> u x" then show ?thesis
-            by (cases "u x") (auto simp: nat_floor_neg mult_nonpos_nonneg)
-        qed
-        also have "\<dots> = nat \<lfloor>real_of_ereal (u x) * 2 ^ Suc i\<rfloor>"
-          by (simp add: ac_simps)
-        finally show "nat \<lfloor>real_of_ereal (u x) * 2 ^ i\<rfloor> * 2 \<le> nat \<lfloor>real_of_ereal (u x) * 2 ^ Suc i\<rfloor>" .
-      qed simp
-      then show "?g i x \<le> ?g (Suc i) x"
-        by (auto simp: field_simps)
-    qed
-  next
-    fix x show "(SUP i. ?g i x) = max 0 (u x)"
-    proof (rule SUP_eqI)
-      fix i show "?g i x \<le> max 0 (u x)" unfolding max_def f_def
-        by (cases "u x") (auto simp: field_simps nat_floor_neg mult_nonpos_nonneg)
+  have "incseq f"
+  proof (intro monoI le_funI)
+    fix m n :: nat and x assume "m \<le> n"
+    moreover
+    { fix d :: nat
+      have "\<lfloor>2^d::real\<rfloor> * \<lfloor>2^m * enn2real (min (of_nat m) (u x))\<rfloor> \<le>
+        \<lfloor>2^d * (2^m * enn2real (min (of_nat m) (u x)))\<rfloor>"
+        by (rule le_mult_floor) (auto simp: enn2real_nonneg)
+      also have "\<dots> \<le> \<lfloor>2^d * (2^m * enn2real (min (of_nat d + of_nat m) (u x)))\<rfloor>"
+        by (intro floor_mono mult_mono enn2real_mono min.mono)
+           (auto simp: enn2real_nonneg min_less_iff_disj of_nat_less_top)
+      finally have "f m x \<le> f (m + d) x"
+        unfolding f_def
+        by (auto simp: field_simps power_add * simp del: of_int_mult) }
+    ultimately show "f m x \<le> f n x"
+      by (auto simp add: le_iff_add)
+  qed
+  then have inc_f: "incseq (\<lambda>i. ennreal (f i x))" for x
+    by (auto simp: incseq_def le_fun_def)
+  then have "incseq (\<lambda>i x. ennreal (f i x))"
+    by (auto simp: incseq_def le_fun_def)
+  moreover
+  have "simple_function M (f i)" for i
+  proof (rule simple_function_borel_measurable)
+    have "\<lfloor>enn2real (min (of_nat i) (u x)) * 2 ^ i\<rfloor> \<le> \<lfloor>int i * 2 ^ i\<rfloor>" for x
+      by (cases "u x" rule: ennreal_cases)
+         (auto split: split_min intro!: floor_mono)
+    then have "f i ` space M \<subseteq> (\<lambda>n. real_of_int n / 2^i) ` {0 .. of_nat i * 2^i}"
+      unfolding floor_of_int by (auto simp: f_def enn2real_nonneg intro!: imageI)
+    then show "finite (f i ` space M)"
+      by (rule finite_subset) auto
+    show "f i \<in> borel_measurable M"
+      unfolding f_def enn2real_def by measurable
+  qed
+  moreover
+  { fix x
+    have "(SUP i. ennreal (f i x)) = u x"
+    proof (cases "u x" rule: ennreal_cases)
+      case top then show ?thesis
+        by (simp add: f_def inf_min[symmetric] ennreal_of_nat_eq_real_of_nat[symmetric]
+                      ennreal_SUP_of_nat_eq_top)
     next
-      fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> ?g i x \<le> y"
-      have "\<And>i. 0 \<le> ?g i x" by auto
-      from order_trans[OF this *] have "0 \<le> y" by simp
-      show "max 0 (u x) \<le> y"
-      proof (cases y)
-        case (real r)
-        with * have *: "\<And>i. f x i \<le> r * 2^i" by (auto simp: divide_le_eq)
-        from reals_Archimedean2[of r] * have "u x \<noteq> \<infinity>" by (auto simp: f_def) (metis less_le_not_le)
-        then have "\<exists>p. max 0 (u x) = ereal p \<and> 0 \<le> p" by (cases "u x") (auto simp: max_def)
-        then guess p .. note ux = this
-        obtain m :: nat where m: "p < real m" using reals_Archimedean2 ..
-        have "p \<le> r"
-        proof (rule ccontr)
-          assume "\<not> p \<le> r"
-          with LIMSEQ_inverse_realpow_zero[unfolded LIMSEQ_iff, rule_format, of 2 "p - r"]
-          obtain N where "\<forall>n\<ge>N. r * 2^n < p * 2^n - 1" by (auto simp: field_simps)
-          then have "r * 2^max N m < p * 2^max N m - 1" by simp
-          moreover
-          have "real (nat \<lfloor>p * 2 ^ max N m\<rfloor>) \<le> r * 2 ^ max N m"
-            using *[of "max N m"] m unfolding real_f using ux
-            by (cases "0 \<le> u x") (simp_all add: max_def split: if_split_asm)
-          then have "p * 2 ^ max N m - 1 < r * 2 ^ max N m"
-            by linarith
-          ultimately show False by auto
-        qed
-        then show "max 0 (u x) \<le> y" using real ux by simp
-      qed (insert \<open>0 \<le> y\<close>, auto)
-    qed
-  qed auto
+      case (real r)
+      obtain n where "r \<le> of_nat n" using real_arch_simple by auto
+      then have min_eq_r: "\<forall>\<^sub>F x in sequentially. min (real x) r = r"
+        by (auto simp: eventually_sequentially intro!: exI[of _ n] split: split_min)
+
+      have "(\<lambda>i. real_of_int \<lfloor>min (real i) r * 2^i\<rfloor> / 2^i) \<longlonglongrightarrow> r"
+      proof (rule tendsto_sandwich)
+        show "(\<lambda>n. r - (1/2)^n) \<longlonglongrightarrow> r"
+          by (auto intro!: tendsto_eq_intros LIMSEQ_power_zero)
+        show "\<forall>\<^sub>F n in sequentially. real_of_int \<lfloor>min (real n) r * 2 ^ n\<rfloor> / 2 ^ n \<le> r"
+          using min_eq_r by eventually_elim (auto simp: field_simps)
+        have *: "r * (2 ^ n * 2 ^ n) \<le> 2^n + 2^n * real_of_int \<lfloor>r * 2 ^ n\<rfloor>" for n
+          using real_of_int_floor_ge_diff_one[of "r * 2^n", THEN mult_left_mono, of "2^n"]
+          by (auto simp: field_simps)
+        show "\<forall>\<^sub>F n in sequentially. r - (1/2)^n \<le> real_of_int \<lfloor>min (real n) r * 2 ^ n\<rfloor> / 2 ^ n"
+          using min_eq_r by eventually_elim (insert *, auto simp: field_simps)
+      qed auto
+      then have "(\<lambda>i. ennreal (f i x)) \<longlonglongrightarrow> ennreal r"
+        by (simp add: real f_def ennreal_of_nat_eq_real_of_nat min_ennreal)
+      from LIMSEQ_unique[OF LIMSEQ_SUP[OF inc_f] this]
+      show ?thesis
+        by (simp add: real)
+    qed }
+  ultimately show ?thesis
+    by (intro exI[of _ "\<lambda>i x. ennreal (f i x)"]) auto
 qed
 
 lemma borel_measurable_implies_simple_function_sequence':
-  fixes u :: "'a \<Rightarrow> ereal"
+  fixes u :: "'a \<Rightarrow> ennreal"
   assumes u: "u \<in> borel_measurable M"
-  obtains f where "\<And>i. simple_function M (f i)" "incseq f" "\<And>i. \<infinity> \<notin> range (f i)"
-    "\<And>x. (SUP i. f i x) = max 0 (u x)" "\<And>i x. 0 \<le> f i x"
-  using borel_measurable_implies_simple_function_sequence[OF u] by auto
+  obtains f where
+    "\<And>i. simple_function M (f i)" "incseq f" "\<And>i x. f i x < top" "\<And>x. (SUP i. f i x) = u x"
+  using borel_measurable_implies_simple_function_sequence[OF u] by (auto simp: fun_eq_iff) blast
 
 lemma simple_function_induct[consumes 1, case_names cong set mult add, induct set: simple_function]:
-  fixes u :: "'a \<Rightarrow> ereal"
+  fixes u :: "'a \<Rightarrow> ennreal"
   assumes u: "simple_function M u"
   assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> simple_function M g \<Longrightarrow> (AE x in M. f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g"
   assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
@@ -360,13 +330,13 @@
     done
 qed fact
 
-lemma simple_function_induct_nn[consumes 2, case_names cong set mult add]:
-  fixes u :: "'a \<Rightarrow> ereal"
-  assumes u: "simple_function M u" and nn: "\<And>x. 0 \<le> u x"
+lemma simple_function_induct_nn[consumes 1, case_names cong set mult add]:
+  fixes u :: "'a \<Rightarrow> ennreal"
+  assumes u: "simple_function M u"
   assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> simple_function M g \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g"
   assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
-  assumes mult: "\<And>u c. 0 \<le> c \<Longrightarrow> simple_function M u \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)"
-  assumes add: "\<And>u v. simple_function M u \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> simple_function M v \<Longrightarrow> (\<And>x. 0 \<le> v x) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = 0 \<or> v x = 0) \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)"
+  assumes mult: "\<And>u c. simple_function M u \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)"
+  assumes add: "\<And>u v. simple_function M u \<Longrightarrow> P u \<Longrightarrow> simple_function M v \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = 0 \<or> v x = 0) \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)"
   shows "P u"
 proof -
   show ?thesis
@@ -381,8 +351,7 @@
       apply (auto intro: u)
       done
   next
-
-    from u nn have "finite (u ` space M)" "\<And>x. x \<in> u ` space M \<Longrightarrow> 0 \<le> x"
+    from u have "finite (u ` space M)"
       unfolding simple_function_def by auto
     then show "P (\<lambda>x. \<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x)"
     proof induct
@@ -392,31 +361,30 @@
       case (insert x S)
       { fix z have "(\<Sum>y\<in>S. y * indicator (u -` {y} \<inter> space M) z) = 0 \<or>
           x * indicator (u -` {x} \<inter> space M) z = 0"
-          using insert by (subst setsum_ereal_0) (auto simp: indicator_def) }
+          using insert by (subst setsum_eq_0_iff) (auto simp: indicator_def) }
       note disj = this
       from insert show ?case
-        by (auto intro!: add mult set simple_functionD u setsum_nonneg simple_function_setsum disj)
+        by (auto intro!: add mult set simple_functionD u simple_function_setsum disj)
     qed
   qed fact
 qed
 
-lemma borel_measurable_induct[consumes 2, case_names cong set mult add seq, induct set: borel_measurable]:
-  fixes u :: "'a \<Rightarrow> ereal"
-  assumes u: "u \<in> borel_measurable M" "\<And>x. 0 \<le> u x"
+lemma borel_measurable_induct[consumes 1, case_names cong set mult add seq, induct set: borel_measurable]:
+  fixes u :: "'a \<Rightarrow> ennreal"
+  assumes u: "u \<in> borel_measurable M"
   assumes cong: "\<And>f g. f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P g \<Longrightarrow> P f"
   assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
-  assumes mult': "\<And>u c. 0 \<le> c \<Longrightarrow> c < \<infinity> \<Longrightarrow> u \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < \<infinity>) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)"
-  assumes add: "\<And>u v. u \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < \<infinity>) \<Longrightarrow> P u \<Longrightarrow> v \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> v x) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> v x < \<infinity>) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = 0 \<or> v x = 0) \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)"
-  assumes seq: "\<And>U. (\<And>i. U i \<in> borel_measurable M) \<Longrightarrow> (\<And>i x. 0 \<le> U i x) \<Longrightarrow> (\<And>i x. x \<in> space M \<Longrightarrow> U i x < \<infinity>) \<Longrightarrow>  (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> u = (SUP i. U i) \<Longrightarrow> P (SUP i. U i)"
+  assumes mult': "\<And>u c. c < top \<Longrightarrow> u \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < top) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)"
+  assumes add: "\<And>u v. u \<in> borel_measurable M\<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < top) \<Longrightarrow> P u \<Longrightarrow> v \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> v x < top) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = 0 \<or> v x = 0) \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)"
+  assumes seq: "\<And>U. (\<And>i. U i \<in> borel_measurable M) \<Longrightarrow> (\<And>i x. x \<in> space M \<Longrightarrow> U i x < top) \<Longrightarrow> (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> u = (SUP i. U i) \<Longrightarrow> P (SUP i. U i)"
   shows "P u"
   using u
 proof (induct rule: borel_measurable_implies_simple_function_sequence')
-  fix U assume U: "\<And>i. simple_function M (U i)" "incseq U" "\<And>i. \<infinity> \<notin> range (U i)" and
-    sup: "\<And>x. (SUP i. U i x) = max 0 (u x)" and nn: "\<And>i x. 0 \<le> U i x"
+  fix U assume U: "\<And>i. simple_function M (U i)" "incseq U" "\<And>i x. U i x < top" and sup: "\<And>x. (SUP i. U i x) = u x"
   have u_eq: "u = (SUP i. U i)"
-    using nn u sup by (auto simp: max_def)
+    using u sup by auto
 
-  have not_inf: "\<And>x i. x \<in> space M \<Longrightarrow> U i x < \<infinity>"
+  have not_inf: "\<And>x i. x \<in> space M \<Longrightarrow> U i x < top"
     using U by (auto simp: image_iff eq_commute)
 
   from U have "\<And>i. U i \<in> borel_measurable M"
@@ -426,26 +394,26 @@
     unfolding u_eq
   proof (rule seq)
     fix i show "P (U i)"
-      using \<open>simple_function M (U i)\<close> nn[of i] not_inf[of _ i]
+      using \<open>simple_function M (U i)\<close> not_inf[of _ i]
     proof (induct rule: simple_function_induct_nn)
       case (mult u c)
       show ?case
       proof cases
         assume "c = 0 \<or> space M = {} \<or> (\<forall>x\<in>space M. u x = 0)"
-        with mult(2) show ?thesis
+        with mult(1) show ?thesis
           by (intro cong[of "\<lambda>x. c * u x" "indicator {}"] set)
              (auto dest!: borel_measurable_simple_function)
       next
         assume "\<not> (c = 0 \<or> space M = {} \<or> (\<forall>x\<in>space M. u x = 0))"
-        with mult obtain x where u_fin: "\<And>x. x \<in> space M \<Longrightarrow> u x < \<infinity>"
-          and x: "x \<in> space M" "u x \<noteq> 0" "c \<noteq> 0"
-          by auto
-        with mult have "P u"
+        then obtain x where "space M \<noteq> {}" and x: "x \<in> space M" "u x \<noteq> 0" "c \<noteq> 0"
           by auto
-        from x mult(5)[OF \<open>x \<in> space M\<close>] mult(1) mult(3)[of x] have "c < \<infinity>"
-          by auto
-        with u_fin mult
-        show ?thesis
+        with mult(3)[of x] have "c < top"
+          by (auto simp: ennreal_mult_less_top)
+        then have u_fin: "x' \<in> space M \<Longrightarrow> u x' < top" for x'
+          using mult(3)[of x'] \<open>c \<noteq> 0\<close> by (auto simp: ennreal_mult_less_top)
+        then have "P u"
+          by (rule mult)
+        with u_fin \<open>c < top\<close> mult(1) show ?thesis
           by (intro mult') (auto dest!: borel_measurable_simple_function)
       qed
     qed (auto intro: cong intro!: set add dest!: borel_measurable_simple_function)
@@ -511,11 +479,11 @@
 
 subsection "Simple integral"
 
-definition simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ereal" ("integral\<^sup>S") where
+definition simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" ("integral\<^sup>S") where
   "integral\<^sup>S M f = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M))"
 
 syntax
-  "_simple_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> 'a measure \<Rightarrow> ereal" ("\<integral>\<^sup>S _. _ \<partial>_" [60,61] 110)
+  "_simple_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" ("\<integral>\<^sup>S _. _ \<partial>_" [60,61] 110)
 
 translations
   "\<integral>\<^sup>S x. f \<partial>M" == "CONST simple_integral M (%x. f)"
@@ -560,7 +528,7 @@
     (\<Sum>y\<in>f`space M. y * (\<Sum>z\<in>g`space M.
       if \<exists>x\<in>space M. y = f x \<and> z = g x then emeasure M {x\<in>space M. g x = z} else 0))"
     unfolding simple_integral_def
-  proof (safe intro!: setsum.cong ereal_right_mult_cong)
+  proof (safe intro!: setsum.cong ennreal_mult_left_cong)
     fix y assume y: "y \<in> space M" "f y \<noteq> 0"
     have [simp]: "g ` space M \<inter> {z. \<exists>x\<in>space M. f y = f x \<and> z = g x} =
         {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
@@ -580,7 +548,7 @@
   qed
   also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M.
       if \<exists>x\<in>space M. y = f x \<and> z = g x then y * emeasure M {x\<in>space M. g x = z} else 0))"
-    by (auto intro!: setsum.cong simp: setsum_ereal_right_distrib emeasure_nonneg)
+    by (auto intro!: setsum.cong simp: setsum_right_distrib)
   also have "\<dots> = ?r"
     by (subst setsum.commute)
        (auto intro!: setsum.cong simp: setsum.If_cases scaleR_setsum_right[symmetric] eq)
@@ -596,7 +564,7 @@
     by (intro simple_function_partition) (auto intro: f g)
   also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * emeasure M {x\<in>space M. (f x, g x) = y}) +
     (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y})"
-    using assms(2,4) by (auto intro!: setsum.cong ereal_left_distrib simp: setsum.distrib[symmetric])
+    using assms(2,4) by (auto intro!: setsum.cong distrib_right simp: setsum.distrib[symmetric])
   also have "(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * emeasure M {x\<in>space M. (f x, g x) = y}) = (\<integral>\<^sup>Sx. f x \<partial>M)"
     by (intro simple_function_partition[symmetric]) (auto intro: f g)
   also have "(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y}) = (\<integral>\<^sup>Sx. g x \<partial>M)"
@@ -615,14 +583,14 @@
 qed auto
 
 lemma simple_integral_mult[simp]:
-  assumes f: "simple_function M f" "\<And>x. 0 \<le> f x"
+  assumes f: "simple_function M f"
   shows "(\<integral>\<^sup>Sx. c * f x \<partial>M) = c * integral\<^sup>S M f"
 proof -
   have "(\<integral>\<^sup>Sx. c * f x \<partial>M) = (\<Sum>y\<in>f ` space M. (c * y) * emeasure M {x\<in>space M. f x = y})"
     using f by (intro simple_function_partition) auto
   also have "\<dots> = c * integral\<^sup>S M f"
     using f unfolding simple_integral_def
-    by (subst setsum_ereal_right_distrib) (auto simp: emeasure_nonneg mult.assoc Int_def conj_commute)
+    by (subst setsum_right_distrib) (auto simp: mult.assoc Int_def conj_commute)
   finally show ?thesis .
 qed
 
@@ -642,13 +610,13 @@
     proof cases
       assume "?M \<noteq> 0"
       then have "0 < ?M"
-        by (simp add: less_le emeasure_nonneg)
+        by (simp add: less_le)
       also have "\<dots> \<le> ?\<mu> (\<lambda>y. f x \<le> g x)"
         using mono by (intro emeasure_mono_AE) auto
       finally have "\<not> \<not> f x \<le> g x"
         by (intro notI) auto
       then show ?thesis
-        by (intro ereal_mult_right_mono) auto
+        by (intro mult_right_mono) auto
     qed simp
   qed
   also have "\<dots> = integral\<^sup>S M g"
@@ -684,7 +652,7 @@
   shows "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) =
     (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
 proof -
-  have eq: "(\<lambda>x. (f x, indicator A x)) ` space M \<inter> {x. snd x = 1} = (\<lambda>x. (f x, 1::ereal))`A"
+  have eq: "(\<lambda>x. (f x, indicator A x)) ` space M \<inter> {x. snd x = 1} = (\<lambda>x. (f x, 1::ennreal))`A"
     using A[THEN sets.sets_into_space] by (auto simp: indicator_def image_iff split: if_split_asm)
   have eq2: "\<And>x. f x \<notin> f ` A \<Longrightarrow> f -` {f x} \<inter> space M \<inter> A = {}"
     by (auto simp: image_iff)
@@ -692,12 +660,12 @@
   have "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) =
     (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x))`space M. (fst y * snd y) * emeasure M {x\<in>space M. (f x, indicator A x) = y})"
     using assms by (intro simple_function_partition) auto
-  also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x::ereal))`space M.
+  also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x::ennreal))`space M.
     if snd y = 1 then fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A) else 0)"
     by (auto simp: indicator_def split: if_split_asm intro!: arg_cong2[where f="op *"] arg_cong2[where f=emeasure] setsum.cong)
-  also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, 1::ereal))`A. fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A))"
+  also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, 1::ennreal))`A. fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A))"
     using assms by (subst setsum.If_cases) (auto intro!: simple_functionD(1) simp: eq)
-  also have "\<dots> = (\<Sum>y\<in>fst`(\<lambda>x. (f x, 1::ereal))`A. y * emeasure M (f -` {y} \<inter> space M \<inter> A))"
+  also have "\<dots> = (\<Sum>y\<in>fst`(\<lambda>x. (f x, 1::ennreal))`A. y * emeasure M (f -` {y} \<inter> space M \<inter> A))"
     by (subst setsum.reindex [of fst]) (auto simp: inj_on_def)
   also have "\<dots> = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
     using A[THEN sets.sets_into_space]
@@ -715,7 +683,7 @@
   assumes "simple_function M u" "\<And>x. 0 \<le> u x" and "N \<in> null_sets M"
   shows "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = 0"
 proof -
-  have "AE x in M. indicator N x = (0 :: ereal)"
+  have "AE x in M. indicator N x = (0 :: ennreal)"
     using \<open>N \<in> null_sets M\<close> by (auto simp: indicator_def intro!: AE_I[of _ _ N])
   then have "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^sup>Sx. 0 \<partial>M)"
     using assms apply (intro simple_integral_cong_AE) by auto
@@ -744,69 +712,47 @@
 
 subsection \<open>Integral on nonnegative functions\<close>
 
-definition nn_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ereal" ("integral\<^sup>N") where
-  "integral\<^sup>N M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f}. integral\<^sup>S M g)"
+definition nn_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" ("integral\<^sup>N") where
+  "integral\<^sup>N M f = (SUP g : {g. simple_function M g \<and> g \<le> f}. integral\<^sup>S M g)"
 
 syntax
-  "_nn_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> 'a measure \<Rightarrow> ereal" ("\<integral>\<^sup>+((2 _./ _)/ \<partial>_)" [60,61] 110)
+  "_nn_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" ("\<integral>\<^sup>+((2 _./ _)/ \<partial>_)" [60,61] 110)
 
 translations
   "\<integral>\<^sup>+x. f \<partial>M" == "CONST nn_integral M (\<lambda>x. f)"
 
-lemma nn_integral_nonneg: "0 \<le> integral\<^sup>N M f"
-  by (auto intro!: SUP_upper2[of "\<lambda>x. 0"] simp: nn_integral_def le_fun_def)
-
-lemma nn_integral_le_0[simp]: "integral\<^sup>N M f \<le> 0 \<longleftrightarrow> integral\<^sup>N M f = 0"
-  using nn_integral_nonneg[of M f] by auto
-
-lemma nn_integral_not_less_0 [simp]: "\<not> nn_integral M f < 0"
-by(simp add: not_less nn_integral_nonneg)
-
-lemma nn_integral_not_MInfty[simp]: "integral\<^sup>N M f \<noteq> -\<infinity>"
-  using nn_integral_nonneg[of M f] by auto
-
 lemma nn_integral_def_finite:
-  "integral\<^sup>N M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f \<and> range g \<subseteq> {0 ..< \<infinity>}}. integral\<^sup>S M g)"
+  "integral\<^sup>N M f = (SUP g : {g. simple_function M g \<and> g \<le> f \<and> (\<forall>x. g x < top)}. integral\<^sup>S M g)"
     (is "_ = SUPREMUM ?A ?f")
   unfolding nn_integral_def
 proof (safe intro!: antisym SUP_least)
-  fix g assume g: "simple_function M g" "g \<le> max 0 \<circ> f"
-  let ?G = "{x \<in> space M. \<not> g x \<noteq> \<infinity>}"
-  note gM = g(1)[THEN borel_measurable_simple_function]
-  have \<mu>_G_pos: "0 \<le> (emeasure M) ?G" using gM by auto
-  let ?g = "\<lambda>y x. if g x = \<infinity> then y else max 0 (g x)"
-  from g gM have g_in_A: "\<And>y. 0 \<le> y \<Longrightarrow> y \<noteq> \<infinity> \<Longrightarrow> ?g y \<in> ?A"
-    apply (safe intro!: simple_function_max simple_function_If)
-    apply (force simp: max_def le_fun_def split: if_split_asm)+
-    done
+  fix g assume g[measurable]: "simple_function M g" "g \<le> f"
+
   show "integral\<^sup>S M g \<le> SUPREMUM ?A ?f"
   proof cases
-    have g0: "?g 0 \<in> ?A" by (intro g_in_A) auto
-    assume "(emeasure M) ?G = 0"
-    with gM have "AE x in M. x \<notin> ?G"
-      by (auto simp add: AE_iff_null intro!: null_setsI)
-    with gM g show ?thesis
-      by (intro SUP_upper2[OF g0] simple_integral_mono_AE)
-         (auto simp: max_def intro!: simple_function_If)
+    assume ae: "AE x in M. g x \<noteq> top"
+    let ?G = "{x \<in> space M. g x \<noteq> top}"
+    have "integral\<^sup>S M g = integral\<^sup>S M (\<lambda>x. g x * indicator ?G x)"
+    proof (rule simple_integral_cong_AE)
+      show "AE x in M. g x = g x * indicator ?G x"
+        using ae AE_space by eventually_elim auto
+    qed (insert g, auto)
+    also have "\<dots> \<le> SUPREMUM ?A ?f"
+      using g by (intro SUP_upper) (auto simp: le_fun_def less_top split: split_indicator)
+    finally show ?thesis .
   next
-    assume \<mu>_G: "(emeasure M) ?G \<noteq> 0"
-    have "SUPREMUM ?A (integral\<^sup>S M) = \<infinity>"
-    proof (intro SUP_PInfty)
-      fix n :: nat
-      let ?y = "ereal (real n) / (if (emeasure M) ?G = \<infinity> then 1 else (emeasure M) ?G)"
-      have "0 \<le> ?y" "?y \<noteq> \<infinity>" using \<mu>_G \<mu>_G_pos by (auto simp: ereal_divide_eq)
-      then have "?g ?y \<in> ?A" by (rule g_in_A)
-      have "real n \<le> ?y * (emeasure M) ?G"
-        using \<mu>_G \<mu>_G_pos by (cases "(emeasure M) ?G") (auto simp: field_simps)
-      also have "\<dots> = (\<integral>\<^sup>Sx. ?y * indicator ?G x \<partial>M)"
-        using \<open>0 \<le> ?y\<close> \<open>?g ?y \<in> ?A\<close> gM
-        by (subst simple_integral_cmult_indicator) auto
-      also have "\<dots> \<le> integral\<^sup>S M (?g ?y)" using \<open>?g ?y \<in> ?A\<close> gM
-        by (intro simple_integral_mono) auto
-      finally show "\<exists>i\<in>?A. real n \<le> integral\<^sup>S M i"
-        using \<open>?g ?y \<in> ?A\<close> by blast
-    qed
-    then show ?thesis by simp
+    assume nAE: "\<not> (AE x in M. g x \<noteq> top)"
+    then have "emeasure M {x\<in>space M. g x = top} \<noteq> 0" (is "emeasure M ?G \<noteq> 0")
+      by (subst (asm) AE_iff_measurable[OF _ refl]) auto
+    then have "top = (SUP n. (\<integral>\<^sup>Sx. of_nat n * indicator ?G x \<partial>M))"
+      by (simp add: ennreal_SUP_of_nat_eq_top ennreal_top_eq_mult_iff SUP_mult_right_ennreal[symmetric])
+    also have "\<dots> \<le> SUPREMUM ?A ?f"
+      using g
+      by (safe intro!: SUP_least SUP_upper)
+         (auto simp: le_fun_def of_nat_less_top top_unique[symmetric] split: split_indicator
+               intro: order_trans[of _ "g x" "f x" for x, OF order_trans[of _ top]])
+    finally show ?thesis
+      by (simp add: top_unique del: SUP_eq_top_iff Sup_eq_top_iff)
   qed
 qed (auto intro: SUP_upper)
 
@@ -814,24 +760,24 @@
   assumes ae: "AE x in M. u x \<le> v x" shows "integral\<^sup>N M u \<le> integral\<^sup>N M v"
   unfolding nn_integral_def
 proof (safe intro!: SUP_mono)
-  fix n assume n: "simple_function M n" "n \<le> max 0 \<circ> u"
+  fix n assume n: "simple_function M n" "n \<le> u"
   from ae[THEN AE_E] guess N . note N = this
   then have ae_N: "AE x in M. x \<notin> N" by (auto intro: AE_not_in)
   let ?n = "\<lambda>x. n x * indicator (space M - N) x"
   have "AE x in M. n x \<le> ?n x" "simple_function M ?n"
     using n N ae_N by auto
   moreover
-  { fix x have "?n x \<le> max 0 (v x)"
+  { fix x have "?n x \<le> v x"
     proof cases
       assume x: "x \<in> space M - N"
       with N have "u x \<le> v x" by auto
       with n(2)[THEN le_funD, of x] x show ?thesis
         by (auto simp: max_def split: if_split_asm)
     qed simp }
-  then have "?n \<le> max 0 \<circ> v" by (auto simp: le_funI)
+  then have "?n \<le> v" by (auto simp: le_funI)
   moreover have "integral\<^sup>S M n \<le> integral\<^sup>S M ?n"
     using ae_N N n by (auto intro!: simple_integral_mono_AE)
-  ultimately show "\<exists>m\<in>{g. simple_function M g \<and> g \<le> max 0 \<circ> v}. integral\<^sup>S M n \<le> integral\<^sup>S M m"
+  ultimately show "\<exists>m\<in>{g. simple_function M g \<and> g \<le> v}. integral\<^sup>S M n \<le> integral\<^sup>S M m"
     by force
 qed
 
@@ -858,13 +804,20 @@
   "M = N \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N N v"
   by (auto intro: nn_integral_cong)
 
+lemma incseq_nn_integral:
+  assumes "incseq f" shows "incseq (\<lambda>i. integral\<^sup>N M (f i))"
+proof -
+  have "\<And>i x. f i x \<le> f (Suc i) x"
+    using assms by (auto dest!: incseq_SucD simp: le_fun_def)
+  then show ?thesis
+    by (auto intro!: incseq_SucI nn_integral_mono)
+qed
+
 lemma nn_integral_eq_simple_integral:
-  assumes f: "simple_function M f" "\<And>x. 0 \<le> f x" shows "integral\<^sup>N M f = integral\<^sup>S M f"
+  assumes f: "simple_function M f" shows "integral\<^sup>N M f = integral\<^sup>S M f"
 proof -
   let ?f = "\<lambda>x. f x * indicator (space M) x"
   have f': "simple_function M ?f" using f by auto
-  with f(2) have [simp]: "max 0 \<circ> ?f = ?f"
-    by (auto simp: fun_eq_iff max_def split: split_indicator)
   have "integral\<^sup>N M ?f \<le> integral\<^sup>S M ?f" using f'
     by (force intro!: SUP_least simple_integral_mono simp: le_fun_def nn_integral_def)
   moreover have "integral\<^sup>S M ?f \<le> integral\<^sup>N M ?f"
@@ -874,153 +827,81 @@
     by (simp cong: nn_integral_cong simple_integral_cong)
 qed
 
-lemma nn_integral_eq_simple_integral_AE:
-  assumes f: "simple_function M f" "AE x in M. 0 \<le> f x" shows "integral\<^sup>N M f = integral\<^sup>S M f"
-proof -
-  have "AE x in M. f x = max 0 (f x)" using f by (auto split: split_max)
-  with f have "integral\<^sup>N M f = integral\<^sup>S M (\<lambda>x. max 0 (f x))"
-    by (simp cong: nn_integral_cong_AE simple_integral_cong_AE
-             add: nn_integral_eq_simple_integral)
-  with assms show ?thesis
-    by (auto intro!: simple_integral_cong_AE split: split_max)
-qed
-
-lemma nn_integral_SUP_approx:
-  assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
-  and u: "simple_function M u" "u \<le> (SUP i. f i)" "u`space M \<subseteq> {0..<\<infinity>}"
-  shows "integral\<^sup>S M u \<le> (SUP i. integral\<^sup>N M (f i))" (is "_ \<le> ?S")
-proof (rule ereal_le_mult_one_interval)
-  have "0 \<le> (SUP i. integral\<^sup>N M (f i))"
-    using f(3) by (auto intro!: SUP_upper2 nn_integral_nonneg)
-  then show "(SUP i. integral\<^sup>N M (f i)) \<noteq> -\<infinity>" by auto
-  have u_range: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> u x \<and> u x \<noteq> \<infinity>"
-    using u(3) by auto
-  fix a :: ereal assume "0 < a" "a < 1"
-  hence "a \<noteq> 0" by auto
-  let ?B = "\<lambda>i. {x \<in> space M. a * u x \<le> f i x}"
-  have B: "\<And>i. ?B i \<in> sets M"
-    using f \<open>simple_function M u\<close>[THEN borel_measurable_simple_function] by auto
-
-  let ?uB = "\<lambda>i x. u x * indicator (?B i) x"
-
-  { fix i have "?B i \<subseteq> ?B (Suc i)"
-    proof safe
-      fix i x assume "a * u x \<le> f i x"
-      also have "\<dots> \<le> f (Suc i) x"
-        using \<open>incseq f\<close>[THEN incseq_SucD] unfolding le_fun_def by auto
-      finally show "a * u x \<le> f (Suc i) x" .
-    qed }
-  note B_mono = this
-
-  note B_u = sets.Int[OF u(1)[THEN simple_functionD(2)] B]
-
-  let ?B' = "\<lambda>i n. (u -` {i} \<inter> space M) \<inter> ?B n"
-  have measure_conv: "\<And>i. (emeasure M) (u -` {i} \<inter> space M) = (SUP n. (emeasure M) (?B' i n))"
-  proof -
-    fix i
-    have 1: "range (?B' i) \<subseteq> sets M" using B_u by auto
-    have 2: "incseq (?B' i)" using B_mono by (auto intro!: incseq_SucI)
-    have "(\<Union>n. ?B' i n) = u -` {i} \<inter> space M"
-    proof safe
-      fix x i assume x: "x \<in> space M"
-      show "x \<in> (\<Union>i. ?B' (u x) i)"
-      proof cases
-        assume "u x = 0" thus ?thesis using \<open>x \<in> space M\<close> f(3) by simp
-      next
-        assume "u x \<noteq> 0"
-        with \<open>a < 1\<close> u_range[OF \<open>x \<in> space M\<close>]
-        have "a * u x < 1 * u x"
-          by (intro ereal_mult_strict_right_mono) (auto simp: image_iff)
-        also have "\<dots> \<le> (SUP i. f i x)" using u(2) by (auto simp: le_fun_def)
-        finally obtain i where "a * u x < f i x"
-          by (auto simp add: less_SUP_iff)
-        hence "a * u x \<le> f i x" by auto
-        thus ?thesis using \<open>x \<in> space M\<close> by auto
-      qed
-    qed
-    then show "?thesis i" using SUP_emeasure_incseq[OF 1 2] by simp
-  qed
-
-  have "integral\<^sup>S M u = (SUP i. integral\<^sup>S M (?uB i))"
-    unfolding simple_integral_indicator[OF B \<open>simple_function M u\<close>]
-  proof (subst SUP_ereal_setsum, safe)
-    fix x n assume "x \<in> space M"
-    with u_range show "incseq (\<lambda>i. u x * (emeasure M) (?B' (u x) i))" "\<And>i. 0 \<le> u x * (emeasure M) (?B' (u x) i)"
-      using B_mono B_u by (auto intro!: emeasure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff)
-  next
-    show "integral\<^sup>S M u = (\<Sum>i\<in>u ` space M. SUP n. i * (emeasure M) (?B' i n))"
-      using measure_conv u_range B_u unfolding simple_integral_def
-      by (auto intro!: setsum.cong SUP_ereal_mult_left [symmetric])
-  qed
-  moreover
-  have "a * (SUP i. integral\<^sup>S M (?uB i)) \<le> ?S"
-    apply (subst SUP_ereal_mult_left [symmetric])
-  proof (safe intro!: SUP_mono bexI)
-    fix i
-    have "a * integral\<^sup>S M (?uB i) = (\<integral>\<^sup>Sx. a * ?uB i x \<partial>M)"
-      using B \<open>simple_function M u\<close> u_range
-      by (subst simple_integral_mult) (auto split: split_indicator)
-    also have "\<dots> \<le> integral\<^sup>N M (f i)"
-    proof -
-      have *: "simple_function M (\<lambda>x. a * ?uB i x)" using B \<open>0 < a\<close> u(1) by auto
-      show ?thesis using f(3) * u_range \<open>0 < a\<close>
-        by (subst nn_integral_eq_simple_integral[symmetric])
-           (auto intro!: nn_integral_mono split: split_indicator)
-    qed
-    finally show "a * integral\<^sup>S M (?uB i) \<le> integral\<^sup>N M (f i)"
-      by auto
-  next
-    fix i show "0 \<le> \<integral>\<^sup>S x. ?uB i x \<partial>M" using B \<open>0 < a\<close> u(1) u_range
-      by (intro simple_integral_nonneg) (auto split: split_indicator)
-  qed (insert \<open>0 < a\<close>, auto)
-  ultimately show "a * integral\<^sup>S M u \<le> ?S" by simp
-qed
-
-lemma incseq_nn_integral:
-  assumes "incseq f" shows "incseq (\<lambda>i. integral\<^sup>N M (f i))"
-proof -
-  have "\<And>i x. f i x \<le> f (Suc i) x"
-    using assms by (auto dest!: incseq_SucD simp: le_fun_def)
-  then show ?thesis
-    by (auto intro!: incseq_SucI nn_integral_mono)
-qed
-
-lemma nn_integral_max_0: "(\<integral>\<^sup>+x. max 0 (f x) \<partial>M) = integral\<^sup>N M f"
-  by (simp add: le_fun_def nn_integral_def)
-
 text \<open>Beppo-Levi monotone convergence theorem\<close>
 lemma nn_integral_monotone_convergence_SUP:
-  assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M"
+  assumes f: "incseq f" and [measurable]: "\<And>i. f i \<in> borel_measurable M"
   shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))"
 proof (rule antisym)
-  show "(SUP j. integral\<^sup>N M (f j)) \<le> (\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M)"
-    by (auto intro!: SUP_least SUP_upper nn_integral_mono)
-next
-  have f': "incseq (\<lambda>i x. max 0 (f i x))"
-    using f by (auto simp: incseq_def le_fun_def not_le split: split_max)
-               (blast intro: order_trans less_imp_le)
-  have "(\<integral>\<^sup>+ x. max 0 (SUP i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (SUP i. max 0 (f i x)) \<partial>M)"
-    unfolding sup_max[symmetric] Complete_Lattices.SUP_sup_distrib[symmetric] by simp
-  also have "\<dots> \<le> (SUP i. (\<integral>\<^sup>+ x. max 0 (f i x) \<partial>M))"
-    unfolding nn_integral_def_finite[of _ "\<lambda>x. SUP i. max 0 (f i x)"]
+  show "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) \<le> (SUP i. (\<integral>\<^sup>+ x. f i x \<partial>M))"
+    unfolding nn_integral_def_finite[of _ "\<lambda>x. SUP i. f i x"]
   proof (safe intro!: SUP_least)
-    fix g assume g: "simple_function M g"
-      and *: "g \<le> max 0 \<circ> (\<lambda>x. SUP i. max 0 (f i x))" "range g \<subseteq> {0..<\<infinity>}"
-    then have "\<And>x. 0 \<le> (SUP i. max 0 (f i x))" and g': "g`space M \<subseteq> {0..<\<infinity>}"
-      using f by (auto intro!: SUP_upper2)
-    with * show "integral\<^sup>S M g \<le> (SUP j. \<integral>\<^sup>+x. max 0 (f j x) \<partial>M)"
-      by (intro nn_integral_SUP_approx[OF f' _ _ g _ g'])
-         (auto simp: le_fun_def max_def intro!: measurable_If f borel_measurable_le)
+    fix u assume sf_u[simp]: "simple_function M u" and
+      u: "u \<le> (\<lambda>x. SUP i. f i x)" and u_range: "\<forall>x. u x < top"
+    note sf_u[THEN borel_measurable_simple_function, measurable]
+    show "integral\<^sup>S M u \<le> (SUP j. \<integral>\<^sup>+x. f j x \<partial>M)"
+    proof (rule ennreal_approx_unit)
+      fix a :: ennreal assume "a < 1"
+      let ?au = "\<lambda>x. a * u x"
+
+      let ?B = "\<lambda>c i. {x\<in>space M. ?au x = c \<and> c \<le> f i x}"
+      have "integral\<^sup>S M ?au = (\<Sum>c\<in>?au`space M. c * (SUP i. emeasure M (?B c i)))"
+        unfolding simple_integral_def
+      proof (intro setsum.cong ennreal_mult_left_cong refl)
+        fix c assume "c \<in> ?au ` space M" "c \<noteq> 0"
+        { fix x' assume x': "x' \<in> space M" "?au x' = c"
+          with \<open>c \<noteq> 0\<close> u_range have "?au x' < 1 * u x'"
+            by (intro ennreal_mult_strict_right_mono \<open>a < 1\<close>) (auto simp: less_le)
+          also have "\<dots> \<le> (SUP i. f i x')"
+            using u by (auto simp: le_fun_def)
+          finally have "\<exists>i. ?au x' \<le> f i x'"
+            by (auto simp: less_SUP_iff intro: less_imp_le) }
+        then have *: "?au -` {c} \<inter> space M = (\<Union>i. ?B c i)"
+          by auto
+        show "emeasure M (?au -` {c} \<inter> space M) = (SUP i. emeasure M (?B c i))"
+          unfolding * using f
+          by (intro SUP_emeasure_incseq[symmetric])
+             (auto simp: incseq_def le_fun_def intro: order_trans)
+      qed
+      also have "\<dots> = (SUP i. \<Sum>c\<in>?au`space M. c * emeasure M (?B c i))"
+        unfolding SUP_mult_left_ennreal using f
+        by (intro ennreal_SUP_setsum[symmetric])
+           (auto intro!: mult_mono emeasure_mono simp: incseq_def le_fun_def intro: order_trans)
+      also have "\<dots> \<le> (SUP i. integral\<^sup>N M (f i))"
+      proof (intro SUP_subset_mono order_refl)
+        fix i
+        have "(\<Sum>c\<in>?au`space M. c * emeasure M (?B c i)) =
+          (\<integral>\<^sup>Sx. (a * u x) * indicator {x\<in>space M. a * u x \<le> f i x} x \<partial>M)"
+          by (subst simple_integral_indicator)
+             (auto intro!: setsum.cong ennreal_mult_left_cong arg_cong2[where f=emeasure])
+        also have "\<dots> = (\<integral>\<^sup>+x. (a * u x) * indicator {x\<in>space M. a * u x \<le> f i x} x \<partial>M)"
+          by (rule nn_integral_eq_simple_integral[symmetric]) simp
+        also have "\<dots> \<le> (\<integral>\<^sup>+x. f i x \<partial>M)"
+          by (intro nn_integral_mono) (auto split: split_indicator)
+        finally show "(\<Sum>c\<in>?au`space M. c * emeasure M (?B c i)) \<le> (\<integral>\<^sup>+x. f i x \<partial>M)" .
+      qed
+      finally show "a * integral\<^sup>S M u \<le> (SUP i. integral\<^sup>N M (f i))"
+        by simp
+    qed
   qed
-  finally show "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) \<le> (SUP j. integral\<^sup>N M (f j))"
-    unfolding nn_integral_max_0 .
+qed (auto intro!: SUP_least SUP_upper nn_integral_mono)
+
+lemma sup_continuous_nn_integral[order_continuous_intros]:
+  assumes f: "\<And>y. sup_continuous (f y)"
+  assumes [measurable]: "\<And>x. (\<lambda>y. f y x) \<in> borel_measurable M"
+  shows "sup_continuous (\<lambda>x. (\<integral>\<^sup>+y. f y x \<partial>M))"
+  unfolding sup_continuous_def
+proof safe
+  fix C :: "nat \<Rightarrow> 'b" assume C: "incseq C"
+  with sup_continuous_mono[OF f] show "(\<integral>\<^sup>+ y. f y (SUPREMUM UNIV C) \<partial>M) = (SUP i. \<integral>\<^sup>+ y. f y (C i) \<partial>M)"
+    unfolding sup_continuousD[OF f C]
+    by (subst nn_integral_monotone_convergence_SUP) (auto simp: mono_def le_fun_def)
 qed
 
 lemma nn_integral_monotone_convergence_SUP_AE:
-  assumes f: "\<And>i. AE x in M. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x" "\<And>i. f i \<in> borel_measurable M"
+  assumes f: "\<And>i. AE x in M. f i x \<le> f (Suc i) x" "\<And>i. f i \<in> borel_measurable M"
   shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))"
 proof -
-  from f have "AE x in M. \<forall>i. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x"
+  from f have "AE x in M. \<forall>i. f i x \<le> f (Suc i) x"
     by (simp add: AE_all_countable)
   from this[THEN AE_E] guess N . note N = this
   let ?f = "\<lambda>i x. if x \<in> space M - N then f i x else 0"
@@ -1038,36 +919,14 @@
   finally show ?thesis .
 qed
 
-lemma nn_integral_monotone_convergence_SUP_AE_incseq:
-  assumes f: "incseq f" "\<And>i. AE x in M. 0 \<le> f i x" and borel: "\<And>i. f i \<in> borel_measurable M"
-  shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))"
-  using f[unfolded incseq_Suc_iff le_fun_def]
-  by (intro nn_integral_monotone_convergence_SUP_AE[OF _ borel])
-     auto
-
 lemma nn_integral_monotone_convergence_simple:
-  assumes f: "incseq f" "\<And>i x. 0 \<le> f i x" "\<And>i. simple_function M (f i)"
-  shows "(SUP i. integral\<^sup>S M (f i)) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)"
-  using assms unfolding nn_integral_monotone_convergence_SUP[OF f(1)
-    f(3)[THEN borel_measurable_simple_function]]
-  by (auto intro!: nn_integral_eq_simple_integral[symmetric] arg_cong[where f="SUPREMUM UNIV"] ext)
-
-lemma nn_integral_cong_pos:
-  assumes "\<And>x. x \<in> space M \<Longrightarrow> f x \<le> 0 \<and> g x \<le> 0 \<or> f x = g x"
-  shows "integral\<^sup>N M f = integral\<^sup>N M g"
-proof -
-  have "integral\<^sup>N M (\<lambda>x. max 0 (f x)) = integral\<^sup>N M (\<lambda>x. max 0 (g x))"
-  proof (intro nn_integral_cong)
-    fix x assume "x \<in> space M"
-    from assms[OF this] show "max 0 (f x) = max 0 (g x)"
-      by (auto split: split_max)
-  qed
-  then show ?thesis by (simp add: nn_integral_max_0)
-qed
+  "incseq f \<Longrightarrow> (\<And>i. simple_function M (f i)) \<Longrightarrow> (SUP i. \<integral>\<^sup>S x. f i x \<partial>M) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)"
+  using assms nn_integral_monotone_convergence_SUP[of f M]
+  by (simp add: nn_integral_eq_simple_integral[symmetric] borel_measurable_simple_function)
 
 lemma SUP_simple_integral_sequences:
-  assumes f: "incseq f" "\<And>i x. 0 \<le> f i x" "\<And>i. simple_function M (f i)"
-  and g: "incseq g" "\<And>i x. 0 \<le> g i x" "\<And>i. simple_function M (g i)"
+  assumes f: "incseq f" "\<And>i. simple_function M (f i)"
+  and g: "incseq g" "\<And>i. simple_function M (g i)"
   and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)"
   shows "(SUP i. integral\<^sup>S M (f i)) = (SUP i. integral\<^sup>S M (g i))"
     (is "SUPREMUM _ ?F = SUPREMUM _ ?G")
@@ -1081,97 +940,59 @@
   finally show ?thesis by simp
 qed
 
-lemma nn_integral_const[simp]:
-  "0 \<le> c \<Longrightarrow> (\<integral>\<^sup>+ x. c \<partial>M) = c * (emeasure M) (space M)"
+lemma nn_integral_const[simp]: "(\<integral>\<^sup>+ x. c \<partial>M) = c * emeasure M (space M)"
   by (subst nn_integral_eq_simple_integral) auto
 
-lemma nn_integral_const_nonpos: "c \<le> 0 \<Longrightarrow> nn_integral M (\<lambda>x. c) = 0"
-  using nn_integral_max_0[of M "\<lambda>x. c"] by (simp add: max_def split: if_split_asm)
-
 lemma nn_integral_linear:
-  assumes f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" and "0 \<le> a"
-  and g: "g \<in> borel_measurable M" "\<And>x. 0 \<le> g x"
+  assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M"
   shows "(\<integral>\<^sup>+ x. a * f x + g x \<partial>M) = a * integral\<^sup>N M f + integral\<^sup>N M g"
     (is "integral\<^sup>N M ?L = _")
 proof -
   from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u .
-  note u = nn_integral_monotone_convergence_simple[OF this(2,5,1)] this
+  note u = nn_integral_monotone_convergence_simple[OF this(2,1)] this
   from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v .
-  note v = nn_integral_monotone_convergence_simple[OF this(2,5,1)] this
+  note v = nn_integral_monotone_convergence_simple[OF this(2,1)] this
   let ?L' = "\<lambda>i x. a * u i x + v i x"
 
   have "?L \<in> borel_measurable M" using assms by auto
   from borel_measurable_implies_simple_function_sequence'[OF this] guess l .
-  note l = nn_integral_monotone_convergence_simple[OF this(2,5,1)] this
+  note l = nn_integral_monotone_convergence_simple[OF this(2,1)] this
 
   have inc: "incseq (\<lambda>i. a * integral\<^sup>S M (u i))" "incseq (\<lambda>i. integral\<^sup>S M (v i))"
-    using u v \<open>0 \<le> a\<close>
-    by (auto simp: incseq_Suc_iff le_fun_def
-             intro!: add_mono ereal_mult_left_mono simple_integral_mono)
-  have pos: "\<And>i. 0 \<le> integral\<^sup>S M (u i)" "\<And>i. 0 \<le> integral\<^sup>S M (v i)" "\<And>i. 0 \<le> a * integral\<^sup>S M (u i)"
-    using u v \<open>0 \<le> a\<close> by (auto simp: simple_integral_nonneg)
-  { fix i from pos[of i] have "a * integral\<^sup>S M (u i) \<noteq> -\<infinity>" "integral\<^sup>S M (v i) \<noteq> -\<infinity>"
-      by (auto split: if_split_asm) }
-  note not_MInf = this
+    using u v by (auto simp: incseq_Suc_iff le_fun_def intro!: add_mono mult_left_mono simple_integral_mono)
 
   have l': "(SUP i. integral\<^sup>S M (l i)) = (SUP i. integral\<^sup>S M (?L' i))"
-  proof (rule SUP_simple_integral_sequences[OF l(3,6,2)])
-    show "incseq ?L'" "\<And>i x. 0 \<le> ?L' i x" "\<And>i. simple_function M (?L' i)"
-      using u v  \<open>0 \<le> a\<close> unfolding incseq_Suc_iff le_fun_def
-      by (auto intro!: add_mono ereal_mult_left_mono)
+  proof (rule SUP_simple_integral_sequences[OF l(3,2)])
+    show "incseq ?L'" "\<And>i. simple_function M (?L' i)"
+      using u v unfolding incseq_Suc_iff le_fun_def
+      by (auto intro!: add_mono mult_left_mono)
     { fix x
-      { fix i have "a * u i x \<noteq> -\<infinity>" "v i x \<noteq> -\<infinity>" "u i x \<noteq> -\<infinity>" using \<open>0 \<le> a\<close> u(6)[of i x] v(6)[of i x]
-          by auto }
-      then have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)"
-        using \<open>0 \<le> a\<close> u(3) v(3) u(6)[of _ x] v(6)[of _ x]
-        by (subst SUP_ereal_mult_left [symmetric, OF _ u(6) \<open>0 \<le> a\<close>])
-           (auto intro!: SUP_ereal_add
-                 simp: incseq_Suc_iff le_fun_def add_mono ereal_mult_left_mono) }
+      have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)"
+        using u(3) v(3) u(4)[of _ x] v(4)[of _ x] unfolding SUP_mult_left_ennreal
+        by (auto intro!: ennreal_SUP_add simp: incseq_Suc_iff le_fun_def add_mono mult_left_mono) }
     then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)"
-      unfolding l(5) using \<open>0 \<le> a\<close> u(5) v(5) l(5) f(2) g(2)
-      by (intro AE_I2) (auto split: split_max)
+      unfolding l(5) using u(5) v(5) by (intro AE_I2) auto
   qed
   also have "\<dots> = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))"
-    using u(2, 6) v(2, 6) \<open>0 \<le> a\<close> by (auto intro!: arg_cong[where f="SUPREMUM UNIV"] ext)
-  finally have "(\<integral>\<^sup>+ x. max 0 (a * f x + g x) \<partial>M) = a * (\<integral>\<^sup>+x. max 0 (f x) \<partial>M) + (\<integral>\<^sup>+x. max 0 (g x) \<partial>M)"
-    unfolding l(5)[symmetric] u(5)[symmetric] v(5)[symmetric]
-    unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric]
-    apply (subst SUP_ereal_mult_left [symmetric, OF _ pos(1) \<open>0 \<le> a\<close>])
-    apply simp
-    apply (subst SUP_ereal_add [symmetric, OF inc not_MInf])
-    .
-  then show ?thesis by (simp add: nn_integral_max_0)
+    using u(2) v(2) by auto
+  finally show ?thesis
+    unfolding l(5)[symmetric] l(1)[symmetric]
+    by (simp add: ennreal_SUP_add[OF inc] v u SUP_mult_left_ennreal[symmetric])
 qed
 
-lemma nn_integral_cmult:
-  assumes f: "f \<in> borel_measurable M" "0 \<le> c"
-  shows "(\<integral>\<^sup>+ x. c * f x \<partial>M) = c * integral\<^sup>N M f"
-proof -
-  have [simp]: "\<And>x. c * max 0 (f x) = max 0 (c * f x)" using \<open>0 \<le> c\<close>
-    by (auto split: split_max simp: ereal_zero_le_0_iff)
-  have "(\<integral>\<^sup>+ x. c * f x \<partial>M) = (\<integral>\<^sup>+ x. c * max 0 (f x) \<partial>M)"
-    by (simp add: nn_integral_max_0)
-  then show ?thesis
-    using nn_integral_linear[OF _ _ \<open>0 \<le> c\<close>, of "\<lambda>x. max 0 (f x)" _ "\<lambda>x. 0"] f
-    by (auto simp: nn_integral_max_0)
-qed
+lemma nn_integral_cmult: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. c * f x \<partial>M) = c * integral\<^sup>N M f"
+  using nn_integral_linear[of f M "\<lambda>x. 0" c] by simp
 
-lemma nn_integral_multc:
-  assumes "f \<in> borel_measurable M" "0 \<le> c"
-  shows "(\<integral>\<^sup>+ x. f x * c \<partial>M) = integral\<^sup>N M f * c"
+lemma nn_integral_multc: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x * c \<partial>M) = integral\<^sup>N M f * c"
   unfolding mult.commute[of _ c] nn_integral_cmult[OF assms] by simp
 
-lemma nn_integral_divide:
-   "\<lbrakk> 0 \<le> c; f \<in> borel_measurable M \<rbrakk> \<Longrightarrow> (\<integral>\<^sup>+ x. f x / c \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M) / c"
-by(simp add: divide_ereal_def nn_integral_multc inverse_ereal_ge0I)
+lemma nn_integral_divide: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x / c \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M) / c"
+   unfolding divide_ennreal_def by (rule nn_integral_multc)
 
-lemma nn_integral_indicator[simp]:
-  "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. indicator A x\<partial>M) = (emeasure M) A"
-  by (subst nn_integral_eq_simple_integral)
-     (auto simp: simple_integral_indicator)
+lemma nn_integral_indicator[simp]: "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. indicator A x\<partial>M) = (emeasure M) A"
+  by (subst nn_integral_eq_simple_integral) (auto simp: simple_integral_indicator)
 
-lemma nn_integral_cmult_indicator:
-  "0 \<le> c \<Longrightarrow> A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. c * indicator A x \<partial>M) = c * (emeasure M) A"
+lemma nn_integral_cmult_indicator: "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. c * indicator A x \<partial>M) = c * emeasure M A"
   by (subst nn_integral_eq_simple_integral)
      (auto simp: simple_function_indicator simple_integral_indicator)
 
@@ -1187,61 +1008,51 @@
 qed
 
 lemma nn_integral_indicator_singleton[simp]:
-  assumes [measurable]: "{y} \<in> sets M"
-  shows "(\<integral>\<^sup>+x. f x * indicator {y} x \<partial>M) = max 0 (f y) * emeasure M {y}"
-proof-
-  have "(\<integral>\<^sup>+x. f x * indicator {y} x \<partial>M) = (\<integral>\<^sup>+x. max 0 (f y) * indicator {y} x \<partial>M)"
-    by (subst nn_integral_max_0[symmetric]) (auto intro!: nn_integral_cong split: split_indicator)
+  assumes [measurable]: "{y} \<in> sets M" shows "(\<integral>\<^sup>+x. f x * indicator {y} x \<partial>M) = f y * emeasure M {y}"
+proof -
+  have "(\<integral>\<^sup>+x. f x * indicator {y} x \<partial>M) = (\<integral>\<^sup>+x. f y * indicator {y} x \<partial>M)"
+    by (auto intro!: nn_integral_cong split: split_indicator)
   then show ?thesis
     by (simp add: nn_integral_cmult)
 qed
 
-lemma nn_integral_set_ereal:
-  "(\<integral>\<^sup>+x. ereal (f x) * indicator A x \<partial>M) = (\<integral>\<^sup>+x. ereal (f x * indicator A x) \<partial>M)"
+lemma nn_integral_set_ennreal:
+  "(\<integral>\<^sup>+x. ennreal (f x) * indicator A x \<partial>M) = (\<integral>\<^sup>+x. ennreal (f x * indicator A x) \<partial>M)"
   by (rule nn_integral_cong) (simp split: split_indicator)
 
 lemma nn_integral_indicator_singleton'[simp]:
   assumes [measurable]: "{y} \<in> sets M"
-  shows "(\<integral>\<^sup>+x. ereal (f x * indicator {y} x) \<partial>M) = max 0 (f y) * emeasure M {y}"
-  by (subst nn_integral_set_ereal[symmetric]) (simp add: nn_integral_indicator_singleton)
+  shows "(\<integral>\<^sup>+x. ennreal (f x * indicator {y} x) \<partial>M) = f y * emeasure M {y}"
+  by (subst nn_integral_set_ennreal[symmetric]) (simp add: nn_integral_indicator_singleton)
 
 lemma nn_integral_add:
-  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
-  and g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
-  shows "(\<integral>\<^sup>+ x. f x + g x \<partial>M) = integral\<^sup>N M f + integral\<^sup>N M g"
-proof -
-  have ae: "AE x in M. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)"
-    using assms by (auto split: split_max)
-  have "(\<integral>\<^sup>+ x. f x + g x \<partial>M) = (\<integral>\<^sup>+ x. max 0 (f x + g x) \<partial>M)"
-    by (simp add: nn_integral_max_0)
-  also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) + max 0 (g x) \<partial>M)"
-    unfolding ae[THEN nn_integral_cong_AE] ..
-  also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) \<partial>M) + (\<integral>\<^sup>+ x. max 0 (g x) \<partial>M)"
-    using nn_integral_linear[of "\<lambda>x. max 0 (f x)" _ 1 "\<lambda>x. max 0 (g x)"] f g
-    by auto
-  finally show ?thesis
-    by (simp add: nn_integral_max_0)
-qed
+  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x + g x \<partial>M) = integral\<^sup>N M f + integral\<^sup>N M g"
+  using nn_integral_linear[of f M g 1] by simp
 
 lemma nn_integral_setsum:
-  assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M" "\<And>i. i\<in>P \<Longrightarrow> AE x in M. 0 \<le> f i x"
-  shows "(\<integral>\<^sup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>N M (f i))"
-proof cases
-  assume f: "finite P"
-  from assms have "AE x in M. \<forall>i\<in>P. 0 \<le> f i x" unfolding AE_finite_all[OF f] by auto
-  from f this assms(1) show ?thesis
-  proof induct
-    case (insert i P)
-    then have "f i \<in> borel_measurable M" "AE x in M. 0 \<le> f i x"
-      "(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M" "AE x in M. 0 \<le> (\<Sum>i\<in>P. f i x)"
-      by (auto intro!: setsum_nonneg)
-    from nn_integral_add[OF this]
-    show ?case using insert by auto
-  qed simp
-qed simp
+  "(\<And>i. i \<in> P \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<integral>\<^sup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>N M (f i))"
+  by (induction P rule: infinite_finite_induct) (auto simp: nn_integral_add)
+
+lemma nn_integral_suminf:
+  assumes f: "\<And>i. f i \<in> borel_measurable M"
+  shows "(\<integral>\<^sup>+ x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^sup>N M (f i))"
+proof -
+  have all_pos: "AE x in M. \<forall>i. 0 \<le> f i x"
+    using assms by (auto simp: AE_all_countable)
+  have "(\<Sum>i. integral\<^sup>N M (f i)) = (SUP n. \<Sum>i<n. integral\<^sup>N M (f i))"
+    by (rule suminf_eq_SUP)
+  also have "\<dots> = (SUP n. \<integral>\<^sup>+x. (\<Sum>i<n. f i x) \<partial>M)"
+    unfolding nn_integral_setsum[OF f] ..
+  also have "\<dots> = \<integral>\<^sup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos
+    by (intro nn_integral_monotone_convergence_SUP_AE[symmetric])
+       (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3)
+  also have "\<dots> = \<integral>\<^sup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos
+    by (intro nn_integral_cong_AE) (auto simp: suminf_eq_SUP)
+  finally show ?thesis by simp
+qed
 
 lemma nn_integral_bound_simple_function:
-  assumes bnd: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x" "\<And>x. x \<in> space M \<Longrightarrow> f x < \<infinity>"
+  assumes bnd: "\<And>x. x \<in> space M \<Longrightarrow> f x < \<infinity>"
   assumes f[measurable]: "simple_function M f"
   assumes supp: "emeasure M {x\<in>space M. f x \<noteq> 0} < \<infinity>"
   shows "nn_integral M f < \<infinity>"
@@ -1262,12 +1073,12 @@
       by (auto split: split_indicator intro!: Max_ge simple_functionD)
   qed
   also have "\<dots> < \<infinity>"
-    using bnd supp by (subst nn_integral_cmult) auto
+    using bnd supp by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top)
   finally show ?thesis .
 qed
 
 lemma nn_integral_Markov_inequality:
-  assumes u: "u \<in> borel_measurable M" "AE x in M. 0 \<le> u x" and "A \<in> sets M" and c: "0 \<le> c"
+  assumes u: "u \<in> borel_measurable M" and "A \<in> sets M"
   shows "(emeasure M) ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^sup>+ x. u x * indicator A x \<partial>M)"
     (is "(emeasure M) ?A \<le> _ * ?PI")
 proof -
@@ -1275,50 +1086,50 @@
     using \<open>A \<in> sets M\<close> u by auto
   hence "(emeasure M) ?A = (\<integral>\<^sup>+ x. indicator ?A x \<partial>M)"
     using nn_integral_indicator by simp
-  also have "\<dots> \<le> (\<integral>\<^sup>+ x. c * (u x * indicator A x) \<partial>M)" using u c
-    by (auto intro!: nn_integral_mono_AE
-      simp: indicator_def ereal_zero_le_0_iff)
+  also have "\<dots> \<le> (\<integral>\<^sup>+ x. c * (u x * indicator A x) \<partial>M)"
+    using u by (auto intro!: nn_integral_mono_AE simp: indicator_def)
   also have "\<dots> = c * (\<integral>\<^sup>+ x. u x * indicator A x \<partial>M)"
-    using assms
-    by (auto intro!: nn_integral_cmult simp: ereal_zero_le_0_iff)
+    using assms by (auto intro!: nn_integral_cmult)
   finally show ?thesis .
 qed
 
 lemma nn_integral_noteq_infinite:
-  assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
-  and "integral\<^sup>N M g \<noteq> \<infinity>"
+  assumes g: "g \<in> borel_measurable M" and "integral\<^sup>N M g \<noteq> \<infinity>"
   shows "AE x in M. g x \<noteq> \<infinity>"
 proof (rule ccontr)
   assume c: "\<not> (AE x in M. g x \<noteq> \<infinity>)"
   have "(emeasure M) {x\<in>space M. g x = \<infinity>} \<noteq> 0"
     using c g by (auto simp add: AE_iff_null)
-  moreover have "0 \<le> (emeasure M) {x\<in>space M. g x = \<infinity>}" using g by (auto intro: measurable_sets)
-  ultimately have "0 < (emeasure M) {x\<in>space M. g x = \<infinity>}" by auto
-  then have "\<infinity> = \<infinity> * (emeasure M) {x\<in>space M. g x = \<infinity>}" by auto
+  then have "0 < (emeasure M) {x\<in>space M. g x = \<infinity>}"
+    by (auto simp: zero_less_iff_neq_zero)
+  then have "\<infinity> = \<infinity> * (emeasure M) {x\<in>space M. g x = \<infinity>}"
+    by (auto simp: ennreal_top_eq_mult_iff)
   also have "\<dots> \<le> (\<integral>\<^sup>+x. \<infinity> * indicator {x\<in>space M. g x = \<infinity>} x \<partial>M)"
     using g by (subst nn_integral_cmult_indicator) auto
   also have "\<dots> \<le> integral\<^sup>N M g"
     using assms by (auto intro!: nn_integral_mono_AE simp: indicator_def)
-  finally show False using \<open>integral\<^sup>N M g \<noteq> \<infinity>\<close> by auto
+  finally show False
+    using \<open>integral\<^sup>N M g \<noteq> \<infinity>\<close> by (auto simp: top_unique)
 qed
 
 lemma nn_integral_PInf:
-  assumes f: "f \<in> borel_measurable M"
-  and not_Inf: "integral\<^sup>N M f \<noteq> \<infinity>"
-  shows "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
+  assumes f: "f \<in> borel_measurable M" and not_Inf: "integral\<^sup>N M f \<noteq> \<infinity>"
+  shows "emeasure M (f -` {\<infinity>} \<inter> space M) = 0"
 proof -
-  have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^sup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
+  have "\<infinity> * emeasure M (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^sup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
     using f by (subst nn_integral_cmult_indicator) (auto simp: measurable_sets)
-  also have "\<dots> \<le> integral\<^sup>N M (\<lambda>x. max 0 (f x))"
-    by (auto intro!: nn_integral_mono simp: indicator_def max_def)
+  also have "\<dots> \<le> integral\<^sup>N M f"
+    by (auto intro!: nn_integral_mono simp: indicator_def)
   finally have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) \<le> integral\<^sup>N M f"
-    by (simp add: nn_integral_max_0)
-  moreover have "0 \<le> (emeasure M) (f -` {\<infinity>} \<inter> space M)"
-    by (rule emeasure_nonneg)
-  ultimately show ?thesis
-    using assms by (auto split: if_split_asm)
+    by simp
+  then show ?thesis
+    using assms by (auto simp: ennreal_top_mult top_unique split: if_split_asm)
 qed
 
+lemma simple_integral_PInf:
+  "simple_function M f \<Longrightarrow> integral\<^sup>S M f \<noteq> \<infinity> \<Longrightarrow> emeasure M (f -` {\<infinity>} \<inter> space M) = 0"
+  by (rule nn_integral_PInf) (auto simp: nn_integral_eq_simple_integral borel_measurable_simple_function)
+
 lemma nn_integral_PInf_AE:
   assumes "f \<in> borel_measurable M" "integral\<^sup>N M f \<noteq> \<infinity>" shows "AE x in M. f x \<noteq> \<infinity>"
 proof (rule AE_I)
@@ -1328,147 +1139,141 @@
     using assms by (auto intro: borel_measurable_vimage)
 qed auto
 
-lemma simple_integral_PInf:
-  assumes "simple_function M f" "\<And>x. 0 \<le> f x"
-  and "integral\<^sup>S M f \<noteq> \<infinity>"
-  shows "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
-proof (rule nn_integral_PInf)
-  show "f \<in> borel_measurable M" using assms by (auto intro: borel_measurable_simple_function)
-  show "integral\<^sup>N M f \<noteq> \<infinity>"
-    using assms by (simp add: nn_integral_eq_simple_integral)
-qed
-
 lemma nn_integral_diff:
   assumes f: "f \<in> borel_measurable M"
-  and g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
+  and g: "g \<in> borel_measurable M"
   and fin: "integral\<^sup>N M g \<noteq> \<infinity>"
   and mono: "AE x in M. g x \<le> f x"
   shows "(\<integral>\<^sup>+ x. f x - g x \<partial>M) = integral\<^sup>N M f - integral\<^sup>N M g"
 proof -
-  have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M" "AE x in M. 0 \<le> f x - g x"
-    using assms by (auto intro: ereal_diff_positive)
-  have pos_f: "AE x in M. 0 \<le> f x" using mono g by auto
-  { fix a b :: ereal assume "0 \<le> a" "a \<noteq> \<infinity>" "0 \<le> b" "a \<le> b" then have "b - a + a = b"
-      by (cases rule: ereal2_cases[of a b]) auto }
-  note * = this
-  then have "AE x in M. f x = f x - g x + g x"
-    using mono nn_integral_noteq_infinite[OF g fin] assms by auto
+  have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M"
+    using assms by auto
+  have "AE x in M. f x = f x - g x + g x"
+    using diff_add_cancel_ennreal mono nn_integral_noteq_infinite[OF g fin] assms by auto
   then have **: "integral\<^sup>N M f = (\<integral>\<^sup>+x. f x - g x \<partial>M) + integral\<^sup>N M g"
     unfolding nn_integral_add[OF diff g, symmetric]
     by (rule nn_integral_cong_AE)
   show ?thesis unfolding **
-    using fin nn_integral_nonneg[of M g]
-    by (cases rule: ereal2_cases[of "\<integral>\<^sup>+ x. f x - g x \<partial>M" "integral\<^sup>N M g"]) auto
-qed
-
-lemma nn_integral_suminf:
-  assumes f: "\<And>i. f i \<in> borel_measurable M" "\<And>i. AE x in M. 0 \<le> f i x"
-  shows "(\<integral>\<^sup>+ x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^sup>N M (f i))"
-proof -
-  have all_pos: "AE x in M. \<forall>i. 0 \<le> f i x"
-    using assms by (auto simp: AE_all_countable)
-  have "(\<Sum>i. integral\<^sup>N M (f i)) = (SUP n. \<Sum>i<n. integral\<^sup>N M (f i))"
-    using nn_integral_nonneg by (rule suminf_ereal_eq_SUP)
-  also have "\<dots> = (SUP n. \<integral>\<^sup>+x. (\<Sum>i<n. f i x) \<partial>M)"
-    unfolding nn_integral_setsum[OF f] ..
-  also have "\<dots> = \<integral>\<^sup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos
-    by (intro nn_integral_monotone_convergence_SUP_AE[symmetric])
-       (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3)
-  also have "\<dots> = \<integral>\<^sup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos
-    by (intro nn_integral_cong_AE) (auto simp: suminf_ereal_eq_SUP)
-  finally show ?thesis by simp
+    using fin
+    by (cases rule: ennreal2_cases[of "\<integral>\<^sup>+ x. f x - g x \<partial>M" "integral\<^sup>N M g"]) auto
 qed
 
 lemma nn_integral_mult_bounded_inf:
-  assumes f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>"
-    and c: "0 \<le> c" "c \<noteq> \<infinity>" and ae: "AE x in M. g x \<le> c * f x"
+  assumes f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>" and c: "c \<noteq> \<infinity>" and ae: "AE x in M. g x \<le> c * f x"
   shows "(\<integral>\<^sup>+x. g x \<partial>M) < \<infinity>"
 proof -
   have "(\<integral>\<^sup>+x. g x \<partial>M) \<le> (\<integral>\<^sup>+x. c * f x \<partial>M)"
     by (intro nn_integral_mono_AE ae)
   also have "(\<integral>\<^sup>+x. c * f x \<partial>M) < \<infinity>"
-    using c f by (subst nn_integral_cmult) auto
+    using c f by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top top_unique not_less)
   finally show ?thesis .
 qed
 
 text \<open>Fatou's lemma: convergence theorem on limes inferior\<close>
 
-lemma nn_integral_liminf:
-  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
-  assumes u: "\<And>i. u i \<in> borel_measurable M" "\<And>i. AE x in M. 0 \<le> u i x"
-  shows "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))"
+lemma nn_integral_monotone_convergence_INF_AE':
+  assumes f: "\<And>i. AE x in M. f (Suc i) x \<le> f i x" and [measurable]: "\<And>i. f i \<in> borel_measurable M"
+    and *: "(\<integral>\<^sup>+ x. f 0 x \<partial>M) < \<infinity>"
+  shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
+proof (rule ennreal_minus_cancel)
+  have "integral\<^sup>N M (f 0) - (\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (\<integral>\<^sup>+x. f 0 x - (INF i. f i x) \<partial>M)"
+  proof (rule nn_integral_diff[symmetric])
+    have "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) \<le> (\<integral>\<^sup>+ x. f 0 x \<partial>M)"
+      by (intro nn_integral_mono INF_lower) simp
+    with * show "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) \<noteq> \<infinity>"
+      by simp
+  qed (auto intro: INF_lower)
+  also have "\<dots> = (\<integral>\<^sup>+x. (SUP i. f 0 x - f i x) \<partial>M)"
+    by (simp add: ennreal_INF_const_minus)
+  also have "\<dots> = (SUP i. (\<integral>\<^sup>+x. f 0 x - f i x \<partial>M))"
+  proof (intro nn_integral_monotone_convergence_SUP_AE)
+    show "AE x in M. f 0 x - f i x \<le> f 0 x - f (Suc i) x" for i
+      using f[of i] by eventually_elim (auto simp: ennreal_mono_minus)
+  qed simp
+  also have "\<dots> = (SUP i. nn_integral M (f 0) - (\<integral>\<^sup>+x. f i x \<partial>M))"
+  proof (subst nn_integral_diff[symmetric])
+    fix i
+    have dec: "AE x in M. \<forall>i. f (Suc i) x \<le> f i x"
+      unfolding AE_all_countable using f by auto
+    then show "AE x in M. f i x \<le> f 0 x"
+      using dec by eventually_elim (auto intro: lift_Suc_antimono_le[of "\<lambda>i. f i x" 0 i for x])
+    then have "(\<integral>\<^sup>+ x. f i x \<partial>M) \<le> (\<integral>\<^sup>+ x. f 0 x \<partial>M)"
+      by (rule nn_integral_mono_AE)
+    with * show "(\<integral>\<^sup>+ x. f i x \<partial>M) \<noteq> \<infinity>"
+      by simp
+  qed (insert f, auto simp: decseq_def le_fun_def)
+  finally show "integral\<^sup>N M (f 0) - (\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) =
+    integral\<^sup>N M (f 0) - (INF i. \<integral>\<^sup>+ x. f i x \<partial>M)"
+    by (simp add: ennreal_INF_const_minus)
+qed (insert *, auto intro!: nn_integral_mono intro: INF_lower)
+
+lemma nn_integral_monotone_convergence_INF_AE:
+  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal"
+  assumes f: "\<And>i. AE x in M. f (Suc i) x \<le> f i x"
+    and [measurable]: "\<And>i. f i \<in> borel_measurable M"
+    and fin: "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>"
+  shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
 proof -
-  have pos: "AE x in M. \<forall>i. 0 \<le> u i x" using u by (auto simp: AE_all_countable)
-  have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) =
-    (SUP n. \<integral>\<^sup>+ x. (INF i:{n..}. u i x) \<partial>M)"
-    unfolding liminf_SUP_INF using pos u
-    by (intro nn_integral_monotone_convergence_SUP_AE)
-       (elim AE_mp, auto intro!: AE_I2 intro: INF_greatest INF_superset_mono)
-  also have "\<dots> \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))"
-    unfolding liminf_SUP_INF
-    by (auto intro!: SUP_mono exI INF_greatest nn_integral_mono INF_lower)
+  { fix f :: "nat \<Rightarrow> ennreal" and j assume "decseq f"
+    then have "(INF i. f i) = (INF i. f (i + j))"
+      apply (intro INF_eq)
+      apply (rule_tac x="i" in bexI)
+      apply (auto simp: decseq_def le_fun_def)
+      done }
+  note INF_shift = this
+  have mono: "AE x in M. \<forall>i. f (Suc i) x \<le> f i x"
+    using f by (auto simp: AE_all_countable)
+  then have "AE x in M. (INF i. f i x) = (INF n. f (n + i) x)"
+    by eventually_elim (auto intro!: decseq_SucI INF_shift)
+  then have "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (INF n. f (n + i) x) \<partial>M)"
+    by (rule nn_integral_cong_AE)
+  also have "\<dots> = (INF n. (\<integral>\<^sup>+ x. f (n + i) x \<partial>M))"
+    by (rule nn_integral_monotone_convergence_INF_AE') (insert assms, auto)
+  also have "\<dots> = (INF n. (\<integral>\<^sup>+ x. f n x \<partial>M))"
+    by (intro INF_shift[symmetric] decseq_SucI nn_integral_mono_AE f)
   finally show ?thesis .
 qed
 
-lemma le_Limsup:
-  "F \<noteq> bot \<Longrightarrow> eventually (\<lambda>x. c \<le> g x) F \<Longrightarrow> c \<le> Limsup F g"
-  using Limsup_mono[of "\<lambda>_. c" g F] by (simp add: Limsup_const)
+lemma nn_integral_monotone_convergence_INF_decseq:
+  assumes f: "decseq f" and *: "\<And>i. f i \<in> borel_measurable M" "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>"
+  shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
+  using nn_integral_monotone_convergence_INF_AE[of f M i, OF _ *] f by (auto simp: decseq_Suc_iff le_fun_def)
 
-lemma Limsup_le:
-  "F \<noteq> bot \<Longrightarrow> eventually (\<lambda>x. f x \<le> c) F \<Longrightarrow> Limsup F f \<le> c"
-  using Limsup_mono[of f "\<lambda>_. c" F] by (simp add: Limsup_const)
-
-lemma ereal_mono_minus_cancel:
-  fixes a b c :: ereal
-  shows "c - a \<le> c - b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c < \<infinity> \<Longrightarrow> b \<le> a"
-  by (cases a b c rule: ereal3_cases) auto
+lemma nn_integral_liminf:
+  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal"
+  assumes u: "\<And>i. u i \<in> borel_measurable M"
+  shows "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))"
+proof -
+  have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) = (SUP n. \<integral>\<^sup>+ x. (INF i:{n..}. u i x) \<partial>M)"
+    unfolding liminf_SUP_INF using u
+    by (intro nn_integral_monotone_convergence_SUP_AE)
+       (auto intro!: AE_I2 intro: INF_greatest INF_superset_mono)
+  also have "\<dots> \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))"
+    by (auto simp: liminf_SUP_INF intro!: SUP_mono INF_greatest nn_integral_mono INF_lower)
+  finally show ?thesis .
+qed
 
 lemma nn_integral_limsup:
-  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
+  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal"
   assumes [measurable]: "\<And>i. u i \<in> borel_measurable M" "w \<in> borel_measurable M"
-  assumes bounds: "\<And>i. AE x in M. 0 \<le> u i x" "\<And>i. AE x in M. u i x \<le> w x" and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>"
+  assumes bounds: "\<And>i. AE x in M. u i x \<le> w x" and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>"
   shows "limsup (\<lambda>n. integral\<^sup>N M (u n)) \<le> (\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M)"
 proof -
-  have bnd: "AE x in M. \<forall>i. 0 \<le> u i x \<and> u i x \<le> w x"
+  have bnd: "AE x in M. \<forall>i. u i x \<le> w x"
     using bounds by (auto simp: AE_all_countable)
-
-  from bounds[of 0] have w_nonneg: "AE x in M. 0 \<le> w x"
-    by auto
-
-  have "(\<integral>\<^sup>+x. w x \<partial>M) - (\<integral>\<^sup>+x. limsup (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+x. w x - limsup (\<lambda>n. u n x) \<partial>M)"
-  proof (intro nn_integral_diff[symmetric])
-    show "AE x in M. 0 \<le> limsup (\<lambda>n. u n x)"
-      using bnd by (auto intro!: le_Limsup)
-    show "AE x in M. limsup (\<lambda>n. u n x) \<le> w x"
-      using bnd by (auto intro!: Limsup_le)
-    then have "(\<integral>\<^sup>+x. limsup (\<lambda>n. u n x) \<partial>M) < \<infinity>"
-      by (intro nn_integral_mult_bounded_inf[OF _ w, of 1]) auto
-    then show "(\<integral>\<^sup>+x. limsup (\<lambda>n. u n x) \<partial>M) \<noteq> \<infinity>"
-      by simp
-  qed auto
-  also have "\<dots> = (\<integral>\<^sup>+x. liminf (\<lambda>n. w x - u n x) \<partial>M)"
-    using w_nonneg
-    by (intro nn_integral_cong_AE, eventually_elim)
-       (auto intro!: liminf_ereal_cminus[symmetric])
-  also have "\<dots> \<le> liminf (\<lambda>n. \<integral>\<^sup>+x. w x - u n x \<partial>M)"
-  proof (rule nn_integral_liminf)
-    fix i show "AE x in M. 0 \<le> w x - u i x"
-      using bounds[of i] by eventually_elim (auto intro: ereal_diff_positive)
-  qed simp
-  also have "(\<lambda>n. \<integral>\<^sup>+x. w x - u n x \<partial>M) = (\<lambda>n. (\<integral>\<^sup>+x. w x \<partial>M) - (\<integral>\<^sup>+x. u n x \<partial>M))"
-  proof (intro ext nn_integral_diff)
-    fix i have "(\<integral>\<^sup>+x. u i x \<partial>M) < \<infinity>"
-      using bounds by (intro nn_integral_mult_bounded_inf[OF _ w, of 1]) auto
-    then show "(\<integral>\<^sup>+x. u i x \<partial>M) \<noteq> \<infinity>" by simp
-  qed (insert bounds, auto)
-  also have "liminf (\<lambda>n. (\<integral>\<^sup>+x. w x \<partial>M) - (\<integral>\<^sup>+x. u n x \<partial>M)) = (\<integral>\<^sup>+x. w x \<partial>M) - limsup (\<lambda>n. \<integral>\<^sup>+x. u n x \<partial>M)"
-    using w by (intro liminf_ereal_cminus) auto
-  finally show ?thesis
-    by (rule ereal_mono_minus_cancel) (intro w nn_integral_nonneg)+
+  then have "(\<integral>\<^sup>+ x. (SUP n. u n x) \<partial>M) \<le> (\<integral>\<^sup>+ x. w x \<partial>M)"
+    by (auto intro!: nn_integral_mono_AE elim: eventually_mono intro: SUP_least)
+  then have "(\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M) = (INF n. \<integral>\<^sup>+ x. (SUP i:{n..}. u i x) \<partial>M)"
+    unfolding limsup_INF_SUP using bnd w
+    by (intro nn_integral_monotone_convergence_INF_AE')
+       (auto intro!: AE_I2 intro: SUP_least SUP_subset_mono)
+  also have "\<dots> \<ge> limsup (\<lambda>n. integral\<^sup>N M (u n))"
+    by (auto simp: limsup_INF_SUP intro!: INF_mono SUP_least exI nn_integral_mono SUP_upper)
+  finally (xtrans) show ?thesis .
 qed
 
 lemma nn_integral_LIMSEQ:
-  assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>n x. 0 \<le> f n x"
+  assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M"
     and u: "\<And>x. (\<lambda>i. f i x) \<longlonglongrightarrow> u x"
   shows "(\<lambda>n. integral\<^sup>N M (f n)) \<longlonglongrightarrow> integral\<^sup>N M u"
 proof -
@@ -1477,14 +1282,14 @@
   also have "(SUP n. integral\<^sup>N M (f n)) = integral\<^sup>N M (\<lambda>x. SUP n. f n x)"
     using f by (intro nn_integral_monotone_convergence_SUP[symmetric])
   also have "integral\<^sup>N M (\<lambda>x. SUP n. f n x) = integral\<^sup>N M (\<lambda>x. u x)"
-    using f by (subst SUP_Lim_ereal[OF _ u]) (auto simp: incseq_def le_fun_def)
+    using f by (subst LIMSEQ_SUP[THEN LIMSEQ_unique, OF _ u]) (auto simp: incseq_def le_fun_def)
   finally show ?thesis .
 qed
 
 lemma nn_integral_dominated_convergence:
   assumes [measurable]:
        "\<And>i. u i \<in> borel_measurable M" "u' \<in> borel_measurable M" "w \<in> borel_measurable M"
-    and bound: "\<And>j. AE x in M. 0 \<le> u j x" "\<And>j. AE x in M. u j x \<le> w x"
+    and bound: "\<And>j. AE x in M. u j x \<le> w x"
     and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>"
     and u': "AE x in M. (\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
   shows "(\<lambda>i. (\<integral>\<^sup>+x. u i x \<partial>M)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. u' x \<partial>M)"
@@ -1496,89 +1301,13 @@
   moreover have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+ x. u' x \<partial>M)"
     using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot)
   moreover have "(\<integral>\<^sup>+x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))"
-    by (intro nn_integral_liminf[OF _ bound(1)]) auto
+    by (intro nn_integral_liminf) auto
   moreover have "liminf (\<lambda>n. integral\<^sup>N M (u n)) \<le> limsup (\<lambda>n. integral\<^sup>N M (u n))"
     by (intro Liminf_le_Limsup sequentially_bot)
   ultimately show ?thesis
     by (intro Liminf_eq_Limsup) auto
 qed
 
-lemma nn_integral_monotone_convergence_INF':
-  assumes f: "decseq f" and [measurable]: "\<And>i. f i \<in> borel_measurable M"
-  assumes "(\<integral>\<^sup>+ x. f 0 x \<partial>M) < \<infinity>" and nn: "\<And>x i. 0 \<le> f i x"
-  shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
-proof (rule LIMSEQ_unique)
-  show "(\<lambda>i. integral\<^sup>N M (f i)) \<longlonglongrightarrow> (INF i. integral\<^sup>N M (f i))"
-    using f by (intro LIMSEQ_INF) (auto intro!: nn_integral_mono simp: decseq_def le_fun_def)
-  show "(\<lambda>i. integral\<^sup>N M (f i)) \<longlonglongrightarrow> \<integral>\<^sup>+ x. (INF i. f i x) \<partial>M"
-  proof (rule nn_integral_dominated_convergence)
-    show "(\<integral>\<^sup>+ x. f 0 x \<partial>M) < \<infinity>" "\<And>i. f i \<in> borel_measurable M" "f 0 \<in> borel_measurable M"
-      by fact+
-    show "\<And>j. AE x in M. 0 \<le> f j x"
-      using nn by auto
-    show "\<And>j. AE x in M. f j x \<le> f 0 x"
-      using f by (auto simp: decseq_def le_fun_def)
-    show "AE x in M. (\<lambda>i. f i x) \<longlonglongrightarrow> (INF i. f i x)"
-      using f by (auto intro!: LIMSEQ_INF simp: decseq_def le_fun_def)
-    show "(\<lambda>x. INF i. f i x) \<in> borel_measurable M"
-      by auto
-  qed
-qed
-
-lemma nn_integral_monotone_convergence_INF:
-  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
-  assumes f: "\<And>i j x. i \<le> j \<Longrightarrow> x \<in> space M \<Longrightarrow> f j x \<le> f i x"
-    and [measurable]: "\<And>i. f i \<in> borel_measurable M"
-    and fin: "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>"
-  shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
-proof -
-  { fix f :: "nat \<Rightarrow> ereal" and j assume "decseq f"
-    then have "(INF i. f i) = (INF i. f (i + j))"
-      apply (intro INF_eq)
-      apply (rule_tac x="i" in bexI)
-      apply (auto simp: decseq_def le_fun_def)
-      done }
-  note INF_shift = this
-
-  have dec: "decseq (\<lambda>j x. max 0 (restrict (f (j + i)) (space M) x))"
-    using f by (intro antimonoI le_funI max.mono) (auto simp: decseq_def le_fun_def)
-
-  have "(\<integral>\<^sup>+ x. max 0 (INF i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (INF i. max 0 (restrict (f i) (space M) x)) \<partial>M)"
-    by (intro nn_integral_cong)
-       (simp add: sup_ereal_def[symmetric] sup_INF del: sup_ereal_def)
-  also have "\<dots> = (\<integral>\<^sup>+ x. (INF j. max 0 (restrict (f (j + i)) (space M) x)) \<partial>M)"
-    using f by (intro nn_integral_cong INF_shift antimonoI le_funI max.mono)
-               (auto simp: decseq_def le_fun_def)
-  also have "\<dots> = (INF j. (\<integral>\<^sup>+ x. max 0 (restrict (f (j + i)) (space M) x) \<partial>M))"
-  proof (rule nn_integral_monotone_convergence_INF')
-    show "(\<lambda>x. max 0 (restrict (f (j + i)) (space M) x)) \<in> borel_measurable M" for j
-      by (subst measurable_cong[where g="\<lambda>x. max 0 (f (j + i) x)"]) measurable
-    show "(\<integral>\<^sup>+ x. max 0 (restrict (f (0 + i)) (space M) x) \<partial>M) < \<infinity>"
-      using fin by (simp add: nn_integral_max_0 cong: nn_integral_cong)
-  qed (intro max.cobounded1 dec f)+
-  also have "\<dots> = (INF j. (\<integral>\<^sup>+ x. max 0 (restrict (f j) (space M) x) \<partial>M))"
-    using f by (intro INF_shift[symmetric] nn_integral_mono antimonoI le_funI max.mono)
-               (auto simp: decseq_def le_fun_def)
-  finally show ?thesis unfolding nn_integral_max_0 by (simp cong: nn_integral_cong)
-qed
-
-lemma nn_integral_monotone_convergence_INF_decseq:
-  assumes f: "decseq f" and *: "\<And>i. f i \<in> borel_measurable M" "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>"
-  shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
-  using nn_integral_monotone_convergence_INF[of M f i, OF _ *] f by (auto simp: antimono_def le_fun_def)
-
-lemma sup_continuous_nn_integral[order_continuous_intros]:
-  assumes f: "\<And>y. sup_continuous (f y)"
-  assumes [measurable]: "\<And>x. (\<lambda>y. f y x) \<in> borel_measurable M"
-  shows "sup_continuous (\<lambda>x. (\<integral>\<^sup>+y. f y x \<partial>M))"
-  unfolding sup_continuous_def
-proof safe
-  fix C :: "nat \<Rightarrow> 'b" assume C: "incseq C"
-  with sup_continuous_mono[OF f] show "(\<integral>\<^sup>+ y. f y (SUPREMUM UNIV C) \<partial>M) = (SUP i. \<integral>\<^sup>+ y. f y (C i) \<partial>M)"
-    unfolding sup_continuousD[OF f C]
-    by (subst nn_integral_monotone_convergence_SUP) (auto simp: mono_def le_fun_def)
-qed
-
 lemma inf_continuous_nn_integral[order_continuous_intros]:
   assumes f: "\<And>y. inf_continuous (f y)"
   assumes [measurable]: "\<And>x. (\<lambda>y. f y x) \<in> borel_measurable M"
@@ -1588,9 +1317,9 @@
 proof safe
   fix C :: "nat \<Rightarrow> 'b" assume C: "decseq C"
   then show "(\<integral>\<^sup>+ y. f y (INFIMUM UNIV C) \<partial>M) = (INF i. \<integral>\<^sup>+ y. f y (C i) \<partial>M)"
-    using inf_continuous_mono[OF f]
-    by (auto simp add: inf_continuousD[OF f C] fun_eq_iff antimono_def mono_def le_fun_def bnd
-             intro!:  nn_integral_monotone_convergence_INF)
+    using inf_continuous_mono[OF f] bnd
+    by (auto simp add: inf_continuousD[OF f C] fun_eq_iff antimono_def mono_def le_fun_def less_top
+             intro!: nn_integral_monotone_convergence_INF_decseq)
 qed
 
 lemma nn_integral_null_set:
@@ -1607,7 +1336,7 @@
 qed
 
 lemma nn_integral_0_iff:
-  assumes u: "u \<in> borel_measurable M" and pos: "AE x in M. 0 \<le> u x"
+  assumes u: "u \<in> borel_measurable M"
   shows "integral\<^sup>N M u = 0 \<longleftrightarrow> emeasure M {x\<in>space M. u x \<noteq> 0} = 0"
     (is "_ \<longleftrightarrow> (emeasure M) ?A = 0")
 proof -
@@ -1619,17 +1348,14 @@
     with nn_integral_null_set[of ?A M u] u
     show "integral\<^sup>N M u = 0" by (simp add: u_eq null_sets_def)
   next
-    { fix r :: ereal and n :: nat assume gt_1: "1 \<le> real n * r"
-      then have "0 < real n * r" by (cases r) (auto split: if_split_asm simp: one_ereal_def)
-      then have "0 \<le> r" by (auto simp add: ereal_zero_less_0_iff) }
-    note gt_1 = this
     assume *: "integral\<^sup>N M u = 0"
     let ?M = "\<lambda>n. {x \<in> space M. 1 \<le> real (n::nat) * u x}"
     have "0 = (SUP n. (emeasure M) (?M n \<inter> ?A))"
     proof -
       { fix n :: nat
-        from nn_integral_Markov_inequality[OF u pos, of ?A "ereal (real n)"]
-        have "(emeasure M) (?M n \<inter> ?A) \<le> 0" unfolding u_eq * using u by simp
+        from nn_integral_Markov_inequality[OF u, of ?A "of_nat n"] u
+        have "(emeasure M) (?M n \<inter> ?A) \<le> 0"
+          by (simp add: ennreal_of_nat_eq_real_of_nat u_eq *)
         moreover have "0 \<le> (emeasure M) (?M n \<inter> ?A)" using u by auto
         ultimately have "(emeasure M) (?M n \<inter> ?A) = 0" by auto }
       thus ?thesis by simp
@@ -1643,88 +1369,75 @@
       proof (safe intro!: incseq_SucI)
         fix n :: nat and x
         assume *: "1 \<le> real n * u x"
-        also from gt_1[OF *] have "real n * u x \<le> real (Suc n) * u x"
-          using \<open>0 \<le> u x\<close> by (auto intro!: ereal_mult_right_mono)
+        also have "real n * u x \<le> real (Suc n) * u x"
+          by (auto intro!: mult_right_mono)
         finally show "1 \<le> real (Suc n) * u x" by auto
       qed
     qed
     also have "\<dots> = (emeasure M) {x\<in>space M. 0 < u x}"
-    proof (safe intro!: arg_cong[where f="(emeasure M)"] dest!: gt_1)
+    proof (safe intro!: arg_cong[where f="(emeasure M)"])
       fix x assume "0 < u x" and [simp, intro]: "x \<in> space M"
       show "x \<in> (\<Union>n. ?M n \<inter> ?A)"
-      proof (cases "u x")
+      proof (cases "u x" rule: ennreal_cases)
         case (real r) with \<open>0 < u x\<close> have "0 < r" by auto
         obtain j :: nat where "1 / r \<le> real j" using real_arch_simple ..
         hence "1 / r * r \<le> real j * r" unfolding mult_le_cancel_right using \<open>0 < r\<close> by auto
         hence "1 \<le> real j * r" using real \<open>0 < r\<close> by auto
-        thus ?thesis using \<open>0 < r\<close> real by (auto simp: one_ereal_def)
-      qed (insert \<open>0 < u x\<close>, auto)
-    qed auto
-    finally have "(emeasure M) {x\<in>space M. 0 < u x} = 0" by simp
-    moreover
-    from pos have "AE x in M. \<not> (u x < 0)" by auto
-    then have "(emeasure M) {x\<in>space M. u x < 0} = 0"
-      using AE_iff_null[of M] u by auto
-    moreover have "(emeasure M) {x\<in>space M. u x \<noteq> 0} = (emeasure M) {x\<in>space M. u x < 0} + (emeasure M) {x\<in>space M. 0 < u x}"
-      using u by (subst plus_emeasure) (auto intro!: arg_cong[where f="emeasure M"])
-    ultimately show "(emeasure M) ?A = 0" by simp
+        thus ?thesis using \<open>0 < r\<close> real
+          by (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_1[symmetric] ennreal_mult[symmetric]
+                   simp del: ennreal_1)
+      qed (insert \<open>0 < u x\<close>, auto simp: ennreal_mult_top)
+    qed (auto simp: zero_less_iff_neq_zero)
+    finally show "emeasure M ?A = 0"
+      by (simp add: zero_less_iff_neq_zero)
   qed
 qed
 
 lemma nn_integral_0_iff_AE:
   assumes u: "u \<in> borel_measurable M"
-  shows "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. u x \<le> 0)"
+  shows "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. u x = 0)"
 proof -
-  have sets: "{x\<in>space M. max 0 (u x) \<noteq> 0} \<in> sets M"
+  have sets: "{x\<in>space M. u x \<noteq> 0} \<in> sets M"
     using u by auto
-  from nn_integral_0_iff[of "\<lambda>x. max 0 (u x)"]
-  have "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. max 0 (u x) = 0)"
-    unfolding nn_integral_max_0
-    using AE_iff_null[OF sets] u by auto
-  also have "\<dots> \<longleftrightarrow> (AE x in M. u x \<le> 0)" by (auto split: split_max)
-  finally show ?thesis .
+  show "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. u x = 0)"
+    using nn_integral_0_iff[of u] AE_iff_null[OF sets] u by auto
 qed
 
 lemma AE_iff_nn_integral:
   "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> integral\<^sup>N M (indicator {x. \<not> P x}) = 0"
-  by (subst nn_integral_0_iff_AE) (auto simp: one_ereal_def zero_ereal_def
-    sets.sets_Collect_neg indicator_def[abs_def] measurable_If)
+  by (subst nn_integral_0_iff_AE) (auto simp: indicator_def[abs_def])
 
 lemma nn_integral_less:
   assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  assumes f: "AE x in M. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>M) \<noteq> \<infinity>"
+  assumes f: "(\<integral>\<^sup>+x. f x \<partial>M) \<noteq> \<infinity>"
   assumes ord: "AE x in M. f x \<le> g x" "\<not> (AE x in M. g x \<le> f x)"
   shows "(\<integral>\<^sup>+x. f x \<partial>M) < (\<integral>\<^sup>+x. g x \<partial>M)"
 proof -
   have "0 < (\<integral>\<^sup>+x. g x - f x \<partial>M)"
-  proof (intro order_le_neq_trans nn_integral_nonneg notI)
+  proof (intro order_le_neq_trans notI)
     assume "0 = (\<integral>\<^sup>+x. g x - f x \<partial>M)"
-    then have "AE x in M. g x - f x \<le> 0"
+    then have "AE x in M. g x - f x = 0"
       using nn_integral_0_iff_AE[of "\<lambda>x. g x - f x" M] by simp
-    with f(1) ord(1) have "AE x in M. g x \<le> f x"
-      by eventually_elim (auto simp: ereal_minus_le_iff)
+    with ord(1) have "AE x in M. g x \<le> f x"
+      by eventually_elim (auto simp: ennreal_minus_eq_0)
     with ord show False
       by simp
-  qed
+  qed simp
   also have "\<dots> = (\<integral>\<^sup>+x. g x \<partial>M) - (\<integral>\<^sup>+x. f x \<partial>M)"
-    by (subst nn_integral_diff) (auto simp: f ord)
+    using f by (subst nn_integral_diff) (auto simp: ord)
   finally show ?thesis
-    by (simp add: ereal_less_minus_iff f nn_integral_nonneg)
+    using f by (auto dest!: ennreal_minus_pos_iff[rotated] simp: less_top)
 qed
 
-lemma nn_integral_const_If:
-  "(\<integral>\<^sup>+x. a \<partial>M) = (if 0 \<le> a then a * (emeasure M) (space M) else 0)"
-  by (auto intro!: nn_integral_0_iff_AE[THEN iffD2])
-
 lemma nn_integral_subalgebra:
-  assumes f: "f \<in> borel_measurable N" "\<And>x. 0 \<le> f x"
+  assumes f: "f \<in> borel_measurable N"
   and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> emeasure N A = emeasure M A"
   shows "integral\<^sup>N N f = integral\<^sup>N M f"
 proof -
-  have [simp]: "\<And>f :: 'a \<Rightarrow> ereal. f \<in> borel_measurable N \<Longrightarrow> f \<in> borel_measurable M"
+  have [simp]: "\<And>f :: 'a \<Rightarrow> ennreal. f \<in> borel_measurable N \<Longrightarrow> f \<in> borel_measurable M"
     using N by (auto simp: measurable_def)
   have [simp]: "\<And>P. (AE x in N. P x) \<Longrightarrow> (AE x in M. P x)"
-    using N by (auto simp add: eventually_ae_filter null_sets_def)
+    using N by (auto simp add: eventually_ae_filter null_sets_def subset_eq)
   have [simp]: "\<And>A. A \<in> sets N \<Longrightarrow> A \<in> sets M"
     using N by auto
   from f show ?thesis
@@ -1737,7 +1450,7 @@
 lemma nn_integral_nat_function:
   fixes f :: "'a \<Rightarrow> nat"
   assumes "f \<in> measurable M (count_space UNIV)"
-  shows "(\<integral>\<^sup>+x. ereal (of_nat (f x)) \<partial>M) = (\<Sum>t. emeasure M {x\<in>space M. t < f x})"
+  shows "(\<integral>\<^sup>+x. of_nat (f x) \<partial>M) = (\<Sum>t. emeasure M {x\<in>space M. t < f x})"
 proof -
   def F \<equiv> "\<lambda>i. {x\<in>space M. i < f x}"
   with assms have [measurable]: "\<And>i. F i \<in> sets M"
@@ -1746,13 +1459,14 @@
   { fix x assume "x \<in> space M"
     have "(\<lambda>i. if i < f x then 1 else 0) sums (of_nat (f x)::real)"
       using sums_If_finite[of "\<lambda>i. i < f x" "\<lambda>_. 1::real"] by simp
-    then have "(\<lambda>i. ereal(if i < f x then 1 else 0)) sums (ereal(of_nat(f x)))"
-      unfolding sums_ereal .
-    moreover have "\<And>i. ereal (if i < f x then 1 else 0) = indicator (F i) x"
-      using \<open>x \<in> space M\<close> by (simp add: one_ereal_def F_def)
-    ultimately have "ereal(of_nat(f x)) = (\<Sum>i. indicator (F i) x)"
+    then have "(\<lambda>i. ennreal (if i < f x then 1 else 0)) sums of_nat(f x)"
+      unfolding ennreal_of_nat_eq_real_of_nat
+      by (subst sums_ennreal) auto
+    moreover have "\<And>i. ennreal (if i < f x then 1 else 0) = indicator (F i) x"
+      using \<open>x \<in> space M\<close> by (simp add: one_ennreal_def F_def)
+    ultimately have "of_nat (f x) = (\<Sum>i. indicator (F i) x :: ennreal)"
       by (simp add: sums_iff) }
-  then have "(\<integral>\<^sup>+x. ereal (of_nat (f x)) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)"
+  then have "(\<integral>\<^sup>+x. of_nat (f x) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)"
     by (simp cong: nn_integral_cong)
   also have "\<dots> = (\<Sum>i. emeasure M (F i))"
     by (simp add: nn_integral_suminf)
@@ -1764,21 +1478,16 @@
   assumes sets[simp]: "\<And>s. sets (M s) = sets N"
   assumes f: "sup_continuous f"
   assumes g: "sup_continuous g"
-  assumes nonneg: "\<And>F s. 0 \<le> g F s"
   assumes meas: "\<And>F. F \<in> borel_measurable N \<Longrightarrow> f F \<in> borel_measurable N"
   assumes step: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> integral\<^sup>N (M s) (f F) = g (\<lambda>s. integral\<^sup>N (M s) F) s"
   shows "(\<integral>\<^sup>+\<omega>. lfp f \<omega> \<partial>M s) = lfp g s"
 proof (subst lfp_transfer_bounded[where \<alpha>="\<lambda>F s. \<integral>\<^sup>+x. F x \<partial>M s" and g=g and f=f and P="\<lambda>f. f \<in> borel_measurable N", symmetric])
-  fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ereal" assume "incseq C" "\<And>i. C i \<in> borel_measurable N"
+  fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ennreal" assume "incseq C" "\<And>i. C i \<in> borel_measurable N"
   then show "(\<lambda>s. \<integral>\<^sup>+x. (SUP i. C i) x \<partial>M s) = (SUP i. (\<lambda>s. \<integral>\<^sup>+x. C i x \<partial>M s))"
     unfolding SUP_apply[abs_def]
     by (subst nn_integral_monotone_convergence_SUP)
        (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets)
-next
-  show "\<And>x. (\<lambda>s. integral\<^sup>N (M s) bot) \<le> g x"
-    by (subst nn_integral_max_0[symmetric])
-       (simp add: sup_ereal_def[symmetric] le_fun_def nonneg del: sup_ereal_def)
-qed (auto simp add: step nonneg le_fun_def SUP_apply[abs_def] bot_fun_def intro!: meas f g)
+qed (auto simp add: step le_fun_def SUP_apply[abs_def] bot_fun_def bot_ennreal intro!: meas f g)
 
 lemma nn_integral_gfp:
   assumes sets[simp]: "\<And>s. sets (M s) = sets N"
@@ -1790,7 +1499,7 @@
   shows "(\<integral>\<^sup>+\<omega>. gfp f \<omega> \<partial>M s) = gfp g s"
 proof (subst gfp_transfer_bounded[where \<alpha>="\<lambda>F s. \<integral>\<^sup>+x. F x \<partial>M s" and g=g and f=f
     and P="\<lambda>F. F \<in> borel_measurable N \<and> (\<forall>s. (\<integral>\<^sup>+x. F x \<partial>M s) < \<infinity>)", symmetric])
-  fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ereal" assume "decseq C" "\<And>i. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)"
+  fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ennreal" assume "decseq C" "\<And>i. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)"
   then show "(\<lambda>s. \<integral>\<^sup>+x. (INF i. C i) x \<partial>M s) = (INF i. (\<lambda>s. \<integral>\<^sup>+x. C i x \<partial>M s))"
     unfolding INF_apply[abs_def]
     by (subst nn_integral_monotone_convergence_INF_decseq)
@@ -1798,15 +1507,14 @@
 next
   show "\<And>x. g x \<le> (\<lambda>s. integral\<^sup>N (M s) (f top))"
     by (subst step)
-       (auto simp add: top_fun_def top_ereal_def less_le emeasure_nonneg non_zero le_fun_def
+       (auto simp add: top_fun_def less_le non_zero le_fun_def ennreal_top_mult
              cong del: if_cong intro!: monoD[OF inf_continuous_mono[OF g], THEN le_funD])
 next
   fix C assume "\<And>i::nat. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)" "decseq C"
   with bound show "INFIMUM UNIV C \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (INFIMUM UNIV C) < \<infinity>)"
     unfolding INF_apply[abs_def]
     by (subst nn_integral_monotone_convergence_INF_decseq)
-       (auto cong: measurable_cong_sets intro!: borel_measurable_INF
-             simp: INF_less_iff simp del: ereal_infty_less(1))
+       (auto simp: INF_less_iff cong: measurable_cong_sets intro!: borel_measurable_INF)
 next
   show "\<And>x. x \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) x < \<infinity>) \<Longrightarrow>
          (\<lambda>s. integral\<^sup>N (M s) (f x)) = g (\<lambda>s. integral\<^sup>N (M s) x)"
@@ -1826,9 +1534,8 @@
 
 subsubsection \<open>Distributions\<close>
 
-lemma nn_integral_distr':
-  assumes T: "T \<in> measurable M M'"
-  and f: "f \<in> borel_measurable (distr M M' T)" "\<And>x. 0 \<le> f x"
+lemma nn_integral_distr:
+  assumes T: "T \<in> measurable M M'" and f: "f \<in> borel_measurable (distr M M' T)"
   shows "integral\<^sup>N (distr M M' T) f = (\<integral>\<^sup>+ x. f (T x) \<partial>M)"
   using f
 proof induct
@@ -1850,11 +1557,6 @@
 qed (simp_all add: measurable_compose[OF T] T nn_integral_cmult nn_integral_add
                    nn_integral_monotone_convergence_SUP le_fun_def incseq_def)
 
-lemma nn_integral_distr:
-  "T \<in> measurable M M' \<Longrightarrow> f \<in> borel_measurable M' \<Longrightarrow> integral\<^sup>N (distr M M' T) f = (\<integral>\<^sup>+ x. f (T x) \<partial>M)"
-  by (subst (1 2) nn_integral_max_0[symmetric])
-     (simp add: nn_integral_distr')
-
 subsubsection \<open>Counting space\<close>
 
 lemma simple_function_count_space[simp]:
@@ -1870,20 +1572,18 @@
     by (auto intro!: nn_integral_cong
              simp add: indicator_def if_distrib setsum.If_cases[OF A] max_def le_less)
   also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. \<integral>\<^sup>+ x. f a * indicator {a} x \<partial>count_space A)"
-    by (subst nn_integral_setsum)
-       (simp_all add: AE_count_space ereal_zero_le_0_iff less_imp_le)
+    by (subst nn_integral_setsum) (simp_all add: AE_count_space  less_imp_le)
   also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)"
-    by (auto intro!: setsum.cong simp: one_ereal_def[symmetric] max_def)
-  finally show ?thesis by (simp add: nn_integral_max_0)
+    by (auto intro!: setsum.cong simp: one_ennreal_def[symmetric] max_def)
+  finally show ?thesis by (simp add: max.absorb2)
 qed
 
 lemma nn_integral_count_space_finite:
-    "finite A \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<Sum>a\<in>A. max 0 (f a))"
-  by (subst nn_integral_max_0[symmetric])
-     (auto intro!: setsum.mono_neutral_left simp: nn_integral_count_space less_le)
+    "finite A \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"
+  by (auto intro!: setsum.mono_neutral_left simp: nn_integral_count_space less_le)
 
 lemma nn_integral_count_space':
-  assumes "finite A" "\<And>x. x \<in> B \<Longrightarrow> x \<notin> A \<Longrightarrow> f x = 0" "\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x" "A \<subseteq> B"
+  assumes "finite A" "\<And>x. x \<in> B \<Longrightarrow> x \<notin> A \<Longrightarrow> f x = 0" "A \<subseteq> B"
   shows "(\<integral>\<^sup>+x. f x \<partial>count_space B) = (\<Sum>x\<in>A. f x)"
 proof -
   have "(\<integral>\<^sup>+x. f x \<partial>count_space B) = (\<Sum>a | a \<in> B \<and> 0 < f a. f a)"
@@ -1902,20 +1602,19 @@
      (auto intro!: nn_integral_distr[symmetric])
 
 lemma nn_integral_indicator_finite:
-  fixes f :: "'a \<Rightarrow> ereal"
-  assumes f: "finite A" and nn: "\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x" and [measurable]: "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> sets M"
+  fixes f :: "'a \<Rightarrow> ennreal"
+  assumes f: "finite A" and [measurable]: "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> sets M"
   shows "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<Sum>x\<in>A. f x * emeasure M {x})"
 proof -
   from f have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>a\<in>A. f a * indicator {a} x) \<partial>M)"
     by (intro nn_integral_cong) (auto simp: indicator_def if_distrib[where f="\<lambda>a. x * a" for x] setsum.If_cases)
   also have "\<dots> = (\<Sum>a\<in>A. f a * emeasure M {a})"
-    using nn by (subst nn_integral_setsum) (auto simp: max_def)
+    by (subst nn_integral_setsum) auto
   finally show ?thesis .
 qed
 
 lemma nn_integral_count_space_nat:
-  fixes f :: "nat \<Rightarrow> ereal"
-  assumes nonneg: "\<And>i. 0 \<le> f i"
+  fixes f :: "nat \<Rightarrow> ennreal"
   shows "(\<integral>\<^sup>+i. f i \<partial>count_space UNIV) = (\<Sum>i. f i)"
 proof -
   have "(\<integral>\<^sup>+i. f i \<partial>count_space UNIV) =
@@ -1929,20 +1628,42 @@
     finally show "f i = (\<Sum>j. f j * indicator {j} i)" .
   qed
   also have "\<dots> = (\<Sum>j. (\<integral>\<^sup>+i. f j * indicator {j} i \<partial>count_space UNIV))"
-    by (rule nn_integral_suminf) (auto simp: nonneg)
-  also have "\<dots> = (\<Sum>j. f j)"
-    by (simp add: nonneg one_ereal_def[symmetric] max_def)
-  finally show ?thesis .
+    by (rule nn_integral_suminf) auto
+  finally show ?thesis
+    by simp
+qed
+
+lemma nn_integral_enat_function:
+  assumes f: "f \<in> measurable M (count_space UNIV)"
+  shows "(\<integral>\<^sup>+ x. ennreal_of_enat (f x) \<partial>M) = (\<Sum>t. emeasure M {x \<in> space M. t < f x})"
+proof -
+  def F \<equiv> "\<lambda>i::nat. {x\<in>space M. i < f x}"
+  with assms have [measurable]: "\<And>i. F i \<in> sets M"
+    by auto
+
+  { fix x assume "x \<in> space M"
+    have "(\<lambda>i::nat. if i < f x then 1 else 0) sums ennreal_of_enat (f x)"
+      using sums_If_finite[of "\<lambda>r. r < f x" "\<lambda>_. 1 :: ennreal"]
+      by (cases "f x") (simp_all add: sums_def of_nat_tendsto_top_ennreal)
+    also have "(\<lambda>i. (if i < f x then 1 else 0)) = (\<lambda>i. indicator (F i) x)"
+      using `x \<in> space M` by (simp add: one_ennreal_def F_def fun_eq_iff)
+    finally have "ennreal_of_enat (f x) = (\<Sum>i. indicator (F i) x)"
+      by (simp add: sums_iff) }
+  then have "(\<integral>\<^sup>+x. ennreal_of_enat (f x) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)"
+    by (simp cong: nn_integral_cong)
+  also have "\<dots> = (\<Sum>i. emeasure M (F i))"
+    by (simp add: nn_integral_suminf)
+  finally show ?thesis
+    by (simp add: F_def)
 qed
 
 lemma nn_integral_count_space_nn_integral:
-  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ereal"
+  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ennreal"
   assumes "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<integral>\<^sup>+x. \<integral>\<^sup>+i. f i x \<partial>count_space I \<partial>M) = (\<integral>\<^sup>+i. \<integral>\<^sup>+x. f i x \<partial>M \<partial>count_space I)"
 proof cases
   assume "finite I" then show ?thesis
-    by (simp add: nn_integral_count_space_finite nn_integral_nonneg max.absorb2 nn_integral_setsum
-                  nn_integral_max_0)
+    by (simp add: nn_integral_count_space_finite nn_integral_setsum)
 next
   assume "infinite I"
   then have [simp]: "I \<noteq> {}"
@@ -1951,10 +1672,7 @@
   have **: "\<And>f. (\<And>i. 0 \<le> f i) \<Longrightarrow> (\<integral>\<^sup>+i. f i \<partial>count_space I) = (\<Sum>n. f (from_nat_into I n))"
     by (simp add: nn_integral_bij_count_space[symmetric, OF *] nn_integral_count_space_nat)
   show ?thesis
-    apply (subst (2) nn_integral_max_0[symmetric])
-    apply (simp add: ** nn_integral_nonneg nn_integral_suminf from_nat_into)
-    apply (simp add: nn_integral_max_0)
-    done
+    by (simp add: ** nn_integral_suminf from_nat_into)
 qed
 
 lemma emeasure_UN_countable:
@@ -1967,7 +1685,7 @@
     fix x assume x: "x \<in> UNION I X"
     then obtain j where j: "x \<in> X j" "j \<in> I"
       by auto
-    with disj have "\<And>i. i \<in> I \<Longrightarrow> indicator (X i) x = (indicator {j} i::ereal)"
+    with disj have "\<And>i. i \<in> I \<Longrightarrow> indicator (X i) x = (indicator {j} i::ennreal)"
       by (auto simp: disjoint_family_on_def split: split_indicator)
     with x j show "?thesis x"
       by (simp cong: nn_integral_cong_simp)
@@ -2030,19 +1748,18 @@
   finally show "emeasure M A = emeasure N A" ..
 qed simp
 
-lemma nn_integral_monotone_convergence_SUP_nat':
-  fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ereal"
+lemma nn_integral_monotone_convergence_SUP_nat:
+  fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ennreal"
   assumes chain: "Complete_Partial_Order.chain op \<le> (f ` Y)"
   and nonempty: "Y \<noteq> {}"
-  and nonneg: "\<And>i n. i \<in> Y \<Longrightarrow> f i n \<ge> 0"
   shows "(\<integral>\<^sup>+ x. (SUP i:Y. f i x) \<partial>count_space UNIV) = (SUP i:Y. (\<integral>\<^sup>+ x. f i x \<partial>count_space UNIV))"
   (is "?lhs = ?rhs" is "integral\<^sup>N ?M _ = _")
 proof (rule order_class.order.antisym)
   show "?rhs \<le> ?lhs"
     by (auto intro!: SUP_least SUP_upper nn_integral_mono)
 next
-  have "\<And>x. \<exists>g. incseq g \<and> range g \<subseteq> (\<lambda>i. f i x) ` Y \<and> (SUP i:Y. f i x) = (SUP i. g i)"
-    by (rule Sup_countable_SUP) (simp add: nonempty)
+  have "\<exists>g. incseq g \<and> range g \<subseteq> (\<lambda>i. f i x) ` Y \<and> (SUP i:Y. f i x) = (SUP i. g i)" for x
+    by (rule ennreal_Sup_countable_SUP) (simp add: nonempty)
   then obtain g where incseq: "\<And>x. incseq (g x)"
     and range: "\<And>x. range (g x) \<subseteq> (\<lambda>i. f i x) ` Y"
     and sup: "\<And>x. (SUP i:Y. f i x) = (SUP i. g x i)" by moura
@@ -2058,21 +1775,17 @@
     have "\<And>x. \<exists>i. g x n = f i x \<and> i \<in> Y" using range by blast
     then obtain I where I: "\<And>x. g x n = f (I x) x" "\<And>x. I x \<in> Y" by moura
 
-    { fix x
-      from range[of x] obtain i where "i \<in> Y" "g x n = f i x" by auto
-      hence "g x n \<ge> 0" using nonneg[of i x] by simp }
-    note nonneg_g = this
-    then have "(\<integral>\<^sup>+ x. g x n \<partial>count_space UNIV) = (\<Sum>x. g x n)"
+    have "(\<integral>\<^sup>+ x. g x n \<partial>count_space UNIV) = (\<Sum>x. g x n)"
       by(rule nn_integral_count_space_nat)
-    also have "\<dots> = (SUP m. \<Sum>x<m. g x n)" using nonneg_g
-      by(rule suminf_ereal_eq_SUP)
+    also have "\<dots> = (SUP m. \<Sum>x<m. g x n)"
+      by(rule suminf_eq_SUP)
     also have "\<dots> \<le> (SUP i:Y. \<integral>\<^sup>+ x. f i x \<partial>?M)"
     proof(rule SUP_mono)
       fix m
       show "\<exists>m'\<in>Y. (\<Sum>x<m. g x n) \<le> (\<integral>\<^sup>+ x. f m' x \<partial>?M)"
       proof(cases "m > 0")
         case False
-        thus ?thesis using nonempty by(auto simp add: nn_integral_nonneg)
+        thus ?thesis using nonempty by auto
       next
         case True
         let ?Y = "I ` {..<m}"
@@ -2096,9 +1809,9 @@
         also have "\<dots> \<le> (SUP m. (\<Sum>x<m. f (I m') x))"
           by(rule SUP_upper) simp
         also have "\<dots> = (\<Sum>x. f (I m') x)"
-          by(rule suminf_ereal_eq_SUP[symmetric])(simp add: nonneg \<open>I m' \<in> Y\<close>)
+          by(rule suminf_eq_SUP[symmetric])
         also have "\<dots> = (\<integral>\<^sup>+ x. f (I m') x \<partial>?M)"
-          by(rule nn_integral_count_space_nat[symmetric])(simp add: nonneg \<open>I m' \<in> Y\<close>)
+          by(rule nn_integral_count_space_nat[symmetric])
         finally show ?thesis using \<open>I m' \<in> Y\<close> by blast
       qed
     qed
@@ -2107,75 +1820,37 @@
   finally show "?lhs \<le> ?rhs" .
 qed
 
-lemma nn_integral_monotone_convergence_SUP_nat:
-  fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ereal"
-  assumes nonempty: "Y \<noteq> {}"
-  and chain: "Complete_Partial_Order.chain op \<le> (f ` Y)"
-  shows "(\<integral>\<^sup>+ x. (SUP i:Y. f i x) \<partial>count_space UNIV) = (SUP i:Y. (\<integral>\<^sup>+ x. f i x \<partial>count_space UNIV))"
-  (is "?lhs = ?rhs")
-proof -
-  let ?f = "\<lambda>i x. max 0 (f i x)"
-  have chain': "Complete_Partial_Order.chain op \<le> (?f ` Y)"
-  proof(rule chainI)
-    fix g h
-    assume "g \<in> ?f ` Y" "h \<in> ?f ` Y"
-    then obtain g' h' where gh: "g' \<in> Y" "h' \<in> Y" "g = ?f g'" "h = ?f h'" by blast
-    hence "f g' \<in> f ` Y" "f h' \<in> f ` Y" by blast+
-    with chain have "f g' \<le> f h' \<or> f h' \<le> f g'" by(rule chainD)
-    thus "g \<le> h \<or> h \<le> g"
-    proof
-      assume "f g' \<le> f h'"
-      hence "g \<le> h" using gh order_trans by(auto simp add: le_fun_def max_def)
-      thus ?thesis ..
-    next
-      assume "f h' \<le> f g'"
-      hence "h \<le> g" using gh order_trans by(auto simp add: le_fun_def max_def)
-      thus ?thesis ..
-    qed
-  qed
-  have "?lhs = (\<integral>\<^sup>+ x. max 0 (SUP i:Y. f i x) \<partial>count_space UNIV)"
-    by(simp add: nn_integral_max_0)
-  also have "\<dots> = (\<integral>\<^sup>+ x. (SUP i:Y. ?f i x) \<partial>count_space UNIV)"
-  proof(rule nn_integral_cong)
-    fix x
-    have "max 0 (SUP i:Y. f i x) \<le> (SUP i:Y. ?f i x)"
-    proof(cases "0 \<le> (SUP i:Y. f i x)")
-      case True
-      have "(SUP i:Y. f i x) \<le> (SUP i:Y. ?f i x)" by(rule SUP_mono)(auto intro: rev_bexI)
-      with True show ?thesis by simp
-    next
-      case False
-      have "0 \<le> (SUP i:Y. ?f i x)" using nonempty by(auto intro: SUP_upper2)
-      thus ?thesis using False by simp
-    qed
-    moreover have "\<dots> \<le> max 0 (SUP i:Y. f i x)"
-    proof(cases "(SUP i:Y. f i x) \<ge> 0")
-      case True
-      show ?thesis
-        by(rule SUP_least)(auto simp add: True max_def intro: SUP_upper)
-    next
-      case False
-      hence "(SUP i:Y. f i x) \<le> 0" by simp
-      hence less: "\<forall>i\<in>Y. f i x \<le> 0" by(simp add: SUP_le_iff)
-      show ?thesis by(rule SUP_least)(auto simp add: max_def less intro: SUP_upper)
-    qed
-    ultimately show "\<dots> = (SUP i:Y. ?f i x)" by(rule order.antisym)
-  qed
-  also have "\<dots> = (SUP i:Y. (\<integral>\<^sup>+ x. ?f i x \<partial>count_space UNIV))"
-    using chain' nonempty by(rule nn_integral_monotone_convergence_SUP_nat') simp
-  also have "\<dots> = ?rhs" by(simp add: nn_integral_max_0)
-  finally show ?thesis .
+lemma power_series_tendsto_at_left:
+  assumes nonneg: "\<And>i. 0 \<le> f i" and summable: "\<And>z. 0 \<le> z \<Longrightarrow> z < 1 \<Longrightarrow> summable (\<lambda>n. f n * z^n)"
+  shows "((\<lambda>z. ennreal (\<Sum>n. f n * z^n)) \<longlongrightarrow> (\<Sum>n. ennreal (f n))) (at_left (1::real))"
+proof (intro tendsto_at_left_sequentially)
+  show "0 < (1::real)" by simp
+  fix S :: "nat \<Rightarrow> real" assume S: "\<And>n. S n < 1" "\<And>n. 0 < S n" "S \<longlonglongrightarrow> 1" "incseq S"
+  then have S_nonneg: "\<And>i. 0 \<le> S i" by (auto intro: less_imp_le)
+
+  have "(\<lambda>i. (\<integral>\<^sup>+n. f n * S i^n \<partial>count_space UNIV)) \<longlonglongrightarrow> (\<integral>\<^sup>+n. ennreal (f n) \<partial>count_space UNIV)"
+  proof (rule nn_integral_LIMSEQ)
+    show "incseq (\<lambda>i n. ennreal (f n * S i^n))"
+      using S by (auto intro!: mult_mono power_mono nonneg ennreal_leI
+                       simp: incseq_def le_fun_def less_imp_le)
+    fix n have "(\<lambda>i. ennreal (f n * S i^n)) \<longlonglongrightarrow> ennreal (f n * 1^n)"
+      by (intro tendsto_intros tendsto_ennrealI S)
+    then show "(\<lambda>i. ennreal (f n * S i^n)) \<longlonglongrightarrow> ennreal (f n)"
+      by simp
+  qed (auto simp: S_nonneg intro!: mult_nonneg_nonneg nonneg)
+  also have "(\<lambda>i. (\<integral>\<^sup>+n. f n * S i^n \<partial>count_space UNIV)) = (\<lambda>i. \<Sum>n. f n * S i^n)"
+    by (subst nn_integral_count_space_nat)
+       (intro ext suminf_ennreal2 mult_nonneg_nonneg nonneg S_nonneg
+              zero_le_power summable S)+
+  also have "(\<integral>\<^sup>+n. ennreal (f n) \<partial>count_space UNIV) = (\<Sum>n. ennreal (f n))"
+    by (simp add: nn_integral_count_space_nat nonneg)
+  finally show "(\<lambda>n. ennreal (\<Sum>na. f na * S n ^ na)) \<longlonglongrightarrow> (\<Sum>n. ennreal (f n))" .
 qed
 
 subsubsection \<open>Measures with Restricted Space\<close>
 
-lemma simple_function_iff_borel_measurable:
-  fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
-  shows "simple_function M f \<longleftrightarrow> finite (f ` space M) \<and> f \<in> borel_measurable M"
-  by (metis borel_measurable_simple_function simple_functionD(1) simple_function_borel_measurable)
-
-lemma simple_function_restrict_space_ereal:
-  fixes f :: "'a \<Rightarrow> ereal"
+lemma simple_function_restrict_space_ennreal:
+  fixes f :: "'a \<Rightarrow> ennreal"
   assumes "\<Omega> \<inter> space M \<in> sets M"
   shows "simple_function (restrict_space M \<Omega>) f \<longleftrightarrow> simple_function M (\<lambda>x. f x * indicator \<Omega> x)"
 proof -
@@ -2188,8 +1863,8 @@
     then have "finite (f ` space (restrict_space M \<Omega>))"
       by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
   ultimately show ?thesis
-    unfolding simple_function_iff_borel_measurable
-      borel_measurable_restrict_space_iff_ereal[OF assms]
+    unfolding
+      simple_function_iff_borel_measurable borel_measurable_restrict_space_iff_ennreal[OF assms]
     by auto
 qed
 
@@ -2212,44 +1887,37 @@
     by auto
 qed
 
-
-lemma split_indicator_asm: "P (indicator Q x) = (\<not> (x \<in> Q \<and> \<not> P 1 \<or> x \<notin> Q \<and> \<not> P 0))"
-  by (auto split: split_indicator)
-
 lemma simple_integral_restrict_space:
   assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M" "simple_function (restrict_space M \<Omega>) f"
   shows "simple_integral (restrict_space M \<Omega>) f = simple_integral M (\<lambda>x. f x * indicator \<Omega> x)"
-  using simple_function_restrict_space_ereal[THEN iffD1, OF \<Omega>, THEN simple_functionD(1)]
+  using simple_function_restrict_space_ennreal[THEN iffD1, OF \<Omega>, THEN simple_functionD(1)]
   by (auto simp add: space_restrict_space emeasure_restrict_space[OF \<Omega>(1)] le_infI2 simple_integral_def
            split: split_indicator split_indicator_asm
-           intro!: setsum.mono_neutral_cong_left ereal_right_mult_cong[OF refl] arg_cong2[where f=emeasure])
-
-lemma one_not_less_zero[simp]: "\<not> 1 < (0::ereal)"
-  by (simp add: zero_ereal_def one_ereal_def)
+           intro!: setsum.mono_neutral_cong_left ennreal_mult_left_cong arg_cong2[where f=emeasure])
 
 lemma nn_integral_restrict_space:
   assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M"
   shows "nn_integral (restrict_space M \<Omega>) f = nn_integral M (\<lambda>x. f x * indicator \<Omega> x)"
 proof -
-  let ?R = "restrict_space M \<Omega>" and ?X = "\<lambda>M f. {s. simple_function M s \<and> s \<le> max 0 \<circ> f \<and> range s \<subseteq> {0 ..< \<infinity>}}"
+  let ?R = "restrict_space M \<Omega>" and ?X = "\<lambda>M f. {s. simple_function M s \<and> s \<le> f \<and> (\<forall>x. s x < top)}"
   have "integral\<^sup>S ?R ` ?X ?R f = integral\<^sup>S M ` ?X M (\<lambda>x. f x * indicator \<Omega> x)"
   proof (safe intro!: image_eqI)
-    fix s assume s: "simple_function ?R s" "s \<le> max 0 \<circ> f" "range s \<subseteq> {0..<\<infinity>}"
+    fix s assume s: "simple_function ?R s" "s \<le> f" "\<forall>x. s x < top"
     from s show "integral\<^sup>S (restrict_space M \<Omega>) s = integral\<^sup>S M (\<lambda>x. s x * indicator \<Omega> x)"
       by (intro simple_integral_restrict_space) auto
     from s show "simple_function M (\<lambda>x. s x * indicator \<Omega> x)"
-      by (simp add: simple_function_restrict_space_ereal)
-    from s show "(\<lambda>x. s x * indicator \<Omega> x) \<le> max 0 \<circ> (\<lambda>x. f x * indicator \<Omega> x)"
-      "\<And>x. s x * indicator \<Omega> x \<in> {0..<\<infinity>}"
+      by (simp add: simple_function_restrict_space_ennreal)
+    from s show "(\<lambda>x. s x * indicator \<Omega> x) \<le> (\<lambda>x. f x * indicator \<Omega> x)"
+      "\<And>x. s x * indicator \<Omega> x < top"
       by (auto split: split_indicator simp: le_fun_def image_subset_iff)
   next
-    fix s assume s: "simple_function M s" "s \<le> max 0 \<circ> (\<lambda>x. f x * indicator \<Omega> x)" "range s \<subseteq> {0..<\<infinity>}"
+    fix s assume s: "simple_function M s" "s \<le> (\<lambda>x. f x * indicator \<Omega> x)" "\<forall>x. s x < top"
     then have "simple_function M (\<lambda>x. s x * indicator (\<Omega> \<inter> space M) x)" (is ?s')
       by (intro simple_function_mult simple_function_indicator) auto
     also have "?s' \<longleftrightarrow> simple_function M (\<lambda>x. s x * indicator \<Omega> x)"
       by (rule simple_function_cong) (auto split: split_indicator)
     finally show sf: "simple_function (restrict_space M \<Omega>) s"
-      by (simp add: simple_function_restrict_space_ereal)
+      by (simp add: simple_function_restrict_space_ennreal)
 
     from s have s_eq: "s = (\<lambda>x. s x * indicator \<Omega> x)"
       by (auto simp add: fun_eq_iff le_fun_def image_subset_iff
@@ -2258,9 +1926,9 @@
 
     show "integral\<^sup>S M s = integral\<^sup>S (restrict_space M \<Omega>) s"
       by (subst s_eq) (rule simple_integral_restrict_space[symmetric, OF \<Omega> sf])
-    show "\<And>x. s x \<in> {0..<\<infinity>}"
+    show "\<And>x. s x < top"
       using s by (auto simp: image_subset_iff)
-    from s show "s \<le> max 0 \<circ> f"
+    from s show "s \<le> f"
       by (subst s_eq) (auto simp: image_subset_iff le_fun_def split: split_indicator split_indicator_asm)
   qed
   then show ?thesis
@@ -2285,15 +1953,14 @@
     by(auto simp add: nn_integral_count_space_finite max_def)
   also have "\<dots> = \<integral>\<^sup>+ x'. p x' * indicator {x} x' \<partial>count_space A"
     using assms by(auto simp add: nn_integral_count_space_indicator indicator_def intro!: nn_integral_cong)
-  also have "\<dots> \<le> \<integral>\<^sup>+ x. max 0 (p x) \<partial>count_space A"
+  also have "\<dots> \<le> \<integral>\<^sup>+ x. p x \<partial>count_space A"
     by(rule nn_integral_mono)(simp add: indicator_def)
-  also have "\<dots> = \<integral>\<^sup>+ x. p x \<partial>count_space A" by(simp add: nn_integral_def o_def)
   finally show ?thesis .
 qed
 
 subsubsection \<open>Measure spaces with an associated density\<close>
 
-definition density :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
+definition density :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
   "density M f = measure_of (space M) (sets M) (\<lambda>A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M)"
 
 lemma
@@ -2315,17 +1982,6 @@
   (AE x in M. f x = f' x) \<Longrightarrow> density M f = density M f'"
   unfolding density_def by (auto intro!: measure_of_eq nn_integral_cong_AE sets.space_closed)
 
-lemma density_max_0: "density M f = density M (\<lambda>x. max 0 (f x))"
-proof -
-  have "\<And>x A. max 0 (f x) * indicator A x = max 0 (f x * indicator A x)"
-    by (auto simp: indicator_def)
-  then show ?thesis
-    unfolding density_def by (simp add: nn_integral_max_0)
-qed
-
-lemma density_ereal_max_0: "density M (\<lambda>x. ereal (f x)) = density M (\<lambda>x. ereal (max 0 (f x)))"
-  by (subst density_max_0) (auto intro!: arg_cong[where f="density M"] split: split_max)
-
 lemma emeasure_density:
   assumes f[measurable]: "f \<in> borel_measurable M" and A[measurable]: "A \<in> sets M"
   shows "emeasure (density M f) A = (\<integral>\<^sup>+ x. f x * indicator A x \<partial>M)"
@@ -2334,49 +1990,34 @@
 proof (rule emeasure_measure_of_sigma)
   show "sigma_algebra (space M) (sets M)" ..
   show "positive (sets M) ?\<mu>"
-    using f by (auto simp: positive_def intro!: nn_integral_nonneg)
-  have \<mu>_eq: "?\<mu> = (\<lambda>A. \<integral>\<^sup>+ x. max 0 (f x) * indicator A x \<partial>M)" (is "?\<mu> = ?\<mu>'")
-    apply (subst nn_integral_max_0[symmetric])
-    apply (intro ext nn_integral_cong_AE AE_I2)
-    apply (auto simp: indicator_def)
-    done
+    using f by (auto simp: positive_def)
   show "countably_additive (sets M) ?\<mu>"
-    unfolding \<mu>_eq
   proof (intro countably_additiveI)
     fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets M"
     then have "\<And>i. A i \<in> sets M" by auto
-    then have *: "\<And>i. (\<lambda>x. max 0 (f x) * indicator (A i) x) \<in> borel_measurable M"
-      by (auto simp: set_eq_iff)
+    then have *: "\<And>i. (\<lambda>x. f x * indicator (A i) x) \<in> borel_measurable M"
+      by auto
     assume disj: "disjoint_family A"
-    have "(\<Sum>n. ?\<mu>' (A n)) = (\<integral>\<^sup>+ x. (\<Sum>n. max 0 (f x) * indicator (A n) x) \<partial>M)"
-      using f * by (simp add: nn_integral_suminf)
-    also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) * (\<Sum>n. indicator (A n) x) \<partial>M)" using f
-      by (auto intro!: suminf_cmult_ereal nn_integral_cong_AE)
-    also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) * indicator (\<Union>n. A n) x \<partial>M)"
+    then have "(\<Sum>n. ?\<mu> (A n)) = (\<integral>\<^sup>+ x. (\<Sum>n. f x * indicator (A n) x) \<partial>M)"
+       using f * by (subst nn_integral_suminf) auto
+    also have "(\<integral>\<^sup>+ x. (\<Sum>n. f x * indicator (A n) x) \<partial>M) = (\<integral>\<^sup>+ x. f x * (\<Sum>n. indicator (A n) x) \<partial>M)"
+      using f by (auto intro!: ennreal_suminf_cmult nn_integral_cong_AE)
+    also have "\<dots> = (\<integral>\<^sup>+ x. f x * indicator (\<Union>n. A n) x \<partial>M)"
       unfolding suminf_indicator[OF disj] ..
-    finally show "(\<Sum>n. ?\<mu>' (A n)) = ?\<mu>' (\<Union>x. A x)" by simp
+    finally show "(\<Sum>i. \<integral>\<^sup>+ x. f x * indicator (A i) x \<partial>M) = \<integral>\<^sup>+ x. f x * indicator (\<Union>i. A i) x \<partial>M" .
   qed
 qed fact
 
 lemma null_sets_density_iff:
   assumes f: "f \<in> borel_measurable M"
-  shows "A \<in> null_sets (density M f) \<longleftrightarrow> A \<in> sets M \<and> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)"
+  shows "A \<in> null_sets (density M f) \<longleftrightarrow> A \<in> sets M \<and> (AE x in M. x \<in> A \<longrightarrow> f x = 0)"
 proof -
   { assume "A \<in> sets M"
-    have eq: "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. max 0 (f x) * indicator A x \<partial>M)"
-      apply (subst nn_integral_max_0[symmetric])
-      apply (intro nn_integral_cong)
-      apply (auto simp: indicator_def)
-      done
-    have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow>
-      emeasure M {x \<in> space M. max 0 (f x) * indicator A x \<noteq> 0} = 0"
-      unfolding eq
-      using f \<open>A \<in> sets M\<close>
-      by (intro nn_integral_0_iff) auto
-    also have "\<dots> \<longleftrightarrow> (AE x in M. max 0 (f x) * indicator A x = 0)"
-      using f \<open>A \<in> sets M\<close>
-      by (intro AE_iff_measurable[OF _ refl, symmetric]) auto
-    also have "(AE x in M. max 0 (f x) * indicator A x = 0) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)"
+    have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> emeasure M {x \<in> space M. f x * indicator A x \<noteq> 0} = 0"
+      using f \<open>A \<in> sets M\<close> by (intro nn_integral_0_iff) auto
+    also have "\<dots> \<longleftrightarrow> (AE x in M. f x * indicator A x = 0)"
+      using f \<open>A \<in> sets M\<close> by (intro AE_iff_measurable[OF _ refl, symmetric]) auto
+    also have "(AE x in M. f x * indicator A x = 0) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)"
       by (auto simp add: indicator_def max_def split: if_split_asm)
     finally have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)" . }
   with f show ?thesis
@@ -2388,7 +2029,7 @@
   shows "(AE x in density M f. P x) \<longleftrightarrow> (AE x in M. 0 < f x \<longrightarrow> P x)"
 proof
   assume "AE x in density M f. P x"
-  with f obtain N where "{x \<in> space M. \<not> P x} \<subseteq> N" "N \<in> sets M" and ae: "AE x in M. x \<in> N \<longrightarrow> f x \<le> 0"
+  with f obtain N where "{x \<in> space M. \<not> P x} \<subseteq> N" "N \<in> sets M" and ae: "AE x in M. x \<in> N \<longrightarrow> f x = 0"
     by (auto simp: eventually_ae_filter null_sets_density_iff)
   then have "AE x in M. x \<notin> N \<longrightarrow> P x" by auto
   with ae show "AE x in M. 0 < f x \<longrightarrow> P x"
@@ -2397,19 +2038,18 @@
   fix N assume ae: "AE x in M. 0 < f x \<longrightarrow> P x"
   then obtain N where "{x \<in> space M. \<not> (0 < f x \<longrightarrow> P x)} \<subseteq> N" "N \<in> null_sets M"
     by (auto simp: eventually_ae_filter)
-  then have *: "{x \<in> space (density M f). \<not> P x} \<subseteq> N \<union> {x\<in>space M. \<not> 0 < f x}"
-    "N \<union> {x\<in>space M. \<not> 0 < f x} \<in> sets M" and ae2: "AE x in M. x \<notin> N"
-    using f by (auto simp: subset_eq intro!: sets.sets_Collect_neg AE_not_in)
+  then have *: "{x \<in> space (density M f). \<not> P x} \<subseteq> N \<union> {x\<in>space M. f x = 0}"
+    "N \<union> {x\<in>space M. f x = 0} \<in> sets M" and ae2: "AE x in M. x \<notin> N"
+    using f by (auto simp: subset_eq zero_less_iff_neq_zero intro!: AE_not_in)
   show "AE x in density M f. P x"
     using ae2
     unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f]
-    by (intro exI[of _ "N \<union> {x\<in>space M. \<not> 0 < f x}"] conjI *)
-       (auto elim: eventually_elim2)
+    by (intro exI[of _ "N \<union> {x\<in>space M. f x = 0}"] conjI *) (auto elim: eventually_elim2)
 qed
 
-lemma nn_integral_density':
-  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
-  assumes g: "g \<in> borel_measurable M" "\<And>x. 0 \<le> g x"
+lemma nn_integral_density:
+  assumes f: "f \<in> borel_measurable M"
+  assumes g: "g \<in> borel_measurable M"
   shows "integral\<^sup>N (density M f) g = (\<integral>\<^sup>+ x. f x * g x \<partial>M)"
 using g proof induct
   case (cong u v)
@@ -2428,28 +2068,21 @@
 next
   case (add u v)
   then have "\<And>x. f x * (v x + u x) = f x * v x + f x * u x"
-    by (simp add: ereal_right_distrib)
+    by (simp add: distrib_left)
   with add f show ?case
-    by (auto simp add: nn_integral_add ereal_zero_le_0_iff intro!: nn_integral_add[symmetric])
+    by (auto simp add: nn_integral_add intro!: nn_integral_add[symmetric])
 next
   case (seq U)
-  from f(2) have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)"
-    by eventually_elim (simp add: SUP_ereal_mult_left seq)
+  have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)"
+    by eventually_elim (simp add: SUP_mult_left_ennreal seq)
   from seq f show ?case
     apply (simp add: nn_integral_monotone_convergence_SUP)
     apply (subst nn_integral_cong_AE[OF eq])
     apply (subst nn_integral_monotone_convergence_SUP_AE)
-    apply (auto simp: incseq_def le_fun_def intro!: ereal_mult_left_mono)
+    apply (auto simp: incseq_def le_fun_def intro!: mult_left_mono)
     done
 qed
 
-lemma nn_integral_density:
-  "f \<in> borel_measurable M \<Longrightarrow> AE x in M. 0 \<le> f x \<Longrightarrow> g' \<in> borel_measurable M \<Longrightarrow>
-    integral\<^sup>N (density M f) g' = (\<integral>\<^sup>+ x. f x * g' x \<partial>M)"
-  by (subst (1 2) nn_integral_max_0[symmetric])
-     (auto intro!: nn_integral_cong_AE
-           simp: measurable_If max_def ereal_zero_le_0_iff nn_integral_density')
-
 lemma density_distr:
   assumes [measurable]: "f \<in> borel_measurable N" "X \<in> measurable M N"
   shows "density (distr M N X) f = distr (density M (\<lambda>x. f (X x))) N X"
@@ -2479,15 +2112,15 @@
   by standard (simp add: emeasure_restricted)
 
 lemma emeasure_density_const:
-  "A \<in> sets M \<Longrightarrow> 0 \<le> c \<Longrightarrow> emeasure (density M (\<lambda>_. c)) A = c * emeasure M A"
+  "A \<in> sets M \<Longrightarrow> emeasure (density M (\<lambda>_. c)) A = c * emeasure M A"
   by (auto simp: nn_integral_cmult_indicator emeasure_density)
 
 lemma measure_density_const:
-  "A \<in> sets M \<Longrightarrow> 0 \<le> c \<Longrightarrow> c \<noteq> \<infinity> \<Longrightarrow> measure (density M (\<lambda>_. c)) A = real_of_ereal c * measure M A"
-  by (auto simp: emeasure_density_const measure_def)
+  "A \<in> sets M \<Longrightarrow> c \<noteq> \<infinity> \<Longrightarrow> measure (density M (\<lambda>_. c)) A = enn2real c * measure M A"
+  by (auto simp: emeasure_density_const measure_def enn2real_mult)
 
 lemma density_density_eq:
-   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> AE x in M. 0 \<le> f x \<Longrightarrow>
+   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
    density (density M f) g = density M (\<lambda>x. f x * g x)"
   by (auto intro!: measure_eqI simp: emeasure_density nn_integral_density ac_simps)
 
@@ -2500,7 +2133,7 @@
   fix A assume A: "A \<in> sets ?R"
   { fix x assume "x \<in> space M"
     with sets.sets_into_space[OF A]
-    have "indicator (T' -` A \<inter> space M') (T x) = (indicator A x :: ereal)"
+    have "indicator (T' -` A \<inter> space M') (T x) = (indicator A x :: ennreal)"
       using T inv by (auto simp: indicator_def measurable_space) }
   with A T T' f show "emeasure ?R A = emeasure ?L A"
     by (simp add: measurable_comp emeasure_density emeasure_distr
@@ -2514,8 +2147,8 @@
   assumes ac: "AE x in M. f x = 0 \<longrightarrow> g x = 0"
   shows "density (density M f) (\<lambda>x. g x / f x) = density M g"
 proof -
-  have "density M g = density M (\<lambda>x. f x * (g x / f x))"
-    using f g ac by (auto intro!: density_cong measurable_If)
+  have "density M g = density M (\<lambda>x. ennreal (f x) * ennreal (g x / f x))"
+    using f g ac by (auto intro!: density_cong measurable_If simp: ennreal_mult[symmetric])
   then show ?thesis
     using f g by (subst density_density_eq) auto
 qed
@@ -2527,8 +2160,6 @@
   assumes X: "X \<in> sets M"
   assumes Mf[measurable]: "f \<in> borel_measurable M"
   assumes Mg[measurable]: "g \<in> borel_measurable M"
-  assumes nonnegf: "\<And>x. x \<in> space M \<Longrightarrow> f x \<ge> 0"
-  assumes nonnegg: "\<And>x. x \<in> space M \<Longrightarrow> g x \<ge> 0"
   shows "emeasure (density M f) X + emeasure (density M g) X =
            emeasure (density M (\<lambda>x. f x + g x)) X"
   using assms
@@ -2539,7 +2170,7 @@
 
 subsubsection \<open>Point measure\<close>
 
-definition point_measure :: "'a set \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
+definition point_measure :: "'a set \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
   "point_measure A f = density (count_space A) f"
 
 lemma
@@ -2571,21 +2202,20 @@
   have "{a. (a \<in> X \<longrightarrow> a \<in> A \<and> 0 < f a) \<and> a \<in> X} = {a\<in>X. 0 < f a}"
     using \<open>X \<subseteq> A\<close> by auto
   with A show ?thesis
-    by (simp add: emeasure_density nn_integral_count_space ereal_zero_le_0_iff
-                  point_measure_def indicator_def)
+    by (simp add: emeasure_density nn_integral_count_space point_measure_def indicator_def)
 qed
 
 lemma emeasure_point_measure_finite:
-  "finite A \<Longrightarrow> (\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)"
+  "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)"
   by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum.mono_neutral_left simp: less_le)
 
 lemma emeasure_point_measure_finite2:
-  "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> (\<And>i. i \<in> X \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)"
+  "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)"
   by (subst emeasure_point_measure)
      (auto dest: finite_subset intro!: setsum.mono_neutral_left simp: less_le)
 
 lemma null_sets_point_measure_iff:
-  "X \<in> null_sets (point_measure A f) \<longleftrightarrow> X \<subseteq> A \<and> (\<forall>x\<in>X. f x \<le> 0)"
+  "X \<in> null_sets (point_measure A f) \<longleftrightarrow> X \<subseteq> A \<and> (\<forall>x\<in>X. f x = 0)"
  by (auto simp: AE_count_space null_sets_density_iff point_measure_def)
 
 lemma AE_point_measure:
@@ -2597,20 +2227,11 @@
   "finite {a\<in>A. 0 < f a \<and> 0 < g a} \<Longrightarrow>
     integral\<^sup>N (point_measure A f) g = (\<Sum>a|a\<in>A \<and> 0 < f a \<and> 0 < g a. f a * g a)"
   unfolding point_measure_def
-  apply (subst density_max_0)
-  apply (subst nn_integral_density)
-  apply (simp_all add: AE_count_space nn_integral_density)
-  apply (subst nn_integral_count_space )
-  apply (auto intro!: setsum.cong simp: max_def ereal_zero_less_0_iff)
-  apply (rule finite_subset)
-  prefer 2
-  apply assumption
-  apply auto
-  done
+  by (subst nn_integral_density)
+     (simp_all add: nn_integral_density nn_integral_count_space ennreal_zero_less_mult_iff)
 
 lemma nn_integral_point_measure_finite:
-  "finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> f a) \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> g a) \<Longrightarrow>
-    integral\<^sup>N (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)"
+  "finite A \<Longrightarrow> integral\<^sup>N (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)"
   by (subst nn_integral_point_measure) (auto intro!: setsum.mono_neutral_left simp: less_le)
 
 subsubsection \<open>Uniform measure\<close>
@@ -2627,11 +2248,11 @@
   shows "emeasure (uniform_measure M A) B = emeasure M (A \<inter> B) / emeasure M A"
 proof -
   from A B have "emeasure (uniform_measure M A) B = (\<integral>\<^sup>+x. (1 / emeasure M A) * indicator (A \<inter> B) x \<partial>M)"
-    by (auto simp add: uniform_measure_def emeasure_density split: split_indicator
+    by (auto simp add: uniform_measure_def emeasure_density divide_ennreal_def split: split_indicator
              intro!: nn_integral_cong)
   also have "\<dots> = emeasure M (A \<inter> B) / emeasure M A"
     using A B
-    by (subst nn_integral_cmult_indicator) (simp_all add: sets.Int emeasure_nonneg)
+    by (subst nn_integral_cmult_indicator) (simp_all add: sets.Int divide_ennreal_def mult.commute)
   finally show ?thesis .
 qed
 
@@ -2639,19 +2260,21 @@
   assumes A: "emeasure M A \<noteq> 0" "emeasure M A \<noteq> \<infinity>" and B: "B \<in> sets M"
   shows "measure (uniform_measure M A) B = measure M (A \<inter> B) / measure M A"
   using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A
-  by (cases "emeasure M A" "emeasure M (A \<inter> B)" rule: ereal2_cases) (simp_all add: measure_def)
+  by (cases "emeasure M A" "emeasure M (A \<inter> B)" rule: ennreal2_cases)
+     (simp_all add: measure_def divide_ennreal top_ennreal.rep_eq top_ereal_def ennreal_top_divide)
 
 lemma AE_uniform_measureI:
   "A \<in> sets M \<Longrightarrow> (AE x in M. x \<in> A \<longrightarrow> P x) \<Longrightarrow> (AE x in uniform_measure M A. P x)"
-  unfolding uniform_measure_def by (auto simp: AE_density)
+  unfolding uniform_measure_def by (auto simp: AE_density divide_ennreal_def)
 
 lemma emeasure_uniform_measure_1:
   "emeasure M S \<noteq> 0 \<Longrightarrow> emeasure M S \<noteq> \<infinity> \<Longrightarrow> emeasure (uniform_measure M S) S = 1"
   by (subst emeasure_uniform_measure)
-     (simp_all add: emeasure_nonneg emeasure_neq_0_sets)
+     (simp_all add: emeasure_neq_0_sets emeasure_eq_ennreal_measure divide_ennreal
+                    zero_less_iff_neq_zero[symmetric])
 
 lemma nn_integral_uniform_measure:
-  assumes f[measurable]: "f \<in> borel_measurable M" and "\<And>x. 0 \<le> f x" and S[measurable]: "S \<in> sets M"
+  assumes f[measurable]: "f \<in> borel_measurable M" and S[measurable]: "S \<in> sets M"
   shows "(\<integral>\<^sup>+x. f x \<partial>uniform_measure M S) = (\<integral>\<^sup>+x. f x * indicator S x \<partial>M) / emeasure M S"
 proof -
   { assume "emeasure M S = \<infinity>"
@@ -2662,22 +2285,16 @@
     then have ae: "AE x in M. x \<notin> S"
       using sets.sets_into_space[OF S]
       by (subst AE_iff_measurable[OF _ refl]) (simp_all add: subset_eq cong: rev_conj_cong)
-    from ae have "(\<integral>\<^sup>+ x. indicator S x * f x / 0 \<partial>M) = 0"
+    from ae have "(\<integral>\<^sup>+ x. indicator S x / 0 * f x \<partial>M) = 0"
       by (subst nn_integral_0_iff_AE) auto
     moreover from ae have "(\<integral>\<^sup>+ x. f x * indicator S x \<partial>M) = 0"
       by (subst nn_integral_0_iff_AE) auto
     ultimately have ?thesis
       by (simp add: uniform_measure_def nn_integral_density f) }
-  moreover
-  { assume "emeasure M S \<noteq> 0" "emeasure M S \<noteq> \<infinity>"
-    moreover then have "0 < emeasure M S"
-      by (simp add: emeasure_nonneg less_le)
-    ultimately have ?thesis
-      unfolding uniform_measure_def
-      apply (subst nn_integral_density)
-      apply (auto simp: f nn_integral_divide intro!: zero_le_divide_ereal)
-      apply (simp add: mult.commute)
-      done }
+  moreover have "emeasure M S \<noteq> 0 \<Longrightarrow> emeasure M S \<noteq> \<infinity> \<Longrightarrow> ?thesis"
+    unfolding uniform_measure_def
+    by (subst nn_integral_density)
+       (auto simp: ennreal_times_divide f nn_integral_divide[symmetric] mult.commute)
   ultimately show ?thesis by blast
 qed
 
@@ -2688,8 +2305,8 @@
   have "A \<in> sets M"
     using \<open>emeasure M A \<noteq> 0\<close> by (metis emeasure_notin_sets)
   moreover have "\<And>x. 0 < indicator A x / emeasure M A \<longleftrightarrow> x \<in> A"
-    using emeasure_nonneg[of M A] assms
-    by (cases "emeasure M A") (auto split: split_indicator simp: one_ereal_def)
+    using assms
+    by (cases "emeasure M A") (auto split: split_indicator simp: ennreal_zero_less_divide)
   ultimately show ?thesis
     unfolding uniform_measure_def by (simp add: AE_density)
 qed
@@ -2700,7 +2317,7 @@
   by (intro measure_eqI) (simp_all add: emeasure_density)
 
 lemma nn_integral_null_measure[simp]: "(\<integral>\<^sup>+x. f x \<partial>null_measure M) = 0"
-  by (auto simp add: nn_integral_def simple_integral_def SUP_constant bot_ereal_def le_fun_def
+  by (auto simp add: nn_integral_def simple_integral_def SUP_constant bot_ennreal_def le_fun_def
            intro!: exI[of _ "\<lambda>x. 0"])
 
 lemma density_null_measure[simp]: "density (null_measure M) f = null_measure M"
@@ -2724,11 +2341,12 @@
 
 lemma emeasure_uniform_count_measure:
   "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (uniform_count_measure A) X = card X / card A"
-  by (simp add: emeasure_point_measure_finite uniform_count_measure_def)
+  by (simp add: emeasure_point_measure_finite uniform_count_measure_def divide_inverse ennreal_mult
+                ennreal_of_nat_eq_real_of_nat)
 
 lemma measure_uniform_count_measure:
   "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> measure (uniform_count_measure A) X = card X / card A"
-  by (simp add: emeasure_point_measure_finite uniform_count_measure_def measure_def)
+  by (simp add: emeasure_point_measure_finite uniform_count_measure_def measure_def enn2real_mult)
 
 lemma space_uniform_count_measure_empty_iff [simp]:
   "space (uniform_count_measure X) = {} \<longleftrightarrow> X = {}"
@@ -2741,9 +2359,9 @@
 
 subsubsection \<open>Scaled measure\<close>
 
-lemma nn_integral_scale_measure':
-  assumes f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x"
-  shows "nn_integral (scale_measure r M) f = max 0 r * nn_integral M f"
+lemma nn_integral_scale_measure:
+  assumes f: "f \<in> borel_measurable M"
+  shows "nn_integral (scale_measure r M) f = r * nn_integral M f"
   using f
 proof induction
   case (cong f g)
@@ -2756,15 +2374,11 @@
 next
   case (add f g)
   thus ?case
-    by(simp add: nn_integral_add ereal_right_distrib nn_integral_nonneg)
+    by(simp add: nn_integral_add distrib_left)
 next
   case (seq U)
   thus ?case
-    by(simp add: nn_integral_monotone_convergence_SUP SUP_ereal_mult_left nn_integral_nonneg)
+    by(simp add: nn_integral_monotone_convergence_SUP SUP_mult_left_ennreal)
 qed simp
 
-lemma nn_integral_scale_measure:
-  "f \<in> borel_measurable M \<Longrightarrow> nn_integral (scale_measure r M) f = max 0 r * nn_integral M f"
-by(subst (1 2) nn_integral_max_0[symmetric])(rule nn_integral_scale_measure', simp_all)
-
 end
--- a/src/HOL/Probability/Probability_Mass_Function.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Probability_Mass_Function.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -22,16 +22,13 @@
     with x_M[THEN sets.sets_into_space] N have "emeasure M {x} \<le> emeasure M N"
       by (intro emeasure_mono) auto
     with x N have False
-      by (auto simp: emeasure_le_0_iff) }
+      by (auto simp:) }
   then show "P x" by auto
 qed
 
 lemma AE_measure_singleton: "measure M {x} \<noteq> 0 \<Longrightarrow> AE x in M. P x \<Longrightarrow> P x"
   by (metis AE_emeasure_singleton measure_def emeasure_empty measure_empty)
 
-lemma ereal_divide': "b \<noteq> 0 \<Longrightarrow> ereal (a / b) = ereal a / ereal b"
-  using ereal_divide[of a b] by simp
-
 lemma (in finite_measure) AE_support_countable:
   assumes [simp]: "sets M = UNIV"
   shows "(AE x in M. measure M {x} \<noteq> 0) \<longleftrightarrow> (\<exists>S. countable S \<and> (AE x in M. x \<in> S))"
@@ -54,7 +51,7 @@
     by (simp add: emeasure_single_in_space cong: rev_conj_cong)
   with finite_measure_compl[of "{x \<in> space M. x\<in>S \<and> emeasure M {x} \<noteq> 0}"]
   have "AE x in M. x \<in> S \<and> emeasure M {x} \<noteq> 0"
-    by (intro AE_I[OF order_refl]) (auto simp: emeasure_eq_measure set_diff_eq cong: conj_cong)
+    by (intro AE_I[OF order_refl]) (auto simp: emeasure_eq_measure measure_nonneg set_diff_eq cong: conj_cong)
   then show "AE x in M. measure M {x} \<noteq> 0"
     by (auto simp: emeasure_eq_measure)
 qed (auto intro!: exI[of _ "{x. measure M {x} \<noteq> 0}"] countable_support)
@@ -153,7 +150,7 @@
 lemma emeasure_pmf_single_eq_zero_iff:
   fixes M :: "'a pmf"
   shows "emeasure M {y} = 0 \<longleftrightarrow> y \<notin> M"
-  by transfer (simp add: finite_measure.emeasure_eq_measure[OF prob_space.finite_measure])
+  unfolding set_pmf.rep_eq by (simp add: measure_pmf.emeasure_eq_measure)
 
 lemma AE_measure_pmf_iff: "(AE x in measure_pmf M. P x) \<longleftrightarrow> (\<forall>y\<in>M. P y)"
   using AE_measure_singleton[of M] AE_measure_pmf[of M]
@@ -166,10 +163,10 @@
   by transfer (metis prob_space.finite_measure finite_measure.countable_support)
 
 lemma pmf_positive: "x \<in> set_pmf p \<Longrightarrow> 0 < pmf p x"
-  by transfer (simp add: less_le measure_nonneg)
+  by transfer (simp add: less_le)
 
-lemma pmf_nonneg: "0 \<le> pmf p x"
-  by transfer (simp add: measure_nonneg)
+lemma pmf_nonneg[simp]: "0 \<le> pmf p x"
+  by transfer simp
 
 lemma pmf_le_1: "pmf p x \<le> 1"
   by (simp add: pmf.rep_eq)
@@ -180,6 +177,9 @@
 lemma set_pmf_iff: "x \<in> set_pmf M \<longleftrightarrow> pmf M x \<noteq> 0"
   by transfer simp
 
+lemma pmf_positive_iff: "0 < pmf p x \<longleftrightarrow> x \<in> set_pmf p"
+  unfolding less_le by (simp add: set_pmf_iff)
+
 lemma set_pmf_eq: "set_pmf M = {x. pmf M x \<noteq> 0}"
   by (auto simp: set_pmf_iff)
 
@@ -189,16 +189,17 @@
   by transfer (simp add: finite_measure.emeasure_eq_measure[OF prob_space.finite_measure])
 
 lemma measure_pmf_single: "measure (measure_pmf M) {x} = pmf M x"
-using emeasure_pmf_single[of M x] by(simp add: measure_pmf.emeasure_eq_measure)
+  using emeasure_pmf_single[of M x] by(simp add: measure_pmf.emeasure_eq_measure pmf_nonneg measure_nonneg)
 
 lemma emeasure_measure_pmf_finite: "finite S \<Longrightarrow> emeasure (measure_pmf M) S = (\<Sum>s\<in>S. pmf M s)"
-  by (subst emeasure_eq_setsum_singleton) (auto simp: emeasure_pmf_single)
+  by (subst emeasure_eq_setsum_singleton) (auto simp: emeasure_pmf_single pmf_nonneg)
 
 lemma measure_measure_pmf_finite: "finite S \<Longrightarrow> measure (measure_pmf M) S = setsum (pmf M) S"
-  using emeasure_measure_pmf_finite[of S M] by(simp add: measure_pmf.emeasure_eq_measure)
+  using emeasure_measure_pmf_finite[of S M]
+  by (simp add: measure_pmf.emeasure_eq_measure measure_nonneg setsum_nonneg pmf_nonneg)
 
 lemma nn_integral_measure_pmf_support:
-  fixes f :: "'a \<Rightarrow> ereal"
+  fixes f :: "'a \<Rightarrow> ennreal"
   assumes f: "finite A" and nn: "\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x" "\<And>x. x \<in> set_pmf M \<Longrightarrow> x \<notin> A \<Longrightarrow> f x = 0"
   shows "(\<integral>\<^sup>+x. f x \<partial>measure_pmf M) = (\<Sum>x\<in>A. f x * pmf M x)"
 proof -
@@ -211,14 +212,15 @@
 qed
 
 lemma nn_integral_measure_pmf_finite:
-  fixes f :: "'a \<Rightarrow> ereal"
+  fixes f :: "'a \<Rightarrow> ennreal"
   assumes f: "finite (set_pmf M)" and nn: "\<And>x. x \<in> set_pmf M \<Longrightarrow> 0 \<le> f x"
   shows "(\<integral>\<^sup>+x. f x \<partial>measure_pmf M) = (\<Sum>x\<in>set_pmf M. f x * pmf M x)"
   using assms by (intro nn_integral_measure_pmf_support) auto
+
 lemma integrable_measure_pmf_finite:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
   shows "finite (set_pmf M) \<Longrightarrow> integrable M f"
-  by (auto intro!: integrableI_bounded simp: nn_integral_measure_pmf_finite)
+  by (auto intro!: integrableI_bounded simp: nn_integral_measure_pmf_finite ennreal_mult_less_top)
 
 lemma integral_measure_pmf:
   assumes [simp]: "finite A" and "\<And>a. a \<in> set_pmf M \<Longrightarrow> f a \<noteq> 0 \<Longrightarrow> a \<in> A"
@@ -227,7 +229,8 @@
   have "(\<integral>x. f x \<partial>measure_pmf M) = (\<integral>x. f x * indicator A x \<partial>measure_pmf M)"
     using assms(2) by (intro integral_cong_AE) (auto split: split_indicator simp: AE_measure_pmf_iff)
   also have "\<dots> = (\<Sum>a\<in>A. f a * pmf M a)"
-    by (subst integral_indicator_finite_real) (auto simp: measure_def emeasure_measure_pmf_finite)
+    by (subst integral_indicator_finite_real)
+       (auto simp: measure_def emeasure_measure_pmf_finite pmf_nonneg)
   finally show ?thesis .
 qed
 
@@ -254,7 +257,7 @@
   also have "\<dots> = emeasure M X"
     by (auto intro!: emeasure_eq_AE simp: AE_measure_pmf_iff)
   finally show ?thesis
-    by (simp add: measure_pmf.emeasure_eq_measure)
+    by (simp add: measure_pmf.emeasure_eq_measure measure_nonneg integral_nonneg pmf_nonneg)
 qed
 
 lemma integral_pmf_restrict:
@@ -336,23 +339,29 @@
     done
 qed
 
-lemma ereal_pmf_bind: "pmf (bind_pmf N f) i = (\<integral>\<^sup>+x. pmf (f x) i \<partial>measure_pmf N)"
+lemma ennreal_pmf_bind: "pmf (bind_pmf N f) i = (\<integral>\<^sup>+x. pmf (f x) i \<partial>measure_pmf N)"
   unfolding pmf.rep_eq bind_pmf.rep_eq
   by (auto simp: measure_pmf.measure_bind[where N="count_space UNIV"] measure_subprob measure_nonneg
            intro!: nn_integral_eq_integral[symmetric] measure_pmf.integrable_const_bound[where B=1])
 
 lemma pmf_bind: "pmf (bind_pmf N f) i = (\<integral>x. pmf (f x) i \<partial>measure_pmf N)"
-  using ereal_pmf_bind[of N f i]
+  using ennreal_pmf_bind[of N f i]
   by (subst (asm) nn_integral_eq_integral)
-     (auto simp: pmf_nonneg pmf_le_1
+     (auto simp: pmf_nonneg pmf_le_1 pmf_nonneg integral_nonneg
            intro!: nn_integral_eq_integral[symmetric] measure_pmf.integrable_const_bound[where B=1])
 
 lemma bind_pmf_const[simp]: "bind_pmf M (\<lambda>x. c) = c"
   by transfer (simp add: bind_const' prob_space_imp_subprob_space)
 
 lemma set_bind_pmf[simp]: "set_pmf (bind_pmf M N) = (\<Union>M\<in>set_pmf M. set_pmf (N M))"
-  unfolding set_pmf_eq ereal_eq_0(1)[symmetric] ereal_pmf_bind
-  by (auto simp add: nn_integral_0_iff_AE AE_measure_pmf_iff set_pmf_eq not_le less_le pmf_nonneg)
+proof -
+  have "set_pmf (bind_pmf M N) = {x. ennreal (pmf (bind_pmf M N) x) \<noteq> 0}"
+    by (simp add: set_pmf_eq pmf_nonneg)
+  also have "\<dots> = (\<Union>M\<in>set_pmf M. set_pmf (N M))"
+    unfolding ennreal_pmf_bind
+    by (subst nn_integral_0_iff_AE) (auto simp: AE_measure_pmf_iff pmf_nonneg set_pmf_eq)
+  finally show ?thesis .
+qed
 
 lemma bind_pmf_cong:
   assumes "p = q"
@@ -372,7 +381,6 @@
 lemma nn_integral_bind_pmf[simp]: "(\<integral>\<^sup>+x. f x \<partial>bind_pmf M N) = (\<integral>\<^sup>+x. \<integral>\<^sup>+y. f y \<partial>N x \<partial>M)"
   using measurable_measure_pmf[of N]
   unfolding measure_pmf_bind
-  apply (subst (1 3) nn_integral_max_0[symmetric])
   apply (intro nn_integral_bind[where B="count_space UNIV"])
   apply auto
   done
@@ -452,18 +460,28 @@
   unfolding map_pmf_rep_eq by (subst emeasure_distr) auto
 
 lemma measure_map_pmf[simp]: "measure (map_pmf f M) X = measure M (f -` X)"
-using emeasure_map_pmf[of f M X] by(simp add: measure_pmf.emeasure_eq_measure)
+using emeasure_map_pmf[of f M X] by(simp add: measure_pmf.emeasure_eq_measure measure_nonneg)
 
 lemma nn_integral_map_pmf[simp]: "(\<integral>\<^sup>+x. f x \<partial>map_pmf g M) = (\<integral>\<^sup>+x. f (g x) \<partial>M)"
   unfolding map_pmf_rep_eq by (intro nn_integral_distr) auto
 
-lemma ereal_pmf_map: "pmf (map_pmf f p) x = (\<integral>\<^sup>+ y. indicator (f -` {x}) y \<partial>measure_pmf p)"
+lemma ennreal_pmf_map: "pmf (map_pmf f p) x = (\<integral>\<^sup>+ y. indicator (f -` {x}) y \<partial>measure_pmf p)"
 proof (transfer fixing: f x)
   fix p :: "'b measure"
   presume "prob_space p"
   then interpret prob_space p .
   presume "sets p = UNIV"
-  then show "ereal (measure (distr p (count_space UNIV) f) {x}) = integral\<^sup>N p (indicator (f -` {x}))"
+  then show "ennreal (measure (distr p (count_space UNIV) f) {x}) = integral\<^sup>N p (indicator (f -` {x}))"
+    by(simp add: measure_distr measurable_def emeasure_eq_measure)
+qed simp_all
+
+lemma pmf_map: "pmf (map_pmf f p) x = measure p (f -` {x})"
+proof (transfer fixing: f x)
+  fix p :: "'b measure"
+  presume "prob_space p"
+  then interpret prob_space p .
+  presume "sets p = UNIV"
+  then show "measure (distr p (count_space UNIV) f) {x} = measure p (f -` {x})"
     by(simp add: measure_distr measurable_def emeasure_eq_measure)
 qed simp_all
 
@@ -480,6 +498,11 @@
   finally show ?thesis .
 qed
 
+lemma integral_map_pmf[simp]:
+  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+  shows "integral\<^sup>L (map_pmf g p) f = integral\<^sup>L p (\<lambda>x. f (g x))"
+  by (simp add: integral_distr map_pmf_rep_eq)
+
 lemma map_return_pmf [simp]: "map_pmf f (return_pmf x) = return_pmf (f x)"
   by transfer (simp add: distr_return)
 
@@ -529,15 +552,13 @@
 
 lemma nn_integral_pair_pmf': "(\<integral>\<^sup>+x. f x \<partial>pair_pmf A B) = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. f (a, b) \<partial>B \<partial>A)"
 proof -
-  have "(\<integral>\<^sup>+x. f x \<partial>pair_pmf A B) = (\<integral>\<^sup>+x. max 0 (f x) * indicator (A \<times> B) x \<partial>pair_pmf A B)"
-    by (subst nn_integral_max_0[symmetric])
-       (auto simp: AE_measure_pmf_iff intro!: nn_integral_cong_AE)
-  also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. max 0 (f (a, b)) * indicator (A \<times> B) (a, b) \<partial>B \<partial>A)"
+  have "(\<integral>\<^sup>+x. f x \<partial>pair_pmf A B) = (\<integral>\<^sup>+x. f x * indicator (A \<times> B) x \<partial>pair_pmf A B)"
+    by (auto simp: AE_measure_pmf_iff intro!: nn_integral_cong_AE)
+  also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. f (a, b) * indicator (A \<times> B) (a, b) \<partial>B \<partial>A)"
     by (simp add: pair_pmf_def)
-  also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. max 0 (f (a, b)) \<partial>B \<partial>A)"
+  also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. f (a, b) \<partial>B \<partial>A)"
     by (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff)
-  finally show ?thesis
-    unfolding nn_integral_max_0 .
+  finally show ?thesis .
 qed
 
 lemma bind_pair_pmf:
@@ -562,7 +583,7 @@
   then show "emeasure ?L X = emeasure ?R X"
     apply (simp add: emeasure_bind[OF _ M' X])
     apply (simp add: nn_integral_bind[where B="count_space UNIV"] pair_pmf_def measure_pmf_bind[of A]
-                     nn_integral_measure_pmf_finite emeasure_nonneg one_ereal_def[symmetric])
+                     nn_integral_measure_pmf_finite)
     apply (subst emeasure_bind[OF _ _ X])
     apply measurable
     apply (subst emeasure_bind[OF _ _ X])
@@ -582,10 +603,10 @@
      (auto simp: bij_betw_def nn_integral_pmf)
 
 lemma pmf_le_0_iff[simp]: "pmf M p \<le> 0 \<longleftrightarrow> pmf M p = 0"
-  using pmf_nonneg[of M p] by simp
+  using pmf_nonneg[of M p] by arith
 
 lemma min_pmf_0[simp]: "min (pmf M p) 0 = 0" "min 0 (pmf M p) = 0"
-  using pmf_nonneg[of M p] by simp_all
+  using pmf_nonneg[of M p] by arith+
 
 lemma pmf_eq_0_set_pmf: "pmf M p = 0 \<longleftrightarrow> p \<notin> set_pmf M"
   unfolding set_pmf_iff by simp
@@ -612,22 +633,22 @@
   assumes prob: "(\<integral>\<^sup>+x. f x \<partial>count_space UNIV) = 1"
 begin
 
-lift_definition embed_pmf :: "'a pmf" is "density (count_space UNIV) (ereal \<circ> f)"
+lift_definition embed_pmf :: "'a pmf" is "density (count_space UNIV) (ennreal \<circ> f)"
 proof (intro conjI)
-  have *[simp]: "\<And>x y. ereal (f y) * indicator {x} y = ereal (f x) * indicator {x} y"
+  have *[simp]: "\<And>x y. ennreal (f y) * indicator {x} y = ennreal (f x) * indicator {x} y"
     by (simp split: split_indicator)
-  show "AE x in density (count_space UNIV) (ereal \<circ> f).
-    measure (density (count_space UNIV) (ereal \<circ> f)) {x} \<noteq> 0"
+  show "AE x in density (count_space UNIV) (ennreal \<circ> f).
+    measure (density (count_space UNIV) (ennreal \<circ> f)) {x} \<noteq> 0"
     by (simp add: AE_density nonneg measure_def emeasure_density max_def)
-  show "prob_space (density (count_space UNIV) (ereal \<circ> f))"
+  show "prob_space (density (count_space UNIV) (ennreal \<circ> f))"
     by standard (simp add: emeasure_density prob)
 qed simp
 
 lemma pmf_embed_pmf: "pmf embed_pmf x = f x"
 proof transfer
-  have *[simp]: "\<And>x y. ereal (f y) * indicator {x} y = ereal (f x) * indicator {x} y"
+  have *[simp]: "\<And>x y. ennreal (f y) * indicator {x} y = ennreal (f x) * indicator {x} y"
     by (simp split: split_indicator)
-  fix x show "measure (density (count_space UNIV) (ereal \<circ> f)) {x} = f x"
+  fix x show "measure (density (count_space UNIV) (ennreal \<circ> f)) {x} = f x"
     by transfer (simp add: measure_def emeasure_density nonneg max_def)
 qed
 
@@ -637,17 +658,17 @@
 end
 
 lemma embed_pmf_transfer:
-  "rel_fun (eq_onp (\<lambda>f. (\<forall>x. 0 \<le> f x) \<and> (\<integral>\<^sup>+x. ereal (f x) \<partial>count_space UNIV) = 1)) pmf_as_measure.cr_pmf (\<lambda>f. density (count_space UNIV) (ereal \<circ> f)) embed_pmf"
+  "rel_fun (eq_onp (\<lambda>f. (\<forall>x. 0 \<le> f x) \<and> (\<integral>\<^sup>+x. ennreal (f x) \<partial>count_space UNIV) = 1)) pmf_as_measure.cr_pmf (\<lambda>f. density (count_space UNIV) (ennreal \<circ> f)) embed_pmf"
   by (auto simp: rel_fun_def eq_onp_def embed_pmf.transfer)
 
 lemma measure_pmf_eq_density: "measure_pmf p = density (count_space UNIV) (pmf p)"
 proof (transfer, elim conjE)
   fix M :: "'a measure" assume [simp]: "sets M = UNIV" and ae: "AE x in M. measure M {x} \<noteq> 0"
   assume "prob_space M" then interpret prob_space M .
-  show "M = density (count_space UNIV) (\<lambda>x. ereal (measure M {x}))"
+  show "M = density (count_space UNIV) (\<lambda>x. ennreal (measure M {x}))"
   proof (rule measure_eqI)
     fix A :: "'a set"
-    have "(\<integral>\<^sup>+ x. ereal (measure M {x}) * indicator A x \<partial>count_space UNIV) =
+    have "(\<integral>\<^sup>+ x. ennreal (measure M {x}) * indicator A x \<partial>count_space UNIV) =
       (\<integral>\<^sup>+ x. emeasure M {x} * indicator (A \<inter> {x. measure M {x} \<noteq> 0}) x \<partial>count_space UNIV)"
       by (auto intro!: nn_integral_cong simp: emeasure_eq_measure split: split_indicator)
     also have "\<dots> = (\<integral>\<^sup>+ x. emeasure M {x} \<partial>count_space (A \<inter> {x. measure M {x} \<noteq> 0}))"
@@ -657,19 +678,19 @@
          (auto simp: disjoint_family_on_def)
     also have "\<dots> = emeasure M A"
       using ae by (intro emeasure_eq_AE) auto
-    finally show " emeasure M A = emeasure (density (count_space UNIV) (\<lambda>x. ereal (measure M {x}))) A"
+    finally show " emeasure M A = emeasure (density (count_space UNIV) (\<lambda>x. ennreal (measure M {x}))) A"
       using emeasure_space_1 by (simp add: emeasure_density)
   qed simp
 qed
 
 lemma td_pmf_embed_pmf:
-  "type_definition pmf embed_pmf {f::'a \<Rightarrow> real. (\<forall>x. 0 \<le> f x) \<and> (\<integral>\<^sup>+x. ereal (f x) \<partial>count_space UNIV) = 1}"
+  "type_definition pmf embed_pmf {f::'a \<Rightarrow> real. (\<forall>x. 0 \<le> f x) \<and> (\<integral>\<^sup>+x. ennreal (f x) \<partial>count_space UNIV) = 1}"
   unfolding type_definition_def
 proof safe
   fix p :: "'a pmf"
   have "(\<integral>\<^sup>+ x. 1 \<partial>measure_pmf p) = 1"
     using measure_pmf.emeasure_space_1[of p] by simp
-  then show *: "(\<integral>\<^sup>+ x. ereal (pmf p x) \<partial>count_space UNIV) = 1"
+  then show *: "(\<integral>\<^sup>+ x. ennreal (pmf p x) \<partial>count_space UNIV) = 1"
     by (simp add: measure_pmf_eq_density nn_integral_density pmf_nonneg del: nn_integral_const)
 
   show "embed_pmf (pmf p) = p"
@@ -683,7 +704,7 @@
 
 end
 
-lemma nn_integral_measure_pmf: "(\<integral>\<^sup>+ x. f x \<partial>measure_pmf p) = \<integral>\<^sup>+ x. ereal (pmf p x) * f x \<partial>count_space UNIV"
+lemma nn_integral_measure_pmf: "(\<integral>\<^sup>+ x. f x \<partial>measure_pmf p) = \<integral>\<^sup>+ x. ennreal (pmf p x) * f x \<partial>count_space UNIV"
 by(simp add: measure_pmf_eq_density nn_integral_density pmf_nonneg)
 
 locale pmf_as_function
@@ -745,31 +766,31 @@
 lemma pair_map_pmf1: "pair_pmf (map_pmf f A) B = map_pmf (apfst f) (pair_pmf A B)"
 proof (safe intro!: pmf_eqI)
   fix a :: "'a" and b :: "'b"
-  have [simp]: "\<And>c d. indicator (apfst f -` {(a, b)}) (c, d) = indicator (f -` {a}) c * (indicator {b} d::ereal)"
+  have [simp]: "\<And>c d. indicator (apfst f -` {(a, b)}) (c, d) = indicator (f -` {a}) c * (indicator {b} d::ennreal)"
     by (auto split: split_indicator)
 
-  have "ereal (pmf (pair_pmf (map_pmf f A) B) (a, b)) =
-         ereal (pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b))"
-    unfolding pmf_pair ereal_pmf_map
+  have "ennreal (pmf (pair_pmf (map_pmf f A) B) (a, b)) =
+         ennreal (pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b))"
+    unfolding pmf_pair ennreal_pmf_map
     by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_multc pmf_nonneg
-                  emeasure_map_pmf[symmetric] del: emeasure_map_pmf)
+                  emeasure_map_pmf[symmetric] ennreal_mult del: emeasure_map_pmf)
   then show "pmf (pair_pmf (map_pmf f A) B) (a, b) = pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b)"
-    by simp
+    by (simp add: pmf_nonneg)
 qed
 
 lemma pair_map_pmf2: "pair_pmf A (map_pmf f B) = map_pmf (apsnd f) (pair_pmf A B)"
 proof (safe intro!: pmf_eqI)
   fix a :: "'a" and b :: "'b"
-  have [simp]: "\<And>c d. indicator (apsnd f -` {(a, b)}) (c, d) = indicator {a} c * (indicator (f -` {b}) d::ereal)"
+  have [simp]: "\<And>c d. indicator (apsnd f -` {(a, b)}) (c, d) = indicator {a} c * (indicator (f -` {b}) d::ennreal)"
     by (auto split: split_indicator)
 
-  have "ereal (pmf (pair_pmf A (map_pmf f B)) (a, b)) =
-         ereal (pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b))"
-    unfolding pmf_pair ereal_pmf_map
+  have "ennreal (pmf (pair_pmf A (map_pmf f B)) (a, b)) =
+         ennreal (pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b))"
+    unfolding pmf_pair ennreal_pmf_map
     by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_cmult nn_integral_multc pmf_nonneg
-                  emeasure_map_pmf[symmetric] del: emeasure_map_pmf)
+                  emeasure_map_pmf[symmetric] ennreal_mult del: emeasure_map_pmf)
   then show "pmf (pair_pmf A (map_pmf f B)) (a, b) = pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b)"
-    by simp
+    by (simp add: pmf_nonneg)
 qed
 
 lemma map_pair: "map_pmf (\<lambda>(a, b). (f a, g b)) (pair_pmf A B) = pair_pmf (map_pmf f A) (map_pmf g B)"
@@ -794,7 +815,7 @@
   fix i
   assume x: "set_pmf p \<subseteq> {x}"
   hence *: "set_pmf p = {x}" using set_pmf_not_empty[of p] by auto
-  have "ereal (pmf p x) = \<integral>\<^sup>+ i. indicator {x} i \<partial>p" by(simp add: emeasure_pmf_single)
+  have "ennreal (pmf p x) = \<integral>\<^sup>+ i. indicator {x} i \<partial>p" by(simp add: emeasure_pmf_single)
   also have "\<dots> = \<integral>\<^sup>+ i. 1 \<partial>p" by(rule nn_integral_cong_AE)(simp add: AE_measure_pmf_iff * )
   also have "\<dots> = 1" by simp
   finally show "pmf p i = pmf (return_pmf x) i" using x
@@ -817,11 +838,14 @@
   show ?lhs
   proof(rule pmf_eqI)
     fix i
-    have "ereal (pmf (bind_pmf p f) i) = \<integral>\<^sup>+ y. ereal (pmf (f y) i) \<partial>p" by(simp add: ereal_pmf_bind)
-    also have "\<dots> = \<integral>\<^sup>+ y. ereal (pmf (return_pmf x) i) \<partial>p"
+    have "ennreal (pmf (bind_pmf p f) i) = \<integral>\<^sup>+ y. ennreal (pmf (f y) i) \<partial>p"
+      by (simp add: ennreal_pmf_bind)
+    also have "\<dots> = \<integral>\<^sup>+ y. ennreal (pmf (return_pmf x) i) \<partial>p"
       by(rule nn_integral_cong_AE)(simp add: AE_measure_pmf_iff * )
-    also have "\<dots> = ereal (pmf (return_pmf x) i)" by simp
-    finally show "pmf (bind_pmf p f) i = pmf (return_pmf x) i" by simp
+    also have "\<dots> = ennreal (pmf (return_pmf x) i)"
+      by simp
+    finally show "pmf (bind_pmf p f) i = pmf (return_pmf x) i"
+      by (simp add: pmf_nonneg)
   qed
 qed
 
@@ -860,7 +884,7 @@
 qed
 
 lemma measure_measure_pmf_not_zero: "measure (measure_pmf p) s \<noteq> 0"
-  using emeasure_measure_pmf_not_zero unfolding measure_pmf.emeasure_eq_measure by simp
+  using emeasure_measure_pmf_not_zero by (simp add: measure_pmf.emeasure_eq_measure measure_nonneg)
 
 lift_definition cond_pmf :: "'a pmf" is
   "uniform_measure (measure_pmf p) s"
@@ -869,7 +893,7 @@
     by (intro prob_space_uniform_measure) (auto simp: emeasure_measure_pmf_not_zero)
   show "AE x in uniform_measure (measure_pmf p) s. measure (uniform_measure (measure_pmf p) s) {x} \<noteq> 0"
     by (simp add: emeasure_measure_pmf_not_zero measure_measure_pmf_not_zero AE_uniform_measure
-                  AE_measure_pmf_iff set_pmf.rep_eq)
+                  AE_measure_pmf_iff set_pmf.rep_eq less_top[symmetric])
 qed simp
 
 lemma pmf_cond: "pmf cond_pmf x = (if x \<in> s then pmf p x / measure p s else 0)"
@@ -887,20 +911,20 @@
   have *: "set_pmf (map_pmf f p) \<inter> s \<noteq> {}"
     using assms by auto
   { fix x
-    have "ereal (pmf (map_pmf f (cond_pmf p (f -` s))) x) =
+    have "ennreal (pmf (map_pmf f (cond_pmf p (f -` s))) x) =
       emeasure p (f -` s \<inter> f -` {x}) / emeasure p (f -` s)"
-      unfolding ereal_pmf_map cond_pmf.rep_eq[OF assms] by (simp add: nn_integral_uniform_measure)
+      unfolding ennreal_pmf_map cond_pmf.rep_eq[OF assms] by (simp add: nn_integral_uniform_measure)
     also have "f -` s \<inter> f -` {x} = (if x \<in> s then f -` {x} else {})"
       by auto
     also have "emeasure p (if x \<in> s then f -` {x} else {}) / emeasure p (f -` s) =
-      ereal (pmf (cond_pmf (map_pmf f p) s) x)"
+      ennreal (pmf (cond_pmf (map_pmf f p) s) x)"
       using measure_measure_pmf_not_zero[OF *]
-      by (simp add: pmf_cond[OF *] ereal_divide' ereal_pmf_map measure_pmf.emeasure_eq_measure[symmetric]
-               del: ereal_divide)
-    finally have "ereal (pmf (cond_pmf (map_pmf f p) s) x) = ereal (pmf (map_pmf f (cond_pmf p (f -` s))) x)"
+      by (simp add: pmf_cond[OF *] ennreal_pmf_map measure_pmf.emeasure_eq_measure
+                    divide_ennreal pmf_nonneg measure_nonneg zero_less_measure_iff pmf_map)
+    finally have "ennreal (pmf (cond_pmf (map_pmf f p) s) x) = ennreal (pmf (map_pmf f (cond_pmf p (f -` s))) x)"
       by simp }
   then show ?thesis
-    by (intro pmf_eqI) simp
+    by (intro pmf_eqI) (simp add: pmf_nonneg)
 qed
 
 lemma bind_cond_pmf_cancel:
@@ -910,16 +934,18 @@
   shows "bind_pmf p (\<lambda>x. cond_pmf q {y. R x y}) = q"
 proof (rule pmf_eqI)
   fix i
-  have "ereal (pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i) =
-    (\<integral>\<^sup>+x. ereal (pmf q i / measure p {x. R x i}) * ereal (indicator {x. R x i} x) \<partial>p)"
-    by (auto simp add: ereal_pmf_bind AE_measure_pmf_iff pmf_cond pmf_eq_0_set_pmf intro!: nn_integral_cong_AE)
+  have "ennreal (pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i) =
+    (\<integral>\<^sup>+x. ennreal (pmf q i / measure p {x. R x i}) * ennreal (indicator {x. R x i} x) \<partial>p)"
+    by (auto simp add: ennreal_pmf_bind AE_measure_pmf_iff pmf_cond pmf_eq_0_set_pmf pmf_nonneg measure_nonneg
+             intro!: nn_integral_cong_AE)
   also have "\<dots> = (pmf q i * measure p {x. R x i}) / measure p {x. R x i}"
-    by (simp add: pmf_nonneg measure_nonneg zero_ereal_def[symmetric] ereal_indicator
-                  nn_integral_cmult measure_pmf.emeasure_eq_measure)
+    by (simp add: pmf_nonneg measure_nonneg zero_ennreal_def[symmetric] ennreal_indicator
+                  nn_integral_cmult measure_pmf.emeasure_eq_measure ennreal_mult[symmetric])
   also have "\<dots> = pmf q i"
-    by (cases "pmf q i = 0") (simp_all add: pmf_eq_0_set_pmf measure_measure_pmf_not_zero)
+    by (cases "pmf q i = 0")
+       (simp_all add: pmf_eq_0_set_pmf measure_measure_pmf_not_zero pmf_nonneg)
   finally show "pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i = pmf q i"
-    by simp
+    by (simp add: pmf_nonneg)
 qed
 
 subsection \<open> Relator \<close>
@@ -1277,8 +1303,8 @@
 lemma pmf_join: "pmf (join_pmf N) i = (\<integral>M. pmf M i \<partial>measure_pmf N)"
   unfolding join_pmf_def pmf_bind ..
 
-lemma ereal_pmf_join: "ereal (pmf (join_pmf N) i) = (\<integral>\<^sup>+M. pmf M i \<partial>measure_pmf N)"
-  unfolding join_pmf_def ereal_pmf_bind ..
+lemma ennreal_pmf_join: "ennreal (pmf (join_pmf N) i) = (\<integral>\<^sup>+M. pmf M i \<partial>measure_pmf N)"
+  unfolding join_pmf_def ennreal_pmf_bind ..
 
 lemma set_pmf_join_pmf[simp]: "set_pmf (join_pmf f) = (\<Union>p\<in>set_pmf f. set_pmf p)"
   by (simp add: join_pmf_def)
@@ -1430,7 +1456,7 @@
 lemma pmf_bernoulli_False[simp]: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> pmf (bernoulli_pmf p) False = 1 - p"
   by transfer simp
 
-lemma set_pmf_bernoulli: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (bernoulli_pmf p) = UNIV"
+lemma set_pmf_bernoulli[simp]: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (bernoulli_pmf p) = UNIV"
   by (auto simp add: set_pmf_iff UNIV_bool)
 
 lemma nn_integral_bernoulli_pmf[simp]:
@@ -1448,7 +1474,10 @@
 by(cases x) simp_all
 
 lemma measure_pmf_bernoulli_half: "measure_pmf (bernoulli_pmf (1 / 2)) = uniform_count_measure UNIV"
-by(rule measure_eqI)(simp_all add: nn_integral_pmf[symmetric] emeasure_uniform_count_measure nn_integral_count_space_finite sets_uniform_count_measure)
+  by (rule measure_eqI)
+     (simp_all add: nn_integral_pmf[symmetric] emeasure_uniform_count_measure ennreal_divide_numeral[symmetric]
+                    nn_integral_count_space_finite sets_uniform_count_measure divide_ennreal_def mult_ac
+                    ennreal_of_nat_eq_real_of_nat)
 
 subsubsection \<open> Geometric Distribution \<close>
 
@@ -1458,9 +1487,9 @@
 
 lift_definition geometric_pmf :: "nat pmf" is "\<lambda>n. (1 - p)^n * p"
 proof
-  have "(\<Sum>i. ereal (p * (1 - p) ^ i)) = ereal (p * (1 / (1 - (1 - p))))"
-    by (intro sums_suminf_ereal sums_mult geometric_sums) auto
-  then show "(\<integral>\<^sup>+ x. ereal ((1 - p)^x * p) \<partial>count_space UNIV) = 1"
+  have "(\<Sum>i. ennreal (p * (1 - p) ^ i)) = ennreal (p * (1 / (1 - (1 - p))))"
+    by (intro suminf_ennreal_eq sums_mult geometric_sums) auto
+  then show "(\<integral>\<^sup>+ x. ennreal ((1 - p)^x * p) \<partial>count_space UNIV) = 1"
     by (simp add: nn_integral_count_space_nat field_simps)
 qed simp
 
@@ -1480,7 +1509,7 @@
 
 lift_definition pmf_of_multiset :: "'a pmf" is "\<lambda>x. count M x / size M"
 proof
-  show "(\<integral>\<^sup>+ x. ereal (real (count M x) / real (size M)) \<partial>count_space UNIV) = 1"
+  show "(\<integral>\<^sup>+ x. ennreal (real (count M x) / real (size M)) \<partial>count_space UNIV) = 1"
     using M_not_empty
     by (simp add: zero_less_divide_iff nn_integral_count_space nonempty_has_size
                   setsum_divide_distrib[symmetric])
@@ -1503,8 +1532,10 @@
 
 lift_definition pmf_of_set :: "'a pmf" is "\<lambda>x. indicator S x / card S"
 proof
-  show "(\<integral>\<^sup>+ x. ereal (indicator S x / real (card S)) \<partial>count_space UNIV) = 1"
-    using S_not_empty S_finite by (subst nn_integral_count_space'[of S]) auto
+  show "(\<integral>\<^sup>+ x. ennreal (indicator S x / real (card S)) \<partial>count_space UNIV) = 1"
+    using S_not_empty S_finite
+    by (subst nn_integral_count_space'[of S])
+       (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_mult[symmetric])
 qed simp
 
 lemma pmf_of_set[simp]: "pmf pmf_of_set x = indicator S x / card S"
@@ -1516,38 +1547,22 @@
 lemma emeasure_pmf_of_set_space[simp]: "emeasure pmf_of_set S = 1"
   by (rule measure_pmf.emeasure_eq_1_AE) (auto simp: AE_measure_pmf_iff)
 
-lemma nn_integral_pmf_of_set':
-  "(\<And>x. x \<in> S \<Longrightarrow> f x \<ge> 0) \<Longrightarrow> nn_integral (measure_pmf pmf_of_set) f = setsum f S / card S"
-apply(subst nn_integral_measure_pmf_finite, simp_all add: S_finite)
-apply(simp add: setsum_ereal_left_distrib[symmetric])
-apply(subst ereal_divide', simp add: S_not_empty S_finite)
-apply(simp add: ereal_times_divide_eq one_ereal_def[symmetric])
-done
+lemma nn_integral_pmf_of_set: "nn_integral (measure_pmf pmf_of_set) f = setsum f S / card S"
+  by (subst nn_integral_measure_pmf_finite)
+     (simp_all add: setsum_left_distrib[symmetric] card_gt_0_iff S_not_empty S_finite divide_ennreal_def
+                divide_ennreal[symmetric] ennreal_of_nat_eq_real_of_nat[symmetric] ennreal_times_divide)
 
-lemma nn_integral_pmf_of_set:
-  "nn_integral (measure_pmf pmf_of_set) f = setsum (max 0 \<circ> f) S / card S"
-apply(subst nn_integral_max_0[symmetric])
-apply(subst nn_integral_pmf_of_set')
-apply simp_all
-done
+lemma integral_pmf_of_set: "integral\<^sup>L (measure_pmf pmf_of_set) f = setsum f S / card S"
+  by (subst integral_measure_pmf[of S]) (auto simp: S_finite setsum_divide_distrib)
 
-lemma integral_pmf_of_set:
-  "integral\<^sup>L (measure_pmf pmf_of_set) f = setsum f S / card S"
-apply(simp add: real_lebesgue_integral_def integrable_measure_pmf_finite nn_integral_pmf_of_set S_finite)
-apply(subst real_of_ereal_minus')
- apply(simp add: ereal_max_0 S_finite del: ereal_max)
-apply(simp add: ereal_max_0 S_finite S_not_empty del: ereal_max)
-apply(simp add: field_simps S_finite S_not_empty)
-apply(subst setsum.distrib[symmetric])
-apply(rule setsum.cong; simp_all)
-done
+lemma emeasure_pmf_of_set: "emeasure (measure_pmf pmf_of_set) A = card (S \<inter> A) / card S"
+  by (subst nn_integral_indicator[symmetric], simp)
+     (simp add: S_finite S_not_empty card_gt_0_iff indicator_def setsum.If_cases divide_ennreal
+                ennreal_of_nat_eq_real_of_nat nn_integral_pmf_of_set)
 
-lemma emeasure_pmf_of_set:
-  "emeasure (measure_pmf pmf_of_set) A = card (S \<inter> A) / card S"
-apply(subst nn_integral_indicator[symmetric], simp)
-apply(subst nn_integral_pmf_of_set)
-apply(simp add: o_def max_def ereal_indicator[symmetric] S_not_empty S_finite real_of_nat_indicator[symmetric] of_nat_setsum[symmetric] setsum_indicator_eq_card del: of_nat_setsum)
-done
+lemma measure_pmf_of_set: "measure (measure_pmf pmf_of_set) A = card (S \<inter> A) / card S"
+  using emeasure_pmf_of_set[OF assms, of A]
+  by (simp add: measure_nonneg measure_pmf.emeasure_eq_measure)
 
 end
 
@@ -1574,15 +1589,7 @@
 qed
 
 lemma bernoulli_pmf_half_conv_pmf_of_set: "bernoulli_pmf (1 / 2) = pmf_of_set UNIV"
-by(rule pmf_eqI) simp_all
-
-
-
-lemma measure_pmf_of_set:
-  assumes "S \<noteq> {}" "finite S"
-  shows "measure (measure_pmf (pmf_of_set S)) A = card (S \<inter> A) / card S"
-using emeasure_pmf_of_set[OF assms, of A]
-unfolding measure_pmf.emeasure_eq_measure by simp
+  by (rule pmf_eqI) simp_all
 
 subsubsection \<open> Poisson Distribution \<close>
 
@@ -1596,14 +1603,14 @@
     by (simp add: field_simps divide_inverse [symmetric])
   have "(\<integral>\<^sup>+(x::nat). rate ^ x / fact x * exp (-rate) \<partial>count_space UNIV) =
           exp (-rate) * (\<integral>\<^sup>+(x::nat). rate ^ x / fact x \<partial>count_space UNIV)"
-    by (simp add: field_simps nn_integral_cmult[symmetric])
+    by (simp add: field_simps nn_integral_cmult[symmetric] ennreal_mult'[symmetric])
   also from rate_pos have "(\<integral>\<^sup>+(x::nat). rate ^ x / fact x \<partial>count_space UNIV) = (\<Sum>x. rate ^ x / fact x)"
-    by (simp_all add: nn_integral_count_space_nat suminf_ereal summable suminf_ereal_finite)
+    by (simp_all add: nn_integral_count_space_nat suminf_ennreal summable ennreal_suminf_neq_top)
   also have "... = exp rate" unfolding exp_def
     by (simp add: field_simps divide_inverse [symmetric])
-  also have "ereal (exp (-rate)) * ereal (exp rate) = 1"
-    by (simp add: mult_exp_exp)
-  finally show "(\<integral>\<^sup>+ x. ereal (rate ^ x / (fact x) * exp (- rate)) \<partial>count_space UNIV) = 1" .
+  also have "ennreal (exp (-rate)) * ennreal (exp rate) = 1"
+    by (simp add: mult_exp_exp ennreal_mult[symmetric])
+  finally show "(\<integral>\<^sup>+ x. ennreal (rate ^ x / (fact x) * exp (- rate)) \<partial>count_space UNIV) = 1" .
 qed (simp add: rate_pos[THEN less_imp_le])
 
 lemma pmf_poisson[simp]: "pmf poisson_pmf k = rate ^ k / fact k * exp (-rate)"
@@ -1622,12 +1629,12 @@
 
 lift_definition binomial_pmf :: "nat pmf" is "\<lambda>k. (n choose k) * p^k * (1 - p)^(n - k)"
 proof
-  have "(\<integral>\<^sup>+k. ereal (real (n choose k) * p ^ k * (1 - p) ^ (n - k)) \<partial>count_space UNIV) =
-    ereal (\<Sum>k\<le>n. real (n choose k) * p ^ k * (1 - p) ^ (n - k))"
+  have "(\<integral>\<^sup>+k. ennreal (real (n choose k) * p ^ k * (1 - p) ^ (n - k)) \<partial>count_space UNIV) =
+    ennreal (\<Sum>k\<le>n. real (n choose k) * p ^ k * (1 - p) ^ (n - k))"
     using p_le_1 p_nonneg by (subst nn_integral_count_space') auto
   also have "(\<Sum>k\<le>n. real (n choose k) * p ^ k * (1 - p) ^ (n - k)) = (p + (1 - p)) ^ n"
     by (subst binomial_ring) (simp add: atLeast0AtMost)
-  finally show "(\<integral>\<^sup>+ x. ereal (real (n choose x) * p ^ x * (1 - p) ^ (n - x)) \<partial>count_space UNIV) = 1"
+  finally show "(\<integral>\<^sup>+ x. ennreal (real (n choose x) * p ^ x * (1 - p) ^ (n - x)) \<partial>count_space UNIV) = 1"
     by simp
 qed (insert p_nonneg p_le_1, simp)
 
--- a/src/HOL/Probability/Probability_Measure.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Probability_Measure.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -60,7 +60,7 @@
 proof -
   interpret finite_measure M
   proof
-    show "emeasure M (space M) \<noteq> \<infinity>" using * by simp 
+    show "emeasure M (space M) \<noteq> \<infinity>" using * by simp
   qed
   show "prob_space M" by standard fact
 qed
@@ -96,7 +96,7 @@
 qed
 
 lemma (in prob_space) prob_space: "prob (space M) = 1"
-  using emeasure_space_1 unfolding measure_def by (simp add: one_ereal_def)
+  using emeasure_space_1 unfolding measure_def by (simp add: one_ennreal.rep_eq)
 
 lemma (in prob_space) prob_le_1[simp, intro]: "prob A \<le> 1"
   using bounded_measure[of A] by (simp add: prob_space)
@@ -104,6 +104,24 @@
 lemma (in prob_space) not_empty: "space M \<noteq> {}"
   using prob_space by auto
 
+lemma (in prob_space) emeasure_eq_1_AE:
+  "S \<in> sets M \<Longrightarrow> AE x in M. x \<in> S \<Longrightarrow> emeasure M S = 1"
+  by (subst emeasure_eq_AE[where B="space M"]) (auto simp: emeasure_space_1)
+
+lemma (in prob_space) emeasure_le_1: "emeasure M S \<le> 1"
+  unfolding ennreal_1[symmetric] emeasure_eq_measure by (subst ennreal_le_iff) auto
+
+lemma (in prob_space) AE_iff_emeasure_eq_1:
+  assumes [measurable]: "Measurable.pred M P"
+  shows "(AE x in M. P x) \<longleftrightarrow> emeasure M {x\<in>space M. P x} = 1"
+proof -
+  have *: "{x \<in> space M. \<not> P x} = space M - {x\<in>space M. P x}"
+    by auto
+  show ?thesis
+    by (auto simp add: ennreal_minus_eq_0 * emeasure_compl emeasure_space_1 AE_iff_measurable[OF _ refl]
+             intro: antisym emeasure_le_1)
+qed
+
 lemma (in prob_space) measure_le_1: "emeasure M X \<le> 1"
   using emeasure_space[of M X] by (simp add: emeasure_space_1)
 
@@ -126,26 +144,12 @@
   using finite_measure_compl[OF A] by (simp add: prob_space)
 
 lemma (in prob_space) AE_in_set_eq_1:
-  assumes "A \<in> events" shows "(AE x in M. x \<in> A) \<longleftrightarrow> prob A = 1"
-proof
-  assume ae: "AE x in M. x \<in> A"
-  have "{x \<in> space M. x \<in> A} = A" "{x \<in> space M. x \<notin> A} = space M - A"
-    using \<open>A \<in> events\<close>[THEN sets.sets_into_space] by auto
-  with AE_E2[OF ae] \<open>A \<in> events\<close> have "1 - emeasure M A = 0"
-    by (simp add: emeasure_compl emeasure_space_1)
-  then show "prob A = 1"
-    using \<open>A \<in> events\<close> by (simp add: emeasure_eq_measure one_ereal_def)
-next
-  assume prob: "prob A = 1"
-  show "AE x in M. x \<in> A"
-  proof (rule AE_I)
-    show "{x \<in> space M. x \<notin> A} \<subseteq> space M - A" by auto
-    show "emeasure M (space M - A) = 0"
-      using \<open>A \<in> events\<close> prob
-      by (simp add: prob_compl emeasure_space_1 emeasure_eq_measure one_ereal_def)
-    show "space M - A \<in> events"
-      using \<open>A \<in> events\<close> by auto
-  qed
+  assumes A[measurable]: "A \<in> events" shows "(AE x in M. x \<in> A) \<longleftrightarrow> prob A = 1"
+proof -
+  have *: "{x\<in>space M. x \<in> A} = A"
+    using A[THEN sets.sets_into_space] by auto
+  show ?thesis
+    by (subst AE_iff_emeasure_eq_1) (auto simp: emeasure_eq_measure *)
 qed
 
 lemma (in prob_space) AE_False: "(AE x in M. False) \<longleftrightarrow> False"
@@ -178,10 +182,6 @@
   then show False by auto
 qed
 
-lemma (in prob_space) emeasure_eq_1_AE:
-  "S \<in> sets M \<Longrightarrow> AE x in M. x \<in> S \<Longrightarrow> emeasure M S = 1"
-  by (subst emeasure_eq_AE[where B="space M"]) (auto simp: emeasure_space_1)
-
 lemma (in prob_space) integral_ge_const:
   fixes c :: real
   shows "integrable M f \<Longrightarrow> (AE x in M. c \<le> f x) \<Longrightarrow> c \<le> (\<integral>x. f x \<partial>M)"
@@ -195,7 +195,7 @@
 lemma (in prob_space) nn_integral_ge_const:
   "(AE x in M. c \<le> f x) \<Longrightarrow> c \<le> (\<integral>\<^sup>+x. f x \<partial>M)"
   using nn_integral_mono_AE[of "\<lambda>x. c" f M] emeasure_space_1
-  by (simp add: nn_integral_const_If split: if_split_asm)
+  by (simp split: if_split_asm)
 
 lemma (in prob_space) expectation_less:
   fixes X :: "_ \<Rightarrow> real"
@@ -297,24 +297,24 @@
       end
 
     fun unnest_tuples
-      (Const (@{syntax_const "_pattern"}, _) $ 
+      (Const (@{syntax_const "_pattern"}, _) $
         t1 $
         (t as (Const (@{syntax_const "_pattern"}, _) $ _ $ _)))
       = let
         val (_ $ t2 $ t3) = unnest_tuples t
       in
-        Syntax.const @{syntax_const "_pattern"} $ 
+        Syntax.const @{syntax_const "_pattern"} $
           unnest_tuples t1 $
           (Syntax.const @{syntax_const "_patterns"} $ t2 $ t3)
       end
     | unnest_tuples pat = pat
 
-    fun tr' [sig_alg, Const (@{const_syntax Collect}, _) $ t] = 
+    fun tr' [sig_alg, Const (@{const_syntax Collect}, _) $ t] =
       let
         val bound_dummyT = Const (@{syntax_const "_bound"}, dummyT)
 
         fun go pattern elem
-          (Const (@{const_syntax "conj"}, _) $ 
+          (Const (@{const_syntax "conj"}, _) $
             (Const (@{const_syntax Set.member}, _) $ elem' $ (Const (@{const_syntax space}, _) $ sig_alg')) $
             u)
           = let
@@ -331,7 +331,7 @@
               go ((x, 0) :: pattern) (bound_dummyT $ tx :: elem) t
             end
         | go pattern elem (Const (@{const_syntax case_prod}, _) $ t) =
-            go 
+            go
               ((Syntax.const @{syntax_const "_pattern"}, 2) :: pattern)
               (Syntax.const @{const_syntax Pair} :: elem)
               t
@@ -359,7 +359,7 @@
   from ae[THEN AE_E] guess N .
   then show thesis
     by (intro that[of "space M - N"])
-       (auto simp: prob_compl prob_space emeasure_eq_measure)
+       (auto simp: prob_compl prob_space emeasure_eq_measure measure_nonneg)
 qed
 
 lemma (in prob_space) prob_neg: "{x\<in>space M. P x} \<in> events \<Longrightarrow> \<P>(x in M. \<not> P x) = 1 - \<P>(x in M. P x)"
@@ -380,7 +380,7 @@
 
 lemma (in prob_space) prob_Collect_eq_0:
   "{x \<in> space M. P x} \<in> sets M \<Longrightarrow> \<P>(x in M. P x) = 0 \<longleftrightarrow> (AE x in M. \<not> P x)"
-  using AE_iff_measurable[OF _ refl, of M "\<lambda>x. \<not> P x"] by (simp add: emeasure_eq_measure)
+  using AE_iff_measurable[OF _ refl, of M "\<lambda>x. \<not> P x"] by (simp add: emeasure_eq_measure measure_nonneg)
 
 lemma (in prob_space) prob_Collect_eq_1:
   "{x \<in> space M. P x} \<in> sets M \<Longrightarrow> \<P>(x in M. P x) = 1 \<longleftrightarrow> (AE x in M. P x)"
@@ -389,7 +389,7 @@
 lemma (in prob_space) prob_eq_0:
   "A \<in> sets M \<Longrightarrow> prob A = 0 \<longleftrightarrow> (AE x in M. x \<notin> A)"
   using AE_iff_measurable[OF _ refl, of M "\<lambda>x. x \<notin> A"]
-  by (auto simp add: emeasure_eq_measure Int_def[symmetric])
+  by (auto simp add: emeasure_eq_measure Int_def[symmetric] measure_nonneg)
 
 lemma (in prob_space) prob_eq_1:
   "A \<in> sets M \<Longrightarrow> prob A = 1 \<longleftrightarrow> (AE x in M. x \<in> A)"
@@ -445,27 +445,27 @@
 qed
 
 lemma (in prob_space) prob_EX_countable:
-  assumes sets: "\<And>i. i \<in> I \<Longrightarrow> {x\<in>space M. P i x} \<in> sets M" and I: "countable I" 
+  assumes sets: "\<And>i. i \<in> I \<Longrightarrow> {x\<in>space M. P i x} \<in> sets M" and I: "countable I"
   assumes disj: "AE x in M. \<forall>i\<in>I. \<forall>j\<in>I. P i x \<longrightarrow> P j x \<longrightarrow> i = j"
   shows "\<P>(x in M. \<exists>i\<in>I. P i x) = (\<integral>\<^sup>+i. \<P>(x in M. P i x) \<partial>count_space I)"
 proof -
   let ?N= "\<lambda>x. \<exists>!i\<in>I. P i x"
-  have "ereal (\<P>(x in M. \<exists>i\<in>I. P i x)) = \<P>(x in M. (\<exists>i\<in>I. P i x \<and> ?N x))"
-    unfolding ereal.inject
+  have "ennreal (\<P>(x in M. \<exists>i\<in>I. P i x)) = \<P>(x in M. (\<exists>i\<in>I. P i x \<and> ?N x))"
+    unfolding ennreal_inj[OF measure_nonneg measure_nonneg]
   proof (rule prob_eq_AE)
     show "AE x in M. (\<exists>i\<in>I. P i x) = (\<exists>i\<in>I. P i x \<and> ?N x)"
       using disj by eventually_elim blast
   qed (auto intro!: sets.sets_Collect_countable_Ex' sets.sets_Collect_conj sets.sets_Collect_countable_Ex1' I sets)+
   also have "\<P>(x in M. (\<exists>i\<in>I. P i x \<and> ?N x)) = emeasure M (\<Union>i\<in>I. {x\<in>space M. P i x \<and> ?N x})"
-    unfolding emeasure_eq_measure by (auto intro!: arg_cong[where f=prob])
+    unfolding emeasure_eq_measure by (auto intro!: arg_cong[where f=prob] simp: measure_nonneg)
   also have "\<dots> = (\<integral>\<^sup>+i. emeasure M {x\<in>space M. P i x \<and> ?N x} \<partial>count_space I)"
     by (rule emeasure_UN_countable)
        (auto intro!: sets.sets_Collect_countable_Ex' sets.sets_Collect_conj sets.sets_Collect_countable_Ex1' I sets
              simp: disjoint_family_on_def)
   also have "\<dots> = (\<integral>\<^sup>+i. \<P>(x in M. P i x) \<partial>count_space I)"
     unfolding emeasure_eq_measure using disj
-    by (intro nn_integral_cong ereal.inject[THEN iffD2] prob_eq_AE)
-       (auto intro!: sets.sets_Collect_countable_Ex' sets.sets_Collect_conj sets.sets_Collect_countable_Ex1' I sets)+
+    by (intro nn_integral_cong ennreal_inj[THEN iffD2] prob_eq_AE)
+       (auto intro!: sets.sets_Collect_countable_Ex' sets.sets_Collect_conj sets.sets_Collect_countable_Ex1' I sets measure_nonneg)+
   finally show ?thesis .
 qed
 
@@ -528,26 +528,29 @@
   assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets (M i)"
   shows "prob (\<Pi>\<^sub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))"
 proof -
-  have "ereal (measure (\<Pi>\<^sub>M i\<in>I. M i) (\<Pi>\<^sub>E i\<in>I. X i)) = emeasure (\<Pi>\<^sub>M i\<in>I. M i) (\<Pi>\<^sub>E i\<in>I. X i)"
+  have "ennreal (measure (\<Pi>\<^sub>M i\<in>I. M i) (\<Pi>\<^sub>E i\<in>I. X i)) = emeasure (\<Pi>\<^sub>M i\<in>I. M i) (\<Pi>\<^sub>E i\<in>I. X i)"
     using X by (simp add: emeasure_eq_measure)
   also have "\<dots> = (\<Prod>i\<in>I. emeasure (M i) (X i))"
     using measure_times X by simp
-  also have "\<dots> = ereal (\<Prod>i\<in>I. measure (M i) (X i))"
-    using X by (simp add: M.emeasure_eq_measure setprod_ereal)
-  finally show ?thesis by simp
+  also have "\<dots> = ennreal (\<Prod>i\<in>I. measure (M i) (X i))"
+    using X by (simp add: M.emeasure_eq_measure setprod_ennreal measure_nonneg)
+  finally show ?thesis by (simp add: measure_nonneg setprod_nonneg)
 qed
 
 subsection \<open>Distributions\<close>
 
-definition "distributed M N X f \<longleftrightarrow> distr M N X = density N f \<and> 
-  f \<in> borel_measurable N \<and> (AE x in N. 0 \<le> f x) \<and> X \<in> measurable M N"
+definition distributed :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> ennreal) \<Rightarrow> bool"
+where
+  "distributed M N X f \<longleftrightarrow>
+  distr M N X = density N f \<and> f \<in> borel_measurable N \<and> X \<in> measurable M N"
+
+term distributed
 
 lemma
   assumes "distributed M N X f"
   shows distributed_distr_eq_density: "distr M N X = density N f"
     and distributed_measurable: "X \<in> measurable M N"
     and distributed_borel_measurable: "f \<in> borel_measurable N"
-    and distributed_AE: "(AE x in N. 0 \<le> f x)"
   using assms by (simp_all add: distributed_def)
 
 lemma
@@ -559,14 +562,14 @@
   using distributed_measurable[OF D] distributed_borel_measurable[OF D]
   by simp_all
 
-lemma
-  shows distributed_real_measurable: "distributed M N X (\<lambda>x. ereal (f x)) \<Longrightarrow> f \<in> borel_measurable N"
-    and distributed_real_AE: "distributed M N X (\<lambda>x. ereal (f x)) \<Longrightarrow> (AE x in N. 0 \<le> f x)"
-  by (simp_all add: distributed_def borel_measurable_ereal_iff)
+lemma distributed_real_measurable:
+  "(\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> distributed M N X (\<lambda>x. ennreal (f x)) \<Longrightarrow> f \<in> borel_measurable N"
+  by (simp_all add: distributed_def)
 
 lemma distributed_real_measurable':
-  "distributed M N X (\<lambda>x. ereal (f x)) \<Longrightarrow> h \<in> measurable L N \<Longrightarrow> (\<lambda>x. f (h x)) \<in> borel_measurable L"
-  by simp
+  "(\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> distributed M N X (\<lambda>x. ennreal (f x)) \<Longrightarrow>
+    h \<in> measurable L N \<Longrightarrow> (\<lambda>x. f (h x)) \<in> borel_measurable L"
+  using distributed_real_measurable[measurable] by simp
 
 lemma joint_distributed_measurable1:
   "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) f \<Longrightarrow> h1 \<in> measurable N M \<Longrightarrow> (\<lambda>x. X (h1 x)) \<in> measurable N S"
@@ -587,7 +590,7 @@
   also have "\<dots> = (\<integral>\<^sup>+x. P a * indicator {a} x \<partial>count_space A)"
     using X a by (auto simp add: emeasure_density distributed_def indicator_def intro!: nn_integral_cong)
   also have "\<dots> = P a"
-    using X a by (subst nn_integral_cmult_indicator) (auto simp: distributed_def one_ereal_def[symmetric] AE_count_space)
+    using X a by (subst nn_integral_cmult_indicator) (auto simp: distributed_def one_ennreal_def[symmetric] AE_count_space)
   finally show ?thesis ..
 qed
 
@@ -596,6 +599,26 @@
     distributed M N X f \<longleftrightarrow> distributed M N X g"
   by (auto simp: distributed_def intro!: density_cong)
 
+lemma (in prob_space) distributed_imp_emeasure_nonzero:
+  assumes X: "distributed M MX X Px"
+  shows "emeasure MX {x \<in> space MX. Px x \<noteq> 0} \<noteq> 0"
+proof
+  note Px = distributed_borel_measurable[OF X]
+  interpret X: prob_space "distr M MX X"
+    using distributed_measurable[OF X] by (rule prob_space_distr)
+
+  assume "emeasure MX {x \<in> space MX. Px x \<noteq> 0} = 0"
+  with Px have "AE x in MX. Px x = 0"
+    by (intro AE_I[OF subset_refl]) (auto simp: borel_measurable_ennreal_iff)
+  moreover
+  from X.emeasure_space_1 have "(\<integral>\<^sup>+x. Px x \<partial>MX) = 1"
+    unfolding distributed_distr_eq_density[OF X] using Px
+    by (subst (asm) emeasure_density)
+       (auto simp: borel_measurable_ennreal_iff intro!: integral_cong cong: nn_integral_cong)
+  ultimately show False
+    by (simp add: nn_integral_cong_AE)
+qed
+
 lemma subdensity:
   assumes T: "T \<in> measurable P Q"
   assumes f: "distributed M P X f"
@@ -621,27 +644,27 @@
   assumes f: "distributed M P X f"
   assumes g: "distributed M Q Y g"
   assumes Y: "Y = T \<circ> X"
-  shows "AE x in P. g (T x) = 0 \<longrightarrow> f x = 0"
-  using subdensity[OF T, of M X "\<lambda>x. ereal (f x)" Y "\<lambda>x. ereal (g x)"] assms by auto
+  shows "(AE x in P. 0 \<le> g (T x)) \<Longrightarrow> (AE x in P. 0 \<le> f x) \<Longrightarrow> AE x in P. g (T x) = 0 \<longrightarrow> f x = 0"
+  using subdensity[OF T, of M X "\<lambda>x. ennreal (f x)" Y "\<lambda>x. ennreal (g x)"] assms
+  by auto
 
 lemma distributed_emeasure:
   "distributed M N X f \<Longrightarrow> A \<in> sets N \<Longrightarrow> emeasure M (X -` A \<inter> space M) = (\<integral>\<^sup>+x. f x * indicator A x \<partial>N)"
-  by (auto simp: distributed_AE
-                 distributed_distr_eq_density[symmetric] emeasure_density[symmetric] emeasure_distr)
+  by (auto simp: distributed_distr_eq_density[symmetric] emeasure_density[symmetric] emeasure_distr)
 
 lemma distributed_nn_integral:
   "distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<integral>\<^sup>+x. f x * g x \<partial>N) = (\<integral>\<^sup>+x. g (X x) \<partial>M)"
-  by (auto simp: distributed_AE
-                 distributed_distr_eq_density[symmetric] nn_integral_density[symmetric] nn_integral_distr)
+  by (auto simp: distributed_distr_eq_density[symmetric] nn_integral_density[symmetric] nn_integral_distr)
 
 lemma distributed_integral:
-  "distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<integral>x. f x * g x \<partial>N) = (\<integral>x. g (X x) \<partial>M)"
-  by (auto simp: distributed_real_AE
-                 distributed_distr_eq_density[symmetric] integral_real_density[symmetric] integral_distr)
-  
+  "distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> f x) \<Longrightarrow>
+    (\<integral>x. f x * g x \<partial>N) = (\<integral>x. g (X x) \<partial>M)"
+  supply distributed_real_measurable[measurable]
+  by (auto simp: distributed_distr_eq_density[symmetric] integral_real_density[symmetric] integral_distr)
+
 lemma distributed_transform_integral:
-  assumes Px: "distributed M N X Px"
-  assumes "distributed M P Y Py"
+  assumes Px: "distributed M N X Px" "\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> Px x"
+  assumes "distributed M P Y Py" "\<And>x. x \<in> space P \<Longrightarrow> 0 \<le> Py x"
   assumes Y: "Y = T \<circ> X" and T: "T \<in> measurable N P" and f: "f \<in> borel_measurable P"
   shows "(\<integral>x. Py x * f x \<partial>P) = (\<integral>x. Px x * f (T x) \<partial>N)"
 proof -
@@ -671,7 +694,7 @@
   assumes "sigma_finite_measure S" "sigma_finite_measure T"
   assumes X[measurable]: "X \<in> measurable M S" and Y[measurable]: "Y \<in> measurable M T"
   assumes [measurable]: "f \<in> borel_measurable (S \<Otimes>\<^sub>M T)" and f: "AE x in S \<Otimes>\<^sub>M T. 0 \<le> f x"
-  assumes eq: "\<And>A B. A \<in> sets S \<Longrightarrow> B \<in> sets T \<Longrightarrow> 
+  assumes eq: "\<And>A B. A \<in> sets S \<Longrightarrow> B \<in> sets T \<Longrightarrow>
     emeasure M {x \<in> space M. X x \<in> A \<and> Y x \<in> B} = (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. f (x, y) * indicator B y \<partial>T) * indicator A x \<partial>S)"
   shows "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) f"
   unfolding distributed_def
@@ -722,19 +745,13 @@
   interpret TS: pair_sigma_finite T S ..
 
   note Pxy[measurable]
-  show ?thesis 
+  show ?thesis
     apply (subst TS.distr_pair_swap)
     unfolding distributed_def
   proof safe
     let ?D = "distr (S \<Otimes>\<^sub>M T) (T \<Otimes>\<^sub>M S) (\<lambda>(x, y). (y, x))"
     show 1: "(\<lambda>(x, y). Pxy (y, x)) \<in> borel_measurable ?D"
       by auto
-    with Pxy
-    show "AE x in distr (S \<Otimes>\<^sub>M T) (T \<Otimes>\<^sub>M S) (\<lambda>(x, y). (y, x)). 0 \<le> (case x of (x, y) \<Rightarrow> Pxy (y, x))"
-      by (subst AE_distr_iff)
-         (auto dest!: distributed_AE
-               simp: measurable_split_conv split_beta
-               intro!: measurable_Pair)
     show 2: "random_variable (distr (S \<Otimes>\<^sub>M T) (T \<Otimes>\<^sub>M S) (\<lambda>(x, y). (y, x))) (\<lambda>x. (Y x, X x))"
       using Pxy by auto
     { fix A assume A: "A \<in> sets (T \<Otimes>\<^sub>M S)"
@@ -777,9 +794,6 @@
 
   interpret Pxy: prob_space "distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))"
     by (intro prob_space_distr) simp
-  have "(\<integral>\<^sup>+ x. max 0 (- Pxy x) \<partial>(S \<Otimes>\<^sub>M T)) = (\<integral>\<^sup>+ x. 0 \<partial>(S \<Otimes>\<^sub>M T))"
-    using Pxy
-    by (intro nn_integral_cong_AE) (auto simp: max_def dest: distributed_AE)
 
   show "distr M S X = density S Px"
   proof (rule measure_eqI)
@@ -793,24 +807,19 @@
       using A borel Pxy
       by (simp add: emeasure_density T.nn_integral_fst[symmetric])
     also have "\<dots> = \<integral>\<^sup>+ x. Px x * indicator A x \<partial>S"
-      apply (rule nn_integral_cong_AE)
-      using Pxy[THEN distributed_AE, THEN ST.AE_pair] AE_space
-    proof eventually_elim
-      fix x assume "x \<in> space S" "AE y in T. 0 \<le> Pxy (x, y)"
+    proof (rule nn_integral_cong)
+      fix x assume "x \<in> space S"
       moreover have eq: "\<And>y. y \<in> space T \<Longrightarrow> indicator (A \<times> space T) (x, y) = indicator A x"
         by (auto simp: indicator_def)
       ultimately have "(\<integral>\<^sup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T) = (\<integral>\<^sup>+ y. Pxy (x, y) \<partial>T) * indicator A x"
         by (simp add: eq nn_integral_multc cong: nn_integral_cong)
       also have "(\<integral>\<^sup>+ y. Pxy (x, y) \<partial>T) = Px x"
-        by (simp add: Px_def ereal_real nn_integral_nonneg)
+        by (simp add: Px_def)
       finally show "(\<integral>\<^sup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T) = Px x * indicator A x" .
     qed
     finally show "emeasure (distr M S X) A = emeasure (density S Px) A"
       using A borel Pxy by (simp add: emeasure_density)
   qed simp
-  
-  show "AE x in S. 0 \<le> Px x"
-    by (simp add: Px_def nn_integral_nonneg real_of_ereal_pos)
 qed
 
 lemma (in prob_space) distr_marginal2:
@@ -858,33 +867,24 @@
 
   show "distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) = density (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). Px x * Py y)"
     unfolding indep[symmetric] distributed_distr_eq_density[OF X] distributed_distr_eq_density[OF Y]
-    using distributed_borel_measurable[OF X] distributed_AE[OF X]
-    using distributed_borel_measurable[OF Y] distributed_AE[OF Y]
-    by (rule pair_measure_density[OF _ _ _ _ T sf_Y])
+    using distributed_borel_measurable[OF X]
+    using distributed_borel_measurable[OF Y]
+    by (rule pair_measure_density[OF _ _ T sf_Y])
 
   show "random_variable (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))" by auto
 
   show Pxy: "(\<lambda>(x, y). Px x * Py y) \<in> borel_measurable (S \<Otimes>\<^sub>M T)" by auto
-
-  show "AE x in S \<Otimes>\<^sub>M T. 0 \<le> (case x of (x, y) \<Rightarrow> Px x * Py y)"
-    apply (intro ST.AE_pair_measure borel_measurable_le Pxy borel_measurable_const)
-    using distributed_AE[OF X]
-    apply eventually_elim
-    using distributed_AE[OF Y]
-    apply eventually_elim
-    apply auto
-    done
 qed
 
 lemma distributed_integrable:
-  "distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow>
+  "distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> f x) \<Longrightarrow>
     integrable N (\<lambda>x. f x * g x) \<longleftrightarrow> integrable M (\<lambda>x. g (X x))"
-  by (auto simp: distributed_real_AE
-                    distributed_distr_eq_density[symmetric] integrable_real_density[symmetric] integrable_distr_eq)
-  
+  supply distributed_real_measurable[measurable]
+  by (auto simp: distributed_distr_eq_density[symmetric] integrable_real_density[symmetric] integrable_distr_eq)
+
 lemma distributed_transform_integrable:
-  assumes Px: "distributed M N X Px"
-  assumes "distributed M P Y Py"
+  assumes Px: "distributed M N X Px" "\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> Px x"
+  assumes "distributed M P Y Py" "\<And>x. x \<in> space P \<Longrightarrow> 0 \<le> Py x"
   assumes Y: "Y = (\<lambda>x. T (X x))" and T: "T \<in> measurable N P" and f: "f \<in> borel_measurable P"
   shows "integrable P (\<lambda>x. Py x * f x) \<longleftrightarrow> integrable N (\<lambda>x. Px x * f (T x))"
 proof -
@@ -899,17 +899,18 @@
 
 lemma distributed_integrable_var:
   fixes X :: "'a \<Rightarrow> real"
-  shows "distributed M lborel X (\<lambda>x. ereal (f x)) \<Longrightarrow> integrable lborel (\<lambda>x. f x * x) \<Longrightarrow> integrable M X"
+  shows "distributed M lborel X (\<lambda>x. ennreal (f x)) \<Longrightarrow> (\<And>x. 0 \<le> f x) \<Longrightarrow>
+    integrable lborel (\<lambda>x. f x * x) \<Longrightarrow> integrable M X"
   using distributed_integrable[of M lborel X f "\<lambda>x. x"] by simp
 
 lemma (in prob_space) distributed_variance:
   fixes f::"real \<Rightarrow> real"
-  assumes D: "distributed M lborel X f"
+  assumes D: "distributed M lborel X f" and [simp]: "\<And>x. 0 \<le> f x"
   shows "variance X = (\<integral>x. x\<^sup>2 * f (x + expectation X) \<partial>lborel)"
 proof (subst distributed_integral[OF D, symmetric])
   show "(\<integral> x. f x * (x - expectation X)\<^sup>2 \<partial>lborel) = (\<integral> x. x\<^sup>2 * f (x + expectation X) \<partial>lborel)"
     by (subst lborel_integral_real_affine[where c=1 and t="expectation X"])  (auto simp: ac_simps)
-qed simp
+qed simp_all
 
 lemma (in prob_space) variance_affine:
   fixes f::"real \<Rightarrow> real"
@@ -917,15 +918,20 @@
   assumes D[intro]: "distributed M lborel X f"
   assumes [simp]: "prob_space (density lborel f)"
   assumes I[simp]: "integrable M X"
-  assumes I2[simp]: "integrable M (\<lambda>x. (X x)\<^sup>2)" 
+  assumes I2[simp]: "integrable M (\<lambda>x. (X x)\<^sup>2)"
   shows "variance (\<lambda>x. a + b * X x) = b\<^sup>2 * variance X"
   by (subst variance_eq)
      (auto simp: power2_sum power_mult_distrib prob_space variance_eq right_diff_distrib)
 
 definition
-  "simple_distributed M X f \<longleftrightarrow> distributed M (count_space (X`space M)) X (\<lambda>x. ereal (f x)) \<and>
+  "simple_distributed M X f \<longleftrightarrow>
+    (\<forall>x. 0 \<le> f x) \<and>
+    distributed M (count_space (X`space M)) X (\<lambda>x. ennreal (f x)) \<and>
     finite (X`space M)"
 
+lemma simple_distributed_nonneg[dest]: "simple_distributed M X f \<Longrightarrow> 0 \<le> f x"
+  by (auto simp: simple_distributed_def)
+
 lemma simple_distributed:
   "simple_distributed M X Px \<Longrightarrow> distributed M (count_space (X`space M)) X Px"
   unfolding simple_distributed_def by auto
@@ -940,9 +946,7 @@
   shows "distributed M S X P'"
   unfolding distributed_def
 proof safe
-  show "(\<lambda>x. ereal (P' x)) \<in> borel_measurable S" unfolding S_def by simp
-  show "AE x in S. 0 \<le> ereal (P' x)"
-    using X by (auto simp: S_def P'_def simple_distributed_def intro!: measure_nonneg)
+  show "(\<lambda>x. ennreal (P' x)) \<in> borel_measurable S" unfolding S_def by simp
   show "distr M S X = density S P'"
   proof (rule measure_eqI_finite)
     show "sets (distr M S X) = Pow A" "sets (density S P') = Pow A"
@@ -954,11 +958,11 @@
       by (subst emeasure_distr)
          (auto simp add: S_def P'_def simple_functionD emeasure_eq_measure measurable_count_space_eq2
                intro!: arg_cong[where f=prob])
-    also have "\<dots> = (\<integral>\<^sup>+x. ereal (P' a) * indicator {a} x \<partial>S)"
+    also have "\<dots> = (\<integral>\<^sup>+x. ennreal (P' a) * indicator {a} x \<partial>S)"
       using A X a
       by (subst nn_integral_cmult_indicator)
          (auto simp: S_def P'_def simple_distributed_def simple_functionD measure_nonneg)
-    also have "\<dots> = (\<integral>\<^sup>+x. ereal (P' x) * indicator {a} x \<partial>S)"
+    also have "\<dots> = (\<integral>\<^sup>+x. ennreal (P' x) * indicator {a} x \<partial>S)"
       by (auto simp: indicator_def intro!: nn_integral_cong)
     also have "\<dots> = emeasure (density S P') {a}"
       using a A by (intro emeasure_density[symmetric]) (auto simp: S_def)
@@ -969,14 +973,16 @@
 qed
 
 lemma (in prob_space) simple_distributedI:
-  assumes X: "simple_function M X" "\<And>x. x \<in> X ` space M \<Longrightarrow> P x = measure M (X -` {x} \<inter> space M)"
+  assumes X: "simple_function M X"
+    "\<And>x. 0 \<le> P x"
+    "\<And>x. x \<in> X ` space M \<Longrightarrow> P x = measure M (X -` {x} \<inter> space M)"
   shows "simple_distributed M X P"
   unfolding simple_distributed_def
-proof
-  have "distributed M (count_space (X ` space M)) X (\<lambda>x. ereal (if x \<in> X`space M then P x else 0))"
+proof (safe intro!: X)
+  have "distributed M (count_space (X ` space M)) X (\<lambda>x. ennreal (if x \<in> X`space M then P x else 0))"
     (is "?A")
-    using simple_functionD[OF X(1)] by (intro distributed_simple_function_superset[OF X]) auto
-  also have "?A \<longleftrightarrow> distributed M (count_space (X ` space M)) X (\<lambda>x. ereal (P x))"
+    using simple_functionD[OF X(1)] by (intro distributed_simple_function_superset[OF X(1,3)]) auto
+  also have "?A \<longleftrightarrow> distributed M (count_space (X ` space M)) X (\<lambda>x. ennreal (P x))"
     by (rule distributed_cong_density) auto
   finally show "\<dots>" .
 qed (rule simple_functionD[OF X(1)])
@@ -1017,9 +1023,6 @@
   using distributed_count_space[of M "X`space M" X P a, symmetric]
   by (auto simp: simple_distributed_def measure_def)
 
-lemma simple_distributed_nonneg: "simple_distributed M X f \<Longrightarrow> x \<in> space M \<Longrightarrow> 0 \<le> f (X x)"
-  by (auto simp: simple_distributed_measure measure_nonneg)
-
 lemma (in prob_space) simple_distributed_joint:
   assumes X: "simple_distributed M (\<lambda>x. (X x, Y x)) Px"
   defines "S \<equiv> count_space (X`space M) \<Otimes>\<^sub>M count_space (Y`space M)"
@@ -1071,7 +1074,7 @@
   also have "\<dots> = prob (space M)"
     by (auto intro!: arg_cong[where f=prob])
   finally show ?thesis
-    using emeasure_space_1 by (simp add: emeasure_eq_measure one_ereal_def)
+    using emeasure_space_1 by (simp add: emeasure_eq_measure)
 qed
 
 lemma (in prob_space) distributed_marginal_eq_joint_simple:
@@ -1081,22 +1084,22 @@
   assumes y: "y \<in> Y`space M"
   shows "Py y = (\<Sum>x\<in>X`space M. if (x, y) \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy (x, y) else 0)"
 proof -
-  note Px = simple_distributedI[OF Px refl]
-  have *: "\<And>f A. setsum (\<lambda>x. max 0 (ereal (f x))) A = ereal (setsum (\<lambda>x. max 0 (f x)) A)"
-    by (simp add: setsum_ereal[symmetric] zero_ereal_def)
-  from distributed_marginal_eq_joint2[OF
-    sigma_finite_measure_count_space_finite
-    sigma_finite_measure_count_space_finite
-    simple_distributed[OF Py] simple_distributed_joint[OF Pxy],
-    OF Py[THEN simple_distributed_finite] Px[THEN simple_distributed_finite]]
-    y
-    Px[THEN simple_distributed_finite]
-    Py[THEN simple_distributed_finite]
-    Pxy[THEN simple_distributed, THEN distributed_real_AE]
-  show ?thesis
-    unfolding AE_count_space
-    apply (auto simp add: nn_integral_count_space_finite * intro!: setsum.cong split: split_max)
-    done
+  note Px = simple_distributedI[OF Px measure_nonneg refl]
+  have "AE y in count_space (Y ` space M). ennreal (Py y) =
+       \<integral>\<^sup>+ x. ennreal (if (x, y) \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy (x, y) else 0) \<partial>count_space (X ` space M)"
+    using sigma_finite_measure_count_space_finite sigma_finite_measure_count_space_finite
+      simple_distributed[OF Py] simple_distributed_joint[OF Pxy]
+    by (rule distributed_marginal_eq_joint2)
+       (auto intro: Py Px simple_distributed_finite)
+  then have "ennreal (Py y) =
+    (\<Sum>x\<in>X`space M. ennreal (if (x, y) \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy (x, y) else 0))"
+    using y Px[THEN simple_distributed_finite]
+    by (auto simp: AE_count_space nn_integral_count_space_finite)
+  also have "\<dots> = (\<Sum>x\<in>X`space M. if (x, y) \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy (x, y) else 0)"
+    using Pxy by (intro setsum_ennreal) auto
+  finally show ?thesis
+    using simple_distributed_nonneg[OF Py] simple_distributed_nonneg[OF Pxy]
+    by (subst (asm) ennreal_inj) (auto intro!: setsum_nonneg)
 qed
 
 lemma distributedI_real:
@@ -1116,9 +1119,8 @@
       then have "A \<in> sets M1"
         using gen by simp
       with f A eq[of A] X show "emeasure (distr M M1 X) A = emeasure (density M1 f) A"
-        by (simp add: emeasure_distr emeasure_density borel_measurable_ereal
-                      times_ereal.simps[symmetric] ereal_indicator
-                 del: times_ereal.simps) }
+        by (auto simp add: emeasure_distr emeasure_density ennreal_indicator
+                 intro!: nn_integral_cong split: split_indicator) }
     note eq_E = this
     show "Int_stable E" by fact
     { fix e assume "e \<in> E"
@@ -1127,7 +1129,7 @@
       then have "e \<subseteq> space M1" by (rule sets.sets_into_space) }
     then show "E \<subseteq> Pow (space M1)" by auto
     show "sets (distr M M1 X) = sigma_sets (space M1) E"
-      "sets (density M1 (\<lambda>x. ereal (f x))) = sigma_sets (space M1) E"
+      "sets (density M1 (\<lambda>x. ennreal (f x))) = sigma_sets (space M1) E"
       unfolding gen[symmetric] by auto
   qed fact+
 qed (insert X f, auto)
@@ -1136,8 +1138,8 @@
   fixes f :: "real \<Rightarrow> real"
   assumes [measurable]: "X \<in> borel_measurable M"
     and [measurable]: "f \<in> borel_measurable borel" and f[simp]: "AE x in lborel. 0 \<le> f x"
-    and g_eq: "\<And>a. (\<integral>\<^sup>+x. f x * indicator {..a} x \<partial>lborel)  = ereal (g a)"
-    and M_eq: "\<And>a. emeasure M {x\<in>space M. X x \<le> a} = ereal (g a)"
+    and g_eq: "\<And>a. (\<integral>\<^sup>+x. f x * indicator {..a} x \<partial>lborel)  = ennreal (g a)"
+    and M_eq: "\<And>a. emeasure M {x\<in>space M. X x \<le> a} = ennreal (g a)"
   shows "distributed M lborel X f"
 proof (rule distributedI_real)
   show "sets (lborel::real measure) = sigma_sets (space lborel) (range atMost)"
@@ -1168,7 +1170,7 @@
     assume "measure MX A = 0"
     with X.emeasure_space_1 X.prob_space distributed_distr_eq_density[OF X]
     show False
-      by (simp add: emeasure_density zero_ereal_def[symmetric])
+      by (simp add: emeasure_density zero_ennreal_def[symmetric])
   qed
   with measure_notin_sets[of A MX] show "A \<in> sets MX"
     by blast
@@ -1181,24 +1183,30 @@
   show "emeasure (uniform_measure M A) (space (uniform_measure M A)) = 1"
     using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)], of "space M"]
     using sets.sets_into_space[OF emeasure_neq_0_sets[OF A(1)]] A
-    by (simp add: Int_absorb2 emeasure_nonneg)
+    by (simp add: Int_absorb2 less_top)
 qed
 
 lemma prob_space_uniform_count_measure: "finite A \<Longrightarrow> A \<noteq> {} \<Longrightarrow> prob_space (uniform_count_measure A)"
-  by standard (auto simp: emeasure_uniform_count_measure space_uniform_count_measure one_ereal_def)
+  by standard (auto simp: emeasure_uniform_count_measure space_uniform_count_measure one_ennreal_def)
 
 lemma (in prob_space) measure_uniform_measure_eq_cond_prob:
   assumes [measurable]: "Measurable.pred M P" "Measurable.pred M Q"
   shows "\<P>(x in uniform_measure M {x\<in>space M. Q x}. P x) = \<P>(x in M. P x \<bar> Q x)"
 proof cases
   assume Q: "measure M {x\<in>space M. Q x} = 0"
-  then have "AE x in M. \<not> Q x"
+  then have *: "AE x in M. \<not> Q x"
     by (simp add: prob_eq_0)
-  then have "AE x in M. indicator {x\<in>space M. Q x} x / ereal 0 = 0"
-    by (auto split: split_indicator)
-  from density_cong[OF _ _ this] show ?thesis
-    by (simp add: uniform_measure_def emeasure_eq_measure cond_prob_def Q measure_density_const)
-qed (auto simp add: emeasure_eq_measure cond_prob_def intro!: arg_cong[where f=prob])
+  then have "density M (\<lambda>x. indicator {x \<in> space M. Q x} x / emeasure M {x \<in> space M. Q x}) = density M (\<lambda>x. 0)"
+    by (intro density_cong) auto
+  with * show ?thesis
+    unfolding uniform_measure_def
+    by (simp add: emeasure_density measure_def cond_prob_def emeasure_eq_0_AE)
+next
+  assume Q: "measure M {x\<in>space M. Q x} \<noteq> 0"
+  then show "\<P>(x in uniform_measure M {x \<in> space M. Q x}. P x) = cond_prob M P Q"
+    by (subst measure_uniform_measure)
+       (auto simp: emeasure_eq_measure cond_prob_def measure_nonneg intro!: arg_cong[where f=prob])
+qed
 
 lemma prob_space_point_measure:
   "finite S \<Longrightarrow> (\<And>s. s \<in> S \<Longrightarrow> 0 \<le> p s) \<Longrightarrow> (\<Sum>s\<in>S. p s) = 1 \<Longrightarrow> prob_space (point_measure S p)"
--- a/src/HOL/Probability/Projective_Family.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Projective_Family.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -88,7 +88,7 @@
     by (intro generator.intros J sets.Diff sets.top X)
   with J show "space (Pi\<^sub>M I M) - emb I J X \<in> generator"
     by (simp add: space_PiM prod_emb_PiE)
-  
+
   fix K Y assume K: "finite K" "K \<subseteq> I" and Y: "Y \<in> sets (PiM K M)"
   have "emb I (J \<union> K) (emb (J \<union> K) J X) \<inter> emb I (J \<union> K) (emb (J \<union> K) K Y) \<in> generator"
     unfolding prod_emb_Int[symmetric]
@@ -150,7 +150,7 @@
     then have "\<mu>G (emb I J {}) = 0"
       by (subst mu_G_spec) auto
     then show "\<mu>G {} = 0" by simp
-  qed (auto simp: mu_G_spec elim!: generator.cases)
+  qed
 qed
 
 lemma additive_mu_G: "additive generator \<mu>G"
@@ -182,7 +182,7 @@
     proof (clarsimp elim!: generator.cases simp: mu_G_spec del: notI)
       fix J assume "finite J" "J \<subseteq> I"
       then interpret prob_space "P J" by (rule prob_space_P)
-      show "\<And>X. X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> emeasure (P J) X \<noteq> \<infinity>"
+      show "\<And>X. X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> emeasure (P J) X \<noteq> top"
         by simp
     qed
   next
@@ -312,7 +312,7 @@
   apply (auto simp add: sets_P atLeastLessThanSuc space_P simp del: nn_integral_indicator
      intro!: nn_integral_cong split: split_indicator)
   done
-  
+
 
 primrec C :: "nat \<Rightarrow> nat \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) measure" where
   "C n 0 \<omega> = return (PiM {0..<n} M) \<omega>"
@@ -446,7 +446,7 @@
   have sets_X[measurable]: "X n \<in> sets (PiM {0..<n} M)" for n
     by (cases "{i. J i \<subseteq> {0..< n}} = {}")
        (simp_all add: X_def, auto intro!: sets.countable_INT' sets.Int)
-  
+
   have dec_X: "n \<le> m \<Longrightarrow> X m \<subseteq> PF.emb {0..<m} {0..<n} (X n)" for n m
     unfolding X_def using ivl_subset[of 0 n 0 m]
     by (cases "{i. J i \<subseteq> {0..< n}} = {}")
@@ -502,24 +502,32 @@
       apply (simp add: bind_return[OF measurable_eP] nn_integral_eP)
       done
     also have "\<dots> = (\<integral>\<^sup>+x. (INF i. ?C' i x) \<partial>P n \<omega>)"
-    proof (rule nn_integral_monotone_convergence_INF[symmetric])
+    proof (rule nn_integral_monotone_convergence_INF_AE[symmetric])
       have "(\<integral>\<^sup>+x. ?C' 0 x \<partial>P n \<omega>) \<le> (\<integral>\<^sup>+x. 1 \<partial>P n \<omega>)"
         by (intro nn_integral_mono) (auto split: split_indicator)
       also have "\<dots> < \<infinity>"
         using prob_space_P[OF \<omega>, THEN prob_space.emeasure_space_1] by simp
       finally show "(\<integral>\<^sup>+x. ?C' 0 x \<partial>P n \<omega>) < \<infinity>" .
     next
-      fix i j :: nat and x assume "i \<le> j" "x \<in> space (P n \<omega>)"
-      with \<omega> have \<omega>': "\<omega>(n := x) \<in> space (PiM {0..<Suc n} M)"
-        by (auto simp: space_P[OF \<omega>] space_PiM PiE_iff extensional_def)
-      show "?C' j x \<le> ?C' i x"
-        using \<open>i \<le> j\<close> by (subst emeasure_C[symmetric, of i]) (auto intro!: emeasure_mono dec_X del: subsetI simp: sets_C space_P \<omega> \<omega>')
+      show "AE x in P n \<omega>. ?C' (Suc i) x \<le> ?C' i x" for i
+      proof (rule AE_I2)
+        fix x assume "x \<in> space (P n \<omega>)"
+        with \<omega> have \<omega>': "\<omega>(n := x) \<in> space (PiM {0..<Suc n} M)"
+          by (auto simp: space_P[OF \<omega>] space_PiM PiE_iff extensional_def)
+        with \<omega> show "?C' (Suc i) x \<le> ?C' i x"
+          apply (subst emeasure_C[symmetric, of i "Suc i"])
+          apply (auto intro!: emeasure_mono[OF dec_X] del: subsetI
+                      simp: sets_C space_P)
+          apply (subst sets_bind[OF sets_eP])
+          apply (simp_all add: space_C space_P)
+          done
+      qed
     qed fact
     finally have "(\<integral>\<^sup>+x. (INF i. ?C' i x) \<partial>P n \<omega>) \<noteq> 0"
       by simp
-    then have "\<exists>\<^sub>F x in ae_filter (P n \<omega>). 0 < (INF i. ?C' i x)"
-       using M by (subst (asm) nn_integral_0_iff_AE)
-         (auto intro!: borel_measurable_INF simp: Filter.not_eventually not_le)
+    with M have "\<exists>\<^sub>F x in ae_filter (P n \<omega>). 0 < (INF i. ?C' i x)"
+       by (subst (asm) nn_integral_0_iff_AE)
+          (auto intro!: borel_measurable_INF simp: Filter.not_eventually not_le zero_less_iff_neq_zero)
     then have "\<exists>\<^sub>F x in ae_filter (P n \<omega>). x \<in> space (M n) \<and> 0 < (INF i. ?C' i x)"
       by (rule frequently_mp[rotated]) (auto simp: space_P \<omega>)
     then obtain x where "x \<in> space (M n)" "0 < (INF i. ?C' i x)"
@@ -585,7 +593,7 @@
   shows "emeasure lim (emb I J X) = emeasure (Pi\<^sub>M J M) X"
 proof (rule emeasure_lim[OF *], goal_cases)
   case (1 J X)
-  
+
   have "\<exists>Q. (\<forall>i. sets Q = PiM (\<Union>i. J i) M \<and> distr Q (PiM (J i) M) (\<lambda>x. restrict x (J i)) = Pi\<^sub>M (J i) M)"
   proof cases
     assume "finite (\<Union>i. J i)"
@@ -626,7 +634,7 @@
 
     let ?Q = "distr IT.PF.lim (PiM (\<Union>i. J i) M) (\<lambda>\<omega>. \<lambda>x\<in>\<Union>i. J i. \<omega> (t x))"
     { fix i
-      have "distr ?Q (Pi\<^sub>M (J i) M) (\<lambda>x. restrict x (J i)) = 
+      have "distr ?Q (Pi\<^sub>M (J i) M) (\<lambda>x. restrict x (J i)) =
         distr IT.PF.lim (Pi\<^sub>M (J i) M) ((\<lambda>\<omega>. \<lambda>n\<in>J i. \<omega> (t n)) \<circ> (\<lambda>\<omega>. restrict \<omega> (t`J i)))"
       proof (subst distr_distr)
         have "(\<lambda>\<omega>. \<omega> (t x)) \<in> measurable (Pi\<^sub>M UNIV (\<lambda>x. M (f x))) (M x)" if x: "x \<in> J i" for x i
--- a/src/HOL/Probability/Projective_Limit.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Projective_Limit.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -115,8 +115,8 @@
     and "0 < (INF i. P (J i) (B i))" (is "0 < ?a")
   moreover have "?a \<le> 1"
     using J by (auto intro!: INF_lower2[of 0] prob_space_P[THEN prob_space.measure_le_1])
-  ultimately obtain r where r: "?a = ereal r" "0 < r" "r \<le> 1"
-    by (cases "?a") auto
+  ultimately obtain r where r: "?a = ennreal r" "0 < r" "r \<le> 1"
+    by (cases "?a") (auto simp: top_unique)
   def Z \<equiv> "\<lambda>n. emb I (J n) (B n)"
   have Z_mono: "n \<le> m \<Longrightarrow> Z m \<subseteq> Z n" for n m
     unfolding Z_def using B[THEN antimonoD, of n m] .
@@ -143,7 +143,7 @@
   interpret P': prob_space "P' n" for n
     unfolding P'_def mapmeasure_def using J
     by (auto intro!: prob_space_distr fm_measurable simp: measurable_cong_sets[OF sets_P])
-  
+
   let ?SUP = "\<lambda>n. SUP K : {K. K \<subseteq> fm n ` (B n) \<and> compact K}. emeasure (P' n) K"
   { fix n
     have "emeasure (P (J n)) (B n) = emeasure (P' n) (fm n ` (B n))"
@@ -156,40 +156,37 @@
       show "fm n ` B n \<in> sets borel"
         unfolding borel_eq_PiF_borel by (auto simp: P'_def fm_image_measurable_finite sets_P J(3))
     qed simp
-    finally have "emeasure (P (J n)) (B n) = ?SUP n" "?SUP n \<noteq> \<infinity>" "?SUP n \<noteq> - \<infinity>" by auto
+    finally have *: "emeasure (P (J n)) (B n) = ?SUP n" .
+    have "?SUP n \<noteq> \<infinity>"
+      unfolding *[symmetric] by simp
+    note * this
   } note R = this
   have "\<forall>n. \<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> 2 powr (-n) * ?a \<and> compact K \<and> K \<subseteq> fm n ` B n"
   proof
-    fix n show "\<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> ereal (2 powr - real n) * ?a \<and>
+    fix n show "\<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> ennreal (2 powr - real n) * ?a \<and>
         compact K \<and> K \<subseteq> fm n ` B n"
       unfolding R[of n]
     proof (rule ccontr)
-      assume H: "\<not> (\<exists>K'. ?SUP n - emeasure (P' n) K' \<le> ereal (2 powr - real n)  * ?a \<and>
-        compact K' \<and> K' \<subseteq> fm n ` B n)"
-      have "?SUP n \<le> ?SUP n - 2 powr (-n) * ?a"
+      assume H: "\<nexists>K'. ?SUP n - emeasure (P' n) K' \<le> ennreal (2 powr - real n)  * ?a \<and>
+        compact K' \<and> K' \<subseteq> fm n ` B n"
+      have "?SUP n + 0 < ?SUP n + 2 powr (-n) * ?a"
+        using R[of n] unfolding ennreal_add_left_cancel_less ennreal_zero_less_mult_iff
+        by (auto intro: \<open>0 < ?a\<close>)
+      also have "\<dots> = (SUP K:{K. K \<subseteq> fm n ` B n \<and> compact K}. emeasure (P' n) K + 2 powr (-n) * ?a)"
+        by (rule ennreal_SUP_add_left[symmetric]) auto
+      also have "\<dots> \<le> ?SUP n"
       proof (intro SUP_least)
-        fix K
-        assume "K \<in> {K. K \<subseteq> fm n ` B n \<and> compact K}"
-        with H have "\<not> ?SUP n - emeasure (P' n) K \<le> 2 powr (-n) * ?a"
+        fix K assume "K \<in> {K. K \<subseteq> fm n ` B n \<and> compact K}"
+        with H have "2 powr (-n) * ?a < ?SUP n - emeasure (P' n) K"
           by auto
-        hence "?SUP n - emeasure (P' n) K > 2 powr (-n) * ?a"
-          unfolding not_less[symmetric] by simp
-        hence "?SUP n - 2 powr (-n) * ?a > emeasure (P' n) K"
-          using \<open>0 < ?a\<close> by (auto simp add: ereal_less_minus_iff ac_simps)
-        thus "?SUP n - 2 powr (-n) * ?a \<ge> emeasure (P' n) K" by simp
+        then show "emeasure (P' n) K + (2 powr (-n)) * ?a \<le> ?SUP n"
+          by (subst (asm) less_diff_eq_ennreal) (auto simp: less_top[symmetric] R(1)[symmetric] ac_simps)
       qed
-      hence "?SUP n + 0 \<le> ?SUP n - (2 powr (-n) * ?a)" using \<open>0 < ?a\<close> by simp
-      hence "?SUP n + 0 \<le> ?SUP n + - (2 powr (-n) * ?a)" unfolding minus_ereal_def .
-      hence "0 \<le> - (2 powr (-n) * ?a)"
-        using \<open>?SUP n \<noteq> \<infinity>\<close> \<open>?SUP n \<noteq> - \<infinity>\<close>
-        by (subst (asm) ereal_add_le_add_iff) (auto simp:)
-      moreover have "ereal (2 powr - real n) * ?a > 0" using \<open>0 < ?a\<close>
-        by (auto simp: ereal_zero_less_0_iff)
-      ultimately show False by simp
+      finally show False by simp
     qed
   qed
   then obtain K' where K':
-    "\<And>n. emeasure (P (J n)) (B n) - emeasure (P' n) (K' n) \<le> ereal (2 powr - real n) * ?a"
+    "\<And>n. emeasure (P (J n)) (B n) - emeasure (P' n) (K' n) \<le> ennreal (2 powr - real n) * ?a"
     "\<And>n. compact (K' n)" "\<And>n. K' n \<subseteq> fm n ` B n"
     unfolding choice_iff by blast
   def K \<equiv> "\<lambda>n. fm n -` K' n \<inter> space (Pi\<^sub>M (J n) (\<lambda>_. borel))"
@@ -252,11 +249,11 @@
       by (auto simp del: prod_emb_INT intro!: generator.intros)
     have *: "\<mu>G (Z n) = P (J n) (B n)"
       unfolding Z_def using J by (intro mu_G_spec) auto
-    then have "\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>" by auto
+    then have "\<mu>G (Z n) \<noteq> \<infinity>" by auto
     note *
     moreover have *: "\<mu>G (Y n) = P (J n) (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))"
       unfolding Y_emb using J J_mono K_sets \<open>n \<ge> 1\<close> by (subst mu_G_spec) auto
-    then have "\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>" by auto
+    then have "\<mu>G (Y n) \<noteq> \<infinity>" by auto
     note *
     moreover have "\<mu>G (Z n - Y n) =
         P (J n) (B n - (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i)))"
@@ -266,8 +263,10 @@
     have "\<mu>G (Z n) - \<mu>G (Y n) = \<mu>G (Z n - Y n)"
       using J J_mono K_sets \<open>n \<ge> 1\<close>
       by (simp only: emeasure_eq_measure Z_def)
-        (auto dest!: bspec[where x=n] intro!: measure_Diff[symmetric] set_mp[OF K_B]
-          simp: extensional_restrict emeasure_eq_measure prod_emb_iff sets_P space_P)
+         (auto dest!: bspec[where x=n] intro!: measure_Diff[symmetric] set_mp[OF K_B]
+               intro!: arg_cong[where f=ennreal]
+               simp: extensional_restrict emeasure_eq_measure prod_emb_iff sets_P space_P
+                     ennreal_minus measure_nonneg)
     also have subs: "Z n - Y n \<subseteq> (\<Union>i\<in>{1..n}. (Z i - Z' i))"
       using \<open>n \<ge> 1\<close> unfolding Y_def UN_extend_simps(7) by (intro UN_mono Diff_mono Z_mono order_refl) auto
     have "Z n - Y n \<in> generator" "(\<Union>i\<in>{1..n}. (Z i - Z' i)) \<in> generator"
@@ -290,32 +289,29 @@
         unfolding K_def P'_def
         by (auto simp: mapmeasure_PiF borel_eq_PiF_borel[symmetric]
           compact_imp_closed[OF \<open>compact (K' _)\<close>] space_PiM PiE_def)
-      also have "\<dots> \<le> ereal (2 powr - real i) * ?a" using K'(1)[of i] .
+      also have "\<dots> \<le> ennreal (2 powr - real i) * ?a" using K'(1)[of i] .
       finally show "\<mu>G (Z i - Z' i) \<le> (2 powr - real i) * ?a" .
     qed
-    also have "\<dots> = ereal ((\<Sum> i\<in>{1..n}. (2 powr -real_of_ereal i)) * real_of_ereal ?a)"
-      by (simp add: setsum_left_distrib r)
-    also have "\<dots> < ereal (1 * real_of_ereal ?a)" unfolding less_ereal.simps
-    proof (rule mult_strict_right_mono)
+    also have "\<dots> = ennreal ((\<Sum> i\<in>{1..n}. (2 powr -enn2real i)) * enn2real ?a)"
+      using r by (simp add: setsum_left_distrib ennreal_mult[symmetric])
+    also have "\<dots> < ennreal (1 * enn2real ?a)"
+    proof (intro ennreal_lessI mult_strict_right_mono)
       have "(\<Sum>i = 1..n. 2 powr - real i) = (\<Sum>i = 1..<Suc n. (1/2) ^ i)"
-        by (rule setsum.cong) (auto simp: powr_realpow powr_divide power_divide powr_minus_divide)  
+        by (rule setsum.cong) (auto simp: powr_realpow powr_divide power_divide powr_minus_divide)
       also have "{1..<Suc n} = {..<Suc n} - {0}" by auto
       also have "setsum (op ^ (1 / 2::real)) ({..<Suc n} - {0}) =
         setsum (op ^ (1 / 2)) ({..<Suc n}) - 1" by (auto simp: setsum_diff1)
       also have "\<dots> < 1" by (subst geometric_sum) auto
-      finally show "(\<Sum>i = 1..n. 2 powr - real_of_ereal i) < 1" by simp
-    qed (auto simp: r ereal_less_real_iff zero_ereal_def[symmetric])
+      finally show "(\<Sum>i = 1..n. 2 powr - enn2real i) < 1" by simp
+    qed (auto simp: r enn2real_positive_iff)
     also have "\<dots> = ?a" by (auto simp: r)
     also have "\<dots> \<le> \<mu>G (Z n)"
       using J by (auto intro: INF_lower simp: Z_def mu_G_spec)
     finally have "\<mu>G (Z n) - \<mu>G (Y n) < \<mu>G (Z n)" .
     hence R: "\<mu>G (Z n) < \<mu>G (Z n) + \<mu>G (Y n)"
-      using \<open>\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>\<close> by (simp add: ereal_minus_less)
-    have "0 \<le> (- \<mu>G (Z n)) + \<mu>G (Z n)" using \<open>\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>\<close> by auto
-    also have "\<dots> < (- \<mu>G (Z n)) + (\<mu>G (Z n) + \<mu>G (Y n))"
-      apply (rule ereal_less_add[OF _ R]) using \<open>\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>\<close> by auto
-    finally have "\<mu>G (Y n) > 0"
-      using \<open>\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>\<close> by (auto simp: ac_simps zero_ereal_def[symmetric])
+      using \<open>\<mu>G (Y n) \<noteq> \<infinity>\<close> by (auto simp: zero_less_iff_neq_zero)
+    then have "\<mu>G (Y n) > 0"
+      by simp
     thus "Y n \<noteq> {}" using positive_mu_G by (auto simp add: positive_def)
   qed
   hence "\<forall>n\<in>{1..}. \<exists>y. y \<in> Y n" by auto
@@ -473,7 +469,7 @@
 proof
   have *: "emb I {} {\<lambda>x. undefined} = space (\<Pi>\<^sub>M i\<in>I. borel)"
     by (auto simp: prod_emb_def space_PiM)
-  interpret prob_space "P {}" 
+  interpret prob_space "P {}"
     using prob_space_P by simp
   show "emeasure lim (space lim) = 1"
     using emeasure_lim_emb[of "{}" "{\<lambda>x. undefined}"] emeasure_space_1
--- a/src/HOL/Probability/Radon_Nikodym.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Radon_Nikodym.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -24,7 +24,7 @@
 proof (rule emeasure_measure_of_sigma)
   show "sigma_algebra (space M) (sets M)" ..
   show "positive (sets M) ?\<mu>"
-    using pos by (simp add: positive_def ereal_diff_positive)
+    using pos by (simp add: positive_def)
   show "countably_additive (sets M) ?\<mu>"
   proof (rule countably_additiveI)
     fix A :: "nat \<Rightarrow> _"  assume A: "range A \<subseteq> sets M" and "disjoint_family A"
@@ -34,9 +34,9 @@
       by (simp_all add: suminf_emeasure sets_eq)
     with A have "(\<Sum>i. emeasure M (A i) - emeasure N (A i)) =
       (\<Sum>i. emeasure M (A i)) - (\<Sum>i. emeasure N (A i))"
-      using fin
-      by (intro suminf_ereal_minus pos emeasure_nonneg)
-         (auto simp: sets_eq finite_measure.emeasure_eq_measure suminf_emeasure)
+      using fin pos[of "A _"]
+      by (intro ennreal_suminf_minus)
+         (auto simp: sets_eq finite_measure.emeasure_eq_measure suminf_emeasure measure_nonneg)
     then show "(\<Sum>i. emeasure M (A i) - emeasure N (A i)) =
       emeasure M (\<Union>i. A i) - emeasure N (\<Union>i. A i) "
       by (simp add: suminf)
@@ -44,7 +44,7 @@
 qed fact
 
 lemma (in sigma_finite_measure) Ex_finite_integrable_function:
-  shows "\<exists>h\<in>borel_measurable M. integral\<^sup>N M h \<noteq> \<infinity> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>) \<and> (\<forall>x. 0 \<le> h x)"
+  "\<exists>h\<in>borel_measurable M. integral\<^sup>N M h \<noteq> \<infinity> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>)"
 proof -
   obtain A :: "nat \<Rightarrow> 'a set" where
     range[measurable]: "range A \<subseteq> sets M" and
@@ -56,8 +56,8 @@
   have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)"
   proof
     fix i show "\<exists>x. 0 < x \<and> x < inverse (?B i)"
-      using measure[of i] emeasure_nonneg[of M "A i"]
-      by (auto intro!: dense simp: ereal_0_gt_inverse ereal_zero_le_0_iff)
+      using measure[of i]
+      by (auto intro!: dense simp: ennreal_inverse_positive ennreal_mult_eq_top_iff power_eq_top_ennreal)
   qed
   from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
     "\<And>i. n i < inverse (2^Suc i * emeasure M (A i))" by auto
@@ -69,34 +69,43 @@
       using range by fastforce+
     then have "integral\<^sup>N M ?h = (\<Sum>i. n i * emeasure M (A i))" using pos
       by (simp add: nn_integral_suminf nn_integral_cmult_indicator)
-    also have "\<dots> \<le> (\<Sum>i. (1 / 2)^Suc i)"
-    proof (rule suminf_le_pos)
+    also have "\<dots> \<le> (\<Sum>i. ennreal ((1/2)^Suc i))"
+    proof (intro suminf_le allI)
       fix N
       have "n N * emeasure M (A N) \<le> inverse (2^Suc N * emeasure M (A N)) * emeasure M (A N)"
-        using n[of N]
-        by (intro ereal_mult_right_mono) auto
-      also have "\<dots> \<le> (1 / 2) ^ Suc N"
-        using measure[of N] n[of N]
-        by (cases rule: ereal2_cases[of "n N" "emeasure M (A N)"])
-           (simp_all add: inverse_eq_divide power_divide one_ereal_def ereal_power_divide)
-      finally show "n N * emeasure M (A N) \<le> (1 / 2) ^ Suc N" .
-      show "0 \<le> n N * emeasure M (A N)" using n[of N] \<open>A N \<in> sets M\<close> by (simp add: emeasure_nonneg)
-    qed
-    finally show "integral\<^sup>N M ?h \<noteq> \<infinity>" unfolding suminf_half_series_ereal by auto
+        using n[of N] by (intro mult_right_mono) auto
+      also have "\<dots> = (1/2)^Suc N * (inverse (emeasure M (A N)) * emeasure M (A N))"
+        using measure[of N]
+        by (simp add: ennreal_inverse_power divide_ennreal_def ennreal_inverse_mult
+                      power_eq_top_ennreal less_top[symmetric] mult_ac
+                 del: power_Suc)
+      also have "\<dots> \<le> inverse (ennreal 2) ^ Suc N"
+        using measure[of N]
+        apply (cases "emeasure M (A N)" rule: ennreal_cases)
+        apply (cases "emeasure M (A N) = 0")
+        apply (auto simp: inverse_ennreal ennreal_mult[symmetric] divide_ennreal_def simp del: power_Suc)
+        done
+      also have "\<dots> = ennreal (inverse 2 ^ Suc N)"
+        by (subst ennreal_power[symmetric], simp) (simp add: inverse_ennreal)
+      finally show "n N * emeasure M (A N) \<le> ennreal ((1/2)^Suc N)"
+        by simp
+    qed auto
+    also have "\<dots> < top"
+      unfolding less_top[symmetric]
+      apply (rule ennreal_suminf_neq_top)
+      apply (subst summable_Suc_iff)
+      apply (subst summable_geometric)
+      apply auto
+      done
+    finally show "integral\<^sup>N M ?h \<noteq> \<infinity>"
+      by (auto simp: top_unique)
   next
     { fix x assume "x \<in> space M"
       then obtain i where "x \<in> A i" using space[symmetric] by auto
       with disjoint n have "?h x = n i"
         by (auto intro!: suminf_cmult_indicator intro: less_imp_le)
-      then show "0 < ?h x" and "?h x < \<infinity>" using n[of i] by auto }
+      then show "0 < ?h x" and "?h x < \<infinity>" using n[of i] by (auto simp: less_top[symmetric]) }
     note pos = this
-    fix x show "0 \<le> ?h x"
-    proof cases
-      assume "x \<in> space M" then show "0 \<le> ?h x" using pos by (auto intro: less_imp_le)
-    next
-      assume "x \<notin> space M" then have "\<And>i. x \<notin> A i" using space by auto
-      then show "0 \<le> ?h x" by auto
-    qed
   qed measurable
 qed
 
@@ -207,7 +216,7 @@
       proof (induct n)
         case (Suc n) with dA_epsilon[of n, OF B] show ?case by (simp del: A_simps add: of_nat_Suc field_simps)
       next
-        case 0 with measure_empty show ?case by (simp add: zero_ereal_def)
+        case 0 with measure_empty show ?case by (simp add: zero_ennreal_def)
       qed } note dA_less = this
     have decseq: "decseq (\<lambda>n. ?d (A n))" unfolding decseq_eq_incseq
     proof (rule incseq_SucI)
@@ -347,7 +356,7 @@
     from this[THEN bspec, OF sets.top] show "integral\<^sup>N M g \<le> N (space M)"
       by (simp cong: nn_integral_cong)
   qed
-  from SUP_countable_SUP [OF \<open>G \<noteq> {}\<close>, of "integral\<^sup>N M"] guess ys .. note ys = this
+  from ennreal_SUP_countable_SUP [OF \<open>G \<noteq> {}\<close>, of "integral\<^sup>N M"] guess ys .. note ys = this
   then have "\<forall>n. \<exists>g. g\<in>G \<and> integral\<^sup>N M g = ys n"
   proof safe
     fix n assume "range ys \<subseteq> integral\<^sup>N M ` G"
@@ -394,14 +403,14 @@
   have emeasure_M: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure ?M A = ?t A"
   proof (subst emeasure_diff_measure)
     from f_le_N[of "space M"] show "finite_measure N" "finite_measure (density M f)"
-      by (auto intro!: finite_measureI simp: emeasure_density cong: nn_integral_cong)
+      by (auto intro!: finite_measureI simp: emeasure_density top_unique cong: nn_integral_cong)
   next
     fix B assume "B \<in> sets N" with f_le_N[of B] show "emeasure (density M f) B \<le> emeasure N B"
       by (auto simp: sets_eq emeasure_density cong: nn_integral_cong)
   qed (auto simp: sets_eq emeasure_density)
-  from emeasure_M[of "space M"] N.finite_emeasure_space nn_integral_nonneg[of M "?F (space M)"]
+  from emeasure_M[of "space M"] N.finite_emeasure_space
   interpret M': finite_measure ?M
-    by (auto intro!: finite_measureI simp: sets_eq_imp_space_eq[OF sets_eq] N.emeasure_eq_measure ereal_minus_eq_PInfty_iff)
+    by (auto intro!: finite_measureI simp: sets_eq_imp_space_eq[OF sets_eq] N.emeasure_eq_measure )
 
   have ac: "absolutely_continuous M ?M" unfolding absolutely_continuous_def
   proof
@@ -410,7 +419,7 @@
       unfolding absolutely_continuous_def by auto
     moreover from A_M A_N have "(\<integral>\<^sup>+ x. ?F A x \<partial>M) \<le> N A" using \<open>f \<in> G\<close> by (auto simp: G_def)
     ultimately have "N A - (\<integral>\<^sup>+ x. ?F A x \<partial>M) = 0"
-      using nn_integral_nonneg[of M] by (auto intro!: antisym)
+      by (auto intro!: antisym)
     then show "A \<in> null_sets ?M"
       using A_M by (simp add: emeasure_M null_sets_def sets_eq)
   qed
@@ -418,7 +427,7 @@
   proof (rule ccontr)
     assume "\<not> ?thesis"
     then obtain A where A: "A \<in> sets M" and pos: "0 < ?M A"
-      by (auto simp: not_le)
+      by (auto simp: zero_less_iff_neq_zero)
     note pos
     also have "?M A \<le> ?M (space M)"
       using emeasure_space[of ?M A] by (simp add: sets_eq[THEN sets_eq_imp_space_eq])
@@ -427,28 +436,30 @@
     from pos_t have "emeasure M (space M) \<noteq> 0"
       using ac unfolding absolutely_continuous_def by (auto simp: null_sets_def)
     then have pos_M: "0 < emeasure M (space M)"
-      using emeasure_nonneg[of M "space M"] by (simp add: le_less)
+      by (simp add: zero_less_iff_neq_zero)
     moreover
     have "(\<integral>\<^sup>+x. f x * indicator (space M) x \<partial>M) \<le> N (space M)"
       using \<open>f \<in> G\<close> unfolding G_def by auto
     hence "(\<integral>\<^sup>+x. f x * indicator (space M) x \<partial>M) \<noteq> \<infinity>"
-      using M'.finite_emeasure_space by auto
+      using M'.finite_emeasure_space by (auto simp: top_unique)
     moreover
     def b \<equiv> "?M (space M) / emeasure M (space M) / 2"
     ultimately have b: "b \<noteq> 0 \<and> 0 \<le> b \<and> b \<noteq> \<infinity>"
-      by (auto simp: ereal_divide_eq)
-    then have b: "b \<noteq> 0" "0 \<le> b" "0 < b"  "b \<noteq> \<infinity>" by auto
+      by (auto simp: ennreal_divide_eq_top_iff)
+    then have b: "b \<noteq> 0" "0 \<le> b" "0 < b"  "b \<noteq> \<infinity>"
+      by (auto simp: less_le)
     let ?Mb = "density M (\<lambda>_. b)"
     have Mb: "finite_measure ?Mb" "sets ?Mb = sets ?M"
-        using b by (auto simp: emeasure_density_const sets_eq intro!: finite_measureI)
+        using b by (auto simp: emeasure_density_const sets_eq ennreal_mult_eq_top_iff intro!: finite_measureI)
     from M'.Radon_Nikodym_aux[OF this] guess A0 ..
     then have "A0 \<in> sets M"
-      and space_less_A0: "measure ?M (space M) - real_of_ereal b * measure M (space M) \<le> measure ?M A0 - real_of_ereal b * measure M A0"
-      and *: "\<And>B. B \<in> sets M \<Longrightarrow> B \<subseteq> A0 \<Longrightarrow> 0 \<le> measure ?M B - real_of_ereal b * measure M B"
+      and space_le_A0: "measure ?M (space M) - enn2real b * measure M (space M) \<le> measure ?M A0 - enn2real b * measure M A0"
+      and *: "\<And>B. B \<in> sets M \<Longrightarrow> B \<subseteq> A0 \<Longrightarrow> 0 \<le> measure ?M B - enn2real b * measure M B"
       using b by (simp_all add: measure_density_const sets_eq_imp_space_eq[OF sets_eq] sets_eq)
     { fix B assume B: "B \<in> sets M" "B \<subseteq> A0"
       with *[OF this] have "b * emeasure M B \<le> ?M B"
-        using b unfolding M'.emeasure_eq_measure emeasure_eq_measure by (cases b) auto }
+        using b unfolding M'.emeasure_eq_measure emeasure_eq_measure
+        by (cases b rule: ennreal_cases) (auto simp: ennreal_mult[symmetric] measure_nonneg) }
     note bM_le_t = this
     let ?f0 = "\<lambda>x. f x + b * indicator A0 x"
     { fix A assume A: "A \<in> sets M"
@@ -473,48 +484,49 @@
         by (auto intro!: add_left_mono simp: sets_eq)
       also have "\<dots> \<le> N A"
         unfolding emeasure_M[OF \<open>A \<in> sets M\<close>]
-        using f_le_v N.emeasure_eq_measure[of A] nn_integral_nonneg[of M "?F A"]
-        by (cases "\<integral>\<^sup>+x. ?F A x \<partial>M", cases "N A") auto
+        using f_le_v N.emeasure_eq_measure[of A]
+        by (cases "\<integral>\<^sup>+x. ?F A x \<partial>M" "N A" rule: ennreal2_cases)
+           (auto simp: top_unique measure_nonneg ennreal_minus ennreal_plus[symmetric] simp del: ennreal_plus)
       finally have "(\<integral>\<^sup>+x. ?f0 x * indicator A x \<partial>M) \<le> N A" . }
     hence "?f0 \<in> G" using \<open>A0 \<in> sets M\<close> b \<open>f \<in> G\<close> by (auto simp: G_def)
     have int_f_finite: "integral\<^sup>N M f \<noteq> \<infinity>"
-      by (metis N.emeasure_finite ereal_infty_less_eq2(1) int_f_eq_y y_le)
-    have  "0 < ?M (space M) - emeasure ?Mb (space M)"
-      using pos_t
-      by (simp add: b emeasure_density_const)
-         (simp add: M'.emeasure_eq_measure emeasure_eq_measure pos_M b_def)
-    also have "\<dots> \<le> ?M A0 - b * emeasure M A0"
-      using space_less_A0 \<open>A0 \<in> sets M\<close> b
-      by (cases b) (auto simp add: b emeasure_density_const sets_eq M'.emeasure_eq_measure emeasure_eq_measure)
-    finally have 1: "b * emeasure M A0 < ?M A0"
-      by (metis M'.emeasure_real \<open>A0 \<in> sets M\<close> bM_le_t diff_self ereal_less(1) ereal_minus(1)
-                less_eq_ereal_def mult_zero_left not_square_less_zero subset_refl zero_ereal_def)
-    with b have "0 < ?M A0"
-      by (metis M'.emeasure_real MInfty_neq_PInfty(1) emeasure_real ereal_less_eq(5) ereal_zero_times
-               ereal_mult_eq_MInfty ereal_mult_eq_PInfty ereal_zero_less_0_iff less_eq_ereal_def)
-    then have "emeasure M A0 \<noteq> 0" using ac \<open>A0 \<in> sets M\<close>
-      by (auto simp: absolutely_continuous_def null_sets_def)
-    then have "0 < emeasure M A0" using emeasure_nonneg[of M A0] by auto
-    hence "0 < b * emeasure M A0" using b by (auto simp: ereal_zero_less_0_iff)
-    with int_f_finite have "?y + 0 < integral\<^sup>N M f + b * emeasure M A0" unfolding int_f_eq_y
-      using \<open>f \<in> G\<close>
-      by (intro ereal_add_strict_mono) (auto intro!: SUP_upper2 nn_integral_nonneg)
-    also have "\<dots> = integral\<^sup>N M ?f0" using f0_eq[OF sets.top] \<open>A0 \<in> sets M\<close> sets.sets_into_space
-      by (simp cong: nn_integral_cong)
-    finally have "?y < integral\<^sup>N M ?f0" by simp
-    moreover from \<open>?f0 \<in> G\<close> have "integral\<^sup>N M ?f0 \<le> ?y" by (auto intro!: SUP_upper)
+      by (metis top_unique infinity_ennreal_def int_f_eq_y y_le N.emeasure_finite)
+    have pos: "0 < measure ?M (space M) - enn2real b * measure M (space M)"
+      using pos_t pos_M
+      by (simp add: M'.emeasure_eq_measure emeasure_eq_measure b_def divide_ennreal ennreal_divide_numeral)
+    also have "\<dots> \<le> measure ?M A0 - enn2real b * measure M A0"
+      by (rule space_le_A0)
+    finally have "enn2real b * measure M A0 < measure ?M A0"
+      by simp
+    with b have "?M A0 \<noteq> 0"
+      by (cases b rule: ennreal_cases)
+         (auto simp: M'.emeasure_eq_measure measure_nonneg mult_less_0_iff not_le[symmetric])
+    then have "emeasure M A0 \<noteq> 0"
+      using ac \<open>A0 \<in> sets M\<close> by (auto simp: absolutely_continuous_def null_sets_def)
+    then have "0 < emeasure M A0"
+      by (auto simp: zero_less_iff_neq_zero)
+    hence "0 < b * emeasure M A0"
+      using b by (auto simp: ennreal_zero_less_mult_iff)
+    with int_f_finite have "?y < integral\<^sup>N M f + b * emeasure M A0"
+      unfolding int_f_eq_y by auto
+    also have "\<dots> = integral\<^sup>N M ?f0"
+      using f0_eq[OF sets.top] \<open>A0 \<in> sets M\<close> sets.sets_into_space by (simp cong: nn_integral_cong)
+    finally have "?y < integral\<^sup>N M ?f0"
+      by simp
+    moreover have "integral\<^sup>N M ?f0 \<le> ?y"
+      using \<open>?f0 \<in> G\<close> by (auto intro!: SUP_upper)
     ultimately show False by auto
   qed
-  let ?f = "\<lambda>x. max 0 (f x)"
   show ?thesis
-  proof (intro bexI[of _ ?f] measure_eqI conjI)
-    show "sets (density M ?f) = sets N"
+  proof (intro bexI[of _ f] measure_eqI conjI)
+    show "sets (density M f) = sets N"
       by (simp add: sets_eq)
-    fix A assume A: "A\<in>sets (density M ?f)"
-    then show "emeasure (density M ?f) A = emeasure N A"
+    fix A assume A: "A\<in>sets (density M f)"
+    then show "emeasure (density M f) A = emeasure N A"
       using \<open>f \<in> G\<close> A upper_bound[THEN bspec, of A] N.emeasure_eq_measure[of A]
       by (cases "integral\<^sup>N M (?F A)")
-         (auto intro!: antisym simp add: emeasure_density G_def emeasure_M density_max_0[symmetric])
+         (auto intro!: antisym simp: emeasure_density G_def emeasure_M ennreal_minus_eq_0 top_unique
+                               simp del: measure_nonneg)
   qed auto
 qed
 
@@ -530,9 +542,10 @@
   then have Q_not_empty: "?Q \<noteq> {}" by blast
   have "?a \<le> emeasure M (space M)" using sets.sets_into_space
     by (auto intro!: SUP_least emeasure_mono)
-  then have "?a \<noteq> \<infinity>" using finite_emeasure_space
-    by auto
-  from SUP_countable_SUP [OF Q_not_empty, of "emeasure M"]
+  then have "?a \<noteq> \<infinity>"
+    using finite_emeasure_space
+    by (auto simp: less_top[symmetric] top_unique simp del: SUP_eq_top_iff Sup_eq_top_iff)
+  from ennreal_SUP_countable_SUP [OF Q_not_empty, of "emeasure M"]
   obtain Q'' where "range Q'' \<subseteq> emeasure M ` ?Q" and a: "?a = (SUP i::nat. Q'' i)"
     by auto
   then have "\<forall>i. \<exists>Q'. Q'' i = emeasure M Q' \<and> Q' \<in> ?Q" by auto
@@ -552,7 +565,7 @@
     fix i have "Q' ` {..i} \<subseteq> sets M" using Q' by auto
     then have "N (?O i) \<le> (\<Sum>i\<le>i. N (Q' i))"
       by (simp add: sets_eq emeasure_subadditive_finite)
-    also have "\<dots> < \<infinity>" using Q' by (simp add: setsum_Pinfty)
+    also have "\<dots> < \<infinity>" using Q' by (simp add: less_top)
     finally show "N (?O i) \<noteq> \<infinity>" by simp
   qed auto
   have O_mono: "\<And>n. ?O n \<subseteq> ?O (Suc n)" by fastforce
@@ -584,7 +597,7 @@
     { fix A assume A: "A \<in> sets M" "A \<subseteq> space M - ?O_0"
       show "emeasure M A = 0 \<and> N A = 0 \<or> 0 < emeasure M A \<and> N A = \<infinity>"
       proof (rule disjCI, simp)
-        assume *: "0 < emeasure M A \<longrightarrow> N A \<noteq> \<infinity>"
+        assume *: "emeasure M A = 0 \<or> N A \<noteq> top"
         show "emeasure M A = 0 \<and> N A = 0"
         proof (cases "emeasure M A = 0")
           case True
@@ -593,7 +606,7 @@
           with True show ?thesis by simp
         next
           case False
-          with * have "N A \<noteq> \<infinity>" using emeasure_nonneg[of M A] by auto
+          with * have "N A \<noteq> \<infinity>" by auto
           with A have "emeasure M ?O_0 + emeasure M A = emeasure M (?O_0 \<union> A)"
             using Q' by (auto intro!: plus_emeasure sets.countable_UN)
           also have "\<dots> = (SUP i. emeasure M (?O i \<union> A))"
@@ -609,7 +622,7 @@
               from O_in_G[of i] have "N (?O i \<union> A) \<le> N (?O i) + N A"
                 using emeasure_subadditive[of "?O i" N A] A O_sets by (auto simp: sets_eq)
               with O_in_G[of i] show "N (?O i \<union> A) \<noteq> \<infinity>"
-                using \<open>N A \<noteq> \<infinity>\<close> by auto
+                using \<open>N A \<noteq> \<infinity>\<close> by (auto simp: top_unique)
             qed
             then show "emeasure M (?O i \<union> A) \<le> ?a" by (rule SUP_upper)
           qed
@@ -624,10 +637,9 @@
           unfolding Q_def using Q'[of 0] by simp
       next
         case (Suc n)
-        with \<open>?O n \<in> ?Q\<close> \<open>?O (Suc n) \<in> ?Q\<close>
-            emeasure_Diff[OF _ _ _ O_mono, of N n] emeasure_nonneg[of N "(\<Union>x\<le>n. Q' x)"]
+        with \<open>?O n \<in> ?Q\<close> \<open>?O (Suc n) \<in> ?Q\<close> emeasure_Diff[OF _ _ _ O_mono, of N n]
         show ?thesis
-          by (auto simp: sets_eq ereal_minus_eq_PInfty_iff Q_def)
+          by (auto simp: sets_eq Q_def)
       qed }
     show "space M - ?O_0 \<in> sets M" using Q'_sets by auto
     { fix j have "(\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q i)"
@@ -700,13 +712,13 @@
       have Qi: "\<And>i. Q i \<in> sets M" using Q by auto
       have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M"
         "\<And>i. AE x in M. 0 \<le> f i x * indicator (Q i \<inter> A) x"
-        using borel Qi Q0(1) \<open>A \<in> sets M\<close> by (auto intro!: borel_measurable_ereal_times)
+        using borel Qi Q0(1) \<open>A \<in> sets M\<close> by auto
       have "(\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) + \<infinity> * indicator (Q0 \<inter> A) x \<partial>M)"
         using borel by (intro nn_integral_cong) (auto simp: indicator_def)
       also have "\<dots> = (\<integral>\<^sup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * emeasure M (Q0 \<inter> A)"
         using borel Qi Q0(1) \<open>A \<in> sets M\<close>
-        by (subst nn_integral_add) (auto simp del: ereal_infty_mult
-            simp add: nn_integral_cmult_indicator sets.Int intro!: suminf_0_le)
+        by (subst nn_integral_add)
+           (auto simp add: nn_integral_cmult_indicator sets.Int intro!: suminf_0_le)
       also have "\<dots> = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M (Q0 \<inter> A)"
         by (subst integral_eq[OF \<open>A \<in> sets M\<close>], subst nn_integral_suminf) auto
       finally have "(\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M) = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M (Q0 \<inter> A)" .
@@ -716,7 +728,7 @@
       moreover have "\<infinity> * emeasure M (Q0 \<inter> A) = N (Q0 \<inter> A)"
       proof -
         have "Q0 \<inter> A \<in> sets M" using Q0(1) \<open>A \<in> sets M\<close> by blast
-        from in_Q0[OF this] show ?thesis by auto
+        from in_Q0[OF this] show ?thesis by (auto simp: ennreal_top_mult)
       qed
       moreover have "Q0 \<inter> A \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M"
         using Q_sets \<open>A \<in> sets M\<close> Q0(1) by auto
@@ -787,7 +799,7 @@
     have N: "?N \<in> sets M" using borel by simp
     have "?P g ?N \<le> integral\<^sup>N M g" using pos
       by (intro nn_integral_mono_AE) (auto split: split_indicator)
-    then have Pg_fin: "?P g ?N \<noteq> \<infinity>" using g_fin by auto
+    then have Pg_fin: "?P g ?N \<noteq> \<infinity>" using g_fin by (auto simp: top_unique)
     have "?P (\<lambda>x. (f x - g x)) ?N = (\<integral>\<^sup>+x. f x * indicator ?N x - g x * indicator ?N x \<partial>M)"
       by (auto intro!: nn_integral_cong simp: indicator_def)
     also have "\<dots> = ?P f ?N - ?P g ?N"
@@ -795,15 +807,14 @@
       show "(\<lambda>x. f x * indicator ?N x) \<in> borel_measurable M" "(\<lambda>x. g x * indicator ?N x) \<in> borel_measurable M"
         using borel N by auto
       show "AE x in M. g x * indicator ?N x \<le> f x * indicator ?N x"
-           "AE x in M. 0 \<le> g x * indicator ?N x"
         using pos by (auto split: split_indicator)
     qed fact
     also have "\<dots> = 0"
-      unfolding eq[THEN bspec, OF N] using nn_integral_nonneg[of M] Pg_fin by auto
+      unfolding eq[THEN bspec, OF N] using Pg_fin by auto
     finally have "AE x in M. f x \<le> g x"
       using pos borel nn_integral_PInf_AE[OF borel(2) g_fin]
       by (subst (asm) nn_integral_0_iff_AE)
-         (auto split: split_indicator simp: not_less ereal_minus_le_iff) }
+         (auto split: split_indicator simp: not_less ennreal_minus_eq_0) }
   from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq
   show "AE x in M. f x = g x" by auto
 qed
@@ -833,31 +844,31 @@
   from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
   let ?D = "{x\<in>space M. f x \<noteq> f' x}"
   have "?D \<in> sets M" using borel by auto
-  have *: "\<And>i x A. \<And>y::ereal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
+  have *: "\<And>i x A. \<And>y::ennreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
     unfolding indicator_def by auto
   have "\<forall>i. AE x in M. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos
     by (intro finite_density_unique[THEN iffD1] allI)
        (auto intro!: f measure_eqI simp: emeasure_density * subset_eq)
   moreover have "AE x in M. ?f Q0 x = ?f' Q0 x"
   proof (rule AE_I')
-    { fix f :: "'a \<Rightarrow> ereal" assume borel: "f \<in> borel_measurable M"
+    { fix f :: "'a \<Rightarrow> ennreal" assume borel: "f \<in> borel_measurable M"
         and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?N A = (\<integral>\<^sup>+x. f x * indicator A x \<partial>M)"
       let ?A = "\<lambda>i. Q0 \<inter> {x \<in> space M. f x < (i::nat)}"
       have "(\<Union>i. ?A i) \<in> null_sets M"
       proof (rule null_sets_UN)
         fix i ::nat have "?A i \<in> sets M"
           using borel Q0(1) by auto
-        have "?N (?A i) \<le> (\<integral>\<^sup>+x. (i::ereal) * indicator (?A i) x \<partial>M)"
+        have "?N (?A i) \<le> (\<integral>\<^sup>+x. (i::ennreal) * indicator (?A i) x \<partial>M)"
           unfolding eq[OF \<open>?A i \<in> sets M\<close>]
           by (auto intro!: nn_integral_mono simp: indicator_def)
         also have "\<dots> = i * emeasure M (?A i)"
           using \<open>?A i \<in> sets M\<close> by (auto intro!: nn_integral_cmult_indicator)
-        also have "\<dots> < \<infinity>" using emeasure_real[of "?A i"] by simp
+        also have "\<dots> < \<infinity>" using emeasure_real[of "?A i"] by (auto simp: ennreal_mult_less_top of_nat_less_top)
         finally have "?N (?A i) \<noteq> \<infinity>" by simp
         then show "?A i \<in> null_sets M" using in_Q0[OF \<open>?A i \<in> sets M\<close>] \<open>?A i \<in> sets M\<close> by auto
       qed
       also have "(\<Union>i. ?A i) = Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>}"
-        by (auto simp: less_PInf_Ex_of_nat)
+        by (auto simp: ennreal_Ex_less_of_nat less_top[symmetric])
       finally have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets M" by simp }
     from this[OF borel(1) refl] this[OF borel(2) f]
     have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets M" "Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>} \<in> null_sets M" by simp_all
@@ -875,8 +886,8 @@
 qed
 
 lemma (in sigma_finite_measure) density_unique:
-  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
-  assumes f': "f' \<in> borel_measurable M" "AE x in M. 0 \<le> f' x"
+  assumes f: "f \<in> borel_measurable M"
+  assumes f': "f' \<in> borel_measurable M"
   assumes density_eq: "density M f = density M f'"
   shows "AE x in M. f x = f' x"
 proof -
@@ -911,23 +922,18 @@
       using \<open>A \<in> sets M\<close> h_borel h_nn f f'
       by (subst (asm) (1 2) nn_integral_density[symmetric]) auto }
   then have "AE x in ?H. f x = f' x" using h_borel h_nn f f'
-    by (intro h.density_unique_finite_measure absolutely_continuous_AE[of M])
-       (auto simp add: AE_density)
-  then show "AE x in M. f x = f' x"
-    unfolding eventually_ae_filter using h_borel pos
-    by (auto simp add: h_null_sets null_sets_density_iff not_less[symmetric]
-                          AE_iff_null_sets[symmetric]) blast
+    by (intro h.density_unique_finite_measure absolutely_continuous_AE[of M]) auto
+  with AE_space[of M] pos show "AE x in M. f x = f' x"
+    unfolding AE_density[OF h_borel] by auto
 qed
 
 lemma (in sigma_finite_measure) density_unique_iff:
-  assumes f: "f \<in> borel_measurable M" and "AE x in M. 0 \<le> f x"
-  assumes f': "f' \<in> borel_measurable M" and "AE x in M. 0 \<le> f' x"
+  assumes f: "f \<in> borel_measurable M" and f': "f' \<in> borel_measurable M"
   shows "density M f = density M f' \<longleftrightarrow> (AE x in M. f x = f' x)"
   using density_unique[OF assms] density_cong[OF f f'] by auto
 
 lemma sigma_finite_density_unique:
   assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x"
   and fin: "sigma_finite_measure (density M f)"
   shows "density M f = density M g \<longleftrightarrow> (AE x in M. f x = g x)"
 proof
@@ -947,9 +953,8 @@
     moreover have "(\<integral>\<^sup>+x. f x * indicator (A i) x \<partial>M) \<noteq> \<infinity>"
       using cover(1) cover(3)[of i] borel by (auto simp: emeasure_density subset_eq)
     ultimately have "AE x in M. f x * indicator (A i) x = g x * indicator (A i) x"
-      using borel pos cover(1) pos
-      by (intro finite_density_unique[THEN iffD1])
-         (auto simp: density_density_eq subset_eq)
+      using borel cover(1)
+      by (intro finite_density_unique[THEN iffD1]) (auto simp: density_density_eq subset_eq)
     then show "AE x in M. x \<in> A i \<longrightarrow> f x = g x"
       by auto
   qed
@@ -961,7 +966,7 @@
 qed
 
 lemma (in sigma_finite_measure) sigma_finite_iff_density_finite':
-  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  assumes f: "f \<in> borel_measurable M"
   shows "sigma_finite_measure (density M f) \<longleftrightarrow> (AE x in M. f x \<noteq> \<infinity>)"
     (is "sigma_finite_measure ?N \<longleftrightarrow> _")
 proof
@@ -969,23 +974,23 @@
   then interpret N: sigma_finite_measure ?N .
   from N.Ex_finite_integrable_function obtain h where
     h: "h \<in> borel_measurable M" "integral\<^sup>N ?N h \<noteq> \<infinity>" and
-    h_nn: "\<And>x. 0 \<le> h x" and
-    fin: "\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>" by auto
+    fin: "\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>"
+    by auto
   have "AE x in M. f x * h x \<noteq> \<infinity>"
   proof (rule AE_I')
-    have "integral\<^sup>N ?N h = (\<integral>\<^sup>+x. f x * h x \<partial>M)" using f h h_nn
-      by (auto intro!: nn_integral_density)
+    have "integral\<^sup>N ?N h = (\<integral>\<^sup>+x. f x * h x \<partial>M)"
+      using f h by (auto intro!: nn_integral_density)
     then have "(\<integral>\<^sup>+x. f x * h x \<partial>M) \<noteq> \<infinity>"
       using h(2) by simp
     then show "(\<lambda>x. f x * h x) -` {\<infinity>} \<inter> space M \<in> null_sets M"
-      using f h(1) by (auto intro!: nn_integral_PInf borel_measurable_vimage)
+      using f h(1) by (auto intro!: nn_integral_PInf[unfolded infinity_ennreal_def] borel_measurable_vimage)
   qed auto
   then show "AE x in M. f x \<noteq> \<infinity>"
-    using fin by (auto elim!: AE_Ball_mp)
+    using fin by (auto elim!: AE_Ball_mp simp: less_top ennreal_mult_less_top)
 next
   assume AE: "AE x in M. f x \<noteq> \<infinity>"
   from sigma_finite guess Q . note Q = this
-  def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. ereal(of_nat (Suc n))}) \<inter> space M"
+  def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. ennreal(of_nat (Suc n))}) \<inter> space M"
   { fix i j have "A i \<inter> Q j \<in> sets M"
     unfolding A_def using f Q
     apply (rule_tac sets.Int)
@@ -1006,15 +1011,16 @@
     proof safe
       fix x assume x: "x \<in> space M"
       show "x \<in> (\<Union>i. A i)"
-      proof (cases "f x")
-        case PInf with x show ?thesis unfolding A_def by (auto intro: exI[of _ 0])
+      proof (cases "f x" rule: ennreal_cases)
+        case top with x show ?thesis unfolding A_def by (auto intro: exI[of _ 0])
       next
         case (real r)
-        with less_PInf_Ex_of_nat[of "f x"] obtain n :: nat where "f x < n" by auto
-        then show ?thesis using x real unfolding A_def by (auto intro!: exI[of _ "Suc n"])
-      next
-        case MInf with x show ?thesis
-          unfolding A_def by (auto intro!: exI[of _ "Suc 0"])
+        with ennreal_Ex_less_of_nat[of "f x"] obtain n :: nat where "f x < n"
+          by auto
+        also have "n < (Suc n :: ennreal)"
+          by simp
+        finally show ?thesis
+          using x real by (auto simp: A_def ennreal_of_nat_eq_real_of_nat intro!: exI[of _ "Suc n"])
       qed
     qed (auto simp: A_def)
     finally show "\<Union>range (\<lambda>(i, j). A i \<inter> Q j) = space ?N" by simp
@@ -1030,12 +1036,12 @@
     next
       case (Suc n)
       then have "(\<integral>\<^sup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<le>
-        (\<integral>\<^sup>+x. (Suc n :: ereal) * indicator (Q j) x \<partial>M)"
-        by (auto intro!: nn_integral_mono simp: indicator_def A_def)
+        (\<integral>\<^sup>+x. (Suc n :: ennreal) * indicator (Q j) x \<partial>M)"
+        by (auto intro!: nn_integral_mono simp: indicator_def A_def ennreal_of_nat_eq_real_of_nat)
       also have "\<dots> = Suc n * emeasure M (Q j)"
         using Q by (auto intro!: nn_integral_cmult_indicator)
       also have "\<dots> < \<infinity>"
-        using Q by auto
+        using Q by (auto simp: ennreal_mult_less_top less_top of_nat_less_top)
       finally show ?thesis by simp
     qed
     then show "emeasure ?N X \<noteq> \<infinity>"
@@ -1045,55 +1051,42 @@
 
 lemma (in sigma_finite_measure) sigma_finite_iff_density_finite:
   "f \<in> borel_measurable M \<Longrightarrow> sigma_finite_measure (density M f) \<longleftrightarrow> (AE x in M. f x \<noteq> \<infinity>)"
-  apply (subst density_max_0)
-  apply (subst sigma_finite_iff_density_finite')
-  apply (auto simp: max_def intro!: measurable_If)
-  done
+  by (subst sigma_finite_iff_density_finite')
+     (auto simp: max_def intro!: measurable_If)
 
 subsection \<open>Radon-Nikodym derivative\<close>
 
-definition RN_deriv :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a \<Rightarrow> ereal" where
+definition RN_deriv :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a \<Rightarrow> ennreal" where
   "RN_deriv M N =
-    (if \<exists>f. f \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> density M f = N
-       then SOME f. f \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> density M f = N
+    (if \<exists>f. f \<in> borel_measurable M \<and> density M f = N
+       then SOME f. f \<in> borel_measurable M \<and> density M f = N
        else (\<lambda>_. 0))"
 
 lemma RN_derivI:
-  assumes "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" "density M f = N"
+  assumes "f \<in> borel_measurable M" "density M f = N"
   shows "density M (RN_deriv M N) = N"
 proof -
-  have "\<exists>f. f \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> density M f = N"
+  have "\<exists>f. f \<in> borel_measurable M \<and> density M f = N"
     using assms by auto
-  moreover then have "density M (SOME f. f \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> density M f = N) = N"
+  moreover then have "density M (SOME f. f \<in> borel_measurable M \<and> density M f = N) = N"
     by (rule someI2_ex) auto
   ultimately show ?thesis
     by (auto simp: RN_deriv_def)
 qed
 
-lemma
-  shows borel_measurable_RN_deriv[measurable]: "RN_deriv M N \<in> borel_measurable M" (is ?m)
-    and RN_deriv_nonneg: "0 \<le> RN_deriv M N x" (is ?nn)
+lemma borel_measurable_RN_deriv[measurable]: "RN_deriv M N \<in> borel_measurable M"
 proof -
-  { assume ex: "\<exists>f. f \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> density M f = N"
-    have 1: "(SOME f. f \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> density M f = N) \<in> borel_measurable M"
-      using ex by (rule someI2_ex) auto
-    moreover
-    have 2: "0 \<le> (SOME f. f \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> density M f = N) x"
-      using ex by (rule someI2_ex) auto
-    note 1 2 }
-  from this show ?m ?nn
+  { assume ex: "\<exists>f. f \<in> borel_measurable M \<and> density M f = N"
+    have 1: "(SOME f. f \<in> borel_measurable M \<and> density M f = N) \<in> borel_measurable M"
+      using ex by (rule someI2_ex) auto }
+  from this show ?thesis
     by (auto simp: RN_deriv_def)
 qed
 
 lemma density_RN_deriv_density:
-  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  assumes f: "f \<in> borel_measurable M"
   shows "density M (RN_deriv M (density M f)) = density M f"
-proof (rule RN_derivI)
-  show "(\<lambda>x. max 0 (f x)) \<in> borel_measurable M" "\<And>x. 0 \<le> max 0 (f x)"
-    using f by auto
-  show "density M (\<lambda>x. max 0 (f x)) = density M f"
-    using f by (intro density_cong) (auto simp: max_def)
-qed
+  by (rule RN_derivI[OF f]) simp
 
 lemma (in sigma_finite_measure) density_RN_deriv:
   "absolutely_continuous M N \<Longrightarrow> sets N = sets M \<Longrightarrow> density M (RN_deriv M N) = N"
@@ -1107,7 +1100,7 @@
   have "integral\<^sup>N N f = integral\<^sup>N (density M (RN_deriv M N)) f"
     using N by (simp add: density_RN_deriv)
   also have "\<dots> = (\<integral>\<^sup>+x. RN_deriv M N x * f x \<partial>M)"
-    using f by (simp add: nn_integral_density RN_deriv_nonneg)
+    using f by (simp add: nn_integral_density)
   finally show ?thesis by simp
 qed
 
@@ -1115,20 +1108,20 @@
   using AE_iff_null_sets[of N M] by auto
 
 lemma (in sigma_finite_measure) RN_deriv_unique:
-  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  assumes f: "f \<in> borel_measurable M"
   and eq: "density M f = N"
   shows "AE x in M. f x = RN_deriv M N x"
   unfolding eq[symmetric]
   by (intro density_unique_iff[THEN iffD1] f borel_measurable_RN_deriv
-            RN_deriv_nonneg[THEN AE_I2] density_RN_deriv_density[symmetric])
+            density_RN_deriv_density[symmetric])
 
 lemma RN_deriv_unique_sigma_finite:
-  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  assumes f: "f \<in> borel_measurable M"
   and eq: "density M f = N" and fin: "sigma_finite_measure N"
   shows "AE x in M. f x = RN_deriv M N x"
   using fin unfolding eq[symmetric]
   by (intro sigma_finite_density_unique[THEN iffD1] f borel_measurable_RN_deriv
-            RN_deriv_nonneg[THEN AE_I2] density_RN_deriv_density[symmetric])
+            density_RN_deriv_density[symmetric])
 
 lemma (in sigma_finite_measure) RN_deriv_distr:
   fixes T :: "'a \<Rightarrow> 'b"
@@ -1175,7 +1168,6 @@
     using T ac by measurable
   then show "(\<lambda>x. RN_deriv ?M' ?N' (T x)) \<in> borel_measurable M"
     by (simp add: comp_def)
-  show "AE x in M. 0 \<le> RN_deriv ?M' ?N' (T x)" by (auto intro: RN_deriv_nonneg)
 
   have "N = distr N M (T' \<circ> T)"
     by (subst measure_of_of_measure[of N, symmetric])
@@ -1196,39 +1188,38 @@
 proof -
   interpret N: sigma_finite_measure N by fact
   from N show ?thesis
-    using sigma_finite_iff_density_finite[OF borel_measurable_RN_deriv, of N]
-      density_RN_deriv[OF ac]
-    by (simp add: RN_deriv_nonneg)
+    using sigma_finite_iff_density_finite[OF borel_measurable_RN_deriv, of N] density_RN_deriv[OF ac]
+    by simp
 qed
 
 lemma (in sigma_finite_measure)
   assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M"
     and f: "f \<in> borel_measurable M"
   shows RN_deriv_integrable: "integrable N f \<longleftrightarrow>
-      integrable M (\<lambda>x. real_of_ereal (RN_deriv M N x) * f x)" (is ?integrable)
-    and RN_deriv_integral: "integral\<^sup>L N f = (\<integral>x. real_of_ereal (RN_deriv M N x) * f x \<partial>M)" (is ?integral)
+      integrable M (\<lambda>x. enn2real (RN_deriv M N x) * f x)" (is ?integrable)
+    and RN_deriv_integral: "integral\<^sup>L N f = (\<integral>x. enn2real (RN_deriv M N x) * f x \<partial>M)" (is ?integral)
 proof -
   note ac(2)[simp] and sets_eq_imp_space_eq[OF ac(2), simp]
   interpret N: sigma_finite_measure N by fact
 
-  have eq: "density M (RN_deriv M N) = density M (\<lambda>x. real_of_ereal (RN_deriv M N x))"
+  have eq: "density M (RN_deriv M N) = density M (\<lambda>x. enn2real (RN_deriv M N x))"
   proof (rule density_cong)
     from RN_deriv_finite[OF assms(1,2,3)]
-    show "AE x in M. RN_deriv M N x = ereal (real_of_ereal (RN_deriv M N x))"
-      by eventually_elim (insert RN_deriv_nonneg[of M N], auto simp: ereal_real)
+    show "AE x in M. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))"
+      by eventually_elim (auto simp: less_top)
   qed (insert ac, auto)
 
   show ?integrable
     apply (subst density_RN_deriv[OF ac, symmetric])
     unfolding eq
-    apply (intro integrable_real_density f AE_I2 real_of_ereal_pos RN_deriv_nonneg)
+    apply (intro integrable_real_density f AE_I2 enn2real_nonneg)
     apply (insert ac, auto)
     done
 
   show ?integral
     apply (subst density_RN_deriv[OF ac, symmetric])
     unfolding eq
-    apply (intro integral_real_density f AE_I2 real_of_ereal_pos RN_deriv_nonneg)
+    apply (intro integral_real_density f AE_I2 enn2real_nonneg)
     apply (insert ac, auto)
     done
 qed
@@ -1237,21 +1228,21 @@
   assumes "finite_measure N"
   assumes ac: "absolutely_continuous M N" "sets N = sets M"
   obtains D where "D \<in> borel_measurable M"
-    and "AE x in M. RN_deriv M N x = ereal (D x)"
+    and "AE x in M. RN_deriv M N x = ennreal (D x)"
     and "AE x in N. 0 < D x"
     and "\<And>x. 0 \<le> D x"
 proof
   interpret N: finite_measure N by fact
 
-  note RN = borel_measurable_RN_deriv density_RN_deriv[OF ac] RN_deriv_nonneg[of M N]
+  note RN = borel_measurable_RN_deriv density_RN_deriv[OF ac]
 
   let ?RN = "\<lambda>t. {x \<in> space M. RN_deriv M N x = t}"
 
-  show "(\<lambda>x. real_of_ereal (RN_deriv M N x)) \<in> borel_measurable M"
+  show "(\<lambda>x. enn2real (RN_deriv M N x)) \<in> borel_measurable M"
     using RN by auto
 
   have "N (?RN \<infinity>) = (\<integral>\<^sup>+ x. RN_deriv M N x * indicator (?RN \<infinity>) x \<partial>M)"
-    using RN(1,3) by (subst RN(2)[symmetric]) (auto simp: emeasure_density)
+    using RN(1) by (subst RN(2)[symmetric]) (auto simp: emeasure_density)
   also have "\<dots> = (\<integral>\<^sup>+ x. \<infinity> * indicator (?RN \<infinity>) x \<partial>M)"
     by (intro nn_integral_cong) (auto simp: indicator_def)
   also have "\<dots> = \<infinity> * emeasure M (?RN \<infinity>)"
@@ -1261,30 +1252,28 @@
   have "emeasure M (?RN \<infinity>) = 0"
   proof (rule ccontr)
     assume "emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>} \<noteq> 0"
-    moreover from RN have "0 \<le> emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>}" by auto
-    ultimately have "0 < emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>}" by auto
-    with eq have "N (?RN \<infinity>) = \<infinity>" by simp
+    then have "0 < emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>}"
+      by (auto simp: zero_less_iff_neq_zero)
+    with eq have "N (?RN \<infinity>) = \<infinity>" by (simp add: ennreal_mult_eq_top_iff)
     with N.emeasure_finite[of "?RN \<infinity>"] RN show False by auto
   qed
   ultimately have "AE x in M. RN_deriv M N x < \<infinity>"
-    using RN by (intro AE_iff_measurable[THEN iffD2]) auto
-  then show "AE x in M. RN_deriv M N x = ereal (real_of_ereal (RN_deriv M N x))"
-    using RN(3) by (auto simp: ereal_real)
-  then have eq: "AE x in N. RN_deriv M N x = ereal (real_of_ereal (RN_deriv M N x))"
+    using RN by (intro AE_iff_measurable[THEN iffD2]) (auto simp: less_top[symmetric])
+  then show "AE x in M. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))"
+    by auto
+  then have eq: "AE x in N. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))"
     using ac absolutely_continuous_AE by auto
 
-  show "\<And>x. 0 \<le> real_of_ereal (RN_deriv M N x)"
-    using RN by (auto intro: real_of_ereal_pos)
 
   have "N (?RN 0) = (\<integral>\<^sup>+ x. RN_deriv M N x * indicator (?RN 0) x \<partial>M)"
-    using RN(1,3) by (subst RN(2)[symmetric]) (auto simp: emeasure_density)
+    by (subst RN(2)[symmetric]) (auto simp: emeasure_density)
   also have "\<dots> = (\<integral>\<^sup>+ x. 0 \<partial>M)"
     by (intro nn_integral_cong) (auto simp: indicator_def)
   finally have "AE x in N. RN_deriv M N x \<noteq> 0"
     using RN by (subst AE_iff_measurable[OF _ refl]) (auto simp: ac cong: sets_eq_imp_space_eq)
-  with RN(3) eq show "AE x in N. 0 < real_of_ereal (RN_deriv M N x)"
-    by (auto simp: zero_less_real_of_ereal le_less)
-qed
+  with eq show "AE x in N. 0 < enn2real (RN_deriv M N x)"
+    by (auto simp: enn2real_positive_iff less_top[symmetric] zero_less_iff_neq_zero)
+qed (rule enn2real_nonneg)
 
 lemma (in sigma_finite_measure) RN_deriv_singleton:
   assumes ac: "absolutely_continuous M N" "sets N = sets M"
@@ -1294,7 +1283,7 @@
   from \<open>{x} \<in> sets M\<close>
   have "density M (RN_deriv M N) {x} = (\<integral>\<^sup>+w. RN_deriv M N x * indicator {x} w \<partial>M)"
     by (auto simp: indicator_def emeasure_density intro!: nn_integral_cong)
-  with x density_RN_deriv[OF ac] RN_deriv_nonneg[of M N] show ?thesis
+  with x density_RN_deriv[OF ac] show ?thesis
     by (auto simp: max_def)
 qed
 
--- a/src/HOL/Probability/Regularity.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Regularity.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -8,101 +8,6 @@
 imports Measure_Space Borel_Space
 begin
 
-lemma ereal_approx_SUP:
-  fixes x::ereal
-  assumes A_notempty: "A \<noteq> {}"
-  assumes f_bound: "\<And>i. i \<in> A \<Longrightarrow> f i \<le> x"
-  assumes f_fin: "\<And>i. i \<in> A \<Longrightarrow> f i \<noteq> \<infinity>"
-  assumes f_nonneg: "\<And>i. 0 \<le> f i"
-  assumes approx: "\<And>e. (e::real) > 0 \<Longrightarrow> \<exists>i \<in> A. x \<le> f i + e"
-  shows "x = (SUP i : A. f i)"
-proof (subst eq_commute, rule SUP_eqI)
-  show "\<And>i. i \<in> A \<Longrightarrow> f i \<le> x" using f_bound by simp
-next
-  fix y :: ereal assume f_le_y: "(\<And>i::'a. i \<in> A \<Longrightarrow> f i \<le> y)"
-  with A_notempty f_nonneg have "y \<ge> 0" by auto (metis order_trans)
-  show "x \<le> y"
-  proof (rule ccontr)
-    assume "\<not> x \<le> y" hence "x > y" by simp
-    hence y_fin: "\<bar>y\<bar> \<noteq> \<infinity>" using \<open>y \<ge> 0\<close> by auto
-    have x_fin: "\<bar>x\<bar> \<noteq> \<infinity>" using \<open>x > y\<close> f_fin approx[where e = 1] by auto
-    def e \<equiv> "real_of_ereal ((x - y) / 2)"
-    have e: "x > y + e" "e > 0" using \<open>x > y\<close> y_fin x_fin by (auto simp: e_def field_simps)
-    note e(1)
-    also from approx[OF \<open>e > 0\<close>] obtain i where i: "i \<in> A" "x \<le> f i + e" by blast
-    note i(2)
-    finally have "y < f i" using y_fin f_fin by (metis add_right_mono linorder_not_le)
-    moreover have "f i \<le> y" by (rule f_le_y) fact
-    ultimately show False by simp
-  qed
-qed
-
-lemma ereal_approx_INF:
-  fixes x::ereal
-  assumes A_notempty: "A \<noteq> {}"
-  assumes f_bound: "\<And>i. i \<in> A \<Longrightarrow> x \<le> f i"
-  assumes f_fin: "\<And>i. i \<in> A \<Longrightarrow> f i \<noteq> \<infinity>"
-  assumes f_nonneg: "\<And>i. 0 \<le> f i"
-  assumes approx: "\<And>e. (e::real) > 0 \<Longrightarrow> \<exists>i \<in> A. f i \<le> x + e"
-  shows "x = (INF i : A. f i)"
-proof (subst eq_commute, rule INF_eqI)
-  show "\<And>i. i \<in> A \<Longrightarrow> x \<le> f i" using f_bound by simp
-next
-  fix y :: ereal assume f_le_y: "(\<And>i::'a. i \<in> A \<Longrightarrow> y \<le> f i)"
-  with A_notempty f_fin have "y \<noteq> \<infinity>" by force
-  show "y \<le> x"
-  proof (rule ccontr)
-    assume "\<not> y \<le> x" hence "y > x" by simp hence "y \<noteq> - \<infinity>" by auto
-    hence y_fin: "\<bar>y\<bar> \<noteq> \<infinity>" using \<open>y \<noteq> \<infinity>\<close> by auto
-    have x_fin: "\<bar>x\<bar> \<noteq> \<infinity>" using \<open>y > x\<close> f_fin f_nonneg approx[where e = 1] A_notempty
-      by auto
-    def e \<equiv> "real_of_ereal ((y - x) / 2)"
-    have e: "y > x + e" "e > 0" using \<open>y > x\<close> y_fin x_fin by (auto simp: e_def field_simps)
-    from approx[OF \<open>e > 0\<close>] obtain i where i: "i \<in> A" "x + e \<ge> f i" by blast
-    note i(2)
-    also note e(1)
-    finally have "y > f i" .
-    moreover have "y \<le> f i" by (rule f_le_y) fact
-    ultimately show False by simp
-  qed
-qed
-
-lemma INF_approx_ereal:
-  fixes x::ereal and e::real
-  assumes "e > 0"
-  assumes INF: "x = (INF i : A. f i)"
-  assumes "\<bar>x\<bar> \<noteq> \<infinity>"
-  shows "\<exists>i \<in> A. f i < x + e"
-proof (rule ccontr, clarsimp)
-  assume "\<forall>i\<in>A. \<not> f i < x + e"
-  moreover
-  from INF have "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> f i) \<Longrightarrow> y \<le> x" by (auto intro: INF_greatest)
-  ultimately
-  have "(INF i : A. f i) = x + e" using \<open>e > 0\<close>
-    by (intro INF_eqI)
-      (force, metis add.comm_neutral add_left_mono ereal_less(1)
-        linorder_not_le not_less_iff_gr_or_eq)
-  thus False using assms by auto
-qed
-
-lemma SUP_approx_ereal:
-  fixes x::ereal and e::real
-  assumes "e > 0"
-  assumes SUP: "x = (SUP i : A. f i)"
-  assumes "\<bar>x\<bar> \<noteq> \<infinity>"
-  shows "\<exists>i \<in> A. x \<le> f i + e"
-proof (rule ccontr, clarsimp)
-  assume "\<forall>i\<in>A. \<not> x \<le> f i + e"
-  moreover
-  from SUP have "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> f i \<le> y) \<Longrightarrow> y \<ge> x" by (auto intro: SUP_least)
-  ultimately
-  have "(SUP i : A. f i) = x - e" using \<open>e > 0\<close> \<open>\<bar>x\<bar> \<noteq> \<infinity>\<close>
-    by (intro SUP_eqI)
-       (metis PInfty_neq_ereal(2) abs_ereal.simps(1) ereal_minus_le linorder_linear,
-        metis ereal_between(1) ereal_less(2) less_eq_ereal_def order_trans)
-  thus False using assms by auto
-qed
-
 lemma
   fixes M::"'a::{second_countable_topology, complete_space} measure"
   assumes sb: "sets M = sets borel"
@@ -117,12 +22,12 @@
   hence sU: "space M = UNIV" by simp
   interpret finite_measure M by rule fact
   have approx_inner: "\<And>A. A \<in> sets M \<Longrightarrow>
-    (\<And>e. e > 0 \<Longrightarrow> \<exists>K. K \<subseteq> A \<and> compact K \<and> emeasure M A \<le> emeasure M K + ereal e) \<Longrightarrow> ?inner A"
-    by (rule ereal_approx_SUP)
+    (\<And>e. e > 0 \<Longrightarrow> \<exists>K. K \<subseteq> A \<and> compact K \<and> emeasure M A \<le> emeasure M K + ennreal e) \<Longrightarrow> ?inner A"
+    by (rule ennreal_approx_SUP)
       (force intro!: emeasure_mono simp: compact_imp_closed emeasure_eq_measure)+
   have approx_outer: "\<And>A. A \<in> sets M \<Longrightarrow>
-    (\<And>e. e > 0 \<Longrightarrow> \<exists>B. A \<subseteq> B \<and> open B \<and> emeasure M B \<le> emeasure M A + ereal e) \<Longrightarrow> ?outer A"
-    by (rule ereal_approx_INF)
+    (\<And>e. e > 0 \<Longrightarrow> \<exists>B. A \<subseteq> B \<and> open B \<and> emeasure M B \<le> emeasure M A + ennreal e) \<Longrightarrow> ?outer A"
+    by (rule ennreal_approx_INF)
        (force intro!: emeasure_mono simp: emeasure_eq_measure sb)+
   from countable_dense_setE guess X::"'a set"  . note X = this
   {
@@ -132,8 +37,7 @@
       by (auto simp add: sU) (metis dist_commute order_less_imp_le)
     let ?U = "\<Union>k. (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)"
     have "(\<lambda>k. emeasure M (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)) \<longlonglongrightarrow> M ?U"
-      by (rule Lim_emeasure_incseq)
-        (auto intro!: borel_closed bexI simp: closed_cball incseq_def Us sb)
+      by (rule Lim_emeasure_incseq) (auto intro!: borel_closed bexI simp: incseq_def Us sb)
     also have "?U = space M"
     proof safe
       fix x from X(2)[OF open_ball[of x r]] \<open>r > 0\<close> obtain d where d: "d\<in>X" "d \<in> ball x r" by auto
@@ -148,7 +52,7 @@
     hence "1/n > 0" "e * 2 powr - n > 0" by (auto)
     from M_space[OF \<open>1/n>0\<close>]
     have "(\<lambda>k. measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n))) \<longlonglongrightarrow> measure M (space M)"
-      unfolding emeasure_eq_measure by simp
+      unfolding emeasure_eq_measure by (auto simp: measure_nonneg)
     from metric_LIMSEQ_D[OF this \<open>0 < e * 2 powr -n\<close>]
     obtain k where "dist (measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n))) (measure M (space M)) <
       e * 2 powr -n"
@@ -169,13 +73,11 @@
     \<le> measure M (\<Union>i\<in>{0..k e n}. cball (from_nat_into X i) (1 / n))"
     unfolding Ball_def by blast
   have approx_space:
-    "\<And>e. e > 0 \<Longrightarrow>
-      \<exists>K \<in> {K. K \<subseteq> space M \<and> compact K}. emeasure M (space M) \<le> emeasure M K + ereal e"
-      (is "\<And>e. _ \<Longrightarrow> ?thesis e")
+    "\<exists>K \<in> {K. K \<subseteq> space M \<and> compact K}. emeasure M (space M) \<le> emeasure M K + ennreal e"
+    (is "?thesis e") if "0 < e" for e :: real
   proof -
-    fix e :: real assume "e > 0"
     def B \<equiv> "\<lambda>n. \<Union>i\<in>{0..k e (Suc n)}. cball (from_nat_into X i) (1 / Suc n)"
-    have "\<And>n. closed (B n)" by (auto simp: B_def closed_cball)
+    have "\<And>n. closed (B n)" by (auto simp: B_def)
     hence [simp]: "\<And>n. B n \<in> sets M" by (simp add: sb)
     from k[OF \<open>e > 0\<close> zero_less_Suc]
     have "\<And>n. measure M (space M) - measure M (B n) \<le> e * 2 powr - real (Suc n)"
@@ -190,18 +92,21 @@
     also have "\<dots> = emeasure M (\<Union>n. space M - B n)" by (auto simp: K_def emeasure_eq_measure)
     also have "\<dots> \<le> (\<Sum>n. emeasure M (space M - B n))"
       by (rule emeasure_subadditive_countably) (auto simp: summable_def)
-    also have "\<dots> \<le> (\<Sum>n. ereal (e*2 powr - real (Suc n)))"
-      using B_compl_le by (intro suminf_le_pos) (simp_all add: measure_nonneg emeasure_eq_measure)
-    also have "\<dots> \<le> (\<Sum>n. ereal (e * (1 / 2) ^ Suc n))"
-      by (simp add: Transcendental.powr_minus powr_realpow field_simps del: of_nat_Suc)
-    also have "\<dots> = (\<Sum>n. ereal e * ((1 / 2) ^ Suc n))"
-      unfolding times_ereal.simps[symmetric] ereal_power[symmetric] one_ereal_def numeral_eq_ereal
-      by simp
-    also have "\<dots> = ereal e * (\<Sum>n. ((1 / 2) ^ Suc n))"
-      by (rule suminf_cmult_ereal) (auto simp: \<open>0 < e\<close> less_imp_le)
-    also have "\<dots> = e" unfolding suminf_half_series_ereal by simp
-    finally have "measure M (space M) \<le> measure M K + e" by simp
-    hence "emeasure M (space M) \<le> emeasure M K + e" by (simp add: emeasure_eq_measure)
+    also have "\<dots> \<le> (\<Sum>n. ennreal (e*2 powr - real (Suc n)))"
+      using B_compl_le by (intro suminf_le) (simp_all add: measure_nonneg emeasure_eq_measure ennreal_leI)
+    also have "\<dots> \<le> (\<Sum>n. ennreal (e * (1 / 2) ^ Suc n))"
+      by (simp add: powr_minus powr_realpow field_simps del: of_nat_Suc)
+    also have "\<dots> = ennreal e * (\<Sum>n. ennreal ((1 / 2) ^ Suc n))"
+      unfolding ennreal_power[symmetric]
+      using \<open>0 < e\<close>
+      by (simp add: ac_simps ennreal_mult' divide_ennreal[symmetric] divide_ennreal_def
+                    ennreal_power[symmetric])
+    also have "\<dots> = e"
+      by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto
+    finally have "measure M (space M) \<le> measure M K + e"
+      using \<open>0 < e\<close> by simp
+    hence "emeasure M (space M) \<le> emeasure M K + e"
+      using \<open>0 < e\<close> by (simp add: emeasure_eq_measure measure_nonneg ennreal_plus[symmetric] del: ennreal_plus)
     moreover have "compact K"
       unfolding compact_eq_totally_bounded
     proof safe
@@ -214,7 +119,7 @@
       ultimately show "\<exists>k. finite k \<and> K \<subseteq> (\<Union>x\<in>k. ball x e')" by blast
     qed
     ultimately
-    show "?thesis e " by (auto simp: sU)
+    show ?thesis by (auto simp: sU)
   qed
   { fix A::"'a set" assume "closed A" hence "A \<in> sets borel" by (simp add: compact_imp_closed)
     hence [simp]: "A \<in> sets M" by (simp add: sb)
@@ -225,20 +130,19 @@
         K: "K \<subseteq> space M" "compact K" "emeasure M (space M) \<le> emeasure M K + e"
         by (auto simp: emeasure_eq_measure)
       hence [simp]: "K \<in> sets M" by (simp add: sb compact_imp_closed)
-      have "M A - M (A \<inter> K) = measure M A - measure M (A \<inter> K)"
-        by (simp add: emeasure_eq_measure)
-      also have "\<dots> = measure M (A - A \<inter> K)"
+      have "measure M A - measure M (A \<inter> K) = measure M (A - A \<inter> K)"
         by (subst finite_measure_Diff) auto
       also have "A - A \<inter> K = A \<union> K - K" by auto
       also have "measure M \<dots> = measure M (A \<union> K) - measure M K"
         by (subst finite_measure_Diff) auto
       also have "\<dots> \<le> measure M (space M) - measure M K"
         by (simp add: emeasure_eq_measure sU sb finite_measure_mono)
-      also have "\<dots> \<le> e" using K by (simp add: emeasure_eq_measure)
-      finally have "emeasure M A \<le> emeasure M (A \<inter> K) + ereal e"
-        by (simp add: emeasure_eq_measure algebra_simps)
+      also have "\<dots> \<le> e"
+        using K \<open>0 < e\<close> by (simp add: emeasure_eq_measure ennreal_plus[symmetric] measure_nonneg del: ennreal_plus)
+      finally have "emeasure M A \<le> emeasure M (A \<inter> K) + ennreal e"
+        using \<open>0<e\<close> by (simp add: emeasure_eq_measure algebra_simps ennreal_plus[symmetric] measure_nonneg del: ennreal_plus)
       moreover have "A \<inter> K \<subseteq> A" "compact (A \<inter> K)" using \<open>closed A\<close> \<open>compact K\<close> by auto
-      ultimately show "\<exists>K \<subseteq> A. compact K \<and> emeasure M A \<le> emeasure M K + ereal e"
+      ultimately show "\<exists>K \<subseteq> A. compact K \<and> emeasure M A \<le> emeasure M K + ennreal e"
         by blast
     qed simp
     have "?outer A"
@@ -313,7 +217,7 @@
     case 2
     have "M (space M - B) = M (space M) - emeasure M B" by (auto simp: emeasure_compl)
     also have "\<dots> = (INF K:{K. K \<subseteq> B \<and> compact K}. M (space M) -  M K)"
-      unfolding inner by (subst INF_ereal_minus_right) force+
+      by (subst ennreal_SUP_const_minus) (auto simp: less_top[symmetric] inner)
     also have "\<dots> = (INF U:{U. U \<subseteq> B \<and> compact U}. M (space M - U))"
       by (rule INF_cong) (auto simp add: emeasure_compl sb compact_imp_closed)
     also have "\<dots> \<ge> (INF U:{U. U \<subseteq> B \<and> closed U}. M (space M - U))"
@@ -321,7 +225,7 @@
     also have "(INF U:{U. U \<subseteq> B \<and> closed U}. M (space M - U)) =
         (INF U:{U. space M - B \<subseteq> U \<and> open U}. emeasure M U)"
       unfolding INF_image [of _ "\<lambda>u. space M - u" _, symmetric, unfolded comp_def]
-        by (rule INF_cong) (auto simp add: sU open_Compl Compl_eq_Diff_UNIV [symmetric, simp])
+        by (rule INF_cong) (auto simp add: sU Compl_eq_Diff_UNIV [symmetric, simp])
     finally have
       "(INF U:{U. space M - B \<subseteq> U \<and> open U}. emeasure M U) \<le> emeasure M (space M - B)" .
     moreover have
@@ -332,7 +236,7 @@
     case 1
     have "M (space M - B) = M (space M) - emeasure M B" by (auto simp: emeasure_compl)
     also have "\<dots> = (SUP U: {U. B \<subseteq> U \<and> open U}. M (space M) -  M U)"
-      unfolding outer by (subst SUP_ereal_minus_right) auto
+      unfolding outer by (subst ennreal_INF_const_minus) auto
     also have "\<dots> = (SUP U:{U. B \<subseteq> U \<and> open U}. M (space M - U))"
       by (rule SUP_cong) (auto simp add: emeasure_compl sb compact_imp_closed)
     also have "\<dots> = (SUP K:{K. K \<subseteq> space M - B \<and> closed K}. emeasure M K)"
@@ -353,9 +257,9 @@
     then have "range D \<subseteq> sets M" by (auto simp: sb borel_eq_closed)
     with union have M[symmetric]: "(\<Sum>i. M (D i)) = M (\<Union>i. D i)" by (intro suminf_emeasure)
     also have "(\<lambda>n. \<Sum>i<n. M (D i)) \<longlonglongrightarrow> (\<Sum>i. M (D i))"
-      by (intro summable_LIMSEQ summable_ereal_pos emeasure_nonneg)
+      by (intro summable_LIMSEQ) auto
     finally have measure_LIMSEQ: "(\<lambda>n. \<Sum>i<n. measure M (D i)) \<longlonglongrightarrow> measure M (\<Union>i. D i)"
-      by (simp add: emeasure_eq_measure)
+      by (simp add: emeasure_eq_measure measure_nonneg setsum_nonneg)
     have "(\<Union>i. D i) \<in> sets M" using \<open>range D \<subseteq> sets M\<close> by auto
 
     case 1
@@ -368,22 +272,22 @@
       hence "\<exists>n0. \<bar>(\<Sum>i<n0. measure M (D i)) - measure M (\<Union>x. D x)\<bar> < e/2" by auto
       then obtain n0 where n0: "\<bar>(\<Sum>i<n0. measure M (D i)) - measure M (\<Union>i. D i)\<bar> < e/2"
         unfolding choice_iff by blast
-      have "ereal (\<Sum>i<n0. measure M (D i)) = (\<Sum>i<n0. M (D i))"
-        by (auto simp add: emeasure_eq_measure)
-      also have "\<dots> \<le> (\<Sum>i. M (D i))" by (rule suminf_upper) (auto simp: emeasure_nonneg)
+      have "ennreal (\<Sum>i<n0. measure M (D i)) = (\<Sum>i<n0. M (D i))"
+        by (auto simp add: emeasure_eq_measure setsum_nonneg measure_nonneg)
+      also have "\<dots> \<le> (\<Sum>i. M (D i))" by (rule setsum_le_suminf) auto
       also have "\<dots> = M (\<Union>i. D i)" by (simp add: M)
       also have "\<dots> = measure M (\<Union>i. D i)" by (simp add: emeasure_eq_measure)
       finally have n0: "measure M (\<Union>i. D i) - (\<Sum>i<n0. measure M (D i)) < e/2"
-        using n0 by auto
+        using n0 by (auto simp: measure_nonneg setsum_nonneg)
       have "\<forall>i. \<exists>K. K \<subseteq> D i \<and> compact K \<and> emeasure M (D i) \<le> emeasure M K + e/(2*Suc n0)"
       proof
         fix i
         from \<open>0 < e\<close> have "0 < e/(2*Suc n0)" by simp
         have "emeasure M (D i) = (SUP K:{K. K \<subseteq> (D i) \<and> compact K}. emeasure M K)"
           using union by blast
-        from SUP_approx_ereal[OF \<open>0 < e/(2*Suc n0)\<close> this]
+        from SUP_approx_ennreal[OF \<open>0 < e/(2*Suc n0)\<close> _ this]
         show "\<exists>K. K \<subseteq> D i \<and> compact K \<and> emeasure M (D i) \<le> emeasure M K + e/(2*Suc n0)"
-          by (auto simp: emeasure_eq_measure)
+          by (auto simp: emeasure_eq_measure intro: less_imp_le compact_empty)
       qed
       then obtain K where K: "\<And>i. K i \<subseteq> D i" "\<And>i. compact (K i)"
         "\<And>i. emeasure M (D i) \<le> emeasure M (K i) + e/(2*Suc n0)"
@@ -395,7 +299,8 @@
         by (intro finite_measure_finite_Union) (auto simp: sb compact_imp_closed)
       have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (D i)) + e/2" using n0 by simp
       also have "(\<Sum>i<n0. measure M (D i)) \<le> (\<Sum>i<n0. measure M (K i) + e/(2*Suc n0))"
-        using K by (auto intro: setsum_mono simp: emeasure_eq_measure)
+        using K \<open>0 < e\<close>
+        by (auto intro: setsum_mono simp: emeasure_eq_measure measure_nonneg ennreal_plus[symmetric] simp del: ennreal_plus)
       also have "\<dots> = (\<Sum>i<n0. measure M (K i)) + (\<Sum>i<n0. e/(2*Suc n0))"
         by (simp add: setsum.distrib)
       also have "\<dots> \<le> (\<Sum>i<n0. measure M (K i)) +  e / 2" using \<open>0 < e\<close>
@@ -403,14 +308,15 @@
       finally
       have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (K i)) + e / 2 + e / 2"
         by auto
-      hence "M (\<Union>i. D i) < M ?K + e" by (auto simp: mK emeasure_eq_measure)
+      hence "M (\<Union>i. D i) < M ?K + e"
+        using \<open>0<e\<close> by (auto simp: mK emeasure_eq_measure measure_nonneg setsum_nonneg ennreal_less_iff ennreal_plus[symmetric] simp del: ennreal_plus)
       moreover
       have "?K \<subseteq> (\<Union>i. D i)" using K by auto
       moreover
       have "compact ?K" using K by auto
       ultimately
-      have "?K\<subseteq>(\<Union>i. D i) \<and> compact ?K \<and> emeasure M (\<Union>i. D i) \<le> emeasure M ?K + ereal e" by simp
-      thus "\<exists>K\<subseteq>\<Union>i. D i. compact K \<and> emeasure M (\<Union>i. D i) \<le> emeasure M K + ereal e" ..
+      have "?K\<subseteq>(\<Union>i. D i) \<and> compact ?K \<and> emeasure M (\<Union>i. D i) \<le> emeasure M ?K + ennreal e" by simp
+      thus "\<exists>K\<subseteq>\<Union>i. D i. compact K \<and> emeasure M (\<Union>i. D i) \<le> emeasure M K + ennreal e" ..
     qed fact
     case 2
     show ?case
@@ -422,40 +328,52 @@
         from \<open>0 < e\<close> have "0 < e/(2 powr Suc i)" by simp
         have "emeasure M (D i) = (INF U:{U. (D i) \<subseteq> U \<and> open U}. emeasure M U)"
           using union by blast
-        from INF_approx_ereal[OF \<open>0 < e/(2 powr Suc i)\<close> this]
+        from INF_approx_ennreal[OF \<open>0 < e/(2 powr Suc i)\<close> this]
         show "\<exists>U. D i \<subseteq> U \<and> open U \<and> e/(2 powr Suc i) > emeasure M U - emeasure M (D i)"
-          by (auto simp: emeasure_eq_measure)
+          using \<open>0<e\<close>
+          by (auto simp: emeasure_eq_measure measure_nonneg setsum_nonneg ennreal_less_iff ennreal_plus[symmetric] ennreal_minus
+                         finite_measure_mono sb
+                   simp del: ennreal_plus)
       qed
       then obtain U where U: "\<And>i. D i \<subseteq> U i" "\<And>i. open (U i)"
         "\<And>i. e/(2 powr Suc i) > emeasure M (U i) - emeasure M (D i)"
         unfolding choice_iff by blast
       let ?U = "\<Union>i. U i"
-      have "M ?U - M (\<Union>i. D i) = M (?U - (\<Union>i. D i))" using U  \<open>(\<Union>i. D i) \<in> sets M\<close>
+      have "ennreal (measure M ?U - measure M (\<Union>i. D i)) = M ?U - M (\<Union>i. D i)"
+        using U(1,2)
+        by (subst ennreal_minus[symmetric])
+           (auto intro!: finite_measure_mono simp: sb measure_nonneg emeasure_eq_measure)
+      also have "\<dots> = M (?U - (\<Union>i. D i))" using U  \<open>(\<Union>i. D i) \<in> sets M\<close>
         by (subst emeasure_Diff) (auto simp: sb)
       also have "\<dots> \<le> M (\<Union>i. U i - D i)" using U  \<open>range D \<subseteq> sets M\<close>
         by (intro emeasure_mono) (auto simp: sb intro!: sets.countable_nat_UN sets.Diff)
       also have "\<dots> \<le> (\<Sum>i. M (U i - D i))" using U  \<open>range D \<subseteq> sets M\<close>
         by (intro emeasure_subadditive_countably) (auto intro!: sets.Diff simp: sb)
-      also have "\<dots> \<le> (\<Sum>i. ereal e/(2 powr Suc i))" using U \<open>range D \<subseteq> sets M\<close>
-        by (intro suminf_le_pos, subst emeasure_Diff)
-           (auto simp: emeasure_Diff emeasure_eq_measure sb measure_nonneg intro: less_imp_le)
-      also have "\<dots> \<le> (\<Sum>n. ereal (e * (1 / 2) ^ Suc n))"
-        by (simp add: powr_minus inverse_eq_divide powr_realpow field_simps power_divide del: of_nat_Suc)
-      also have "\<dots> = (\<Sum>n. ereal e * ((1 / 2) ^  Suc n))"
-        unfolding times_ereal.simps[symmetric] ereal_power[symmetric] one_ereal_def numeral_eq_ereal
-        by simp
-      also have "\<dots> = ereal e * (\<Sum>n. ((1 / 2) ^ Suc n))"
-        by (rule suminf_cmult_ereal) (auto simp: \<open>0 < e\<close> less_imp_le)
-      also have "\<dots> = e" unfolding suminf_half_series_ereal by simp
-      finally
-      have "emeasure M ?U \<le> emeasure M (\<Union>i. D i) + ereal e" by (simp add: emeasure_eq_measure)
+      also have "\<dots> \<le> (\<Sum>i. ennreal e/(2 powr Suc i))" using U \<open>range D \<subseteq> sets M\<close>
+        using \<open>0<e\<close>
+        by (intro suminf_le, subst emeasure_Diff)
+           (auto simp: emeasure_Diff emeasure_eq_measure sb measure_nonneg ennreal_minus
+                       finite_measure_mono divide_ennreal ennreal_less_iff
+                 intro: less_imp_le)
+      also have "\<dots> \<le> (\<Sum>n. ennreal (e * (1 / 2) ^ Suc n))"
+        using \<open>0<e\<close>
+        by (simp add: powr_minus powr_realpow field_simps divide_ennreal del: of_nat_Suc)
+      also have "\<dots> = ennreal e * (\<Sum>n. ennreal ((1 / 2) ^  Suc n))"
+        unfolding ennreal_power[symmetric]
+        using \<open>0 < e\<close>
+        by (simp add: ac_simps ennreal_mult' divide_ennreal[symmetric] divide_ennreal_def
+                      ennreal_power[symmetric])
+      also have "\<dots> = ennreal e"
+        by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto
+      finally have "emeasure M ?U \<le> emeasure M (\<Union>i. D i) + ennreal e"
+        using \<open>0<e\<close> by (simp add: emeasure_eq_measure ennreal_plus[symmetric] measure_nonneg del: ennreal_plus)
       moreover
       have "(\<Union>i. D i) \<subseteq> ?U" using U by auto
       moreover
       have "open ?U" using U by auto
       ultimately
-      have "(\<Union>i. D i) \<subseteq> ?U \<and> open ?U \<and> emeasure M ?U \<le> emeasure M (\<Union>i. D i) + ereal e" by simp
-      thus "\<exists>B. (\<Union>i. D i) \<subseteq> B \<and> open B \<and> emeasure M B \<le> emeasure M (\<Union>i. D i) + ereal e" ..
+      have "(\<Union>i. D i) \<subseteq> ?U \<and> open ?U \<and> emeasure M ?U \<le> emeasure M (\<Union>i. D i) + ennreal e" by simp
+      thus "\<exists>B. (\<Union>i. D i) \<subseteq> B \<and> open B \<and> emeasure M B \<le> emeasure M (\<Union>i. D i) + ennreal e" ..
     qed
   qed
 qed
--- a/src/HOL/Probability/Sigma_Algebra.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Sigma_Algebra.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -13,7 +13,7 @@
   "~~/src/HOL/Library/Countable_Set"
   "~~/src/HOL/Library/FuncSet"
   "~~/src/HOL/Library/Indicator_Function"
-  "~~/src/HOL/Library/Extended_Real"
+  "~~/src/HOL/Library/Extended_Nonnegative_Real"
   "~~/src/HOL/Library/Disjoint_Sets"
 begin
 
@@ -379,7 +379,7 @@
   shows "{x\<in>\<Omega>. \<exists>i\<in>I. P i x} \<in> M"
 proof -
   have "{x\<in>\<Omega>. \<exists>i\<in>I. P i x} = (\<Union>i\<in>I. {x\<in>\<Omega>. P i x})" by auto
-  with assms show ?thesis 
+  with assms show ?thesis
     by (auto intro!: countable_UN')
 qed
 
@@ -389,7 +389,7 @@
   shows "{x\<in>\<Omega>. \<forall>i\<in>I. P i x} \<in> M"
 proof -
   have "{x\<in>\<Omega>. \<forall>i\<in>I. P i x} = (\<Inter>i\<in>I. {x\<in>\<Omega>. P i x}) \<inter> \<Omega>" by auto
-  with assms show ?thesis 
+  with assms show ?thesis
     by (cases "I = {}") (auto intro!: countable_INT')
 qed
 
@@ -400,7 +400,7 @@
 proof -
   have "{x\<in>\<Omega>. \<exists>!i\<in>I. P i x} = {x\<in>\<Omega>. \<exists>i\<in>I. P i x \<and> (\<forall>j\<in>I. P j x \<longrightarrow> i = j)}"
     by auto
-  with assms show ?thesis 
+  with assms show ?thesis
     by (auto intro!: sets_Collect_countable_All' sets_Collect_countable_Ex' sets_Collect_conj sets_Collect_imp sets_Collect_const)
 qed
 
@@ -888,7 +888,7 @@
 
   { fix a assume a: "a \<in> ?R" then guess Ca .. note Ca = this
     fix b assume b: "b \<in> ?R" then guess Cb .. note Cb = this
-  
+
     show "a - b \<in> ?R"
     proof cases
       assume "Cb = {}" with Cb \<open>a \<in> ?R\<close> show ?thesis
@@ -929,7 +929,7 @@
 
 subsubsection \<open>A Two-Element Series\<close>
 
-definition binaryset :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> 'a set "
+definition binaryset :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> 'a set"
   where "binaryset A B = (\<lambda>x. {})(0 := A, Suc 0 := B)"
 
 lemma range_binaryset_eq: "range(binaryset A B) = {A,B,{}}"
@@ -1448,14 +1448,14 @@
 
 subsection \<open>Measure type\<close>
 
-definition positive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
-  "positive M \<mu> \<longleftrightarrow> \<mu> {} = 0 \<and> (\<forall>A\<in>M. 0 \<le> \<mu> A)"
+definition positive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
+  "positive M \<mu> \<longleftrightarrow> \<mu> {} = 0"
 
-definition countably_additive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
+definition countably_additive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
   "countably_additive M f \<longleftrightarrow> (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow>
     (\<Sum>i. f (A i)) = f (\<Union>i. A i))"
 
-definition measure_space :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
+definition measure_space :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
   "measure_space \<Omega> A \<mu> \<longleftrightarrow> sigma_algebra \<Omega> A \<and> positive A \<mu> \<and> countably_additive A \<mu>"
 
 typedef 'a measure = "{(\<Omega>::'a set, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu> }"
@@ -1472,11 +1472,11 @@
 definition sets :: "'a measure \<Rightarrow> 'a set set" where
   "sets M = fst (snd (Rep_measure M))"
 
-definition emeasure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ereal" where
+definition emeasure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ennreal" where
   "emeasure M = snd (snd (Rep_measure M))"
 
 definition measure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> real" where
-  "measure M A = real_of_ereal (emeasure M A)"
+  "measure M A = enn2real (emeasure M A)"
 
 declare [[coercion sets]]
 
@@ -1490,7 +1490,7 @@
 interpretation sets: sigma_algebra "space M" "sets M" for M :: "'a measure"
   using measure_space[of M] by (auto simp: measure_space_def)
 
-definition measure_of :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
+definition measure_of :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
   "measure_of \<Omega> A \<mu> = Abs_measure (\<Omega>, if A \<subseteq> Pow \<Omega> then sigma_sets \<Omega> A else {{}, \<Omega>},
     \<lambda>a. if a \<in> sigma_sets \<Omega> A \<and> measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> then \<mu> a else 0)"
 
@@ -1546,8 +1546,8 @@
   shows space_measure_of_conv: "space (measure_of \<Omega> A \<mu>) = \<Omega>" (is ?space)
   and sets_measure_of_conv:
   "sets (measure_of \<Omega> A \<mu>) = (if A \<subseteq> Pow \<Omega> then sigma_sets \<Omega> A else {{}, \<Omega>})" (is ?sets)
-  and emeasure_measure_of_conv: 
-  "emeasure (measure_of \<Omega> A \<mu>) = 
+  and emeasure_measure_of_conv:
+  "emeasure (measure_of \<Omega> A \<mu>) =
   (\<lambda>B. if B \<in> sigma_sets \<Omega> A \<and> measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> then \<mu> B else 0)" (is ?emeasure)
 proof -
   have "?space \<and> ?sets \<and> ?emeasure"
@@ -1663,7 +1663,7 @@
   using emeasure_notin_sets[of A M] by blast
 
 lemma measure_notin_sets: "A \<notin> sets M \<Longrightarrow> measure M A = 0"
-  by (simp add: measure_def emeasure_notin_sets)
+  by (simp add: measure_def emeasure_notin_sets zero_ennreal.rep_eq)
 
 lemma measure_eqI:
   fixes M N :: "'a measure"
@@ -1737,7 +1737,7 @@
 proof -
   interpret A: sigma_algebra \<Omega> "sigma_sets \<Omega> A" using B(2) by (rule sigma_algebra_sigma_sets)
   from B sets.top[of N] A.top sets.space_closed[of N] A.space_closed have \<Omega>: "\<Omega> = space N" by force
-  
+
   { fix X assume "X \<in> sigma_sets \<Omega> A"
     then have "f -` X \<inter> space M \<in> sets M \<and> X \<subseteq> \<Omega>"
       proof induct
@@ -1867,9 +1867,9 @@
 subsubsection \<open>Counting space\<close>
 
 definition count_space :: "'a set \<Rightarrow> 'a measure" where
-  "count_space \<Omega> = measure_of \<Omega> (Pow \<Omega>) (\<lambda>A. if finite A then ereal (card A) else \<infinity>)"
+  "count_space \<Omega> = measure_of \<Omega> (Pow \<Omega>) (\<lambda>A. if finite A then of_nat (card A) else \<infinity>)"
 
-lemma 
+lemma
   shows space_count_space[simp]: "space (count_space \<Omega>) = \<Omega>"
     and sets_count_space[simp]: "sets (count_space \<Omega>) = Pow \<Omega>"
   using sigma_sets_into_sp[of "Pow \<Omega>" \<Omega>]
@@ -1938,7 +1938,7 @@
   shows "(\<lambda>x. f (g x)) \<in> measurable M N"
   using measurable_compose[OF g f] .
 
-lemma measurable_empty_iff: 
+lemma measurable_empty_iff:
   "space N = {} \<Longrightarrow> f \<in> measurable M N \<longleftrightarrow> space M = {}"
   by (auto simp add: measurable_def Pi_iff)
 
@@ -2021,11 +2021,11 @@
 lemma in_Sup_sigma: "m \<in> M \<Longrightarrow> A \<in> sets m \<Longrightarrow> A \<in> sets (Sup_sigma M)"
   unfolding sets_Sup_sigma by auto
 
-lemma SUP_sigma_cong: 
+lemma SUP_sigma_cong:
   assumes *: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i)" shows "sets (\<Squnion>\<^sub>\<sigma> i\<in>I. M i) = sets (\<Squnion>\<^sub>\<sigma> i\<in>I. N i)"
   using * sets_eq_imp_space_eq[OF *] by (simp add: Sup_sigma_def)
 
-lemma sets_Sup_in_sets: 
+lemma sets_Sup_in_sets:
   assumes "M \<noteq> {}"
   assumes "\<And>m. m \<in> M \<Longrightarrow> space m = space N"
   assumes "\<And>m. m \<in> M \<Longrightarrow> sets m \<subseteq> sets N"
@@ -2248,7 +2248,7 @@
   apply (auto simp add: sets_vimage_algebra intro!: arg_cong2[where f=sigma_sets])
   done
 
-lemma sets_Collect_restrict_space_iff: 
+lemma sets_Collect_restrict_space_iff:
   assumes "S \<in> sets M"
   shows "{x\<in>space (restrict_space M S). P x} \<in> sets (restrict_space M S) \<longleftrightarrow> {x\<in>space M. x \<in> S \<and> P x} \<in> sets M"
 proof -
@@ -2365,4 +2365,3 @@
      (auto simp: eq[symmetric] space_restrict_space cong: measurable_cong' intro: f measurable_restrict_space1)
 
 end
-
--- a/src/HOL/Probability/Sinc_Integral.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Sinc_Integral.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -1,6 +1,6 @@
 (*
   Theory: Sinc_Integral.thy
-  Authors: Jeremy Avigad, Luke Serafin, Johannes Hölzl 
+  Authors: Jeremy Avigad, Luke Serafin, Johannes Hölzl
 *)
 
 section \<open>Integral of sinc\<close>
@@ -14,7 +14,7 @@
 text \<open> Naming convention
 The theorem name consists of the following parts:
   \<^item> Kind of integral: @{text has_bochner_integral} / @{text integrable} / @{text LBINT}
-  \<^item> Interval: Interval (0 / infinity / open / closed) (infinity / open / closed) 
+  \<^item> Interval: Interval (0 / infinity / open / closed) (infinity / open / closed)
   \<^item> Name of the occurring constants: power, exp, m (for minus), scale, sin, $\ldots$
 \<close>
 
@@ -22,7 +22,7 @@
   "has_bochner_integral lborel (\<lambda>x. x^k * exp (-x) * indicator {0 ..} x::real) (fact k)"
   using nn_intergal_power_times_exp_Ici[of k]
   by (intro has_bochner_integral_nn_integral)
-     (auto simp: ereal_indicator[symmetric] split: split_indicator)
+     (auto simp: nn_integral_set_ennreal split: split_indicator)
 
 lemma has_bochner_integral_I0i_power_exp_m:
   "has_bochner_integral lborel (\<lambda>x. x^k * exp (-x) * indicator {0 <..} x::real) (fact k)"
@@ -30,12 +30,12 @@
   by (intro has_bochner_integral_cong_AE[THEN iffD1, OF _ _ _ has_bochner_integral_I0i_power_exp_m'])
      (auto split: split_indicator)
 
-lemma integrable_I0i_exp_mscale: "0 < (u::real) \<Longrightarrow> set_integrable lborel {0 <..} (\<lambda>x. exp (-(x * u)))" 
+lemma integrable_I0i_exp_mscale: "0 < (u::real) \<Longrightarrow> set_integrable lborel {0 <..} (\<lambda>x. exp (-(x * u)))"
   using lborel_integrable_real_affine_iff[of u "\<lambda>x. indicator {0 <..} x *\<^sub>R exp (- x)" 0]
         has_bochner_integral_I0i_power_exp_m[of 0]
   by (simp add: indicator_def zero_less_mult_iff mult_ac integrable.intros)
 
-lemma LBINT_I0i_exp_mscale: "0 < (u::real) \<Longrightarrow> LBINT x=0..\<infinity>. exp (-(x * u)) = 1 / u" 
+lemma LBINT_I0i_exp_mscale: "0 < (u::real) \<Longrightarrow> LBINT x=0..\<infinity>. exp (-(x * u)) = 1 / u"
   using lborel_integral_real_affine[of u "\<lambda>x. indicator {0<..} x *\<^sub>R exp (- x)" 0]
         has_bochner_integral_I0i_power_exp_m[of 0]
   by (auto simp: indicator_def zero_less_mult_iff interval_lebesgue_integral_0_infty field_simps
@@ -54,7 +54,7 @@
 
 lemma LBINT_I0i_exp_mscale_sin:
   assumes "0 < x"
-  shows "LBINT u=0..\<infinity>. \<bar>exp (-u * x) * sin x\<bar> = \<bar>sin x\<bar> / x" 
+  shows "LBINT u=0..\<infinity>. \<bar>exp (-u * x) * sin x\<bar> = \<bar>sin x\<bar> / x"
 proof (subst interval_integral_FTC_nonneg)
   let ?F = "\<lambda>u. 1 / x * (1 - exp (- u * x)) * \<bar>sin x\<bar>"
   show "\<And>t. (?F has_real_derivative \<bar>exp (- t * x) * sin x\<bar>) (at t)"
@@ -90,7 +90,7 @@
              simp add: ereal_tendsto_simps filterlim_tan_at_left add_nonneg_eq_0_iff)
 qed
 
-lemma 
+lemma
   shows integrable_I0i_1_div_plus_square:
     "interval_lebesgue_integrable lborel 0 \<infinity> (\<lambda>x. 1 / (1 + x^2))"
   and LBINT_I0i_1_div_plus_square:
@@ -105,7 +105,7 @@
        (auto intro: derivative_eq_intros 1 2 tendsto_eq_intros
              simp add: ereal_tendsto_simps filterlim_tan_at_left zero_ereal_def add_nonneg_eq_0_iff)
   show "interval_lebesgue_integrable lborel 0 \<infinity> (\<lambda>x. 1 / (1 + x^2))"
-    unfolding interval_lebesgue_integrable_def 
+    unfolding interval_lebesgue_integrable_def
     by (subst interval_integral_substitution_nonneg[of "0" "pi/2" tan "\<lambda>x. 1 + (tan x)^2"])
        (auto intro: derivative_eq_intros 1 2 tendsto_eq_intros
              simp add: ereal_tendsto_simps filterlim_tan_at_left zero_ereal_def add_nonneg_eq_0_iff)
@@ -159,7 +159,7 @@
        (auto simp: Si_def intro!: interval_lebesgue_integral_cong_AE)
 qed
 
-lemma Si_neg: 
+lemma Si_neg:
   assumes "T \<ge> 0" shows "Si (- T) = - Si T"
 proof -
   have "LBINT x=ereal 0..T. -1 *\<^sub>R sinc (- x) = LBINT y= ereal (- 0)..ereal (- T). sinc y"
@@ -222,7 +222,7 @@
     then have **: "1 \<le> t \<Longrightarrow> 0 < x \<Longrightarrow> \<bar>?F x t\<bar> \<le> exp (- x) * (x + 1)" for x t :: real
         by (auto simp add: abs_mult times_divide_eq_right[symmetric] simp del: times_divide_eq_right
                  intro!:  mult_mono)
-  
+
     show "\<forall>\<^sub>F i in at_top. AE x in lborel. 0 < x \<longrightarrow> \<bar>?F x i\<bar> \<le> exp (- x) * (x + 1)"
       using eventually_ge_at_top[of "1::real"] ** by (auto elim: eventually_mono)
     show "AE x in lborel. 0 < x \<longrightarrow> (?F x \<longlongrightarrow> 0) at_top"
@@ -289,11 +289,11 @@
         using AE_lborel_singleton[of 0] AE_lborel_singleton[of t]
       proof eventually_elim
         fix x :: real assume x: "x \<noteq> 0" "x \<noteq> t"
-        have "LBINT y. \<bar>indicator ({0<..} \<times> {0<..<t}) (y, x) *\<^sub>R (sin x * exp (- (y * x)))\<bar> = 
+        have "LBINT y. \<bar>indicator ({0<..} \<times> {0<..<t}) (y, x) *\<^sub>R (sin x * exp (- (y * x)))\<bar> =
             LBINT y. \<bar>sin x\<bar> * exp (- (y * x)) * indicator {0<..} y * indicator {0<..<t} x"
           by (intro integral_cong) (auto split: split_indicator simp: abs_mult)
         also have "\<dots> = \<bar>sin x\<bar> * indicator {0<..<t} x * (LBINT y=0..\<infinity>.  exp (- (y * x)))"
-          by (cases "x > 0")                                       
+          by (cases "x > 0")
              (auto simp: field_simps interval_lebesgue_integral_0_infty split: split_indicator)
         also have "\<dots> = \<bar>sin x\<bar> * indicator {0<..<t} x * (1 / x)"
           by (cases "x > 0") (auto simp add: LBINT_I0i_exp_mscale)
@@ -306,7 +306,7 @@
         by (auto intro!: borel_integrable_compact continuous_intros simp del: real_scaleR_def)
       ultimately show "integrable lborel (\<lambda>x. LBINT y. norm (?f (x, y)))"
         by (rule integrable_cong_AE_imp[rotated 2]) simp
-  
+
       have "0 < x \<Longrightarrow> set_integrable lborel {0<..} (\<lambda>y. sin x * exp (- (y * x)))" for x :: real
           by (intro set_integrable_mult_right integrable_I0i_exp_mscale)
       then show "AE x in lborel. integrable lborel (\<lambda>y. ?f (x, y))"
@@ -350,7 +350,7 @@
     by auto
   ultimately show ?thesis
     by (intro exI[of _ "max M (pi/2 + 1)"]) (meson le_max_iff_disj linorder_not_le order_le_less)
-qed    
+qed
 
 lemma LBINT_I0c_sin_scale_divide:
   assumes "T \<ge> 0"
@@ -369,12 +369,12 @@
       by simp
     hence "LBINT x. indicator {0<..<T} x * sin (x * \<theta>) / x =
         LBINT x. indicator {0<..<T * \<theta>} x * sin x / x"
-      using assms `0 < \<theta>` unfolding interval_lebesgue_integral_def einterval_eq zero_ereal_def 
+      using assms `0 < \<theta>` unfolding interval_lebesgue_integral_def einterval_eq zero_ereal_def
         by (auto simp: ac_simps)
   } note aux1 = this
   { assume "0 > \<theta>"
     have [simp]: "integrable lborel (\<lambda>x. sin (x * \<theta>) * indicator {0<..<T} x / x)"
-      using integrable_sinc' [of T \<theta>] assms 
+      using integrable_sinc' [of T \<theta>] assms
       by (simp add: interval_lebesgue_integrable_def ac_simps)
     have "(LBINT t=ereal 0..T. sin (t * -\<theta>) / t) = (LBINT t=ereal 0..T. -\<theta> *\<^sub>R sinc (t * -\<theta>))"
       by (rule interval_integral_discrete_difference[of "{0}"]) auto
--- a/src/HOL/Probability/Weak_Convergence.thy	Thu Apr 14 12:17:44 2016 +0200
+++ b/src/HOL/Probability/Weak_Convergence.thy	Thu Apr 14 15:48:11 2016 +0200
@@ -104,7 +104,7 @@
              intro!: arg_cong2[where f="measure"])
   also have "\<dots> = measure lborel {0 <..< C x}"
     using cdf_bounded_prob[of x] AE_lborel_singleton[of "C x"]
-    by (auto intro!: arg_cong[where f=real_of_ereal] emeasure_eq_AE simp: measure_def)
+    by (auto intro!: arg_cong[where f=enn2real] emeasure_eq_AE simp: measure_def)
   also have "\<dots> = C x"
     by (simp add: cdf_nonneg)
   finally show "cdf (distr ?\<Omega> borel I) x = C x" .
@@ -334,11 +334,13 @@
   have "cdf M x = integral\<^sup>L M (indicator {..x})"
     by (simp add: cdf_def)
   also have "\<dots> \<le> expectation (cts_step x y)"
-    by (intro integral_mono integrable_cts_step) (auto simp: cts_step_def split: split_indicator)
+    by (intro integral_mono integrable_cts_step)
+       (auto simp: cts_step_def less_top[symmetric] split: split_indicator)
   finally show "cdf M x \<le> expectation (cts_step x y)" .
 next
   have "expectation (cts_step x y) \<le> integral\<^sup>L M (indicator {..y})"
-    by (intro integral_mono integrable_cts_step) (auto simp: cts_step_def split: split_indicator)
+    by (intro integral_mono integrable_cts_step)
+       (auto simp: cts_step_def less_top[symmetric] split: split_indicator)
   also have "\<dots> = cdf M y"
     by (simp add: cdf_def)
   finally show "expectation (cts_step x y) \<le> cdf M y" .