Theory Nonnegative_Lebesgue_Integration

(*  Title:      HOL/Analysis/Nonnegative_Lebesgue_Integration.thy
    Author:     Johannes Hölzl, TU München
    Author:     Armin Heller, TU München
*)

section ‹Lebesgue Integration for Nonnegative Functions›

theory Nonnegative_Lebesgue_Integration
  imports Measure_Space Borel_Space
begin

subsectiontag unimportant› ‹Approximating functions›

lemma AE_upper_bound_inf_ennreal:
  fixes F G::"'a  ennreal"
  assumes "e. (e::real) > 0  AE x in M. F x  G x + e"
  shows "AE x in M. F x  G x"
proof -
  have "AE x in M. n::nat. F x  G x + ennreal (1 / Suc n)"
    using assms by (auto simp: AE_all_countable)
  then show ?thesis
  proof (eventually_elim)
    fix x assume x: "n::nat. F x  G x + ennreal (1 / Suc n)"
    show "F x  G x"
    proof (rule ennreal_le_epsilon)
      fix e :: real assume "0 < e"
      then obtain n where n: "1 / Suc n < e"
        by (blast elim: nat_approx_posE)
      have "F x  G x + 1 / Suc n"
        using x by simp
      also have "  G x + e"
        using n by (intro add_mono ennreal_leI) auto
      finally show "F x  G x + ennreal e" .
    qed
  qed
qed

lemma AE_upper_bound_inf:
  fixes F G::"'a  real"
  assumes "e. e > 0  AE x in M. F x  G x + e"
  shows "AE x in M. F x  G x"
proof -
  have "AE x in M. F x  G x + 1/real (n+1)" for n::nat
    by (rule assms, auto)
  then have "AE x in M. n::nat  UNIV. F x  G x + 1/real (n+1)"
    by (rule AE_ball_countable', auto)
  moreover
  {
    fix x assume i: "n::nat  UNIV. F x  G x + 1/real (n+1)"
    have "(λn. G x + 1/real (n+1))  G x + 0"
      by (rule tendsto_add, simp, rule LIMSEQ_ignore_initial_segment[OF lim_1_over_n, of 1])
    then have "F x  G x" using i LIMSEQ_le_const by fastforce
  }
  ultimately show ?thesis by auto
qed

lemma not_AE_zero_ennreal_E:
  fixes f::"'a  ennreal"
  assumes "¬ (AE x in M. f x = 0)" and [measurable]: "f  borel_measurable M"
  shows "Asets M. e::real>0. emeasure M A > 0  (x  A. f x  e)"
proof -
  { assume "¬ (e::real>0. {x  space M. f x  e}  null_sets M)"
    then have "0 < e  AE x in M. f x  e" for e :: real
      by (auto simp: not_le less_imp_le dest!: AE_not_in)
    then have "AE x in M. f x  0"
      by (intro AE_upper_bound_inf_ennreal[where G="λ_. 0"]) simp
    then have False
      using assms by auto }
  then obtain e::real where e: "e > 0" "{x  space M. f x  e}  null_sets M" by auto
  define A where "A = {x  space M. f x  e}"
  have 1 [measurable]: "A  sets M" unfolding A_def by auto
  have 2: "emeasure M A > 0"
    using e(2) A_def A  sets M by auto
  have 3: "x. x  A  f x  e" unfolding A_def by auto
  show ?thesis using e(1) 1 2 3 by blast
qed

lemma not_AE_zero_E:
  fixes f::"'a  real"
  assumes "AE x in M. f x  0"
          "¬(AE x in M. f x = 0)"
      and [measurable]: "f  borel_measurable M"
  shows "A e. A  sets M  e>0  emeasure M A > 0  (x  A. f x  e)"
proof -
  have "e. e > 0  {x  space M. f x  e}  null_sets M"
  proof (rule ccontr)
    assume *: "¬(e. e > 0  {x  space M. f x  e}  null_sets M)"
    {
      fix e::real assume "e > 0"
      then have "{x  space M. f x  e}  null_sets M" using * by blast
      then have "AE x in M. x  {x  space M. f x  e}" using AE_not_in by blast
      then have "AE x in M. f x  e" by auto
    }
    then have "AE x in M. f x  0" by (rule AE_upper_bound_inf, auto)
    then have "AE x in M. f x = 0" using assms(1) by auto
    then show False using assms(2) by auto
  qed
  then obtain e where e: "e>0" "{x  space M. f x  e}  null_sets M" by auto
  define A where "A = {x  space M. f x  e}"
  have 1 [measurable]: "A  sets M" unfolding A_def by auto
  have 2: "emeasure M A > 0"
    using e(2) A_def A  sets M by auto
  have 3: "x. x  A  f x  e" unfolding A_def by auto
  show ?thesis
    using e(1) 1 2 3 by blast
qed

subsection "Simple function"

text ‹

Our simple functions are not restricted to nonnegative real numbers. Instead
they are just functions with a finite range and are measurable when singleton
sets are measurable.

›

definitiontag important› "simple_function M g 
    finite (g ` space M) 
    (x  g ` space M. g -` {x}  space M  sets M)"

lemma simple_functionD:
  assumes "simple_function M g"
  shows "finite (g ` space M)" and "g -` X  space M  sets M"
proof -
  show "finite (g ` space M)"
    using assms unfolding simple_function_def by auto
  have "g -` X  space M = g -` (X  g`space M)  space M" by auto
  also have " = (xX  g`space M. g-`{x}  space M)" by auto
  finally show "g -` X  space M  sets M" using assms
    by (auto simp del: UN_simps simp: simple_function_def)
qed

lemma measurable_simple_function[measurable_dest]:
  "simple_function M f  f  measurable M (count_space UNIV)"
  unfolding simple_function_def measurable_def
proof safe
  fix A assume "finite (f ` space M)" "xf ` space M. f -` {x}  space M  sets M"
  then have "(xf ` space M. if x  A then f -` {x}  space M else {})  sets M"
    by (intro sets.finite_UN) auto
  also have "(xf ` space M. if x  A then f -` {x}  space M else {}) = f -` A  space M"
    by (auto split: if_split_asm)
  finally show "f -` A  space M  sets M" .
qed simp

lemma borel_measurable_simple_function:
  "simple_function M f  f  borel_measurable M"
  by (auto dest!: measurable_simple_function simp: measurable_def)

lemma simple_function_measurable2[intro]:
  assumes "simple_function M f" "simple_function M g"
  shows "f -` A  g -` B  space M  sets M"
proof -
  have "f -` A  g -` B  space M = (f -` A  space M)  (g -` B  space M)"
    by auto
  then show ?thesis using assms[THEN simple_functionD(2)] by auto
qed

lemma simple_function_indicator_representation:
  fixes f ::"'a  ennreal"
  assumes f: "simple_function M f" and x: "x  space M"
  shows "f x = (y  f ` space M. y * indicator (f -` {y}  space M) x)"
  (is "?l = ?r")
proof -
  have "?r = (y  f ` space M.
    (if y = f x then y * indicator (f -` {y}  space M) x else 0))"
    by (auto intro!: sum.cong)
  also have "... =  f x *  indicator (f -` {f x}  space M) x"
    using assms by (auto dest: simple_functionD)
  also have "... = f x" using x by (auto simp: indicator_def)
  finally show ?thesis by auto
qed

lemma simple_function_notspace:
  "simple_function M (λx. h x * indicator (- space M) x::ennreal)" (is "simple_function M ?h")
proof -
  have "?h ` space M  {0}" unfolding indicator_def by auto
  hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset)
  have "?h -` {0}  space M = space M" by auto
  thus ?thesis unfolding simple_function_def by (auto simp add: image_constant_conv)
qed

lemma simple_function_cong:
  assumes "t. t  space M  f t = g t"
  shows "simple_function M f  simple_function M g"
proof -
  have "x. f -` {x}  space M = g -` {x}  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:
  assumes "sets N = sets M" "space N = space M"
  shows "simple_function M f  simple_function N f"
  unfolding simple_function_def assms ..

lemma simple_function_borel_measurable:
  fixes f :: "'a  'x::{t2_space}"
  assumes "f  borel_measurable M" and "finite (f ` space M)"
  shows "simple_function M f"
  using assms unfolding simple_function_def
  by (auto intro: borel_measurable_vimage)

lemma simple_function_iff_borel_measurable:
  fixes f :: "'a  'x::{t2_space}"
  shows "simple_function M f  finite (f ` space M)  f  borel_measurable M"
  by (metis borel_measurable_simple_function simple_functionD(1) simple_function_borel_measurable)

lemma simple_function_eq_measurable:
  "simple_function M f  finite (f`space M)  f  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 (λx. c)"
  by (auto intro: finite_subset simp: simple_function_def)
lemma simple_function_compose[intro, simp]:
  assumes "simple_function M f"
  shows "simple_function M (g  f)"
  unfolding simple_function_def
proof safe
  show "finite ((g  f) ` space M)"
    using assms unfolding simple_function_def image_comp [symmetric] by auto
next
  fix x assume "x  space M"
  let ?G = "g -` {g (f x)}  (f`space M)"
  have *: "(g  f) -` {(g  f) x}  space M =
    (x?G. f -` {x}  space M)" by auto
  show "(g  f) -` {(g  f) x}  space M  sets M"
    using assms unfolding simple_function_def *
    by (rule_tac sets.finite_UN) auto
qed

lemma simple_function_indicator[intro, simp]:
  assumes "A  sets M"
  shows "simple_function M (indicator A)"
proof -
  have "indicator A ` space M  {0, 1}" (is "?S  _")
    by (auto simp: indicator_def)
  hence "finite ?S" by (rule finite_subset) simp
  moreover have "- A  space M = space M - A" by auto
  ultimately show ?thesis unfolding simple_function_def
    using assms by (auto simp: indicator_def [abs_def])
qed

lemma simple_function_Pair[intro, simp]:
  assumes "simple_function M f"
  assumes "simple_function M g"
  shows "simple_function M (λx. (f x, g x))" (is "simple_function M ?p")
  unfolding simple_function_def
proof safe
  show "finite (?p ` space M)"
    using assms unfolding simple_function_def
    by (rule_tac finite_subset[of _ "f`space M × g`space M"]) auto
next
  fix x assume "x  space M"
  have "(λx. (f x, g x)) -` {(f x, g x)}  space M =
      (f -` {f x}  space M)  (g -` {g x}  space M)"
    by auto
  with x  space M show "(λx. (f x, g x)) -` {(f x, g x)}  space M  sets M"
    using assms unfolding simple_function_def by auto
qed

lemma simple_function_compose1:
  assumes "simple_function M f"
  shows "simple_function M (λx. g (f x))"
  using simple_function_compose[OF assms, of g]
  by (simp add: comp_def)

lemma simple_function_compose2:
  assumes "simple_function M f" and "simple_function M g"
  shows "simple_function M (λx. h (f x) (g x))"
proof -
  have "simple_function M ((λ(x, y). h x y)  (λx. (f x, g x)))"
    using assms by auto
  thus ?thesis by (simp_all add: comp_def)
qed

lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="(+)"]
  and simple_function_diff[intro, simp] = simple_function_compose2[where h="(-)"]
  and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"]
  and simple_function_mult[intro, simp] = simple_function_compose2[where h="(*)"]
  and simple_function_div[intro, simp] = simple_function_compose2[where h="(/)"]
  and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"]
  and simple_function_max[intro, simp] = simple_function_compose2[where h=max]

lemma simple_function_sum[intro, simp]:
  assumes "i. i  P  simple_function M (f i)"
  shows "simple_function M (λx. iP. f i x)"
proof cases
  assume "finite P" from this assms show ?thesis by induct auto
qed auto

lemma simple_function_ennreal[intro, simp]:
  fixes f g :: "'a  real" assumes sf: "simple_function M f"
  shows "simple_function M (λx. ennreal (f x))"
  by (rule simple_function_compose1[OF sf])

lemma simple_function_real_of_nat[intro, simp]:
  fixes f g :: "'a  nat" assumes sf: "simple_function M f"
  shows "simple_function M (λx. real (f x))"
  by (rule simple_function_compose1[OF sf])

lemmatag important› borel_measurable_implies_simple_function_sequence:
  fixes u :: "'a  ennreal"
  assumes u[measurable]: "u  borel_measurable M"
  shows "f. incseq f  (i. (x. f i x < top)  simple_function M (f i))  u = (SUP i. f i)"
proof -
  define f where [abs_def]:
    "f i x = real_of_int (floor (enn2real (min i (u x)) * 2^i)) / 2^i" for i x

  have [simp]: "0  f i x" for i x
    by (auto simp: f_def intro!: divide_nonneg_nonneg mult_nonneg_nonneg enn2real_nonneg)

  have *: "2^n * real_of_int x = real_of_int (2^n * x)" for n x
    by simp

  have "real_of_int real i * 2 ^ i = real_of_int i * 2 ^ i" for i
    by (intro arg_cong[where f=real_of_int]) simp
  then have [simp]: "real_of_int real i * 2 ^ i = i * 2 ^ i" for i
    unfolding floor_of_nat by simp

  have "incseq f"
  proof (intro monoI le_funI)
    fix m n :: nat and x assume "m  n"
    moreover
    { fix d :: nat
      have "2^d::real * 2^m * enn2real (min (of_nat m) (u x)) 
        2^d * (2^m * enn2real (min (of_nat m) (u x)))"
        by (rule le_mult_floor) (auto)
      also have "  2^d * (2^m * enn2real (min (of_nat d + of_nat m) (u x)))"
        by (intro floor_mono mult_mono enn2real_mono min.mono)
           (auto simp: min_less_iff_disj of_nat_less_top)
      finally have "f m x  f (m + d) x"
        unfolding f_def
        by (auto simp: field_simps power_add * simp del: of_int_mult) }
    ultimately show "f m x  f n x"
      by (auto simp add: le_iff_add)
  qed
  then have inc_f: "incseq (λi. ennreal (f i x))" for x
    by (auto simp: incseq_def le_fun_def)
  then have "incseq (λ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 "enn2real (min (of_nat i) (u x)) * 2 ^ i  int i * 2 ^ i" for x
      by (cases "u x" rule: ennreal_cases)
         (auto split: split_min intro!: floor_mono)
    then have "f i ` space M  (λn. real_of_int n / 2^i) ` {0 .. of_nat i * 2^i}"
      unfolding floor_of_int by (auto simp: f_def intro!: imageI)
    then show "finite (f i ` space M)"
      by (rule finite_subset) auto
    show "f i  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
      case (real r)
      obtain n where "r  of_nat n" using real_arch_simple by auto
      then have min_eq_r: "F x in sequentially. min (real x) r = r"
        by (auto simp: eventually_sequentially intro!: exI[of _ n] split: split_min)

      have "(λi. real_of_int min (real i) r * 2^i / 2^i)  r"
      proof (rule tendsto_sandwich)
        show "(λn. r - (1/2)^n)  r"
          by (auto intro!: tendsto_eq_intros LIMSEQ_power_zero)
        show "F n in sequentially. real_of_int min (real n) r * 2 ^ n / 2 ^ n  r"
          using min_eq_r by eventually_elim (auto simp: field_simps)
        have *: "r * (2 ^ n * 2 ^ n)  2^n + 2^n * real_of_int r * 2 ^ n" 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 "F n in sequentially. r - (1/2)^n  real_of_int min (real n) r * 2 ^ n / 2 ^ n"
          using min_eq_r by eventually_elim (insert *, auto simp: field_simps)
      qed auto
      then have "(λi. ennreal (f i x))  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 _ "λi x. ennreal (f i x)"]) (auto simp add: image_comp)
