Theory Measure_Space

theory Measure_Space
imports Measurable Multivariate_Analysis
(*  Title:      HOL/Probability/Measure_Space.thy
    Author:     Lawrence C Paulson
    Author:     Johannes Hölzl, TU München
    Author:     Armin Heller, TU München
*)

section ‹Measure spaces and their properties›

theory Measure_Space
imports
  Measurable "~~/src/HOL/Multivariate_Analysis/Multivariate_Analysis"
begin

subsection "Relate extended reals and the indicator function"

lemma suminf_cmult_indicator:
  fixes f :: "nat ⇒ ennreal"
  assumes "disjoint_family A" "x ∈ A i"
  shows "(∑n. f n * indicator (A n) x) = f i"
proof -
  have **: "⋀n. f n * indicator (A n) x = (if n = i then f n else 0 :: ennreal)"
    using ‹x ∈ A i› assms unfolding disjoint_family_on_def indicator_def by auto
  then have "⋀n. (∑j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ennreal)"
    by (auto simp: setsum.If_cases)
  moreover have "(SUP n. if i < n then f i else 0) = (f i :: ennreal)"
  proof (rule SUP_eqI)
    fix y :: ennreal assume "⋀n. n ∈ UNIV ⟹ (if i < n then f i else 0) ≤ y"
    from this[of "Suc i"] show "f i ≤ y" by auto
  qed (insert assms, simp add: zero_le)
  ultimately show ?thesis using assms
    by (subst suminf_eq_SUP) (auto simp: indicator_def)
qed

lemma suminf_indicator:
  assumes "disjoint_family A"
  shows "(∑n. indicator (A n) x :: ennreal) = indicator (⋃i. A i) x"
proof cases
  assume *: "x ∈ (⋃i. A i)"
  then obtain i where "x ∈ A i" by auto
  from suminf_cmult_indicator[OF assms(1), OF ‹x ∈ A i›, of "λk. 1"]
  show ?thesis using * by simp
qed simp

lemma setsum_indicator_disjoint_family:
  fixes f :: "'d ⇒ 'e::semiring_1"
  assumes d: "disjoint_family_on A P" and "x ∈ A j" and "finite P" and "j ∈ P"
  shows "(∑i∈P. f i * indicator (A i) x) = f j"
proof -
  have "P ∩ {i. x ∈ A i} = {j}"
    using d ‹x ∈ A j› ‹j ∈ P› unfolding disjoint_family_on_def
    by auto
  thus ?thesis
    unfolding indicator_def
    by (simp add: if_distrib setsum.If_cases[OF ‹finite P›])
qed

text ‹
  The type for emeasure spaces is already defined in @{theory Sigma_Algebra}, as it is also used to
  represent sigma algebras (with an arbitrary emeasure).
›

subsection "Extend binary sets"

lemma LIMSEQ_binaryset:
  assumes f: "f {} = 0"
  shows  "(λn. ∑i<n. f (binaryset A B i)) ⇢ f A + f B"
proof -
  have "(λn. ∑i < Suc (Suc n). f (binaryset A B i)) = (λn. f A + f B)"
    proof
      fix n
      show "(∑i < Suc (Suc n). f (binaryset A B i)) = f A + f B"
        by (induct n)  (auto simp add: binaryset_def f)
    qed
  moreover
  have "... ⇢ f A + f B" by (rule tendsto_const)
  ultimately
  have "(λn. ∑i< Suc (Suc n). f (binaryset A B i)) ⇢ f A + f B"
    by metis
  hence "(λn. ∑i< n+2. f (binaryset A B i)) ⇢ f A + f B"
    by simp
  thus ?thesis by (rule LIMSEQ_offset [where k=2])
qed

lemma binaryset_sums:
  assumes f: "f {} = 0"
  shows  "(λn. f (binaryset A B n)) sums (f A + f B)"
    by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan)

lemma suminf_binaryset_eq:
  fixes f :: "'a set ⇒ 'b::{comm_monoid_add, t2_space}"
  shows "f {} = 0 ⟹ (∑n. f (binaryset A B n)) = f A + f B"
  by (metis binaryset_sums sums_unique)

subsection ‹Properties of a premeasure @{term μ}›

text ‹
  The definitions for @{const positive} and @{const countably_additive} should be here, by they are
  necessary to define @{typ "'a measure"} in @{theory Sigma_Algebra}.
›

definition subadditive where
  "subadditive M f ⟷ (∀x∈M. ∀y∈M. x ∩ y = {} ⟶ f (x ∪ y) ≤ f x + f y)"

lemma subadditiveD: "subadditive M f ⟹ x ∩ y = {} ⟹ x ∈ M ⟹ y ∈ M ⟹ f (x ∪ y) ≤ f x + f y"
  by (auto simp add: subadditive_def)

definition countably_subadditive where
  "countably_subadditive M f ⟷
    (∀A. range A ⊆ M ⟶ disjoint_family A ⟶ (⋃i. A i) ∈ M ⟶ (f (⋃i. A i) ≤ (∑i. f (A i))))"

lemma (in ring_of_sets) countably_subadditive_subadditive:
  fixes f :: "'a set ⇒ ennreal"
  assumes f: "positive M f" and cs: "countably_subadditive M f"
  shows  "subadditive M f"
proof (auto simp add: subadditive_def)
  fix x y
  assume x: "x ∈ M" and y: "y ∈ M" and "x ∩ y = {}"
  hence "disjoint_family (binaryset x y)"
    by (auto simp add: disjoint_family_on_def binaryset_def)
  hence "range (binaryset x y) ⊆ M ⟶
         (⋃i. binaryset x y i) ∈ M ⟶
         f (⋃i. binaryset x y i) ≤ (∑ n. f (binaryset x y n))"
    using cs by (auto simp add: countably_subadditive_def)
  hence "{x,y,{}} ⊆ M ⟶ x ∪ y ∈ M ⟶
         f (x ∪ y) ≤ (∑ n. f (binaryset x y n))"
    by (simp add: range_binaryset_eq UN_binaryset_eq)
  thus "f (x ∪ y) ≤  f x + f y" using f x y
    by (auto simp add: Un o_def suminf_binaryset_eq positive_def)
qed

definition additive where
  "additive M μ ⟷ (∀x∈M. ∀y∈M. x ∩ y = {} ⟶ μ (x ∪ y) = μ x + μ y)"

definition increasing where
  "increasing M μ ⟷ (∀x∈M. ∀y∈M. x ⊆ y ⟶ μ x ≤ μ y)"

lemma positiveD1: "positive M f ⟹ f {} = 0" by (auto simp: positive_def)

lemma positiveD_empty:
  "positive M f ⟹ f {} = 0"
  by (auto simp add: positive_def)

lemma additiveD:
  "additive M f ⟹ x ∩ y = {} ⟹ x ∈ M ⟹ y ∈ M ⟹ f (x ∪ y) = f x + f y"
  by (auto simp add: additive_def)

lemma increasingD:
  "increasing M f ⟹ x ⊆ y ⟹ x∈M ⟹ y∈M ⟹ f x ≤ f y"
  by (auto simp add: increasing_def)

lemma countably_additiveI[case_names countably]:
  "(⋀A. range A ⊆ M ⟹ disjoint_family A ⟹ (⋃i. A i) ∈ M ⟹ (∑i. f (A i)) = f (⋃i. A i))
  ⟹ countably_additive M f"
  by (simp add: countably_additive_def)

lemma (in ring_of_sets) disjointed_additive:
  assumes f: "positive M f" "additive M f" and A: "range A ⊆ M" "incseq A"
  shows "(∑i≤n. f (disjointed A i)) = f (A n)"
proof (induct n)
  case (Suc n)
  then have "(∑i≤Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))"
    by simp
  also have "… = f (A n ∪ disjointed A (Suc n))"
    using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_mono)
  also have "A n ∪ disjointed A (Suc n) = A (Suc n)"
    using ‹incseq A› by (auto dest: incseq_SucD simp: disjointed_mono)
  finally show ?case .
qed simp

lemma (in ring_of_sets) additive_sum:
  fixes A:: "'i ⇒ 'a set"
  assumes f: "positive M f" and ad: "additive M f" and "finite S"
      and A: "A`S ⊆ M"
      and disj: "disjoint_family_on A S"
  shows  "(∑i∈S. f (A i)) = f (⋃i∈S. A i)"
  using ‹finite S› disj A
proof induct
  case empty show ?case using f by (simp add: positive_def)
next
  case (insert s S)
  then have "A s ∩ (⋃i∈S. A i) = {}"
    by (auto simp add: disjoint_family_on_def neq_iff)
  moreover
  have "A s ∈ M" using insert by blast
  moreover have "(⋃i∈S. A i) ∈ M"
    using insert ‹finite S› by auto
  ultimately have "f (A s ∪ (⋃i∈S. A i)) = f (A s) + f(⋃i∈S. A i)"
    using ad UNION_in_sets A by (auto simp add: additive_def)
  with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A]
    by (auto simp add: additive_def subset_insertI)
qed

lemma (in ring_of_sets) additive_increasing:
  fixes f :: "'a set ⇒ ennreal"
  assumes posf: "positive M f" and addf: "additive M f"
  shows "increasing M f"
proof (auto simp add: increasing_def)
  fix x y
  assume xy: "x ∈ M" "y ∈ M" "x ⊆ y"
  then have "y - x ∈ M" by auto
  then have "f x + 0 ≤ f x + f (y-x)" by (intro add_left_mono zero_le)
  also have "... = f (x ∪ (y-x))" using addf
    by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
  also have "... = f y"
    by (metis Un_Diff_cancel Un_absorb1 xy(3))
  finally show "f x ≤ f y" by simp
qed

lemma (in ring_of_sets) subadditive:
  fixes f :: "'a set ⇒ ennreal"
  assumes f: "positive M f" "additive M f" and A: "A`S ⊆ M" and S: "finite S"
  shows "f (⋃i∈S. A i) ≤ (∑i∈S. f (A i))"
using S A
proof (induct S)
  case empty thus ?case using f by (auto simp: positive_def)
next
  case (insert x F)
  hence in_M: "A x ∈ M" "(⋃i∈F. A i) ∈ M" "(⋃i∈F. A i) - A x ∈ M" using A by force+
  have subs: "(⋃i∈F. A i) - A x ⊆ (⋃i∈F. A i)" by auto
  have "(⋃i∈(insert x F). A i) = A x ∪ ((⋃i∈F. A i) - A x)" by auto
  hence "f (⋃i∈(insert x F). A i) = f (A x ∪ ((⋃i∈F. A i) - A x))"
    by simp
  also have "… = f (A x) + f ((⋃i∈F. A i) - A x)"
    using f(2) by (rule additiveD) (insert in_M, auto)
  also have "… ≤ f (A x) + f (⋃i∈F. A i)"
    using additive_increasing[OF f] in_M subs by (auto simp: increasing_def intro: add_left_mono)
  also have "… ≤ f (A x) + (∑i∈F. f (A i))" using insert by (auto intro: add_left_mono)
  finally show "f (⋃i∈(insert x F). A i) ≤ (∑i∈(insert x F). f (A i))" using insert by simp
qed

lemma (in ring_of_sets) countably_additive_additive:
  fixes f :: "'a set ⇒ ennreal"
  assumes posf: "positive M f" and ca: "countably_additive M f"
  shows "additive M f"
proof (auto simp add: additive_def)
  fix x y
  assume x: "x ∈ M" and y: "y ∈ M" and "x ∩ y = {}"
  hence "disjoint_family (binaryset x y)"
    by (auto simp add: disjoint_family_on_def binaryset_def)
  hence "range (binaryset x y) ⊆ M ⟶
         (⋃i. binaryset x y i) ∈ M ⟶
         f (⋃i. binaryset x y i) = (∑ n. f (binaryset x y n))"
    using ca
    by (simp add: countably_additive_def)
  hence "{x,y,{}} ⊆ M ⟶ x ∪ y ∈ M ⟶
         f (x ∪ y) = (∑n. f (binaryset x y n))"
    by (simp add: range_binaryset_eq UN_binaryset_eq)
  thus "f (x ∪ y) = f x + f y" using posf x y
    by (auto simp add: Un suminf_binaryset_eq positive_def)
qed

lemma (in algebra) increasing_additive_bound:
  fixes A:: "nat ⇒ 'a set" and  f :: "'a set ⇒ ennreal"
  assumes f: "positive M f" and ad: "additive M f"
      and inc: "increasing M f"
      and A: "range A ⊆ M"
      and disj: "disjoint_family A"
  shows  "(∑i. f (A i)) ≤ f Ω"
proof (safe intro!: suminf_le_const)
  fix N
  note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
  have "(∑i<N. f (A i)) = f (⋃i∈{..<N}. A i)"
    using A by (intro additive_sum [OF f ad _ _]) (auto simp: disj_N)
  also have "... ≤ f Ω" using space_closed A
    by (intro increasingD[OF inc] finite_UN) auto
  finally show "(∑i<N. f (A i)) ≤ f Ω" by simp
qed (insert f A, auto simp: positive_def)

lemma (in ring_of_sets) countably_additiveI_finite:
  fixes μ :: "'a set ⇒ ennreal"
  assumes "finite Ω" "positive M μ" "additive M μ"
  shows "countably_additive M μ"
proof (rule countably_additiveI)
  fix F :: "nat ⇒ 'a set" assume F: "range F ⊆ M" "(⋃i. F i) ∈ M" and disj: "disjoint_family F"

  have "∀i∈{i. F i ≠ {}}. ∃x. x ∈ F i" by auto
  from bchoice[OF this] obtain f where f: "⋀i. F i ≠ {} ⟹ f i ∈ F i" by auto

  have inj_f: "inj_on f {i. F i ≠ {}}"
  proof (rule inj_onI, simp)
    fix i j a b assume *: "f i = f j" "F i ≠ {}" "F j ≠ {}"
    then have "f i ∈ F i" "f j ∈ F j" using f by force+
    with disj * show "i = j" by (auto simp: disjoint_family_on_def)
  qed
  have "finite (⋃i. F i)"
    by (metis F(2) assms(1) infinite_super sets_into_space)

  have F_subset: "{i. μ (F i) ≠ 0} ⊆ {i. F i ≠ {}}"
    by (auto simp: positiveD_empty[OF ‹positive M μ›])
  moreover have fin_not_empty: "finite {i. F i ≠ {}}"
  proof (rule finite_imageD)
    from f have "f`{i. F i ≠ {}} ⊆ (⋃i. F i)" by auto
    then show "finite (f`{i. F i ≠ {}})"
      by (rule finite_subset) fact
  qed fact
  ultimately have fin_not_0: "finite {i. μ (F i) ≠ 0}"
    by (rule finite_subset)

  have disj_not_empty: "disjoint_family_on F {i. F i ≠ {}}"
    using disj by (auto simp: disjoint_family_on_def)

  from fin_not_0 have "(∑i. μ (F i)) = (∑i | μ (F i) ≠ 0. μ (F i))"
    by (rule suminf_finite) auto
  also have "… = (∑i | F i ≠ {}. μ (F i))"
    using fin_not_empty F_subset by (rule setsum.mono_neutral_left) auto
  also have "… = μ (⋃i∈{i. F i ≠ {}}. F i)"
    using ‹positive M μ› ‹additive M μ› fin_not_empty disj_not_empty F by (intro additive_sum) auto
  also have "… = μ (⋃i. F i)"
    by (rule arg_cong[where f=μ]) auto
  finally show "(∑i. μ (F i)) = μ (⋃i. F i)" .
qed

