Theory Set_Integral
chapter ‹Integrals over a Set›
theory Set_Integral
imports Radon_Nikodym
begin
section ‹Notation›
definition "set_borel_measurable M A f ≡ (λx. indicator A x *⇩R f x) ∈ borel_measurable M"
definition "set_integrable M A f ≡ integrable M (λx. indicator A x *⇩R f x)"
definition "set_lebesgue_integral M A f ≡ lebesgue_integral M (λx. indicator A x *⇩R f x)"
syntax
"_ascii_set_lebesgue_integral" :: "pttrn ⇒ 'a set ⇒ 'a measure ⇒ real ⇒ real"
("(4LINT (_):(_)/|(_)./ _)" [0,0,0,10] 10)
translations
"LINT x:A|M. f" == "CONST set_lebesgue_integral M A (λx. f)"
syntax
"_lebesgue_borel_integral" :: "pttrn ⇒ real ⇒ real"
("(2LBINT _./ _)" [0,10] 10)
syntax
"_set_lebesgue_borel_integral" :: "pttrn ⇒ real set ⇒ real ⇒ real"
("(3LBINT _:_./ _)" [0,0,10] 10)
section ‹Basic properties›
lemma set_integrable_cong:
assumes "M = M'" "A = A'" "⋀x. x ∈ A ⟹ f x = f' x"
shows "set_integrable M A f = set_integrable M' A' f'"
proof -
have "(λx. indicator A x *⇩R f x) = (λx. indicator A' x *⇩R f' x)"
using assms by (auto simp: indicator_def of_bool_def)
thus ?thesis by (simp add: set_integrable_def assms)
qed
lemma set_borel_measurable_sets:
fixes f :: "_ ⇒ _::real_normed_vector"
assumes "set_borel_measurable M X f" "B ∈ sets borel" "X ∈ sets M"
shows "f -` B ∩ X ∈ sets M"
proof -
have "f ∈ borel_measurable (restrict_space M X)"
using assms unfolding set_borel_measurable_def by (subst borel_measurable_restrict_space_iff) auto
then have "f -` B ∩ space (restrict_space M X) ∈ sets (restrict_space M X)"
by (rule measurable_sets) fact
with ‹X ∈ sets M› show ?thesis
by (subst (asm) sets_restrict_space_iff) (auto simp: space_restrict_space)
qed
lemma set_integrable_bound:
fixes f :: "'a ⇒ 'b::{banach, second_countable_topology}"
and g :: "'a ⇒ 'c::{banach, second_countable_topology}"
assumes "set_integrable M A f" "set_borel_measurable M A g"
assumes "AE x in M. x ∈ A ⟶ norm (g x) ≤ norm (f x)"
shows "set_integrable M A g"
unfolding set_integrable_def
proof (rule Bochner_Integration.integrable_bound)
from assms(1) show "integrable M (λx. indicator A x *⇩R f x)"
by (simp add: set_integrable_def)
from assms(2) show "(λx. indicat_real A x *⇩R g x) ∈ borel_measurable M"
by (simp add: set_borel_measurable_def)
from assms(3) show "AE x in M. norm (indicat_real A x *⇩R g x) ≤ norm (indicat_real A x *⇩R f x)"
by eventually_elim (simp add: indicator_def)
qed
lemma set_lebesgue_integral_zero [simp]: "set_lebesgue_integral M A (λx. 0) = 0"
by (auto simp: set_lebesgue_integral_def)
lemma set_lebesgue_integral_cong:
assumes "A ∈ sets M" and "∀x. x ∈ A ⟶ f x = g x"
shows "(LINT x:A|M. f x) = (LINT x:A|M. g x)"
unfolding set_lebesgue_integral_def
using assms
by (metis indicator_simps(2) real_vector.scale_zero_left)
lemma set_lebesgue_integral_cong_AE:
assumes [measurable]: "A ∈ sets M" "f ∈ borel_measurable M" "g ∈ borel_measurable M"
assumes "AE x ∈ A in M. f x = g x"
shows "(LINT x:A|M. f x) = (LINT x:A|M. g x)"
proof-
have "AE x in M. indicator A x *⇩R f x = indicator A x *⇩R g x"
using assms by auto
thus ?thesis
unfolding set_lebesgue_integral_def by (intro integral_cong_AE) auto
qed
lemma set_integrable_cong_AE:
"f ∈ borel_measurable M ⟹ g ∈ borel_measurable M ⟹
AE x ∈ A in M. f x = g x ⟹ A ∈ sets M ⟹
set_integrable M A f = set_integrable M A g"
unfolding set_integrable_def
by (rule integrable_cong_AE) auto
lemma set_integrable_subset:
fixes M A B and f :: "_ ⇒ _ :: {banach, second_countable_topology}"
assumes "set_integrable M A f" "B ∈ sets M" "B ⊆ A"
shows "set_integrable M B f"
proof -
have "set_integrable M B (λx. indicator A x *⇩R f x)"
using assms integrable_mult_indicator set_integrable_def by blast
with ‹B ⊆ A› show ?thesis
unfolding set_integrable_def
by (simp add: indicator_inter_arith[symmetric] Int_absorb2)
qed
lemma set_integrable_restrict_space:
fixes f :: "'a ⇒ 'b::{banach, second_countable_topology}"
assumes f: "set_integrable M S f" and T: "T ∈ sets (restrict_space M S)"
shows "set_integrable M T f"
proof -
obtain T' where T_eq: "T = S ∩ T'" and "T' ∈ sets M"
using T by (auto simp: sets_restrict_space)
have ‹integrable M (λx. indicator T' x *⇩R (indicator S x *⇩R f x))›
using ‹T' ∈ sets M› f integrable_mult_indicator set_integrable_def by blast
then show ?thesis
unfolding set_integrable_def
unfolding T_eq indicator_inter_arith by (simp add: ac_simps)
qed
lemma set_integral_scaleR_left:
assumes "A ∈ sets M" "c ≠ 0 ⟹ integrable M f"
shows "(LINT t:A|M. f t *⇩R c) = (LINT t:A|M. f t) *⇩R c"
unfolding set_lebesgue_integral_def
using integrable_mult_indicator[OF assms]
by (subst integral_scaleR_left[symmetric], auto)
lemma set_integral_scaleR_right [simp]: "(LINT t:A|M. a *⇩R f t) = a *⇩R (LINT t:A|M. f t)"
unfolding set_lebesgue_integral_def
by (subst integral_scaleR_right[symmetric]) (auto intro!: Bochner_Integration.integral_cong)
lemma set_integral_mult_right [simp]:
fixes a :: "'a::{real_normed_field, second_countable_topology}"
shows "(LINT t:A|M. a * f t) = a * (LINT t:A|M. f t)"
unfolding set_lebesgue_integral_def
by (subst integral_mult_right_zero[symmetric]) auto
lemma set_integral_mult_left [simp]:
fixes a :: "'a::{real_normed_field, second_countable_topology}"
shows "(LINT t:A|M. f t * a) = (LINT t:A|M. f t) * a"
unfolding set_lebesgue_integral_def
by (subst integral_mult_left_zero[symmetric]) auto
lemma set_integral_divide_zero [simp]:
fixes a :: "'a::{real_normed_field, field, second_countable_topology}"
shows "(LINT t:A|M. f t / a) = (LINT t:A|M. f t) / a"
unfolding set_lebesgue_integral_def
by (subst integral_divide_zero[symmetric], intro Bochner_Integration.integral_cong)
(auto split: split_indicator)
lemma set_integrable_scaleR_right [simp, intro]:
shows "(a ≠ 0 ⟹ set_integrable M A f) ⟹ set_integrable M A (λt. a *⇩R f t)"
unfolding set_integrable_def
unfolding scaleR_left_commute by (rule integrable_scaleR_right)
lemma set_integrable_scaleR_left [simp, intro]:
fixes a :: "_ :: {banach, second_countable_topology}"
shows "(a ≠ 0 ⟹ set_integrable M A f) ⟹ set_integrable M A (λt. f t *⇩R a)"
unfolding set_integrable_def
using integrable_scaleR_left[of a M "λx. indicator A x *⇩R f x"] by simp
lemma set_integrable_mult_right [simp, intro]:
fixes a :: "'a::{real_normed_field, second_countable_topology}"
shows "(a ≠ 0 ⟹ set_integrable M A f) ⟹ set_integrable M A (λt. a * f t)"
unfolding set_integrable_def
using integrable_mult_right[of a M "λx. indicator A x *⇩R f x"] by simp
lemma set_integrable_mult_right_iff [simp]:
fixes a :: "'a::{real_normed_field, second_countable_topology}"
assumes "a ≠ 0"
shows "set_integrable M A (λt. a * f t) ⟷ set_integrable M A f"
proof
assume "set_integrable M A (λt. a * f t)"
then have "set_integrable M A (λt. 1/a * (a * f t))"
using set_integrable_mult_right by blast
then show "set_integrable M A f"
using assms by auto
qed auto
lemma set_integrable_mult_left [simp, intro]:
fixes a :: "'a::{real_normed_field, second_countable_topology}"
shows "(a ≠ 0 ⟹ set_integrable M A f) ⟹ set_integrable M A (λt. f t * a)"
unfolding set_integrable_def
using integrable_mult_left[of a M "λx. indicator A x *⇩R f x"] by simp
lemma set_integrable_mult_left_iff [simp]:
fixes a :: "'a::{real_normed_field, second_countable_topology}"
assumes "a ≠ 0"
shows "set_integrable M A (λt. f t * a) ⟷ set_integrable M A f"
using assms by (subst set_integrable_mult_right_iff [symmetric]) (auto simp: mult.commute)
lemma set_integrable_divide [simp, intro]:
fixes a :: "'a::{real_normed_field, field, second_countable_topology}"
assumes "a ≠ 0 ⟹ set_integrable M A f"
shows "set_integrable M A (λt. f t / a)"
proof -
have "integrable M (λx. indicator A x *⇩R f x / a)"
using assms unfolding set_integrable_def by (rule integrable_divide_zero)
also have "(λx. indicator A x *⇩R f x / a) = (λx. indicator A x *⇩R (f x / a))"
by (auto split: split_indicator)
finally show ?thesis
unfolding set_integrable_def .
qed
lemma set_integrable_mult_divide_iff [simp]:
fixes a :: "'a::{real_normed_field, second_countable_topology}"
assumes "a ≠ 0"
shows "set_integrable M A (λt. f t / a) ⟷ set_integrable M A f"
by (simp add: divide_inverse assms)
lemma set_integral_add [simp, intro]:
fixes f g :: "_ ⇒ _ :: {banach, second_countable_topology}"
assumes "set_integrable M A f" "set_integrable M A g"
shows "set_integrable M A (λx. f x + g x)"
and "(LINT x:A|M. f x + g x) = (LINT x:A|M. f x) + (LINT x:A|M. g x)"
using assms unfolding set_integrable_def set_lebesgue_integral_def by (simp_all add: scaleR_add_right)
lemma set_integral_diff [simp, intro]:
assumes "set_integrable M A f" "set_integrable M A g"
shows "set_integrable M A (λx. f x - g x)" and "(LINT x:A|M. f x - g x) = (LINT x:A|M. f x) - (LINT x:A|M. g x)"
using assms unfolding set_integrable_def set_lebesgue_integral_def by (simp_all add: scaleR_diff_right)
lemma set_integral_uminus: "set_integrable M A f ⟹ (LINT x:A|M. - f x) = - (LINT x:A|M. f x)"
unfolding set_integrable_def set_lebesgue_integral_def
by (subst integral_minus[symmetric]) simp_all
lemma set_integral_complex_of_real:
"(LINT x:A|M. complex_of_real (f x)) = of_real (LINT x:A|M. f x)"
unfolding set_lebesgue_integral_def
by (subst integral_complex_of_real[symmetric])
(auto intro!: Bochner_Integration.integral_cong split: split_indicator)
lemma set_integral_mono:
fixes f g :: "_ ⇒ real"
assumes "set_integrable M A f" "set_integrable M A g"
"⋀x. x ∈ A ⟹ f x ≤ g x"
shows "(LINT x:A|M. f x) ≤ (LINT x:A|M. g x)"
using assms unfolding set_integrable_def set_lebesgue_integral_def
by (auto intro: integral_mono split: split_indicator)
lemma set_integral_mono_AE:
fixes f g :: "_ ⇒ real"
assumes "set_integrable M A f" "set_integrable M A g"
"AE x ∈ A in M. f x ≤ g x"
shows "(LINT x:A|M. f x) ≤ (LINT x:A|M. g x)"
using assms unfolding set_integrable_def set_lebesgue_integral_def
by (auto intro: integral_mono_AE split: split_indicator)
lemma set_integrable_abs: "set_integrable M A f ⟹ set_integrable M A (λx. ¦f x¦ :: real)"
using integrable_abs[of M "λx. f x * indicator A x"]unfolding set_integrable_def by (simp add: abs_mult ac_simps)
lemma set_integrable_abs_iff:
fixes f :: "_ ⇒ real"
shows "set_borel_measurable M A f ⟹ set_integrable M A (λx. ¦f x¦) = set_integrable M A f"
unfolding set_integrable_def set_borel_measurable_def
by (subst (2) integrable_abs_iff[symmetric]) (simp_all add: abs_mult ac_simps)
lemma set_integrable_abs_iff':
fixes f :: "_ ⇒ real"
shows "f ∈ borel_measurable M ⟹ A ∈ sets M ⟹
set_integrable M A (λx. ¦f x¦) = set_integrable M A f"
by (simp add: set_borel_measurable_def set_integrable_abs_iff)
lemma set_integrable_discrete_difference:
fixes f :: "'a ⇒ 'b::{banach, second_countable_topology}"
assumes "countable X"
assumes diff: "(A - B) ∪ (B - A) ⊆ X"
assumes "⋀x. x ∈ X ⟹ emeasure M {x} = 0" "⋀x. x ∈ X ⟹ {x} ∈ sets M"
shows "set_integrable M A f ⟷ set_integrable M B f"
unfolding set_integrable_def
proof (rule integrable_discrete_difference[where X=X])
show "⋀x. x ∈ space M ⟹ x ∉ X ⟹ indicator A x *⇩R f x = indicator B x *⇩R f x"
using diff by (auto split: split_indicator)
qed fact+
lemma set_integral_discrete_difference:
fixes f :: "'a ⇒ 'b::{banach, second_countable_topology}"
assumes "countable X"
assumes diff: "(A - B) ∪ (B - A) ⊆ X"
assumes "⋀x. x ∈ X ⟹ emeasure M {x} = 0" "⋀x. x ∈ X ⟹ {x} ∈ sets M"
shows "set_lebesgue_integral M A f = set_lebesgue_integral M B f"
unfolding set_lebesgue_integral_def
proof (rule integral_discrete_difference[where X=X])
show "⋀x. x ∈ space M ⟹ x ∉ X ⟹ indicator A x *⇩R f x = indicator B x *⇩R f x"
using diff by (auto split: split_indicator)
qed fact+
lemma set_integrable_Un:
fixes f g :: "_ ⇒ _ :: {banach, second_countable_topology}"
assumes f_A: "set_integrable M A f" and f_B: "set_integrable M B f"
and [measurable]: "A ∈ sets M" "B ∈ sets M"
shows "set_integrable M (A ∪ B) f"
proof -
have "set_integrable M (A - B) f"
using f_A by (rule set_integrable_subset) auto
with f_B have "integrable M (λx. indicator (A - B) x *⇩R f x + indicator B x *⇩R f x)"
unfolding set_integrable_def using integrable_add by blast
then show ?thesis
unfolding set_integrable_def
by (rule integrable_cong[THEN iffD1, rotated 2]) (auto split: split_indicator)
qed
lemma set_integrable_empty [simp]: "set_integrable M {} f"
by (auto simp: set_integrable_def)
lemma set_integrable_UN:
fixes f :: "_ ⇒ _ :: {banach, second_countable_topology}"
assumes "finite I" "⋀i. i∈I ⟹ set_integrable M (A i) f"
"⋀i. i∈I ⟹ A i ∈ sets M"
shows "set_integrable M (⋃i∈I. A i) f"
using assms
by (induct I) (auto simp: set_integrable_Un sets.finite_UN)
lemma set_integral_Un:
fixes f :: "_ ⇒ _ :: {banach, second_countable_topology}"
assumes "A ∩ B = {}"
and "set_integrable M A f"
and "set_integrable M B f"
shows "(LINT x:A∪B|M. f x) = (LINT x:A|M. f x) + (LINT x:B|M. f x)"
using assms
unfolding set_integrable_def set_lebesgue_integral_def
by (auto simp add: indicator_union_arith indicator_inter_arith[symmetric] scaleR_add_left)
lemma set_integral_cong_set:
fixes f :: "_ ⇒ _ :: {banach, second_countable_topology}"
assumes "set_borel_measurable M A f" "set_borel_measurable M B f"
and ae: "AE x in M. x ∈ A ⟷ x ∈ B"
shows "(LINT x:B|M. f x) = (LINT x:A|M. f x)"
unfolding set_lebesgue_integral_def
proof (rule integral_cong_AE)
show "AE x in M. indicator B x *⇩R f x = indicator A x *⇩R f x"
using ae by (auto simp: subset_eq split: split_indicator)
qed (use assms in ‹auto simp: set_borel_measurable_def›)
proposition set_borel_measurable_subset:
fixes f :: "_ ⇒ _ :: {banach, second_countable_topology}"
assumes [measurable]: "set_borel_measurable M A f" "B ∈ sets M" and "B ⊆ A"
shows "set_borel_measurable M B f"
proof-
have "set_borel_measurable M B (λx. indicator A x *⇩R f x)"
using assms unfolding set_borel_measurable_def
using borel_measurable_indicator borel_measurable_scaleR by blast
moreover have "(λx. indicator B x *⇩R indicator A x *⇩R f x) = (λx. indicator B x *⇩R f x)"
using ‹B ⊆ A› by (auto simp: fun_eq_iff split: split_indicator)
ultimately show ?thesis
unfolding set_borel_measurable_def by simp
qed
lemma set_integral_Un_AE:
fixes f :: "_ ⇒ _ :: {banach, second_countable_topology}"
assumes ae: "AE x in M. ¬ (x ∈ A ∧ x ∈ B)" and [measurable]: "A ∈ sets M" "B ∈ sets M"
and "set_integrable M A f"
and "set_integrable M B f"
shows "(LINT x:A∪B|M. f x) = (LINT x:A|M. f x) + (LINT x:B|M. f x)"
proof -
have f: "set_integrable M (A ∪ B) f"
by (intro set_integrable_Un assms)
then have f': "set_borel_measurable M (A ∪ B) f"
using integrable_iff_bounded set_borel_measurable_def set_integrable_def by blast
have "(LINT x:A∪B|M. f x) = (LINT x:(A - A ∩ B) ∪ (B - A ∩ B)|M. f x)"
proof (rule set_integral_cong_set)
show "AE x in M. (x ∈ A - A ∩ B ∪ (B - A ∩ B)) = (x ∈ A ∪ B)"
using ae by auto
show "set_borel_measurable M (A - A ∩ B ∪ (B - A ∩ B)) f"
using f' by (rule set_borel_measurable_subset) auto
qed fact
also have "… = (LINT x:(A - A ∩ B)|M. f x) + (LINT x:(B - A ∩ B)|M. f x)"
by (auto intro!: set_integral_Un set_integrable_subset[OF f])
also have "… = (LINT x:A|M. f x) + (LINT x:B|M. f x)"
using ae
by (intro arg_cong2[where f="(+)"] set_integral_cong_set)
(auto intro!: set_borel_measurable_subset[OF f'])
finally show ?thesis .
qed
lemma set_integral_finite_Union:
fixes f :: "_ ⇒ _ :: {banach, second_countable_topology}"
assumes "finite I" "disjoint_family_on A I"
and "⋀i. i ∈ I ⟹ set_integrable M (A i) f" "⋀i. i ∈ I ⟹ A i ∈ sets M"
shows "(LINT x:(⋃i∈I. A i)|M. f x) = (∑i∈I. LINT x:A i|M. f x)"
using assms
proof induction
case (insert x F)
then have "A x ∩ ⋃(A ` F) = {}"
by (meson disjoint_family_on_insert)
with insert show ?case
by (simp add: set_integral_Un set_integrable_Un set_integrable_UN disjoint_family_on_insert)
qed (simp add: set_lebesgue_integral_def)
lemma pos_integrable_to_top:
fixes l::real
assumes "⋀i. A i ∈ sets M" "mono A"
assumes nneg: "⋀x i. x ∈ A i ⟹ 0 ≤ f x"
and intgbl: "⋀i::nat. set_integrable M (A i) f"
and lim: "(λi::nat. LINT x:A i|M. f x) ⇢ l"
shows "set_integrable M (⋃i. A i) f"
unfolding set_integrable_def
apply (rule integrable_monotone_convergence[where f = "λi::nat. λx. indicator (A i) x *⇩R f x" and x = l])
apply (rule intgbl [unfolded set_integrable_def])
prefer 3 apply (rule lim [unfolded set_lebesgue_integral_def])
apply (rule AE_I2)
using ‹mono A› apply (auto simp: mono_def nneg split: split_indicator) []
proof (rule AE_I2)
{ fix x assume "x ∈ space M"
show "(λi. indicator (A i) x *⇩R f x) ⇢ indicator (⋃i. A i) x *⇩R f x"
proof cases
assume "∃i. x ∈ A i"
then obtain i where "x ∈ A i" ..
then have "∀⇩F i in sequentially. x ∈ A i"
using ‹x ∈ A i› ‹mono A› by (auto simp: eventually_sequentially mono_def)
with eventually_mono have "∀⇩F i in sequentially. indicat_real (A i) x *⇩R f x = indicat_real (⋃ (range A)) x *⇩R f x"
by fastforce
then show ?thesis
by (intro tendsto_eventually)
qed auto }
then show "(λx. indicator (⋃i. A i) x *⇩R f x) ∈ borel_measurable M"
apply (rule borel_measurable_LIMSEQ_real)
apply assumption
using intgbl set_integrable_def by blast
qed
text ‹Proof from Royden, \emph{Real Analysis}, p. 91.›
lemma lebesgue_integral_countable_add:
fixes f :: "_ ⇒ 'a :: {banach, second_countable_topology}"
assumes meas[intro]: "⋀i::nat. A i ∈ sets M"
and disj: "⋀i j. i ≠ j ⟹ A i ∩ A j = {}"
and intgbl: "set_integrable M (⋃i. A i) f"
shows "(LINT x:(⋃i. A i)|M. f x) = (∑i. (LINT x:(A i)|M. f x))"
unfolding set_lebesgue_integral_def
proof (subst integral_suminf[symmetric])
show int_A: "integrable M (λx. indicat_real (A i) x *⇩R f x)" for i
using intgbl unfolding set_integrable_def [symmetric]
by (rule set_integrable_subset) auto
{ fix x assume "x ∈ space M"
have "(λi. indicator (A i) x *⇩R f x) sums (indicator (⋃i. A i) x *⇩R f x)"
by (intro sums_scaleR_left indicator_sums) fact }
note sums = this
have norm_f: "⋀i. set_integrable M (A i) (λx. norm (f x))"
using int_A[THEN integrable_norm] unfolding set_integrable_def by auto
show "AE x in M. summable (λi. norm (indicator (A i) x *⇩R f x))"
using disj by (intro AE_I2) (auto intro!: summable_mult2 sums_summable[OF indicator_sums])
show "summable (λi. LINT x|M. norm (indicator (A i) x *⇩R f x))"
proof (rule summableI_nonneg_bounded)
fix n
show "0 ≤ LINT x|M. norm (indicator (A n) x *⇩R f x)"
using norm_f by (auto intro!: integral_nonneg_AE)
have "(∑i<n. LINT x|M. norm (indicator (A i) x *⇩R f x)) = (∑i<n. LINT x:A i|M. norm (f x))"
by (simp add: abs_mult set_lebesgue_integral_def)
also have "… = set_lebesgue_integral M (⋃i<n. A i) (λx. norm (f x))"
using norm_f
by (subst set_integral_finite_Union) (auto simp: disjoint_family_on_def disj)
also have "… ≤ set_lebesgue_integral M (⋃i. A i) (λx. norm (f x))"
using intgbl[unfolded set_integrable_def, THEN integrable_norm] norm_f
unfolding set_lebesgue_integral_def set_integrable_def
apply (intro integral_mono set_integrable_UN[of "{..<n}", unfolded set_integrable_def])
apply (auto split: split_indicator)
done
finally show "(∑i<n. LINT x|M. norm (indicator (A i) x *⇩R f x)) ≤
set_lebesgue_integral M (⋃i. A i) (λx. norm (f x))"
by simp
qed
show "LINT x|M. indicator (⋃(A ` UNIV)) x *⇩R f x = LINT x|M. (∑i. indicator (A i) x *⇩R f x)"
by (metis (no_types, lifting) integral_cong sums sums_unique)
qed
lemma set_integral_cont_up:
fixes f :: "_ ⇒ 'a :: {banach, second_countable_topology}"
assumes [measurable]: "⋀i. A i ∈ sets M" and A: "incseq A"
and intgbl: "set_integrable M (⋃i. A i) f"
shows "(λi. LINT x:(A i)|M. f x) ⇢ (LINT x:(⋃i. A i)|M. f x)"
unfolding set_lebesgue_integral_def
proof (intro integral_dominated_convergence[where w="λx. indicator (⋃i. A i) x *⇩R norm (f x)"])
have int_A: "⋀i. set_integrable M (A i) f"
using intgbl by (rule set_integrable_subset) auto
show "⋀i. (λx. indicator (A i) x *⇩R f x) ∈ borel_measurable M"
using int_A integrable_iff_bounded set_integrable_def by blast
show "(λx. indicator (⋃(A ` UNIV)) x *⇩R f x) ∈ borel_measurable M"
using integrable_iff_bounded intgbl set_integrable_def by blast
show "integrable M (λx. indicator (⋃i. A i) x *⇩R norm (f x))"
using int_A intgbl integrable_norm unfolding set_integrable_def
by fastforce
{ fix x i assume "x ∈ A i"
with A have "(λxa. indicator (A xa) x::real) ⇢ 1 ⟷ (λxa. 1::real) ⇢ 1"
by (intro filterlim_cong refl)
(fastforce simp: eventually_sequentially incseq_def subset_eq intro!: exI[of _ i]) }
then show "AE x in M. (λi. indicator (A i) x *⇩R f x) ⇢ indicator (⋃i. A i) x *⇩R f x"
by (intro AE_I2 tendsto_intros) (auto split: split_indicator)
qed (auto split: split_indicator)
lemma set_integral_cont_down:
fixes f :: "_ ⇒ 'a :: {banach, second_countable_topology}"
assumes [measurable]: "⋀i. A i ∈ sets M" and A: "decseq A"
and int0: "set_integrable M (A 0) f"
shows "(λi::nat. LINT x:(A i)|M. f x) ⇢ (LINT x:(⋂i. A i)|M. f x)"
unfolding set_lebesgue_integral_def
proof (rule integral_dominated_convergence)
have int_A: "⋀i. set_integrable M (A i) f"
using int0 by (rule set_integrable_subset) (insert A, auto simp: decseq_def)
have "integrable M (λc. norm (indicat_real (A 0) c *⇩R f c))"
by (metis (no_types) int0 integrable_norm set_integrable_def)
then show "integrable M (λx. indicator (A 0) x *⇩R norm (f x))"
by force
have "set_integrable M (⋂i. A i) f"
using int0 by (rule set_integrable_subset) (insert A, auto simp: decseq_def)
with int_A show "(λx. indicat_real (⋂(A ` UNIV)) x *⇩R f x) ∈ borel_measurable M"
"⋀i. (λx. indicat_real (A i) x *⇩R f x) ∈ borel_measurable M"
by (auto simp: set_integrable_def)
show "⋀i. AE x in M. norm (indicator (A i) x *⇩R f x) ≤ indicator (A 0) x *⇩R norm (f x)"
using A by (auto split: split_indicator simp: decseq_def)
{ fix x i assume "x ∈ space M" "x ∉ A i"
with A have "(λi. indicator (A i) x::real) ⇢ 0 ⟷ (λi. 0::real) ⇢ 0"
by (intro filterlim_cong refl)
(auto split: split_indicator simp: eventually_sequentially decseq_def intro!: exI[of _ i]) }
then show "AE x in M. (λi. indicator (A i) x *⇩R f x) ⇢ indicator (⋂i. A i) x *⇩R f x"
by (intro AE_I2 tendsto_intros) (auto split: split_indicator)
qed
lemma set_integral_at_point:
fixes a :: real
assumes "set_integrable M {a} f"
and [simp]: "{a} ∈ sets M" and "(emeasure M) {a} ≠ ∞"
shows "(LINT x:{a} | M. f x) = f a * measure M {a}"
proof-
have "set_lebesgue_integral M {a} f = set_lebesgue_integral M {a} (%x. f a)"
by (intro set_lebesgue_integral_cong) simp_all
then show ?thesis using assms
unfolding set_lebesgue_integral_def by simp
qed
section ‹Complex integrals›
abbreviation complex_integrable :: "'a measure ⇒ ('a ⇒ complex) ⇒ bool" where
"complex_integrable M f ≡ integrable M f"
abbreviation complex_lebesgue_integral :: "'a measure ⇒ ('a ⇒ complex) ⇒ complex" ("integral⇧C") where
"integral⇧C M f == integral⇧L M f"
syntax
"_complex_lebesgue_integral" :: "pttrn ⇒ complex ⇒ 'a measure ⇒ complex"
("∫⇧C _. _ ∂_" [0,0] 110)
translations
"∫⇧Cx. f ∂M" == "CONST complex_lebesgue_integral M (λx. f)"
syntax
"_ascii_complex_lebesgue_integral" :: "pttrn ⇒ 'a measure ⇒ real ⇒ real"
("(3CLINT _|_. _)" [0,0,10] 10)
translations
"CLINT x|M. f" == "CONST complex_lebesgue_integral M (λx. f)"
lemma complex_integrable_cnj [simp]:
"complex_integrable M (λx. cnj (f x)) ⟷ complex_integrable M f"
proof
assume "complex_integrable M (λx. cnj (f x))"
then have "complex_integrable M (λx. cnj (cnj (f x)))"
by (rule integrable_cnj)
then show "complex_integrable M f"
by simp
qed simp
lemma complex_of_real_integrable_eq:
"complex_integrable M (λx. complex_of_real (f x)) ⟷ integrable M f"
proof
assume "complex_integrable M (λx. complex_of_real (f x))"
then have "integrable M (λx. Re (complex_of_real (f x)))"
by (rule integrable_Re)
then show "integrable M f"
by simp
qed simp
abbreviation complex_set_integrable :: "'a measure ⇒ 'a set ⇒ ('a ⇒ complex) ⇒ bool" where
"complex_set_integrable M A f ≡ set_integrable M A f"
abbreviation complex_set_lebesgue_integral :: "'a measure ⇒ 'a set ⇒ ('a ⇒ complex) ⇒ complex" where
"complex_set_lebesgue_integral M A f ≡ set_lebesgue_integral M A f"
syntax
"_ascii_complex_set_lebesgue_integral" :: "pttrn ⇒ 'a set ⇒ 'a measure ⇒ real ⇒ real"
("(4CLINT _:_|_. _)" [0,0,0,10] 10)
translations
"CLINT x:A|M. f" == "CONST complex_set_lebesgue_integral M A (λx. f)"
lemma set_measurable_continuous_on_ivl:
assumes "continuous_on {a..b} (f :: real ⇒ real)"
shows "set_borel_measurable borel {a..b} f"
unfolding set_borel_measurable_def
by (rule borel_measurable_continuous_on_indicator[OF _ assms]) simp
section ‹NN Set Integrals›
text‹This notation is from Sébastien Gouëzel: His use is not directly in line with the
notations in this file, they are more in line with sum, and more readable he thinks.›
abbreviation "set_nn_integral M A f ≡ nn_integral M (λx. f x * indicator A x)"
syntax
"_set_nn_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal"
("(∫⇧+((_)∈(_)./ _)/∂_)" [0,0,0,110] 10)
"_set_lebesgue_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal"
("(∫((_)∈(_)./ _)/∂_)" [0,0,0,110] 10)
translations
"∫⇧+x ∈ A. f ∂M" == "CONST set_nn_integral M A (λx. f)"
"∫x ∈ A. f ∂M" == "CONST set_lebesgue_integral M A (λx. f)"
lemma set_nn_integral_cong:
assumes "M = M'" "A = B" "⋀x. x ∈ space M ∩ A ⟹ f x = g x"
shows "set_nn_integral M A f = set_nn_integral M' B g"
by (metis (mono_tags, lifting) IntI assms indicator_simps(2) mult_eq_0_iff nn_integral_cong)
lemma nn_integral_disjoint_pair:
assumes [measurable]: "f ∈ borel_measurable M"
"B ∈ sets M" "C ∈ sets M"
"B ∩ C = {}"
shows "(∫⇧+x ∈ B ∪ C. f x ∂M) = (∫⇧+x ∈ B. f x ∂M) + (∫⇧+x ∈ C. f x ∂M)"
proof -
have mes: "⋀D. D ∈ sets M ⟹ (λx. f x * indicator D x) ∈ borel_measurable M" by simp
have pos: "⋀D. AE x in M. f x * indicator D x ≥ 0" using assms(2) by auto
have "⋀x. f x * indicator (B ∪ C) x = f x * indicator B x + f x * indicator C x" using assms(4)
by (auto split: split_indicator)
then have "(∫⇧+x. f x * indicator (B ∪ C) x ∂M) = (∫⇧+x. f x * indicator B x + f x * indicator C x ∂M)"
by simp
also have "… = (∫⇧+x. f x * indicator B x ∂M) + (∫⇧+x. f x * indicator C x ∂M)"
by (rule nn_integral_add) (auto simp add: assms mes pos)
finally show ?thesis by simp
qed
lemma nn_integral_disjoint_pair_countspace:
assumes "B ∩ C = {}"
shows "(∫⇧+x ∈ B ∪ C. f x ∂count_space UNIV) = (∫⇧+x ∈ B. f x ∂count_space UNIV) + (∫⇧+x ∈ C. f x ∂count_space UNIV)"
by (rule nn_integral_disjoint_pair) (simp_all add: assms)
lemma nn_integral_null_delta:
assumes "A ∈ sets M" "B ∈ sets M"
"(A - B) ∪ (B - A) ∈ null_sets M"
shows "(∫⇧+x ∈ A. f x ∂M) = (∫⇧+x ∈ B. f x ∂M)"
proof (rule nn_integral_cong_AE)
have *: "AE a in M. a ∉ (A - B) ∪ (B - A)"
using assms(3) AE_not_in by blast
then show ‹AE x in M. f x * indicator A x = f x * indicator B x›
by auto
qed
proposition nn_integral_disjoint_family:
assumes [measurable]: "f ∈ borel_measurable M" "⋀(n::nat). B n ∈ sets M"
and "disjoint_family B"
shows "(∫⇧+x ∈ (⋃n. B n). f x ∂M) = (∑n. (∫⇧+x ∈ B n. f x ∂M))"
proof -
have "(∫⇧+ x. (∑n. f x * indicator (B n) x) ∂M) = (∑n. (∫⇧+ x. f x * indicator (B n) x ∂M))"
by (rule nn_integral_suminf) simp
moreover have "(∑n. f x * indicator (B n) x) = f x * indicator (⋃n. B n) x" for x
proof (cases)
assume "x ∈ (⋃n. B n)"
then obtain n where "x ∈ B n" by blast
have a: "finite {n}" by simp
have "⋀i. i ≠ n ⟹ x ∉ B i" using ‹x ∈ B n› assms(3) disjoint_family_on_def
by (metis IntI UNIV_I empty_iff)
then have "⋀i. i ∉ {n} ⟹ indicator (B i) x = (0::ennreal)" using indicator_def by simp
then have b: "⋀i. i ∉ {n} ⟹ f x * indicator (B i) x = (0::ennreal)" by simp
define h where "h = (λi. f x * indicator (B i) x)"
then have "⋀i. i ∉ {n} ⟹ h i = 0" using b by simp
then have "(∑i. h i) = (∑i∈{n}. h i)"
by (metis sums_unique[OF sums_finite[OF a]])
then have "(∑i. h i) = h n" by simp
then have "(∑n. f x * indicator (B n) x) = f x * indicator (B n) x" using h_def by simp
then have "(∑n. f x * indicator (B n) x) = f x" using ‹x ∈ B n› indicator_def by simp
then show ?thesis using ‹x ∈ (⋃n. B n)› by auto
next
assume "x ∉ (⋃n. B n)"
then have "⋀n. f x * indicator (B n) x = 0" by simp
have "(∑n. f x * indicator (B n) x) = 0"
by (simp add: ‹⋀n. f x * indicator (B n) x = 0›)
then show ?thesis using ‹x ∉ (⋃n. B n)› by auto
qed
ultimately show ?thesis by simp
qed
lemma nn_set_integral_add:
assumes [measurable]: "f ∈ borel_measurable M" "g ∈ borel_measurable M"
"A ∈ sets M"
shows "(∫⇧+x ∈ A. (f x + g x) ∂M) = (∫⇧+x ∈ A. f x ∂M) + (∫⇧+x ∈ A. g x ∂M)"
proof -
have "(∫⇧+x ∈ A. (f x + g x) ∂M) = (∫⇧+x. (f x * indicator A x + g x * indicator A x) ∂M)"
by (auto simp add: indicator_def intro!: nn_integral_cong)
also have "… = (∫⇧+x. f x * indicator A x ∂M) + (∫⇧+x. g x * indicator A x ∂M)"
apply (rule nn_integral_add) using assms(1) assms(2) by auto
finally show ?thesis by simp
qed
lemma nn_set_integral_cong:
assumes "AE x in M. f x = g x"
shows "(∫⇧+x ∈ A. f x ∂M) = (∫⇧+x ∈ A. g x ∂M)"
apply (rule nn_integral_cong_AE) using assms(1) by auto
lemma nn_set_integral_set_mono:
"A ⊆ B ⟹ (∫⇧+ x ∈ A. f x ∂M) ≤ (∫⇧+ x ∈ B. f x ∂M)"
by (auto intro!: nn_integral_mono split: split_indicator)
lemma nn_set_integral_mono:
assumes [measurable]: "f ∈ borel_measurable M" "g ∈ borel_measurable M"
"A ∈ sets M"
and "AE x∈A in M. f x ≤ g x"
shows "(∫⇧+x ∈ A. f x ∂M) ≤ (∫⇧+x ∈ A. g x ∂M)"
by (auto intro!: nn_integral_mono_AE split: split_indicator simp: assms)
lemma nn_set_integral_space [simp]:
shows "(∫⇧+ x ∈ space M. f x ∂M) = (∫⇧+x. f x ∂M)"
by (metis (mono_tags, lifting) indicator_simps(1) mult.right_neutral nn_integral_cong)
lemma nn_integral_count_compose_inj:
assumes "inj_on g A"
shows "(∫⇧+x ∈ A. f (g x) ∂count_space UNIV) = (∫⇧+y ∈ g`A. f y ∂count_space UNIV)"
proof -
have "(∫⇧+x ∈ A. f (g x) ∂count_space UNIV) = (∫⇧+x. f (g x) ∂count_space A)"
by (auto simp add: nn_integral_count_space_indicator[symmetric])
also have "… = (∫⇧+y. f y ∂count_space (g`A))"
by (simp add: assms nn_integral_bij_count_space inj_on_imp_bij_betw)
also have "… = (∫⇧+y ∈ g`A. f y ∂count_space UNIV)"
by (auto simp add: nn_integral_count_space_indicator[symmetric])
finally show ?thesis by simp
qed
lemma nn_integral_count_compose_bij:
assumes "bij_betw g A B"
shows "(∫⇧+x ∈ A. f (g x) ∂count_space UNIV) = (∫⇧+y ∈ B. f y ∂count_space UNIV)"
proof -
have "inj_on g A" using assms bij_betw_def by auto
then have "(∫⇧+x ∈ A. f (g x) ∂count_space UNIV) = (∫⇧+y ∈ g`A. f y ∂count_space UNIV)"
by (rule nn_integral_count_compose_inj)
then show ?thesis using assms by (simp add: bij_betw_def)
qed
lemma set_integral_null_delta:
fixes f::"_ ⇒ _ :: {banach, second_countable_topology}"
assumes [measurable]: "integrable M f" "A ∈ sets M" "B ∈ sets M"
and null: "(A - B) ∪ (B - A) ∈ null_sets M"
shows "(∫x ∈ A. f x ∂M) = (∫x ∈ B. f x ∂M)"
proof (rule set_integral_cong_set)
have *: "AE a in M. a ∉ (A - B) ∪ (B - A)"
using null AE_not_in by blast
then show "AE x in M. (x ∈ B) = (x ∈ A)"
by auto
qed (simp_all add: set_borel_measurable_def)
lemma set_integral_space:
assumes "integrable M f"
shows "(∫x ∈ space M. f x ∂M) = (∫x. f x ∂M)"
by (metis (no_types, lifting) indicator_simps(1) integral_cong scaleR_one set_lebesgue_integral_def)
lemma null_if_pos_func_has_zero_nn_int:
fixes f::"'a ⇒ ennreal"
assumes [measurable]: "f ∈ borel_measurable M" "A ∈ sets M"
and "AE x∈A in M. f x > 0" "(∫⇧+x∈A. f x ∂M) = 0"
shows "A ∈ null_sets M"
proof -
have "AE x in M. f x * indicator A x = 0"
by (subst nn_integral_0_iff_AE[symmetric], auto simp add: assms(4))
then have "AE x∈A in M. False" using assms(3) by auto
then show "A ∈ null_sets M" using assms(2) by (simp add: AE_iff_null_sets)
qed
lemma null_if_pos_func_has_zero_int:
assumes [measurable]: "integrable M f" "A ∈ sets M"
and "AE x∈A in M. f x > 0" "(∫x∈A. f x ∂M) = (0::real)"
shows "A ∈ null_sets M"
proof -
have "AE x in M. indicator A x * f x = 0"
apply (subst integral_nonneg_eq_0_iff_AE[symmetric])
using assms integrable_mult_indicator[OF ‹A ∈ sets M› assms(1)]
by (auto simp: set_lebesgue_integral_def)
then have "AE x∈A in M. f x = 0" by auto
then have "AE x∈A in M. False" using assms(3) by auto
then show "A ∈ null_sets M" using assms(2) by (simp add: AE_iff_null_sets)
qed
text‹The next lemma is a variant of ‹density_unique›. Note that it uses the notation
for nonnegative set integrals introduced earlier.›
lemma (in sigma_finite_measure) density_unique2:
assumes [measurable]: "f ∈ borel_measurable M" "f' ∈ borel_measurable M"
assumes density_eq: "⋀A. A ∈ sets M ⟹ (∫⇧+ x ∈ A. f x ∂M) = (∫⇧+ x ∈ A. f' x ∂M)"
shows "AE x in M. f x = f' x"
proof (rule density_unique)
show "density M f = density M f'"
by (intro measure_eqI) (auto simp: emeasure_density intro!: density_eq)
qed (auto simp add: assms)
text ‹The next lemma implies the same statement for Banach-space valued functions
using Hahn-Banach theorem and linear forms. Since they are not yet easily available, I
only formulate it for real-valued functions.›
lemma density_unique_real:
fixes f f'::"_ ⇒ real"
assumes M[measurable]: "integrable M f" "integrable M f'"
assumes density_eq: "⋀A. A ∈ sets M ⟹ (∫x ∈ A. f x ∂M) = (∫x ∈ A. f' x ∂M)"
shows "AE x in M. f x = f' x"
proof -
define A where "A = {x ∈ space M. f x < f' x}"
then have [measurable]: "A ∈ sets M" by simp
have "(∫x ∈ A. (f' x - f x) ∂M) = (∫x ∈ A. f' x ∂M) - (∫x ∈ A. f x ∂M)"
using ‹A ∈ sets M› M integrable_mult_indicator set_integrable_def by blast
then have "(∫x ∈ A. (f' x - f x) ∂M) = 0" using assms(3) by simp
then have "A ∈ null_sets M"
using A_def null_if_pos_func_has_zero_int[where ?f = "λx. f' x - f x" and ?A = A] assms by auto
then have "AE x in M. x ∉ A" by (simp add: AE_not_in)
then have *: "AE x in M. f' x ≤ f x" unfolding A_def by auto
define B where "B = {x ∈ space M. f' x < f x}"
then have [measurable]: "B ∈ sets M" by simp
have "(∫x ∈ B. (f x - f' x) ∂M) = (∫x ∈ B. f x ∂M) - (∫x ∈ B. f' x ∂M)"
using ‹B ∈ sets M› M integrable_mult_indicator set_integrable_def by blast
then have "(∫x ∈ B. (f x - f' x) ∂M) = 0" using assms(3) by simp
then have "B ∈ null_sets M"
using B_def null_if_pos_func_has_zero_int[where ?f = "λx. f x - f' x" and ?A = B] assms by auto
then have "AE x in M. x ∉ B" by (simp add: AE_not_in)
then have "AE x in M. f' x ≥ f x" unfolding B_def by auto
then show ?thesis using * by auto
qed
section ‹Scheffé's lemma›
text ‹The next lemma shows that ‹L⇧1› convergence of a sequence of functions follows from almost
everywhere convergence and the weaker condition of the convergence of the integrated norms (or even
just the nontrivial inequality about them). Useful in a lot of contexts! This statement (or its
variations) are known as Scheffe lemma.
The formalization is more painful as one should jump back and forth between reals and ereals and justify
all the time positivity or integrability (thankfully, measurability is handled more or less automatically).›
proposition Scheffe_lemma1:
assumes "⋀n. integrable M (F n)" "integrable M f"
"AE x in M. (λn. F n x) ⇢ f x"
"limsup (λn. ∫⇧+ x. norm(F n x) ∂M) ≤ (∫⇧+ x. norm(f x) ∂M)"
shows "(λn. ∫⇧+ x. norm(F n x - f x) ∂M) ⇢ 0"
proof -
have [measurable]: "⋀n. F n ∈ borel_measurable M" "f ∈ borel_measurable M"
using assms(1) assms(2) by simp_all
define G where "G = (λn x. norm(f x) + norm(F n x) - norm(F n x - f x))"
have [measurable]: "⋀n. G n ∈ borel_measurable M" unfolding G_def by simp
have G_pos[simp]: "⋀n x. G n x ≥ 0"
unfolding G_def by (metis ge_iff_diff_ge_0 norm_minus_commute norm_triangle_ineq4)
have finint: "(∫⇧+ x. norm(f x) ∂M) ≠ ∞"
using has_bochner_integral_implies_finite_norm[OF has_bochner_integral_integrable[OF ‹integrable M f›]]
by simp
then have fin2: "2 * (∫⇧+ x. norm(f x) ∂M) ≠ ∞"
by (auto simp: ennreal_mult_eq_top_iff)
{
fix x assume *: "(λn. F n x) ⇢ f x"
then have "(λn. norm(F n x)) ⇢ norm(f x)" using tendsto_norm by blast
moreover have "(λn. norm(F n x - f x)) ⇢ 0" using * Lim_null tendsto_norm_zero_iff by fastforce
ultimately have a: "(λn. norm(F n x) - norm(F n x - f x)) ⇢ norm(f x)" using tendsto_diff by fastforce
have "(λn. norm(f x) + (norm(F n x) - norm(F n x - f x))) ⇢ norm(f x) + norm(f x)"
by (rule tendsto_add) (auto simp add: a)
moreover have "⋀n. G n x = norm(f x) + (norm(F n x) - norm(F n x - f x))" unfolding G_def by simp
ultimately have "(λn. G n x) ⇢ 2 * norm(f x)" by simp
then have "(λn. ennreal(G n x)) ⇢ ennreal(2 * norm(f x))" by simp
then have "liminf (λn. ennreal(G n x)) = ennreal(2 * norm(f x))"
using sequentially_bot tendsto_iff_Liminf_eq_Limsup by blast
}
then have "AE x in M. liminf (λn. ennreal(G n x)) = ennreal(2 * norm(f x))" using assms(3) by auto
then have "(∫⇧+ x. liminf (λn. ennreal (G n x)) ∂M) = (∫⇧+ x. 2 * ennreal(norm(f x)) ∂M)"
by (simp add: nn_integral_cong_AE ennreal_mult)
also have "… = 2 * (∫⇧+ x. norm(f x) ∂M)" by (rule nn_integral_cmult) auto
finally have int_liminf: "(∫⇧+ x. liminf (λn. ennreal (G n x)) ∂M) = 2 * (∫⇧+ x. norm(f x) ∂M)"
by simp
have "(∫⇧+x. ennreal(norm(f x)) + ennreal(norm(F n x)) ∂M) = (∫⇧+x. norm(f x) ∂M) + (∫⇧+x. norm(F n x) ∂M)" for n
by (rule nn_integral_add) (auto simp add: assms)
then have "limsup (λn. (∫⇧+x. ennreal(norm(f x)) + ennreal(norm(F n x)) ∂M)) =
limsup (λn. (∫⇧+x. norm(f x) ∂M) + (∫⇧+x. norm(F n x) ∂M))"
by simp
also have "… = (∫⇧+x. norm(f x) ∂M) + limsup (λn. (∫⇧+x. norm(F n x) ∂M))"
by (rule Limsup_const_add, auto simp add: finint)
also have "… ≤ (∫⇧+x. norm(f x) ∂M) + (∫⇧+x. norm(f x) ∂M)"
using assms(4) by (simp add: add_left_mono)
also have "… = 2 * (∫⇧+x. norm(f x) ∂M)"
unfolding one_add_one[symmetric] distrib_right by simp
ultimately have a: "limsup (λn. (∫⇧+x. ennreal(norm(f x)) + ennreal(norm(F n x)) ∂M)) ≤
2 * (∫⇧+x. norm(f x) ∂M)" by simp
have le: "ennreal (norm (F n x - f x)) ≤ ennreal (norm (f x)) + ennreal (norm (F n x))" for n x
by (simp add: norm_minus_commute norm_triangle_ineq4 ennreal_minus flip: ennreal_plus)
then have le2: "(∫⇧+ x. ennreal (norm (F n x - f x)) ∂M) ≤ (∫⇧+ x. ennreal (norm (f x)) + ennreal (norm (F n x)) ∂M)" for n
by (rule nn_integral_mono)
have "2 * (∫⇧+ x. norm(f x) ∂M) = (∫⇧+ x. liminf (λn. ennreal (G n x)) ∂M)"
by (simp add: int_liminf)
also have "… ≤ liminf (λn. (∫⇧+x. G n x ∂M))"
by (rule nn_integral_liminf) auto
also have "liminf (λn. (∫⇧+x. G n x ∂M)) =
liminf (λn. (∫⇧+x. ennreal(norm(f x)) + ennreal(norm(F n x)) ∂M) - (∫⇧+x. norm(F n x - f x) ∂M))"
proof (intro arg_cong[where f=liminf] ext)
fix n
have "⋀x. ennreal(G n x) = ennreal(norm(f x)) + ennreal(norm(F n x)) - ennreal(norm(F n x - f x))"
unfolding G_def by (simp add: ennreal_minus flip: ennreal_plus)
moreover have "(∫⇧+x. ennreal(norm(f x)) + ennreal(norm(F n x)) - ennreal(norm(F n x - f x)) ∂M)
= (∫⇧+x. ennreal(norm(f x)) + ennreal(norm(F n x)) ∂M) - (∫⇧+x. norm(F n x - f x) ∂M)"
proof (rule nn_integral_diff)
from le show "AE x in M. ennreal (norm (F n x - f x)) ≤ ennreal (norm (f x)) + ennreal (norm (F n x))"
by simp
from le2 have "(∫⇧+x. ennreal (norm (F n x - f x)) ∂M) < ∞" using assms(1) assms(2)
by (metis has_bochner_integral_implies_finite_norm integrable.simps Bochner_Integration.integrable_diff)
then show "(∫⇧+x. ennreal (norm (F n x - f x)) ∂M) ≠ ∞" by simp
qed (auto simp add: assms)
ultimately show "(∫⇧+x. G n x ∂M) = (∫⇧+x. ennreal(norm(f x)) + ennreal(norm(F n x)) ∂M) - (∫⇧+x. norm(F n x - f x) ∂M)"
by simp
qed
finally have "2 * (∫⇧+ x. norm(f x) ∂M) + limsup (λn. (∫⇧+x. norm(F n x - f x) ∂M)) ≤
liminf (λn. (∫⇧+x. ennreal(norm(f x)) + ennreal(norm(F n x)) ∂M) - (∫⇧+x. norm(F n x - f x) ∂M)) +
limsup (λn. (∫⇧+x. norm(F n x - f x) ∂M))"
by (intro add_mono) auto
also have "… ≤ (limsup (λn. ∫⇧+x. ennreal(norm(f x)) + ennreal(norm(F n x)) ∂M) - limsup (λn. ∫⇧+x. norm (F n x - f x) ∂M)) +
limsup (λn. (∫⇧+x. norm(F n x - f x) ∂M))"
by (intro add_mono liminf_minus_ennreal le2) auto
also have "… = limsup (λn. (∫⇧+x. ennreal(norm(f x)) + ennreal(norm(F n x)) ∂M))"
by (intro diff_add_cancel_ennreal Limsup_mono always_eventually allI le2)
also have "… ≤ 2 * (∫⇧+x. norm(f x) ∂M)"
by fact
finally have "limsup (λn. (∫⇧+x. norm(F n x - f x) ∂M)) = 0"
using fin2 by simp
then show ?thesis
by (rule tendsto_0_if_Limsup_eq_0_ennreal)
qed
proposition Scheffe_lemma2:
fixes F::"nat ⇒ 'a ⇒ 'b::{banach, second_countable_topology}"
assumes "⋀ n::nat. F n ∈ borel_measurable M" "integrable M f"
"AE x in M. (λn. F n x) ⇢ f x"
"⋀n. (∫⇧+ x. norm(F n x) ∂M) ≤ (∫⇧+ x. norm(f x) ∂M)"
shows "(λn. ∫⇧+ x. norm(F n x - f x) ∂M) ⇢ 0"
proof (rule Scheffe_lemma1)
fix n::nat
have "(∫⇧+ x. norm(f x) ∂M) < ∞" using assms(2) by (metis has_bochner_integral_implies_finite_norm integrable.cases)
then have "(∫⇧+ x. norm(F n x) ∂M) < ∞" using assms(4)[of n] by auto
then show "integrable M (F n)" by (subst integrable_iff_bounded, simp add: assms(1)[of n])
qed (auto simp add: assms Limsup_bounded)
lemma tendsto_set_lebesgue_integral_at_right:
fixes a b :: real and f :: "real ⇒ 'a :: {banach,second_countable_topology}"
assumes "a < b" and sets: "⋀a'. a' ∈ {a<..b} ⟹ {a'..b} ∈ sets M"
and "set_integrable M {a<..b} f"
shows "((λa'. set_lebesgue_integral M {a'..b} f) ⤏
set_lebesgue_integral M {a<..b} f) (at_right a)"
proof (rule tendsto_at_right_sequentially[OF assms(1)], goal_cases)
case (1 S)
have eq: "(⋃n. {S n..b}) = {a<..b}"
proof safe
fix x n assume "x ∈ {S n..b}"
with 1(1,2)[of n] show "x ∈ {a<..b}" by auto
next
fix x assume "x ∈ {a<..b}"
with order_tendstoD[OF ‹S ⇢ a›, of x] show "x ∈ (⋃n. {S n..b})"
by (force simp: eventually_at_top_linorder dest: less_imp_le)
qed
have "(λn. set_lebesgue_integral M {S n..b} f) ⇢ set_lebesgue_integral M (⋃n. {S n..b}) f"
by (rule set_integral_cont_up) (insert assms 1, auto simp: eq incseq_def decseq_def less_imp_le)
with eq show ?case by simp
qed
section ‹Convergence of integrals over an interval›
text ‹
The next lemmas relate convergence of integrals over an interval to
improper integrals.
›
lemma tendsto_set_lebesgue_integral_at_left:
fixes a b :: real and f :: "real ⇒ 'a :: {banach,second_countable_topology}"
assumes "a < b" and sets: "⋀b'. b' ∈ {a..<b} ⟹ {a..b'} ∈ sets M"
and "set_integrable M {a..<b} f"
shows "((λb'. set_lebesgue_integral M {a..b'} f) ⤏
set_lebesgue_integral M {a..<b} f) (at_left b)"
proof (rule tendsto_at_left_sequentially[OF assms(1)], goal_cases)
case (1 S)
have eq: "(⋃n. {a..S n}) = {a..<b}"
proof safe
fix x n assume "x ∈ {a..S n}"
with 1(1,2)[of n] show "x ∈ {a..<b}" by auto
next
fix x assume "x ∈ {a..<b}"
with order_tendstoD[OF ‹S ⇢ b›, of x] show "x ∈ (⋃n. {a..S n})"
by (force simp: eventually_at_top_linorder dest: less_imp_le)
qed
have "(λn. set_lebesgue_integral M {a..S n} f) ⇢ set_lebesgue_integral M (⋃n. {a..S n}) f"
by (rule set_integral_cont_up) (insert assms 1, auto simp: eq incseq_def decseq_def less_imp_le)
with eq show ?case by simp
qed
proposition tendsto_set_lebesgue_integral_at_top:
fixes f :: "real ⇒ 'a::{banach, second_countable_topology}"
assumes sets: "⋀b. b ≥ a ⟹ {a..b} ∈ sets M"
and int: "set_integrable M {a..} f"
shows "((λb. set_lebesgue_integral M {a..b} f) ⤏ set_lebesgue_integral M {a..} f) at_top"
proof (rule tendsto_at_topI_sequentially)
fix X :: "nat ⇒ real" assume "filterlim X at_top sequentially"
show "(λn. set_lebesgue_integral M {a..X n} f) ⇢ set_lebesgue_integral M {a..} f"
unfolding set_lebesgue_integral_def
proof (rule integral_dominated_convergence)
show "integrable M (λx. indicat_real {a..} x *⇩R norm (f x))"
using integrable_norm[OF int[unfolded set_integrable_def]] by simp
show "AE x in M. (λn. indicator {a..X n} x *⇩R f x) ⇢ indicat_real {a..} x *⇩R f x"
proof
fix x
from ‹filterlim X at_top sequentially›
have "eventually (λn. x ≤ X n) sequentially"
unfolding filterlim_at_top_ge[where c=x] by auto
then show "(λn. indicator {a..X n} x *⇩R f x) ⇢ indicat_real {a..} x *⇩R f x"
by (intro tendsto_eventually) (auto split: split_indicator elim!: eventually_mono)
qed
fix n show "AE x in M. norm (indicator {a..X n} x *⇩R f x) ≤
indicator {a..} x *⇩R norm (f x)"
by (auto split: split_indicator)
next
from int show "(λx. indicat_real {a..} x *⇩R f x) ∈ borel_measurable M"
by (simp add: set_integrable_def)
next
fix n :: nat
from sets have "{a..X n} ∈ sets M" by (cases "X n ≥ a") auto
with int have "set_integrable M {a..X n} f"
by (rule set_integrable_subset) auto
thus "(λx. indicat_real {a..X n} x *⇩R f x) ∈ borel_measurable M"
by (simp add: set_integrable_def)
qed
qed
proposition tendsto_set_lebesgue_integral_at_bot:
fixes f :: "real ⇒ 'a::{banach, second_countable_topology}"
assumes sets: "⋀a. a ≤ b ⟹ {a..b} ∈ sets M"
and int: "set_integrable M {..b} f"
shows "((λa. set_lebesgue_integral M {a..b} f) ⤏ set_lebesgue_integral M {..b} f) at_bot"
proof (rule tendsto_at_botI_sequentially)
fix X :: "nat ⇒ real" assume "filterlim X at_bot sequentially"
show "(λn. set_lebesgue_integral M {X n..b} f) ⇢ set_lebesgue_integral M {..b} f"
unfolding set_lebesgue_integral_def
proof (rule integral_dominated_convergence)
show "integrable M (λx. indicat_real {..b} x *⇩R norm (f x))"
using integrable_norm[OF int[unfolded set_integrable_def]] by simp
show "AE x in M. (λn. indicator {X n..b} x *⇩R f x) ⇢ indicat_real {..b} x *⇩R f x"
proof
fix x
from ‹filterlim X at_bot sequentially›
have "eventually (λn. x ≥ X n) sequentially"
unfolding filterlim_at_bot_le[where c=x] by auto
then show "(λn. indicator {X n..b} x *⇩R f x) ⇢ indicat_real {..b} x *⇩R f x"
by (intro tendsto_eventually) (auto split: split_indicator elim!: eventually_mono)
qed
fix n show "AE x in M. norm (indicator {X n..b} x *⇩R f x) ≤
indicator {..b} x *⇩R norm (f x)"
by (auto split: split_indicator)
next
from int show "(λx. indicat_real {..b} x *⇩R f x) ∈ borel_measurable M"
by (simp add: set_integrable_def)
next
fix n :: nat
from sets have "{X n..b} ∈ sets M" by (cases "X n ≤ b") auto
with int have "set_integrable M {X n..b} f"
by (rule set_integrable_subset) auto
thus "(λx. indicat_real {X n..b} x *⇩R f x) ∈ borel_measurable M"
by (simp add: set_integrable_def)
qed
qed
theorem integral_Markov_inequality':
fixes u :: "'a ⇒ real"
assumes [measurable]: "set_integrable M A u" and "A ∈ sets M"
assumes "AE x in M. x ∈ A ⟶ u x ≥ 0" and "0 < (c::real)"
shows "emeasure M {x∈A. u x ≥ c} ≤ (1/c::real) * (∫x∈A. u x ∂M)"
proof -
have "(λx. u x * indicator A x) ∈ borel_measurable M"
using assms by (auto simp: set_integrable_def mult_ac)
hence "(λx. ennreal (u x * indicator A x)) ∈ borel_measurable M"
by measurable
also have "(λx. ennreal (u x * indicator A x)) = (λx. ennreal (u x) * indicator A x)"
by (intro ext) (auto simp: indicator_def)
finally have meas: "… ∈ borel_measurable M" .
from assms(3) have AE: "AE x in M. 0 ≤ u x * indicator A x"
by eventually_elim (auto simp: indicator_def)
have nonneg: "set_lebesgue_integral M A u ≥ 0"
unfolding set_lebesgue_integral_def
by (intro Bochner_Integration.integral_nonneg_AE eventually_mono[OF AE]) (auto simp: mult_ac)
have A: "A ⊆ space M"
using ‹A ∈ sets M› by (simp add: sets.sets_into_space)
have "{x ∈ A. u x ≥ c} = {x ∈ A. ennreal(1/c) * u x ≥ 1}"
using ‹c>0› A by (auto simp: ennreal_mult'[symmetric])
then have "emeasure M {x ∈ A. u x ≥ c} = emeasure M ({x ∈ A. ennreal(1/c) * u x ≥ 1})"
by simp
also have "… ≤ ennreal(1/c) * (∫⇧+ x. ennreal(u x) * indicator A x ∂M)"
by (intro nn_integral_Markov_inequality meas assms)
also have "(∫⇧+ x. ennreal(u x) * indicator A x ∂M) = ennreal (set_lebesgue_integral M A u)"
unfolding set_lebesgue_integral_def nn_integral_set_ennreal using assms AE
by (subst nn_integral_eq_integral) (simp_all add: mult_ac set_integrable_def)
finally show ?thesis
using ‹c > 0› nonneg by (subst ennreal_mult) auto
qed
theorem integral_Markov_inequality'_measure:
assumes [measurable]: "set_integrable M A u" and "A ∈ sets M"
and "AE x in M. x ∈ A ⟶ 0 ≤ u x" "0 < (c::real)"
shows "measure M {x∈A. u x ≥ c} ≤ (∫x∈A. u x ∂M) / c"
proof -
have nonneg: "set_lebesgue_integral M A u ≥ 0"
unfolding set_lebesgue_integral_def
by (intro Bochner_Integration.integral_nonneg_AE eventually_mono[OF assms(3)])
(auto simp: mult_ac)
have le: "emeasure M {x∈A. u x ≥ c} ≤ ennreal ((1/c) * (∫x∈A. u x ∂M))"
by (rule integral_Markov_inequality') (use assms in auto)
also have "… < top"
by auto
finally have "ennreal (measure M {x∈A. u x ≥ c}) = emeasure M {x∈A. u x ≥ c}"
by (intro emeasure_eq_ennreal_measure [symmetric]) auto
also note le
finally show ?thesis using nonneg
by (subst (asm) ennreal_le_iff)
(auto intro!: divide_nonneg_pos Bochner_Integration.integral_nonneg_AE assms)
qed
theorem%important (in finite_measure) Chernoff_ineq_ge:
assumes s: "s > 0"
assumes integrable: "set_integrable M A (λx. exp (s * f x))" and "A ∈ sets M"
shows "measure M {x∈A. f x ≥ a} ≤ exp (-s * a) * (∫x∈A. exp (s * f x) ∂M)"
proof -
have "{x∈A. f x ≥ a} = {x∈A. exp (s * f x) ≥ exp (s * a)}"
using s by auto
also have "measure M … ≤ set_lebesgue_integral M A (λx. exp (s * f x)) / exp (s * a)"
by (intro integral_Markov_inequality'_measure assms) auto
finally show ?thesis
by (simp add: exp_minus field_simps)
qed
theorem%important (in finite_measure) Chernoff_ineq_le:
assumes s: "s > 0"
assumes integrable: "set_integrable M A (λx. exp (-s * f x))" and "A ∈ sets M"
shows "measure M {x∈A. f x ≤ a} ≤ exp (s * a) * (∫x∈A. exp (-s * f x) ∂M)"
proof -
have "{x∈A. f x ≤ a} = {x∈A. exp (-s * f x) ≥ exp (-s * a)}"
using s by auto
also have "measure M … ≤ set_lebesgue_integral M A (λx. exp (-s * f x)) / exp (-s * a)"
by (intro integral_Markov_inequality'_measure assms) auto
finally show ?thesis
by (simp add: exp_minus field_simps)
qed
section ‹Integrable Simple Functions›
text ‹This section is from the Martingales AFP entry, by Ata Keskin, TU München›
text ‹We restate some basic results concerning Bochner-integrable functions.›
lemma integrable_implies_simple_function_sequence:
fixes f :: "'a ⇒ 'b::{banach, second_countable_topology}"
assumes "integrable M f"
obtains s where "⋀i. simple_function M (s i)"
and "⋀i. emeasure M {y ∈ space M. s i y ≠ 0} ≠ ∞"
and "⋀x. x ∈ space M ⟹ (λi. s i x) ⇢ f x"
and "⋀x i. x ∈ space M ⟹ norm (s i x) ≤ 2 * norm (f x)"
proof-
have f: "f ∈ borel_measurable M" "(∫⇧+x. norm (f x) ∂M) < ∞" using assms unfolding integrable_iff_bounded by auto
obtain s where s: "⋀i. simple_function M (s i)" "⋀x. x ∈ space M ⟹ (λi. s i x) ⇢ f x" "⋀i x. x ∈ space M ⟹ norm (s i x) ≤ 2 * norm (f x)" using borel_measurable_implies_sequence_metric[OF f(1)] unfolding norm_conv_dist by metis
{
fix i
have "(∫⇧+x. norm (s i x) ∂M) ≤ (∫⇧+x. ennreal (2 * norm (f x)) ∂M)" using s by (intro nn_integral_mono, auto)
also have "… < ∞" using f by (simp add: nn_integral_cmult ennreal_mult_less_top ennreal_mult)
finally have sbi: "Bochner_Integration.simple_bochner_integrable M (s i)" using s by (intro simple_bochner_integrableI_bounded) auto
hence "emeasure M {y ∈ space M. s i y ≠ 0} ≠ ∞" by (auto intro: integrableI_simple_bochner_integrable simple_bochner_integrable.cases)
}
thus ?thesis using that s by blast
qed
text ‹Simple functions can be represented by sums of indicator functions.›
lemma simple_function_indicator_representation_banach:
fixes f ::"'a ⇒ 'b :: {second_countable_topology, banach}"
assumes "simple_function M f" "x ∈ space M"
shows "f x = (∑y ∈ f ` space M. indicator (f -` {y} ∩ space M) x *⇩R y)"
(is "?l = ?r")
proof -
have "?r = (∑y ∈ f ` space M. (if y = f x then indicator (f -` {y} ∩ space M) x *⇩R y else 0))"
by (auto intro!: sum.cong)
also have "… = indicator (f -` {f x} ∩ space M) x *⇩R f x" using assms by (auto dest: simple_functionD)
also have "… = f x" using assms by (auto simp: indicator_def)
finally show ?thesis by auto
qed
lemma simple_function_indicator_representation_AE:
fixes f ::"'a ⇒ 'b :: {second_countable_topology, banach}"
assumes "simple_function M f"
shows "AE x in M. f x = (∑