qed

lemma borel_measurable_implies_simple_function_sequence':
  fixes u :: "'a  ennreal"
  assumes u: "u  borel_measurable M"
  obtains f where
    "i. simple_function M (f i)" "incseq f" "i x. f i x < top" "x. (SUP i. f i x) = u x"
  using borel_measurable_implies_simple_function_sequence [OF u]
  by (metis SUP_apply)

lemmatag important› simple_function_induct
    [consumes 1, case_names cong set mult add, induct set: simple_function]:
  fixes u :: "'a  ennreal"
  assumes u: "simple_function M u"
  assumes cong: "f g. simple_function M f  simple_function M g  (AE x in M. f x = g x)  P f  P g"
  assumes set: "A. A  sets M  P (indicator A)"
  assumes mult: "u c. P u  P (λx. c * u x)"
  assumes add: "u v. P u  P v  P (λx. v x + u x)"
  shows "P u"
proof (rule cong)
  from AE_space show "AE x in M. (yu ` space M. y * indicator (u -` {y}  space M) x) = u x"
  proof eventually_elim
    fix x assume x: "x  space M"
    from simple_function_indicator_representation[OF u x]
    show "(yu ` space M. y * indicator (u -` {y}  space M) x) = u x" ..
  qed
next
  from u have "finite (u ` space M)"
    unfolding simple_function_def by auto
  then show "P (λx. yu ` space M. y * indicator (u -` {y}  space M) x)"
  proof induct
    case empty show ?case
      using set[of "{}"] by (simp add: indicator_def[abs_def])
  qed (auto intro!: add mult set simple_functionD u)
next
  show "simple_function M (λx. (yu ` space M. y * indicator (u -` {y}  space M) x))"
    apply (subst simple_function_cong)
    apply (rule simple_function_indicator_representation[symmetric])
    apply (auto intro: u)
    done
qed fact

lemma simple_function_induct_nn[consumes 1, case_names cong set mult add]:
  fixes u :: "'a  ennreal"
  assumes u: "simple_function M u"
  assumes cong: "f g. simple_function M f  simple_function M g  (x. x  space M  f x = g x)  P f  P g"
  assumes set: "A. A  sets M  P (indicator A)"
  assumes mult: "u c. simple_function M u  P u  P (λx. c * u x)"
  assumes add: "u v. simple_function M u  P u  simple_function M v  (x. x  space M  u x = 0  v x = 0)  P v  P (λx. v x + u x)"
  shows "P u"
proof -
  show ?thesis
  proof (rule cong)
    fix x assume x: "x  space M"
    from simple_function_indicator_representation[OF u x]
    show "(yu ` space M. y * indicator (u -` {y}  space M) x) = u x" ..
  next
    show "simple_function M (λx. (yu ` space M. y * indicator (u -` {y}  space M) x))"
      apply (subst simple_function_cong)
      apply (rule simple_function_indicator_representation[symmetric])
      apply (auto intro: u)
      done
  next
    from u have "finite (u ` space M)"
      unfolding simple_function_def by auto
    then show "P (λx. yu ` space M. y * indicator (u -` {y}  space M) x)"
    proof induct
      case empty show ?case
        using set[of "{}"] by (simp add: indicator_def[abs_def])
    next
      case (insert x S)
      { fix z have "(yS. y * indicator (u -` {y}  space M) z) = 0 
          x * indicator (u -` {x}  space M) z = 0"
          using insert by (subst sum_eq_0_iff) (auto simp: indicator_def) }
      note disj = this
      from insert show ?case
        by (auto intro!: add mult set simple_functionD u simple_function_sum disj)
    qed
  qed fact
qed

lemmatag important› borel_measurable_induct
    [consumes 1, case_names cong set mult add seq, induct set: borel_measurable]:
  fixes u :: "'a  ennreal"
  assumes u: "u  borel_measurable M"
  assumes cong: "f g. f  borel_measurable M  g  borel_measurable M  (x. x  space M  f x = g x)  P g  P f"
  assumes set: "A. A  sets M  P (indicator A)"
  assumes mult': "u c. c < top  u  borel_measurable M  (x. x  space M  u x < top)  P u  P (λx. c * u x)"
  assumes add: "u v. u  borel_measurable M (x. x  space M  u x < top)  P u  v  borel_measurable M  (x. x  space M  v x < top)  (x. x  space M  u x = 0  v x = 0)  P v  P (λx. v x + u x)"
  assumes seq: "U. (i. U i  borel_measurable M)  (i x. x  space M  U i x < top)  (i. P (U i))  incseq U  u = (SUP i. U i)  P (SUP i. U i)"
  shows "P u"
  using u
proof (induct rule: borel_measurable_implies_simple_function_sequence')
  fix U assume U: "i. simple_function M (U i)" "incseq U" "i x. U i x < top" and sup: "x. (SUP i. U i x) = u x"
  have u_eq: "u = (SUP i. U i)"
    using u by (auto simp add: image_comp sup)

  have not_inf: "x i. x  space M  U i x < top"
    using U by (auto simp: image_iff eq_commute)

  from U have "i. U i  borel_measurable M"
    by (simp add: borel_measurable_simple_function)

  show "P u"
    unfolding u_eq
  proof (rule seq)
    fix i show "P (U i)"
      using simple_function M (U i) not_inf[of _ i]
    proof (induct rule: simple_function_induct_nn)
      case (mult u c)
      show ?case
      proof cases
        assume "c = 0  space M = {}  (xspace M. u x = 0)"
        with mult(1) show ?thesis
          by (intro cong[of "λx. c * u x" "indicator {}"] set)
             (auto dest!: borel_measurable_simple_function)
      next
        assume "¬ (c = 0  space M = {}  (xspace M. u x = 0))"
        then obtain x where "space M  {}" and x: "x  space M" "u x  0" "c  0"
          by auto
        with mult(3)[of x] have "c < top"
          by (auto simp: ennreal_mult_less_top)
        then have u_fin: "x'  space M  u x' < top" for x'
          using mult(3)[of x'] c  0 by (auto simp: ennreal_mult_less_top)
        then have "P u"
          by (rule mult)
        with u_fin c < top 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)
  qed fact+
qed

lemma simple_function_If_set:
  assumes sf: "simple_function M f" "simple_function M g" and A: "A  space M  sets M"
  shows "simple_function M (λx. if x  A then f x else g x)" (is "simple_function M ?IF")
proof -
  define F where "F x = f -` {x}  space M" for x
  define G where "G x = g -` {x}  space M" for x
  show ?thesis unfolding simple_function_def
  proof safe
    have "?IF ` space M  f ` space M  g ` space M" by auto
    from finite_subset[OF this] assms
    show "finite (?IF ` space M)" unfolding simple_function_def by auto
  next
    fix x assume "x  space M"
    then have *: "?IF -` {?IF x}  space M = (if x  A
      then ((F (f x)  (A  space M))  (G (f x) - (G (f x)  (A  space M))))
      else ((F (g x)  (A  space M))  (G (g x) - (G (g x)  (A  space M)))))"
      using sets.sets_into_space[OF A] by (auto split: if_split_asm simp: G_def F_def)
    have [intro]: "x. F x  sets M" "x. G x  sets M"
      unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto
    show "?IF -` {?IF x}  space M  sets M" unfolding * using A by auto
  qed
qed

lemma simple_function_If:
  assumes sf: "simple_function M f" "simple_function M g" and P: "{xspace M. P x}  sets M"
  shows "simple_function M (λx. if P x then f x else g x)"
proof -
  have "{xspace M. P x} = {x. P x}  space M" by auto
  with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp
qed

lemma simple_function_subalgebra:
  assumes "simple_function N f"
  and N_subalgebra: "sets N  sets M" "space N = space M"
  shows "simple_function M f"
  using assms unfolding simple_function_def by auto

lemma simple_function_comp:
  assumes T: "T  measurable M M'"
    and f: "simple_function M' f"
  shows "simple_function M (λx. f (T x))"
proof (intro simple_function_def[THEN iffD2] conjI ballI)
  have "(λx. f (T x)) ` space M  f ` space M'"
    using T unfolding measurable_def by auto
  then show "finite ((λx. f (T x)) ` space M)"
    using f unfolding simple_function_def by (auto intro: finite_subset)
  fix i assume i: "i  (λx. f (T x)) ` space M"
  then have "i  f ` space M'"
    using T unfolding measurable_def by auto
  then have "f -` {i}  space M'  sets M'"
    using f unfolding simple_function_def by auto
  then have "T -` (f -` {i}  space M')  space M  sets M"
    using T unfolding measurable_def by auto
  also have "T -` (f -` {i}  space M')  space M = (λx. f (T x)) -` {i}  space M"
    using T unfolding measurable_def by auto
  finally show "(λx. f (T x)) -` {i}  space M  sets M" .
qed

subsection "Simple integral"

definitiontag important› simple_integral :: "'a measure  ('a  ennreal)  ennreal" ("integralS") where
  "integralS M f = (x  f ` space M. x * emeasure M (f -` {x}  space M))"

syntax
  "_simple_integral" :: "pttrn  ennreal  'a measure  ennreal" ("S _. _ _" [60,61] 110)

translations
  "S x. f M" == "CONST simple_integral M (%x. f)"

lemma simple_integral_cong:
  assumes "t. t  space M  f t = g t"
  shows "integralS M f = integralS M g"
proof -
  have "f ` space M = g ` space M"
    "x. f -` {x}  space M = g -` {x}  space M"
    using assms by (auto intro!: image_eqI)
  thus ?thesis unfolding simple_integral_def by simp
qed

lemma simple_integral_const[simp]:
  "(Sx. c M) = c * (emeasure M) (space M)"
proof (cases "space M = {}")
  case True thus ?thesis unfolding simple_integral_def by simp
next
  case False hence "(λx. c) ` space M = {c}" by auto
  thus ?thesis unfolding simple_integral_def by simp
qed

lemma simple_function_partition:
  assumes f: "simple_function M f" and g: "simple_function M g"
  assumes sub: "x y. x  space M  y  space M  g x = g y  f x = f y"
  assumes v: "x. x  space M  f x = v (g x)"
  shows "integralS M f = (yg ` space M. v y * emeasure M {xspace M. g x = y})"
    (is "_ = ?r")
proof -
  from f g have [simp]: "finite (f`space M)" "finite (g`space M)"
    by (auto simp: simple_function_def)
  from f g have [measurable]: "f  measurable M (count_space UNIV)" "g  measurable M (count_space UNIV)"
    by (auto intro: measurable_simple_function)

  { fix y assume "y  space M"
    then have "f ` space M  {i. xspace M. i = f x  g y = g x} = {v (g y)}"
      by (auto cong: sub simp: v[symmetric]) }
  note eq = this

  have "integralS M f =
    (yf`space M. y * (zg`space M.
      if xspace M. y = f x  z = g x then emeasure M {xspace M. g x = z} else 0))"
    unfolding simple_integral_def
  proof (safe intro!: sum.cong ennreal_mult_left_cong)
    fix y assume y: "y  space M" "f y  0"
    have [simp]: "g ` space M  {z. xspace M. f y = f x  z = g x} =
        {z. xspace M. f y = f x  z = g x}"
      by auto
    have eq:"(i{z. xspace M. f y = f x  z = g x}. {x  space M. g x = i}) =
        f -` {f y}  space M"
      by (auto simp: eq_commute cong: sub rev_conj_cong)
    have "finite (g`space M)" by simp
    then have "finite {z. xspace M. f y = f x  z = g x}"
      by (rule rev_finite_subset) auto
    then show "emeasure M (f -` {f y}  space M) =
      (zg ` space M. if xspace M. f y = f x  z = g x then emeasure M {x  space M. g x = z} else 0)"
      apply (simp add: sum.If_cases)
      apply (subst sum_emeasure)
      apply (auto simp: disjoint_family_on_def eq)
      done
  qed
  also have " = (yf`space M. (zg`space M.
      if xspace M. y = f x  z = g x then y * emeasure M {xspace M. g x = z} else 0))"
    by (auto intro!: sum.cong simp: sum_distrib_left)
  also have " = ?r"
    by (subst sum.swap)
       (auto intro!: sum.cong simp: sum.If_cases scaleR_sum_right[symmetric] eq)
  finally show "integralS M f = ?r" .
qed

lemma simple_integral_add[simp]:
  assumes f: "simple_function M f" and "x. 0  f x" and g: "simple_function M g" and "x. 0  g x"
  shows "(Sx. f x + g x M) = integralS M f + integralS M g"
proof -
  have "(Sx. f x + g x M) =
    (y(λx. (f x, g x))`space M. (fst y + snd y) * emeasure M {xspace M. (f x, g x) = y})"
    by (intro simple_function_partition) (auto intro: f g)
  also have " = (y(λx. (f x, g x))`space M. fst y * emeasure M {xspace M. (f x, g x) = y}) +
    (y(λx. (f x, g x))`space M. snd y * emeasure M {xspace M. (f x, g x) = y})"
    using assms(2,4) by (auto intro!: sum.cong distrib_right simp: sum.distrib[symmetric])
  also have "(y(λx. (f x, g x))`space M. fst y * emeasure M {xspace M. (f x, g x) = y}) = (Sx. f x M)"
    by (intro simple_function_partition[symmetric]) (auto intro: f g)
  also have "(y(λx. (f x, g x))`space M. snd y * emeasure M {xspace M. (f x, g x) = y}) = (Sx. g x M)"
    by (intro simple_function_partition[symmetric]) (auto intro: f g)
  finally show ?thesis .
qed

lemma simple_integral_sum[simp]:
  assumes "i x. i  P  0  f i x"
  assumes "i. i  P  simple_function M (f i)"
  shows "(Sx. (iP. f i x) M) = (iP. integralS M (f i))"
proof cases
  assume "finite P"
  from this assms show ?thesis
    by induct (auto)
qed auto

lemma simple_integral_mult[simp]:
  assumes f: "simple_function M f"
  shows "(Sx. c * f x M) = c * integralS M f"
proof -
  have "(Sx. c * f x M) = (yf ` space M. (c * y) * emeasure M {xspace M. f x = y})"
    using f by (intro simple_function_partition) auto
  also have " = c * integralS M f"
    using f unfolding simple_integral_def
    by (subst sum_distrib_left) (auto simp: mult.assoc Int_def conj_commute)
  finally show ?thesis .
qed

lemma simple_integral_mono_AE:
  assumes f[measurable]: "simple_function M f" and g[measurable]: "simple_function M g"
  and mono: "AE x in M. f x  g x"
  shows "integralS M f  integralS M g"
proof -
  let  = "λP. emeasure M {xspace M. P x}"
  have "integralS M f = (y(λx. (f x, g x))`space M. fst y *  (λx. (f x, g x) = y))"
    using f g by (intro simple_function_partition) auto
  also have "  (y(λx. (f x, g x))`space M. snd y *  (λx. (f x, g x) = y))"
  proof (clarsimp intro!: sum_mono)
    fix x assume "x  space M"
    let ?M = " (λy. f y = f x  g y = g x)"
    show "f x * ?M  g x * ?M"
    proof cases
      assume "?M  0"
      then have "0 < ?M"
        by (simp add: less_le)
      also have "   (λy. f x  g x)"
        using mono by (intro emeasure_mono_AE) auto
      finally have "¬ ¬ f x  g x"
        by (intro notI) auto
      then show ?thesis
        by (intro mult_right_mono) auto
    qed simp
  qed
  also have " = integralS M g"
    using f g by (intro simple_function_partition[symmetric]) auto
  finally show ?thesis .
qed

lemma simple_integral_mono:
  assumes "simple_function M f" and "simple_function M g"
  and mono: " x. x  space M  f x  g x"
  shows "integralS M f  integralS M g"
  using assms by (intro simple_integral_mono_AE) auto

lemma simple_integral_cong_AE:
  assumes "simple_function M f" and "simple_function M g"
  and "AE x in M. f x = g x"
  shows "integralS M f = integralS M g"
  using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE)