lemma (in ring_of_sets) countably_additive_iff_continuous_from_below:
  fixes f :: "'a set ⇒ ennreal"
  assumes f: "positive M f" "additive M f"
  shows "countably_additive M f ⟷
    (∀A. range A ⊆ M ⟶ incseq A ⟶ (⋃i. A i) ∈ M ⟶ (λi. f (A i)) ⇢ f (⋃i. A i))"
  unfolding countably_additive_def
proof safe
  assume count_sum: "∀A. range A ⊆ M ⟶ disjoint_family A ⟶ UNION UNIV A ∈ M ⟶ (∑i. f (A i)) = f (UNION UNIV A)"
  fix A :: "nat ⇒ 'a set" assume A: "range A ⊆ M" "incseq A" "(⋃i. A i) ∈ M"
  then have dA: "range (disjointed A) ⊆ M" by (auto simp: range_disjointed_sets)
  with count_sum[THEN spec, of "disjointed A"] A(3)
  have f_UN: "(∑i. f (disjointed A i)) = f (⋃i. A i)"
    by (auto simp: UN_disjointed_eq disjoint_family_disjointed)
  moreover have "(λn. (∑i<n. f (disjointed A i))) ⇢ (∑i. f (disjointed A i))"
    using f(1)[unfolded positive_def] dA
    by (auto intro!: summable_LIMSEQ summableI)
  from LIMSEQ_Suc[OF this]
  have "(λn. (∑i≤n. f (disjointed A i))) ⇢ (∑i. f (disjointed A i))"
    unfolding lessThan_Suc_atMost .
  moreover have "⋀n. (∑i≤n. f (disjointed A i)) = f (A n)"
    using disjointed_additive[OF f A(1,2)] .
  ultimately show "(λi. f (A i)) ⇢ f (⋃i. A i)" by simp
next
  assume cont: "∀A. range A ⊆ M ⟶ incseq A ⟶ (⋃i. A i) ∈ M ⟶ (λi. f (A i)) ⇢ f (⋃i. A i)"
  fix A :: "nat ⇒ 'a set" assume A: "range A ⊆ M" "disjoint_family A" "(⋃i. A i) ∈ M"
  have *: "(⋃n. (⋃i<n. A i)) = (⋃i. A i)" by auto
  have "(λn. f (⋃i<n. A i)) ⇢ f (⋃i. A i)"
  proof (unfold *[symmetric], intro cont[rule_format])
    show "range (λi. ⋃i<i. A i) ⊆ M" "(⋃i. ⋃i<i. A i) ∈ M"
      using A * by auto
  qed (force intro!: incseq_SucI)
  moreover have "⋀n. f (⋃i<n. A i) = (∑i<n. f (A i))"
    using A
    by (intro additive_sum[OF f, of _ A, symmetric])
       (auto intro: disjoint_family_on_mono[where B=UNIV])
  ultimately
  have "(λi. f (A i)) sums f (⋃i. A i)"
    unfolding sums_def by simp
  from sums_unique[OF this]
  show "(∑i. f (A i)) = f (⋃i. A i)" by simp
qed

lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous:
  fixes f :: "'a set ⇒ ennreal"
  assumes f: "positive M f" "additive M f"
  shows "(∀A. range A ⊆ M ⟶ decseq A ⟶ (⋂i. A i) ∈ M ⟶ (∀i. f (A i) ≠ ∞) ⟶ (λi. f (A i)) ⇢ f (⋂i. A i))
     ⟷ (∀A. range A ⊆ M ⟶ decseq A ⟶ (⋂i. A i) = {} ⟶ (∀i. f (A i) ≠ ∞) ⟶ (λi. f (A i)) ⇢ 0)"
proof safe
  assume cont: "(∀A. range A ⊆ M ⟶ decseq A ⟶ (⋂i. A i) ∈ M ⟶ (∀i. f (A i) ≠ ∞) ⟶ (λi. f (A i)) ⇢ f (⋂i. A i))"
  fix A :: "nat ⇒ 'a set" assume A: "range A ⊆ M" "decseq A" "(⋂i. A i) = {}" "∀i. f (A i) ≠ ∞"
  with cont[THEN spec, of A] show "(λi. f (A i)) ⇢ 0"
    using ‹positive M f›[unfolded positive_def] by auto
next
  assume cont: "∀A. range A ⊆ M ⟶ decseq A ⟶ (⋂i. A i) = {} ⟶ (∀i. f (A i) ≠ ∞) ⟶ (λi. f (A i)) ⇢ 0"
  fix A :: "nat ⇒ 'a set" assume A: "range A ⊆ M" "decseq A" "(⋂i. A i) ∈ M" "∀i. f (A i) ≠ ∞"

  have f_mono: "⋀a b. a ∈ M ⟹ b ∈ M ⟹ a ⊆ b ⟹ f a ≤ f b"
    using additive_increasing[OF f] unfolding increasing_def by simp

  have decseq_fA: "decseq (λi. f (A i))"
    using A by (auto simp: decseq_def intro!: f_mono)
  have decseq: "decseq (λi. A i - (⋂i. A i))"
    using A by (auto simp: decseq_def)
  then have decseq_f: "decseq (λi. f (A i - (⋂i. A i)))"
    using A unfolding decseq_def by (auto intro!: f_mono Diff)
  have "f (⋂x. A x) ≤ f (A 0)"
    using A by (auto intro!: f_mono)
  then have f_Int_fin: "f (⋂x. A x) ≠ ∞"
    using A by (auto simp: top_unique)
  { fix i
    have "f (A i - (⋂i. A i)) ≤ f (A i)" using A by (auto intro!: f_mono)
    then have "f (A i - (⋂i. A i)) ≠ ∞"
      using A by (auto simp: top_unique) }
  note f_fin = this
  have "(λi. f (A i - (⋂i. A i))) ⇢ 0"
  proof (intro cont[rule_format, OF _ decseq _ f_fin])
    show "range (λi. A i - (⋂i. A i)) ⊆ M" "(⋂i. A i - (⋂i. A i)) = {}"
      using A by auto
  qed
  from INF_Lim_ereal[OF decseq_f this]
  have "(INF n. f (A n - (⋂i. A i))) = 0" .
  moreover have "(INF n. f (⋂i. A i)) = f (⋂i. A i)"
    by auto
  ultimately have "(INF n. f (A n - (⋂i. A i)) + f (⋂i. A i)) = 0 + f (⋂i. A i)"
    using A(4) f_fin f_Int_fin
    by (subst INF_ennreal_add_const) (auto simp: decseq_f)
  moreover {
    fix n
    have "f (A n - (⋂i. A i)) + f (⋂i. A i) = f ((A n - (⋂i. A i)) ∪ (⋂i. A i))"
      using A by (subst f(2)[THEN additiveD]) auto
    also have "(A n - (⋂i. A i)) ∪ (⋂i. A i) = A n"
      by auto
    finally have "f (A n - (⋂i. A i)) + f (⋂i. A i) = f (A n)" . }
  ultimately have "(INF n. f (A n)) = f (⋂i. A i)"
    by simp
  with LIMSEQ_INF[OF decseq_fA]
  show "(λi. f (A i)) ⇢ f (⋂i. A i)" by simp
qed

lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below:
  fixes f :: "'a set ⇒ ennreal"
  assumes f: "positive M f" "additive M f" "∀A∈M. f A ≠ ∞"
  assumes cont: "∀A. range A ⊆ M ⟶ decseq A ⟶ (⋂i. A i) = {} ⟶ (λi. f (A i)) ⇢ 0"
  assumes A: "range A ⊆ M" "incseq A" "(⋃i. A i) ∈ M"
  shows "(λi. f (A i)) ⇢ f (⋃i. A i)"
proof -
  from A have "(λi. f ((⋃i. A i) - A i)) ⇢ 0"
    by (intro cont[rule_format]) (auto simp: decseq_def incseq_def)
  moreover
  { fix i
    have "f ((⋃i. A i) - A i ∪ A i) = f ((⋃i. A i) - A i) + f (A i)"
      using A by (intro f(2)[THEN additiveD]) auto
    also have "((⋃i. A i) - A i) ∪ A i = (⋃i. A i)"
      by auto
    finally have "f ((⋃i. A i) - A i) = f (⋃i. A i) - f (A i)"
      using f(3)[rule_format, of "A i"] A by (auto simp: ennreal_add_diff_cancel subset_eq) }
  moreover have "∀F i in sequentially. f (A i) ≤ f (⋃i. A i)"
    using increasingD[OF additive_increasing[OF f(1, 2)], of "A _" "⋃i. A i"] A
    by (auto intro!: always_eventually simp: subset_eq)
  ultimately show "(λi. f (A i)) ⇢ f (⋃i. A i)"
    by (auto intro: ennreal_tendsto_const_minus)
qed

lemma (in ring_of_sets) empty_continuous_imp_countably_additive:
  fixes f :: "'a set ⇒ ennreal"
  assumes f: "positive M f" "additive M f" and fin: "∀A∈M. f A ≠ ∞"
  assumes cont: "⋀A. range A ⊆ M ⟹ decseq A ⟹ (⋂i. A i) = {} ⟹ (λi. f (A i)) ⇢ 0"
  shows "countably_additive M f"
  using countably_additive_iff_continuous_from_below[OF f]
  using empty_continuous_imp_continuous_from_below[OF f fin] cont
  by blast

subsection ‹Properties of @{const emeasure}›

lemma emeasure_positive: "positive (sets M) (emeasure M)"
  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)

lemma emeasure_empty[simp, intro]: "emeasure M {} = 0"
  using emeasure_positive[of M] by (simp add: positive_def)

lemma emeasure_single_in_space: "emeasure M {x} ≠ 0 ⟹ x ∈ space M"
  using emeasure_notin_sets[of "{x}" M] by (auto dest: sets.sets_into_space zero_less_iff_neq_zero[THEN iffD2])

lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)"
  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)

lemma suminf_emeasure:
  "range A ⊆ sets M ⟹ disjoint_family A ⟹ (∑i. emeasure M (A i)) = emeasure M (⋃i. A i)"
  using sets.countable_UN[of A UNIV M] emeasure_countably_additive[of M]
  by (simp add: countably_additive_def)

lemma sums_emeasure:
  "disjoint_family F ⟹ (⋀i. F i ∈ sets M) ⟹ (λi. emeasure M (F i)) sums emeasure M (⋃i. F i)"
  unfolding sums_iff by (intro conjI suminf_emeasure) auto

lemma emeasure_additive: "additive (sets M) (emeasure M)"
  by (metis sets.countably_additive_additive emeasure_positive emeasure_countably_additive)

lemma plus_emeasure:
  "a ∈ sets M ⟹ b ∈ sets M ⟹ a ∩ b = {} ⟹ emeasure M a + emeasure M b = emeasure M (a ∪ b)"
  using additiveD[OF emeasure_additive] ..

lemma setsum_emeasure:
  "F`I ⊆ sets M ⟹ disjoint_family_on F I ⟹ finite I ⟹
    (∑i∈I. emeasure M (F i)) = emeasure M (⋃i∈I. F i)"
  by (metis sets.additive_sum emeasure_positive emeasure_additive)

lemma emeasure_mono:
  "a ⊆ b ⟹ b ∈ sets M ⟹ emeasure M a ≤ emeasure M b"
  by (metis zero_le sets.additive_increasing emeasure_additive emeasure_notin_sets emeasure_positive increasingD)

lemma emeasure_space:
  "emeasure M A ≤ emeasure M (space M)"
  by (metis emeasure_mono emeasure_notin_sets sets.sets_into_space sets.top zero_le)

lemma emeasure_Diff:
  assumes finite: "emeasure M B ≠ ∞"
  and [measurable]: "A ∈ sets M" "B ∈ sets M" and "B ⊆ A"
  shows "emeasure M (A - B) = emeasure M A - emeasure M B"
proof -
  have "(A - B) ∪ B = A" using ‹B ⊆ A› by auto
  then have "emeasure M A = emeasure M ((A - B) ∪ B)" by simp
  also have "… = emeasure M (A - B) + emeasure M B"
    by (subst plus_emeasure[symmetric]) auto
  finally show "emeasure M (A - B) = emeasure M A - emeasure M B"
    using finite by simp
qed

lemma emeasure_compl:
  "s ∈ sets M ⟹ emeasure M s ≠ ∞ ⟹ emeasure M (space M - s) = emeasure M (space M) - emeasure M s"
  by (rule emeasure_Diff) (auto dest: sets.sets_into_space)

lemma Lim_emeasure_incseq:
  "range A ⊆ sets M ⟹ incseq A ⟹ (λi. (emeasure M (A i))) ⇢ emeasure M (⋃i. A i)"
  using emeasure_countably_additive
  by (auto simp add: sets.countably_additive_iff_continuous_from_below emeasure_positive
    emeasure_additive)

lemma incseq_emeasure:
  assumes "range B ⊆ sets M" "incseq B"
  shows "incseq (λi. emeasure M (B i))"
  using assms by (auto simp: incseq_def intro!: emeasure_mono)

lemma SUP_emeasure_incseq:
  assumes A: "range A ⊆ sets M" "incseq A"
  shows "(SUP n. emeasure M (A n)) = emeasure M (⋃i. A i)"
  using LIMSEQ_SUP[OF incseq_emeasure, OF A] Lim_emeasure_incseq[OF A]
  by (simp add: LIMSEQ_unique)

lemma decseq_emeasure:
  assumes "range B ⊆ sets M" "decseq B"
  shows "decseq (λi. emeasure M (B i))"
  using assms by (auto simp: decseq_def intro!: emeasure_mono)

lemma INF_emeasure_decseq:
  assumes A: "range A ⊆ sets M" and "decseq A"
  and finite: "⋀i. emeasure M (A i) ≠ ∞"
  shows "(INF n. emeasure M (A n)) = emeasure M (⋂i. A i)"
proof -
  have le_MI: "emeasure M (⋂i. A i) ≤ emeasure M (A 0)"
    using A by (auto intro!: emeasure_mono)
  hence *: "emeasure M (⋂i. A i) ≠ ∞" using finite[of 0] by (auto simp: top_unique)

  have "emeasure M (A 0) - (INF n. emeasure M (A n)) = (SUP n. emeasure M (A 0) - emeasure M (A n))"
    by (simp add: ennreal_INF_const_minus)
  also have "… = (SUP n. emeasure M (A 0 - A n))"
    using A finite ‹decseq A›[unfolded decseq_def] by (subst emeasure_Diff) auto
  also have "… = emeasure M (⋃i. A 0 - A i)"
  proof (rule SUP_emeasure_incseq)
    show "range (λn. A 0 - A n) ⊆ sets M"
      using A by auto
    show "incseq (λn. A 0 - A n)"
      using ‹decseq A› by (auto simp add: incseq_def decseq_def)
  qed
  also have "… = emeasure M (A 0) - emeasure M (⋂i. A i)"
    using A finite * by (simp, subst emeasure_Diff) auto
  finally show ?thesis
    by (rule ennreal_minus_cancel[rotated 3])
       (insert finite A, auto intro: INF_lower emeasure_mono)
qed

lemma emeasure_INT_decseq_subset:
  fixes F :: "nat ⇒ 'a set"
  assumes I: "I ≠ {}" and F: "⋀i j. i ∈ I ⟹ j ∈ I ⟹ i ≤ j ⟹ F j ⊆ F i"
  assumes F_sets[measurable]: "⋀i. i ∈ I ⟹ F i ∈ sets M"
    and fin: "⋀i. i ∈ I ⟹ emeasure M (F i) ≠ ∞"
  shows "emeasure M (⋂i∈I. F i) = (INF i:I. emeasure M (F i))"
