Theory Measure_Space
section ‹Measure Spaces›
theory Measure_Space
imports
Measurable "HOL-Library.Extended_Nonnegative_Real"
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: sum.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 (use assms in simp)
ultimately show ?thesis using assms
by (simp add: suminf_eq_SUP)
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 sum_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
with ‹finite P› show ?thesis
by (simp add: indicator_def)
qed
text ‹
The type for emeasure spaces is already defined in \<^theory>‹HOL-Analysis.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< 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)"
using LIMSEQ_binaryset f sums_def by blast
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>‹HOL-Analysis.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))"
by (metis addf Diff_disjoint ‹y - x ∈ M› additiveD xy(1))
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 (simp add: increasingD)
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))"
by (simp add: insert)
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. F i ≠ {} ⟶ (∃x. x ∈ F i)" by auto
then obtain f where f: "⋀i. F i ≠ {} ⟹ f i ∈ F i" by metis
have finU: "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 (simp add: finU finite_subset)
show inj_f: "inj_on f {i. F i ≠ {}}"
using f disj
by (simp add: inj_on_def disjoint_family_on_def disjoint_iff) metis
qed
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 sum.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 ⟶ ⋃(A ` UNIV) ∈ M ⟶ (∑i. f (A i)) = f (⋃(A ` UNIV))"
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))"
by (simp add: summable_LIMSEQ)
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[rule_format]: "∀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 "range (λi. ⋃i<i. A i) ⊆ M" "(⋃i. ⋃i<i. A i) ∈ M"
using A * by auto
then have "(λn. f (⋃i<n. A i)) ⇢ f (⋃i. A i)"
unfolding *[symmetric] by (force intro!: cont incseq_SucI)+
moreover have "⋀n. f (⋃i<n. A i) = (∑i<n. f (A i))"
using A
by (intro additive_sum[OF f, symmetric]) (auto intro: disjoint_family_on_mono)
ultimately
have "(λi. f (A i)) sums f (⋃i. A i)"
unfolding sums_def by simp
then show "(∑i. f (A i)) = f (⋃i. A i)"
by (metis sums_unique)
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[rule_format]: "(∀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[of A] show "(λi. f (A i)) ⇢ 0"
using ‹positive M f›[unfolded positive_def] by auto
next
assume cont[rule_format]: "∀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)
have f_fin: "f (A i - (⋂i. A i)) ≠ ∞" for i
using A by (metis Diff Diff_subset f_mono infinity_ennreal_def range_subsetD top_unique)
have "(λi. f (A i - (⋂i. A i))) ⇢ 0"
proof (intro cont[ OF _ decseq _ f_fin])
show "range (λi. A i - (⋂i. A i)) ⊆ M" "(⋂i. A i - (⋂i. A i)) = {}"
using A by auto
qed
with INF_Lim decseq_f have "(INF n. f (A n - (⋂i. A i))) = 0" by metis
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
using INF_ennreal_add_const by presburger
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 assms f by fastforce
}
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 emeasure_Un:
"A ∈ sets M ⟹ B ∈ sets M ⟹ emeasure M (A ∪ B) = emeasure M A + emeasure M (B - A)"
using plus_emeasure[of A M "B - A"] by auto
lemma emeasure_Un_Int:
assumes "A ∈ sets M" "B ∈ sets M"
shows "emeasure M A + emeasure M B = emeasure M (A ∪ B) + emeasure M (A ∩ B)"
proof -
have "A = (A-B) ∪ (A ∩ B)" by auto
then have "emeasure M A = emeasure M (A-B) + emeasure M (A ∩ B)"
by (metis Diff_Diff_Int Diff_disjoint assms plus_emeasure sets.Diff)
moreover have "A ∪ B = (A-B) ∪ B" by auto
then have "emeasure M (A ∪ B) = emeasure M (A-B) + emeasure M B"
by (metis Diff_disjoint Int_commute assms plus_emeasure sets.Diff)
ultimately show ?thesis by (metis add.assoc add.commute)
qed
lemma sum_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 "emeasure M B ≠ ∞"
and "A ∈ sets M" "B ∈ sets M" and "B ⊆ A"
shows "emeasure M (A - B) = emeasure M A - emeasure M B"
by (smt (verit, best) add_diff_self_ennreal assms emeasure_Un emeasure_mono
ennreal_add_left_cancel le_iff_sup)
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 (smt (verit, best) Inf_lower diff_diff_ennreal le_MI finite range_eqI)
qed
lemma INF_emeasure_decseq':
assumes A: "⋀i. A i ∈ 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 -
from finite obtain i where i: "emeasure M (A i) < ∞"
by (auto simp: less_top)
have fin: "i ≤ j ⟹ emeasure M (A j) < ∞" for j
by (rule le_less_trans[OF emeasure_mono i]) (auto intro!: decseqD[OF ‹decseq A›] A)
have "(INF n. emeasure M (A n)) = (INF n. emeasure M (A (n + i)))"
proof (rule INF_eq)
show "∃j∈UNIV. emeasure M (A (j + i)) ≤ emeasure M (A i')" for i'
by (meson A ‹decseq A› decseq_def emeasure_mono iso_tuple_UNIV_I nat_le_iff_add)
qed auto
also have "… = emeasure M (INF n. (A (n + i)))"
using A ‹decseq A› fin by (intro INF_emeasure_decseq) (auto simp: decseq_def less_top)
also have "(INF n. (A (n + i))) = (INF n. A n)"
by (meson INF_eq UNIV_I assms(2) decseqD le_add1)
finally show ?thesis .
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"
define L where "L n = (LEAST i. i ∈ I ∧ i ≥ n)" for 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 f=F , 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
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 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_sum_singleton:
assumes "finite S" "⋀x. x ∈ S ⟹ {x} ∈ sets M"
shows "emeasure M S = (∑x∈S. emeasure M {x})"
using sum_emeasure[of "λx. {x}" S M] assms
by (auto simp: disjoint_family_on_def subset_eq)
lemma sum_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 sum_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_sum_singleton) (auto dest: finite_subset)
also have "… = (∑a∈X. emeasure N {a})"
using X eq by (auto intro!: sum.cong)
also have "… = emeasure N X"
using X ‹finite A› by (subst emeasure_eq_sum_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 space_empty: "space M = {} ⟹ M = count_space {}"
by (rule measure_eqI) (simp_all add: space_empty_iff)
lemma measure_eqI_generator_eq_countable:
fixes M N :: "'a measure" and E :: "'a set set" and A :: "'a set set"
assumes E: "Int_stable E" "E ⊆ Pow Ω" "⋀X. X ∈ E ⟹ emeasure M X = emeasure N X"
and sets: "sets M = sigma_sets Ω E" "sets N = sigma_sets Ω E"
and A: "A ⊆ E" "(⋃A) = Ω" "countable A" "⋀a. a ∈ A ⟹ emeasure M a ≠ ∞"
shows "M = N"
proof cases
assume "Ω = {}"
have *: "sigma_sets Ω E = sets (sigma Ω E)"
using E(2) by simp
have "space M = Ω" "space N = Ω"
using sets E(2) unfolding * by (auto dest: sets_eq_imp_space_eq simp del: sets_measure_of)
then show "M = N"
unfolding ‹Ω = {}› by (auto dest: space_empty)
next
assume "Ω ≠ {}" with ‹⋃A = Ω› have "A ≠ {}" by auto
from this ‹countable A› have rng: "range (from_nat_into A) = A"
by (rule range_from_nat_into)
show "M = N"
proof (rule measure_eqI_generator_eq[OF E sets])
show "range (from_nat_into A) ⊆ E"
unfolding rng using ‹A ⊆ E› .
show "(⋃i. from_nat_into A i) = Ω"
unfolding rng using ‹⋃A = Ω› .
show "emeasure M (from_nat_into A i) ≠ ∞" for i
using rng by (intro A) auto
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
lemma emeasure_Un':
assumes "A ∈ sets M" "B ∈ sets M" "A ∩ B ∈ null_sets M"
shows "emeasure M (A ∪ B) = emeasure M A + emeasure M B"
proof -
have "A ∪ B = A ∪ (B - A ∩ B)" by blast
also have "emeasure M … = emeasure M A + emeasure M (B - A ∩ B)"
using assms by (subst plus_emeasure) auto
also have "emeasure M (B - A ∩ B) = emeasure M B"
using assms by (intro emeasure_Diff_null_set) auto
finally show ?thesis .
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"
(‹(‹open_block notation=‹binder AE››AE _ in _. _)› [0,0,10] 10)
syntax_consts
"_almost_everywhere" ⇌ almost_everywhere
translations
"AE x in M. P" ⇌ "CONST almost_everywhere M (λx. P)"
abbreviation
"set_almost_everywhere A M P ≡ AE x in M. x ∈ A ⟶ P x"
syntax
"_set_almost_everywhere" :: "pttrn ⇒ 'a set ⇒ 'a ⇒ bool ⇒ bool"
(‹(‹open_block notation=‹binder AE››AE _∈_ in _./ _)› [0,0,0,10] 10)
syntax_consts
"_set_almost_everywhere" ⇌ set_almost_everywhere
translations
"AE x∈A in M. P" ⇌ "CONST set_almost_everywhere A 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_filter_eq_bot_iff: "ae_filter M = bot ⟷ emeasure M (space M) = 0"
proof -
have "ae_filter M = bot ⟷ (AE x in M. False)"
using trivial_limit_def by blast
also have "… ⟷ space M ∈ null_sets M"
by (simp add: AE_iff_null_sets eventually_ae_filter)
also have "… ⟷ emeasure M (space M) = 0"
by auto
finally show ?thesis .
qed
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"
shows "emeasure M {x∈space M. ¬ P x} = 0"
by (metis (mono_tags, lifting) AE_iff_null assms emeasure_notin_sets null_setsD1)
lemma AE_E3:
assumes "AE x in M. P x"
obtains N where "⋀x. x ∈ space M - N ⟹ P x" "N ∈ null_sets M"
using assms unfolding eventually_ae_filter by auto
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"
using assms by (fact eventually_rev_mp)
text ‹The next lemma is convenient to combine with a lemma whose conclusion is of the
form ‹AE x in M. P x = Q x›: for such a lemma, there is no ‹[symmetric]› variant,
but using ‹AE_symmetric[OF…]› will replace it.›
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_symmetric:
assumes "AE x in M. P x = Q x"
shows "AE x in M. Q x = P x"
using assms by auto
lemma AE_impI:
"(P ⟹ AE x in M. Q x) ⟹ AE x in M. P ⟶ Q x"
by fastforce
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] obtain N
where N: "{x ∈ space M. ¬ P x} ⊆ N" "emeasure M N = 0" "N ∈ sets M"
by auto
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_cong_simp: "M = N ⟹ (⋀x. x ∈ space N =simp=> P x = Q x) ⟹ (AE x in M. P x) ⟷ (AE x in N. Q x)"
by (auto simp: simp_implies_def)
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_ball_countable':
"(⋀N. N ∈ I ⟹ AE x in M. P N x) ⟹ countable I ⟹ AE x in M. ∀N ∈ I. P N x"
unfolding AE_ball_countable by simp
lemma AE_pairwise: "countable F ⟹ pairwise (λA B. AE x in M. R x A B) F ⟷ (AE x in M. pairwise (R x) F)"
unfolding pairwise_alt by (simp add: AE_ball_countable)
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_0_AE:
assumes "emeasure M (space M) = 0"
shows "AE x in M. P x"
using eventually_ae_filter assms by blast
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: sum_Pinfty less_top)
finally show ?thesis by simp
qed
show "incseq (λn. ⋃i≤n. F i)"
by (force simp: incseq_def)
qed
qed
lemma (in sigma_finite_measure) approx_PInf_emeasure_with_finite:
fixes C::real
assumes W_meas: "W ∈ sets M"
and W_inf: "emeasure M W = ∞"
obtains Z where "Z ∈ sets M" "Z ⊆ W" "emeasure M Z < ∞" "emeasure M Z > C"
proof -
obtain A :: "nat ⇒ 'a set"
where A: "range A ⊆ sets M" "(⋃i. A i) = space M" "⋀i. emeasure M (A i) ≠ ∞" "incseq A"
using sigma_finite_incseq by blast
define B where "B = (λi. W ∩ A i)"
have B_meas: "⋀i. B i ∈ sets M" using W_meas ‹range A ⊆ sets M› B_def by blast
have b: "⋀i. B i ⊆ W" using B_def by blast
{ fix i
have "emeasure M (B i) ≤ emeasure M (A i)"
using A by (intro emeasure_mono) (auto simp: B_def)
also have "emeasure M (A i) < ∞"
using ‹⋀i. emeasure M (A i) ≠ ∞› by (simp add: less_top)
finally have "emeasure M (B i) < ∞" . }
note c = this
have "W = (⋃i. B i)" using B_def ‹(⋃i. A i) = space M› W_meas by auto
moreover have "incseq B" using B_def ‹incseq A› by (simp add: incseq_def subset_eq)
ultimately have "(λi. emeasure M (B i)) ⇢ emeasure M W" using W_meas B_meas
by (simp add: B_meas Lim_emeasure_incseq image_subset_iff)
then have "(λi. emeasure M (B i)) ⇢ ∞" using W_inf by simp
from order_tendstoD(1)[OF this, of C]
obtain i where d: "emeasure M (B i) > C"
by (auto simp: eventually_sequentially)
have "B i ∈ sets M" "B i ⊆ W" "emeasure M (B i) < ∞" "emeasure M (B i) > C"
using B_meas b c d by auto
then show ?thesis using that by blast
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 distr_id2: "sets M = sets N ⟹ distr N M (λ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] obtain N
where "{x ∈ space (distr M M' f). ¬ P x} ⊆ N"
"emeasure (distr M M' f) N = 0"
"N ∈ sets (distr M M' f)"
by auto
with f show ?thesis
by (simp add: eventually_ae_filter, 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)
proposition 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: top_unique)
qed (auto dest: sets.sets_into_space)
lemma measure_nonneg[simp]: "0 ≤ measure M A"
unfolding measure_def by auto
lemma measure_nonneg' [simp]: "¬ measure M A < 0"
using measure_nonneg not_le by blast
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)
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 less_top
del: real_of_ereal_enn2ereal)
lemma enn2real_sum:"(⋀i. i ∈ I ⟹ f i < top) ⟹ enn2real (sum f I) = sum (enn2real ∘ f) I"
by (induction I rule: infinite_finite_induct) (auto simp: enn2real_plus)
lemma 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: measure_def)
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 plus_emeasure[symmetric] enn2real_plus less_top)
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 sum_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
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)"
unfolding measure_def
by (metis emeasure_subadditive[OF measurable] fin enn2real_mono enn2real_plus
ennreal_add_less_top infinity_ennreal_def less_top)
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: sum_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 -
have **: "⋀i. emeasure M (A i) ≠ top"
using fin ennreal_suminf_lessD[of "λi. emeasure M (A i)"] by (simp add: less_top)
have ge0: "(∑i. Sigma_Algebra.measure M (A i)) ≥ 0"
using fin emeasure_eq_ennreal_measure[OF **]
by (metis infinity_ennreal_def measure_nonneg suminf_cong suminf_nonneg summable_suminf_not_top)
have "emeasure M (⋃i. A i) ≠ top"
by (metis A emeasure_subadditive_countably fin infinity_ennreal_def neq_top_trans)
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
using ge0 ennreal_le_iff by blast
qed
lemma measure_Un_null_set: "A ∈ sets M ⟹ B ∈ null_sets M ⟹ measure M (A ∪ B) = measure M A"
by (simp add: measure_def emeasure_Un_null_set)
lemma measure_Diff_null_set: "A ∈ sets M ⟹ B ∈ null_sets M ⟹ measure M (A - B) = measure M A"
by (simp add: measure_def emeasure_Diff_null_set)
lemma measure_eq_sum_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_sum_singleton[of S M]
by (intro measure_eq_emeasure_eq_ennreal) (auto simp: sum_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 (measure M (A x))) ⇢ ennreal (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 (measure M (A x))) ⇢ ennreal (measure M (⋂i. A i))"
using fin A by (auto intro!: Lim_emeasure_decseq)
qed auto
subsection ‹Set of measurable sets with finite measure›
definition fmeasurable :: "'a measure ⇒ 'a set set" where
"fmeasurable M = {A∈sets M. emeasure M A < ∞}"
lemma fmeasurableD[dest, measurable_dest]: "A ∈ fmeasurable M ⟹ A ∈ sets M"
by (auto simp: fmeasurable_def)
lemma fmeasurableD2: "A ∈ fmeasurable M ⟹ emeasure M A ≠ top"
by (auto simp: fmeasurable_def)
lemma fmeasurableI: "A ∈ sets M ⟹ emeasure M A < ∞ ⟹ A ∈ fmeasurable M"
by (auto simp: fmeasurable_def)
lemma fmeasurableI_null_sets: "A ∈ null_sets M ⟹ A ∈ fmeasurable M"
by (auto simp: fmeasurable_def)
lemma fmeasurableI2: "A ∈ fmeasurable M ⟹ B ⊆ A ⟹ B ∈ sets M ⟹ B ∈ fmeasurable M"
using emeasure_mono[of B A M] by (auto simp: fmeasurable_def)
lemma measure_mono_fmeasurable:
"A ⊆ B ⟹ A ∈ sets M ⟹ B ∈ fmeasurable M ⟹ measure M A ≤ measure M B"
by (auto simp: measure_def fmeasurable_def intro!: emeasure_mono enn2real_mono)
lemma emeasure_eq_measure2: "A ∈ fmeasurable M ⟹ emeasure M A = measure M A"
by (simp add: emeasure_eq_ennreal_measure fmeasurable_def less_top)
interpretation fmeasurable: ring_of_sets "space M" "fmeasurable M"
proof (rule ring_of_setsI)
show "fmeasurable M ⊆ Pow (space M)" "{} ∈ fmeasurable M"
by (auto simp: fmeasurable_def dest: sets.sets_into_space)
fix a b assume *: "a ∈ fmeasurable M" "b ∈ fmeasurable M"
then have "emeasure M (a ∪ b) ≤ emeasure M a + emeasure M b"
by (intro emeasure_subadditive) auto
also have "… < top"
using * by (auto simp: fmeasurable_def)
finally show "a ∪ b ∈ fmeasurable M"
using * by (auto intro: fmeasurableI)
show "a - b ∈ fmeasurable M"
using emeasure_mono[of "a - b" a M] * by (auto simp: fmeasurable_def)
qed
subsection‹Measurable sets formed by unions and intersections›
lemma fmeasurable_Diff: "A ∈ fmeasurable M ⟹ B ∈ sets M ⟹ A - B ∈ fmeasurable M"
using fmeasurableI2[of A M "A - B"] by auto
lemma fmeasurable_Int_fmeasurable:
"⟦S ∈ fmeasurable M; T ∈ sets M⟧ ⟹ (S ∩ T) ∈ fmeasurable M"
by (meson fmeasurableD fmeasurableI2 inf_le1 sets.Int)
lemma fmeasurable_UN:
assumes "countable I" "⋀i. i ∈ I ⟹ F i ⊆ A" "⋀i. i ∈ I ⟹ F i ∈ sets M" "A ∈ fmeasurable M"
shows "(⋃i∈I. F i) ∈ fmeasurable M"
proof (rule fmeasurableI2)
show "A ∈ fmeasurable M" "(⋃i∈I. F i) ⊆ A" using assms by auto
show "(⋃i∈I. F i) ∈ sets M"
using assms by (intro sets.countable_UN') auto
qed
lemma fmeasurable_INT:
assumes "countable I" "i ∈ I" "⋀i. i ∈ I ⟹ F i ∈ sets M" "F i ∈ fmeasurable M"
shows "(⋂i∈I. F i) ∈ fmeasurable M"
proof (rule fmeasurableI2)
show "F i ∈ fmeasurable M" "(⋂i∈I. F i) ⊆ F i"
using assms by auto
show "(⋂i∈I. F i) ∈ sets M"
using assms by (intro sets.countable_INT') auto
qed
lemma measurable_measure_Diff:
assumes "A ∈ fmeasurable M" "B ∈ sets M" "B ⊆ A"
shows "measure M (A - B) = measure M A - measure M B"
by (simp add: assms fmeasurableD fmeasurableD2 measure_Diff)
lemma measurable_Un_null_set:
assumes "B ∈ null_sets M"
shows "(A ∪ B ∈ fmeasurable M ∧ A ∈ sets M) ⟷ A ∈ fmeasurable M"
using assms by (fastforce simp add: fmeasurable.Un fmeasurableI_null_sets intro: fmeasurableI2)
lemma measurable_Diff_null_set:
assumes "B ∈ null_sets M"
shows "(A - B) ∈ fmeasurable M ∧ A ∈ sets M ⟷ A ∈ fmeasurable M"
using assms
by (metis Un_Diff_cancel2 fmeasurable.Diff fmeasurableD fmeasurableI_null_sets measurable_Un_null_set)
lemma fmeasurable_Diff_D:
assumes m: "T - S ∈ fmeasurable M" "S ∈ fmeasurable M" and sub: "S ⊆ T"
shows "T ∈ fmeasurable M"
proof -
have "T = S ∪ (T - S)"
using assms by blast
then show ?thesis
by (metis m fmeasurable.Un)
qed
lemma measure_Un2:
"A ∈ fmeasurable M ⟹ B ∈ fmeasurable M ⟹ measure M (A ∪ B) = measure M A + measure M (B - A)"
using measure_Union[of M A "B - A"] by (auto simp: fmeasurableD2 fmeasurable.Diff)
lemma measure_Un3:
assumes "A ∈ fmeasurable M" "B ∈ fmeasurable M"
shows "measure M (A ∪ B) = measure M A + measure M B - measure M (A ∩ B)"
proof -
have "measure M (A ∪ B) = measure M A + measure M (B - A)"
using assms by (rule measure_Un2)
also have "B - A = B - (A ∩ B)"
by auto
also have "measure M (B - (A ∩ B)) = measure M B - measure M (A ∩ B)"
using assms by (intro measure_Diff) (auto simp: fmeasurable_def)
finally show ?thesis
by simp
qed
lemma measure_Un_AE:
"AE x in M. x ∉ A ∨ x ∉ B ⟹ A ∈ fmeasurable M ⟹ B ∈ fmeasurable M ⟹
measure M (A ∪ B) = measure M A + measure M B"
by (subst measure_Un2) (auto intro!: measure_eq_AE)
lemma measure_UNION_AE:
assumes I: "finite I"
shows "(⋀i. i ∈ I ⟹ F i ∈ fmeasurable M) ⟹ pairwise (λi j. AE x in M. x ∉ F i ∨ x ∉ F j) I ⟹
measure M (⋃i∈I. F i) = (∑i∈I. measure M (F i))"
unfolding AE_pairwise[OF countable_finite, OF I]
using I
proof (induction I rule: finite_induct)
case (insert x I)
have "measure M (F x ∪ ⋃(F ` I)) = measure M (F x) + measure M (⋃(F ` I))"
by (rule measure_Un_AE) (use insert in ‹auto simp: pairwise_insert›)
with insert show ?case
by (simp add: pairwise_insert )
qed simp
lemma measure_UNION':
"finite I ⟹ (⋀i. i ∈ I ⟹ F i ∈ fmeasurable M) ⟹ pairwise (λi j. disjnt (F i) (F j)) I ⟹
measure M (⋃i∈I. F i) = (∑i∈I. measure M (F i))"
by (intro measure_UNION_AE) (auto simp: disjnt_def elim!: pairwise_mono intro!: always_eventually)
lemma measure_Union_AE:
"finite F ⟹ (⋀S. S ∈ F ⟹ S ∈ fmeasurable M) ⟹ pairwise (λS T. AE x in M. x ∉ S ∨ x ∉ T) F ⟹
measure M (⋃F) = (∑S∈F. measure M S)"
using measure_UNION_AE[of F "λx. x" M] by simp
lemma measure_Union':
"finite F ⟹ (⋀S. S ∈ F ⟹ S ∈ fmeasurable M) ⟹ pairwise disjnt F ⟹ measure M (⋃F) = (∑S∈F. measure M S)"
using measure_UNION'[of F "λx. x" M] by simp
lemma measure_Un_le:
assumes "A ∈ sets M" "B ∈ sets M" shows "measure M (A ∪ B) ≤ measure M A + measure M B"
proof cases
assume "A ∈ fmeasurable M ∧ B ∈ fmeasurable M"
with measure_subadditive[of A M B] assms show ?thesis
by (auto simp: fmeasurableD2)
next
assume "¬ (A ∈ fmeasurable M ∧ B ∈ fmeasurable M)"
then have "A ∪ B ∉ fmeasurable M"
using fmeasurableI2[of "A ∪ B" M A] fmeasurableI2[of "A ∪ B" M B] assms by auto
with assms show ?thesis
by (auto simp: fmeasurable_def measure_def less_top[symmetric])
qed
lemma measure_UNION_le:
"finite I ⟹ (⋀i. i ∈ I ⟹ F i ∈ sets M) ⟹ measure M (⋃i∈I. F i) ≤ (∑i∈I. measure M (F i))"
proof (induction I rule: finite_induct)
case (insert i I)
then have "measure M (⋃i∈insert i I. F i) = measure M (F i ∪ ⋃ (F ` I))"
by simp
also from insert have "measure M (F i ∪ ⋃ (F ` I)) ≤ measure M (F i) + measure M (⋃ (F ` I))"
by (intro measure_Un_le sets.finite_Union) auto
also have "measure M (⋃i∈I. F i) ≤ (∑i∈I. measure M (F i))"
using insert by auto
finally show ?case
using insert by simp
qed simp
lemma measure_Union_le:
"finite F ⟹ (⋀S. S ∈ F ⟹ S ∈ sets M) ⟹ measure M (⋃F) ≤ (∑S∈F. measure M S)"
using measure_UNION_le[of F "λx. x" M] by simp
text‹Version for indexed union over a countable set›
lemma
assumes "countable I" and I: "⋀i. i ∈ I ⟹ A i ∈ fmeasurable M"
and bound: "⋀I'. I' ⊆ I ⟹ finite I' ⟹ measure M (⋃i∈I'. A i) ≤ B"
shows fmeasurable_UN_bound: "(⋃i∈I. A i) ∈ fmeasurable M" (is ?fm)
and measure_UN_bound: "measure M (⋃i∈I. A i) ≤ B" (is ?m)
proof -
have "B ≥ 0"
using bound by force
have "?fm ∧ ?m"
proof cases
assume "I = {}"
with ‹B ≥ 0› show ?thesis
by simp
next
assume "I ≠ {}"
have "(⋃i∈I. A i) = (⋃i. (⋃n≤i. A (from_nat_into I n)))"
by (subst range_from_nat_into[symmetric, OF ‹I ≠ {}› ‹countable I›]) auto
then have "emeasure M (⋃i∈I. A i) = emeasure M (⋃i. (⋃n≤i. A (from_nat_into I n)))" by simp
also have "… = (SUP i. emeasure M (⋃n≤i. A (from_nat_into I n)))"
using I ‹I ≠ {}›[THEN from_nat_into] by (intro SUP_emeasure_incseq[symmetric]) (fastforce simp: incseq_Suc_iff)+
also have "… ≤ B"
proof (intro SUP_least)
fix i :: nat
have "emeasure M (⋃n≤i. A (from_nat_into I n)) = measure M (⋃n≤i. A (from_nat_into I n))"
using I ‹I ≠ {}›[THEN from_nat_into] by (intro emeasure_eq_measure2 fmeasurable.finite_UN) auto
also have "… = measure M (⋃n∈from_nat_into I ` {..i}. A n)"
by simp
also have "… ≤ B"
by (intro ennreal_leI bound) (auto intro: from_nat_into[OF ‹I ≠ {}›])
finally show "emeasure M (⋃n≤i. A (from_nat_into I n)) ≤ ennreal B" .
qed
finally have *: "emeasure M (⋃i∈I. A i) ≤ B" .
then have ?fm
using I ‹countable I› by (intro fmeasurableI conjI) (auto simp: less_top[symmetric] top_unique)
with * ‹0≤B› show ?thesis
by (simp add: emeasure_eq_measure2)
qed
then show ?fm ?m by auto
qed
text‹Version for big union of a countable set›
lemma
assumes "countable 𝒟"
and meas: "⋀D. D ∈ 𝒟 ⟹ D ∈ fmeasurable M"
and bound: "⋀ℰ. ⟦ℰ ⊆ 𝒟; finite ℰ⟧ ⟹ measure M (⋃ℰ) ≤ B"
shows fmeasurable_Union_bound: "⋃𝒟 ∈ fmeasurable M" (is ?fm)
and measure_Union_bound: "measure M (⋃𝒟) ≤ B" (is ?m)
proof -
have "B ≥ 0"
using bound by force
have "?fm ∧ ?m"
proof (cases "𝒟 = {}")
case True
with ‹B ≥ 0› show ?thesis
by auto
next
case False
then obtain D :: "nat ⇒ 'a set" where D: "𝒟 = range D"
using ‹countable 𝒟› uncountable_def by force
have 1: "⋀i. D i ∈ fmeasurable M"
by (simp add: D meas)
have 2: "⋀I'. finite I' ⟹ measure M (⋃x∈I'. D x) ≤ B"
by (simp add: D bound image_subset_iff)
show ?thesis
unfolding D
by (intro conjI fmeasurable_UN_bound [OF _ 1 2] measure_UN_bound [OF _ 1 2]) auto
qed
then show ?fm ?m by auto
qed
text‹Version for indexed union over the type of naturals›
lemma
fixes S :: "nat ⇒ 'a set"
assumes S: "⋀i. S i ∈ fmeasurable M" and B: "⋀n. measure M (⋃i≤n. S i) ≤ B"
shows fmeasurable_countable_Union: "(⋃i. S i) ∈ fmeasurable M"
and measure_countable_Union_le: "measure M (⋃i. S i) ≤ B"
proof -
have mB: "measure M (⋃i∈I. S i) ≤ B" if "finite I" for I
proof -
have "(⋃i∈I. S i) ⊆ (⋃i≤Max I. S i)"
using Max_ge that by force
then have "measure M (⋃i∈I. S i) ≤ measure M (⋃i ≤ Max I. S i)"
by (rule measure_mono_fmeasurable) (use S in ‹blast+›)
then show ?thesis
using B order_trans by blast
qed
show "(⋃i. S i) ∈ fmeasurable M"
by (auto intro: fmeasurable_UN_bound [OF _ S mB])
show "measure M (⋃n. S n) ≤ B"
by (auto intro: measure_UN_bound [OF _ S mB])
qed
lemma measure_diff_le_measure_setdiff:
assumes "S ∈ fmeasurable M" "T ∈ fmeasurable M"
shows "measure M S - measure M T ≤ measure M (S - T)"
proof -
have "measure M S ≤ measure M ((S - T) ∪ T)"
by (simp add: assms fmeasurable.Un fmeasurableD measure_mono_fmeasurable)
also have "… ≤ measure M (S - T) + measure M T"
using assms by (blast intro: measure_Un_le)
finally show ?thesis
by (simp add: algebra_simps)
qed
lemma suminf_exist_split2:
fixes f :: "nat ⇒ 'a::real_normed_vector"
assumes "summable f"
shows "(λn. (∑k. f(k+n))) ⇢ 0"
by (subst lim_sequentially, auto simp add: dist_norm suminf_exist_split[OF _ assms])
lemma emeasure_union_summable:
assumes [measurable]: "⋀n. A n ∈ sets M"
and "⋀n. emeasure M (A n) < ∞" "summable (λn. measure M (A n))"
shows "emeasure M (⋃n. A n) < ∞" "emeasure M (⋃n. A n) ≤ (∑n. measure M (A n))"
proof -
define B where "B = (λN. (⋃n∈{..<N}. A n))"
have [measurable]: "B N ∈ sets M" for N unfolding B_def by auto
have "(λN. emeasure M (B N)) ⇢ emeasure M (⋃N. B N)"
apply (rule Lim_emeasure_incseq) unfolding B_def by (auto simp add: SUP_subset_mono incseq_def)
moreover have "emeasure M (B N) ≤ ennreal (∑n. measure M (A n))" for N
proof -
have *: "(∑n<N. measure M (A n)) ≤ (∑n. measure M (A n))"
using assms(3) measure_nonneg sum_le_suminf by blast
have "emeasure M (B N) ≤ (∑n<N. emeasure M (A n))"
unfolding B_def by (rule emeasure_subadditive_finite, auto)
also have "… = (∑n<N. ennreal(measure M (A n)))"
using assms(2) by (simp add: emeasure_eq_ennreal_measure less_top)
also have "… = ennreal (∑n<N. measure M (A n))"
by auto
also have "… ≤ ennreal (∑n. measure M (A n))"
using * by (auto simp: ennreal_leI)
finally show ?thesis by simp
qed
ultimately have "emeasure M (⋃N. B N) ≤ ennreal (∑n. measure M (A n))"
by (simp add: Lim_bounded)
then show "emeasure M (⋃n. A n) ≤ (∑n. measure M (A n))"
unfolding B_def by (metis UN_UN_flatten UN_lessThan_UNIV)
then show "emeasure M (⋃n. A n) < ∞"
by (auto simp: less_top[symmetric] top_unique)
qed
lemma borel_cantelli_limsup1:
assumes [measurable]: "⋀n. A n ∈ sets M"
and "⋀n. emeasure M (A n) < ∞" "summable (λn. measure M (A n))"
shows "limsup A ∈ null_sets M"
proof -
have "emeasure M (limsup A) ≤ 0"
proof (rule LIMSEQ_le_const)
have "(λn. (∑k. measure M (A (k+n)))) ⇢ 0" by (rule suminf_exist_split2[OF assms(3)])
then show "(λn. ennreal (∑k. measure M (A (k+n)))) ⇢ 0"
unfolding ennreal_0[symmetric] by (intro tendsto_ennrealI)
have "emeasure M (limsup A) ≤ (∑k. measure M (A (k+n)))" for n
proof -
have I: "(⋃k∈{n..}. A k) = (⋃k. A (k+n))" by (auto, metis le_add_diff_inverse2, fastforce)
have "emeasure M (limsup A) ≤ emeasure M (⋃k∈{n..}. A k)"
by (rule emeasure_mono, auto simp add: limsup_INF_SUP)
also have "… = emeasure M (⋃k. A (k+n))"
using I by auto
also have "… ≤ (∑k. measure M (A (k+n)))"
apply (rule emeasure_union_summable)
using assms summable_ignore_initial_segment[OF assms(3), of n] by auto
finally show ?thesis by simp
qed
then show "∃N. ∀n≥N. emeasure M (limsup A) ≤ (∑k. measure M (A (k+n)))"
by auto
qed
then show ?thesis using assms(1) measurable_limsup by auto
qed
lemma borel_cantelli_AE1:
assumes [measurable]: "⋀n. A n ∈ sets M"
and "⋀n. emeasure M (A n) < ∞" "summable (λn. measure M (A n))"
shows "AE x in M. eventually (λn. x ∈ space M - A n) sequentially"
proof -
have "AE x in M. x ∉ limsup A"
using borel_cantelli_limsup1[OF assms] unfolding eventually_ae_filter by auto
moreover have "∀⇩F n in sequentially. x ∉ A n" if "x ∉ limsup A" for x
using that by (auto simp: limsup_INF_SUP eventually_sequentially)
ultimately show ?thesis by auto
qed
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) fmeasurable_eq_sets: "fmeasurable M = sets M"
by (auto simp: fmeasurable_def less_top[symmetric])
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_sum_singleton:
assumes "finite S" and *: "⋀x. x ∈ S ⟹ {x} ∈ sets M"
shows "measure M S = (∑x∈S. measure M {x})"
using measure_eq_sum_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_sum_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 sum_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 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
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 "(⋀n. f n < f (Suc n))" shows "strict_mono f"
by (simp add: assms strict_mono_Suc_iff)
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 obtain i where i: "∀j≥i. F i = F j" ..
with ‹incseq F› have "F j ⊆ F i" for j
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_transform_eventually[OF tendsto_const] 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 obtain k where "of_nat n ≤ ?M (F k)" ..
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 (simp only:) simp
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 from X 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 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 "A ∈ sets M")
case True
show ?thesis
proof (rule emeasure_measure_of[OF restrict_space_def])
show "(∩) Ω ` 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
case False
with assms have "A ∉ sets (restrict_space M Ω)"
by (simp add: sets_restrict_space_iff)
with False 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 = "(∩) 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: SUP_cong_simp)
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 :: "'a measure ⇒ 'a measure" where
"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 ‹Complete lattice structure on measures›
lemma (in finite_measure) finite_measure_Diff':
"A ∈ sets M ⟹ B ∈ sets M ⟹ measure M (A - B) = measure M A - measure M (A ∩ B)"
using finite_measure_Diff[of A "A ∩ B"] by (auto simp: Diff_Int)
lemma (in finite_measure) finite_measure_Union':
"A ∈ sets M ⟹ B ∈ sets M ⟹ measure M (A ∪ B) = measure M A + measure M (B - A)"
using finite_measure_Union[of A "B - A"] by auto
lemma finite_unsigned_Hahn_decomposition:
assumes "finite_measure M" "finite_measure N" and [simp]: "sets N = sets M"
shows "∃Y∈sets M. (∀X∈sets M. X ⊆ Y ⟶ N X ≤ M X) ∧ (∀X∈sets M. X ∩ Y = {} ⟶ M X ≤ N X)"
proof -
interpret M: finite_measure M by fact
interpret N: finite_measure N by fact
define d where "d X = measure M X - measure N X" for X
have [intro]: "bdd_above (d`sets M)"
using sets.sets_into_space[of _ M]
by (intro bdd_aboveI[where M="measure M (space M)"])
(auto simp: d_def field_simps subset_eq intro!: add_increasing M.finite_measure_mono)
define γ where "γ = (SUP X∈sets M. d X)"
have le_γ[intro]: "X ∈ sets M ⟹ d X ≤ γ" for X
by (auto simp: γ_def intro!: cSUP_upper)
have "∃f. ∀n. f n ∈ sets M ∧ d (f n) > γ - 1 / 2^n"
proof (intro choice_iff[THEN iffD1] allI)
fix n
have "∃X∈sets M. γ - 1 / 2^n < d X"
unfolding γ_def by (intro less_cSUP_iff[THEN iffD1]) auto
then show "∃y. y ∈ sets M ∧ γ - 1 / 2 ^ n < d y"
by auto
qed
then obtain E where [measurable]: "E n ∈ sets M" and E: "d (E n) > γ - 1 / 2^n" for n
by auto
define F where "F m n = (if m ≤ n then ⋂i∈{m..n}. E i else space M)" for m n
have [measurable]: "m ≤ n ⟹ F m n ∈ sets M" for m n
by (auto simp: F_def)
have 1: "γ - 2 / 2 ^ m + 1 / 2 ^ n ≤ d (F m n)" if "m ≤ n" for m n
using that
proof (induct rule: dec_induct)
case base with E[of m] show ?case
by (simp add: F_def field_simps)
next
case (step i)
have F_Suc: "F m (Suc i) = F m i ∩ E (Suc i)"
using ‹m ≤ i› by (auto simp: F_def le_Suc_eq)
have "γ + (γ - 2 / 2^m + 1 / 2 ^ Suc i) ≤ (γ - 1 / 2^Suc i) + (γ - 2 / 2^m + 1 / 2^i)"
by (simp add: field_simps)
also have "… ≤ d (E (Suc i)) + d (F m i)"
using E[of "Suc i"] by (intro add_mono step) auto
also have "… = d (E (Suc i)) + d (F m i - E (Suc i)) + d (F m (Suc i))"
using ‹m ≤ i› by (simp add: d_def field_simps F_Suc M.finite_measure_Diff' N.finite_measure_Diff')
also have "… = d (E (Suc i) ∪ F m i) + d (F m (Suc i))"
using ‹m ≤ i› by (simp add: d_def field_simps M.finite_measure_Union' N.finite_measure_Union')
also have "… ≤ γ + d (F m (Suc i))"
using ‹m ≤ i› by auto
finally show ?case
by auto
qed
define F' where "F' m = (⋂i∈{m..}. E i)" for m
have F'_eq: "F' m = (⋂i. F m (i + m))" for m
by (fastforce simp: le_iff_add[of m] F'_def F_def)
have [measurable]: "F' m ∈ sets M" for m
by (auto simp: F'_def)
have γ_le: "γ - 0 ≤ d (⋃m. F' m)"
proof (rule LIMSEQ_le)
show "(λn. γ - 2 / 2 ^ n) ⇢ γ - 0"
by (intro tendsto_intros LIMSEQ_divide_realpow_zero) auto
have "incseq F'"
by (auto simp: incseq_def F'_def)
then show "(λm. d (F' m)) ⇢ d (⋃m. F' m)"
unfolding d_def
by (intro tendsto_diff M.finite_Lim_measure_incseq N.finite_Lim_measure_incseq) auto
have "γ - 2 / 2 ^ m + 0 ≤ d (F' m)" for m
proof (rule LIMSEQ_le)
have *: "decseq (λn. F m (n + m))"
by (auto simp: decseq_def F_def)
show "(λn. d (F m n)) ⇢ d (F' m)"
unfolding d_def F'_eq
by (rule LIMSEQ_offset[where k=m])
(auto intro!: tendsto_diff M.finite_Lim_measure_decseq N.finite_Lim_measure_decseq *)
show "(λn. γ - 2 / 2 ^ m + 1 / 2 ^ n) ⇢ γ - 2 / 2 ^ m + 0"
by (intro tendsto_add LIMSEQ_divide_realpow_zero tendsto_const) auto
show "∃N. ∀n≥N. γ - 2 / 2 ^ m + 1 / 2 ^ n ≤ d (F m n)"
using 1[of m] by (intro exI[of _ m]) auto
qed
then show "∃N. ∀n≥N. γ - 2 / 2 ^ n ≤ d (F' n)"
by auto
qed
show ?thesis
proof (safe intro!: bexI[of _ "⋃m. F' m"])
fix X assume [measurable]: "X ∈ sets M" and X: "X ⊆ (⋃m. F' m)"
have "d (⋃m. F' m) - d X = d ((⋃m. F' m) - X)"
using X by (auto simp: d_def M.finite_measure_Diff N.finite_measure_Diff)
also have "… ≤ γ"
by auto
finally have "0 ≤ d X"
using γ_le by auto
then show "emeasure N X ≤ emeasure M X"
by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure)
next
fix X assume [measurable]: "X ∈ sets M" and X: "X ∩ (⋃m. F' m) = {}"
then have "d (⋃m. F' m) + d X = d (X ∪ (⋃m. F' m))"
by (auto simp: d_def M.finite_measure_Union N.finite_measure_Union)
also have "… ≤ γ"
by auto
finally have "d X ≤ 0"
using γ_le by auto
then show "emeasure M X ≤ emeasure N X"
by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure)
qed auto
qed
proposition unsigned_Hahn_decomposition:
assumes [simp]: "sets N = sets M" and [measurable]: "A ∈ sets M"
and [simp]: "emeasure M A ≠ top" "emeasure N A ≠ top"
shows "∃Y∈sets M. Y ⊆ A ∧ (∀X∈sets M. X ⊆ Y ⟶ N X ≤ M X) ∧ (∀X∈sets M. X ⊆ A ⟶ X ∩ Y = {} ⟶ M X ≤ N X)"
proof -
have "∃Y∈sets (restrict_space M A).
(∀X∈sets (restrict_space M A). X ⊆ Y ⟶ (restrict_space N A) X ≤ (restrict_space M A) X) ∧
(∀X∈sets (restrict_space M A). X ∩ Y = {} ⟶ (restrict_space M A) X ≤ (restrict_space N A) X)"
proof (rule finite_unsigned_Hahn_decomposition)
show "finite_measure (restrict_space M A)" "finite_measure (restrict_space N A)"
by (auto simp: space_restrict_space emeasure_restrict_space less_top intro!: finite_measureI)
qed (simp add: sets_restrict_space)
with assms show ?thesis
by (metis Int_subset_iff emeasure_restrict_space sets.Int_space_eq2 sets_restrict_space_iff space_restrict_space)
qed
text ‹
Define a lexicographical order on \<^type>‹measure›, in the order space, sets and measure. The parts
of the lexicographical order are point-wise ordered.
›
instantiation measure :: (type) order_bot
begin
inductive less_eq_measure :: "'a measure ⇒ 'a measure ⇒ bool" where
"space M ⊂ space N ⟹ less_eq_measure M N"
| "space M = space N ⟹ sets M ⊂ sets N ⟹ less_eq_measure M N"
| "space M = space N ⟹ sets M = sets N ⟹ emeasure M ≤ emeasure N ⟹ less_eq_measure M N"
lemma le_measure_iff:
"M ≤ N ⟷ (if space M = space N then
if sets M = sets N then emeasure M ≤ emeasure N else sets M ⊆ sets N else space M ⊆ space N)"
by (auto elim: less_eq_measure.cases intro: less_eq_measure.intros)
definition less_measure :: "'a measure ⇒ 'a measure ⇒ bool" where
"less_measure M N ⟷ (M ≤ N ∧ ¬ N ≤ M)"
definition bot_measure :: "'a measure" where
"bot_measure = sigma {} {}"
lemma
shows space_bot[simp]: "space bot = {}"
and sets_bot[simp]: "sets bot = {{}}"
and emeasure_bot[simp]: "emeasure bot X = 0"
by (auto simp: bot_measure_def sigma_sets_empty_eq emeasure_sigma)
instance
proof standard
show "bot ≤ a" for a :: "'a measure"
by (simp add: le_measure_iff bot_measure_def sigma_sets_empty_eq emeasure_sigma le_fun_def)
qed (auto simp: le_measure_iff less_measure_def split: if_split_asm intro: measure_eqI)
end
proposition le_measure: "sets M = sets N ⟹ M ≤ N ⟷ (∀A∈sets M. emeasure M A ≤ emeasure N A)"
by (metis emeasure_neq_0_sets le_fun_def le_measure_iff order_class.order_eq_iff sets_eq_imp_space_eq)
definition sup_measure' :: "'a measure ⇒ 'a measure ⇒ 'a measure" where
"sup_measure' A B =
measure_of (space A) (sets A)
(λX. SUP Y∈sets A. emeasure A (X ∩ Y) + emeasure B (X ∩ - Y))"
lemma assumes [simp]: "sets B = sets A"
shows space_sup_measure'[simp]: "space (sup_measure' A B) = space A"
and sets_sup_measure'[simp]: "sets (sup_measure' A B) = sets A"
using sets_eq_imp_space_eq[OF assms] by (simp_all add: sup_measure'_def)
lemma emeasure_sup_measure':
assumes sets_eq[simp]: "sets B = sets A" and [simp, intro]: "X ∈ sets A"
shows "emeasure (sup_measure' A B) X = (SUP Y∈sets A. emeasure A (X ∩ Y) + emeasure B (X ∩ - Y))"
(is "_ = ?S X")
proof -
note sets_eq_imp_space_eq[OF sets_eq, simp]
show ?thesis
using sup_measure'_def
proof (rule emeasure_measure_of)
let ?d = "λX Y. emeasure A (X ∩ Y) + emeasure B (X ∩ - Y)"
show "countably_additive (sets (sup_measure' A B)) (λX. SUP Y ∈ sets A. emeasure A (X ∩ Y) + emeasure B (X ∩ - Y))"
proof (rule countably_additiveI, goal_cases)
case (1 X)
then have [measurable]: "⋀i. X i ∈ sets A" and "disjoint_family X"
by auto
have disjoint: "disjoint_family (λi. X i ∩ Y)" "disjoint_family (λi. X i - Y)" for Y
using "1"(2) disjoint_family_subset by fastforce+
have "(∑i. ?S (X i)) = (SUP Y∈sets A. ∑i. ?d (X i) Y)"
proof (rule ennreal_suminf_SUP_eq_directed)
fix J :: "nat set" and a b assume "finite J" and [measurable]: "a ∈ sets A" "b ∈ sets A"
have "∃c∈sets A. c ⊆ X i ∧ (∀a∈sets A. ?d (X i) a ≤ ?d (X i) c)" for i
proof cases
assume "emeasure A (X i) = top ∨ emeasure B (X i) = top"
then show ?thesis
by force
next
assume finite: "¬ (emeasure A (X i) = top ∨ emeasure B (X i) = top)"
then have "∃Y∈sets A. Y ⊆ X i ∧ (∀C∈sets A. C ⊆ Y ⟶ B C ≤ A C) ∧ (∀C∈sets A. C ⊆ X i ⟶ C ∩ Y = {} ⟶ A C ≤ B C)"
using unsigned_Hahn_decomposition[of B A "X i"] by simp
then obtain Y where [measurable]: "Y ∈ sets A" and [simp]: "Y ⊆ X i"
and B_le_A: "⋀C. C ∈ sets A ⟹ C ⊆ Y ⟹ B C ≤ A C"
and A_le_B: "⋀C. C ∈ sets A ⟹ C ⊆ X i ⟹ C ∩ Y = {} ⟹ A C ≤ B C"
by auto
show ?thesis
proof (intro bexI ballI conjI)
fix a assume [measurable]: "a ∈ sets A"
have *: "(X i ∩ a ∩ Y ∪ (X i ∩ a - Y)) = X i ∩ a" "(X i - a) ∩ Y ∪ (X i - a - Y) = X i ∩ - a"
for a Y by auto
then have "?d (X i) a =
(A (X i ∩ a ∩ Y) + A (X i ∩ a ∩ - Y)) + (B (X i ∩ - a ∩ Y) + B (X i ∩ - a ∩ - Y))"
by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric])
also have "… ≤ (A (X i ∩ a ∩ Y) + B (X i ∩ a ∩ - Y)) + (A (X i ∩ - a ∩ Y) + B (X i ∩ - a ∩ - Y))"
by (intro add_mono order_refl B_le_A A_le_B) (auto simp: Diff_eq[symmetric])
also have "… ≤ (A (X i ∩ Y ∩ a) + A (X i ∩ Y ∩ - a)) + (B (X i ∩ - Y ∩ a) + B (X i ∩ - Y ∩ - a))"
by (simp add: ac_simps)
also have "… ≤ A (X i ∩ Y) + B (X i ∩ - Y)"
by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric] *)
finally show "?d (X i) a ≤ ?d (X i) Y" .
qed auto
qed
then obtain C where [measurable]: "C i ∈ sets A" and "C i ⊆ X i"
and C: "⋀a. a ∈ sets A ⟹ ?d (X i) a ≤ ?d (X i) (C i)" for i
by metis
have *: "X i ∩ (⋃i. C i) = X i ∩ C i" for i
using ‹disjoint_family X› ‹⋀i. C i ⊆ X i›
by (simp add: disjoint_family_on_def disjoint_iff_not_equal set_eq_iff) (metis subsetD)
then have **: "X i ∩ - (⋃i. C i) = X i ∩ - C i" for i by blast
moreover have "(⋃i. C i) ∈ sets A"
by fastforce
ultimately show "∃c∈sets A. ∀i∈J. ?d (X i) a ≤ ?d (X i) c ∧ ?d (X i) b ≤ ?d (X i) c"
by (metis "*" C ‹a ∈ sets A› ‹b ∈ sets A›)
qed
also have "… = ?S (⋃i. X i)"
proof -
have "⋀Y. Y ∈ sets A ⟹ (∑i. emeasure A (X i ∩ Y) + emeasure B (X i ∩ -Y))
= emeasure A (⋃i. X i ∩ Y) + emeasure B (⋃i. X i ∩ -Y)"
using disjoint
by (auto simp flip: suminf_add Diff_eq simp add: image_subset_iff suminf_emeasure)
then show ?thesis by force
qed
finally show "(∑i. ?S (X i)) = ?S (⋃i. X i)" .
qed
qed (auto dest: sets.sets_into_space simp: positive_def intro!: SUP_const)
qed
lemma le_emeasure_sup_measure'1:
assumes "sets B = sets A" "X ∈ sets A" shows "emeasure A X ≤ emeasure (sup_measure' A B) X"
by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "X"] assms)
lemma le_emeasure_sup_measure'2:
assumes "sets B = sets A" "X ∈ sets A" shows "emeasure B X ≤ emeasure (sup_measure' A B) X"
by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "{}"] assms)
lemma emeasure_sup_measure'_le2:
assumes [simp]: "sets B = sets C" "sets A = sets C" and [measurable]: "X ∈ sets C"
assumes A: "⋀Y. Y ⊆ X ⟹ Y ∈ sets A ⟹ emeasure A Y ≤ emeasure C Y"
assumes B: "⋀Y. Y ⊆ X ⟹ Y ∈ sets A ⟹ emeasure B Y ≤ emeasure C Y"
shows "emeasure (sup_measure' A B) X ≤ emeasure C X"
proof (subst emeasure_sup_measure')
show "(SUP Y∈sets A. emeasure A (X ∩ Y) + emeasure B (X ∩ - Y)) ≤ emeasure C X"
unfolding ‹sets A = sets C›
proof (intro SUP_least)
fix Y assume [measurable]: "Y ∈ sets C"
have [simp]: "X ∩ Y ∪ (X - Y) = X"
by auto
have "emeasure A (X ∩ Y) + emeasure B (X ∩ - Y) ≤ emeasure C (X ∩ Y) + emeasure C (X ∩ - Y)"
by (intro add_mono A B) (auto simp: Diff_eq[symmetric])
also have "… = emeasure C X"
by (subst plus_emeasure) (auto simp: Diff_eq[symmetric])
finally show "emeasure A (X ∩ Y) + emeasure B (X ∩ - Y) ≤ emeasure C X" .
qed
qed simp_all
definition sup_lexord :: "'a ⇒ 'a ⇒ ('a ⇒ 'b::order) ⇒ 'a ⇒ 'a ⇒ 'a" where
"sup_lexord A B k s c =
(if k A = k B then c else
if ¬ k A ≤ k B ∧ ¬ k B ≤ k A then s else
if k B ≤ k A then A else B)"
lemma sup_lexord:
"(k A < k B ⟹ P B) ⟹ (k B < k A ⟹ P A) ⟹ (k A = k B ⟹ P c) ⟹
(¬ k B ≤ k A ⟹ ¬ k A ≤ k B ⟹ P s) ⟹ P (sup_lexord A B k s c)"
by (auto simp: sup_lexord_def)
lemmas le_sup_lexord = sup_lexord[where P="λa. c ≤ a" for c]
lemma sup_lexord1: "k A = k B ⟹ sup_lexord A B k s c = c"
by (simp add: sup_lexord_def)
lemma sup_lexord_commute: "sup_lexord A B k s c = sup_lexord B A k s c"
by (auto simp: sup_lexord_def)
lemma sigma_sets_le_sets_iff: "(sigma_sets (space x) 𝒜 ⊆ sets x) = (𝒜 ⊆ sets x)"
using sets.sigma_sets_subset[of 𝒜 x] by auto
lemma sigma_le_iff: "𝒜 ⊆ Pow Ω ⟹ sigma Ω 𝒜 ≤ x ⟷ (Ω ⊆ space x ∧ (space x = Ω ⟶ 𝒜 ⊆ sets x))"
by (cases "Ω = space x")
(simp_all add: eq_commute[of _ "sets x"] le_measure_iff emeasure_sigma le_fun_def
sigma_sets_superset_generator sigma_sets_le_sets_iff)
instantiation measure :: (type) semilattice_sup
begin
definition sup_measure :: "'a measure ⇒ 'a measure ⇒ 'a measure" where
"sup_measure A B =
sup_lexord A B space (sigma (space A ∪ space B) {})
(sup_lexord A B sets (sigma (space A) (sets A ∪ sets B)) (sup_measure' A B))"
instance
proof
fix x y z :: "'a measure"
show "x ≤ sup x y"
unfolding sup_measure_def
proof (intro le_sup_lexord)
assume "space x = space y"
then have *: "sets x ∪ sets y ⊆ Pow (space x)"
using sets.space_closed by auto
assume "¬ sets y ⊆ sets x" "¬ sets x ⊆ sets y"
then have "sets x ⊂ sets x ∪ sets y"
by auto
also have "… ≤ sigma (space x) (sets x ∪ sets y)"
by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator)
finally show "x ≤ sigma (space x) (sets x ∪ sets y)"
by (simp add: space_measure_of[OF *] less_eq_measure.intros(2))
next
assume "¬ space y ⊆ space x" "¬ space x ⊆ space y"
then show "x ≤ sigma (space x ∪ space y) {}"
by (intro less_eq_measure.intros) auto
next
assume "sets x = sets y" then show "x ≤ sup_measure' x y"
by (simp add: le_measure le_emeasure_sup_measure'1)
qed (auto intro: less_eq_measure.intros)
show "y ≤ sup x y"
unfolding sup_measure_def
proof (intro le_sup_lexord)
assume **: "space x = space y"
then have *: "sets x ∪ sets y ⊆ Pow (space y)"
using sets.space_closed by auto
assume "¬ sets y ⊆ sets x" "¬ sets x ⊆ sets y"
then have "sets y ⊂ sets x ∪ sets y"
by auto
also have "… ≤ sigma (space y) (sets x ∪ sets y)"
by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator)
finally show "y ≤ sigma (space x) (sets x ∪ sets y)"
by (simp add: ** space_measure_of[OF *] less_eq_measure.intros(2))
next
assume "¬ space y ⊆ space x" "¬ space x ⊆ space y"
then show "y ≤ sigma (space x ∪ space y) {}"
by (intro less_eq_measure.intros) auto
next
assume "sets x = sets y" then show "y ≤ sup_measure' x y"
by (simp add: le_measure le_emeasure_sup_measure'2)
qed (auto intro: less_eq_measure.intros)
show "x ≤ y ⟹ z ≤ y ⟹ sup x z ≤ y"
unfolding sup_measure_def
proof (intro sup_lexord[where P="λx. x ≤ y"])
assume "x ≤ y" "z ≤ y" and [simp]: "space x = space z" "sets x = sets z"
from ‹x ≤ y› show "sup_measure' x z ≤ y"
proof cases
case 1 then show ?thesis
by (intro less_eq_measure.intros(1)) simp
next
case 2 then show ?thesis
by (intro less_eq_measure.intros(2)) simp_all
next
case 3 with ‹z ≤ y› ‹x ≤ y› show ?thesis
by (auto simp add: le_measure intro!: emeasure_sup_measure'_le2)
qed
next
assume **: "x ≤ y" "z ≤ y" "space x = space z" "¬ sets z ⊆ sets x" "¬ sets x ⊆ sets z"
then have *: "sets x ∪ sets z ⊆ Pow (space x)"
using sets.space_closed by auto
show "sigma (space x) (sets x ∪ sets z) ≤ y"
unfolding sigma_le_iff[OF *] using ** by (auto simp: le_measure_iff split: if_split_asm)
next
assume "x ≤ y" "z ≤ y" "¬ space z ⊆ space x" "¬ space x ⊆ space z"
then have "space x ⊆ space y" "space z ⊆ space y"
by (auto simp: le_measure_iff split: if_split_asm)
then show "sigma (space x ∪ space z) {} ≤ y"
by (simp add: sigma_le_iff)
qed
qed
end
lemma space_empty_eq_bot: "space a = {} ⟷ a = bot"
using space_empty[of a] by (auto intro!: measure_eqI)
lemma sets_eq_iff_bounded: "A ≤ B ⟹ B ≤ C ⟹ sets A = sets C ⟹ sets B = sets A"
by (auto dest: sets_eq_imp_space_eq simp add: le_measure_iff split: if_split_asm)
lemma sets_sup: "sets A = sets M ⟹ sets B = sets M ⟹ sets (sup A B) = sets M"
by (auto simp add: sup_measure_def sup_lexord_def dest: sets_eq_imp_space_eq)
lemma le_measureD1: "A ≤ B ⟹ space A ≤ space B"
by (auto simp: le_measure_iff split: if_split_asm)
lemma le_measureD2: "A ≤ B ⟹ space A = space B ⟹ sets A ≤ sets B"
by (auto simp: le_measure_iff split: if_split_asm)
lemma le_measureD3: "A ≤ B ⟹ sets A = sets B ⟹ emeasure A X ≤ emeasure B X"
by (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq split: if_split_asm)
lemma UN_space_closed: "⋃(sets ` S) ⊆ Pow (⋃(space ` S))"
using sets.space_closed by auto
definition
Sup_lexord :: "('a ⇒ 'b::complete_lattice) ⇒ ('a set ⇒ 'a) ⇒ ('a set ⇒ 'a) ⇒ 'a set ⇒ 'a"
where
"Sup_lexord k c s A =
(let U = (SUP a∈A. k a)
in if ∃a∈A. k a = U then c {a∈A. k a = U} else s A)"
lemma Sup_lexord:
"(⋀a S. a ∈ A ⟹ k a = (SUP a∈A. k a) ⟹ S = {a'∈A. k a' = k a} ⟹ P (c S)) ⟹ ((⋀a. a ∈ A ⟹ k a ≠ (SUP a∈A. k a)) ⟹ P (s A)) ⟹
P (Sup_lexord k c s A)"
by (auto simp: Sup_lexord_def Let_def)
lemma Sup_lexord1:
assumes A: "A ≠ {}" "(⋀a. a ∈ A ⟹ k a = (⋃a∈A. k a))" "P (c A)"
shows "P (Sup_lexord k c s A)"
unfolding Sup_lexord_def Let_def
proof (clarsimp, safe)
show "∀a∈A. k a ≠ (⋃x∈A. k x) ⟹ P (s A)"
by (metis assms(1,2) ex_in_conv)
next
fix a assume "a ∈ A" "k a = (⋃x∈A. k x)"
then have "{a ∈ A. k a = (⋃x∈A. k x)} = {a ∈ A. k a = k a}"
by (metis A(2)[symmetric])
then show "P (c {a ∈ A. k a = (⋃x∈A. k x)})"
by (simp add: A(3))
qed
instantiation measure :: (type) complete_lattice
begin
interpretation sup_measure: comm_monoid_set sup "bot :: 'a measure"
by standard (auto intro!: antisym)
lemma sup_measure_F_mono':
"finite J ⟹ finite I ⟹ sup_measure.F id I ≤ sup_measure.F id (I ∪ J)"
proof (induction J rule: finite_induct)
case empty then show ?case
by simp
next
case (insert i J)
show ?case
proof cases
assume "i ∈ I" with insert show ?thesis
by (auto simp: insert_absorb)
next
assume "i ∉ I"
have "sup_measure.F id I ≤ sup_measure.F id (I ∪ J)"
by (intro insert)
also have "… ≤ sup_measure.F id (insert i (I ∪ J))"
using insert ‹i ∉ I› by (subst sup_measure.insert) auto
finally show ?thesis
by auto
qed
qed
lemma sup_measure_F_mono: "finite I ⟹ J ⊆ I ⟹ sup_measure.F id J ≤ sup_measure.F id I"
using sup_measure_F_mono'[of I J] by (auto simp: finite_subset Un_absorb1)
lemma sets_sup_measure_F:
"finite I ⟹ I ≠ {} ⟹ (⋀i. i ∈ I ⟹ sets i = sets M) ⟹ sets (sup_measure.F id I) = sets M"
by (induction I rule: finite_ne_induct) (simp_all add: sets_sup)
definition Sup_measure' :: "'a measure set ⇒ 'a measure" where
"Sup_measure' M =
measure_of (⋃a∈M. space a) (⋃a∈M. sets a)
(λX. (SUP P∈{P. finite P ∧ P ⊆ M }. sup_measure.F id P X))"
lemma space_Sup_measure'2: "space (Sup_measure' M) = (⋃m∈M. space m)"
unfolding Sup_measure'_def by (intro space_measure_of[OF UN_space_closed])
lemma sets_Sup_measure'2: "sets (Sup_measure' M) = sigma_sets (⋃m∈M. space m) (⋃m∈M. sets m)"
unfolding Sup_measure'_def by (intro sets_measure_of[OF UN_space_closed])
lemma sets_Sup_measure':
assumes sets_eq[simp]: "⋀m. m ∈ M ⟹ sets m = sets A" and "M ≠ {}"
shows "sets (Sup_measure' M) = sets A"
using sets_eq[THEN sets_eq_imp_space_eq, simp] ‹M ≠ {}› by (simp add: Sup_measure'_def)
lemma space_Sup_measure':
assumes sets_eq[simp]: "⋀m. m ∈ M ⟹ sets m = sets A" and "M ≠ {}"
shows "space (Sup_measure' M) = space A"
using sets_eq[THEN sets_eq_imp_space_eq, simp] ‹M ≠ {}›
by (simp add: Sup_measure'_def )
lemma emeasure_Sup_measure':
assumes sets_eq[simp]: "⋀m. m ∈ M ⟹ sets m = sets A" and "X ∈ sets A" "M ≠ {}"
shows "emeasure (Sup_measure' M) X = (SUP P∈{P. finite P ∧ P ⊆ M}. sup_measure.F id P X)"
(is "_ = ?S X")
using Sup_measure'_def
proof (rule emeasure_measure_of)
note sets_eq[THEN sets_eq_imp_space_eq, simp]
have *: "sets (Sup_measure' M) = sets A" "space (Sup_measure' M) = space A"
using ‹M ≠ {}› by (simp_all add: Sup_measure'_def)
let ?μ = "sup_measure.F id"
show "countably_additive (sets (Sup_measure' M)) ?S"
proof (rule countably_additiveI, goal_cases)
case (1 F)
then have **: "range F ⊆ sets A"
by (auto simp: *)
show "(∑i. ?S (F i)) = ?S (⋃i. F i)"
proof (subst ennreal_suminf_SUP_eq_directed)
fix i j and N :: "nat set" assume ij: "i ∈ {P. finite P ∧ P ⊆ M}" "j ∈ {P. finite P ∧ P ⊆ M}"
have "(i ≠ {} ⟶ sets (?μ i) = sets A) ∧ (j ≠ {} ⟶ sets (?μ j) = sets A) ∧
(i ≠ {} ∨ j ≠ {} ⟶ sets (?μ (i ∪ j)) = sets A)"
using ij by (intro impI sets_sup_measure_F conjI) auto
then have "?μ j (F n) ≤ ?μ (i ∪ j) (F n) ∧ ?μ i (F n) ≤ ?μ (i ∪ j) (F n)" for n
using ij
by (cases "i = {}"; cases "j = {}")
(auto intro!: le_measureD3 sup_measure_F_mono simp: sets_sup_measure_F
simp del: id_apply)
with ij show "∃k∈{P. finite P ∧ P ⊆ M}. ∀n∈N. ?μ i (F n) ≤ ?μ k (F n) ∧ ?μ j (F n) ≤ ?μ k (F n)"
by (safe intro!: bexI[of _ "i ∪ j"]) auto
next
show "(SUP P ∈ {P. finite P ∧ P ⊆ M}. ∑n. ?μ P (F n)) = (SUP P ∈ {P. finite P ∧ P ⊆ M}. ?μ P (⋃(F ` UNIV)))"
proof (intro arg_cong [of _ _ Sup] image_cong refl)
fix i assume i: "i ∈ {P. finite P ∧ P ⊆ M}"
show "(∑n. ?μ i (F n)) = ?μ i (⋃(F ` UNIV))"
proof cases
assume "i ≠ {}" with i ** show ?thesis
by (smt (verit, best) "1"(2) Measure_Space.sets_sup_measure_F assms(1) mem_Collect_eq subset_eq suminf_cong suminf_emeasure)
qed simp
qed
qed
qed
show "positive (sets (Sup_measure' M)) ?S"
by (auto simp: positive_def bot_ennreal[symmetric])
show "X ∈ sets (Sup_measure' M)"
using assms * by auto
qed (rule UN_space_closed)
definition Sup_measure :: "'a measure set ⇒ 'a measure" where
"Sup_measure =
Sup_lexord space
(Sup_lexord sets Sup_measure'
(λU. sigma (⋃u∈U. space u) (⋃u∈U. sets u)))
(λU. sigma (⋃u∈U. space u) {})"
definition Inf_measure :: "'a measure set ⇒ 'a measure" where
"Inf_measure A = Sup {x. ∀a∈A. x ≤ a}"
definition inf_measure :: "'a measure ⇒ 'a measure ⇒ 'a measure" where
"inf_measure a b = Inf {a, b}"
definition top_measure :: "'a measure" where
"top_measure = Inf {}"
instance
proof
note UN_space_closed [simp]
show upper: "x ≤ Sup A" if x: "x ∈ A" for x :: "'a measure" and A
unfolding Sup_measure_def
proof (intro Sup_lexord[where P="λy. x ≤ y"])
assume "⋀a. a ∈ A ⟹ space a ≠ (⋃a∈A. space a)"
from this[OF ‹x ∈ A›] ‹x ∈ A› show "x ≤ sigma (⋃a∈A. space a) {}"
by (intro less_eq_measure.intros) auto
next
fix a S assume "a ∈ A" and a: "space a = (⋃a∈A. space a)" and S: "S = {a' ∈ A. space a' = space a}"
and neq: "⋀aa. aa ∈ S ⟹ sets aa ≠ (⋃a∈S. sets a)"
have sp_a: "space a = (⋃(space ` S))"
using ‹a∈A› by (auto simp: S)
show "x ≤ sigma (⋃(space ` S)) (⋃(sets ` S))"
proof cases
assume [simp]: "space x = space a"
have "sets x ⊂ (⋃a∈S. sets a)"
using ‹x∈A› neq[of x] by (auto simp: S)
also have "… ⊆ sigma_sets (⋃x∈S. space x) (⋃x∈S. sets x)"
by (rule sigma_sets_superset_generator)
finally show ?thesis
by (intro less_eq_measure.intros(2)) (simp_all add: sp_a)
next
assume "space x ≠ space a"
moreover have "space x ≤ space a"
unfolding a using ‹x∈A› by auto
ultimately show ?thesis
by (intro less_eq_measure.intros) (simp add: less_le sp_a)
qed
next
fix a b S S' assume "a ∈ A" and a: "space a = (⋃a∈A. space a)" and S: "S = {a' ∈ A. space a' = space a}"
and "b ∈ S" and b: "sets b = (⋃a∈S. sets a)" and S': "S' = {a' ∈ S. sets a' = sets b}"
then have "S' ≠ {}" "space b = space a"
by auto
have sets_eq: "⋀x. x ∈ S' ⟹ sets x = sets b"
by (auto simp: S')
note sets_eq[THEN sets_eq_imp_space_eq, simp]
have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b"
using ‹S' ≠ {}› by (simp_all add: Sup_measure'_def sets_eq)
show "x ≤ Sup_measure' S'"
proof cases
assume "x ∈ S"
with ‹b ∈ S› have "space x = space b"
by (simp add: S)
show ?thesis
proof cases
assume "x ∈ S'"
show "x ≤ Sup_measure' S'"
proof (intro le_measure[THEN iffD2] ballI)
show "sets x = sets (Sup_measure' S')"
using ‹x∈S'› * by (simp add: S')
fix X assume "X ∈ sets x"
show "emeasure x X ≤ emeasure (Sup_measure' S') X"
proof (subst emeasure_Sup_measure'[OF _ ‹X ∈ sets x›])
show "emeasure x X ≤ (SUP P ∈ {P. finite P ∧ P ⊆ S'}. emeasure (sup_measure.F id P) X)"
using ‹x∈S'› by (intro SUP_upper2[where i="{x}"]) auto
qed (insert ‹x∈S'› S', auto)
qed
next
assume "x ∉ S'"
then have "sets x ≠ sets b"
using ‹x∈S› by (auto simp: S')
moreover have "sets x ≤ sets b"
using ‹x∈S› unfolding b by auto
ultimately show ?thesis
using * ‹x ∈ S›
by (intro less_eq_measure.intros(2))
(simp_all add: * ‹space x = space b› less_le)
qed
next
assume "x ∉ S"
with ‹x∈A› ‹x ∉ S› ‹space b = space a› show ?thesis
by (intro less_eq_measure.intros)
(simp_all add: * less_le a SUP_upper S)
qed
qed
show least: "Sup A ≤ x" if x: "⋀z. z ∈ A ⟹ z ≤ x" for x :: "'a measure" and A
unfolding Sup_measure_def
proof (intro Sup_lexord[where P="λy. y ≤ x"])
assume "⋀a. a ∈ A ⟹ space a ≠ (⋃a∈A. space a)"
show "sigma (⋃(space ` A)) {} ≤ x"
using x[THEN le_measureD1] by (subst sigma_le_iff) auto
next
fix a S assume "a ∈ A" "space a = (⋃a∈A. space a)" and S: "S = {a' ∈ A. space a' = space a}"
"⋀a. a ∈ S ⟹ sets a ≠ (⋃a∈S. sets a)"
have "⋃(space ` S) ⊆ space x"
using S le_measureD1[OF x] by auto
moreover
have "⋃(space ` S) = space a"
using ‹a∈A› S by auto
then have "space x = ⋃(space ` S) ⟹ ⋃(sets ` S) ⊆ sets x"
using ‹a ∈ A› le_measureD2[OF x] by (auto simp: S)
ultimately show "sigma (⋃(space ` S)) (⋃(sets ` S)) ≤ x"
by (subst sigma_le_iff) simp_all
next
fix a b S S' assume "a ∈ A" and a: "space a = (⋃a∈A. space a)" and S: "S = {a' ∈ A. space a' = space a}"
and "b ∈ S" and b: "sets b = (⋃a∈S. sets a)" and S': "S' = {a' ∈ S. sets a' = sets b}"
then have "S' ≠ {}" "space b = space a"
by auto
have sets_eq: "⋀x. x ∈ S' ⟹ sets x = sets b"
by (auto simp: S')
note sets_eq[THEN sets_eq_imp_space_eq, simp]
have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b"
using ‹S' ≠ {}› by (simp_all add: Sup_measure'_def sets_eq)
show "Sup_measure' S' ≤ x"
proof cases
assume "space x = space a"
show ?thesis
proof cases
assume **: "sets x = sets b"
show ?thesis
proof (intro le_measure[THEN iffD2] ballI)
show ***: "sets (Sup_measure' S') = sets x"
by (simp add: * **)
fix X assume "X ∈ sets (Sup_measure' S')"
show "emeasure (Sup_measure' S') X ≤ emeasure x X"
unfolding ***
proof (subst emeasure_Sup_measure'[OF _ ‹X ∈ sets (Sup_measure' S')›])
show "(SUP P ∈ {P. finite P ∧ P ⊆ S'}. emeasure (sup_measure.F id P) X) ≤ emeasure x X"
proof (safe intro!: SUP_least)
fix P assume P: "finite P" "P ⊆ S'"
show "emeasure (sup_measure.F id P) X ≤ emeasure x X"
proof cases
assume "P = {}" then show ?thesis
by auto
next
assume "P ≠ {}"
from P have "finite P" "P ⊆ A"
unfolding S' S by (simp_all add: subset_eq)
then have "sup_measure.F id P ≤ x"
by (induction P) (auto simp: x)
moreover have "sets (sup_measure.F id P) = sets x"
using ‹finite P› ‹P ≠ {}› ‹P ⊆ S'› ‹sets x = sets b›
by (intro sets_sup_measure_F) (auto simp: S')
ultimately show "emeasure (sup_measure.F id P) X ≤ emeasure x X"
by (rule le_measureD3)
qed
qed
show "m ∈ S' ⟹ sets m = sets (Sup_measure' S')" for m
unfolding * by (simp add: S')
qed fact
qed
next
assume "sets x ≠ sets b"
moreover have "sets b ≤ sets x"
unfolding b S using x[THEN le_measureD2] ‹space x = space a› by auto
ultimately show "Sup_measure' S' ≤ x"
using ‹space x = space a› ‹b ∈ S›
by (intro less_eq_measure.intros(2)) (simp_all add: * S)
qed
next
assume "space x ≠ space a"
then have "space a < space x"
using le_measureD1[OF x[OF ‹a∈A›]] by auto
then show "Sup_measure' S' ≤ x"
by (intro less_eq_measure.intros) (simp add: * ‹space b = space a›)
qed
qed
show "Sup {} = (bot::'a measure)" "Inf {} = (top::'a measure)"
by (auto intro!: antisym least simp: top_measure_def)
show lower: "x ∈ A ⟹ Inf A ≤ x" for x :: "'a measure" and A
unfolding Inf_measure_def by (intro least) auto
show greatest: "(⋀z. z ∈ A ⟹ x ≤ z) ⟹ x ≤ Inf A" for x :: "'a measure" and A
unfolding Inf_measure_def by (intro upper) auto
show "inf x y ≤ x" "inf x y ≤ y" "x ≤ y ⟹ x ≤ z ⟹ x ≤ inf y z" for x y z :: "'a measure"
by (auto simp: inf_measure_def intro!: lower greatest)
qed
end
lemma sets_SUP:
assumes "⋀x. x ∈ I ⟹ sets (M x) = sets N"
shows "I ≠ {} ⟹ sets (SUP i∈I. M i) = sets N"
unfolding Sup_measure_def
using assms assms[THEN sets_eq_imp_space_eq]
sets_Sup_measure'[where A=N and M="M`I"]
by (intro Sup_lexord1[where P="λx. sets x = sets N"]) auto
lemma emeasure_SUP:
assumes sets: "⋀i. i ∈ I ⟹ sets (M i) = sets N" "X ∈ sets N" "I ≠ {}"
shows "emeasure (SUP i∈I. M i) X = (SUP J∈{J. J ≠ {} ∧ finite J ∧ J ⊆ I}. emeasure (SUP i∈J. M i) X)"
proof -
interpret sup_measure: comm_monoid_set sup "bot :: 'b measure"
by standard (auto intro!: antisym)
have eq: "finite J ⟹ sup_measure.F id J = (SUP i∈J. i)" for J :: "'b measure set"
by (induction J rule: finite_induct) auto
have 1: "J ≠ {} ⟹ J ⊆ I ⟹ sets (SUP x∈J. M x) = sets N" for J
by (intro sets_SUP sets) (auto )
from ‹I ≠ {}› obtain i where "i∈I" by auto
have "Sup_measure' (M`I) X = (SUP P∈{P. finite P ∧ P ⊆ M`I}. sup_measure.F id P X)"
using sets by (intro emeasure_Sup_measure') auto
also have "Sup_measure' (M`I) = (SUP i∈I. M i)"
unfolding Sup_measure_def using ‹I ≠ {}› sets sets(1)[THEN sets_eq_imp_space_eq]
by (intro Sup_lexord1[where P="λx. _ = x"]) auto
also have "(SUP P∈{P. finite P ∧ P ⊆ M`I}. sup_measure.F id P X) =
(SUP J∈{J. J ≠ {} ∧ finite J ∧ J ⊆ I}. (SUP i∈J. M i) X)"
proof (intro SUP_eq)
fix J assume "J ∈ {P. finite P ∧ P ⊆ M`I}"
then obtain J' where J': "J' ⊆ I" "finite J'" and J: "J = M`J'" and "finite J"
using finite_subset_image[of J M I] by auto
show "∃j∈{J. J ≠ {} ∧ finite J ∧ J ⊆ I}. sup_measure.F id J X ≤ (SUP i∈j. M i) X"
proof cases
assume "J' = {}" with ‹i ∈ I› show ?thesis
by (auto simp add: J)
next
assume "J' ≠ {}" with J J' show ?thesis
by (intro bexI[of _ "J'"]) (auto simp add: eq simp del: id_apply)
qed
next
fix J assume J: "J ∈ {P. P ≠ {} ∧ finite P ∧ P ⊆ I}"
show "∃J'∈{J. finite J ∧ J ⊆ M`I}. (SUP i∈J. M i) X ≤ sup_measure.F id J' X"
using J by (intro bexI[of _ "M`J"]) (auto simp add: eq simp del: id_apply)
qed
finally show ?thesis .
qed
lemma emeasure_SUP_chain:
assumes sets: "⋀i. i ∈ A ⟹ sets (M i) = sets N" "X ∈ sets N"
assumes ch: "Complete_Partial_Order.chain (≤) (M ` A)" and "A ≠ {}"
shows "emeasure (SUP i∈A. M i) X = (SUP i∈A. emeasure (M i) X)"
proof (subst emeasure_SUP[OF sets ‹A ≠ {}›])
show "(SUP J∈{J. J ≠ {} ∧ finite J ∧ J ⊆ A}. emeasure (Sup (M ` J)) X) = (SUP i∈A. emeasure (M i) X)"
proof (rule SUP_eq)
fix J assume "J ∈ {J. J ≠ {} ∧ finite J ∧ J ⊆ A}"
then have J: "Complete_Partial_Order.chain (≤) (M ` J)" "finite J" "J ≠ {}" and "J ⊆ A"
using ch[THEN chain_subset, of "M`J"] by auto
with in_chain_finite[OF J(1)] obtain j where "j ∈ J" "(SUP j∈J. M j) = M j"
by auto
with ‹J ⊆ A› show "∃j∈A. emeasure (Sup (M ` J)) X ≤ emeasure (M j) X"
by auto
next
fix j assume "j∈A" then show "∃i∈{J. J ≠ {} ∧ finite J ∧ J ⊆ A}. emeasure (M j) X ≤ emeasure (Sup (M ` i)) X"
by (intro bexI[of _ "{j}"]) auto
qed
qed
subsubsection ‹Supremum of a set of ‹σ›-algebras›
lemma space_Sup_eq_UN: "space (Sup M) = (⋃x∈M. space x)" (is "?L=?R")
proof
show "?L ⊆ ?R"
using Sup_lexord[where P="λx. space x = _"]
apply (clarsimp simp: Sup_measure_def)
by (smt (verit) Sup_lexord_def UN_E mem_Collect_eq space_Sup_measure'2 space_measure_of_conv)
qed (use Sup_upper le_measureD1 in fastforce)
lemma sets_Sup_eq:
assumes *: "⋀m. m ∈ M ⟹ space m = X" and "M ≠ {}"
shows "sets (Sup M) = sigma_sets X (⋃x∈M. sets x)"
unfolding Sup_measure_def
proof (rule Sup_lexord1 [OF ‹M ≠ {}›])
show "sets (Sup_lexord sets Sup_measure' (λU. sigma (⋃ (space ` U)) (⋃ (sets ` U))) M)
= sigma_sets X (⋃ (sets ` M))"
apply (rule Sup_lexord)
apply (metis (mono_tags, lifting) "*" empty_iff mem_Collect_eq sets.sigma_sets_eq sets_Sup_measure')
by (metis "*" SUP_eq_const UN_space_closed assms(2) sets_measure_of)
qed (use * in blast)
lemma in_sets_Sup: "(⋀m. m ∈ M ⟹ space m = X) ⟹ m ∈ M ⟹ A ∈ sets m ⟹ A ∈ sets (Sup M)"
by (subst sets_Sup_eq[where X=X]) auto
lemma Sup_lexord_rel:
assumes "⋀i. i ∈ I ⟹ k (A i) = k (B i)"
"R (c (A ` {a ∈ I. k (B a) = (SUP x∈I. k (B x))})) (c (B ` {a ∈ I. k (B a) = (SUP x∈I. k (B x))}))"
"R (s (A`I)) (s (B`I))"
shows "R (Sup_lexord k c s (A`I)) (Sup_lexord k c s (B`I))"
proof -
have "A ` {a ∈ I. k (B a) = (SUP x∈I. k (B x))} = {a ∈ A ` I. k a = (SUP x∈I. k (B x))}"
using assms(1) by auto
moreover have "B ` {a ∈ I. k (B a) = (SUP x∈I. k (B x))} = {a ∈ B ` I. k a = (SUP x∈I. k (B x))}"
by auto
ultimately show ?thesis
using assms by (auto simp: Sup_lexord_def Let_def image_comp)
qed
lemma sets_SUP_cong:
assumes eq: "⋀i. i ∈ I ⟹ sets (M i) = sets (N i)"
shows "sets (SUP i∈I. M i) = sets (SUP i∈I. N i)"
unfolding Sup_measure_def
using eq eq[THEN sets_eq_imp_space_eq]
by (intro Sup_lexord_rel[where R="λx y. sets x = sets y"], simp_all add: sets_Sup_measure'2)
lemma sets_Sup_in_sets:
assumes "M ≠ {}"
assumes "⋀m. m ∈ M ⟹ space m = space N"
assumes "⋀m. m ∈ M ⟹ sets m ⊆ sets N"
shows "sets (Sup M) ⊆ sets N"
proof -
have *: "⋃(space ` M) = space N"
using assms by auto
show ?thesis
unfolding * using assms by (subst sets_Sup_eq[of M "space N"]) (auto intro!: sets.sigma_sets_subset)
qed
lemma measurable_Sup1:
assumes m: "m ∈ M" and f: "f ∈ measurable m N"
and const_space: "⋀m n. m ∈ M ⟹ n ∈ M ⟹ space m = space n"
shows "f ∈ measurable (Sup M) N"
proof -
have "space (Sup M) = space m"
using m by (auto simp add: space_Sup_eq_UN dest: const_space)
then show ?thesis
using m f unfolding measurable_def by (auto intro: in_sets_Sup[OF const_space])
qed
lemma measurable_Sup2:
assumes M: "M ≠ {}"
assumes f: "⋀m. m ∈ M ⟹ f ∈ measurable N m"
and const_space: "⋀m n. m ∈ M ⟹ n ∈ M ⟹ space m = space n"
shows "f ∈ measurable N (Sup M)"
proof -
from M obtain m where "m ∈ M" by auto
have space_eq: "⋀n. n ∈ M ⟹ space n = space m"
by (intro const_space ‹m ∈ M›)
have eq: "sets (sigma (⋃ (space ` M)) (⋃ (sets ` M))) = sets (Sup M)"
by (metis M SUP_eq_const UN_space_closed sets_Sup_eq sets_measure_of space_eq)
have "f ∈ measurable N (sigma (⋃m∈M. space m) (⋃m∈M. sets m))"
proof (rule measurable_measure_of)
show "f ∈ space N → ⋃(space ` M)"
using measurable_space[OF f] M by auto
qed (auto intro: measurable_sets f dest: sets.sets_into_space)
also have "measurable N (sigma (⋃m∈M. space m) (⋃m∈M. sets m)) = measurable N (Sup M)"
using eq measurable_cong_sets by blast
finally show ?thesis .
qed
lemma measurable_SUP2:
"I ≠ {} ⟹ (⋀i. i ∈ I ⟹ f ∈ measurable N (M i)) ⟹
(⋀i j. i ∈ I ⟹ j ∈ I ⟹ space (M i) = space (M j)) ⟹ f ∈ measurable N (SUP i∈I. M i)"
by (auto intro!: measurable_Sup2)
lemma sets_Sup_sigma:
assumes [simp]: "M ≠ {}" and M: "⋀m. m ∈ M ⟹ m ⊆ Pow Ω"
shows "sets (SUP m∈M. sigma Ω m) = sets (sigma Ω (⋃M))"
proof -
{ fix a m assume "a ∈ sigma_sets Ω m" "m ∈ M"
then have "a ∈ sigma_sets Ω (⋃M)"
by induction (auto intro: sigma_sets.intros(2-)) }
then have "sigma_sets Ω (⋃ (sigma_sets Ω ` M)) = sigma_sets Ω (⋃ M)"
by (smt (verit, best) UN_iff Union_iff sigma_sets.Basic sigma_sets_eqI)
then show "sets (SUP m∈M. sigma Ω m) = sets (sigma Ω (⋃M))"
by (subst sets_Sup_eq) (fastforce simp add: M Union_least)+
qed
lemma Sup_sigma:
assumes [simp]: "M ≠ {}" and M: "⋀m. m ∈ M ⟹ m ⊆ Pow Ω"
shows "(SUP m∈M. sigma Ω m) = (sigma Ω (⋃M))"
proof (intro antisym SUP_least)
have *: "⋃M ⊆ Pow Ω"
using M by auto
show "sigma Ω (⋃M) ≤ (SUP m∈M. sigma Ω m)"
proof (intro less_eq_measure.intros(3))
show "space (sigma Ω (⋃M)) = space (SUP m∈M. sigma Ω m)"
"sets (sigma Ω (⋃M)) = sets (SUP m∈M. sigma Ω m)"
by (auto simp add: M sets_Sup_sigma sets_eq_imp_space_eq space_measure_of_conv)
qed (simp add: emeasure_sigma le_fun_def)
fix m assume "m ∈ M" then show "sigma Ω m ≤ sigma Ω (⋃M)"
by (subst sigma_le_iff) (auto simp add: M *)
qed
lemma SUP_sigma_sigma:
"M ≠ {} ⟹ (⋀m. m ∈ M ⟹ f m ⊆ Pow Ω) ⟹ (SUP m∈M. sigma Ω (f m)) = sigma Ω (⋃m∈M. f m)"
using Sup_sigma[of "f`M" Ω] by (auto simp: image_comp)
lemma sets_vimage_Sup_eq:
assumes *: "M ≠ {}" "f ∈ X → Y" "⋀m. m ∈ M ⟹ space m = Y"
shows "sets (vimage_algebra X f (Sup M)) = sets (SUP m ∈ M. vimage_algebra X f m)"
(is "?L = ?R")
proof
have "⋀m. m ∈ M ⟹ f ∈ Sup (vimage_algebra X f ` M) →⇩M m"
using assms
by (smt (verit, del_insts) Pi_iff imageE image_eqI measurable_Sup1
measurable_vimage_algebra1 space_vimage_algebra)
then show "?L ⊆ ?R"
by (intro sets_image_in_sets measurable_Sup2) (simp_all add: space_Sup_eq_UN *)
show "?R ⊆ ?L"
apply (intro sets_Sup_in_sets)
apply (force simp add: * space_Sup_eq_UN sets_vimage_algebra2 intro: in_sets_Sup)+
done
qed
lemma restrict_space_eq_vimage_algebra':
"sets (restrict_space M Ω) = sets (vimage_algebra (Ω ∩ space M) (λx. x) M)"
proof -
have *: "{A ∩ (Ω ∩ space M) |A. A ∈ sets M} = {A ∩ Ω |A. A ∈ sets M}"
using sets.sets_into_space[of _ M] by blast
show ?thesis
unfolding restrict_space_def
by (subst sets_measure_of)
(auto simp add: image_subset_iff sets_vimage_algebra * dest: sets.sets_into_space intro!: arg_cong2[where f=sigma_sets])
qed
lemma sigma_le_sets:
assumes [simp]: "A ⊆ Pow X" shows "sets (sigma X A) ⊆ sets N ⟷ X ∈ sets N ∧ A ⊆ sets N"
proof
have "X ∈ sigma_sets X A" "A ⊆ sigma_sets X A"
by (auto intro: sigma_sets_top)
moreover assume "sets (sigma X A) ⊆ sets N"
ultimately show "X ∈ sets N ∧ A ⊆ sets N"
by auto
next
assume *: "X ∈ sets N ∧ A ⊆ sets N"
{ fix Y assume "Y ∈ sigma_sets X A" from this * have "Y ∈ sets N"
by induction auto }
then show "sets (sigma X A) ⊆ sets N"
by auto
qed
lemma measurable_iff_sets:
"f ∈ measurable M N ⟷ (f ∈ space M → space N ∧ sets (vimage_algebra (space M) f N) ⊆ sets M)"
unfolding measurable_def
by (smt (verit, ccfv_threshold) mem_Collect_eq sets_vimage_algebra sigma_sets_le_sets_iff subset_eq)
lemma sets_vimage_algebra_space: "X ∈ sets (vimage_algebra X f M)"
using sets.top[of "vimage_algebra X f M"] by simp
lemma measurable_mono:
assumes N: "sets N' ≤ sets N" "space N = space N'"
assumes M: "sets M ≤ sets M'" "space M = space M'"
shows "measurable M N ⊆ measurable M' N'"
unfolding measurable_def
proof safe
fix f A assume "f ∈ space M → space N" "A ∈ sets N'"
moreover assume "∀y∈sets N. f -` y ∩ space M ∈ sets M" note this[THEN bspec, of A]
ultimately show "f -` A ∩ space M' ∈ sets M'"
using assms by auto
qed (use N M in auto)
lemma measurable_Sup_measurable:
assumes f: "f ∈ space N → A"
shows "f ∈ measurable N (Sup {M. space M = A ∧ f ∈ measurable N M})"
proof (rule measurable_Sup2)
show "{M. space M = A ∧ f ∈ measurable N M} ≠ {}"
using f unfolding ex_in_conv[symmetric]
by (intro exI[of _ "sigma A {}"]) (auto intro!: measurable_measure_of)
qed auto
lemma (in sigma_algebra) sigma_sets_subset':
assumes a: "a ⊆ M" "Ω' ∈ M"
shows "sigma_sets Ω' a ⊆ M"
proof
show "x ∈ M" if x: "x ∈ sigma_sets Ω' a" for x
using x by (induct rule: sigma_sets.induct) (use a in auto)
qed
lemma in_sets_SUP: "i ∈ I ⟹ (⋀i. i ∈ I ⟹ space (M i) = Y) ⟹ X ∈ sets (M i) ⟹ X ∈ sets (SUP i∈I. M i)"
by (intro in_sets_Sup[where X=Y]) auto
lemma measurable_SUP1:
"i ∈ I ⟹ f ∈ measurable (M i) N ⟹ (⋀m n. m ∈ I ⟹ n ∈ I ⟹ space (M m) = space (M n)) ⟹
f ∈ measurable (SUP i∈I. M i) N"
by (auto intro: measurable_Sup1)
lemma sets_image_in_sets':
assumes X: "X ∈ sets N"
assumes f: "⋀A. A ∈ sets M ⟹ f -` A ∩ X ∈ sets N"
shows "sets (vimage_algebra X f M) ⊆ sets N"
unfolding sets_vimage_algebra
by (rule sets.sigma_sets_subset') (auto intro!: measurable_sets X f)
lemma mono_vimage_algebra:
"sets M ≤ sets N ⟹ sets (vimage_algebra X f M) ⊆ sets (vimage_algebra X f N)"
using sets.top[of "sigma X {f -` A ∩ X |A. A ∈ sets N}"]
unfolding vimage_algebra_def
by (smt (verit, del_insts) space_measure_of sigma_le_sets Pow_iff inf_le2 mem_Collect_eq subset_eq)
lemma mono_restrict_space: "sets M ≤ sets N ⟹ sets (restrict_space M X) ⊆ sets (restrict_space N X)"
unfolding sets_restrict_space by (rule image_mono)
lemma sets_eq_bot: "sets M = {{}} ⟷ M = bot"
by (metis measure_eqI emeasure_empty sets_bot singletonD)
lemma sets_eq_bot2: "{{}} = sets M ⟷ M = bot"
using sets_eq_bot[of M] by blast
lemma (in finite_measure) countable_support:
"countable {x. measure M {x} ≠ 0}"
proof cases
assume "measure M (space M) = 0"
then show ?thesis
by (metis (mono_tags, lifting) bounded_measure measure_le_0_iff Collect_empty_eq countable_empty)
next
let ?M = "measure M (space M)" and ?m = "λx. measure M {x}"
assume "?M ≠ 0"
then have *: "{x. ?m x ≠ 0} = (⋃n. {x. ?M / Suc n < ?m x})"
using reals_Archimedean[of "?m x / ?M" for x]
by (auto simp: field_simps not_le[symmetric] divide_le_0_iff measure_le_0_iff)
have **: "⋀n. finite {x. ?M / Suc n < ?m x}"
proof (rule ccontr)
fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X")
then obtain X where "finite X" "card X = Suc (Suc n)" "X ⊆ ?X"
by (metis infinite_arbitrarily_large)
then have *: "⋀x. x ∈ X ⟹ ?M / Suc n ≤ ?m x"
by auto
{ fix x assume "x ∈ X"
from ‹?M ≠ 0› *[OF this] have "?m x ≠ 0" by (auto simp: field_simps measure_le_0_iff)
then have "{x} ∈ sets M" by (auto dest: measure_notin_sets) }
note singleton_sets = this
have "?M < (∑x∈X. ?M / Suc n)"
using ‹?M ≠ 0›
by (simp add: ‹card X = Suc (Suc n)› field_simps less_le)
also have "… ≤ (∑x∈X. ?m x)"
by (rule sum_mono) fact
also have "… = measure M (⋃x∈X. {x})"
using singleton_sets ‹finite X›
by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def)
finally have "?M < measure M (⋃x∈X. {x})" .
moreover have "measure M (⋃x∈X. {x}) ≤ ?M"
using singleton_sets[THEN sets.sets_into_space] by (intro finite_measure_mono) auto
ultimately show False by simp
qed
show ?thesis
unfolding * by (intro countable_UN countableI_type countable_finite[OF **])
qed
end