lemma simple_integral_cong':
  assumes sf: "simple_function M f" "simple_function M g"
  and mea: "(emeasure M) {xspace M. f x  g x} = 0"
  shows "integralS M f = integralS M g"
proof (intro simple_integral_cong_AE sf AE_I)
  show "(emeasure M) {xspace M. f x  g x} = 0" by fact
  show "{x  space M. f x  g x}  sets M"
    using sf[THEN borel_measurable_simple_function] by auto
qed simp

lemma simple_integral_indicator:
  assumes A: "A  sets M"
  assumes f: "simple_function M f"
  shows "(Sx. f x * indicator A x M) =
    (x  f ` space M. x * emeasure M (f -` {x}  space M  A))"
proof -
  have eq: "(λx. (f x, indicator A x)) ` space M  {x. snd x = 1} = (λ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: "x. f x  f ` A  f -` {f x}  space M  A = {}"
    by (auto simp: image_iff)

  have "(Sx. f x * indicator A x M) =
    (y(λx. (f x, indicator A x))`space M. (fst y * snd y) * emeasure M {xspace M. (f x, indicator A x) = y})"
    using assms by (intro simple_function_partition) auto
  also have " = (y(λx. (f x, indicator A x::ennreal))`space M.
    if snd y = 1 then fst y * emeasure M (f -` {fst y}  space M  A) else 0)"
    by (auto simp: indicator_def split: if_split_asm intro!: arg_cong2[where f="(*)"] arg_cong2[where f=emeasure] sum.cong)
  also have " = (y(λx. (f x, 1::ennreal))`A. fst y * emeasure M (f -` {fst y}  space M  A))"
    using assms by (subst sum.If_cases) (auto intro!: simple_functionD(1) simp: eq)
  also have " = (yfst`(λx. (f x, 1::ennreal))`A. y * emeasure M (f -` {y}  space M  A))"
    by (subst sum.reindex [of fst]) (auto simp: inj_on_def)
  also have " = (x  f ` space M. x * emeasure M (f -` {x}  space M  A))"
    using A[THEN sets.sets_into_space]
    by (intro sum.mono_neutral_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2)
  finally show ?thesis .
qed

lemma simple_integral_indicator_only[simp]:
  assumes "A  sets M"
  shows "integralS M (indicator A) = emeasure M A"
  using simple_integral_indicator[OF assms, of "λx. 1"] sets.sets_into_space[OF assms]
  by (simp_all add: image_constant_conv Int_absorb1 split: if_split_asm)

lemma simple_integral_null_set:
  assumes "simple_function M u" "x. 0  u x" and "N  null_sets M"
  shows "(Sx. u x * indicator N x M) = 0"
proof -
  have "AE x in M. indicator N x = (0 :: ennreal)"
    using N  null_sets M by (auto simp: indicator_def intro!: AE_I[of _ _ N])
  then have "(Sx. u x * indicator N x M) = (Sx. 0 M)"
    using assms apply (intro simple_integral_cong_AE) by auto
  then show ?thesis by simp
qed

lemma simple_integral_cong_AE_mult_indicator:
  assumes sf: "simple_function M f" and eq: "AE x in M. x  S" and "S  sets M"
  shows "integralS M f = (Sx. f x * indicator S x M)"
  using assms by (intro simple_integral_cong_AE) auto

lemma simple_integral_cmult_indicator:
  assumes A: "A  sets M"
  shows "(Sx. c * indicator A x M) = c * emeasure M A"
  using simple_integral_mult[OF simple_function_indicator[OF A]]
  unfolding simple_integral_indicator_only[OF A] by simp

lemma simple_integral_nonneg:
  assumes f: "simple_function M f" and ae: "AE x in M. 0  f x"
  shows "0  integralS M f"
proof -
  have "integralS M (λx. 0)  integralS M f"
    using simple_integral_mono_AE[OF _ f ae] by auto
  then show ?thesis by simp
qed

subsection ‹Integral on nonnegative functions›

definitiontag important› nn_integral :: "'a measure  ('a  ennreal)  ennreal" ("integralN") where
  "integralN M f = (SUP g  {g. simple_function M g  g  f}. integralS M g)"

syntax
  "_nn_integral" :: "pttrn  ennreal  'a measure  ennreal" ("+((2 _./ _)/ _)" [60,61] 110)

translations
  "+x. f M" == "CONST nn_integral M (λx. f)"

lemma nn_integral_def_finite:
  "integralN M f = (SUP g  {g. simple_function M g  g  f  (x. g x < top)}. integralS M g)"
    (is "_ = Sup (?A ` ?f)")
  unfolding nn_integral_def
proof (safe intro!: antisym SUP_least)
  fix g assume g[measurable]: "simple_function M g" "g  f"

  show "integralS M g  Sup (?A ` ?f)"
  proof cases
    assume ae: "AE x in M. g x  top"
    let ?G = "{x  space M. g x  top}"
    have "integralS M g = integralS M (λ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 "  Sup (?A ` ?f)"
      using g by (intro SUP_upper) (auto simp: le_fun_def less_top split: split_indicator)
    finally show ?thesis .
  next
    assume nAE: "¬ (AE x in M. g x  top)"
    then have "emeasure M {xspace M. g x = top}  0" (is "emeasure M ?G  0")
      by (subst (asm) AE_iff_measurable[OF _ refl]) auto
    then have "top = (SUP n. (Sx. of_nat n * indicator ?G x M))"
      by (simp add: ennreal_SUP_of_nat_eq_top ennreal_top_eq_mult_iff SUP_mult_right_ennreal[symmetric])
    also have "  Sup (?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)

lemma nn_integral_mono_AE:
  assumes ae: "AE x in M. u x  v x" shows "integralN M u  integralN M v"
  unfolding nn_integral_def
proof (safe intro!: SUP_mono)
  fix n assume n: "simple_function M n" "n  u"
  from ae[THEN AE_E] obtain N
    where N: "{x  space M. ¬ u x  v x}  N" "emeasure M N = 0" "N  sets M"
    by auto
  then have ae_N: "AE x in M. x  N" by (auto intro: AE_not_in)
  let ?n = "λx. n x * indicator (space M - N) x"
  have "AE x in M. n x  ?n x" "simple_function M ?n"
    using n N ae_N by auto
  moreover
  { fix x have "?n x  v x"
    proof cases
      assume x: "x  space M - N"
      with N have "u x  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  v" by (auto simp: le_funI)
  moreover have "integralS M n  integralS M ?n"
    using ae_N N n by (auto intro!: simple_integral_mono_AE)
  ultimately show "m{g. simple_function M g  g  v}. integralS M n  integralS M m"
    by force
qed

lemma nn_integral_mono:
  "(x. x  space M  u x  v x)  integralN M u  integralN M v"
  by (auto intro: nn_integral_mono_AE)

lemma mono_nn_integral: "mono F  mono (λx. integralN M (F x))"
  by (auto simp add: mono_def le_fun_def intro!: nn_integral_mono)

lemma nn_integral_cong_AE:
  "AE x in M. u x = v x  integralN M u = integralN M v"
  by (auto simp: eq_iff intro!: nn_integral_mono_AE)

lemma nn_integral_cong:
  "(x. x  space M  u x = v x)  integralN M u = integralN M v"
  by (auto intro: nn_integral_cong_AE)

lemma nn_integral_cong_simp:
  "(x. x  space M =simp=> u x = v x)  integralN M u = integralN M v"
  by (auto intro: nn_integral_cong simp: simp_implies_def)

lemma incseq_nn_integral:
  assumes "incseq f" shows "incseq (λi. integralN M (f i))"
proof -
  have "i x. f i x  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" shows "integralN M f = integralS M f"
proof -
  let ?f = "λx. f x * indicator (space M) x"
  have f': "simple_function M ?f" using f by auto
  have "integralN M ?f  integralS M ?f" using f'
    by (force intro!: SUP_least simple_integral_mono simp: le_fun_def nn_integral_def)
  moreover have "integralS M ?f  integralN M ?f"
    unfolding nn_integral_def
    using f' by (auto intro!: SUP_upper)
  ultimately show ?thesis
    by (simp cong: nn_integral_cong simple_integral_cong)
qed

text ‹Beppo-Levi monotone convergence theorem›
lemma nn_integral_monotone_convergence_SUP:
  assumes f: "incseq f" and [measurable]: "i. f i  borel_measurable M"
  shows "(+ x. (SUP i. f i x) M) = (SUP i. integralN M (f i))"
proof (rule antisym)
  show "(+ x. (SUP i. f i x) M)  (SUP i. (+ x. f i x M))"
    unfolding nn_integral_def_finite[of _ "λx. SUP i. f i x"]
  proof (safe intro!: SUP_least)
    fix u assume sf_u[simp]: "simple_function M u" and
      u: "u  (λx. SUP i. f i x)" and u_range: "x. u x < top"
    note sf_u[THEN borel_measurable_simple_function, measurable]
    show "integralS M u  (SUP j. +x. f j x M)"
    proof (rule ennreal_approx_unit)
      fix a :: ennreal assume "a < 1"
      let ?au = "λx. a * u x"

      let ?B = "λc i. {xspace M. ?au x = c  c  f i x}"
      have "integralS M ?au = (c?au`space M. c * (SUP i. emeasure M (?B c i)))"
        unfolding simple_integral_def
      proof (intro sum.cong ennreal_mult_left_cong refl)
        fix c assume "c  ?au ` space M" "c  0"
        { fix x' assume x': "x'  space M" "?au x' = c"
          with c  0 u_range have "?au x' < 1 * u x'"
            by (intro ennreal_mult_strict_right_mono a < 1) (auto simp: less_le)
          also have "  (SUP i. f i x')"
            using u by (auto simp: le_fun_def)
          finally have "i. ?au x'  f i x'"
            by (auto simp: less_SUP_iff intro: less_imp_le) }
        then have *: "?au -` {c}  space M = (i. ?B c i)"
          by auto
        show "emeasure M (?au -` {c}  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 " = (SUP i. c?au`space M. c * emeasure M (?B c i))"
        unfolding SUP_mult_left_ennreal using f
        by (intro ennreal_SUP_sum[symmetric])
           (auto intro!: mult_mono emeasure_mono simp: incseq_def le_fun_def intro: order_trans)
      also have "  (SUP i. integralN M (f i))"
      proof (intro SUP_subset_mono order_refl)
        fix i
        have "(c?au`space M. c * emeasure M (?B c i)) =
          (Sx. (a * u x) * indicator {xspace M. a * u x  f i x} x M)"
          by (subst simple_integral_indicator)
             (auto intro!: sum.cong ennreal_mult_left_cong arg_cong2[where f=emeasure])
        also have " = (+x. (a * u x) * indicator {xspace M. a * u x  f i x} x M)"
          by (rule nn_integral_eq_simple_integral[symmetric]) simp
        also have "  (+x. f i x M)"
          by (intro nn_integral_mono) (auto split: split_indicator)
        finally show "(c?au`space M. c * emeasure M (?B c i))  (+x. f i x M)" .
      qed
      finally show "a * integralS M u  (SUP i. integralN M (f i))"
        by simp
    qed
  qed
qed (auto intro!: SUP_least SUP_upper nn_integral_mono)

lemma sup_continuous_nn_integral[order_continuous_intros]:
  assumes f: "y. sup_continuous (f y)"
  assumes [measurable]: "x. (λy. f y x)  borel_measurable M"
  shows "sup_continuous (λx. (+y. f y x M))"
  unfolding sup_continuous_def
proof safe
  fix C :: "nat  'b" assume C: "incseq C"
  with sup_continuous_mono[OF f] show "(+ y. f y (Sup (C ` UNIV)) M) = (SUP i. + y. f y (C i) M)"
    unfolding sup_continuousD[OF f C]
    by (subst nn_integral_monotone_convergence_SUP) (auto simp: mono_def le_fun_def)
qed

theorem nn_integral_monotone_convergence_SUP_AE:
  assumes f: "i. AE x in M. f i x  f (Suc i) x" "i. f i  borel_measurable M"
  shows "(+ x. (SUP i. f i x) M) = (SUP i. integralN M (f i))"
proof -
  from f have "AE x in M. i. f i x  f (Suc i) x"
    by (simp add: AE_all_countable)
  from this[THEN AE_E] obtain N
    where N: "{x  space M. ¬ (i. f i x  f (Suc i) x)}  N" "emeasure M N = 0" "N  sets M"
    by auto
  let ?f = "λi x. if x  space M - N then f i x else 0"
  have f_eq: "AE x in M. i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N])
  then have "(+ x. (SUP i. f i x) M) = (+ x. (SUP i. ?f i x) M)"
    by (auto intro!: nn_integral_cong_AE)
  also have " = (SUP i. (+ x. ?f i x M))"
  proof (rule nn_integral_monotone_convergence_SUP)
    show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI)
    { fix i show "(λx. if x  space M - N then f i x else 0)  borel_measurable M"
        using f N(3) by (intro measurable_If_set) auto }
  qed
  also have " = (SUP i. (+ x. f i x M))"
    using f_eq by (force intro!: arg_cong[where f = "λf. Sup (range f)"] nn_integral_cong_AE ext)
  finally show ?thesis .