proof cases
  assume "finite I"
  have "(⋂i∈I. F i) = F (Max I)"
    using I ‹finite I› by (intro antisym INF_lower INF_greatest F) auto
  moreover have "(INF i:I. emeasure M (F i)) = emeasure M (F (Max I))"
    using I ‹finite I› by (intro antisym INF_lower INF_greatest F emeasure_mono) auto
  ultimately show ?thesis
    by simp
next
  assume "infinite I"
  def L  "λn. LEAST i. i ∈ I ∧ i ≥ n"
  have L: "L n ∈ I ∧ n ≤ L n" for n
    unfolding L_def
  proof (rule LeastI_ex)
    show "∃x. x ∈ I ∧ n ≤ x"
      using ‹infinite I› finite_subset[of I "{..< n}"]
      by (rule_tac ccontr) (auto simp: not_le)
  qed
  have L_eq[simp]: "i ∈ I ⟹ L i = i" for i
    unfolding L_def by (intro Least_equality) auto
  have L_mono: "i ≤ j ⟹ L i ≤ L j" for i j
    using L[of j] unfolding L_def by (intro Least_le) (auto simp: L_def)

  have "emeasure M (⋂i. F (L i)) = (INF i. emeasure M (F (L i)))"
  proof (intro INF_emeasure_decseq[symmetric])
    show "decseq (λi. F (L i))"
      using L by (intro antimonoI F L_mono) auto
  qed (insert L fin, auto)
  also have "… = (INF i:I. emeasure M (F i))"
  proof (intro antisym INF_greatest)
    show "i ∈ I ⟹ (INF i. emeasure M (F (L i))) ≤ emeasure M (F i)" for i
      by (intro INF_lower2[of i]) auto
  qed (insert L, auto intro: INF_lower)
  also have "(⋂i. F (L i)) = (⋂i∈I. F i)"
  proof (intro antisym INF_greatest)
    show "i ∈ I ⟹ (⋂i. F (L i)) ⊆ F i" for i
      by (intro INF_lower2[of i]) auto
  qed (insert L, auto)
  finally show ?thesis .
qed

lemma Lim_emeasure_decseq:
  assumes A: "range A ⊆ sets M" "decseq A" and fin: "⋀i. emeasure M (A i) ≠ ∞"
  shows "(λi. emeasure M (A i)) ⇢ emeasure M (⋂i. A i)"
  using LIMSEQ_INF[OF decseq_emeasure, OF A]
  using INF_emeasure_decseq[OF A fin] by simp

lemma emeasure_lfp'[consumes 1, case_names cont measurable]:
  assumes "P M"
  assumes cont: "sup_continuous F"
  assumes *: "⋀M A. P M ⟹ (⋀N. P N ⟹ Measurable.pred N A) ⟹ Measurable.pred M (F A)"
  shows "emeasure M {x∈space M. lfp F x} = (SUP i. emeasure M {x∈space M. (F ^^ i) (λx. False) x})"
proof -
  have "emeasure M {x∈space M. lfp F x} = emeasure M (⋃i. {x∈space M. (F ^^ i) (λx. False) x})"
    using sup_continuous_lfp[OF cont] by (auto simp add: bot_fun_def intro!: arg_cong2[where f=emeasure])
  moreover { fix i from ‹P M› have "{x∈space M. (F ^^ i) (λx. False) x} ∈ sets M"
    by (induct i arbitrary: M) (auto simp add: pred_def[symmetric] intro: *) }
  moreover have "incseq (λi. {x∈space M. (F ^^ i) (λx. False) x})"
  proof (rule incseq_SucI)
    fix i
    have "(F ^^ i) (λx. False) ≤ (F ^^ (Suc i)) (λx. False)"
    proof (induct i)
      case 0 show ?case by (simp add: le_fun_def)
    next
      case Suc thus ?case using monoD[OF sup_continuous_mono[OF cont] Suc] by auto
    qed
    then show "{x ∈ space M. (F ^^ i) (λx. False) x} ⊆ {x ∈ space M. (F ^^ Suc i) (λx. False) x}"
      by auto
  qed
  ultimately show ?thesis
    by (subst SUP_emeasure_incseq) auto
qed

lemma emeasure_lfp:
  assumes [simp]: "⋀s. sets (M s) = sets N"
  assumes cont: "sup_continuous F" "sup_continuous f"
  assumes meas: "⋀P. Measurable.pred N P ⟹ Measurable.pred N (F P)"
  assumes iter: "⋀P s. Measurable.pred N P ⟹ P ≤ lfp F ⟹ emeasure (M s) {x∈space N. F P x} = f (λs. emeasure (M s) {x∈space N. P x}) s"
  shows "emeasure (M s) {x∈space N. lfp F x} = lfp f s"
proof (subst lfp_transfer_bounded[where α="λF s. emeasure (M s) {x∈space N. F x}" and g=f and f=F and P="Measurable.pred N", symmetric])
  fix C assume "incseq C" "⋀i. Measurable.pred N (C i)"
  then show "(λs. emeasure (M s) {x ∈ space N. (SUP i. C i) x}) = (SUP i. (λs. emeasure (M s) {x ∈ space N. C i x}))"
    unfolding SUP_apply[abs_def]
    by (subst SUP_emeasure_incseq) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure])
qed (auto simp add: iter le_fun_def SUP_apply[abs_def] intro!: meas cont)

lemma emeasure_subadditive_finite:
  "finite I ⟹ A ` I ⊆ sets M ⟹ emeasure M (⋃i∈I. A i) ≤ (∑i∈I. emeasure M (A i))"
  by (rule sets.subadditive[OF emeasure_positive emeasure_additive]) auto

lemma emeasure_subadditive:
  "A ∈ sets M ⟹ B ∈ sets M ⟹ emeasure M (A ∪ B) ≤ emeasure M A + emeasure M B"
  using emeasure_subadditive_finite[of "{True, False}" "λTrue ⇒ A | False ⇒ B" M] by simp

lemma emeasure_subadditive_countably:
  assumes "range f ⊆ sets M"
  shows "emeasure M (⋃i. f i) ≤ (∑i. emeasure M (f i))"
proof -
  have "emeasure M (⋃i. f i) = emeasure M (⋃i. disjointed f i)"
    unfolding UN_disjointed_eq ..
  also have "… = (∑i. emeasure M (disjointed f i))"
    using sets.range_disjointed_sets[OF assms] suminf_emeasure[of "disjointed f"]
    by (simp add:  disjoint_family_disjointed comp_def)
  also have "… ≤ (∑i. emeasure M (f i))"
    using sets.range_disjointed_sets[OF assms] assms
    by (auto intro!: suminf_le emeasure_mono disjointed_subset)
  finally show ?thesis .
qed

lemma emeasure_insert:
  assumes sets: "{x} ∈ sets M" "A ∈ sets M" and "x ∉ A"
  shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A"
proof -
  have "{x} ∩ A = {}" using ‹x ∉ A› by auto
  from plus_emeasure[OF sets this] show ?thesis by simp
qed

lemma emeasure_insert_ne:
  "A ≠ {} ⟹ {x} ∈ sets M ⟹ A ∈ sets M ⟹ x ∉ A ⟹ emeasure M (insert x A) = emeasure M {x} + emeasure M A"
  by (rule emeasure_insert)

lemma emeasure_eq_setsum_singleton:
  assumes "finite S" "⋀x. x ∈ S ⟹ {x} ∈ sets M"
  shows "emeasure M S = (∑x∈S. emeasure M {x})"
  using setsum_emeasure[of "λx. {x}" S M] assms
  by (auto simp: disjoint_family_on_def subset_eq)

lemma setsum_emeasure_cover:
  assumes "finite S" and "A ∈ sets M" and br_in_M: "B ` S ⊆ sets M"
  assumes A: "A ⊆ (⋃i∈S. B i)"
  assumes disj: "disjoint_family_on B S"
  shows "emeasure M A = (∑i∈S. emeasure M (A ∩ (B i)))"
proof -
  have "(∑i∈S. emeasure M (A ∩ (B i))) = emeasure M (⋃i∈S. A ∩ (B i))"
  proof (rule setsum_emeasure)
    show "disjoint_family_on (λi. A ∩ B i) S"
      using ‹disjoint_family_on B S›
      unfolding disjoint_family_on_def by auto
  qed (insert assms, auto)
  also have "(⋃i∈S. A ∩ (B i)) = A"
    using A by auto
  finally show ?thesis by simp
qed

lemma emeasure_eq_0:
  "N ∈ sets M ⟹ emeasure M N = 0 ⟹ K ⊆ N ⟹ emeasure M K = 0"
  by (metis emeasure_mono order_eq_iff zero_le)

lemma emeasure_UN_eq_0:
  assumes "⋀i::nat. emeasure M (N i) = 0" and "range N ⊆ sets M"
  shows "emeasure M (⋃i. N i) = 0"
proof -
  have "emeasure M (⋃i. N i) ≤ 0"
    using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp
  then show ?thesis
    by (auto intro: antisym zero_le)
qed

lemma measure_eqI_finite:
  assumes [simp]: "sets M = Pow A" "sets N = Pow A" and "finite 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
  then have "emeasure M X = (∑a∈X. emeasure M {a})"
    using ‹finite A› by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
  also have "… = (∑a∈X. emeasure N {a})"
    using X eq by (auto intro!: setsum.cong)
  also have "… = emeasure N X"
    using X ‹finite A› by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
  finally show "emeasure M X = emeasure N X" .
qed simp

lemma measure_eqI_generator_eq:
  fixes M N :: "'a measure" and E :: "'a set set" and A :: "nat ⇒ 'a set"
  assumes "Int_stable E" "E ⊆ Pow Ω"
  and eq: "⋀X. X ∈ E ⟹ emeasure M X = emeasure N X"
  and M: "sets M = sigma_sets Ω E"
  and N: "sets N = sigma_sets Ω E"
  and A: "range A ⊆ E" "(⋃i. A i) = Ω" "⋀i. emeasure M (A i) ≠ ∞"
  shows "M = N"
proof -
  let   = "emeasure M" and  = "emeasure N"
  interpret S: sigma_algebra Ω "sigma_sets Ω E" by (rule sigma_algebra_sigma_sets) fact
  have "space M = Ω"
    using sets.top[of M] sets.space_closed[of M] S.top S.space_closed ‹sets M = sigma_sets Ω E›
    by blast

  { fix F D assume "F ∈ E" and "?μ F ≠ ∞"
    then have [intro]: "F ∈ sigma_sets Ω E" by auto
    have "?ν F ≠ ∞" using ‹?μ F ≠ ∞› ‹F ∈ E› eq by simp
    assume "D ∈ sets M"
    with ‹Int_stable E› ‹E ⊆ Pow Ω› have "emeasure M (F ∩ D) = emeasure N (F ∩ D)"
      unfolding M
    proof (induct rule: sigma_sets_induct_disjoint)
      case (basic A)
      then have "F ∩ A ∈ E" using ‹Int_stable E› ‹F ∈ E› by (auto simp: Int_stable_def)
      then show ?case using eq by auto
    next
      case empty then show ?case by simp
    next
      case (compl A)
      then have **: "F ∩ (Ω - A) = F - (F ∩ A)"
        and [intro]: "F ∩ A ∈ sigma_sets Ω E"
        using ‹F ∈ E› S.sets_into_space by (auto simp: M)
      have "?ν (F ∩ A) ≤ ?ν F" by (auto intro!: emeasure_mono simp: M N)
      then have "?ν (F ∩ A) ≠ ∞" using ‹?ν F ≠ ∞› by (auto simp: top_unique)
      have "?μ (F ∩ A) ≤ ?μ F" by (auto intro!: emeasure_mono simp: M N)
      then have "?μ (F ∩ A) ≠ ∞" using ‹?μ F ≠ ∞› by (auto simp: top_unique)
      then have "?μ (F ∩ (Ω - A)) = ?μ F - ?μ (F ∩ A)" unfolding **
        using ‹F ∩ A ∈ sigma_sets Ω E› by (auto intro!: emeasure_Diff simp: M N)
      also have "… = ?ν F - ?ν (F ∩ A)" using eq ‹F ∈ E› compl by simp
      also have "… = ?ν (F ∩ (Ω - A))" unfolding **
        using ‹F ∩ A ∈ sigma_sets Ω E› ‹?ν (F ∩ A) ≠ ∞›
        by (auto intro!: emeasure_Diff[symmetric] simp: M N)
      finally show ?case
        using ‹space M = Ω› by auto
    next
      case (union A)
      then have "?μ (⋃x. F ∩ A x) = ?ν (⋃x. F ∩ A x)"
        by (subst (1 2) suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def subset_eq M N)
      with A show ?case
        by auto
    qed }
  note * = this
  show "M = N"
  proof (rule measure_eqI)
    show "sets M = sets N"
      using M N by simp
    have [simp, intro]: "⋀i. A i ∈ sets M"
      using A(1) by (auto simp: subset_eq M)
    fix F assume "F ∈ sets M"
    let ?D = "disjointed (λi. F ∩ A i)"
    from ‹space M = Ω› have F_eq: "F = (⋃i. ?D i)"
      using ‹F ∈ sets M›[THEN sets.sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq)
    have [simp, intro]: "⋀i. ?D i ∈ sets M"
      using sets.range_disjointed_sets[of "λi. F ∩ A i" M] ‹F ∈ sets M›
      by (auto simp: subset_eq)
    have "disjoint_family ?D"
      by (auto simp: disjoint_family_disjointed)
    moreover
    have "(∑i. emeasure M (?D i)) = (∑i. emeasure N (?D i))"
    proof (intro arg_cong[where f=suminf] ext)
      fix i
      have "A i ∩ ?D i = ?D i"
        by (auto simp: disjointed_def)
      then show "emeasure M (?D i) = emeasure N (?D i)"
        using *[of "A i" "?D i", OF _ A(3)] A(1) by auto
    qed
    ultimately show "emeasure M F = emeasure N F"
      by (simp add: image_subset_iff ‹sets M = sets N›[symmetric] F_eq[symmetric] suminf_emeasure)
  qed
qed

lemma measure_of_of_measure: "measure_of (space M) (sets M) (emeasure M) = M"
proof (intro measure_eqI emeasure_measure_of_sigma)
  show "sigma_algebra (space M) (sets M)" ..
  show "positive (sets M) (emeasure M)"
    by (simp add: positive_def)
  show "countably_additive (sets M) (emeasure M)"
    by (simp add: emeasure_countably_additive)
qed simp_all

subsection ‹‹μ›-null sets›

definition null_sets :: "'a measure ⇒ 'a set set" where
  "null_sets M = {N∈sets M. emeasure M N = 0}"

lemma null_setsD1[dest]: "A ∈ null_sets M ⟹ emeasure M A = 0"
  by (simp add: null_sets_def)

lemma null_setsD2[dest]: "A ∈ null_sets M ⟹ A ∈ sets M"
  unfolding null_sets_def by simp

lemma null_setsI[intro]: "emeasure M A = 0 ⟹ A ∈ sets M ⟹ A ∈ null_sets M"
  unfolding null_sets_def by simp

interpretation null_sets: ring_of_sets "space M" "null_sets M" for M
proof (rule ring_of_setsI)
  show "null_sets M ⊆ Pow (space M)"
    using sets.sets_into_space by auto
  show "{} ∈ null_sets M"
    by auto
  fix A B assume null_sets: "A ∈ null_sets M" "B ∈ null_sets M"
  then have sets: "A ∈ sets M" "B ∈ sets M"
    by auto
  then have *: "emeasure M (A ∪ B) ≤ emeasure M A + emeasure M B"
    "emeasure M (A - B) ≤ emeasure M A"
    by (auto intro!: emeasure_subadditive emeasure_mono)
  then have "emeasure M B = 0" "emeasure M A = 0"
    using null_sets by auto
  with sets * show "A - B ∈ null_sets M" "A ∪ B ∈ null_sets M"
    by (auto intro!: antisym zero_le)
qed

lemma UN_from_nat_into:
  assumes I: "countable I" "I ≠ {}"
  shows "(⋃i∈I. N i) = (⋃i. N (from_nat_into I i))"
proof -
  have "(⋃i∈I. N i) = ⋃(N ` range (from_nat_into I))"
    using I by simp
  also have "… = (⋃i. (N ∘ from_nat_into I) i)"
    by simp
  finally show ?thesis by simp
qed

lemma null_sets_UN':
  assumes "countable I"
  assumes "⋀i. i ∈ I ⟹ N i ∈ null_sets M"
  shows "(⋃i∈I. N i) ∈ null_sets M"
proof cases
  assume "I = {}" then show ?thesis by simp