qed

lemma nn_integral_monotone_convergence_simple:
  "incseq f  (i. simple_function M (f i))  (SUP i. S x. f i x M) = (+x. (SUP i. f i x) M)"
  using 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" "i. simple_function M (f i)"
  and g: "incseq g" "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. integralS M (f i)) = (SUP i. integralS M (g i))"
    (is "Sup (?F ` _) = Sup (?G ` _)")
proof -
  have "(SUP i. integralS M (f i)) = (+x. (SUP i. f i x) M)"
    using f by (rule nn_integral_monotone_convergence_simple)
  also have " = (+x. (SUP i. g i x) M)"
    unfolding eq[THEN nn_integral_cong_AE] ..
  also have " = (SUP i. ?G i)"
    using g by (rule nn_integral_monotone_convergence_simple[symmetric])
  finally show ?thesis by simp
qed

lemma nn_integral_const[simp]: "(+ x. c M) = c * emeasure M (space M)"
  by (subst nn_integral_eq_simple_integral) auto

lemma nn_integral_linear:
  assumes f: "f  borel_measurable M" and g: "g  borel_measurable M"
  shows "(+ x. a * f x + g x M) = a * integralN M f + integralN M g"
    (is "integralN M ?L = _")
proof -
  obtain u
    where "i. simple_function M (u i)" "incseq u" "i x. u i x < top" "x. (SUP i. u i x) = f x"
    using borel_measurable_implies_simple_function_sequence' f(1)
    by auto
  note u = nn_integral_monotone_convergence_simple[OF this(2,1)] this

  obtain v where
    "i. simple_function M (v i)" "incseq v" "i x. v i x < top" "x. (SUP i. v i x) = g x"
    using borel_measurable_implies_simple_function_sequence' g(1)
    by auto
  note v = nn_integral_monotone_convergence_simple[OF this(2,1)] this

  let ?L' = "λi x. a * u i x + v i x"

  have "?L  borel_measurable M" using assms by auto
  from borel_measurable_implies_simple_function_sequence'[OF this]
  obtain l where "i. simple_function M (l i)" "incseq l" "i x. l i x < top" "x. (SUP i. l i x) = a * f x + g x"
    by auto
  note l = nn_integral_monotone_convergence_simple[OF this(2,1)] this

  have inc: "incseq (λi. a * integralS M (u i))" "incseq (λi. integralS M (v i))"
    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. integralS M (l i)) = (SUP i. integralS M (?L' i))"
  proof (rule SUP_simple_integral_sequences[OF l(3,2)])
    show "incseq ?L'" "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
      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 u(5) v(5) by (intro AE_I2) auto
  qed
  also have " = (SUP i. a * integralS M (u i) + integralS M (v i))"
    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: "f  borel_measurable M  (+ x. c * f x M) = c * integralN M f"
  using nn_integral_linear[of f M "λx. 0" c] by simp

lemma nn_integral_multc: "f  borel_measurable M  (+ x. f x * c M) = integralN M f * c"
  unfolding mult.commute[of _ c] nn_integral_cmult by simp

lemma nn_integral_divide: "f  borel_measurable M  (+ x. f x / c M) = (+ x. f x M) / c"
   unfolding divide_ennreal_def by (rule nn_integral_multc)

lemma nn_integral_indicator[simp]: "A  sets M  (+ x. indicator A xM) = (emeasure M) A"
  by (subst nn_integral_eq_simple_integral) (auto simp: simple_integral_indicator)

lemma nn_integral_cmult_indicator: "A  sets M  (+ x. c * indicator A x M) = c * emeasure M A"
  by (subst nn_integral_eq_simple_integral) (auto)

lemma nn_integral_indicator':
  assumes [measurable]: "A  space M  sets M"
  shows "(+ x. indicator A x M) = emeasure M (A  space M)"
proof -
  have "(+ x. indicator A x M) = (+ x. indicator (A  space M) x M)"
    by (intro nn_integral_cong) (simp split: split_indicator)
  also have " = emeasure M (A  space M)"
    by simp
  finally show ?thesis .
qed

lemma nn_integral_indicator_singleton[simp]:
  assumes [measurable]: "{y}  sets M" shows "(+x. f x * indicator {y} x M) = f y * emeasure M {y}"
proof -
  have "(+x. f x * indicator {y} x M) = (+x. f y * indicator {y} x M)"
    by (auto intro!: nn_integral_cong split: split_indicator)
  then show ?thesis
    by (simp add: nn_integral_cmult)
qed

lemma nn_integral_set_ennreal:
  "(+x. ennreal (f x) * indicator A x M) = (+x. ennreal (f x * indicator A x) M)"
  by (rule nn_integral_cong) (simp split: split_indicator)

lemma nn_integral_indicator_singleton'[simp]:
  assumes [measurable]: "{y}  sets M"
  shows "(+x. ennreal (f x * indicator {y} x) M) = f y * emeasure M {y}"
  by (subst nn_integral_set_ennreal[symmetric]) (simp)

lemma nn_integral_add:
  "f  borel_measurable M  g  borel_measurable M  (+ x. f x + g x M) = integralN M f + integralN M g"
  using nn_integral_linear[of f M g 1] by simp

lemma nn_integral_sum:
  "(i. i  P  f i  borel_measurable M)  (+ x. (iP. f i x) M) = (iP. integralN M (f i))"
  by (induction P rule: infinite_finite_induct) (auto simp: nn_integral_add)

theorem nn_integral_suminf:
  assumes f: "i. f i  borel_measurable M"
  shows "(+ x. (i. f i x) M) = (i. integralN M (f i))"
proof -
  have all_pos: "AE x in M. i. 0  f i x"
    using assms by (auto simp: AE_all_countable)
  have "(i. integralN M (f i)) = (SUP n. i<n. integralN M (f i))"
    by (rule suminf_eq_SUP)
  also have " = (SUP n. +x. (i<n. f i x) M)"
    unfolding nn_integral_sum[OF f] ..
  also have " = +x. (SUP n. i<n. f i x) M" using f all_pos
    by (intro nn_integral_monotone_convergence_SUP_AE[symmetric])
       (elim AE_mp, auto simp: sum_nonneg simp del: sum.lessThan_Suc intro!: AE_I2 sum_mono2)
  also have " = +x. (i. f i x) 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: "x. x  space M  f x < "
  assumes f[measurable]: "simple_function M f"
  assumes supp: "emeasure M {xspace M. f x  0} < "
  shows "nn_integral M f < "
proof cases
  assume "space M = {}"
  then have "nn_integral M f = (+x. 0 M)"
    by (intro nn_integral_cong) auto
  then show ?thesis by simp
next
  assume "space M  {}"
  with simple_functionD(1)[OF f] bnd have bnd: "0  Max (f`space M)  Max (f`space M) < "
    by (subst Max_less_iff) (auto simp: Max_ge_iff)

  have "nn_integral M f  (+x. Max (f`space M) * indicator {xspace M. f x  0} x M)"
  proof (rule nn_integral_mono)
    fix x assume "x  space M"
    with f show "f x  Max (f ` space M) * indicator {x  space M. f x  0} x"
      by (auto split: split_indicator intro!: Max_ge simple_functionD)
  qed
  also have " < "
    using bnd supp by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top)
  finally show ?thesis .
qed

theorem nn_integral_Markov_inequality:
  assumes u: "(λx. u x * indicator A x)  borel_measurable M" and "A  sets M"
  shows "(emeasure M) ({xA. 1  c * u x})  c * (+ x. u x * indicator A x M)"
    (is "(emeasure M) ?A  _ * ?PI")
proof -
  define u' where "u' = (λx. u x * indicator A x)"
  have [measurable]: "u'  borel_measurable M"
    using u unfolding u'_def .
  have "{xspace M. c * u' x  1}  sets M"
    by measurable
  also have "{xspace M. c * u' x  1} = ?A"
    using sets.sets_into_space[OF A  sets M] by (auto simp: u'_def indicator_def)
  finally have "(emeasure M) ?A = (+ x. indicator ?A x M)"
    using nn_integral_indicator by simp
  also have "  (+ x. c * (u x * indicator A x) M)"
    using u by (auto intro!: nn_integral_mono_AE simp: indicator_def)
  also have " = c * (+ x. u x * indicator A x M)"
    using assms by (auto intro!: nn_integral_cmult)
  finally show ?thesis .
qed

lemma Chernoff_ineq_nn_integral_ge:
  assumes s: "s > 0" and [measurable]: "A  sets M"
  assumes [measurable]: "(λx. f x * indicator A x)  borel_measurable M"
  shows   "emeasure M {xA. f x  a} 
           ennreal (exp (-s * a)) * nn_integral M (λx. ennreal (exp (s * f x)) * indicator A x)"
proof -
  define f' where "f' = (λx. f x * indicator A x)"
  have [measurable]: "f'  borel_measurable M"
    using assms(3) unfolding f'_def by assumption
  have "(λx. ennreal (exp (s * f' x)) * indicator A x)  borel_measurable M"
    by simp
  also have "(λx. ennreal (exp (s * f' x)) * indicator A x) =
             (λx. ennreal (exp (s * f x)) * indicator A x)"
    by (auto simp: f'_def indicator_def fun_eq_iff)
  finally have meas: "  borel_measurable M" .

  have "{xA. f x  a} = {xA. ennreal (exp (-s * a)) * ennreal (exp (s * f x))  1}"
    using s by (auto simp: exp_minus field_simps simp flip: ennreal_mult)
  also have "emeasure M   ennreal (exp (-s * a)) *
               (+x. ennreal (exp (s * f x)) * indicator A x M)"
    by (intro order.trans[OF nn_integral_Markov_inequality] meas) auto
  finally show ?thesis .
qed

lemma Chernoff_ineq_nn_integral_le:
  assumes s: "s > 0" and [measurable]: "A  sets M"
  assumes [measurable]: "f  borel_measurable M"
  shows   "emeasure M {xA. f x  a} 
           ennreal (exp (s * a)) * nn_integral M (λx. ennreal (exp (-s * f x)) * indicator A x)"
  using Chernoff_ineq_nn_integral_ge[of s A M "λx. -f x" "-a"] assms by simp

lemma nn_integral_noteq_infinite:
  assumes g: "g  borel_measurable M" and "integralN M g  "
  shows "AE x in M. g x  "
proof (rule ccontr)
  assume c: "¬ (AE x in M. g x  )"
  have "(emeasure M) {xspace M. g x = }  0"
    using c g by (auto simp add: AE_iff_null)
  then have "0 < (emeasure M) {xspace M. g x = }"
    by (auto simp: zero_less_iff_neq_zero)
  then have " =  * (emeasure M) {xspace M. g x = }"
    by (auto simp: ennreal_top_eq_mult_iff)
  also have "  (+x.  * indicator {xspace M. g x = } x M)"
    using g by (subst nn_integral_cmult_indicator) auto
  also have "  integralN M g"
    using assms by (auto intro!: nn_integral_mono_AE simp: indicator_def)
  finally show False
    using integralN M g   by (auto simp: top_unique)
qed

lemma nn_integral_PInf:
  assumes f: "f  borel_measurable M" and not_Inf: "integralN M f  "
  shows "emeasure M (f -` {}  space M) = 0"
proof -
  have " * emeasure M (f -` {}  space M) = (+ x.  * indicator (f -` {}  space M) x M)"
    using f by (subst nn_integral_cmult_indicator) (auto simp: measurable_sets)
  also have "  integralN M f"
    by (auto intro!: nn_integral_mono simp: indicator_def)
  finally have " * (emeasure M) (f -` {}  space M)  integralN M f"
    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  integralS M f    emeasure M (f -` {}  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  borel_measurable M" "integralN M f  " shows "AE x in M. f x  "
proof (rule AE_I)
  show "(emeasure M) (f -` {}  space M) = 0"
    by (rule nn_integral_PInf[OF assms])
  show "f -` {}  space M  sets M"
    using assms by (auto intro: borel_measurable_vimage)
qed auto

lemma nn_integral_diff:
  assumes f: "f  borel_measurable M"
  and g: "g  borel_measurable M"
  and fin: "integralN M g  "
  and mono: "AE x in M. g x  f x"
  shows "(+ x. f x - g x M) = integralN M f - integralN M g"
proof -
  have diff: "(λx. f x - g x)  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 **: "integralN M f = (+x. f x - g x M) + integralN M g"
    unfolding nn_integral_add[OF diff g, symmetric]
    by (rule nn_integral_cong_AE)
  show ?thesis unfolding **
    using fin
    by (cases rule: ennreal2_cases[of "+ x. f x - g x M" "integralN M g"]) auto
qed

lemma nn_integral_mult_bounded_inf:
  assumes f: "f  borel_measurable M" "(+x. f x M) < " and c: "c  " and ae: "AE x in M. g x  c * f x"
  shows "(+x. g x M) < "
proof -
  have "(+x. g x M)  (+x. c * f x M)"
    by (intro nn_integral_mono_AE ae)
  also have "(+x. c * f x M) < "
    using c f by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top top_unique not_less)
  finally show ?thesis .
qed

text ‹Fatou's lemma: convergence theorem on limes inferior›

lemma nn_integral_monotone_convergence_INF_AE':
  assumes f: "i. AE x in M. f (Suc i) x  f i x" and [measurable]: "i. f i  borel_measurable M"
    and *: "(+ x. f 0 x M) < "
  shows "(+ x. (INF i. f i x) M) = (INF i. integralN M (f i))"
proof (rule ennreal_minus_cancel)
  have "integralN M (f 0) - (+ x. (INF i. f i x) M) = (+x. f 0 x - (INF i. f i x) M)"
  proof (rule nn_integral_diff[symmetric])
    have "(+ x. (INF i. f i x) M)  (+ x. f 0 x M)"
      by (intro nn_integral_mono INF_lower) simp
    with * show "(+ x. (INF i. f i x) M)  "
      by simp
  qed (auto intro: INF_lower)
  also have " = (+x. (SUP i. f 0 x - f i x) M)"
    by (simp add: ennreal_INF_const_minus)
  also have " = (SUP i. (+x. f 0 x - f i x M))"
  proof (intro nn_integral_monotone_convergence_SUP_AE)
    show "AE x in M. f 0 x - f i x  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 " = (SUP i. nn_integral M (f 0) - (+x. f i x M))"
  proof (subst nn_integral_diff[symmetric])
    fix i
    have dec: "AE x in M. i. f (Suc i) x  f i x"
      unfolding AE_all_countable using f by auto
    then show "AE x in M. f i x  f 0 x"
      using dec by eventually_elim (auto intro: lift_Suc_antimono_le[of "λi. f i x" 0 i for x])
    then have "(+ x. f i x M)  (+ x. f 0 x M)"
      by (rule nn_integral_mono_AE)
    with * show "(+ x. f i x M)  "
      by simp
  qed (insert f, auto simp: decseq_def le_fun_def)
  finally show "integralN M (f 0) - (+ x. (INF i. f i x) M) =
    integralN M (f 0) - (INF i. + x. f i x M)"
    by (simp add: ennreal_INF_const_minus)
qed (insert *, auto intro!: nn_integral_mono intro: INF_lower)

theorem nn_integral_monotone_convergence_INF_AE:
  fixes f :: "nat  'a  ennreal"
  assumes f: "i. AE x in M. f (Suc i) x  f i x"
    and [measurable]: "i. f i  borel_measurable M"
    and fin: "(+ x. f i x M) < "
  shows "(+ x. (INF i. f i x) M) = (INF i. integralN M (f i))"
proof -
  { fix f :: "nat  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. i. f (Suc i) x  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 "(+ x. (INF i. f i x) M) = (+ x. (INF n. f (n + i) x) M)"
    by (rule nn_integral_cong_AE)
  also have " = (INF n. (+ x. f (n + i) x M))"
    by (rule nn_integral_monotone_convergence_INF_AE') (insert assms, auto)
  also have " = (INF n. (+ x. f n x M))"
    by (intro INF_shift[symmetric] decseq_SucI nn_integral_mono_AE f)
  finally show ?thesis .
qed

lemma nn_integral_monotone_convergence_INF_decseq:
  assumes f: "decseq f" and *: "i. f i  borel_measurable M" "(+ x. f i x M) < "
  shows "(+ x. (INF i. f i x) M) = (INF i. integralN M (f i))"
  using nn_integral_monotone_convergence_INF_AE[of f M i, OF _ *] f by (simp add: decseq_SucD le_funD)

theorem nn_integral_liminf:
  fixes u :: "nat  'a  ennreal"
  assumes u: "i. u i  borel_measurable M"
  shows "(+ x. liminf (λn. u n x) M)  liminf (λn. integralN M (u n))"
proof -
  have "(+ x. liminf (λn. u n x) M) = (SUP n. + x. (INF i{n..}. u i x) 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 "  liminf (λn. integralN M (u n))"
    by (auto simp: liminf_SUP_INF intro!: SUP_mono INF_greatest nn_integral_mono INF_lower)
  finally show ?thesis .
qed

theorem nn_integral_limsup:
  fixes u :: "nat  'a  ennreal"
  assumes [measurable]: "i. u i  borel_measurable M" "w  borel_measurable M"
  assumes bounds: "i. AE x in M. u i x  w x" and w: "(+x. w x M) < "
  shows "limsup (λn. integralN M (u n))  (+ x. limsup (λn. u n x) M)"
proof -
  have bnd: "AE x in M. i. u i x  w x"
    using bounds by (auto simp: AE_all_countable)
  then have "(+ x. (SUP n. u n x) M)  (+ x. w x M)"
    by (auto intro!: nn_integral_mono_AE elim: eventually_mono intro: SUP_least)
  then have "(+ x. limsup (λn. u n x) M) = (INF n. + x. (SUP i{n..}. u i x) 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 "  limsup (λn. integralN 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" "i. f i  borel_measurable M"
    and u: "x. (λi. f i x)  u x"
  shows "(λn. integralN M (f n))  integralN M u"
proof -
  have "(λn. integralN M (f n))  (SUP n. integralN M (f n))"
    using f by (intro LIMSEQ_SUP[of "λn. integralN M (f n)"] incseq_nn_integral)
  also have "(SUP n. integralN M (f n)) = integralN M (λx. SUP n. f n x)"
    using f by (intro nn_integral_monotone_convergence_SUP[symmetric])
  also have "integralN M (λx. SUP n. f n x) = integralN M (λx. u x)"
    using f by (subst LIMSEQ_SUP[THEN LIMSEQ_unique, OF _ u]) (auto simp: incseq_def le_fun_def)
  finally show ?thesis .
qed

theorem nn_integral_dominated_convergence:
  assumes [measurable]:
       "i. u i  borel_measurable M" "u'  borel_measurable M" "w  borel_measurable M"
    and bound: "j. AE x in M. u j x  w x"
    and w: "(+x. w x M) < "
    and u': "AE x in M. (λi. u i x)  u' x"
  shows "(λi. (+x. u i x M))  (+x. u' x M)"
proof -
  have "limsup (λn. integralN M (u n))  (+ x. limsup (λn. u n x) M)"
    by (intro nn_integral_limsup[OF _ _ bound w]) auto
  moreover have "(+ x. limsup (λn. u n x) M) = (+ x. u' x M)"
    using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot)
  moreover have "(+ x. liminf (λn. u n x) M) = (+ x. u' x M)"
    using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot)
  moreover have "(+x. liminf (λn. u n x) M)  liminf (λn. integralN M (u n))"
    by (intro nn_integral_liminf) auto
  moreover have "liminf (λn. integralN M (u n))  limsup (λn. integralN M (u n))"
    by (intro Liminf_le_Limsup sequentially_bot)
  ultimately show ?thesis
    by (intro Liminf_eq_Limsup) auto
qed

lemma inf_continuous_nn_integral[order_continuous_intros]:
  assumes f: "y. inf_continuous (f y)"
  assumes [measurable]: "x. (λy. f y x)  borel_measurable M"
  assumes bnd: "x. (+ y. f y x M)  "
  shows "inf_continuous (λx. (+y. f y x M))"
  unfolding inf_continuous_def
proof safe
  fix C :: "nat  'b" assume C: "decseq C"
  then show "(+ y. f y (Inf (C ` UNIV)) M) = (INF i. + y. f y (C i) M)"
    using inf_continuous_mono[OF f] bnd
    by (auto simp add: inf_continuousD[OF f C] fun_eq_iff monotone_def le_fun_def less_top
             intro!: nn_integral_monotone_convergence_INF_decseq)
qed

lemma nn_integral_null_set:
  assumes "N  null_sets M" shows "(+ x. u x * indicator N x M) = 0"
proof -
  have "(+ x. u x * indicator N x M) = (+ x. 0 M)"
  proof (intro nn_integral_cong_AE AE_I)
    show "{x  space M. u x * indicator N x  0}  N"
      by (auto simp: indicator_def)
    show "(emeasure M) N = 0" "N  sets M"
      using assms by auto
  qed
  then show ?thesis by simp
qed

lemma nn_integral_0_iff:
  assumes u [measurable]: "u  borel_measurable M"
  shows "integralN M u = 0  emeasure M {xspace M. u x  0} = 0"
    (is "_  (emeasure M) ?A = 0")
proof -
  have u_eq: "(+ x. u x * indicator ?A x M) = integralN M u"
    by (auto intro!: nn_integral_cong simp: indicator_def)
  show ?thesis
  proof
    assume "(emeasure M) ?A = 0"
    with nn_integral_null_set[of ?A M u] u
    show "integralN M u = 0" by (simp add: u_eq null_sets_def)
  next
    assume *: "integralN M u = 0"
    let ?M = "λn. {x  space M. 1  real (n::nat) * u x}"
    have "0 = (SUP n. (emeasure M) (?M n  ?A))"
    proof -
      { fix n :: nat
        have "emeasure M {x  ?A. 1  of_nat n * u x} 
                of_nat n * + x. u x * indicator ?A x M"
          by (intro nn_integral_Markov_inequality) auto
        also have "{x  ?A. 1  of_nat n * u x} = (?M n  ?A)"
          by (auto simp: ennreal_of_nat_eq_real_of_nat u_eq * )
        finally have "emeasure M (?M n  ?A)  0"
          by (simp add: ennreal_of_nat_eq_real_of_nat u_eq * )
        moreover have "0  (emeasure M) (?M n  ?A)" using u by auto
        ultimately have "(emeasure M) (?M n  ?A) = 0" by auto }
      thus ?thesis by simp
    qed
    also have " = (emeasure M) (n. ?M n  ?A)"
    proof (safe intro!: SUP_emeasure_incseq)
      fix n show "?M n  ?A  sets M"
        using u by (auto intro!: sets.Int)
    next
      show "incseq (λn. {x  space M. 1  real n * u x}  {x  space M. u x  0})"
      proof (safe intro!: incseq_SucI)
        fix n :: nat and x
        assume *: "1  real n * u x"
        also have "real n * u x  real (Suc n) * u x"
          by (auto intro!: mult_right_mono)
        finally show "1  real (Suc n) * u x" by auto
      qed
    qed
    also have " = (emeasure M) {xspace M. 0 < u x}"
    proof (safe intro!: arg_cong[where f="(emeasure M)"])
      fix x assume "0 < u x" and [simp, intro]: "x  space M"
      show "x  (n. ?M n  ?A)"
      proof (cases "u x" rule: ennreal_cases)
        case (real r) with 0 < u x have "0 < r" by auto
        obtain j :: nat where "1 / r  real j" using real_arch_simple ..
        hence "1 / r * r  real j * r" unfolding mult_le_cancel_right using 0 < r by auto
        hence "1  real j * r" using real 0 < r by auto
        thus ?thesis using 0 < r real
          by (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_1[symmetric] ennreal_mult[symmetric]
                   simp del: ennreal_1)
      qed (insert 0 < u x, 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  borel_measurable M"
  shows "integralN M u = 0  (AE x in M. u x = 0)"
proof -
  have sets: "{xspace M. u x  0}  sets M"
    using u by auto
  show "integralN M u = 0  (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:
  "{xspace M. P x}  sets M  (AE x in M. P x)  integralN M (indicator {x. ¬ P x}) = 0"
  by (subst nn_integral_0_iff_AE) (auto simp: indicator_def[abs_def])

lemma nn_integral_less:
  assumes [measurable]: "f  borel_measurable M" "g  borel_measurable M"
  assumes f: "(+x. f x M)  "
  assumes ord: "AE x in M. f x  g x" "¬ (AE x in M. g x  f x)"
  shows "(+x. f x M) < (+x. g x M)"
proof -
  have "0 < (+x. g x - f x M)"
  proof (intro order_le_neq_trans notI)
    assume "0 = (+x. g x - f x M)"
    then have "AE x in M. g x - f x = 0"
      using nn_integral_0_iff_AE[of "λx. g x - f x" M] by simp
    with ord(1) have "AE x in M. g x  f x"
      by eventually_elim (auto simp: ennreal_minus_eq_0)
    with ord show False
      by simp
  qed simp
  also have " = (+x. g x M) - (+x. f x M)"
    using f by (subst nn_integral_diff) (auto simp: ord)
  finally show ?thesis
    using f by (auto dest!: ennreal_minus_pos_iff[rotated] simp: less_top)
qed

lemma nn_integral_subalgebra:
  assumes f: "f  borel_measurable N"
  and N: "sets N  sets M" "space N = space M" "A. A  sets N  emeasure N A = emeasure M A"
  shows "integralN N f = integralN M f"
proof -
  have [simp]: "f :: 'a  ennreal. f  borel_measurable N  f  borel_measurable M"
    using N by (auto simp: measurable_def)
  have [simp]: "P. (AE x in N. P x)  (AE x in M. P x)"
    using N by (auto simp add: eventually_ae_filter null_sets_def subset_eq)
  have [simp]: "A. A  sets N  A  sets M"
    using N by auto
  from f show ?thesis
    apply induct
    apply (simp_all add: nn_integral_add nn_integral_cmult nn_integral_monotone_convergence_SUP N image_comp)
    apply (auto intro!: nn_integral_cong cong: nn_integral_cong simp: N(2)[symmetric])
    done
qed

lemma nn_integral_nat_function:
  fixes f :: "'a  nat"
  assumes "f  measurable M (count_space UNIV)"
  shows "(+x. of_nat (f x) M) = (t. emeasure M {xspace M. t < f x})"
proof -
  define F where "F i = {xspace M. i < f x}" for i
  with assms have [measurable]: "i. F i  sets M"
    by auto

  { fix x assume "x  space M"
    have "(λi. if i < f x then 1 else 0) sums (of_nat (f x)::real)"
      using sums_If_finite[of "λi. i < f x" "λ_. 1::real"] by simp
    then have "(λ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 "i. ennreal (if i < f x then 1 else 0) = indicator (F i) x"
      using x  space M by (simp add: one_ennreal_def F_def)
    ultimately have "of_nat (f x) = (i. indicator (F i) x :: ennreal)"
      by (simp add: sums_iff) }
  then have "(+x. of_nat (f x) M) = (+x. (i. indicator (F i) x) M)"
    by (simp cong: nn_integral_cong)
  also have " = (i. emeasure M (F i))"
    by (simp add: nn_integral_suminf)
  finally show ?thesis
    by (simp add: F_def)
qed

theorem nn_integral_lfp:
  assumes sets[simp]: "s. sets (M s) = sets N"
  assumes f: "sup_continuous f"
  assumes g: "sup_continuous g"
  assumes meas: "F. F  borel_measurable N  f F  borel_measurable N"
  assumes step: "F s. F  borel_measurable N  integralN (M s) (f F) = g (λs. integralN (M s) F) s"
  shows "(+ω. lfp f ω M s) = lfp g s"
proof (subst lfp_transfer_bounded[where α="λF s. +x. F x M s" and g=g and f=f and P="λf. f  borel_measurable N", symmetric])
  fix C :: "nat  'b  ennreal" assume "incseq C" "i. C i  borel_measurable N"
  then show "(λs. +x. (SUP i. C i) x M s) = (SUP i. (λs. +x. C i x 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)
qed (auto simp add: step le_fun_def SUP_apply[abs_def] bot_fun_def bot_ennreal intro!: meas f g)

theorem nn_integral_gfp:
  assumes sets[simp]: "s. sets (M s) = sets N"
  assumes f: "inf_continuous f" and g: "inf_continuous g"
  assumes meas: "F. F  borel_measurable N  f F  borel_measurable N"
  assumes bound: "F s. F  borel_measurable N  (+x. f F x M s) < "
  assumes non_zero: "s. emeasure (M s) (space (M s))  0"
  assumes step: "F s. F  borel_measurable N  integralN (M s) (f F) = g (λs. integralN (M s) F) s"
  shows "(+ω. gfp f ω M s) = gfp g s"
proof (subst gfp_transfer_bounded[where α="λF s. +x. F x M s" and g=g and f=f
    and P="λF. F  borel_measurable N  (s. (+x. F x M s) < )", symmetric])
  fix C :: "nat  'b  ennreal" assume "decseq C" "i. C i  borel_measurable N  (s. integralN (M s) (C i) < )"
  then show "(λs. +x. (INF i. C i) x M s) = (INF i. (λs. +x. C i x M s))"
    unfolding INF_apply[abs_def]
    by (subst nn_integral_monotone_convergence_INF_decseq)
       (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets)
next
  show "x. g x  (λs. integralN (M s) (f top))"
    by (subst step)
       (auto simp add: top_fun_def less_le non_zero le_fun_def ennreal_top_mult
             cong del: if_weak_cong intro!: monoD[OF inf_continuous_mono[OF g], THEN le_funD])
next
  fix C assume "i::nat. C i  borel_measurable N  (s. integralN (M s) (C i) < )" "decseq C"
  with bound show "Inf (C ` UNIV)  borel_measurable N  (s. integralN (M s) (Inf (C ` UNIV)) < )"
    unfolding INF_apply[abs_def]
    by (subst nn_integral_monotone_convergence_INF_decseq)
       (auto simp: INF_less_iff cong: measurable_cong_sets intro!: borel_measurable_INF)
next
  show "x. x  borel_measurable N  (s. integralN (M s) x < ) 
         (λs. integralN (M s) (f x)) = g (λs. integralN (M s) x)"
    by (subst step) auto
qed (insert bound, auto simp add: le_fun_def INF_apply[abs_def] top_fun_def intro!: meas f g)


text ‹Cauchy--Schwarz inequality for constnn_integral

lemma sum_of_squares_ge_ennreal:
  fixes a b :: ennreal
  shows "2 * a * b  a2 + b2"
proof (cases a; cases b)
  fix x y
  assume xy: "x  0" "y  0" and [simp]: "a = ennreal x" "b = ennreal y"
  have "0  (x - y)2"
    by simp
  also have " = x2 + y2 - 2 * x * y"
    by (simp add: algebra_simps power2_eq_square)
  finally have "2 * x * y  x2 + y2"
    by simp
  hence "ennreal (2 * x * y)  ennreal (x2 + y2)"
    by (intro ennreal_leI)
  thus ?thesis using xy
    by (simp add: ennreal_mult ennreal_power)
qed auto

lemma Cauchy_Schwarz_nn_integral:
  assumes [measurable]: "f  borel_measurable M" "g  borel_measurable M"
  shows "(+x. f x * g x M)2  (+x. f x ^ 2 M) * (+x. g x ^ 2 M)"
proof (cases "(+x. f x * g x M) = 0")
  case False
  define F where "F = nn_integral M (λx. f x ^ 2)"
  define G where "G = nn_integral M (λx. g x ^ 2)"
  from False have "¬(AE x in M. f x = 0  g x = 0)"
    by (auto simp: nn_integral_0_iff_AE)
  hence "¬(AE x in M. f x = 0)" and "¬(AE x in M. g x = 0)"
    by (auto intro: AE_disjI1 AE_disjI2)
  hence nz: "F  0" "G  0"
    by (auto simp: nn_integral_0_iff_AE F_def G_def)

  show ?thesis
  proof (cases "F =   G = ")
    case True
    thus ?thesis using nz
      by (auto simp: F_def G_def)
  next
    case False
    define F' where "F' = ennreal (sqrt (enn2real F))"
    define G' where "G' = ennreal (sqrt (enn2real G))"
    from False have fin: "F < top" "G < top"
      by (simp_all add: top.not_eq_extremum)
    have F'_sqr: "F'2 = F"
      using False by (cases F) (auto simp: F'_def ennreal_power)
    have G'_sqr: "G'2 = G"
      using False by (cases G) (auto simp: G'_def ennreal_power)
    have nz': "F'  0" "G'  0" and fin': "F'  " "G'  "
      using F'_sqr G'_sqr nz fin by auto
    from fin' have fin'': "F' < top" "G' < top"
      by (auto simp: top.not_eq_extremum)

    have "2 * (F' / F') * (G' / G') * (+x. f x * g x M) =
          F' * G' * (+x. 2 * (f x / F') * (g x / G') M)"
      using nz' fin''
      by (simp add: divide_ennreal_def algebra_simps ennreal_inverse_mult flip: nn_integral_cmult)
    also have "F'/ F' = 1"
      using nz' fin'' by simp
    also have "G'/ G' = 1"
      using nz' fin'' by simp
    also have "2 * 1 * 1 = (2 :: ennreal)" by simp
    also have "F' * G' * (+ x. 2 * (f x / F') * (g x / G') M) 
               F' * G' * (+x. (f x / F')2 + (g x / G')2 M)"
      by (intro mult_left_mono nn_integral_mono sum_of_squares_ge_ennreal) auto
    also have " = F' * G' * (F / F'2 + G / G'2)" using nz
      by (auto simp: nn_integral_add algebra_simps nn_integral_divide F_def G_def)
    also have "F / F'2 = 1"
      using nz F'_sqr fin by simp
    also have "G / G'2 = 1"
      using nz G'_sqr fin by simp
    also have "F' * G' * (1 + 1) = 2 * (F' * G')"
      by (simp add: mult_ac)
    finally have "(+x. f x * g x M)  F' * G'"
      by (subst (asm) ennreal_mult_le_mult_iff) auto
    hence "(+x. f x * g x M)2  (F' * G')2"
      by (intro power_mono_ennreal)
    also have " = F * G"
      by (simp add: algebra_simps F'_sqr G'_sqr)
    finally show ?thesis
      by (simp add: F_def G_def)
  qed
qed auto


(* TODO: rename? *)
subsection ‹Integral under concrete measures›

lemma nn_integral_mono_measure:
  assumes "sets M = sets N" "M  N" shows "nn_integral M f  nn_integral N f"
  unfolding nn_integral_def
proof (intro SUP_subset_mono)
  note sets M = sets N[simp]  sets M = sets N[THEN sets_eq_imp_space_eq, simp]
  show "{g. simple_function M g  g  f}  {g. simple_function N g  g  f}"
    by (simp add: simple_function_def)
  show "integralS M x  integralS N x" for x
    using le_measureD3[OF M  N]
    by (auto simp add: simple_integral_def intro!: sum_mono mult_mono)
qed

lemma nn_integral_empty:
  assumes "space M = {}"
  shows "nn_integral M f = 0"
proof -
  have "(+ x. f x M) = (+ x. 0 M)"
    by(rule nn_integral_cong)(simp add: assms)
  thus ?thesis by simp
qed

lemma nn_integral_bot[simp]: "nn_integral bot f = 0"
  by (simp add: nn_integral_empty)

subsubsectiontag unimportant› ‹Distributions›

lemma nn_integral_distr:
  assumes T: "T  measurable M M'" and f: "f  borel_measurable (distr M M' T)"
  shows "integralN (distr M M' T) f = (+ x. f (T x) M)"
  using f
proof induct
  case (cong f g)
  with T show ?case
    apply (subst nn_integral_cong[of _ f g])
    apply simp
    apply (subst nn_integral_cong[of _ "λx. f (T x)" "λx. g (T x)"])
    apply (simp add: measurable_def Pi_iff)
    apply simp
    done
next
  case (set A)
  then have eq: "x. x  space M  indicator A (T x) = indicator (T -` A  space M) x"
    by (auto simp: indicator_def)
  from set T show ?case
    by (subst nn_integral_cong[OF eq])
       (auto simp add: emeasure_distr intro!: nn_integral_indicator[symmetric] measurable_sets)
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 image_comp)

subsubsectiontag unimportant› ‹Counting space›

lemma simple_function_count_space[simp]:
  "simple_function (count_space A) f  finite (f ` A)"
  unfolding simple_function_def by simp

lemma nn_integral_count_space:
  assumes A: "finite {aA. 0 < f a}"
  shows "integralN (count_space A) f = (a|aA  0 < f a. f a)"
proof -
  have *: "(+x. max 0 (f x) count_space A) =
    (+ x. (a|aA  0 < f a. f a * indicator {a} x) count_space A)"
    by (auto intro!: nn_integral_cong
             simp add: indicator_def of_bool_def if_distrib sum.If_cases[OF A] max_def le_less)
  also have " = (a|aA  0 < f a. + x. f a * indicator {a} x count_space A)"
    by (subst nn_integral_sum) (simp_all add: AE_count_space  less_imp_le)
  also have " = (a|aA  0 < f a. f a)"
    by (auto intro!: sum.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  (+x. f x count_space A) = (aA. f a)"
  by (auto intro!: sum.mono_neutral_left simp: nn_integral_count_space less_le)

lemma nn_integral_count_space':
  assumes "finite A" "x. x  B  x  A  f x = 0" "A  B"
  shows "(+x. f x count_space B) = (xA. f x)"
proof -
  have "(+x. f x count_space B) = (a | a  B  0 < f a. f a)"
    using assms(2,3)
    by (intro nn_integral_count_space finite_subset[OF _ finite A]) (auto simp: less_le)
  also have " = (aA. f a)"
    using assms by (intro sum.mono_neutral_cong_left) (auto simp: less_le)
  finally show ?thesis .
qed

lemma nn_integral_bij_count_space:
  assumes g: "bij_betw g A B"
  shows "(+x. f (g x) count_space A) = (+x. f x count_space B)"
  using g[THEN bij_betw_imp_funcset]
  by (subst distr_bij_count_space[OF g, symmetric])
     (auto intro!: nn_integral_distr[symmetric])

lemma nn_integral_indicator_finite:
  fixes f :: "'a  ennreal"
  assumes f: "finite A" and [measurable]: "a. a  A  {a}  sets M"
  shows "(+x. f x * indicator A x M) = (xA. f x * emeasure M {x})"
proof -
  from f have "(+x. f x * indicator A x M) = (+x. (aA. f a * indicator {a} x) M)"
    by (intro nn_integral_cong) (auto simp: indicator_def if_distrib[where f="λa. x * a" for x] sum.If_cases)
  also have " = (aA. f a * emeasure M {a})"
    by (subst nn_integral_sum) auto
  finally show ?thesis .
qed

lemma nn_integral_count_space_nat:
  fixes f :: "nat  ennreal"
  shows "(+i. f i count_space UNIV) = (i. f i)"
proof -
  have "(+i. f i count_space UNIV) =
    (+i. (j. f j * indicator {j} i) count_space UNIV)"
  proof (intro nn_integral_cong)
    fix i
    have "f i = (j{i}. f j * indicator {j} i)"
      by simp
    also have " = (j. f j * indicator {j} i)"
      by (rule suminf_finite[symmetric]) auto
    finally show "f i = (j. f j * indicator {j} i)" .
  qed
  also have " = (j. (+i. f j * indicator {j} i count_space UNIV))"
    by (rule nn_integral_suminf) auto
  finally show ?thesis
    by simp
qed

lemma nn_integral_enat_function:
  assumes f: "f  measurable M (count_space UNIV)"
  shows "(+ x. ennreal_of_enat (f x) M) = (t. emeasure M {x  space M. t < f x})"
proof -
  define F where "F i = {xspace M. i < f x}" for i :: nat
  with assms have [measurable]: "i. F i  sets M"
    by auto

  { fix x assume "x  space M"
    have "(λi::nat. if i < f x then 1 else 0) sums ennreal_of_enat (f x)"
      using sums_If_finite[of "λr. r < f x" "λ_. 1 :: ennreal"]
      by (cases "f x") (simp_all add: sums_def of_nat_tendsto_top_ennreal)
    also have "(λi. (if i < f x then 1 else 0)) = (λi. indicator (F i) x)"
      using x  space M by (simp add: one_ennreal_def F_def fun_eq_iff)
    finally have "ennreal_of_enat (f x) = (i. indicator (F i) x)"
      by (simp add: sums_iff) }
  then have "(+x. ennreal_of_enat (f x) M) = (+x. (i. indicator (F i) x) M)"
    by (simp cong: nn_integral_cong)
  also have " = (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  'a  ennreal"
  assumes "countable I" and [measurable]: "i. i  I  f i  borel_measurable M"
  shows "(+x. +i. f i x count_space I M) = (+i. +x. f i x M count_space I)"
proof cases
  assume "finite I" then show ?thesis
    by (simp add: nn_integral_count_space_finite nn_integral_sum)
next
  assume "infinite I"
  then have [simp]: "I  {}"
    by auto
  note * = bij_betw_from_nat_into[OF countable I infinite I]
  have **: "f. (i. 0  f i)  (+i. f i count_space I) = (n. f (from_nat_into I n))"
    by (simp add: nn_integral_bij_count_space[symmetric, OF *] nn_integral_count_space_nat)
  show ?thesis
    by (simp add: ** nn_integral_suminf from_nat_into)
qed

lemma of_bool_Bex_eq_nn_integral:
  assumes unique: "x y. x  X  y  X  P x  P y  x = y"
  shows "of_bool (yX. P y) = (+y. of_bool (P y) count_space X)"
proof cases
  assume "yX. P y"
  then obtain y where "P y" "y  X" by auto
  then show ?thesis
    by (subst nn_integral_count_space'[where A="{y}"]) (auto dest: unique)
qed (auto cong: nn_integral_cong_simp)

lemma emeasure_UN_countable:
  assumes sets[measurable]: "i. i  I  X i  sets M" and I[simp]: "countable I"
  assumes disj: "disjoint_family_on X I"
  shows "emeasure M ((X ` I)) = (+i. emeasure M (X i) count_space I)"
proof -
  have eq: "x. indicator ((X ` I)) x = + i. indicator (X i) x count_space I"
  proof cases
    fix x assume x: "x  (X ` I)"
    then obtain j where j: "x  X j" "j  I"
      by auto
    with disj have "i. i  I  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)
  qed (auto simp: nn_integral_0_iff_AE)

  note sets.countable_UN'[unfolded subset_eq, measurable]
  have "emeasure M ((X ` I)) = (+x. indicator ((X ` I)) x M)"
    by simp
  also have " = (+i. +x. indicator (X i) x M count_space I)"
    by (simp add: eq nn_integral_count_space_nn_integral)
  finally show ?thesis
    by (simp cong: nn_integral_cong_simp)
qed

lemma emeasure_countable_singleton:
  assumes sets: "x. x  X  {x}  sets M" and X: "countable X"
  shows "emeasure M X = (+x. emeasure M {x} count_space X)"
proof -
  have "emeasure M (iX. {i}) = (+x. emeasure M {x} count_space X)"
    using assms by (intro emeasure_UN_countable) (auto simp: disjoint_family_on_def)
  also have "(iX. {i}) = X" by auto
  finally show ?thesis .
qed

lemma measure_eqI_countable:
  assumes [simp]: "sets M = Pow A" "sets N = Pow A" and A: "countable A"
  assumes eq: "a. a  A  emeasure M {a} = emeasure N {a}"
  shows "M = N"
proof (rule measure_eqI)
  fix X assume "X  sets M"
  then have X: "X  A" by auto
  moreover from A X have "countable X" by (auto dest: countable_subset)
  ultimately have
    "emeasure M X = (+a. emeasure M {a} count_space X)"
    "emeasure N X = (+a. emeasure N {a} count_space X)"
    by (auto intro!: emeasure_countable_singleton)
  moreover have "(+a. emeasure M {a} count_space X) = (+a. emeasure N {a} count_space X)"
    using X by (intro nn_integral_cong eq) auto
  ultimately show "emeasure M X = emeasure N X"
    by simp
qed simp

lemma measure_eqI_countable_AE:
  assumes [simp]: "sets M = UNIV" "sets N = UNIV"
  assumes ae: "AE x in M. x  Ω" "AE x in N. x  Ω" and [simp]: "countable Ω"
  assumes eq: "x. x  Ω  emeasure M {x} = emeasure N {x}"
  shows "M = N"
proof (rule measure_eqI)
  fix A
  have "emeasure N A = emeasure N {xΩ. x  A}"
    using ae by (intro emeasure_eq_AE) auto
  also have " = (+x. emeasure N {x} count_space {xΩ. x  A})"
    by (intro emeasure_countable_singleton) auto
  also have " = (+x. emeasure M {x} count_space {xΩ. x  A})"
    by (intro nn_integral_cong eq[symmetric]) auto
  also have " = emeasure M {xΩ. x  A}"
    by (intro emeasure_countable_singleton[symmetric]) auto
  also have " = emeasure M A"
    using ae by (intro emeasure_eq_AE) auto
  finally show "emeasure M A = emeasure N A" ..
qed simp

lemma nn_integral_monotone_convergence_SUP_nat:
  fixes f :: "'a  nat  ennreal"
  assumes chain: "Complete_Partial_Order.chain (≤) (f ` Y)"
  and nonempty: "Y  {}"
  shows "(+ x. (SUP iY. f i x) count_space UNIV) = (SUP iY. (+ x. f i x count_space UNIV))"
  (is "?lhs = ?rhs" is "integralN ?M _ = _")
proof (rule order_class.order.antisym)
  show "?rhs  ?lhs"
    by (auto intro!: SUP_least SUP_upper nn_integral_mono)
next
  have "g. incseq g  range g  (λi. f i x) ` Y  (SUP iY. f i x) = (SUP i. g i)" for x
    by (rule ennreal_Sup_countable_SUP) (simp add: nonempty)
  then obtain g where incseq: "x. incseq (g x)"
    and range: "x. range (g x)  (λi. f i x) ` Y"
    and sup: "x. (SUP iY. f i x) = (SUP i. g x i)" by moura
  from incseq have incseq': "incseq (λi x. g x i)"
    by(blast intro: incseq_SucI le_funI dest: incseq_SucD)

  have "?lhs = + x. (SUP i. g x i) ?M" by(simp add: sup)
  also have " = (SUP i. + x. g x i ?M)" using incseq'
    by(rule nn_integral_monotone_convergence_SUP) simp
  also have "  (SUP iY. + x. f i x ?M)"
  proof(rule SUP_least)
    fix n
    have "x. i. g x n = f i x  i  Y" using range by blast
    then obtain I where I: "x. g x n = f (I x) x" "x. I x  Y" by moura

    have "(+ x. g x n count_space UNIV) = (x. g x n)"
      by(rule nn_integral_count_space_nat)
    also have " = (SUP m. x<m. g x n)"
      by(rule suminf_eq_SUP)
    also have "  (SUP iY. + x. f i x ?M)"
    proof(rule SUP_mono)
      fix m
      show "m'Y. (x<m. g x n)  (+ x. f m' x ?M)"
      proof(cases "m > 0")
        case False
        thus ?thesis using nonempty by auto
      next
        case True
        let ?Y = "I ` {..<m}"
        have "f ` ?Y  f ` Y" using I by auto
        with chain have chain': "Complete_Partial_Order.chain (≤) (f ` ?Y)" by(rule chain_subset)
        hence "Sup (f ` ?Y)  f ` ?Y"
          by(rule ccpo_class.in_chain_finite)(auto simp add: True lessThan_empty_iff)
        then obtain m' where "m' < m" and m': "(SUP i?Y. f i) = f (I m')" by auto
        have "I m'  Y" using I by blast
        have "(x<m. g x n)  (x<m. f (I m') x)"
        proof(rule sum_mono)
          fix x
          assume "x  {..<m}"
          hence "x < m" by simp
          have "g x n = f (I x) x" by(simp add: I)
          also have "  (SUP i?Y. f i) x" unfolding Sup_fun_def image_image
            using x  {..<m} by (rule Sup_upper [OF imageI])
          also have " = f (I m') x" unfolding m' by simp
          finally show "g x n  f (I m') x" .
        qed
        also have "  (SUP m. (x<m. f (I m') x))"
          by(rule SUP_upper) simp
        also have " = (x. f (I m') x)"
          by(rule suminf_eq_SUP[symmetric])
        also have " = (+ x. f (I m') x ?M)"
          by(rule nn_integral_count_space_nat[symmetric])
        finally show ?thesis using I m'  Y by blast
      qed
    qed
    finally show "(+ x. g x n count_space UNIV)  " .
  qed
  finally show "?lhs  ?rhs" .
qed

lemma power_series_tendsto_at_left:
  assumes nonneg: "i. 0  f i" and summable: "z. 0  z  z < 1  summable (λn. f n * z^n)"
  shows "((λz. ennreal (n. f n * z^n))  (n. ennreal (f n))) (at_left (1::real))"
proof (intro tendsto_at_left_sequentially)
  show "0 < (1::real)" by simp
  fix S :: "nat  real" assume S: "n. S n < 1" "n. 0 < S n" "S  1" "incseq S"
  then have S_nonneg: "i. 0  S i" by (auto intro: less_imp_le)

  have "(λi. (+n. f n * S i^n count_space UNIV))  (+n. ennreal (f n) count_space UNIV)"
  proof (rule nn_integral_LIMSEQ)
    show "incseq (λ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 "(λi. ennreal (f n * S i^n))  ennreal (f n * 1^n)"
      by (intro tendsto_intros tendsto_ennrealI S)
    then show "(λi. ennreal (f n * S i^n))  ennreal (f n)"
      by simp
  qed (auto simp: S_nonneg intro!: mult_nonneg_nonneg nonneg)
  also have "(λi. (+n. f n * S i^n count_space UNIV)) = (λi. 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 "(+n. ennreal (f n) count_space UNIV) = (n. ennreal (f n))"
    by (simp add: nn_integral_count_space_nat nonneg)
  finally show "(λn. ennreal (na. f na * S n ^ na))  (n. ennreal (f n))" .
qed

subsubsection ‹Measures with Restricted Space›

lemma simple_function_restrict_space_ennreal:
  fixes f :: "'a  ennreal"
  assumes "Ω  space M  sets M"
  shows "simple_function (restrict_space M Ω) f  simple_function M (λx. f x * indicator Ω x)"
proof -
  { assume "finite (f ` space (restrict_space M Ω))"
    then have "finite (f ` space (restrict_space M Ω)  {0})" by simp
    then have "finite ((λx. f x * indicator Ω x) ` space M)"
      by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
  moreover
  { assume "finite ((λx. f x * indicator Ω x) ` space M)"
    then have "finite (f ` space (restrict_space M Ω))"
      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_ennreal[OF assms]
    by auto
qed

lemma simple_function_restrict_space:
  fixes f :: "'a  'b::real_normed_vector"
  assumes "Ω  space M  sets M"
  shows "simple_function (restrict_space M Ω) f  simple_function M (λx. indicator Ω x *R f x)"
proof -
  { assume "finite (f ` space (restrict_space M Ω))"
    then have "finite (f ` space (restrict_space M Ω)  {0})" by simp
    then have "finite ((λx. indicator Ω x *R f x) ` space M)"
      by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
  moreover
  { assume "finite ((λx. indicator Ω x *R f x) ` space M)"
    then have "finite (f ` space (restrict_space M Ω))"
      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[OF assms]
    by auto
qed

lemma simple_integral_restrict_space:
  assumes Ω: "Ω  space M  sets M" "simple_function (restrict_space M Ω) f"
  shows "simple_integral (restrict_space M Ω) f = simple_integral M (λx. f x * indicator Ω x)"
  using simple_function_restrict_space_ennreal[THEN iffD1, OF Ω, THEN simple_functionD(1)]
  by (auto simp add: space_restrict_space emeasure_restrict_space[OF Ω(1)] le_infI2 simple_integral_def
           split: split_indicator split_indicator_asm
           intro!: sum.mono_neutral_cong_left ennreal_mult_left_cong arg_cong2[where f=emeasure])

lemma nn_integral_restrict_space:
  assumes Ω[simp]: "Ω  space M  sets M"
  shows "nn_integral (restrict_space M Ω) f = nn_integral M (λx. f x * indicator Ω x)"
proof -
  let ?R = "restrict_space M Ω" and ?X = "λM f. {s. simple_function M s  s  f  (x. s x < top)}"
  have "integralS ?R ` ?X ?R f = integralS M ` ?X M (λx. f x * indicator Ω x)"
  proof (safe intro!: image_eqI)
    fix s assume s: "simple_function ?R s" "s  f" "x. s x < top"
    from s show "integralS (restrict_space M Ω) s = integralS M (λx. s x * indicator Ω x)"
      by (intro simple_integral_restrict_space) auto
    from s show "simple_function M (λx. s x * indicator Ω x)"
      by (simp add: simple_function_restrict_space_ennreal)
    from s show "(λx. s x * indicator Ω x)  (λx. f x * indicator Ω x)"
      "x. s x * indicator Ω x < top"
      by (auto split: split_indicator simp: le_fun_def image_subset_iff)
  next
    fix s assume s: "simple_function M s" "s  (λx. f x * indicator Ω x)" "x. s x < top"
    then have "simple_function M (λx. s x * indicator (Ω  space M) x)" (is ?s')
      by (intro simple_function_mult simple_function_indicator) auto
    also have "?s'  simple_function M (λx. s x * indicator Ω x)"
      by (rule simple_function_cong) (auto split: split_indicator)
    finally show sf: "simple_function (restrict_space M Ω) s"
      by (simp add: simple_function_restrict_space_ennreal)

    from s have s_eq: "s = (λx. s x * indicator Ω x)"
      by (auto simp add: fun_eq_iff le_fun_def image_subset_iff
                  split: split_indicator split_indicator_asm
                  intro: antisym)

    show "integralS M s = integralS (restrict_space M Ω) s"
      by (subst s_eq) (rule simple_integral_restrict_space[symmetric, OF Ω sf])
    show "x. s x < top"
      using s by (auto simp: image_subset_iff)
    from s show "s  f"
      by (subst s_eq) (auto simp: image_subset_iff le_fun_def split: split_indicator split_indicator_asm)
  qed
  then show ?thesis
    unfolding nn_integral_def_finite by (simp cong del: SUP_cong_simp)
qed

lemma nn_integral_count_space_indicator:
  assumes "NO_MATCH (UNIV::'a set) (X::'a set)"
  shows "(+x. f x count_space X) = (+x. f x * indicator X x count_space UNIV)"
  by (simp add: nn_integral_restrict_space[symmetric] restrict_count_space)

lemma nn_integral_count_space_eq:
  "(x. x  A - B  f x = 0)  (x. x  B - A  f x = 0) 
    (+x. f x count_space A) = (+x. f x count_space B)"
  by (auto simp: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator)

lemma nn_integral_ge_point:
  assumes "x  A"
  shows "p x  + x. p x count_space A"
proof -
  from assms have "p x  + x. p x count_space {x}"
    by(auto simp add: nn_integral_count_space_finite max_def)
  also have " = + x'. p x' * indicator {x} x' count_space A"
    using assms by(auto simp add: nn_integral_count_space_indicator indicator_def intro!: nn_integral_cong)
  also have "  + x. p x count_space A"
    by(rule nn_integral_mono)(simp add: indicator_def)
  finally show ?thesis .
qed

subsubsection ‹Measure spaces with an associated density›

definitiontag important› density :: "'a measure  ('a  ennreal)  'a measure" where
  "density M f = measure_of (space M) (sets M) (λA. + x. f x * indicator A x M)"

lemma
  shows sets_density[simp, measurable_cong]: "sets (density M f) = sets M"
    and space_density[simp]: "space (density M f) = space M"
  by (auto simp: density_def)

(* FIXME: add conversion to simplify space, sets and measurable *)
lemma space_density_imp[measurable_dest]:
  "x M f. x  space (density M f)  x  space M" by auto

lemma
  shows measurable_density_eq1[simp]: "g  measurable (density Mg f) Mg'  g  measurable Mg Mg'"
    and measurable_density_eq2[simp]: "h  measurable Mh (density Mh' f)  h  measurable Mh Mh'"
    and simple_function_density_eq[simp]: "simple_function (density Mu f) u  simple_function Mu u"
  unfolding measurable_def simple_function_def by simp_all

lemma density_cong: "f  borel_measurable M  f'  borel_measurable M 
  (AE x in M. f x = f' x)  density M f = density M f'"
  unfolding density_def by (auto intro!: measure_of_eq nn_integral_cong_AE sets.space_closed)

lemma emeasure_density:
  assumes f[measurable]: "f  borel_measurable M" and A[measurable]: "A  sets M"
  shows "emeasure (density M f) A = (+ x. f x * indicator A x M)"
    (is "_ =  A")
  unfolding density_def
proof (rule emeasure_measure_of_sigma)
  show "sigma_algebra (space M) (sets M)" ..
  show "positive (sets M) "
    using f by (auto simp: positive_def)
  show "countably_additive (sets M) "
  proof (intro countably_additiveI)
    fix A :: "nat  'a set" assume "range A  sets M"
    then have "i. A i  sets M" by auto
    then have *: "i. (λx. f x * indicator (A i) x)  borel_measurable M"
      by auto
    assume disj: "disjoint_family A"
    then have "(n.  (A n)) = (+ x. (n. f x * indicator (A n) x) M)"
       using f * by (subst nn_integral_suminf) auto
    also have "(+ x. (n. f x * indicator (A n) x) M) = (+ x. f x * (n. indicator (A n) x) M)"
      using f by (auto intro!: ennreal_suminf_cmult nn_integral_cong_AE)
    also have " = (+ x. f x * indicator (n. A n) x M)"
      unfolding suminf_indicator[OF disj] ..
    finally show "(i. + x. f x * indicator (A i) x M) = + x. f x * indicator (i. A i) x M" .
  qed
qed fact

lemma null_sets_density_iff:
  assumes f: "f  borel_measurable M"
  shows "A  null_sets (density M f)  A  sets M  (AE x in M. x  A  f x = 0)"
proof -
  { assume "A  sets M"
    have "(+x. f x * indicator A x M) = 0  emeasure M {x  space M. f x * indicator A x  0} = 0"
      using f A  sets M by (intro nn_integral_0_iff) auto
    also have "  (AE x in M. f x * indicator A x = 0)"
      using f A  sets M by (intro AE_iff_measurable[OF _ refl, symmetric]) auto
    also have "(AE x in M. f x * indicator A x = 0)  (AE x in M. x  A  f x  0)"
      by (auto simp add: indicator_def max_def split: if_split_asm)
    finally have "(+x. f x * indicator A x M) = 0  (AE x in M. x  A  f x  0)" . }
  with f show ?thesis
    by (simp add: null_sets_def emeasure_density cong: conj_cong)
qed

lemma AE_density:
  assumes f: "f  borel_measurable M"
  shows "(AE x in density M f. P x)  (AE x in M. 0 < f x  P x)"
proof
  assume "AE x in density M f. P x"
  with f obtain N where "{x  space M. ¬ P x}  N" "N  sets M" and ae: "AE x in M. x  N  f x = 0"
    by (auto simp: eventually_ae_filter null_sets_density_iff)
  then have "AE x in M. x  N  P x" by auto
  with ae show "AE x in M. 0 < f x  P x"
    by (rule eventually_elim2) auto
next
  fix N assume ae: "AE x in M. 0 < f x  P x"
  then obtain N where "{x  space M. ¬ (0 < f x  P x)}  N" "N  null_sets M"
    by (auto simp: eventually_ae_filter)
  then have *: "{x  space (density M f). ¬ P x}  N  {xspace M. f x = 0}"
    "N  {xspace M. f x = 0}  sets M" and ae2: "AE x in M. x  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  {xspace M. f x = 0}"] conjI *) (auto elim: eventually_elim2)
qed

lemmatag important› nn_integral_density:
  assumes f: "f  borel_measurable M"
  assumes g: "g  borel_measurable M"
  shows "integralN (density M f) g = (+ x. f x * g x M)"
using g proof induct
  case (cong u v)
  then show ?case
    apply (subst nn_integral_cong[OF cong(3)])
    apply (simp_all cong: nn_integral_cong)
    done
next
  case (set A) then show ?case
    by (simp add: emeasure_density f)
next
  case (mult u c)
  moreover have "x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps)
  ultimately show ?case
    using f by (simp add: nn_integral_cmult)
next
  case (add u v)
  then have "x. f x * (v x + u x) = f x * v x + f x * u x"
    by (simp add: distrib_left)
  with add f show ?case
    by (auto simp add: nn_integral_add intro!: nn_integral_add[symmetric])
next
  case (seq U)
  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 image_comp)
    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!: mult_left_mono)
    done
qed

lemma density_distr:
  assumes [measurable]: "f  borel_measurable N" "X  measurable M N"
  shows "density (distr M N X) f = distr (density M (λx. f (X x))) N X"
  by (intro measure_eqI)
     (auto simp add: emeasure_density nn_integral_distr emeasure_distr
           split: split_indicator intro!: nn_integral_cong)

lemma emeasure_restricted:
  assumes S: "S  sets M" and X: "X  sets M"
  shows "emeasure (density M (indicator S)) X = emeasure M (S  X)"
proof -
  have "emeasure (density M (indicator S)) X = (+x. indicator S x * indicator X x M)"
    using S X by (simp add: emeasure_density)
  also have " = (+x. indicator (S  X) x M)"
    by (auto intro!: nn_integral_cong simp: indicator_def)
  also have " = emeasure M (S  X)"
    using S X by (simp add: sets.Int)
  finally show ?thesis .
qed

lemma measure_restricted:
  "S  sets M  X  sets M  measure (density M (indicator S)) X = measure M (S  X)"
  by (simp add: emeasure_restricted measure_def)

lemma (in finite_measure) finite_measure_restricted:
  "S  sets M  finite_measure (density M (indicator S))"
  by standard (simp add: emeasure_restricted)

lemma emeasure_density_const:
  "A  sets M  emeasure (density M (λ_. c)) A = c * emeasure M A"
  by (auto simp: nn_integral_cmult_indicator emeasure_density)

lemma measure_density_const:
  "A  sets M  c    measure (density M (λ_. c)) A = enn2real c * measure M A"
  by (auto simp: emeasure_density_const measure_def enn2real_mult)

lemma density_density_eq:
   "f  borel_measurable M  g  borel_measurable M 
   density (density M f) g = density M (λx. f x * g x)"
  by (auto intro!: measure_eqI simp: emeasure_density nn_integral_density ac_simps)

lemma distr_density_distr:
  assumes T: "T  measurable M M'" and T': "T'  measurable M' M"
    and inv: "xspace M. T' (T x) = x"
  assumes f: "f  borel_measurable M'"
  shows "distr (density (distr M M' T) f) M T' = density M (f  T)" (is "?R = ?L")
proof (rule measure_eqI)
  fix A assume A: "A  sets ?R"
  { fix x assume "x  space M"
    with sets.sets_into_space[OF A]
    have "indicator (T' -` A  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
                  nn_integral_distr measurable_sets cong: nn_integral_cong)
qed simp

lemma density_density_divide:
  fixes f g :: "'a  real"
  assumes f: "f  borel_measurable M" "AE x in M. 0  f x"
  assumes g: "g  borel_measurable M" "AE x in M. 0  g x"
  assumes ac: "AE x in M. f x = 0  g x = 0"
  shows "density (density M f) (λx. g x / f x) = density M g"
proof -
  have "density M g = density M (λ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

lemma density_1: "density M (λ_. 1) = M"
  by (intro measure_eqI) (auto simp: emeasure_density)

lemma emeasure_density_add:
  assumes X: "X  sets M"
  assumes Mf[measurable]: "f  borel_measurable M"
  assumes Mg[measurable]: "g  borel_measurable M"
  shows "emeasure (density M f) X + emeasure (density M g) X =
           emeasure (density M (λx. f x + g x)) X"
  using assms
  apply (subst (1 2 3) emeasure_density, simp_all) []
  apply (subst nn_integral_add[symmetric], simp_all) []
  apply (intro nn_integral_cong, simp split: split_indicator)
  done

subsubsection ‹Point measure›

definitiontag important› point_measure :: "'a set  ('a  ennreal)  'a measure" where
  "point_measure A f = density (count_space A) f"

lemma
  shows space_point_measure: "space (point_measure A f) = A"
    and sets_point_measure: "sets (point_measure A f) = Pow A"
  by (auto simp: point_measure_def)

lemma sets_point_measure_count_space[measurable_cong]: "sets (point_measure A f) = sets (count_space A)"
  by (simp add: sets_point_measure)

lemma measurable_point_measure_eq1[simp]:
  "g  measurable (point_measure A f) M  g  A  space M"
  unfolding point_measure_def by simp

lemma measurable_point_measure_eq2_finite[simp]:
  "finite A 
   g  measurable M (point_measure A f) 
    (g  space M  A  (aA. g -` {a}  space M  sets M))"
  unfolding point_measure_def by (simp add: measurable_count_space_eq2)

lemma simple_function_point_measure[simp]:
  "simple_function (point_measure A f) g  finite (g ` A)"
  by (simp add: point_measure_def)

lemma emeasure_point_measure:
  assumes A: "finite {aX. 0 < f a}" "X  A"
  shows "emeasure (point_measure A f) X = (a|aX  0 < f a. f a)"
proof -
  have "{a. (a  X  a  A  0 < f a)  a  X} = {aX. 0 < f a}"
    using X  A by auto
  with A show ?thesis
    by (simp add: emeasure_density nn_integral_count_space point_measure_def indicator_def of_bool_def)
qed

lemma emeasure_point_measure_finite:
  "finite A  X  A  emeasure (point_measure A f) X = (aX. f a)"
  by (subst emeasure_point_measure) (auto dest: finite_subset intro!: sum.mono_neutral_left simp: less_le)

lemma emeasure_point_measure_finite_if:
  "finite A  emeasure (point_measure A f) X = (if X  A then aX. f a else 0)"
  by (simp add: emeasure_point_measure_finite emeasure_notin_sets sets_point_measure)

lemma measure_point_measure_finite_if:
  assumes "finite A" and "x. x  A  f x  0"
  shows "measure (point_measure A f) X = (if X  A then aX. f a else 0)"
  by (simp add: Sigma_Algebra.measure_def assms emeasure_point_measure_finite_if subset_eq sum_nonneg)

lemma emeasure_point_measure_finite2:
  "X  A  finite X  emeasure (point_measure A f) X = (aX. f a)"
  by (subst emeasure_point_measure)
     (auto dest: finite_subset intro!: sum.mono_neutral_left simp: less_le)

lemma null_sets_point_measure_iff:
  "X  null_sets (point_measure A f)  X  A  (xX. f x = 0)"
 by (auto simp: AE_count_space null_sets_density_iff point_measure_def)

lemma AE_point_measure:
  "(AE x in point_measure A f. P x)  (xA. 0 < f x  P x)"
  unfolding point_measure_def
  by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def)

lemma nn_integral_point_measure:
  "finite {aA. 0 < f a  0 < g a} 
    integralN (point_measure A f) g = (a|aA  0 < f a  0 < g a. f a * g a)"
  unfolding point_measure_def
  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  integralN (point_measure A f) g = (aA. f a * g a)"
  by (subst nn_integral_point_measure) (auto intro!: sum.mono_neutral_left simp: less_le)

subsubsection ‹Uniform measure›

definitiontag important› "uniform_measure M A = density M (λx. indicator A x / emeasure M A)"

lemma
  shows sets_uniform_measure[simp, measurable_cong]: "sets (uniform_measure M A) = sets M"
    and space_uniform_measure[simp]: "space (uniform_measure M A) = space M"
  by (auto simp: uniform_measure_def)

lemma emeasure_uniform_measure[simp]:
  assumes A: "A  sets M" and B: "B  sets M"
  shows "emeasure (uniform_measure M A) B = emeasure M (A  B) / emeasure M A"
proof -
  from A B have "emeasure (uniform_measure M A) B = (+x. (1 / emeasure M A) * indicator (A  B) x M)"
    by (auto simp add: uniform_measure_def emeasure_density divide_ennreal_def split: split_indicator
             intro!: nn_integral_cong)
  also have " = emeasure M (A  B) / emeasure M A"
    using A B
    by (subst nn_integral_cmult_indicator) (simp_all add: sets.Int divide_ennreal_def mult.commute)
  finally show ?thesis .
qed

lemma measure_uniform_measure[simp]:
  assumes A: "emeasure M A  0" "emeasure M A  " and B: "B  sets M"
  shows "measure (uniform_measure M A) B = measure M (A  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  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  sets M  (AE x in M. x  A  P x)  (AE x in uniform_measure M A. P x)"
  unfolding uniform_measure_def by (auto simp: AE_density divide_ennreal_def)

lemma emeasure_uniform_measure_1:
  "emeasure M S  0  emeasure M S    emeasure (uniform_measure M S) S = 1"
  by (subst emeasure_uniform_measure)
     (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  borel_measurable M" and S[measurable]: "S  sets M"
  shows "(+x. f x uniform_measure M S) = (+x. f x * indicator S x M) / emeasure M S"
proof -
  { assume "emeasure M S = "
    then have ?thesis
      by (simp add: uniform_measure_def nn_integral_density f) }
  moreover
  { assume [simp]: "emeasure M S = 0"
    then have ae: "AE x in M. x  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 "(+ x. indicator S x / 0 * f x M) = 0"
      by (subst nn_integral_0_iff_AE) auto
    moreover from ae have "(+ x. f x * indicator S x M) = 0"
      by (subst nn_integral_0_iff_AE) auto
    ultimately have ?thesis
      by (simp add: uniform_measure_def nn_integral_density f) }
  moreover have "emeasure M S  0  emeasure M S    ?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

lemma AE_uniform_measure:
  assumes "emeasure M A  0" "emeasure M A < "
  shows "(AE x in uniform_measure M A. P x)  (AE x in M. x  A  P x)"
proof -
  have "A  sets M"
    using emeasure M A  0 by (metis emeasure_notin_sets)
  moreover have "x. 0 < indicator A x / emeasure M A  x  A"
    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

subsubsectiontag unimportant› ‹Null measure›

lemma null_measure_eq_density: "null_measure M = density M (λ_. 0)"
  by (intro measure_eqI) (simp_all add: emeasure_density)

lemma nn_integral_null_measure[simp]: "(+x. f x null_measure M) = 0"
  by (auto simp add: nn_integral_def simple_integral_def SUP_constant bot_ennreal_def le_fun_def
           intro!: exI[of _ "λx. 0"])

lemma density_null_measure[simp]: "density (null_measure M) f = null_measure M"
proof (intro measure_eqI)
  fix A show "emeasure (density (null_measure M) f) A = emeasure (null_measure M) A"
    by (simp add: density_def) (simp only: null_measure_def[symmetric] emeasure_null_measure)
qed simp

subsubsection ‹Uniform count measure›

definitiontag important› "uniform_count_measure A = point_measure A (λx. 1 / card A)"

lemma
  shows space_uniform_count_measure: "space (uniform_count_measure A) = A"
    and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A"
    unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure)

lemma sets_uniform_count_measure_count_space[measurable_cong]:
  "sets (uniform_count_measure A) = sets (count_space A)"
  by (simp add: sets_uniform_count_measure)

lemma emeasure_uniform_count_measure:
  "finite A  X  A  emeasure (uniform_count_measure A) X = card X / card A"
  by (simp add: emeasure_point_measure_finite uniform_count_measure_def divide_inverse ennreal_mult
                ennreal_of_nat_eq_real_of_nat)

lemma emeasure_uniform_count_measure_if:
  "finite A  emeasure (uniform_count_measure A) X = (if X  A then card X / card A else 0)"
  by (simp add: emeasure_notin_sets emeasure_uniform_count_measure sets_uniform_count_measure)

lemma measure_uniform_count_measure:
  "finite A  X  A  measure (uniform_count_measure A) X = card X / card A"
  by (simp add: emeasure_point_measure_finite uniform_count_measure_def measure_def enn2real_mult)

lemma measure_uniform_count_measure_if:
  "finite A  measure (uniform_count_measure A) X = (if X  A then card X / card A else 0)"
  by (simp add: measure_uniform_count_measure measure_notin_sets sets_uniform_count_measure)

lemma space_uniform_count_measure_empty_iff [simp]:
  "space (uniform_count_measure X) = {}  X = {}"
by(simp add: space_uniform_count_measure)

lemma sets_uniform_count_measure_eq_UNIV [simp]:
  "sets (uniform_count_measure UNIV) = UNIV  True"
  "UNIV = sets (uniform_count_measure UNIV)  True"
  by(simp_all add: sets_uniform_count_measure)

subsubsectiontag unimportant› ‹Scaled measure›

lemma nn_integral_scale_measure:
  assumes f: "f  borel_measurable M"
  shows "nn_integral (scale_measure r M) f = r * nn_integral M f"
  using f
proof induction
  case (cong f g)
  thus ?case
    by(simp add: cong.hyps space_scale_measure cong: nn_integral_cong_simp)
next
  case (mult f c)
  thus ?case
    by(simp add: nn_integral_cmult max_def mult.assoc mult.left_commute)
next
  case (add f g)
  thus ?case
    by(simp add: nn_integral_add distrib_left)
next
  case (seq U)
  thus ?case
    by(simp add: nn_integral_monotone_convergence_SUP SUP_mult_left_ennreal image_comp)
qed simp

end