next
  assume "I ≠ {}"
  show ?thesis
  proof (intro conjI CollectI null_setsI)
    show "(⋃i∈I. N i) ∈ sets M"
      using assms by (intro sets.countable_UN') auto
    have "emeasure M (⋃i∈I. N i) ≤ (∑n. emeasure M (N (from_nat_into I n)))"
      unfolding UN_from_nat_into[OF ‹countable I› ‹I ≠ {}›]
      using assms ‹I ≠ {}› by (intro emeasure_subadditive_countably) (auto intro: from_nat_into)
    also have "(λn. emeasure M (N (from_nat_into I n))) = (λ_. 0)"
      using assms ‹I ≠ {}› by (auto intro: from_nat_into)
    finally show "emeasure M (⋃i∈I. N i) = 0"
      by (intro antisym zero_le) simp
  qed
qed

lemma null_sets_UN[intro]:
  "(⋀i::'i::countable. N i ∈ null_sets M) ⟹ (⋃i. N i) ∈ null_sets M"
  by (rule null_sets_UN') auto

lemma null_set_Int1:
  assumes "B ∈ null_sets M" "A ∈ sets M" shows "A ∩ B ∈ null_sets M"
proof (intro CollectI conjI null_setsI)
  show "emeasure M (A ∩ B) = 0" using assms
    by (intro emeasure_eq_0[of B _ "A ∩ B"]) auto
qed (insert assms, auto)

lemma null_set_Int2:
  assumes "B ∈ null_sets M" "A ∈ sets M" shows "B ∩ A ∈ null_sets M"
  using assms by (subst Int_commute) (rule null_set_Int1)

lemma emeasure_Diff_null_set:
  assumes "B ∈ null_sets M" "A ∈ sets M"
  shows "emeasure M (A - B) = emeasure M A"
proof -
  have *: "A - B = (A - (A ∩ B))" by auto
  have "A ∩ B ∈ null_sets M" using assms by (rule null_set_Int1)
  then show ?thesis
    unfolding * using assms
    by (subst emeasure_Diff) auto
qed

lemma null_set_Diff:
  assumes "B ∈ null_sets M" "A ∈ sets M" shows "B - A ∈ null_sets M"
proof (intro CollectI conjI null_setsI)
  show "emeasure M (B - A) = 0" using assms by (intro emeasure_eq_0[of B _ "B - A"]) auto
qed (insert assms, auto)

lemma emeasure_Un_null_set:
  assumes "A ∈ sets M" "B ∈ null_sets M"
  shows "emeasure M (A ∪ B) = emeasure M A"
proof -
  have *: "A ∪ B = A ∪ (B - A)" by auto
  have "B - A ∈ null_sets M" using assms(2,1) by (rule null_set_Diff)
  then show ?thesis
    unfolding * using assms
    by (subst plus_emeasure[symmetric]) auto
qed

subsection ‹The almost everywhere filter (i.e.\ quantifier)›

definition ae_filter :: "'a measure ⇒ 'a filter" where
  "ae_filter M = (INF N:null_sets M. principal (space M - N))"

abbreviation almost_everywhere :: "'a measure ⇒ ('a ⇒ bool) ⇒ bool" where
  "almost_everywhere M P ≡ eventually P (ae_filter M)"

syntax
  "_almost_everywhere" :: "pttrn ⇒ 'a ⇒ bool ⇒ bool" ("AE _ in _. _" [0,0,10] 10)

translations
  "AE x in M. P"  "CONST almost_everywhere M (λx. P)"

lemma eventually_ae_filter: "eventually P (ae_filter M) ⟷ (∃N∈null_sets M. {x ∈ space M. ¬ P x} ⊆ N)"
  unfolding ae_filter_def by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq)

lemma AE_I':
  "N ∈ null_sets M ⟹ {x∈space M. ¬ P x} ⊆ N ⟹ (AE x in M. P x)"
  unfolding eventually_ae_filter by auto

lemma AE_iff_null:
  assumes "{x∈space M. ¬ P x} ∈ sets M" (is "?P ∈ sets M")
  shows "(AE x in M. P x) ⟷ {x∈space M. ¬ P x} ∈ null_sets M"
proof
  assume "AE x in M. P x" then obtain N where N: "N ∈ sets M" "?P ⊆ N" "emeasure M N = 0"
    unfolding eventually_ae_filter by auto
  have "emeasure M ?P ≤ emeasure M N"
    using assms N(1,2) by (auto intro: emeasure_mono)
  then have "emeasure M ?P = 0"
    unfolding ‹emeasure M N = 0› by auto
  then show "?P ∈ null_sets M" using assms by auto
next
  assume "?P ∈ null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I')
qed

lemma AE_iff_null_sets:
  "N ∈ sets M ⟹ N ∈ null_sets M ⟷ (AE x in M. x ∉ N)"
  using Int_absorb1[OF sets.sets_into_space, of N M]
  by (subst AE_iff_null) (auto simp: Int_def[symmetric])

lemma AE_not_in:
  "N ∈ null_sets M ⟹ AE x in M. x ∉ N"
  by (metis AE_iff_null_sets null_setsD2)

lemma AE_iff_measurable:
  "N ∈ sets M ⟹ {x∈space M. ¬ P x} = N ⟹ (AE x in M. P x) ⟷ emeasure M N = 0"
  using AE_iff_null[of _ P] by auto

lemma AE_E[consumes 1]:
  assumes "AE x in M. P x"
  obtains N where "{x ∈ space M. ¬ P x} ⊆ N" "emeasure M N = 0" "N ∈ sets M"
  using assms unfolding eventually_ae_filter by auto

lemma AE_E2:
  assumes "AE x in M. P x" "{x∈space M. P x} ∈ sets M"
  shows "emeasure M {x∈space M. ¬ P x} = 0" (is "emeasure M ?P = 0")
proof -
  have "{x∈space M. ¬ P x} = space M - {x∈space M. P x}" by auto
  with AE_iff_null[of M P] assms show ?thesis by auto
qed

lemma AE_I:
  assumes "{x ∈ space M. ¬ P x} ⊆ N" "emeasure M N = 0" "N ∈ sets M"
  shows "AE x in M. P x"
  using assms unfolding eventually_ae_filter by auto

lemma AE_mp[elim!]:
  assumes AE_P: "AE x in M. P x" and AE_imp: "AE x in M. P x ⟶ Q x"
  shows "AE x in M. Q x"
proof -
  from AE_P obtain A where P: "{x∈space M. ¬ P x} ⊆ A"
    and A: "A ∈ sets M" "emeasure M A = 0"
    by (auto elim!: AE_E)

  from AE_imp obtain B where imp: "{x∈space M. P x ∧ ¬ Q x} ⊆ B"
    and B: "B ∈ sets M" "emeasure M B = 0"
    by (auto elim!: AE_E)

  show ?thesis
  proof (intro AE_I)
    have "emeasure M (A ∪ B) ≤ 0"
      using emeasure_subadditive[of A M B] A B by auto
    then show "A ∪ B ∈ sets M" "emeasure M (A ∪ B) = 0"
      using A B by auto
    show "{x∈space M. ¬ Q x} ⊆ A ∪ B"
      using P imp by auto
  qed
qed

(* depricated replace by laws about eventually *)
lemma
  shows AE_iffI: "AE x in M. P x ⟹ AE x in M. P x ⟷ Q x ⟹ AE x in M. Q x"
    and AE_disjI1: "AE x in M. P x ⟹ AE x in M. P x ∨ Q x"
    and AE_disjI2: "AE x in M. Q x ⟹ AE x in M. P x ∨ Q x"
    and AE_conjI: "AE x in M. P x ⟹ AE x in M. Q x ⟹ AE x in M. P x ∧ Q x"
    and AE_conj_iff[simp]: "(AE x in M. P x ∧ Q x) ⟷ (AE x in M. P x) ∧ (AE x in M. Q x)"
  by auto

lemma AE_impI:
  "(P ⟹ AE x in M. Q x) ⟹ AE x in M. P ⟶ Q x"
  by (cases P) auto

lemma AE_measure:
  assumes AE: "AE x in M. P x" and sets: "{x∈space M. P x} ∈ sets M" (is "?P ∈ sets M")
  shows "emeasure M {x∈space M. P x} = emeasure M (space M)"
proof -
  from AE_E[OF AE] guess N . note N = this
  with sets have "emeasure M (space M) ≤ emeasure M (?P ∪ N)"
    by (intro emeasure_mono) auto
  also have "… ≤ emeasure M ?P + emeasure M N"
    using sets N by (intro emeasure_subadditive) auto
  also have "… = emeasure M ?P" using N by simp
  finally show "emeasure M ?P = emeasure M (space M)"
    using emeasure_space[of M "?P"] by auto
qed

lemma AE_space: "AE x in M. x ∈ space M"
  by (rule AE_I[where N="{}"]) auto

lemma AE_I2[simp, intro]:
  "(⋀x. x ∈ space M ⟹ P x) ⟹ AE x in M. P x"
  using AE_space by force

lemma AE_Ball_mp:
  "∀x∈space M. P x ⟹ AE x in M. P x ⟶ Q x ⟹ AE x in M. Q x"
  by auto

lemma AE_cong[cong]:
  "(⋀x. x ∈ space M ⟹ P x ⟷ Q x) ⟹ (AE x in M. P x) ⟷ (AE x in M. Q x)"
  by auto

lemma AE_all_countable:
  "(AE x in M. ∀i. P i x) ⟷ (∀i::'i::countable. AE x in M. P i x)"
proof
  assume "∀i. AE x in M. P i x"
  from this[unfolded eventually_ae_filter Bex_def, THEN choice]
  obtain N where N: "⋀i. N i ∈ null_sets M" "⋀i. {x∈space M. ¬ P i x} ⊆ N i" by auto
  have "{x∈space M. ¬ (∀i. P i x)} ⊆ (⋃i. {x∈space M. ¬ P i x})" by auto
  also have "… ⊆ (⋃i. N i)" using N by auto
  finally have "{x∈space M. ¬ (∀i. P i x)} ⊆ (⋃i. N i)" .
  moreover from N have "(⋃i. N i) ∈ null_sets M"
    by (intro null_sets_UN) auto
  ultimately show "AE x in M. ∀i. P i x"
    unfolding eventually_ae_filter by auto
qed auto

lemma AE_ball_countable:
  assumes [intro]: "countable X"
  shows "(AE x in M. ∀y∈X. P x y) ⟷ (∀y∈X. AE x in M. P x y)"
proof
  assume "∀y∈X. AE x in M. P x y"
  from this[unfolded eventually_ae_filter Bex_def, THEN bchoice]
  obtain N where N: "⋀y. y ∈ X ⟹ N y ∈ null_sets M" "⋀y. y ∈ X ⟹ {x∈space M. ¬ P x y} ⊆ N y"
    by auto
  have "{x∈space M. ¬ (∀y∈X. P x y)} ⊆ (⋃y∈X. {x∈space M. ¬ P x y})"
    by auto
  also have "… ⊆ (⋃y∈X. N y)"
    using N by auto
  finally have "{x∈space M. ¬ (∀y∈X. P x y)} ⊆ (⋃y∈X. N y)" .
  moreover from N have "(⋃y∈X. N y) ∈ null_sets M"
    by (intro null_sets_UN') auto
  ultimately show "AE x in M. ∀y∈X. P x y"
    unfolding eventually_ae_filter by auto
qed auto

lemma AE_discrete_difference:
  assumes X: "countable X"
  assumes null: "⋀x. x ∈ X ⟹ emeasure M {x} = 0"
  assumes sets: "⋀x. x ∈ X ⟹ {x} ∈ sets M"
  shows "AE x in M. x ∉ X"
proof -
  have "(⋃x∈X. {x}) ∈ null_sets M"
    using assms by (intro null_sets_UN') auto
  from AE_not_in[OF this] show "AE x in M. x ∉ X"
    by auto
qed

lemma AE_finite_all:
  assumes f: "finite S" shows "(AE x in M. ∀i∈S. P i x) ⟷ (∀i∈S. AE x in M. P i x)"
  using f by induct auto

lemma AE_finite_allI:
  assumes "finite S"
  shows "(⋀s. s ∈ S ⟹ AE x in M. Q s x) ⟹ AE x in M. ∀s∈S. Q s x"
  using AE_finite_all[OF ‹finite S›] by auto

lemma emeasure_mono_AE:
  assumes imp: "AE x in M. x ∈ A ⟶ x ∈ B"
    and B: "B ∈ sets M"
  shows "emeasure M A ≤ emeasure M B"
proof cases
  assume A: "A ∈ sets M"
  from imp obtain N where N: "{x∈space M. ¬ (x ∈ A ⟶ x ∈ B)} ⊆ N" "N ∈ null_sets M"
    by (auto simp: eventually_ae_filter)
  have "emeasure M A = emeasure M (A - N)"
    using N A by (subst emeasure_Diff_null_set) auto
  also have "emeasure M (A - N) ≤ emeasure M (B - N)"
    using N A B sets.sets_into_space by (auto intro!: emeasure_mono)
  also have "emeasure M (B - N) = emeasure M B"
    using N B by (subst emeasure_Diff_null_set) auto
  finally show ?thesis .
qed (simp add: emeasure_notin_sets)

lemma emeasure_eq_AE:
  assumes iff: "AE x in M. x ∈ A ⟷ x ∈ B"
  assumes A: "A ∈ sets M" and B: "B ∈ sets M"
  shows "emeasure M A = emeasure M B"
  using assms by (safe intro!: antisym emeasure_mono_AE) auto

lemma emeasure_Collect_eq_AE:
  "AE x in M. P x ⟷ Q x ⟹ Measurable.pred M Q ⟹ Measurable.pred M P ⟹
   emeasure M {x∈space M. P x} = emeasure M {x∈space M. Q x}"
   by (intro emeasure_eq_AE) auto

lemma emeasure_eq_0_AE: "AE x in M. ¬ P x ⟹ emeasure M {x∈space M. P x} = 0"
  using AE_iff_measurable[OF _ refl, of M "λx. ¬ P x"]
  by (cases "{x∈space M. P x} ∈ sets M") (simp_all add: emeasure_notin_sets)

lemma emeasure_add_AE:
  assumes [measurable]: "A ∈ sets M" "B ∈ sets M" "C ∈ sets M"
  assumes 1: "AE x in M. x ∈ C ⟷ x ∈ A ∨ x ∈ B"
  assumes 2: "AE x in M. ¬ (x ∈ A ∧ x ∈ B)"
  shows "emeasure M C = emeasure M A + emeasure M B"
proof -
  have "emeasure M C = emeasure M (A ∪ B)"
    by (rule emeasure_eq_AE) (insert 1, auto)
  also have "… = emeasure M A + emeasure M (B - A)"
    by (subst plus_emeasure) auto
  also have "emeasure M (B - A) = emeasure M B"
    by (rule emeasure_eq_AE) (insert 2, auto)
  finally show ?thesis .
qed

subsection ‹‹σ›-finite Measures›

locale sigma_finite_measure =
  fixes M :: "'a measure"
  assumes sigma_finite_countable:
    "∃A::'a set set. countable A ∧ A ⊆ sets M ∧ (⋃A) = space M ∧ (∀a∈A. emeasure M a ≠ ∞)"

lemma (in sigma_finite_measure) sigma_finite:
  obtains A :: "nat ⇒ 'a set"
  where "range A ⊆ sets M" "(⋃i. A i) = space M" "⋀i. emeasure M (A i) ≠ ∞"
proof -
  obtain A :: "'a set set" where
    [simp]: "countable A" and
    A: "A ⊆ sets M" "(⋃A) = space M" "⋀a. a ∈ A ⟹ emeasure M a ≠ ∞"
    using sigma_finite_countable by metis
  show thesis
  proof cases
    assume "A = {}" with ‹(⋃A) = space M› show thesis
      by (intro that[of "λ_. {}"]) auto
  next
    assume "A ≠ {}"
    show thesis
    proof
      show "range (from_nat_into A) ⊆ sets M"
        using ‹A ≠ {}› A by auto
      have "(⋃i. from_nat_into A i) = ⋃A"
        using range_from_nat_into[OF ‹A ≠ {}› ‹countable A›] by auto
      with A show "(⋃i. from_nat_into A i) = space M"
        by auto
    qed (intro A from_nat_into ‹A ≠ {}›)
  qed
qed

lemma (in sigma_finite_measure) sigma_finite_disjoint:
  obtains A :: "nat ⇒ 'a set"
  where "range A ⊆ sets M" "(⋃i. A i) = space M" "⋀i. emeasure M (A i) ≠ ∞" "disjoint_family A"
proof -
  obtain A :: "nat ⇒ 'a set" where
    range: "range A ⊆ sets M" and
    space: "(⋃i. A i) = space M" and
    measure: "⋀i. emeasure M (A i) ≠ ∞"
    using sigma_finite by blast
  show thesis
  proof (rule that[of "disjointed A"])
    show "range (disjointed A) ⊆ sets M"
      by (rule sets.range_disjointed_sets[OF range])
    show "(⋃i. disjointed A i) = space M"
      and "disjoint_family (disjointed A)"
      using disjoint_family_disjointed UN_disjointed_eq[of A] space range
      by auto
    show "emeasure M (disjointed A i) ≠ ∞" for i
    proof -
      have "emeasure M (disjointed A i) ≤ emeasure M (A i)"
        using range disjointed_subset[of A i] by (auto intro!: emeasure_mono)
      then show ?thesis using measure[of i] by (auto simp: top_unique)
    qed
  qed
qed

lemma (in sigma_finite_measure) sigma_finite_incseq:
  obtains A :: "nat ⇒ 'a set"
  where "range A ⊆ sets M" "(⋃i. A i) = space M" "⋀i. emeasure M (A i) ≠ ∞" "incseq A"
proof -
  obtain F :: "nat ⇒ 'a set" where
    F: "range F ⊆ sets M" "(⋃i. F i) = space M" "⋀i. emeasure M (F i) ≠ ∞"
    using sigma_finite by blast
  show thesis
  proof (rule that[of "λn. ⋃i≤n. F i"])
    show "range (λn. ⋃i≤n. F i) ⊆ sets M"
      using F by (force simp: incseq_def)
    show "(⋃n. ⋃i≤n. F i) = space M"
    proof -
      from F have "⋀x. x ∈ space M ⟹ ∃i. x ∈ F i" by auto
      with F show ?thesis by fastforce
    qed
    show "emeasure M (⋃i≤n. F i) ≠ ∞" for n
    proof -
      have "emeasure M (⋃i≤n. F i) ≤ (∑i≤n. emeasure M (F i))"
        using F by (auto intro!: emeasure_subadditive_finite)
      also have "… < ∞"
        using F by (auto simp: setsum_Pinfty less_top)
      finally show ?thesis by simp
    qed
    show "incseq (λn. ⋃i≤n. F i)"
      by (force simp: incseq_def)
  qed
qed

subsection ‹Measure space induced by distribution of @{const measurable}-functions›

definition distr :: "'a measure ⇒ 'b measure ⇒ ('a ⇒ 'b) ⇒ 'b measure" where
  "distr M N f = measure_of (space N) (sets N) (λA. emeasure M (f -` A ∩ space M))"

lemma
  shows sets_distr[simp, measurable_cong]: "sets (distr M N f) = sets N"
    and space_distr[simp]: "space (distr M N f) = space N"
  by (auto simp: distr_def)

lemma
  shows measurable_distr_eq1[simp]: "measurable (distr Mf Nf f) Mf' = measurable Nf Mf'"
    and measurable_distr_eq2[simp]: "measurable Mg' (distr Mg Ng g) = measurable Mg' Ng"
  by (auto simp: measurable_def)

lemma distr_cong:
  "M = K ⟹ sets N = sets L ⟹ (⋀x. x ∈ space M ⟹ f x = g x) ⟹ distr M N f = distr K L g"
  using sets_eq_imp_space_eq[of N L] by (simp add: distr_def Int_def cong: rev_conj_cong)

lemma emeasure_distr:
  fixes f :: "'a ⇒ 'b"
  assumes f: "f ∈ measurable M N" and A: "A ∈ sets N"
  shows "emeasure (distr M N f) A = emeasure M (f -` A ∩ space M)" (is "_ = ?μ A")
  unfolding distr_def
proof (rule emeasure_measure_of_sigma)
  show "positive (sets N) ?μ"
    by (auto simp: positive_def)

  show "countably_additive (sets N) ?μ"
  proof (intro countably_additiveI)
    fix A :: "nat ⇒ 'b set" assume "range A ⊆ sets N" "disjoint_family A"
    then have A: "⋀i. A i ∈ sets N" "(⋃i. A i) ∈ sets N" by auto
    then have *: "range (λi. f -` (A i) ∩ space M) ⊆ sets M"
      using f by (auto simp: measurable_def)
    moreover have "(⋃i. f -`  A i ∩ space M) ∈ sets M"
      using * by blast
    moreover have **: "disjoint_family (λi. f -` A i ∩ space M)"
      using ‹disjoint_family A› by (auto simp: disjoint_family_on_def)
    ultimately show "(∑i. ?μ (A i)) = ?μ (⋃i. A i)"
      using suminf_emeasure[OF _ **] A f
      by (auto simp: comp_def vimage_UN)
  qed
  show "sigma_algebra (space N) (sets N)" ..
qed fact

lemma emeasure_Collect_distr:
  assumes X[measurable]: "X ∈ measurable M N" "Measurable.pred N P"
  shows "emeasure (distr M N X) {x∈space N. P x} = emeasure M {x∈space M. P (X x)}"
  by (subst emeasure_distr)
     (auto intro!: arg_cong2[where f=emeasure] X(1)[THEN measurable_space])

lemma emeasure_lfp2[consumes 1, case_names cont f measurable]:
  assumes "P M"
  assumes cont: "sup_continuous F"
  assumes f: "⋀M. P M ⟹ f ∈ measurable M' M"
  assumes *: "⋀M A. P M ⟹ (⋀N. P N ⟹ Measurable.pred N A) ⟹ Measurable.pred M (F A)"
  shows "emeasure M' {x∈space M'. lfp F (f x)} = (SUP i. emeasure M' {x∈space M'. (F ^^ i) (λx. False) (f x)})"
proof (subst (1 2) emeasure_Collect_distr[symmetric, where X=f])
  show "f ∈ measurable M' M"  "f ∈ measurable M' M"
    using f[OF ‹P M›] by auto
  { fix i show "Measurable.pred M ((F ^^ i) (λx. False))"
    using ‹P M› by (induction i arbitrary: M) (auto intro!: *) }
  show "Measurable.pred M (lfp F)"
    using ‹P M› cont * by (rule measurable_lfp_coinduct[of P])

  have "emeasure (distr M' M f) {x ∈ space (distr M' M f). lfp F x} =
    (SUP i. emeasure (distr M' M f) {x ∈ space (distr M' M f). (F ^^ i) (λx. False) x})"
    using ‹P M›
  proof (coinduction arbitrary: M rule: emeasure_lfp')
    case (measurable A N) then have "⋀N. P N ⟹ Measurable.pred (distr M' N f) A"
      by metis
    then have "⋀N. P N ⟹ Measurable.pred N A"
      by simp
    with ‹P N›[THEN *] show ?case
      by auto
  qed fact
  then show "emeasure (distr M' M f) {x ∈ space M. lfp F x} =
    (SUP i. emeasure (distr M' M f) {x ∈ space M. (F ^^ i) (λx. False) x})"
   by simp
qed

lemma distr_id[simp]: "distr N N (λx. x) = N"
  by (rule measure_eqI) (auto simp: emeasure_distr)

lemma measure_distr:
  "f ∈ measurable M N ⟹ S ∈ sets N ⟹ measure (distr M N f) S = measure M (f -` S ∩ space M)"
  by (simp add: emeasure_distr measure_def)

lemma distr_cong_AE:
  assumes 1: "M = K" "sets N = sets L" and
    2: "(AE x in M. f x = g x)" and "f ∈ measurable M N" and "g ∈ measurable K L"
  shows "distr M N f = distr K L g"
proof (rule measure_eqI)
  fix A assume "A ∈ sets (distr M N f)"
  with assms show "emeasure (distr M N f) A = emeasure (distr K L g) A"
    by (auto simp add: emeasure_distr intro!: emeasure_eq_AE measurable_sets)
qed (insert 1, simp)

lemma AE_distrD:
  assumes f: "f ∈ measurable M M'"
    and AE: "AE x in distr M M' f. P x"
  shows "AE x in M. P (f x)"
proof -
  from AE[THEN AE_E] guess N .
  with f show ?thesis
    unfolding eventually_ae_filter
    by (intro bexI[of _ "f -` N ∩ space M"])
       (auto simp: emeasure_distr measurable_def)
qed

lemma AE_distr_iff:
  assumes f[measurable]: "f ∈ measurable M N" and P[measurable]: "{x ∈ space N. P x} ∈ sets N"
  shows "(AE x in distr M N f. P x) ⟷ (AE x in M. P (f x))"
proof (subst (1 2) AE_iff_measurable[OF _ refl])
  have "f -` {x∈space N. ¬ P x} ∩ space M = {x ∈ space M. ¬ P (f x)}"
    using f[THEN measurable_space] by auto
  then show "(emeasure (distr M N f) {x ∈ space (distr M N f). ¬ P x} = 0) =
    (emeasure M {x ∈ space M. ¬ P (f x)} = 0)"
    by (simp add: emeasure_distr)
qed auto

lemma null_sets_distr_iff:
  "f ∈ measurable M N ⟹ A ∈ null_sets (distr M N f) ⟷ f -` A ∩ space M ∈ null_sets M ∧ A ∈ sets N"
  by (auto simp add: null_sets_def emeasure_distr)

lemma distr_distr:
  "g ∈ measurable N L ⟹ f ∈ measurable M N ⟹ distr (distr M N f) L g = distr M L (g ∘ f)"
  by (auto simp add: emeasure_distr measurable_space
           intro!: arg_cong[where f="emeasure M"] measure_eqI)

subsection ‹Real measure values›

lemma ring_of_finite_sets: "ring_of_sets (space M) {A∈sets M. emeasure M A ≠ top}"
proof (rule ring_of_setsI)
  show "a ∈ {A ∈ sets M. emeasure M A ≠ top} ⟹ b ∈ {A ∈ sets M. emeasure M A ≠ top} ⟹
    a ∪ b ∈ {A ∈ sets M. emeasure M A ≠ top}" for a b
    using emeasure_subadditive[of a M b] by (auto simp: top_unique)

  show "a ∈ {A ∈ sets M. emeasure M A ≠ top} ⟹ b ∈ {A ∈ sets M. emeasure M A ≠ top} ⟹
    a - b ∈ {A ∈ sets M. emeasure M A ≠ top}" for a b
    using emeasure_mono[of "a - b" a M] by (auto simp: Diff_subset top_unique)
qed (auto dest: sets.sets_into_space)

lemma measure_nonneg[simp]: "0 ≤ measure M A"
  unfolding measure_def by (auto simp: enn2real_nonneg)

lemma zero_less_measure_iff: "0 < measure M A ⟷ measure M A ≠ 0"
  using measure_nonneg[of M A] by (auto simp add: le_less)

lemma measure_le_0_iff: "measure M X ≤ 0 ⟷ measure M X = 0"
  using measure_nonneg[of M X] by linarith

lemma measure_empty[simp]: "measure M {} = 0"
  unfolding measure_def by (simp add: zero_ennreal.rep_eq)

lemma emeasure_eq_ennreal_measure:
  "emeasure M A ≠ top ⟹ emeasure M A = ennreal (measure M A)"
  by (cases "emeasure M A" rule: ennreal_cases) (auto simp: measure_def)

lemma measure_zero_top: "emeasure M A = top ⟹ measure M A = 0"
  by (simp add: measure_def enn2ereal_top)

lemma measure_eq_emeasure_eq_ennreal: "0 ≤ x ⟹ emeasure M A = ennreal x ⟹ measure M A = x"
  using emeasure_eq_ennreal_measure[of M A]
  by (cases "A ∈ M") (auto simp: measure_notin_sets emeasure_notin_sets)

lemma enn2real_plus:"a < top ⟹ b < top ⟹ enn2real (a + b) = enn2real a + enn2real b"
  by (simp add: enn2real_def plus_ennreal.rep_eq real_of_ereal_add enn2ereal_nonneg less_top
           del: real_of_ereal_enn2ereal)

lemma measure_Union:
  "emeasure M A ≠ ∞ ⟹ emeasure M B ≠ ∞ ⟹ A ∈ sets M ⟹ B ∈ sets M ⟹ A ∩ B = {} ⟹
    measure M (A ∪ B) = measure M A + measure M B"
  by (simp add: measure_def enn2ereal_nonneg plus_emeasure[symmetric] enn2real_plus less_top)

lemma disjoint_family_on_insert:
  "i ∉ I ⟹ disjoint_family_on A (insert i I) ⟷ A i ∩ (⋃i∈I. A i) = {} ∧ disjoint_family_on A I"
  by (fastforce simp: disjoint_family_on_def)

lemma measure_finite_Union:
  "finite S ⟹ A`S ⊆ sets M ⟹ disjoint_family_on A S ⟹ (⋀i. i ∈ S ⟹ emeasure M (A i) ≠ ∞) ⟹
    measure M (⋃i∈S. A i) = (∑i∈S. measure M (A i))"
  by (induction S rule: finite_induct)
     (auto simp: disjoint_family_on_insert measure_Union setsum_emeasure[symmetric] sets.countable_UN'[OF countable_finite])

lemma measure_Diff:
  assumes finite: "emeasure M A ≠ ∞"
  and measurable: "A ∈ sets M" "B ∈ sets M" "B ⊆ A"
  shows "measure M (A - B) = measure M A - measure M B"
proof -
  have "emeasure M (A - B) ≤ emeasure M A" "emeasure M B ≤ emeasure M A"
    using measurable by (auto intro!: emeasure_mono)
  hence "measure M ((A - B) ∪ B) = measure M (A - B) + measure M B"
    using measurable finite by (rule_tac measure_Union) (auto simp: top_unique)
  thus ?thesis using ‹B ⊆ A› by (auto simp: Un_absorb2)
qed

lemma measure_UNION:
  assumes measurable: "range A ⊆ sets M" "disjoint_family A"
  assumes finite: "emeasure M (⋃i. A i) ≠ ∞"
  shows "(λi. measure M (A i)) sums (measure M (⋃i. A i))"
proof -
  have "(λi. emeasure M (A i)) sums (emeasure M (⋃i. A i))"
    unfolding suminf_emeasure[OF measurable, symmetric] by (simp add: summable_sums)
  moreover
  { fix i
    have "emeasure M (A i) ≤ emeasure M (⋃i. A i)"
      using measurable by (auto intro!: emeasure_mono)
    then have "emeasure M (A i) = ennreal ((measure M (A i)))"
      using finite by (intro emeasure_eq_ennreal_measure) (auto simp: top_unique) }
  ultimately show ?thesis using finite
    by (subst (asm) (2) emeasure_eq_ennreal_measure)
       (simp_all add: measure_nonneg)
qed

lemma measure_subadditive:
  assumes measurable: "A ∈ sets M" "B ∈ sets M"
  and fin: "emeasure M A ≠ ∞" "emeasure M B ≠ ∞"
  shows "measure M (A ∪ B) ≤ measure M A + measure M B"
proof -
  have "emeasure M (A ∪ B) ≠ ∞"
    using emeasure_subadditive[OF measurable] fin by (auto simp: top_unique)
  then show "(measure M (A ∪ B)) ≤ (measure M A) + (measure M B)"
    using emeasure_subadditive[OF measurable] fin
    apply simp
    apply (subst (asm) (2 3 4) emeasure_eq_ennreal_measure)
    apply (auto simp add: ennreal_plus[symmetric] simp del: ennreal_plus)
    done
qed

lemma measure_subadditive_finite:
  assumes A: "finite I" "A`I ⊆ sets M" and fin: "⋀i. i ∈ I ⟹ emeasure M (A i) ≠ ∞"
  shows "measure M (⋃i∈I. A i) ≤ (∑i∈I. measure M (A i))"
proof -
  { have "emeasure M (⋃i∈I. A i) ≤ (∑i∈I. emeasure M (A i))"
      using emeasure_subadditive_finite[OF A] .
    also have "… < ∞"
      using fin by (simp add: less_top A)
    finally have "emeasure M (⋃i∈I. A i) ≠ top" by simp }
  note * = this
  show ?thesis
    using emeasure_subadditive_finite[OF A] fin
    unfolding emeasure_eq_ennreal_measure[OF *]
    by (simp_all add: setsum_ennreal measure_nonneg setsum_nonneg emeasure_eq_ennreal_measure)
qed

lemma measure_subadditive_countably:
  assumes A: "range A ⊆ sets M" and fin: "(∑i. emeasure M (A i)) ≠ ∞"
  shows "measure M (⋃i. A i) ≤ (∑i. measure M (A i))"
proof -
  from fin have **: "⋀i. emeasure M (A i) ≠ top"
    using ennreal_suminf_lessD[of "λi. emeasure M (A i)"] by (simp add: less_top)
  { have "emeasure M (⋃i. A i) ≤ (∑i. emeasure M (A i))"
      using emeasure_subadditive_countably[OF A] .
    also have "… < ∞"
      using fin by (simp add: less_top)
    finally have "emeasure M (⋃i. A i) ≠ top" by simp }
  then have "ennreal (measure M (⋃i. A i)) = emeasure M (⋃i. A i)"
    by (rule emeasure_eq_ennreal_measure[symmetric])
  also have "… ≤ (∑i. emeasure M (A i))"
    using emeasure_subadditive_countably[OF A] .
  also have "… = ennreal (∑i. measure M (A i))"
    using fin unfolding emeasure_eq_ennreal_measure[OF **]
    by (subst suminf_ennreal) (auto simp: **)
  finally show ?thesis
    apply (rule ennreal_le_iff[THEN iffD1, rotated])
    apply (intro suminf_nonneg allI measure_nonneg summable_suminf_not_top)
    using fin
    apply (simp add: emeasure_eq_ennreal_measure[OF **])
    done
qed

lemma measure_eq_setsum_singleton:
  "finite S ⟹ (⋀x. x ∈ S ⟹ {x} ∈ sets M) ⟹ (⋀x. x ∈ S ⟹ emeasure M {x} ≠ ∞) ⟹
    measure M S = (∑x∈S. measure M {x})"
  using emeasure_eq_setsum_singleton[of S M]
  by (intro measure_eq_emeasure_eq_ennreal) (auto simp: setsum_nonneg emeasure_eq_ennreal_measure)

lemma Lim_measure_incseq:
  assumes A: "range A ⊆ sets M" "incseq A" and fin: "emeasure M (⋃i. A i) ≠ ∞"
  shows "(λi. measure M (A i)) ⇢ measure M (⋃i. A i)"
proof (rule tendsto_ennrealD)
  have "ennreal (measure M (⋃i. A i)) = emeasure M (⋃i. A i)"
    using fin by (auto simp: emeasure_eq_ennreal_measure)
  moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i
    using assms emeasure_mono[of "A _" "⋃i. A i" M]
    by (intro emeasure_eq_ennreal_measure[symmetric]) (auto simp: less_top UN_upper intro: le_less_trans)
  ultimately show "(λx. ennreal (Sigma_Algebra.measure M (A x))) ⇢ ennreal (Sigma_Algebra.measure M (⋃i. A i))"
    using A by (auto intro!: Lim_emeasure_incseq)
qed auto

lemma Lim_measure_decseq:
  assumes A: "range A ⊆ sets M" "decseq A" and fin: "⋀i. emeasure M (A i) ≠ ∞"
  shows "(λn. measure M (A n)) ⇢ measure M (⋂i. A i)"
proof (rule tendsto_ennrealD)
  have "ennreal (measure M (⋂i. A i)) = emeasure M (⋂i. A i)"
    using fin[of 0] A emeasure_mono[of "⋂i. A i" "A 0" M]
    by (auto intro!: emeasure_eq_ennreal_measure[symmetric] simp: INT_lower less_top intro: le_less_trans)
  moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i
    using A fin[of i] by (intro emeasure_eq_ennreal_measure[symmetric]) auto
  ultimately show "(λx. ennreal (Sigma_Algebra.measure M (A x))) ⇢ ennreal (Sigma_Algebra.measure M (⋂i. A i))"
    using fin A by (auto intro!: Lim_emeasure_decseq)
qed auto

subsection ‹Measure spaces with @{term "emeasure M (space M) < ∞"}›

locale finite_measure = sigma_finite_measure M for M +
  assumes finite_emeasure_space: "emeasure M (space M) ≠ top"

lemma finite_measureI[Pure.intro!]:
  "emeasure M (space M) ≠ ∞ ⟹ finite_measure M"
  proof qed (auto intro!: exI[of _ "{space M}"])

lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A ≠ top"
  using finite_emeasure_space emeasure_space[of M A] by (auto simp: top_unique)

lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ennreal (measure M A)"
  by (intro emeasure_eq_ennreal_measure) simp

lemma (in finite_measure) emeasure_real: "∃r. 0 ≤ r ∧ emeasure M A = ennreal r"
  using emeasure_finite[of A] by (cases "emeasure M A" rule: ennreal_cases) auto

lemma (in finite_measure) bounded_measure: "measure M A ≤ measure M (space M)"
  using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def)

lemma (in finite_measure) finite_measure_Diff:
  assumes sets: "A ∈ sets M" "B ∈ sets M" and "B ⊆ A"
  shows "measure M (A - B) = measure M A - measure M B"
  using measure_Diff[OF _ assms] by simp

lemma (in finite_measure) finite_measure_Union:
  assumes sets: "A ∈ sets M" "B ∈ sets M" and "A ∩ B = {}"
  shows "measure M (A ∪ B) = measure M A + measure M B"
  using measure_Union[OF _ _ assms] by simp

lemma (in finite_measure) finite_measure_finite_Union:
  assumes measurable: "finite S" "A`S ⊆ sets M" "disjoint_family_on A S"
  shows "measure M (⋃i∈S. A i) = (∑i∈S. measure M (A i))"
  using measure_finite_Union[OF assms] by simp

lemma (in finite_measure) finite_measure_UNION:
  assumes A: "range A ⊆ sets M" "disjoint_family A"
  shows "(λi. measure M (A i)) sums (measure M (⋃i. A i))"
  using measure_UNION[OF A] by simp

lemma (in finite_measure) finite_measure_mono:
  assumes "A ⊆ B" "B ∈ sets M" shows "measure M A ≤ measure M B"
  using emeasure_mono[OF assms] emeasure_real[of A] emeasure_real[of B] by (auto simp: measure_def)

lemma (in finite_measure) finite_measure_subadditive:
  assumes m: "A ∈ sets M" "B ∈ sets M"
  shows "measure M (A ∪ B) ≤ measure M A + measure M B"
  using measure_subadditive[OF m] by simp

lemma (in finite_measure) finite_measure_subadditive_finite:
  assumes "finite I" "A`I ⊆ sets M" shows "measure M (⋃i∈I. A i) ≤ (∑i∈I. measure M (A i))"
  using measure_subadditive_finite[OF assms] by simp

lemma (in finite_measure) finite_measure_subadditive_countably:
  "range A ⊆ sets M ⟹ summable (λi. measure M (A i)) ⟹ measure M (⋃i. A i) ≤ (∑i. measure M (A i))"
  by (rule measure_subadditive_countably)
     (simp_all add: ennreal_suminf_neq_top emeasure_eq_measure)

lemma (in finite_measure) finite_measure_eq_setsum_singleton:
  assumes "finite S" and *: "⋀x. x ∈ S ⟹ {x} ∈ sets M"
  shows "measure M S = (∑x∈S. measure M {x})"
  using measure_eq_setsum_singleton[OF assms] by simp

lemma (in finite_measure) finite_Lim_measure_incseq:
  assumes A: "range A ⊆ sets M" "incseq A"
  shows "(λi. measure M (A i)) ⇢ measure M (⋃i. A i)"
  using Lim_measure_incseq[OF A] by simp

lemma (in finite_measure) finite_Lim_measure_decseq:
  assumes A: "range A ⊆ sets M" "decseq A"
  shows "(λn. measure M (A n)) ⇢ measure M (⋂i. A i)"
  using Lim_measure_decseq[OF A] by simp

lemma (in finite_measure) finite_measure_compl:
  assumes S: "S ∈ sets M"
  shows "measure M (space M - S) = measure M (space M) - measure M S"
  using measure_Diff[OF _ sets.top S sets.sets_into_space] S by simp

lemma (in finite_measure) finite_measure_mono_AE:
  assumes imp: "AE x in M. x ∈ A ⟶ x ∈ B" and B: "B ∈ sets M"
  shows "measure M A ≤ measure M B"
  using assms emeasure_mono_AE[OF imp B]
  by (simp add: emeasure_eq_measure)

lemma (in finite_measure) finite_measure_eq_AE:
  assumes iff: "AE x in M. x ∈ A ⟷ x ∈ B"
  assumes A: "A ∈ sets M" and B: "B ∈ sets M"
  shows "measure M A = measure M B"
  using assms emeasure_eq_AE[OF assms] by (simp add: emeasure_eq_measure)

lemma (in finite_measure) measure_increasing: "increasing M (measure M)"
  by (auto intro!: finite_measure_mono simp: increasing_def)

lemma (in finite_measure) measure_zero_union:
  assumes "s ∈ sets M" "t ∈ sets M" "measure M t = 0"
  shows "measure M (s ∪ t) = measure M s"
using assms
proof -
  have "measure M (s ∪ t) ≤ measure M s"
    using finite_measure_subadditive[of s t] assms by auto
  moreover have "measure M (s ∪ t) ≥ measure M s"
    using assms by (blast intro: finite_measure_mono)
  ultimately show ?thesis by simp
qed

lemma (in finite_measure) measure_eq_compl:
  assumes "s ∈ sets M" "t ∈ sets M"
  assumes "measure M (space M - s) = measure M (space M - t)"
  shows "measure M s = measure M t"
  using assms finite_measure_compl by auto

lemma (in finite_measure) measure_eq_bigunion_image:
  assumes "range f ⊆ sets M" "range g ⊆ sets M"
  assumes "disjoint_family f" "disjoint_family g"
  assumes "⋀ n :: nat. measure M (f n) = measure M (g n)"
  shows "measure M (⋃i. f i) = measure M (⋃i. g i)"
using assms
proof -
  have a: "(λ i. measure M (f i)) sums (measure M (⋃i. f i))"
    by (rule finite_measure_UNION[OF assms(1,3)])
  have b: "(λ i. measure M (g i)) sums (measure M (⋃i. g i))"
    by (rule finite_measure_UNION[OF assms(2,4)])
  show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp
qed

lemma (in finite_measure) measure_countably_zero:
  assumes "range c ⊆ sets M"
  assumes "⋀ i. measure M (c i) = 0"
  shows "measure M (⋃i :: nat. c i) = 0"
proof (rule antisym)
  show "measure M (⋃i :: nat. c i) ≤ 0"
    using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2))
qed simp

lemma (in finite_measure) measure_space_inter:
  assumes events:"s ∈ sets M" "t ∈ sets M"
  assumes "measure M t = measure M (space M)"
  shows "measure M (s ∩ t) = measure M s"
proof -
  have "measure M ((space M - s) ∪ (space M - t)) = measure M (space M - s)"
    using events assms finite_measure_compl[of "t"] by (auto intro!: measure_zero_union)
  also have "(space M - s) ∪ (space M - t) = space M - (s ∩ t)"
    by blast
  finally show "measure M (s ∩ t) = measure M s"
    using events by (auto intro!: measure_eq_compl[of "s ∩ t" s])
qed

lemma (in finite_measure) measure_equiprobable_finite_unions:
  assumes s: "finite s" "⋀x. x ∈ s ⟹ {x} ∈ sets M"
  assumes "⋀ x y. ⟦x ∈ s; y ∈ s⟧ ⟹ measure M {x} = measure M {y}"
  shows "measure M s = real (card s) * measure M {SOME x. x ∈ s}"
proof cases
  assume "s ≠ {}"
  then have "∃ x. x ∈ s" by blast
  from someI_ex[OF this] assms
  have prob_some: "⋀ x. x ∈ s ⟹ measure M {x} = measure M {SOME y. y ∈ s}" by blast
  have "measure M s = (∑ x ∈ s. measure M {x})"
    using finite_measure_eq_setsum_singleton[OF s] by simp
  also have "… = (∑ x ∈ s. measure M {SOME y. y ∈ s})" using prob_some by auto
  also have "… = real (card s) * measure M {(SOME x. x ∈ s)}"
    using setsum_constant assms by simp
  finally show ?thesis by simp
qed simp

lemma (in finite_measure) measure_real_sum_image_fn:
  assumes "e ∈ sets M"
  assumes "⋀ x. x ∈ s ⟹ e ∩ f x ∈ sets M"
  assumes "finite s"
  assumes disjoint: "⋀ x y. ⟦x ∈ s ; y ∈ s ; x ≠ y⟧ ⟹ f x ∩ f y = {}"
  assumes upper: "space M ⊆ (⋃i ∈ s. f i)"
  shows "measure M e = (∑ x ∈ s. measure M (e ∩ f x))"
proof -
  have "e ⊆ (⋃i∈s. f i)"
    using ‹e ∈ sets M› sets.sets_into_space upper by blast
  then have e: "e = (⋃i ∈ s. e ∩ f i)"
    by auto
  hence "measure M e = measure M (⋃i ∈ s. e ∩ f i)" by simp
  also have "… = (∑ x ∈ s. measure M (e ∩ f x))"
  proof (rule finite_measure_finite_Union)
    show "finite s" by fact
    show "(λi. e ∩ f i)`s ⊆ sets M" using assms(2) by auto
    show "disjoint_family_on (λi. e ∩ f i) s"
      using disjoint by (auto simp: disjoint_family_on_def)
  qed
  finally show ?thesis .
qed

lemma (in finite_measure) measure_exclude:
  assumes "A ∈ sets M" "B ∈ sets M"
  assumes "measure M A = measure M (space M)" "A ∩ B = {}"
  shows "measure M B = 0"
  using measure_space_inter[of B A] assms by (auto simp: ac_simps)
lemma (in finite_measure) finite_measure_distr:
  assumes f: "f ∈ measurable M M'"
  shows "finite_measure (distr M M' f)"
proof (rule finite_measureI)
  have "f -` space M' ∩ space M = space M" using f by (auto dest: measurable_space)
  with f show "emeasure (distr M M' f) (space (distr M M' f)) ≠ ∞" by (auto simp: emeasure_distr)
qed

lemma emeasure_gfp[consumes 1, case_names cont measurable]:
  assumes sets[simp]: "⋀s. sets (M s) = sets N"
  assumes "⋀s. finite_measure (M s)"
  assumes cont: "inf_continuous F" "inf_continuous f"
  assumes meas: "⋀P. Measurable.pred N P ⟹ Measurable.pred N (F P)"
  assumes iter: "⋀P s. Measurable.pred N P ⟹ emeasure (M s) {x∈space N. F P x} = f (λs. emeasure (M s) {x∈space N. P x}) s"
  assumes bound: "⋀P. f P ≤ f (λs. emeasure (M s) (space (M s)))"
  shows "emeasure (M s) {x∈space N. gfp F x} = gfp f s"
proof (subst gfp_transfer_bounded[where α="λF s. emeasure (M s) {x∈space N. F x}" and g=f and f=F and
    P="Measurable.pred N", symmetric])
  interpret finite_measure "M s" for s by fact
  fix C assume "decseq C" "⋀i. Measurable.pred N (C i)"
  then show "(λs. emeasure (M s) {x ∈ space N. (INF i. C i) x}) = (INF i. (λs. emeasure (M s) {x ∈ space N. C i x}))"
    unfolding INF_apply[abs_def]
    by (subst INF_emeasure_decseq) (auto simp: antimono_def fun_eq_iff intro!: arg_cong2[where f=emeasure])
next
  show "f x ≤ (λs. emeasure (M s) {x ∈ space N. F top x})" for x
    using bound[of x] sets_eq_imp_space_eq[OF sets] by (simp add: iter)
qed (auto simp add: iter le_fun_def INF_apply[abs_def] intro!: meas cont)

subsection ‹Counting space›

lemma strict_monoI_Suc:
  assumes ord [simp]: "(⋀n. f n < f (Suc n))" shows "strict_mono f"
  unfolding strict_mono_def
proof safe
  fix n m :: nat assume "n < m" then show "f n < f m"
    by (induct m) (auto simp: less_Suc_eq intro: less_trans ord)
qed

lemma emeasure_count_space:
  assumes "X ⊆ A" shows "emeasure (count_space A) X = (if finite X then of_nat (card X) else ∞)"
    (is "_ = ?M X")
  unfolding count_space_def
proof (rule emeasure_measure_of_sigma)
  show "X ∈ Pow A" using ‹X ⊆ A› by auto
  show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow)
  show positive: "positive (Pow A) ?M"
    by (auto simp: positive_def)
  have additive: "additive (Pow A) ?M"
    by (auto simp: card_Un_disjoint additive_def)

  interpret ring_of_sets A "Pow A"
    by (rule ring_of_setsI) auto
  show "countably_additive (Pow A) ?M"
    unfolding countably_additive_iff_continuous_from_below[OF positive additive]
  proof safe
    fix F :: "nat ⇒ 'a set" assume "incseq F"
    show "(λi. ?M (F i)) ⇢ ?M (⋃i. F i)"
    proof cases
      assume "∃i. ∀j≥i. F i = F j"
      then guess i .. note i = this
      { fix j from i ‹incseq F› have "F j ⊆ F i"
          by (cases "i ≤ j") (auto simp: incseq_def) }
      then have eq: "(⋃i. F i) = F i"
        by auto
      with i show ?thesis
        by (auto intro!: Lim_eventually eventually_sequentiallyI[where c=i])
    next
      assume "¬ (∃i. ∀j≥i. F i = F j)"
      then obtain f where f: "⋀i. i ≤ f i" "⋀i. F i ≠ F (f i)" by metis
      then have "⋀i. F i ⊆ F (f i)" using ‹incseq F› by (auto simp: incseq_def)
      with f have *: "⋀i. F i ⊂ F (f i)" by auto

      have "incseq (λi. ?M (F i))"
        using ‹incseq F› unfolding incseq_def by (auto simp: card_mono dest: finite_subset)
      then have "(λi. ?M (F i)) ⇢ (SUP n. ?M (F n))"
        by (rule LIMSEQ_SUP)

      moreover have "(SUP n. ?M (F n)) = top"
      proof (rule ennreal_SUP_eq_top)
        fix n :: nat show "∃k::nat∈UNIV. of_nat n ≤ ?M (F k)"
        proof (induct n)
          case (Suc n)
          then guess k .. note k = this
          moreover have "finite (F k) ⟹ finite (F (f k)) ⟹ card (F k) < card (F (f k))"
            using ‹F k ⊂ F (f k)› by (simp add: psubset_card_mono)
          moreover have "finite (F (f k)) ⟹ finite (F k)"
            using ‹k ≤ f k› ‹incseq F› by (auto simp: incseq_def dest: finite_subset)
          ultimately show ?case
            by (auto intro!: exI[of _ "f k"] simp del: of_nat_Suc)
        qed auto
      qed

      moreover
      have "inj (λn. F ((f ^^ n) 0))"
        by (intro strict_mono_imp_inj_on strict_monoI_Suc) (simp add: *)
      then have 1: "infinite (range (λi. F ((f ^^ i) 0)))"
        by (rule range_inj_infinite)
      have "infinite (Pow (⋃i. F i))"
        by (rule infinite_super[OF _ 1]) auto
      then have "infinite (⋃i. F i)"
        by auto

      ultimately show ?thesis by auto
    qed
  qed
qed

lemma distr_bij_count_space:
  assumes f: "bij_betw f A B"
  shows "distr (count_space A) (count_space B) f = count_space B"
proof (rule measure_eqI)
  have f': "f ∈ measurable (count_space A) (count_space B)"
    using f unfolding Pi_def bij_betw_def by auto
  fix X assume "X ∈ sets (distr (count_space A) (count_space B) f)"
  then have X: "X ∈ sets (count_space B)" by auto
  moreover then have "f -` X ∩ A = the_inv_into A f ` X"
    using f by (auto simp: bij_betw_def subset_image_iff image_iff the_inv_into_f_f intro: the_inv_into_f_f[symmetric])
  moreover have "inj_on (the_inv_into A f) B"
    using X f by (auto simp: bij_betw_def inj_on_the_inv_into)
  with X have "inj_on (the_inv_into A f) X"
    by (auto intro: subset_inj_on)
  ultimately show "emeasure (distr (count_space A) (count_space B) f) X = emeasure (count_space B) X"
    using f unfolding emeasure_distr[OF f' X]
    by (subst (1 2) emeasure_count_space) (auto simp: card_image dest: finite_imageD)
qed simp

lemma emeasure_count_space_finite[simp]:
  "X ⊆ A ⟹ finite X ⟹ emeasure (count_space A) X = of_nat (card X)"
  using emeasure_count_space[of X A] by simp

lemma emeasure_count_space_infinite[simp]:
  "X ⊆ A ⟹ infinite X ⟹ emeasure (count_space A) X = ∞"
  using emeasure_count_space[of X A] by simp

lemma measure_count_space: "measure (count_space A) X = (if X ⊆ A then of_nat (card X) else 0)"
  by (cases "finite X") (auto simp: measure_notin_sets ennreal_of_nat_eq_real_of_nat
                                    measure_zero_top measure_eq_emeasure_eq_ennreal)

lemma emeasure_count_space_eq_0:
  "emeasure (count_space A) X = 0 ⟷ (X ⊆ A ⟶ X = {})"
proof cases
  assume X: "X ⊆ A"
  then show ?thesis
  proof (intro iffI impI)
    assume "emeasure (count_space A) X = 0"
    with X show "X = {}"
      by (subst (asm) emeasure_count_space) (auto split: if_split_asm)
  qed simp
qed (simp add: emeasure_notin_sets)

lemma space_empty: "space M = {} ⟹ M = count_space {}"
  by (rule measure_eqI) (simp_all add: space_empty_iff)

lemma null_sets_count_space: "null_sets (count_space A) = { {} }"
  unfolding null_sets_def by (auto simp add: emeasure_count_space_eq_0)

lemma AE_count_space: "(AE x in count_space A. P x) ⟷ (∀x∈A. P x)"
  unfolding eventually_ae_filter by (auto simp add: null_sets_count_space)

lemma sigma_finite_measure_count_space_countable:
  assumes A: "countable A"
  shows "sigma_finite_measure (count_space A)"
  proof qed (insert A, auto intro!: exI[of _ "(λa. {a}) ` A"])

lemma sigma_finite_measure_count_space:
  fixes A :: "'a::countable set" shows "sigma_finite_measure (count_space A)"
  by (rule sigma_finite_measure_count_space_countable) auto

lemma finite_measure_count_space:
  assumes [simp]: "finite A"
  shows "finite_measure (count_space A)"
  by rule simp

lemma sigma_finite_measure_count_space_finite:
  assumes A: "finite A" shows "sigma_finite_measure (count_space A)"
proof -
  interpret finite_measure "count_space A" using A by (rule finite_measure_count_space)
  show "sigma_finite_measure (count_space A)" ..
qed

subsection ‹Measure restricted to space›

lemma emeasure_restrict_space:
  assumes "Ω ∩ space M ∈ sets M" "A ⊆ Ω"
  shows "emeasure (restrict_space M Ω) A = emeasure M A"
proof cases
  assume "A ∈ sets M"
  show ?thesis
  proof (rule emeasure_measure_of[OF restrict_space_def])
    show "op ∩ Ω ` sets M ⊆ Pow (Ω ∩ space M)" "A ∈ sets (restrict_space M Ω)"
      using ‹A ⊆ Ω› ‹A ∈ sets M› sets.space_closed by (auto simp: sets_restrict_space)
    show "positive (sets (restrict_space M Ω)) (emeasure M)"
      by (auto simp: positive_def)
    show "countably_additive (sets (restrict_space M Ω)) (emeasure M)"
    proof (rule countably_additiveI)
      fix A :: "nat ⇒ _" assume "range A ⊆ sets (restrict_space M Ω)" "disjoint_family A"
      with assms have "⋀i. A i ∈ sets M" "⋀i. A i ⊆ space M" "disjoint_family A"
        by (fastforce simp: sets_restrict_space_iff[OF assms(1)] image_subset_iff
                      dest: sets.sets_into_space)+
      then show "(∑i. emeasure M (A i)) = emeasure M (⋃i. A i)"
        by (subst suminf_emeasure) (auto simp: disjoint_family_subset)
    qed
  qed
next
  assume "A ∉ sets M"
  moreover with assms have "A ∉ sets (restrict_space M Ω)"
    by (simp add: sets_restrict_space_iff)
  ultimately show ?thesis
    by (simp add: emeasure_notin_sets)
qed

lemma measure_restrict_space:
  assumes "Ω ∩ space M ∈ sets M" "A ⊆ Ω"
  shows "measure (restrict_space M Ω) A = measure M A"
  using emeasure_restrict_space[OF assms] by (simp add: measure_def)

lemma AE_restrict_space_iff:
  assumes "Ω ∩ space M ∈ sets M"
  shows "(AE x in restrict_space M Ω. P x) ⟷ (AE x in M. x ∈ Ω ⟶ P x)"
proof -
  have ex_cong: "⋀P Q f. (⋀x. P x ⟹ Q x) ⟹ (⋀x. Q x ⟹ P (f x)) ⟹ (∃x. P x) ⟷ (∃x. Q x)"
    by auto
  { fix X assume X: "X ∈ sets M" "emeasure M X = 0"
    then have "emeasure M (Ω ∩ space M ∩ X) ≤ emeasure M X"
      by (intro emeasure_mono) auto
    then have "emeasure M (Ω ∩ space M ∩ X) = 0"
      using X by (auto intro!: antisym) }
  with assms show ?thesis
    unfolding eventually_ae_filter
    by (auto simp add: space_restrict_space null_sets_def sets_restrict_space_iff
                       emeasure_restrict_space cong: conj_cong
             intro!: ex_cong[where f="λX. (Ω ∩ space M) ∩ X"])
qed

lemma restrict_restrict_space:
  assumes "A ∩ space M ∈ sets M" "B ∩ space M ∈ sets M"
  shows "restrict_space (restrict_space M A) B = restrict_space M (A ∩ B)" (is "?l = ?r")
proof (rule measure_eqI[symmetric])
  show "sets ?r = sets ?l"
    unfolding sets_restrict_space image_comp by (intro image_cong) auto
next
  fix X assume "X ∈ sets (restrict_space M (A ∩ B))"
  then obtain Y where "Y ∈ sets M" "X = Y ∩ A ∩ B"
    by (auto simp: sets_restrict_space)
  with assms sets.Int[OF assms] show "emeasure ?r X = emeasure ?l X"
    by (subst (1 2) emeasure_restrict_space)
       (auto simp: space_restrict_space sets_restrict_space_iff emeasure_restrict_space ac_simps)
qed

lemma restrict_count_space: "restrict_space (count_space B) A = count_space (A ∩ B)"
proof (rule measure_eqI)
  show "sets (restrict_space (count_space B) A) = sets (count_space (A ∩ B))"
    by (subst sets_restrict_space) auto
  moreover fix X assume "X ∈ sets (restrict_space (count_space B) A)"
  ultimately have "X ⊆ A ∩ B" by auto
  then show "emeasure (restrict_space (count_space B) A) X = emeasure (count_space (A ∩ B)) X"
    by (cases "finite X") (auto simp add: emeasure_restrict_space)
qed

lemma sigma_finite_measure_restrict_space:
  assumes "sigma_finite_measure M"
  and A: "A ∈ sets M"
  shows "sigma_finite_measure (restrict_space M A)"
proof -
  interpret sigma_finite_measure M by fact
  from sigma_finite_countable obtain C
    where C: "countable C" "C ⊆ sets M" "(⋃C) = space M" "∀a∈C. emeasure M a ≠ ∞"
    by blast
  let ?C = "op ∩ A ` C"
  from C have "countable ?C" "?C ⊆ sets (restrict_space M A)" "(⋃?C) = space (restrict_space M A)"
    by(auto simp add: sets_restrict_space space_restrict_space)
  moreover {
    fix a
    assume "a ∈ ?C"
    then obtain a' where "a = A ∩ a'" "a' ∈ C" ..
    then have "emeasure (restrict_space M A) a ≤ emeasure M a'"
      using A C by(auto simp add: emeasure_restrict_space intro: emeasure_mono)
    also have "… < ∞" using C(4)[rule_format, of a'] ‹a' ∈ C› by (simp add: less_top)
    finally have "emeasure (restrict_space M A) a ≠ ∞" by simp }
  ultimately show ?thesis
    by unfold_locales (rule exI conjI|assumption|blast)+
qed

lemma finite_measure_restrict_space:
  assumes "finite_measure M"
  and A: "A ∈ sets M"
  shows "finite_measure (restrict_space M A)"
proof -
  interpret finite_measure M by fact
  show ?thesis
    by(rule finite_measureI)(simp add: emeasure_restrict_space A space_restrict_space)
qed

lemma restrict_distr:
  assumes [measurable]: "f ∈ measurable M N"
  assumes [simp]: "Ω ∩ space N ∈ sets N" and restrict: "f ∈ space M → Ω"
  shows "restrict_space (distr M N f) Ω = distr M (restrict_space N Ω) f"
  (is "?l = ?r")
proof (rule measure_eqI)
  fix A assume "A ∈ sets (restrict_space (distr M N f) Ω)"
  with restrict show "emeasure ?l A = emeasure ?r A"
    by (subst emeasure_distr)
       (auto simp: sets_restrict_space_iff emeasure_restrict_space emeasure_distr
             intro!: measurable_restrict_space2)
qed (simp add: sets_restrict_space)

lemma measure_eqI_restrict_generator:
  assumes E: "Int_stable E" "E ⊆ Pow Ω" "⋀X. X ∈ E ⟹ emeasure M X = emeasure N X"
  assumes sets_eq: "sets M = sets N" and Ω: "Ω ∈ sets M"
  assumes "sets (restrict_space M Ω) = sigma_sets Ω E"
  assumes "sets (restrict_space N Ω) = sigma_sets Ω E"
  assumes ae: "AE x in M. x ∈ Ω" "AE x in N. x ∈ Ω"
  assumes A: "countable A" "A ≠ {}" "A ⊆ E" "⋃A = Ω" "⋀a. a ∈ A ⟹ emeasure M a ≠ ∞"
  shows "M = N"
proof (rule measure_eqI)
  fix X assume X: "X ∈ sets M"
  then have "emeasure M X = emeasure (restrict_space M Ω) (X ∩ Ω)"
    using ae Ω by (auto simp add: emeasure_restrict_space intro!: emeasure_eq_AE)
  also have "restrict_space M Ω = restrict_space N Ω"
  proof (rule measure_eqI_generator_eq)
    fix X assume "X ∈ E"
    then show "emeasure (restrict_space M Ω) X = emeasure (restrict_space N Ω) X"
      using E Ω by (subst (1 2) emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq])
  next
    show "range (from_nat_into A) ⊆ E" "(⋃i. from_nat_into A i) = Ω"
      using A by (auto cong del: strong_SUP_cong)
  next
    fix i
    have "emeasure (restrict_space M Ω) (from_nat_into A i) = emeasure M (from_nat_into A i)"
      using A Ω by (subst emeasure_restrict_space)
                   (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq] intro: from_nat_into)
    with A show "emeasure (restrict_space M Ω) (from_nat_into A i) ≠ ∞"
      by (auto intro: from_nat_into)
  qed fact+
  also have "emeasure (restrict_space N Ω) (X ∩ Ω) = emeasure N X"
    using X ae Ω by (auto simp add: emeasure_restrict_space sets_eq intro!: emeasure_eq_AE)
  finally show "emeasure M X = emeasure N X" .
qed fact

subsection ‹Null measure›

definition "null_measure M = sigma (space M) (sets M)"

lemma space_null_measure[simp]: "space (null_measure M) = space M"
  by (simp add: null_measure_def)

lemma sets_null_measure[simp, measurable_cong]: "sets (null_measure M) = sets M"
  by (simp add: null_measure_def)

lemma emeasure_null_measure[simp]: "emeasure (null_measure M) X = 0"
  by (cases "X ∈ sets M", rule emeasure_measure_of)
     (auto simp: positive_def countably_additive_def emeasure_notin_sets null_measure_def
           dest: sets.sets_into_space)

lemma measure_null_measure[simp]: "measure (null_measure M) X = 0"
  by (intro measure_eq_emeasure_eq_ennreal) auto

lemma null_measure_idem [simp]: "null_measure (null_measure M) = null_measure M"
  by(rule measure_eqI) simp_all

subsection ‹Scaling a measure›

definition scale_measure :: "ennreal ⇒ 'a measure ⇒ 'a measure"
where
  "scale_measure r M = measure_of (space M) (sets M) (λA. r * emeasure M A)"

lemma space_scale_measure: "space (scale_measure r M) = space M"
  by (simp add: scale_measure_def)

lemma sets_scale_measure [simp, measurable_cong]: "sets (scale_measure r M) = sets M"
  by (simp add: scale_measure_def)

lemma emeasure_scale_measure [simp]:
  "emeasure (scale_measure r M) A = r * emeasure M A"
  (is "_ = ?μ A")
proof(cases "A ∈ sets M")
  case True
  show ?thesis unfolding scale_measure_def
  proof(rule emeasure_measure_of_sigma)
    show "sigma_algebra (space M) (sets M)" ..
    show "positive (sets M) ?μ" by (simp add: positive_def)
    show "countably_additive (sets M) ?μ"
    proof (rule countably_additiveI)
      fix A :: "nat ⇒ _"  assume *: "range A ⊆ sets M" "disjoint_family A"
      have "(∑i. ?μ (A i)) = r * (∑i. emeasure M (A i))"
        by simp
      also have "… = ?μ (⋃i. A i)" using * by(simp add: suminf_emeasure)
      finally show "(∑i. ?μ (A i)) = ?μ (⋃i. A i)" .
    qed
  qed(fact True)
qed(simp add: emeasure_notin_sets)

lemma scale_measure_1 [simp]: "scale_measure 1 M = M"
  by(rule measure_eqI) simp_all

lemma scale_measure_0[simp]: "scale_measure 0 M = null_measure M"
  by(rule measure_eqI) simp_all

lemma measure_scale_measure [simp]: "0 ≤ r ⟹ measure (scale_measure r M) A = r * measure M A"
  using emeasure_scale_measure[of r M A]
    emeasure_eq_ennreal_measure[of M A]
    measure_eq_emeasure_eq_ennreal[of _ "scale_measure r M" A]
  by (cases "emeasure (scale_measure r M) A = top")
     (auto simp del: emeasure_scale_measure
           simp: ennreal_top_eq_mult_iff ennreal_mult_eq_top_iff measure_zero_top ennreal_mult[symmetric])

lemma scale_scale_measure [simp]:
  "scale_measure r (scale_measure r' M) = scale_measure (r * r') M"
  by (rule measure_eqI) (simp_all add: max_def mult.assoc)

lemma scale_null_measure [simp]: "scale_measure r (null_measure M) = null_measure M"
  by (rule measure_eqI) simp_all

subsection ‹Measures form a chain-complete partial order›

instantiation measure :: (type) order_bot
begin

definition bot_measure :: "'a measure" where
  "bot_measure = sigma {} {{}}"

lemma space_bot[simp]: "space bot = {}"
  unfolding bot_measure_def by (rule space_measure_of) auto

lemma sets_bot[simp]: "sets bot = {{}}"
  unfolding bot_measure_def by (subst sets_measure_of) auto

lemma emeasure_bot[simp]: "emeasure bot = (λx. 0)"
  unfolding bot_measure_def by (rule emeasure_sigma)

inductive less_eq_measure :: "'a measure ⇒ 'a measure ⇒ bool" where
  "sets N = sets M ⟹ (⋀A. A ∈ sets M ⟹ emeasure M A ≤ emeasure N A) ⟹ less_eq_measure M N"
| "less_eq_measure bot N"

definition less_measure :: "'a measure ⇒ 'a measure ⇒ bool" where
  "less_measure M N ⟷ (M ≤ N ∧ ¬ N ≤ M)"

instance
proof (standard, goal_cases)
  case 1 then show ?case
    unfolding less_measure_def ..
next
  case (2 M) then show ?case
    by (intro less_eq_measure.intros) auto
next
  case (3 M N L) then show ?case
    apply (safe elim!: less_eq_measure.cases)
    apply (simp_all add: less_eq_measure.intros)
    apply (rule less_eq_measure.intros)
    apply simp
    apply (blast intro: order_trans) []
    unfolding less_eq_measure.simps
    apply (rule disjI2)
    apply simp
    apply (rule measure_eqI)
    apply (auto intro!: antisym)
    done
next
  case (4 M N) then show ?case
    apply (safe elim!: less_eq_measure.cases intro!: measure_eqI)
    apply simp
    apply simp
    apply (blast intro: antisym)
    apply (simp)
    apply simp
    done
qed (rule less_eq_measure.intros)
end

lemma le_emeasureD: "M ≤ N ⟹ emeasure M A ≤ emeasure N A"
  by (cases "A ∈ sets M") (auto elim!: less_eq_measure.cases simp: emeasure_notin_sets)

lemma le_sets: "N ≤ M ⟹ sets N ≤ sets M"
  unfolding less_eq_measure.simps by auto

instantiation measure :: (type) ccpo
begin

definition Sup_measure :: "'a measure set ⇒ 'a measure" where
  "Sup_measure A = measure_of (SUP a:A. space a) (SUP a:A. sets a) (SUP a:A. emeasure a)"

lemma
  assumes A: "Complete_Partial_Order.chain op ≤ A" and a: "a ≠ bot" "a ∈ A"
  shows space_Sup: "space (Sup A) = space a"
    and sets_Sup: "sets (Sup A) = sets a"
proof -
  have sets: "(SUP a:A. sets a) = sets a"
  proof (intro antisym SUP_least)
    fix a' show "a' ∈ A ⟹ sets a' ⊆ sets a"
      using a chainD[OF A, of a a'] by (auto elim!: less_eq_measure.cases)
  qed (insert ‹a∈A›, auto)
  have space: "(SUP a:A. space a) = space a"
  proof (intro antisym SUP_least)
    fix a' show "a' ∈ A ⟹ space a' ⊆ space a"
      using a chainD[OF A, of a a'] by (intro sets_le_imp_space_le) (auto elim!: less_eq_measure.cases)
  qed (insert ‹a∈A›, auto)
  show "space (Sup A) = space a"
    unfolding Sup_measure_def sets space sets.space_measure_of_eq ..
  show "sets (Sup A) = sets a"
    unfolding Sup_measure_def sets space sets.sets_measure_of_eq ..
qed

lemma emeasure_Sup:
  assumes A: "Complete_Partial_Order.chain op ≤ A" "A ≠ {}"
  assumes "X ∈ sets (Sup A)"
  shows "emeasure (Sup A) X = (SUP a:A. emeasure a) X"
proof (rule emeasure_measure_of[OF Sup_measure_def])
  show "countably_additive (sets (Sup A)) (SUP a:A. emeasure a)"
    unfolding countably_additive_def
  proof safe
    fix F :: "nat ⇒ 'a set" assume F: "range F ⊆ sets (Sup A)" "disjoint_family F"
    show "(∑i. (SUP a:A. emeasure a) (F i)) = SUPREMUM A emeasure (UNION UNIV F)"
      unfolding SUP_apply
    proof (subst ennreal_suminf_SUP_eq_directed)
      fix N i j assume "i ∈ A" "j ∈ A"
      with A(1)
      show "∃k∈A. ∀n∈N. emeasure i (F n) ≤ emeasure k (F n) ∧ emeasure j (F n) ≤ emeasure k (F n)"
        by (blast elim: chainE dest: le_emeasureD)
    next
      show "(SUP n:A. ∑i. emeasure n (F i)) = (SUP y:A. emeasure y (UNION UNIV F))"
      proof (intro SUP_cong refl)
        fix a assume "a ∈ A" then show "(∑i. emeasure a (F i)) = emeasure a (UNION UNIV F)"
          using sets_Sup[OF A(1), of a] F by (cases "a = bot") (auto simp: suminf_emeasure)
      qed
    qed
  qed
qed (insert ‹A ≠ {}› ‹X ∈ sets (Sup A)›, auto simp: positive_def dest: sets.sets_into_space intro: SUP_upper2)

instance
proof
  fix A and x :: "'a measure" assume A: "Complete_Partial_Order.chain op ≤ A" and x: "x ∈ A"
  show "x ≤ Sup A"
  proof cases
    assume "x ≠ bot"
    show ?thesis
    proof
      show "sets (Sup A) = sets x"
        using A ‹x ≠ bot› x by (rule sets_Sup)
      with x show "⋀a. a ∈ sets x ⟹ emeasure x a ≤ emeasure (Sup A) a"
        by (subst emeasure_Sup[OF A]) (auto intro: SUP_upper)
    qed
  qed simp
next
  fix A and x :: "'a measure" assume A: "Complete_Partial_Order.chain op ≤ A" and x: "⋀z. z ∈ A ⟹ z ≤ x"
  consider "A = {}" | "A = {bot}" | x where "x∈A" "x ≠ bot"
    by blast
  then show "Sup A ≤ x"
  proof cases
    assume "A = {}"
    moreover have "Sup ({}::'a measure set) = bot"
      by (auto simp add: Sup_measure_def sigma_sets_empty_eq intro!: measure_eqI)
    ultimately show ?thesis
      by simp
  next
    assume "A = {bot}"
    moreover have "Sup ({bot}::'a measure set) = bot"
      by (auto simp add: Sup_measure_def sigma_sets_empty_eq intro!: measure_eqI)
     ultimately show ?thesis
      by simp
  next
    fix a assume "a ∈ A" "a ≠ bot"
    then have "a ≤ x" "x ≠ bot" "a ≠ bot"
      using x[OF ‹a ∈ A›] by (auto simp: bot_unique)
    then have "sets x = sets a"
      by (auto elim: less_eq_measure.cases)

    show "Sup A ≤ x"
    proof (rule less_eq_measure.intros)
      show "sets x = sets (Sup A)"
        by (subst sets_Sup[OF A ‹a ≠ bot› ‹a ∈ A›]) fact
    next
      fix X assume "X ∈ sets (Sup A)"
      then have "emeasure (Sup A) X ≤ (SUP a:A. emeasure a X)"
        using ‹a∈A› by (subst emeasure_Sup[OF A _]) auto
      also have "… ≤ emeasure x X"
        by (intro SUP_least le_emeasureD x)
      finally show "emeasure (Sup A) X ≤ emeasure x X" .
    qed
  qed
qed
end

lemma
  assumes A: "Complete_Partial_Order.chain op ≤ (f ` A)" and a: "a ∈ A" "f a ≠ bot"
  shows space_SUP: "space (SUP M:A. f M) = space (f a)"
    and sets_SUP: "sets (SUP M:A. f M) = sets (f a)"
by(rule space_Sup[OF A a(2) imageI[OF a(1)]] sets_Sup[OF A a(2) imageI[OF a(1)]])+

lemma emeasure_SUP:
  assumes A: "Complete_Partial_Order.chain op ≤ (f ` A)" "A ≠ {}"
  assumes "X ∈ sets (SUP M:A. f M)"
  shows "emeasure (SUP M:A. f M) X = (SUP M:A. emeasure (f M)) X"
using ‹X ∈ _› by(subst emeasure_Sup[OF A(1)]; simp add: A)

end