reworked Probability theory: measures are not type restricted to positive extended reals
--- a/src/HOL/IsaMakefile Mon Mar 14 14:37:47 2011 +0100
+++ b/src/HOL/IsaMakefile Mon Mar 14 14:37:49 2011 +0100
@@ -1168,9 +1168,10 @@
Multivariate_Analysis/Topology_Euclidean_Space.thy \
Multivariate_Analysis/document/root.tex \
Multivariate_Analysis/normarith.ML Library/Glbs.thy \
- Library/Indicator_Function.thy Library/Inner_Product.thy \
- Library/Numeral_Type.thy Library/Convex.thy Library/FrechetDeriv.thy \
- Library/Product_Vector.thy Library/Product_plus.thy
+ Library/Extended_Reals.thy Library/Indicator_Function.thy \
+ Library/Inner_Product.thy Library/Numeral_Type.thy Library/Convex.thy \
+ Library/FrechetDeriv.thy Library/Product_Vector.thy \
+ Library/Product_plus.thy
@cd Multivariate_Analysis; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-Multivariate_Analysis
@@ -1185,7 +1186,6 @@
Probability/ex/Koepf_Duermuth_Countermeasure.thy \
Probability/Information.thy Probability/Lebesgue_Integration.thy \
Probability/Lebesgue_Measure.thy Probability/Measure.thy \
- Probability/Positive_Extended_Real.thy \
Probability/Probability_Space.thy Probability/Probability.thy \
Probability/Product_Measure.thy Probability/Radon_Nikodym.thy \
Probability/ROOT.ML Probability/Sigma_Algebra.thy \
--- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy Mon Mar 14 14:37:47 2011 +0100
+++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy Mon Mar 14 14:37:49 2011 +0100
@@ -1028,7 +1028,7 @@
qed simp
qed
-lemma setsum_of_pextreal:
+lemma setsum_real_of_extreal:
assumes "\<And>x. x \<in> S \<Longrightarrow> \<bar>f x\<bar> \<noteq> \<infinity>"
shows "(\<Sum>x\<in>S. real (f x)) = real (setsum f S)"
proof -
@@ -1062,17 +1062,6 @@
by induct (auto simp: extreal_right_distrib setsum_nonneg)
qed simp
-lemma setsum_real_of_extreal:
- assumes "\<And>x. x \<in> A \<Longrightarrow> \<bar>f x\<bar> \<noteq> \<infinity>"
- shows "real (\<Sum>x\<in>A. f x) = (\<Sum>x\<in>A. real (f x))"
-proof cases
- assume "finite A" from this assms show ?thesis
- proof induct
- case (insert a A) then show ?case
- by (simp add: real_of_extreal_add setsum_Inf)
- qed simp
-qed simp
-
lemma sums_extreal_positive:
fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>i. 0 \<le> f i" shows "f sums (SUP n. \<Sum>i<n. f i)"
proof -
--- a/src/HOL/Probability/Borel_Space.thy Mon Mar 14 14:37:47 2011 +0100
+++ b/src/HOL/Probability/Borel_Space.thy Mon Mar 14 14:37:49 2011 +0100
@@ -3,13 +3,9 @@
header {*Borel spaces*}
theory Borel_Space
- imports Sigma_Algebra Positive_Extended_Real Multivariate_Analysis
+ imports Sigma_Algebra Multivariate_Analysis
begin
-lemma LIMSEQ_max:
- "u ----> (x::real) \<Longrightarrow> (\<lambda>i. max (u i) 0) ----> max x 0"
- by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D)
-
section "Generic Borel spaces"
definition "borel = sigma \<lparr> space = UNIV::'a::topological_space set, sets = open\<rparr>"
@@ -112,26 +108,8 @@
ultimately show "?I \<in> borel_measurable M" by auto
qed
-lemma borel_measurable_translate:
- assumes "A \<in> sets borel" and trans: "\<And>B. open B \<Longrightarrow> f -` B \<in> sets borel"
- shows "f -` A \<in> sets borel"
-proof -
- have "A \<in> sigma_sets UNIV open" using assms
- by (simp add: borel_def sigma_def)
- thus ?thesis
- proof (induct rule: sigma_sets.induct)
- case (Basic a) thus ?case using trans[of a] by (simp add: mem_def)
- next
- case (Compl a)
- moreover have "UNIV \<in> sets borel"
- using borel.top by simp
- ultimately show ?case
- by (auto simp: vimage_Diff borel.Diff)
- qed (auto simp add: vimage_UN)
-qed
-
lemma (in sigma_algebra) borel_measurable_restricted:
- fixes f :: "'a \<Rightarrow> 'x\<Colon>{topological_space, semiring_1}" assumes "A \<in> sets M"
+ fixes f :: "'a \<Rightarrow> extreal" assumes "A \<in> sets M"
shows "f \<in> borel_measurable (restricted_space A) \<longleftrightarrow>
(\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
(is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M")
@@ -142,7 +120,7 @@
show ?thesis unfolding *
unfolding in_borel_measurable_borel
proof (simp, safe)
- fix S :: "'x set" assume "S \<in> sets borel"
+ fix S :: "extreal set" assume "S \<in> sets borel"
"\<forall>S\<in>sets borel. ?f -` S \<inter> A \<in> op \<inter> A ` sets M"
then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto
then have f: "?f -` S \<inter> A \<in> sets M"
@@ -161,7 +139,7 @@
then show ?thesis using f by auto
qed
next
- fix S :: "'x set" assume "S \<in> sets borel"
+ fix S :: "extreal set" assume "S \<in> sets borel"
"\<forall>S\<in>sets borel. ?f -` S \<inter> space M \<in> sets M"
then have f: "?f -` S \<inter> space M \<in> sets M" by auto
then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M"
@@ -1024,103 +1002,6 @@
using borel_measurable_euclidean_component
unfolding nth_conv_component by auto
-section "Borel space over the real line with infinity"
-
-lemma borel_Real_measurable:
- "A \<in> sets borel \<Longrightarrow> Real -` A \<in> sets borel"
-proof (rule borel_measurable_translate)
- fix B :: "pextreal set" assume "open B"
- then obtain T x where T: "open T" "Real ` (T \<inter> {0..}) = B - {\<omega>}" and
- x: "\<omega> \<in> B \<Longrightarrow> 0 \<le> x" "\<omega> \<in> B \<Longrightarrow> {Real x <..} \<subseteq> B"
- unfolding open_pextreal_def by blast
- have "Real -` B = Real -` (B - {\<omega>})" by auto
- also have "\<dots> = Real -` (Real ` (T \<inter> {0..}))" using T by simp
- also have "\<dots> = (if 0 \<in> T then T \<union> {.. 0} else T \<inter> {0..})"
- apply (auto simp add: Real_eq_Real image_iff)
- apply (rule_tac x="max 0 x" in bexI)
- by (auto simp: max_def)
- finally show "Real -` B \<in> sets borel"
- using `open T` by auto
-qed simp
-
-lemma borel_real_measurable:
- "A \<in> sets borel \<Longrightarrow> (real -` A :: pextreal set) \<in> sets borel"
-proof (rule borel_measurable_translate)
- fix B :: "real set" assume "open B"
- { fix x have "0 < real x \<longleftrightarrow> (\<exists>r>0. x = Real r)" by (cases x) auto }
- note Ex_less_real = this
- have *: "real -` B = (if 0 \<in> B then real -` (B \<inter> {0 <..}) \<union> {0, \<omega>} else real -` (B \<inter> {0 <..}))"
- by (force simp: Ex_less_real)
-
- have "open (real -` (B \<inter> {0 <..}) :: pextreal set)"
- unfolding open_pextreal_def using `open B`
- by (auto intro!: open_Int exI[of _ "B \<inter> {0 <..}"] simp: image_iff Ex_less_real)
- then show "(real -` B :: pextreal set) \<in> sets borel" unfolding * by auto
-qed simp
-
-lemma (in sigma_algebra) borel_measurable_Real[intro, simp]:
- assumes "f \<in> borel_measurable M"
- shows "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
- unfolding in_borel_measurable_borel
-proof safe
- fix S :: "pextreal set" assume "S \<in> sets borel"
- from borel_Real_measurable[OF this]
- have "(Real \<circ> f) -` S \<inter> space M \<in> sets M"
- using assms
- unfolding vimage_compose in_borel_measurable_borel
- by auto
- thus "(\<lambda>x. Real (f x)) -` S \<inter> space M \<in> sets M" by (simp add: comp_def)
-qed
-
-lemma (in sigma_algebra) borel_measurable_real[intro, simp]:
- fixes f :: "'a \<Rightarrow> pextreal"
- assumes "f \<in> borel_measurable M"
- shows "(\<lambda>x. real (f x)) \<in> borel_measurable M"
- unfolding in_borel_measurable_borel
-proof safe
- fix S :: "real set" assume "S \<in> sets borel"
- from borel_real_measurable[OF this]
- have "(real \<circ> f) -` S \<inter> space M \<in> sets M"
- using assms
- unfolding vimage_compose in_borel_measurable_borel
- by auto
- thus "(\<lambda>x. real (f x)) -` S \<inter> space M \<in> sets M" by (simp add: comp_def)
-qed
-
-lemma (in sigma_algebra) borel_measurable_Real_eq:
- assumes "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
- shows "(\<lambda>x. Real (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
-proof
- have [simp]: "(\<lambda>x. Real (f x)) -` {\<omega>} \<inter> space M = {}"
- by auto
- assume "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
- hence "(\<lambda>x. real (Real (f x))) \<in> borel_measurable M"
- by (rule borel_measurable_real)
- moreover have "\<And>x. x \<in> space M \<Longrightarrow> real (Real (f x)) = f x"
- using assms by auto
- ultimately show "f \<in> borel_measurable M"
- by (simp cong: measurable_cong)
-qed auto
-
-lemma (in sigma_algebra) borel_measurable_pextreal_eq_real:
- "f \<in> borel_measurable M \<longleftrightarrow>
- ((\<lambda>x. real (f x)) \<in> borel_measurable M \<and> f -` {\<omega>} \<inter> space M \<in> sets M)"
-proof safe
- assume "f \<in> borel_measurable M"
- then show "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<omega>} \<inter> space M \<in> sets M"
- by (auto intro: borel_measurable_vimage borel_measurable_real)
-next
- assume *: "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<omega>} \<inter> space M \<in> sets M"
- have "f -` {\<omega>} \<inter> space M = {x\<in>space M. f x = \<omega>}" by auto
- with * have **: "{x\<in>space M. f x = \<omega>} \<in> sets M" by simp
- have f: "f = (\<lambda>x. if f x = \<omega> then \<omega> else Real (real (f x)))"
- by (simp add: fun_eq_iff Real_real)
- show "f \<in> borel_measurable M"
- apply (subst f)
- apply (rule measurable_If)
- using * ** by auto
-qed
-
lemma borel_measurable_continuous_on1:
fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space"
assumes "continuous_on UNIV f"
@@ -1187,206 +1068,213 @@
using measurable_comp[OF f borel_measurable_borel_log[OF `1 < b`]]
by (simp add: comp_def)
+subsection "Borel space on the extended reals"
+
+lemma borel_measurable_extreal_borel:
+ "extreal \<in> borel_measurable borel"
+ unfolding borel_def[where 'a=extreal]
+proof (rule borel.measurable_sigma)
+ fix X :: "extreal set" assume "X \<in> sets \<lparr>space = UNIV, sets = open \<rparr>"
+ then have "open X" by (auto simp: mem_def)
+ then have "open (extreal -` X \<inter> space borel)"
+ by (simp add: open_extreal_vimage)
+ then show "extreal -` X \<inter> space borel \<in> sets borel" by auto
+qed auto
+
+lemma (in sigma_algebra) borel_measurable_extreal[simp, intro]:
+ assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. extreal (f x)) \<in> borel_measurable M"
+ using measurable_comp[OF f borel_measurable_extreal_borel] unfolding comp_def .
+
+lemma borel_measurable_real_of_extreal_borel:
+ "(real :: extreal \<Rightarrow> real) \<in> borel_measurable borel"
+ unfolding borel_def[where 'a=real]
+proof (rule borel.measurable_sigma)
+ fix B :: "real set" assume "B \<in> sets \<lparr>space = UNIV, sets = open \<rparr>"
+ then have "open B" by (auto simp: mem_def)
+ have *: "extreal -` real -` (B - {0}) = B - {0}" by auto
+ have open_real: "open (real -` (B - {0}) :: extreal set)"
+ unfolding open_extreal_def * using `open B` by auto
+ show "(real -` B \<inter> space borel :: extreal set) \<in> sets borel"
+ proof cases
+ assume "0 \<in> B"
+ then have *: "real -` B = real -` (B - {0}) \<union> {-\<infinity>, \<infinity>, 0}"
+ by (auto simp add: real_of_extreal_eq_0)
+ then show "(real -` B :: extreal set) \<inter> space borel \<in> sets borel"
+ using open_real by auto
+ next
+ assume "0 \<notin> B"
+ then have *: "(real -` B :: extreal set) = real -` (B - {0})"
+ by (auto simp add: real_of_extreal_eq_0)
+ then show "(real -` B :: extreal set) \<inter> space borel \<in> sets borel"
+ using open_real by auto
+ qed
+qed auto
+
+lemma (in sigma_algebra) borel_measurable_real_of_extreal[simp, intro]:
+ assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. real (f x :: extreal)) \<in> borel_measurable M"
+ using measurable_comp[OF f borel_measurable_real_of_extreal_borel] unfolding comp_def .
+
+lemma (in sigma_algebra) borel_measurable_extreal_iff:
+ shows "(\<lambda>x. extreal (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
+proof
+ assume "(\<lambda>x. extreal (f x)) \<in> borel_measurable M"
+ from borel_measurable_real_of_extreal[OF this]
+ show "f \<in> borel_measurable M" by auto
+qed auto
+
+lemma (in sigma_algebra) borel_measurable_extreal_iff_real:
+ "f \<in> borel_measurable M \<longleftrightarrow>
+ ((\<lambda>x. real (f x)) \<in> borel_measurable M \<and> f -` {\<infinity>} \<inter> space M \<in> sets M \<and> f -` {-\<infinity>} \<inter> space M \<in> sets M)"
+proof safe
+ assume *: "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<infinity>} \<inter> space M \<in> sets M" "f -` {-\<infinity>} \<inter> space M \<in> sets M"
+ have "f -` {\<infinity>} \<inter> space M = {x\<in>space M. f x = \<infinity>}" "f -` {-\<infinity>} \<inter> space M = {x\<in>space M. f x = -\<infinity>}" by auto
+ with * have **: "{x\<in>space M. f x = \<infinity>} \<in> sets M" "{x\<in>space M. f x = -\<infinity>} \<in> sets M" by simp_all
+ let "?f x" = "if f x = \<infinity> then \<infinity> else if f x = -\<infinity> then -\<infinity> else extreal (real (f x))"
+ have "?f \<in> borel_measurable M" using * ** by (intro measurable_If) auto
+ also have "?f = f" by (auto simp: fun_eq_iff extreal_real)
+ finally show "f \<in> borel_measurable M" .
+qed (auto intro: measurable_sets borel_measurable_real_of_extreal)
lemma (in sigma_algebra) less_eq_ge_measurable:
fixes f :: "'a \<Rightarrow> 'c::linorder"
- shows "{x\<in>space M. a < f x} \<in> sets M \<longleftrightarrow> {x\<in>space M. f x \<le> a} \<in> sets M"
+ shows "f -` {a <..} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {..a} \<inter> space M \<in> sets M"
proof
- assume "{x\<in>space M. f x \<le> a} \<in> sets M"
- moreover have "{x\<in>space M. a < f x} = space M - {x\<in>space M. f x \<le> a}" by auto
- ultimately show "{x\<in>space M. a < f x} \<in> sets M" by auto
+ assume "f -` {a <..} \<inter> space M \<in> sets M"
+ moreover have "f -` {..a} \<inter> space M = space M - f -` {a <..} \<inter> space M" by auto
+ ultimately show "f -` {..a} \<inter> space M \<in> sets M" by auto
next
- assume "{x\<in>space M. a < f x} \<in> sets M"
- moreover have "{x\<in>space M. f x \<le> a} = space M - {x\<in>space M. a < f x}" by auto
- ultimately show "{x\<in>space M. f x \<le> a} \<in> sets M" by auto
+ assume "f -` {..a} \<inter> space M \<in> sets M"
+ moreover have "f -` {a <..} \<inter> space M = space M - f -` {..a} \<inter> space M" by auto
+ ultimately show "f -` {a <..} \<inter> space M \<in> sets M" by auto
qed
lemma (in sigma_algebra) greater_eq_le_measurable:
fixes f :: "'a \<Rightarrow> 'c::linorder"
- shows "{x\<in>space M. f x < a} \<in> sets M \<longleftrightarrow> {x\<in>space M. a \<le> f x} \<in> sets M"
+ shows "f -` {..< a} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {a ..} \<inter> space M \<in> sets M"
proof
- assume "{x\<in>space M. a \<le> f x} \<in> sets M"
- moreover have "{x\<in>space M. f x < a} = space M - {x\<in>space M. a \<le> f x}" by auto
- ultimately show "{x\<in>space M. f x < a} \<in> sets M" by auto
+ assume "f -` {a ..} \<inter> space M \<in> sets M"
+ moreover have "f -` {..< a} \<inter> space M = space M - f -` {a ..} \<inter> space M" by auto
+ ultimately show "f -` {..< a} \<inter> space M \<in> sets M" by auto
next
- assume "{x\<in>space M. f x < a} \<in> sets M"
- moreover have "{x\<in>space M. a \<le> f x} = space M - {x\<in>space M. f x < a}" by auto
- ultimately show "{x\<in>space M. a \<le> f x} \<in> sets M" by auto
+ assume "f -` {..< a} \<inter> space M \<in> sets M"
+ moreover have "f -` {a ..} \<inter> space M = space M - f -` {..< a} \<inter> space M" by auto
+ ultimately show "f -` {a ..} \<inter> space M \<in> sets M" by auto
qed
-lemma (in sigma_algebra) less_eq_le_pextreal_measurable:
- fixes f :: "'a \<Rightarrow> pextreal"
- shows "(\<forall>a. {x\<in>space M. a < f x} \<in> sets M) \<longleftrightarrow> (\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M)"
+lemma (in sigma_algebra) borel_measurable_uminus_borel_extreal:
+ "(uminus :: extreal \<Rightarrow> extreal) \<in> borel_measurable borel"
+proof (subst borel_def, rule borel.measurable_sigma)
+ fix X :: "extreal set" assume "X \<in> sets \<lparr>space = UNIV, sets = open\<rparr>"
+ then have "open X" by (simp add: mem_def)
+ have "uminus -` X = uminus ` X" by (force simp: image_iff)
+ then have "open (uminus -` X)" using `open X` extreal_open_uminus by auto
+ then show "uminus -` X \<inter> space borel \<in> sets borel" by auto
+qed auto
+
+lemma (in sigma_algebra) borel_measurable_uminus_extreal[intro]:
+ assumes "f \<in> borel_measurable M"
+ shows "(\<lambda>x. - f x :: extreal) \<in> borel_measurable M"
+ using measurable_comp[OF assms borel_measurable_uminus_borel_extreal] by (simp add: comp_def)
+
+lemma (in sigma_algebra) borel_measurable_uminus_eq_extreal[simp]:
+ "(\<lambda>x. - f x :: extreal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
proof
- assume a: "\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M"
- show "\<forall>a. {x \<in> space M. a < f x} \<in> sets M"
- proof
- fix a show "{x \<in> space M. a < f x} \<in> sets M"
- proof (cases a)
- case (preal r)
- have "{x\<in>space M. a < f x} = (\<Union>i. {x\<in>space M. a + inverse (of_nat (Suc i)) \<le> f x})"
- proof safe
- fix x assume "a < f x" and [simp]: "x \<in> space M"
- with ex_pextreal_inverse_of_nat_Suc_less[of "f x - a"]
- obtain n where "a + inverse (of_nat (Suc n)) < f x"
- by (cases "f x", auto simp: pextreal_minus_order)
- then have "a + inverse (of_nat (Suc n)) \<le> f x" by simp
- then show "x \<in> (\<Union>i. {x \<in> space M. a + inverse (of_nat (Suc i)) \<le> f x})"
- by auto
- next
- fix i x assume [simp]: "x \<in> space M"
- have "a < a + inverse (of_nat (Suc i))" using preal by auto
- also assume "a + inverse (of_nat (Suc i)) \<le> f x"
- finally show "a < f x" .
- qed
- with a show ?thesis by auto
- qed simp
+ assume ?l from borel_measurable_uminus_extreal[OF this] show ?r by simp
+qed auto
+
+lemma (in sigma_algebra) borel_measurable_eq_atMost_extreal:
+ "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
+proof (intro iffI allI)
+ assume pos[rule_format]: "\<forall>a. f -` {..a} \<inter> space M \<in> sets M"
+ show "f \<in> borel_measurable M"
+ unfolding borel_measurable_extreal_iff_real borel_measurable_iff_le
+ proof (intro conjI allI)
+ fix a :: real
+ { fix x :: extreal assume *: "\<forall>i::nat. real i < x"
+ have "x = \<infinity>"
+ proof (rule extreal_top)
+ fix B from real_arch_lt[of B] guess n ..
+ then have "extreal B < real n" by auto
+ with * show "B \<le> x" by (metis less_trans less_imp_le)
+ qed }
+ then have "f -` {\<infinity>} \<inter> space M = space M - (\<Union>i::nat. f -` {.. real i} \<inter> space M)"
+ by (auto simp: not_le)
+ then show "f -` {\<infinity>} \<inter> space M \<in> sets M" using pos by (auto simp del: UN_simps intro!: Diff)
+ moreover
+ have "{-\<infinity>} = {..-\<infinity>}" by auto
+ then show "f -` {-\<infinity>} \<inter> space M \<in> sets M" using pos by auto
+ moreover have "{x\<in>space M. f x \<le> extreal a} \<in> sets M"
+ using pos[of "extreal a"] by (simp add: vimage_def Int_def conj_commute)
+ moreover have "{w \<in> space M. real (f w) \<le> a} =
+ (if a < 0 then {w \<in> space M. f w \<le> extreal a} - f -` {-\<infinity>} \<inter> space M
+ else {w \<in> space M. f w \<le> extreal a} \<union> (f -` {\<infinity>} \<inter> space M) \<union> (f -` {-\<infinity>} \<inter> space M))" (is "?l = ?r")
+ proof (intro set_eqI) fix x show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" by (cases "f x") auto qed
+ ultimately show "{w \<in> space M. real (f w) \<le> a} \<in> sets M" by auto
qed
-next
- assume a': "\<forall>a. {x \<in> space M. a < f x} \<in> sets M"
- then have a: "\<forall>a. {x \<in> space M. f x \<le> a} \<in> sets M" unfolding less_eq_ge_measurable .
- show "\<forall>a. {x \<in> space M. a \<le> f x} \<in> sets M" unfolding greater_eq_le_measurable[symmetric]
- proof
- fix a show "{x \<in> space M. f x < a} \<in> sets M"
- proof (cases a)
- case (preal r)
- show ?thesis
- proof cases
- assume "a = 0" then show ?thesis by simp
- next
- assume "a \<noteq> 0"
- have "{x\<in>space M. f x < a} = (\<Union>i. {x\<in>space M. f x \<le> a - inverse (of_nat (Suc i))})"
- proof safe
- fix x assume "f x < a" and [simp]: "x \<in> space M"
- with ex_pextreal_inverse_of_nat_Suc_less[of "a - f x"]
- obtain n where "inverse (of_nat (Suc n)) < a - f x"
- using preal by (cases "f x") auto
- then have "f x \<le> a - inverse (of_nat (Suc n)) "
- using preal by (cases "f x") (auto split: split_if_asm)
- then show "x \<in> (\<Union>i. {x \<in> space M. f x \<le> a - inverse (of_nat (Suc i))})"
- by auto
- next
- fix i x assume [simp]: "x \<in> space M"
- assume "f x \<le> a - inverse (of_nat (Suc i))"
- also have "\<dots> < a" using `a \<noteq> 0` preal by auto
- finally show "f x < a" .
- qed
- with a show ?thesis by auto
- qed
- next
- case infinite
- have "f -` {\<omega>} \<inter> space M = (\<Inter>n. {x\<in>space M. of_nat n < f x})"
- proof (safe, simp_all, safe)
- fix x assume *: "\<forall>n::nat. Real (real n) < f x"
- show "f x = \<omega>" proof (rule ccontr)
- assume "f x \<noteq> \<omega>"
- with real_arch_lt[of "real (f x)"] obtain n where "f x < of_nat n"
- by (auto simp: pextreal_noteq_omega_Ex)
- with *[THEN spec, of n] show False by auto
- qed
- qed
- with a' have \<omega>: "f -` {\<omega>} \<inter> space M \<in> sets M" by auto
- moreover have "{x \<in> space M. f x < a} = space M - f -` {\<omega>} \<inter> space M"
- using infinite by auto
- ultimately show ?thesis by auto
- qed
- qed
-qed
+qed (simp add: measurable_sets)
-lemma (in sigma_algebra) borel_measurable_pextreal_iff_greater:
- "(f::'a \<Rightarrow> pextreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a < f x} \<in> sets M)"
-proof safe
- fix a assume f: "f \<in> borel_measurable M"
- have "{x\<in>space M. a < f x} = f -` {a <..} \<inter> space M" by auto
- with f show "{x\<in>space M. a < f x} \<in> sets M"
- by (auto intro!: measurable_sets)
-next
- assume *: "\<forall>a. {x\<in>space M. a < f x} \<in> sets M"
- hence **: "\<forall>a. {x\<in>space M. f x < a} \<in> sets M"
- unfolding less_eq_le_pextreal_measurable
- unfolding greater_eq_le_measurable .
- show "f \<in> borel_measurable M" unfolding borel_measurable_pextreal_eq_real borel_measurable_iff_greater
- proof safe
- have "f -` {\<omega>} \<inter> space M = space M - {x\<in>space M. f x < \<omega>}" by auto
- then show \<omega>: "f -` {\<omega>} \<inter> space M \<in> sets M" using ** by auto
- fix a
- have "{w \<in> space M. a < real (f w)} =
- (if 0 \<le> a then {w\<in>space M. Real a < f w} - (f -` {\<omega>} \<inter> space M) else space M)"
- proof (split split_if, safe del: notI)
- fix x assume "0 \<le> a"
- { assume "a < real (f x)" then show "Real a < f x" "x \<notin> f -` {\<omega>} \<inter> space M"
- using `0 \<le> a` by (cases "f x", auto) }
- { assume "Real a < f x" "x \<notin> f -` {\<omega>}" then show "a < real (f x)"
- using `0 \<le> a` by (cases "f x", auto) }
- next
- fix x assume "\<not> 0 \<le> a" then show "a < real (f x)" by (cases "f x") auto
- qed
- then show "{w \<in> space M. a < real (f w)} \<in> sets M"
- using \<omega> * by (auto intro!: Diff)
- qed
-qed
+lemma (in sigma_algebra) borel_measurable_eq_atLeast_extreal:
+ "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)"
+proof
+ assume pos: "\<forall>a. f -` {a..} \<inter> space M \<in> sets M"
+ moreover have "\<And>a. (\<lambda>x. - f x) -` {..a} = f -` {-a ..}"
+ by (auto simp: extreal_uminus_le_reorder)
+ ultimately have "(\<lambda>x. - f x) \<in> borel_measurable M"
+ unfolding borel_measurable_eq_atMost_extreal by auto
+ then show "f \<in> borel_measurable M" by simp
+qed (simp add: measurable_sets)
-lemma (in sigma_algebra) borel_measurable_pextreal_iff_less:
- "(f::'a \<Rightarrow> pextreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x < a} \<in> sets M)"
- using borel_measurable_pextreal_iff_greater unfolding less_eq_le_pextreal_measurable greater_eq_le_measurable .
+lemma (in sigma_algebra) borel_measurable_extreal_iff_less:
+ "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
+ unfolding borel_measurable_eq_atLeast_extreal greater_eq_le_measurable ..
-lemma (in sigma_algebra) borel_measurable_pextreal_iff_le:
- "(f::'a \<Rightarrow> pextreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x \<le> a} \<in> sets M)"
- using borel_measurable_pextreal_iff_greater unfolding less_eq_ge_measurable .
+lemma (in sigma_algebra) borel_measurable_extreal_iff_ge:
+ "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
+ unfolding borel_measurable_eq_atMost_extreal less_eq_ge_measurable ..
-lemma (in sigma_algebra) borel_measurable_pextreal_iff_ge:
- "(f::'a \<Rightarrow> pextreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M)"
- using borel_measurable_pextreal_iff_greater unfolding less_eq_le_pextreal_measurable .
-
-lemma (in sigma_algebra) borel_measurable_pextreal_eq_const:
- fixes f :: "'a \<Rightarrow> pextreal" assumes "f \<in> borel_measurable M"
+lemma (in sigma_algebra) borel_measurable_extreal_eq_const:
+ fixes f :: "'a \<Rightarrow> extreal" assumes "f \<in> borel_measurable M"
shows "{x\<in>space M. f x = c} \<in> sets M"
proof -
have "{x\<in>space M. f x = c} = (f -` {c} \<inter> space M)" by auto
then show ?thesis using assms by (auto intro!: measurable_sets)
qed
-lemma (in sigma_algebra) borel_measurable_pextreal_neq_const:
- fixes f :: "'a \<Rightarrow> pextreal"
- assumes "f \<in> borel_measurable M"
+lemma (in sigma_algebra) borel_measurable_extreal_neq_const:
+ fixes f :: "'a \<Rightarrow> extreal" assumes "f \<in> borel_measurable M"
shows "{x\<in>space M. f x \<noteq> c} \<in> sets M"
proof -
have "{x\<in>space M. f x \<noteq> c} = space M - (f -` {c} \<inter> space M)" by auto
then show ?thesis using assms by (auto intro!: measurable_sets)
qed
-lemma (in sigma_algebra) borel_measurable_pextreal_less[intro,simp]:
- fixes f g :: "'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_extreal_le[intro,simp]:
+ fixes f g :: "'a \<Rightarrow> extreal"
+ assumes f: "f \<in> borel_measurable M"
+ assumes g: "g \<in> borel_measurable M"
+ shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
+proof -
+ have "{x \<in> space M. f x \<le> g x} =
+ {x \<in> space M. real (f x) \<le> real (g x)} - (f -` {\<infinity>, -\<infinity>} \<inter> space M \<union> g -` {\<infinity>, -\<infinity>} \<inter> space M) \<union>
+ f -` {-\<infinity>} \<inter> space M \<union> g -` {\<infinity>} \<inter> space M" (is "?l = ?r")
+ proof (intro set_eqI)
+ fix x show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" by (cases rule: extreal2_cases[of "f x" "g x"]) auto
+ qed
+ with f g show ?thesis by (auto intro!: Un simp: measurable_sets)
+qed
+
+lemma (in sigma_algebra) borel_measurable_extreal_less[intro,simp]:
+ fixes f :: "'a \<Rightarrow> extreal"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "{x \<in> space M. f x < g x} \<in> sets M"
proof -
- have "(\<lambda>x. real (f x)) \<in> borel_measurable M"
- "(\<lambda>x. real (g x)) \<in> borel_measurable M"
- using assms by (auto intro!: borel_measurable_real)
- from borel_measurable_less[OF this]
- have "{x \<in> space M. real (f x) < real (g x)} \<in> sets M" .
- moreover have "{x \<in> space M. f x \<noteq> \<omega>} \<in> sets M" using f by (rule borel_measurable_pextreal_neq_const)
- moreover have "{x \<in> space M. g x = \<omega>} \<in> sets M" using g by (rule borel_measurable_pextreal_eq_const)
- moreover have "{x \<in> space M. g x \<noteq> \<omega>} \<in> sets M" using g by (rule borel_measurable_pextreal_neq_const)
- moreover have "{x \<in> space M. f x < g x} = ({x \<in> space M. g x = \<omega>} \<inter> {x \<in> space M. f x \<noteq> \<omega>}) \<union>
- ({x \<in> space M. g x \<noteq> \<omega>} \<inter> {x \<in> space M. f x \<noteq> \<omega>} \<inter> {x \<in> space M. real (f x) < real (g x)})"
- by (auto simp: real_of_pextreal_strict_mono_iff)
- ultimately show ?thesis by auto
-qed
-
-lemma (in sigma_algebra) borel_measurable_pextreal_le[intro,simp]:
- fixes f :: "'a \<Rightarrow> pextreal"
- assumes f: "f \<in> borel_measurable M"
- assumes g: "g \<in> borel_measurable M"
- shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
-proof -
- have "{x \<in> space M. f x \<le> g x} = space M - {x \<in> space M. g x < f x}" by auto
+ have "{x \<in> space M. f x < g x} = space M - {x \<in> space M. g x \<le> f x}" by auto
then show ?thesis using g f by auto
qed
-lemma (in sigma_algebra) borel_measurable_pextreal_eq[intro,simp]:
- fixes f :: "'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_extreal_eq[intro,simp]:
+ fixes f :: "'a \<Rightarrow> extreal"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "{w \<in> space M. f w = g w} \<in> sets M"
@@ -1395,8 +1283,8 @@
then show ?thesis using g f by auto
qed
-lemma (in sigma_algebra) borel_measurable_pextreal_neq[intro,simp]:
- fixes f :: "'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_extreal_neq[intro,simp]:
+ fixes f :: "'a \<Rightarrow> extreal"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
@@ -1405,20 +1293,28 @@
thus ?thesis using f g by auto
qed
-lemma (in sigma_algebra) borel_measurable_pextreal_add[intro, simp]:
- fixes f :: "'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) split_sets:
+ "{x\<in>space M. P x \<or> Q x} = {x\<in>space M. P x} \<union> {x\<in>space M. Q x}"
+ "{x\<in>space M. P x \<and> Q x} = {x\<in>space M. P x} \<inter> {x\<in>space M. Q x}"
+ by auto
+
+lemma (in sigma_algebra) borel_measurable_extreal_add[intro, simp]:
+ fixes f :: "'a \<Rightarrow> extreal"
assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
proof -
- have *: "(\<lambda>x. f x + g x) =
- (\<lambda>x. if f x = \<omega> then \<omega> else if g x = \<omega> then \<omega> else Real (real (f x) + real (g x)))"
- by (auto simp: fun_eq_iff pextreal_noteq_omega_Ex)
- show ?thesis using assms unfolding *
- by (auto intro!: measurable_If)
+ { fix x assume "x \<in> space M" then have "f x + g x =
+ (if f x = \<infinity> \<or> g x = \<infinity> then \<infinity>
+ else if f x = -\<infinity> \<or> g x = -\<infinity> then -\<infinity>
+ else extreal (real (f x) + real (g x)))"
+ by (cases rule: extreal2_cases[of "f x" "g x"]) auto }
+ with assms show ?thesis
+ by (auto cong: measurable_cong simp: split_sets
+ intro!: Un measurable_If measurable_sets)
qed
-lemma (in sigma_algebra) borel_measurable_pextreal_setsum[simp, intro]:
- fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_extreal_setsum[simp, intro]:
+ fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> extreal"
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
proof cases
@@ -1427,20 +1323,49 @@
by induct auto
qed (simp add: borel_measurable_const)
-lemma (in sigma_algebra) borel_measurable_pextreal_times[intro, simp]:
- fixes f :: "'a \<Rightarrow> pextreal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+lemma abs_extreal_ge0[simp]: "0 \<le> x \<Longrightarrow> \<bar>x :: extreal\<bar> = x"
+ by (cases x) auto
+
+lemma abs_extreal_less0[simp]: "x < 0 \<Longrightarrow> \<bar>x :: extreal\<bar> = -x"
+ by (cases x) auto
+
+lemma abs_extreal_pos[simp]: "0 \<le> \<bar>x :: extreal\<bar>"
+ by (cases x) auto
+
+lemma (in sigma_algebra) borel_measurable_extreal_abs[intro, simp]:
+ fixes f :: "'a \<Rightarrow> extreal" assumes "f \<in> borel_measurable M"
+ shows "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M"
+proof -
+ { fix x have "\<bar>f x\<bar> = (if 0 \<le> f x then f x else - f x)" by auto }
+ then show ?thesis using assms by (auto intro!: measurable_If)
+qed
+
+lemma (in sigma_algebra) borel_measurable_extreal_times[intro, simp]:
+ fixes f :: "'a \<Rightarrow> extreal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
proof -
+ { fix f g :: "'a \<Rightarrow> extreal"
+ assume b: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+ and pos: "\<And>x. 0 \<le> f x" "\<And>x. 0 \<le> g x"
+ { fix x have *: "f x * g x = (if f x = 0 \<or> g x = 0 then 0
+ else if f x = \<infinity> \<or> g x = \<infinity> then \<infinity>
+ else extreal (real (f x) * real (g x)))"
+ apply (cases rule: extreal2_cases[of "f x" "g x"])
+ using pos[of x] by auto }
+ with b have "(\<lambda>x. f x * g x) \<in> borel_measurable M"
+ by (auto cong: measurable_cong simp: split_sets
+ intro!: Un measurable_If measurable_sets) }
+ note pos_times = this
have *: "(\<lambda>x. f x * g x) =
- (\<lambda>x. if f x = 0 then 0 else if g x = 0 then 0 else if f x = \<omega> then \<omega> else if g x = \<omega> then \<omega> else
- Real (real (f x) * real (g x)))"
- by (auto simp: fun_eq_iff pextreal_noteq_omega_Ex)
+ (\<lambda>x. if 0 \<le> f x \<and> 0 \<le> g x \<or> f x < 0 \<and> g x < 0 then \<bar>f x\<bar> * \<bar>g x\<bar> else - (\<bar>f x\<bar> * \<bar>g x\<bar>))"
+ by (auto simp: fun_eq_iff)
show ?thesis using assms unfolding *
- by (auto intro!: measurable_If)
+ by (intro measurable_If pos_times borel_measurable_uminus_extreal)
+ (auto simp: split_sets intro!: Int)
qed
-lemma (in sigma_algebra) borel_measurable_pextreal_setprod[simp, intro]:
- fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_extreal_setprod[simp, intro]:
+ fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> extreal"
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
proof cases
@@ -1448,64 +1373,73 @@
thus ?thesis using assms by induct auto
qed simp
-lemma (in sigma_algebra) borel_measurable_pextreal_min[simp, intro]:
- fixes f g :: "'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_extreal_min[simp, intro]:
+ fixes f g :: "'a \<Rightarrow> extreal"
assumes "f \<in> borel_measurable M"
assumes "g \<in> borel_measurable M"
shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"
using assms unfolding min_def by (auto intro!: measurable_If)
-lemma (in sigma_algebra) borel_measurable_pextreal_max[simp, intro]:
- fixes f g :: "'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_extreal_max[simp, intro]:
+ fixes f g :: "'a \<Rightarrow> extreal"
assumes "f \<in> borel_measurable M"
and "g \<in> borel_measurable M"
shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"
using assms unfolding max_def by (auto intro!: measurable_If)
lemma (in sigma_algebra) borel_measurable_SUP[simp, intro]:
- fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> pextreal"
+ fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> extreal"
assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
shows "(\<lambda>x. SUP i : A. f i x) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M")
- unfolding borel_measurable_pextreal_iff_greater
-proof safe
+ unfolding borel_measurable_extreal_iff_ge
+proof
fix a
- have "{x\<in>space M. a < ?sup x} = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
+ have "?sup -` {a<..} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
by (auto simp: less_SUP_iff SUPR_apply)
- then show "{x\<in>space M. a < ?sup x} \<in> sets M"
+ then show "?sup -` {a<..} \<inter> space M \<in> sets M"
using assms by auto
qed
lemma (in sigma_algebra) borel_measurable_INF[simp, intro]:
- fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> pextreal"
+ fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> extreal"
assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
shows "(\<lambda>x. INF i : A. f i x) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M")
- unfolding borel_measurable_pextreal_iff_less
-proof safe
+ unfolding borel_measurable_extreal_iff_less
+proof
fix a
- have "{x\<in>space M. ?inf x < a} = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
+ have "?inf -` {..<a} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
by (auto simp: INF_less_iff INFI_apply)
- then show "{x\<in>space M. ?inf x < a} \<in> sets M"
+ then show "?inf -` {..<a} \<inter> space M \<in> sets M"
using assms by auto
qed
-lemma (in sigma_algebra) borel_measurable_pextreal_diff[simp, intro]:
- fixes f g :: "'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_liminf[simp, intro]:
+ fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
+ assumes "\<And>i. f i \<in> borel_measurable M"
+ shows "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M"
+ unfolding liminf_SUPR_INFI using assms by auto
+
+lemma (in sigma_algebra) borel_measurable_limsup[simp, intro]:
+ fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
+ assumes "\<And>i. f i \<in> borel_measurable M"
+ shows "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M"
+ unfolding limsup_INFI_SUPR using assms by auto
+
+lemma (in sigma_algebra) borel_measurable_extreal_diff[simp, intro]:
+ fixes f g :: "'a \<Rightarrow> extreal"
assumes "f \<in> borel_measurable M"
assumes "g \<in> borel_measurable M"
shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
- unfolding borel_measurable_pextreal_iff_greater
-proof safe
- fix a
- have "{x \<in> space M. a < f x - g x} = {x \<in> space M. g x + a < f x}"
- by (simp add: pextreal_less_minus_iff)
- then show "{x \<in> space M. a < f x - g x} \<in> sets M"
- using assms by auto
-qed
+ unfolding minus_extreal_def using assms by auto
lemma (in sigma_algebra) borel_measurable_psuminf[simp, intro]:
- assumes "\<And>i. f i \<in> borel_measurable M"
- shows "(\<lambda>x. (\<Sum>\<^isub>\<infinity> i. f i x)) \<in> borel_measurable M"
- using assms unfolding psuminf_def by auto
+ fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
+ assumes "\<And>i. f i \<in> borel_measurable M" and pos: "\<And>i x. x \<in> space M \<Longrightarrow> 0 \<le> f i x"
+ shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M"
+ apply (subst measurable_cong)
+ apply (subst suminf_extreal_eq_SUPR)
+ apply (rule pos)
+ using assms by auto
section "LIMSEQ is borel measurable"
@@ -1515,28 +1449,11 @@
and u: "\<And>i. u i \<in> borel_measurable M"
shows "u' \<in> borel_measurable M"
proof -
- let "?pu x i" = "max (u i x) 0"
- let "?nu x i" = "max (- u i x) 0"
- { fix x assume x: "x \<in> space M"
- have "(?pu x) ----> max (u' x) 0"
- "(?nu x) ----> max (- u' x) 0"
- using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus)
- from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)]
- have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)"
- "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)"
- by (simp_all add: Real_max'[symmetric]) }
- note eq = this
- have *: "\<And>x. real (Real (u' x)) - real (Real (- u' x)) = u' x"
+ have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. extreal (u n x)) = extreal (u' x)"
+ using u' by (simp add: lim_imp_Liminf trivial_limit_sequentially lim_extreal)
+ moreover from u have "(\<lambda>x. liminf (\<lambda>n. extreal (u n x))) \<in> borel_measurable M"
by auto
- have "(\<lambda>x. SUP n. INF m. Real (u (n + m) x)) \<in> borel_measurable M"
- "(\<lambda>x. SUP n. INF m. Real (- u (n + m) x)) \<in> borel_measurable M"
- using u by auto
- with eq[THEN measurable_cong, of M "\<lambda>x. x" borel]
- have "(\<lambda>x. Real (u' x)) \<in> borel_measurable M"
- "(\<lambda>x. Real (- u' x)) \<in> borel_measurable M" by auto
- note this[THEN borel_measurable_real]
- from borel_measurable_diff[OF this]
- show ?thesis unfolding * .
+ ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_extreal_iff)
qed
end
--- a/src/HOL/Probability/Caratheodory.thy Mon Mar 14 14:37:47 2011 +0100
+++ b/src/HOL/Probability/Caratheodory.thy Mon Mar 14 14:37:49 2011 +0100
@@ -1,36 +1,66 @@
header {*Caratheodory Extension Theorem*}
theory Caratheodory
- imports Sigma_Algebra Positive_Extended_Real
+ imports Sigma_Algebra Extended_Real_Limits
begin
+lemma suminf_extreal_2dimen:
+ fixes f:: "nat \<times> nat \<Rightarrow> extreal"
+ assumes pos: "\<And>p. 0 \<le> f p"
+ assumes "\<And>m. g m = (\<Sum>n. f (m,n))"
+ shows "(\<Sum>i. f (prod_decode i)) = suminf g"
+proof -
+ have g_def: "g = (\<lambda>m. (\<Sum>n. f (m,n)))"
+ using assms by (simp add: fun_eq_iff)
+ have reindex: "\<And>B. (\<Sum>x\<in>B. f (prod_decode x)) = setsum f (prod_decode ` B)"
+ by (simp add: setsum_reindex[OF inj_prod_decode] comp_def)
+ { fix n
+ let ?M = "\<lambda>f. Suc (Max (f ` prod_decode ` {..<n}))"
+ { fix a b x assume "x < n" and [symmetric]: "(a, b) = prod_decode x"
+ then have "a < ?M fst" "b < ?M snd"
+ by (auto intro!: Max_ge le_imp_less_Suc image_eqI) }
+ then have "setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<?M fst} \<times> {..<?M snd})"
+ by (auto intro!: setsum_mono3 simp: pos)
+ then have "\<exists>a b. setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<a} \<times> {..<b})" by auto }
+ moreover
+ { fix a b
+ let ?M = "prod_decode ` {..<Suc (Max (prod_encode ` ({..<a} \<times> {..<b})))}"
+ { fix a' b' assume "a' < a" "b' < b" then have "(a', b') \<in> ?M"
+ by (auto intro!: Max_ge le_imp_less_Suc image_eqI[where x="prod_encode (a', b')"]) }
+ then have "setsum f ({..<a} \<times> {..<b}) \<le> setsum f ?M"
+ by (auto intro!: setsum_mono3 simp: pos) }
+ ultimately
+ show ?thesis unfolding g_def using pos
+ by (auto intro!: SUPR_eq simp: setsum_cartesian_product reindex le_SUPI2
+ setsum_nonneg suminf_extreal_eq_SUPR SUPR_pair
+ SUPR_extreal_setsum[symmetric] incseq_setsumI setsum_nonneg)
+qed
+
text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*}
subsection {* Measure Spaces *}
record 'a measure_space = "'a algebra" +
- measure :: "'a set \<Rightarrow> pextreal"
+ measure :: "'a set \<Rightarrow> extreal"
-definition positive where "positive M f \<longleftrightarrow> f {} = (0::pextreal)"
- -- "Positive is enforced by the type"
+definition positive where "positive M f \<longleftrightarrow> f {} = (0::extreal) \<and> (\<forall>A\<in>sets M. 0 \<le> f A)"
definition additive where "additive M f \<longleftrightarrow>
(\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) = f x + f y)"
-definition countably_additive where "countably_additive M f \<longleftrightarrow>
- (\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow>
- (\<Sum>\<^isub>\<infinity> n. f (A n)) = f (\<Union>i. A i))"
+definition countably_additive :: "('a, 'b) algebra_scheme \<Rightarrow> ('a set \<Rightarrow> extreal) \<Rightarrow> bool" where
+ "countably_additive M f \<longleftrightarrow> (\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow>
+ (\<Sum>i. f (A i)) = f (\<Union>i. A i))"
definition increasing where "increasing M f \<longleftrightarrow>
(\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<subseteq> y \<longrightarrow> f x \<le> f y)"
definition subadditive where "subadditive M f \<longleftrightarrow>
- (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {} \<longrightarrow>
- f (x \<union> y) \<le> f x + f y)"
+ (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
definition countably_subadditive where "countably_subadditive M f \<longleftrightarrow>
(\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow>
- f (\<Union>i. A i) \<le> (\<Sum>\<^isub>\<infinity> n. f (A n)))"
+ (f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))))"
definition lambda_system where "lambda_system M f = {l \<in> sets M.
\<forall>x \<in> sets M. f (l \<inter> x) + f ((space M - l) \<inter> x) = f x}"
@@ -39,14 +69,19 @@
positive M f \<and> increasing M f \<and> countably_subadditive M f"
definition measure_set where "measure_set M f X = {r.
- \<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>\<^isub>\<infinity> i. f (A i)) = r}"
+ \<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) = r}"
locale measure_space = sigma_algebra M for M :: "('a, 'b) measure_space_scheme" +
- assumes empty_measure [simp]: "measure M {} = 0"
+ assumes measure_positive: "positive M (measure M)"
and ca: "countably_additive M (measure M)"
abbreviation (in measure_space) "\<mu> \<equiv> measure M"
+lemma (in measure_space)
+ shows empty_measure[simp, intro]: "\<mu> {} = 0"
+ and positive_measure[simp, intro!]: "\<And>A. A \<in> sets M \<Longrightarrow> 0 \<le> \<mu> A"
+ using measure_positive unfolding positive_def by auto
+
lemma increasingD:
"increasing M f \<Longrightarrow> x \<subseteq> y \<Longrightarrow> x\<in>sets M \<Longrightarrow> y\<in>sets M \<Longrightarrow> f x \<le> f y"
by (auto simp add: increasing_def)
@@ -61,39 +96,30 @@
\<Longrightarrow> f (x \<union> y) = f x + f y"
by (auto simp add: additive_def)
-lemma countably_additiveD:
- "countably_additive M f \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> disjoint_family A
- \<Longrightarrow> (\<Union>i. A i) \<in> sets M \<Longrightarrow> (\<Sum>\<^isub>\<infinity> n. f (A n)) = f (\<Union>i. A i)"
- by (simp add: countably_additive_def)
-
-lemma countably_subadditiveD:
- "countably_subadditive M f \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow>
- (\<Union>i. A i) \<in> sets M \<Longrightarrow> f (\<Union>i. A i) \<le> psuminf (f o A)"
- by (auto simp add: countably_subadditive_def o_def)
-
lemma countably_additiveI:
- "(\<And>A. range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> sets M
- \<Longrightarrow> (\<Sum>\<^isub>\<infinity> n. f (A n)) = f (\<Union>i. A i)) \<Longrightarrow> countably_additive M f"
- by (simp add: countably_additive_def)
+ assumes "\<And>A x. range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> sets M
+ \<Longrightarrow> (\<Sum>i. f (A i)) = f (\<Union>i. A i)"
+ shows "countably_additive M f"
+ using assms by (simp add: countably_additive_def)
section "Extend binary sets"
lemma LIMSEQ_binaryset:
assumes f: "f {} = 0"
- shows "(\<lambda>n. \<Sum>i = 0..<n. f (binaryset A B i)) ----> f A + f B"
+ shows "(\<lambda>n. \<Sum>i<n. f (binaryset A B i)) ----> f A + f B"
proof -
- have "(\<lambda>n. \<Sum>i = 0..< Suc (Suc n). f (binaryset A B i)) = (\<lambda>n. f A + f B)"
+ have "(\<lambda>n. \<Sum>i < Suc (Suc n). f (binaryset A B i)) = (\<lambda>n. f A + f B)"
proof
fix n
- show "(\<Sum>i\<Colon>nat = 0\<Colon>nat..<Suc (Suc n). f (binaryset A B i)) = f A + f B"
+ show "(\<Sum>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 LIMSEQ_const)
ultimately
- have "(\<lambda>n. \<Sum>i = 0..< Suc (Suc n). f (binaryset A B i)) ----> f A + f B"
+ have "(\<lambda>n. \<Sum>i< Suc (Suc n). f (binaryset A B i)) ----> f A + f B"
by metis
- hence "(\<lambda>n. \<Sum>i = 0..< n+2. f (binaryset A B i)) ----> f A + f B"
+ hence "(\<lambda>n. \<Sum>i< n+2. f (binaryset A B i)) ----> f A + f B"
by simp
thus ?thesis by (rule LIMSEQ_offset [where k=2])
qed
@@ -101,28 +127,13 @@
lemma binaryset_sums:
assumes f: "f {} = 0"
shows "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)"
- by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f])
+ by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan)
lemma suminf_binaryset_eq:
- fixes f :: "'a set \<Rightarrow> real"
+ fixes f :: "'a set \<Rightarrow> 'b::{comm_monoid_add, t2_space}"
shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
by (metis binaryset_sums sums_unique)
-lemma binaryset_psuminf:
- assumes "f {} = 0"
- shows "(\<Sum>\<^isub>\<infinity> n. f (binaryset A B n)) = f A + f B" (is "?suminf = ?sum")
-proof -
- have *: "{..<2} = {0, 1::nat}" by auto
- have "\<forall>n\<ge>2. f (binaryset A B n) = 0"
- unfolding binaryset_def
- using assms by auto
- hence "?suminf = (\<Sum>N<2. f (binaryset A B N))"
- by (rule psuminf_finite)
- also have "... = ?sum" unfolding * binaryset_def
- by simp
- finally show ?thesis .
-qed
-
subsection {* Lambda Systems *}
lemma (in algebra) lambda_system_eq:
@@ -144,7 +155,7 @@
by (simp add: lambda_system_def)
lemma (in algebra) lambda_system_Compl:
- fixes f:: "'a set \<Rightarrow> pextreal"
+ fixes f:: "'a set \<Rightarrow> extreal"
assumes x: "x \<in> lambda_system M f"
shows "space M - x \<in> lambda_system M f"
proof -
@@ -157,7 +168,7 @@
qed
lemma (in algebra) lambda_system_Int:
- fixes f:: "'a set \<Rightarrow> pextreal"
+ fixes f:: "'a set \<Rightarrow> extreal"
assumes xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
shows "x \<inter> y \<in> lambda_system M f"
proof -
@@ -191,7 +202,7 @@
qed
lemma (in algebra) lambda_system_Un:
- fixes f:: "'a set \<Rightarrow> pextreal"
+ fixes f:: "'a set \<Rightarrow> extreal"
assumes xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
shows "x \<union> y \<in> lambda_system M f"
proof -
@@ -250,53 +261,54 @@
by (auto simp add: disjoint_family_on_def binaryset_def)
hence "range (binaryset x y) \<subseteq> sets M \<longrightarrow>
(\<Union>i. binaryset x y i) \<in> sets M \<longrightarrow>
- f (\<Union>i. binaryset x y i) \<le> (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
- using cs by (simp add: countably_subadditive_def)
+ f (\<Union>i. binaryset x y i) \<le> (\<Sum> n. f (binaryset x y n))"
+ using cs by (auto simp add: countably_subadditive_def)
hence "{x,y,{}} \<subseteq> sets M \<longrightarrow> x \<union> y \<in> sets M \<longrightarrow>
- f (x \<union> y) \<le> (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
+ f (x \<union> y) \<le> (\<Sum> n. f (binaryset x y n))"
by (simp add: range_binaryset_eq UN_binaryset_eq)
thus "f (x \<union> y) \<le> f x + f y" using f x y
- by (auto simp add: Un o_def binaryset_psuminf positive_def)
+ by (auto simp add: Un o_def suminf_binaryset_eq positive_def)
qed
lemma (in algebra) additive_sum:
fixes A:: "nat \<Rightarrow> 'a set"
- assumes f: "positive M f" and ad: "additive M f"
+ assumes f: "positive M f" and ad: "additive M f" and "finite S"
and A: "range A \<subseteq> sets M"
- and disj: "disjoint_family A"
- shows "setsum (f \<circ> A) {0..<n} = f (\<Union>i\<in>{0..<n}. A i)"
-proof (induct n)
- case 0 show ?case using f by (simp add: positive_def)
+ and disj: "disjoint_family_on A S"
+ shows "(\<Sum>i\<in>S. f (A i)) = f (\<Union>i\<in>S. A i)"
+using `finite S` disj proof induct
+ case empty show ?case using f by (simp add: positive_def)
next
- case (Suc n)
- have "A n \<inter> (\<Union>i\<in>{0..<n}. A i) = {}" using disj
- by (auto simp add: disjoint_family_on_def neq_iff) blast
+ case (insert s S)
+ then have "A s \<inter> (\<Union>i\<in>S. A i) = {}"
+ by (auto simp add: disjoint_family_on_def neq_iff)
moreover
- have "A n \<in> sets M" using A by blast
- moreover have "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
- by (metis A UNION_in_sets atLeast0LessThan)
+ have "A s \<in> sets M" using A by blast
+ moreover have "(\<Union>i\<in>S. A i) \<in> sets M"
+ using A `finite S` by auto
moreover
- ultimately have "f (A n \<union> (\<Union>i\<in>{0..<n}. A i)) = f (A n) + f(\<Union>i\<in>{0..<n}. A i)"
+ ultimately have "f (A s \<union> (\<Union>i\<in>S. A i)) = f (A s) + f(\<Union>i\<in>S. A i)"
using ad UNION_in_sets A by (auto simp add: additive_def)
- with Suc.hyps show ?case using ad
- by (auto simp add: atLeastLessThanSuc 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 algebra) increasing_additive_bound:
- fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> pextreal"
+ fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> extreal"
assumes f: "positive M f" and ad: "additive M f"
and inc: "increasing M f"
and A: "range A \<subseteq> sets M"
and disj: "disjoint_family A"
- shows "psuminf (f \<circ> A) \<le> f (space M)"
-proof (safe intro!: psuminf_bound)
+ shows "(\<Sum>i. f (A i)) \<le> f (space M)"
+proof (safe intro!: suminf_bound)
fix N
- have "setsum (f \<circ> A) {0..<N} = f (\<Union>i\<in>{0..<N}. A i)"
- by (rule additive_sum [OF f ad A disj])
+ note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
+ have "(\<Sum>i<N. f (A i)) = f (\<Union>i\<in>{..<N}. A i)"
+ by (rule additive_sum [OF f ad _ A]) (auto simp: disj_N)
also have "... \<le> f (space M)" using space_closed A
- by (blast intro: increasingD [OF inc] UNION_in_sets top)
- finally show "setsum (f \<circ> A) {..<N} \<le> f (space M)" by (simp add: atLeast0LessThan)
-qed
+ by (intro increasingD[OF inc] finite_UN) auto
+ finally show "(\<Sum>i<N. f (A i)) \<le> f (space M)" by simp
+qed (insert f A, auto simp: positive_def)
lemma lambda_system_increasing:
"increasing M f \<Longrightarrow> increasing (M (|sets := lambda_system M f|)) f"
@@ -307,7 +319,7 @@
by (simp add: positive_def lambda_system_def)
lemma (in algebra) lambda_system_strong_sum:
- fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> pextreal"
+ fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> extreal"
assumes f: "positive M f" and a: "a \<in> sets M"
and A: "range A \<subseteq> lambda_system M f"
and disj: "disjoint_family A"
@@ -331,7 +343,7 @@
assumes oms: "outer_measure_space M f"
and A: "range A \<subseteq> lambda_system M f"
and disj: "disjoint_family A"
- shows "(\<Union>i. A i) \<in> lambda_system M f \<and> psuminf (f \<circ> A) = f (\<Union>i. A i)"
+ shows "(\<Union>i. A i) \<in> lambda_system M f \<and> (\<Sum>i. f (A i)) = f (\<Union>i. A i)"
proof -
have pos: "positive M f" and inc: "increasing M f"
and csa: "countably_subadditive M f"
@@ -347,15 +359,16 @@
have U_in: "(\<Union>i. A i) \<in> sets M"
by (metis A'' countable_UN)
- have U_eq: "f (\<Union>i. A i) = psuminf (f o A)"
+ have U_eq: "f (\<Union>i. A i) = (\<Sum>i. f (A i))"
proof (rule antisym)
- show "f (\<Union>i. A i) \<le> psuminf (f \<circ> A)"
- by (rule countably_subadditiveD [OF csa A'' disj U_in])
- show "psuminf (f \<circ> A) \<le> f (\<Union>i. A i)"
- by (rule psuminf_bound, unfold atLeast0LessThan[symmetric])
- (metis algebra.additive_sum [OF alg_ls] pos disj UN_Un Un_UNIV_right
- lambda_system_positive lambda_system_additive
- subset_Un_eq increasingD [OF inc] A' A'' UNION_in_sets U_in)
+ show "f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))"
+ using csa[unfolded countably_subadditive_def] A'' disj U_in by auto
+ have *: "\<And>i. 0 \<le> f (A i)" using pos A'' unfolding positive_def by auto
+ have dis: "\<And>N. disjoint_family_on A {..<N}" by (intro disjoint_family_on_mono[OF _ disj]) auto
+ show "(\<Sum>i. f (A i)) \<le> f (\<Union>i. A i)"
+ using algebra.additive_sum [OF alg_ls lambda_system_positive[OF pos] lambda_system_additive _ A' dis]
+ using A''
+ by (intro suminf_bound[OF _ *]) (auto intro!: increasingD[OF inc] allI countable_UN)
qed
{
fix a
@@ -373,15 +386,15 @@
have "a \<inter> (\<Union>i. A i) \<in> sets M"
by (metis Int U_in a)
ultimately
- have "f (a \<inter> (\<Union>i. A i)) \<le> psuminf (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A)"
- using countably_subadditiveD [OF csa, of "(\<lambda>i. a \<inter> A i)"]
+ have "f (a \<inter> (\<Union>i. A i)) \<le> (\<Sum>i. f (a \<inter> A i))"
+ using csa[unfolded countably_subadditive_def, rule_format, of "(\<lambda>i. a \<inter> A i)"]
by (simp add: o_def)
hence "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le>
- psuminf (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) + f (a - (\<Union>i. A i))"
+ (\<Sum>i. f (a \<inter> A i)) + f (a - (\<Union>i. A i))"
by (rule add_right_mono)
moreover
- have "psuminf (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) + f (a - (\<Union>i. A i)) \<le> f a"
- proof (safe intro!: psuminf_bound_add)
+ have "(\<Sum>i. f (a \<inter> A i)) + f (a - (\<Union>i. A i)) \<le> f a"
+ proof (intro suminf_bound_add allI)
fix n
have UNION_in: "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
by (metis A'' UNION_in_sets)
@@ -395,8 +408,14 @@
have "f (a - (\<Union>i. A i)) \<le> f (a - (\<Union>i\<in>{0..<n}. A i))"
by (blast intro: increasingD [OF inc] UNION_eq_Union_image
UNION_in U_in)
- thus "setsum (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) {..<n} + f (a - (\<Union>i. A i)) \<le> f a"
+ thus "(\<Sum>i<n. f (a \<inter> A i)) + f (a - (\<Union>i. A i)) \<le> f a"
by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric])
+ next
+ have "\<And>i. a \<inter> A i \<in> sets M" using A'' by auto
+ then show "\<And>i. 0 \<le> f (a \<inter> A i)" using pos[unfolded positive_def] by auto
+ have "\<And>i. a - (\<Union>i. A i) \<in> sets M" using A'' by auto
+ then have "\<And>i. 0 \<le> f (a - (\<Union>i. A i))" using pos[unfolded positive_def] by auto
+ then show "f (a - (\<Union>i. A i)) \<noteq> -\<infinity>" by auto
qed
ultimately show "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> f a"
by (rule order_trans)
@@ -443,12 +462,14 @@
proof (auto simp add: increasing_def)
fix x y
assume xy: "x \<in> sets M" "y \<in> sets M" "x \<subseteq> y"
- have "f x \<le> f x + f (y-x)" ..
+ then have "y - x \<in> sets M" by auto
+ then have "0 \<le> f (y-x)" using posf[unfolded positive_def] by auto
+ then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono) auto
also have "... = f (x \<union> (y-x))" using addf
by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
also have "... = f y"
by (metis Un_Diff_cancel Un_absorb1 xy(3))
- finally show "f x \<le> f y" .
+ finally show "f x \<le> f y" by simp
qed
lemma (in algebra) countably_additive_additive:
@@ -461,27 +482,27 @@
by (auto simp add: disjoint_family_on_def binaryset_def)
hence "range (binaryset x y) \<subseteq> sets M \<longrightarrow>
(\<Union>i. binaryset x y i) \<in> sets M \<longrightarrow>
- f (\<Union>i. binaryset x y i) = (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
+ f (\<Union>i. binaryset x y i) = (\<Sum> n. f (binaryset x y n))"
using ca
by (simp add: countably_additive_def)
hence "{x,y,{}} \<subseteq> sets M \<longrightarrow> x \<union> y \<in> sets M \<longrightarrow>
- f (x \<union> y) = (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
+ f (x \<union> y) = (\<Sum>n. f (binaryset x y n))"
by (simp add: range_binaryset_eq UN_binaryset_eq)
thus "f (x \<union> y) = f x + f y" using posf x y
- by (auto simp add: Un binaryset_psuminf positive_def)
+ by (auto simp add: Un suminf_binaryset_eq positive_def)
qed
lemma inf_measure_nonempty:
assumes f: "positive M f" and b: "b \<in> sets M" and a: "a \<subseteq> b" "{} \<in> sets M"
shows "f b \<in> measure_set M f a"
proof -
- have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = setsum (f \<circ> (\<lambda>i. {})(0 := b)) {..<1::nat}"
- by (rule psuminf_finite) (simp add: f[unfolded positive_def])
+ let ?A = "\<lambda>i::nat. (if i = 0 then b else {})"
+ have "(\<Sum>i. f (?A i)) = (\<Sum>i<1::nat. f (?A i))"
+ by (rule suminf_finite) (simp add: f[unfolded positive_def])
also have "... = f b"
by simp
- finally have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = f b" .
- thus ?thesis using assms
- by (auto intro!: exI [of _ "(\<lambda>i. {})(0 := b)"]
+ finally show ?thesis using assms
+ by (auto intro!: exI [of _ ?A]
simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def)
qed
@@ -489,19 +510,19 @@
assumes posf: "positive M f" and ca: "countably_additive M f"
and s: "s \<in> sets M"
shows "Inf (measure_set M f s) = f s"
- unfolding Inf_pextreal_def
+ unfolding Inf_extreal_def
proof (safe intro!: Greatest_equality)
fix z
assume z: "z \<in> measure_set M f s"
from this obtain A where
A: "range A \<subseteq> sets M" and disj: "disjoint_family A"
- and "s \<subseteq> (\<Union>x. A x)" and si: "psuminf (f \<circ> A) = z"
+ and "s \<subseteq> (\<Union>x. A x)" and si: "(\<Sum>i. f (A i)) = z"
by (auto simp add: measure_set_def comp_def)
hence seq: "s = (\<Union>i. A i \<inter> s)" by blast
have inc: "increasing M f"
by (metis additive_increasing ca countably_additive_additive posf)
- have sums: "psuminf (\<lambda>i. f (A i \<inter> s)) = f (\<Union>i. A i \<inter> s)"
- proof (rule countably_additiveD [OF ca])
+ have sums: "(\<Sum>i. f (A i \<inter> s)) = f (\<Union>i. A i \<inter> s)"
+ proof (rule ca[unfolded countably_additive_def, rule_format])
show "range (\<lambda>n. A n \<inter> s) \<subseteq> sets M" using A s
by blast
show "disjoint_family (\<lambda>n. A n \<inter> s)" using disj
@@ -509,12 +530,14 @@
show "(\<Union>i. A i \<inter> s) \<in> sets M" using A s
by (metis UN_extend_simps(4) s seq)
qed
- hence "f s = psuminf (\<lambda>i. f (A i \<inter> s))"
+ hence "f s = (\<Sum>i. f (A i \<inter> s))"
using seq [symmetric] by (simp add: sums_iff)
- also have "... \<le> psuminf (f \<circ> A)"
- proof (rule psuminf_le)
- fix n show "f (A n \<inter> s) \<le> (f \<circ> A) n" using A s
+ also have "... \<le> (\<Sum>i. f (A i))"
+ proof (rule suminf_le_pos)
+ fix n show "f (A n \<inter> s) \<le> f (A n)" using A s
by (force intro: increasingD [OF inc])
+ fix N have "A N \<inter> s \<in> sets M" using A s by auto
+ then show "0 \<le> f (A N \<inter> s)" using posf unfolding positive_def by auto
qed
also have "... = z" by (rule si)
finally show "f s \<le> z" .
@@ -525,18 +548,40 @@
by (blast intro: inf_measure_nonempty [of _ f, OF posf s subset_refl])
qed
+lemma measure_set_pos:
+ assumes posf: "positive M f" "r \<in> measure_set M f X"
+ shows "0 \<le> r"
+proof -
+ obtain A where "range A \<subseteq> sets M" and r: "r = (\<Sum>i. f (A i))"
+ using `r \<in> measure_set M f X` unfolding measure_set_def by auto
+ then show "0 \<le> r" using posf unfolding r positive_def
+ by (intro suminf_0_le) auto
+qed
+
+lemma inf_measure_pos:
+ assumes posf: "positive M f"
+ shows "0 \<le> Inf (measure_set M f X)"
+proof (rule complete_lattice_class.Inf_greatest)
+ fix r assume "r \<in> measure_set M f X" with posf show "0 \<le> r"
+ by (rule measure_set_pos)
+qed
+
lemma inf_measure_empty:
- assumes posf: "positive M f" "{} \<in> sets M"
+ assumes posf: "positive M f" and "{} \<in> sets M"
shows "Inf (measure_set M f {}) = 0"
proof (rule antisym)
show "Inf (measure_set M f {}) \<le> 0"
by (metis complete_lattice_class.Inf_lower `{} \<in> sets M`
inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def])
-qed simp
+qed (rule inf_measure_pos[OF posf])
lemma (in algebra) inf_measure_positive:
- "positive M f \<Longrightarrow> positive M (\<lambda>x. Inf (measure_set M f x))"
- by (simp add: positive_def inf_measure_empty)
+ assumes p: "positive M f" and "{} \<in> sets M"
+ shows "positive M (\<lambda>x. Inf (measure_set M f x))"
+proof (unfold positive_def, intro conjI ballI)
+ show "Inf (measure_set M f {}) = 0" using inf_measure_empty[OF assms] by auto
+ fix A assume "A \<in> sets M"
+qed (rule inf_measure_pos[OF p])
lemma (in algebra) inf_measure_increasing:
assumes posf: "positive M f"
@@ -548,25 +593,25 @@
apply (clarsimp simp add: measure_set_def, rule_tac x=A in exI, blast)
done
-
lemma (in algebra) inf_measure_le:
assumes posf: "positive M f" and inc: "increasing M f"
- and x: "x \<in> {r . \<exists>A. range A \<subseteq> sets M \<and> s \<subseteq> (\<Union>i. A i) \<and> psuminf (f \<circ> A) = r}"
+ and x: "x \<in> {r . \<exists>A. range A \<subseteq> sets M \<and> s \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) = r}"
shows "Inf (measure_set M f s) \<le> x"
proof -
- from x
obtain A where A: "range A \<subseteq> sets M" and ss: "s \<subseteq> (\<Union>i. A i)"
- and xeq: "psuminf (f \<circ> A) = x"
- by auto
+ and xeq: "(\<Sum>i. f (A i)) = x"
+ using x by auto
have dA: "range (disjointed A) \<subseteq> sets M"
by (metis A range_disjointed_sets)
- have "\<forall>n.(f o disjointed A) n \<le> (f \<circ> A) n" unfolding comp_def
+ have "\<forall>n. f (disjointed A n) \<le> f (A n)"
by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A comp_def)
- hence sda: "psuminf (f o disjointed A) \<le> psuminf (f \<circ> A)"
- by (blast intro: psuminf_le)
- hence ley: "psuminf (f o disjointed A) \<le> x"
+ moreover have "\<forall>i. 0 \<le> f (disjointed A i)"
+ using posf dA unfolding positive_def by auto
+ ultimately have sda: "(\<Sum>i. f (disjointed A i)) \<le> (\<Sum>i. f (A i))"
+ by (blast intro!: suminf_le_pos)
+ hence ley: "(\<Sum>i. f (disjointed A i)) \<le> x"
by (metis xeq)
- hence y: "psuminf (f o disjointed A) \<in> measure_set M f s"
+ hence y: "(\<Sum>i. f (disjointed A i)) \<in> measure_set M f s"
apply (auto simp add: measure_set_def)
apply (rule_tac x="disjointed A" in exI)
apply (simp add: disjoint_family_disjointed UN_disjointed_eq ss dA comp_def)
@@ -576,13 +621,16 @@
qed
lemma (in algebra) inf_measure_close:
+ fixes e :: extreal
assumes posf: "positive M f" and e: "0 < e" and ss: "s \<subseteq> (space M)"
shows "\<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> s \<subseteq> (\<Union>i. A i) \<and>
- psuminf (f \<circ> A) \<le> Inf (measure_set M f s) + e"
-proof (cases "Inf (measure_set M f s) = \<omega>")
+ (\<Sum>i. f (A i)) \<le> Inf (measure_set M f s) + e"
+proof (cases "Inf (measure_set M f s) = \<infinity>")
case False
+ then have fin: "\<bar>Inf (measure_set M f s)\<bar> \<noteq> \<infinity>"
+ using inf_measure_pos[OF posf, of s] by auto
obtain l where "l \<in> measure_set M f s" "l \<le> Inf (measure_set M f s) + e"
- using Inf_close[OF False e] by auto
+ using Inf_extreal_close[OF fin e] by auto
thus ?thesis
by (auto intro!: exI[of _ l] simp: measure_set_def comp_def)
next
@@ -600,9 +648,8 @@
shows "countably_subadditive (| space = space M, sets = Pow (space M) |)
(\<lambda>x. Inf (measure_set M f x))"
unfolding countably_subadditive_def o_def
-proof (safe, simp, rule pextreal_le_epsilon)
- fix A :: "nat \<Rightarrow> 'a set" and e :: pextreal
-
+proof (safe, simp, rule extreal_le_epsilon, safe)
+ fix A :: "nat \<Rightarrow> 'a set" and e :: extreal
let "?outer n" = "Inf (measure_set M f (A n))"
assume A: "range A \<subseteq> Pow (space M)"
and disj: "disjoint_family A"
@@ -610,21 +657,27 @@
and e: "0 < e"
hence "\<exists>BB. \<forall>n. range (BB n) \<subseteq> sets M \<and> disjoint_family (BB n) \<and>
A n \<subseteq> (\<Union>i. BB n i) \<and>
- psuminf (f o BB n) \<le> ?outer n + e * (1/2)^(Suc n)"
- apply (safe intro!: choice inf_measure_close [of f, OF posf _])
- using e sb by (cases e, auto simp add: not_le mult_pos_pos)
+ (\<Sum>i. f (BB n i)) \<le> ?outer n + e * (1/2)^(Suc n)"
+ apply (safe intro!: choice inf_measure_close [of f, OF posf])
+ using e sb by (cases e) (auto simp add: not_le mult_pos_pos one_extreal_def)
then obtain BB
where BB: "\<And>n. (range (BB n) \<subseteq> sets M)"
and disjBB: "\<And>n. disjoint_family (BB n)"
and sbBB: "\<And>n. A n \<subseteq> (\<Union>i. BB n i)"
- and BBle: "\<And>n. psuminf (f o BB n) \<le> ?outer n + e * (1/2)^(Suc n)"
+ and BBle: "\<And>n. (\<Sum>i. f (BB n i)) \<le> ?outer n + e * (1/2)^(Suc n)"
by auto blast
- have sll: "(\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n)) \<le> psuminf ?outer + e"
+ have sll: "(\<Sum>n. \<Sum>i. (f (BB n i))) \<le> suminf ?outer + e"
proof -
- have "(\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n)) \<le> (\<Sum>\<^isub>\<infinity> n. ?outer n + e*(1/2) ^ Suc n)"
- by (rule psuminf_le[OF BBle])
- also have "... = psuminf ?outer + e"
- using psuminf_half_series by simp
+ have sum_eq_1: "(\<Sum>n. e*(1/2) ^ Suc n) = e"
+ using suminf_half_series_extreal e
+ by (simp add: extreal_zero_le_0_iff zero_le_divide_extreal suminf_cmult_extreal)
+ have "\<And>n i. 0 \<le> f (BB n i)" using posf[unfolded positive_def] BB by auto
+ then have "\<And>n. 0 \<le> (\<Sum>i. f (BB n i))" by (rule suminf_0_le)
+ then have "(\<Sum>n. \<Sum>i. (f (BB n i))) \<le> (\<Sum>n. ?outer n + e*(1/2) ^ Suc n)"
+ by (rule suminf_le_pos[OF BBle])
+ also have "... = suminf ?outer + e"
+ using sum_eq_1 inf_measure_pos[OF posf] e
+ by (subst suminf_add_extreal) (auto simp add: extreal_zero_le_0_iff)
finally show ?thesis .
qed
def C \<equiv> "(split BB) o prod_decode"
@@ -644,23 +697,25 @@
qed
have "(f \<circ> C) = (f \<circ> (\<lambda>(x, y). BB x y)) \<circ> prod_decode"
by (rule ext) (auto simp add: C_def)
- moreover have "psuminf ... = (\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n))" using BBle
- by (force intro!: psuminf_2dimen simp: o_def)
- ultimately have Csums: "psuminf (f \<circ> C) = (\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n))" by simp
- have "Inf (measure_set M f (\<Union>i. A i)) \<le> (\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n))"
+ moreover have "suminf ... = (\<Sum>n. \<Sum>i. f (BB n i))" using BBle
+ using BB posf[unfolded positive_def]
+ by (force intro!: suminf_extreal_2dimen simp: o_def)
+ ultimately have Csums: "(\<Sum>i. f (C i)) = (\<Sum>n. \<Sum>i. f (BB n i))" by (simp add: o_def)
+ have "Inf (measure_set M f (\<Union>i. A i)) \<le> (\<Sum>n. \<Sum>i. f (BB n i))"
apply (rule inf_measure_le [OF posf(1) inc], auto)
apply (rule_tac x="C" in exI)
apply (auto simp add: C sbC Csums)
done
- also have "... \<le> (\<Sum>\<^isub>\<infinity>n. Inf (measure_set M f (A n))) + e" using sll
+ also have "... \<le> (\<Sum>n. Inf (measure_set M f (A n))) + e" using sll
by blast
- finally show "Inf (measure_set M f (\<Union>i. A i)) \<le> psuminf ?outer + e" .
+ finally show "Inf (measure_set M f (\<Union>i. A i)) \<le> suminf ?outer + e" .
qed
lemma (in algebra) inf_measure_outer:
"\<lbrakk> positive M f ; increasing M f \<rbrakk>
\<Longrightarrow> outer_measure_space \<lparr> space = space M, sets = Pow (space M) \<rparr>
(\<lambda>x. Inf (measure_set M f x))"
+ using inf_measure_pos[of M f]
by (simp add: outer_measure_space_def inf_measure_empty
inf_measure_increasing inf_measure_countably_subadditive positive_def)
@@ -680,13 +735,13 @@
by blast
have "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
\<le> Inf (measure_set M f s)"
- proof (rule pextreal_le_epsilon)
- fix e :: pextreal
+ proof (rule extreal_le_epsilon, intro allI impI)
+ fix e :: extreal
assume e: "0 < e"
from inf_measure_close [of f, OF posf e s]
obtain A where A: "range A \<subseteq> sets M" and disj: "disjoint_family A"
and sUN: "s \<subseteq> (\<Union>i. A i)"
- and l: "psuminf (f \<circ> A) \<le> Inf (measure_set M f s) + e"
+ and l: "(\<Sum>i. f (A i)) \<le> Inf (measure_set M f s) + e"
by auto
have [simp]: "!!x. x \<in> sets M \<Longrightarrow>
(f o (\<lambda>z. z \<inter> (space M - x)) o A) = (f o (\<lambda>z. z - x) o A)"
@@ -698,9 +753,9 @@
assume u: "u \<in> sets M"
have [simp]: "\<And>n. f (A n \<inter> u) \<le> f (A n)"
by (simp add: increasingD [OF inc] u Int range_subsetD [OF A])
- have 2: "Inf (measure_set M f (s \<inter> u)) \<le> psuminf (f \<circ> (\<lambda>z. z \<inter> u) \<circ> A)"
+ have 2: "Inf (measure_set M f (s \<inter> u)) \<le> (\<Sum>i. f (A i \<inter> u))"
proof (rule complete_lattice_class.Inf_lower)
- show "psuminf (f \<circ> (\<lambda>z. z \<inter> u) \<circ> A) \<in> measure_set M f (s \<inter> u)"
+ show "(\<Sum>i. f (A i \<inter> u)) \<in> measure_set M f (s \<inter> u)"
apply (simp add: measure_set_def)
apply (rule_tac x="(\<lambda>z. z \<inter> u) o A" in exI)
apply (auto simp add: disjoint_family_subset [OF disj] o_def)
@@ -709,15 +764,16 @@
done
qed
} note lesum = this
- have inf1: "Inf (measure_set M f (s\<inter>x)) \<le> psuminf (f o (\<lambda>z. z\<inter>x) o A)"
+ have [simp]: "\<And>i. A i \<inter> (space M - x) = A i - x" using A sets_into_space by auto
+ have inf1: "Inf (measure_set M f (s\<inter>x)) \<le> (\<Sum>i. f (A i \<inter> x))"
and inf2: "Inf (measure_set M f (s \<inter> (space M - x)))
- \<le> psuminf (f o (\<lambda>z. z \<inter> (space M - x)) o A)"
+ \<le> (\<Sum>i. f (A i \<inter> (space M - x)))"
by (metis Diff lesum top x)+
hence "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
- \<le> psuminf (f o (\<lambda>s. s\<inter>x) o A) + psuminf (f o (\<lambda>s. s-x) o A)"
- by (simp add: x add_mono)
- also have "... \<le> psuminf (f o A)"
- by (simp add: x psuminf_add[symmetric] o_def)
+ \<le> (\<Sum>i. f (A i \<inter> x)) + (\<Sum>i. f (A i - x))"
+ by (simp add: add_mono x)
+ also have "... \<le> (\<Sum>i. f (A i))" using posf[unfolded positive_def] A x
+ by (subst suminf_add_extreal[symmetric]) auto
also have "... \<le> Inf (measure_set M f s) + e"
by (rule l)
finally show "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
@@ -732,7 +788,7 @@
also have "... \<le> Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))"
apply (rule subadditiveD)
apply (rule algebra.countably_subadditive_subadditive[OF algebra_Pow])
- apply (simp add: positive_def inf_measure_empty[OF posf])
+ apply (simp add: positive_def inf_measure_empty[OF posf] inf_measure_pos[OF posf])
apply (rule inf_measure_countably_subadditive)
using s by (auto intro!: posf inc)
finally show ?thesis .
@@ -751,7 +807,7 @@
theorem (in algebra) caratheodory:
assumes posf: "positive M f" and ca: "countably_additive M f"
- shows "\<exists>\<mu> :: 'a set \<Rightarrow> pextreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
+ shows "\<exists>\<mu> :: 'a set \<Rightarrow> extreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
measure_space \<lparr> space = space M, sets = sets (sigma M), measure = \<mu> \<rparr>"
proof -
have inc: "increasing M f"
--- a/src/HOL/Probability/Complete_Measure.thy Mon Mar 14 14:37:47 2011 +0100
+++ b/src/HOL/Probability/Complete_Measure.thy Mon Mar 14 14:37:49 2011 +0100
@@ -1,7 +1,6 @@
-(* Title: HOL/Probability/Complete_Measure.thy
+(* Title: Complete_Measure.thy
Author: Robert Himmelmann, Johannes Hoelzl, TU Muenchen
*)
-
theory Complete_Measure
imports Product_Measure
begin
@@ -177,7 +176,8 @@
proof -
show "measure_space completion"
proof
- show "measure completion {} = 0" by (auto simp: completion_def)
+ show "positive completion (measure completion)"
+ by (auto simp: completion_def positive_def)
next
show "countably_additive completion (measure completion)"
proof (intro countably_additiveI)
@@ -189,9 +189,9 @@
using A by (subst (1 2) main_part_null_part_Un) auto
then show "main_part (A n) \<inter> main_part (A m) = {}" by auto
qed
- then have "(\<Sum>\<^isub>\<infinity>n. measure completion (A n)) = \<mu> (\<Union>i. main_part (A i))"
+ then have "(\<Sum>n. measure completion (A n)) = \<mu> (\<Union>i. main_part (A i))"
unfolding completion_def using A by (auto intro!: measure_countably_additive)
- then show "(\<Sum>\<^isub>\<infinity>n. measure completion (A n)) = measure completion (UNION UNIV A)"
+ then show "(\<Sum>n. measure completion (A n)) = measure completion (UNION UNIV A)"
by (simp add: completion_def \<mu>_main_part_UN[OF A(1)])
qed
qed
@@ -251,30 +251,52 @@
qed
qed
-lemma (in completeable_measure_space) completion_ex_borel_measurable:
- fixes g :: "'a \<Rightarrow> pextreal"
- assumes g: "g \<in> borel_measurable completion"
+lemma (in completeable_measure_space) completion_ex_borel_measurable_pos:
+ fixes g :: "'a \<Rightarrow> extreal"
+ assumes g: "g \<in> borel_measurable completion" and "\<And>x. 0 \<le> g x"
shows "\<exists>g'\<in>borel_measurable M. (AE x. g x = g' x)"
proof -
- from g[THEN completion.borel_measurable_implies_simple_function_sequence]
- obtain f where "\<And>i. simple_function completion (f i)" "f \<up> g" by auto
- then have "\<forall>i. \<exists>f'. simple_function M f' \<and> (AE x. f i x = f' x)"
- using completion_ex_simple_function by auto
+ from g[THEN completion.borel_measurable_implies_simple_function_sequence'] guess f . note f = this
+ from this(1)[THEN completion_ex_simple_function]
+ have "\<forall>i. \<exists>f'. simple_function M f' \<and> (AE x. f i x = f' x)" ..
from this[THEN choice] obtain f' where
sf: "\<And>i. simple_function M (f' i)" and
AE: "\<forall>i. AE x. f i x = f' i x" by auto
show ?thesis
proof (intro bexI)
- from AE[unfolded all_AE_countable]
+ from AE[unfolded AE_all_countable[symmetric]]
show "AE x. g x = (SUP i. f' i x)" (is "AE x. g x = ?f x")
proof (elim AE_mp, safe intro!: AE_I2)
fix x assume eq: "\<forall>i. f i x = f' i x"
- moreover have "g = SUPR UNIV f" using `f \<up> g` unfolding isoton_def by simp
- ultimately show "g x = ?f x" by (simp add: SUPR_apply)
+ moreover have "g x = (SUP i. f i x)"
+ unfolding f using `0 \<le> g x` by (auto split: split_max)
+ ultimately show "g x = ?f x" by auto
qed
show "?f \<in> borel_measurable M"
using sf by (auto intro: borel_measurable_simple_function)
qed
qed
+lemma (in completeable_measure_space) completion_ex_borel_measurable:
+ fixes g :: "'a \<Rightarrow> extreal"
+ assumes g: "g \<in> borel_measurable completion"
+ shows "\<exists>g'\<in>borel_measurable M. (AE x. g x = g' x)"
+proof -
+ have "(\<lambda>x. max 0 (g x)) \<in> borel_measurable completion" "\<And>x. 0 \<le> max 0 (g x)" using g by auto
+ from completion_ex_borel_measurable_pos[OF this] guess g_pos ..
+ moreover
+ have "(\<lambda>x. max 0 (- g x)) \<in> borel_measurable completion" "\<And>x. 0 \<le> max 0 (- g x)" using g by auto
+ from completion_ex_borel_measurable_pos[OF this] guess g_neg ..
+ ultimately
+ show ?thesis
+ proof (safe intro!: bexI[of _ "\<lambda>x. g_pos x - g_neg x"])
+ show "AE x. max 0 (- g x) = g_neg x \<longrightarrow> max 0 (g x) = g_pos x \<longrightarrow> g x = g_pos x - g_neg x"
+ proof (intro AE_I2 impI)
+ fix x assume g: "max 0 (- g x) = g_neg x" "max 0 (g x) = g_pos x"
+ show "g x = g_pos x - g_neg x" unfolding g[symmetric]
+ by (cases "g x") (auto split: split_max)
+ qed
+ qed auto
+qed
+
end
--- a/src/HOL/Probability/Information.thy Mon Mar 14 14:37:47 2011 +0100
+++ b/src/HOL/Probability/Information.thy Mon Mar 14 14:37:49 2011 +0100
@@ -2,9 +2,12 @@
imports
Probability_Space
"~~/src/HOL/Library/Convex"
- Lebesgue_Measure
begin
+lemma (in prob_space) not_zero_less_distribution[simp]:
+ "(\<not> 0 < distribution X A) \<longleftrightarrow> distribution X A = 0"
+ using distribution_positive[of X A] by arith
+
lemma log_le: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> x \<le> y \<Longrightarrow> log a x \<le> log a y"
by (subst log_le_cancel_iff) auto
@@ -238,7 +241,7 @@
have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
show "(\<Sum>x \<in> space M. log b (real (RN_deriv M \<nu> x)) * real (\<nu> {x})) = ?sum"
using RN_deriv_finite_measure[OF ms ac]
- by (auto intro!: setsum_cong simp: field_simps real_of_pextreal_mult[symmetric])
+ by (auto intro!: setsum_cong simp: field_simps)
qed
lemma (in finite_prob_space) KL_divergence_positive_finite:
@@ -254,7 +257,8 @@
proof (subst KL_divergence_eq_finite[OF ms ac], safe intro!: log_setsum_divide not_empty)
show "finite (space M)" using finite_space by simp
show "1 < b" by fact
- show "(\<Sum>x\<in>space M. real (\<nu> {x})) = 1" using v.finite_sum_over_space_eq_1 by simp
+ show "(\<Sum>x\<in>space M. real (\<nu> {x})) = 1"
+ using v.finite_sum_over_space_eq_1 by (simp add: v.\<mu>'_def)
fix x assume "x \<in> space M"
then have x: "{x} \<in> sets M" unfolding sets_eq_Pow by auto
@@ -262,17 +266,19 @@
then have "\<nu> {x} \<noteq> 0" by auto
then have "\<mu> {x} \<noteq> 0"
using ac[unfolded absolutely_continuous_def, THEN bspec, of "{x}"] x by auto
- thus "0 < prob {x}" using finite_measure[of "{x}"] x by auto }
- qed auto
- thus "0 \<le> KL_divergence b M \<nu>" using finite_sum_over_space_eq_1 by simp
+ thus "0 < real (\<mu> {x})" using real_measure[OF x] by auto }
+ show "0 \<le> real (\<mu> {x})" "0 \<le> real (\<nu> {x})"
+ using real_measure[OF x] v.real_measure[of "{x}"] x by auto
+ qed
+ thus "0 \<le> KL_divergence b M \<nu>" using finite_sum_over_space_eq_1 by (simp add: \<mu>'_def)
qed
subsection {* Mutual Information *}
definition (in prob_space)
"mutual_information b S T X Y =
- KL_divergence b (S\<lparr>measure := distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := distribution Y\<rparr>)
- (joint_distribution X Y)"
+ KL_divergence b (S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>)
+ (extreal\<circ>joint_distribution X Y)"
definition (in prob_space)
"entropy b s X = mutual_information b s s X X"
@@ -280,38 +286,33 @@
abbreviation (in information_space)
mutual_information_Pow ("\<I>'(_ ; _')") where
"\<I>(X ; Y) \<equiv> mutual_information b
- \<lparr> space = X`space M, sets = Pow (X`space M), measure = distribution X \<rparr>
- \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = distribution Y \<rparr> X Y"
+ \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr>
+ \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = extreal\<circ>distribution Y \<rparr> X Y"
lemma (in prob_space) finite_variables_absolutely_continuous:
assumes X: "finite_random_variable S X" and Y: "finite_random_variable T Y"
shows "measure_space.absolutely_continuous
- (S\<lparr>measure := distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := distribution Y\<rparr>)
- (joint_distribution X Y)"
+ (S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>)
+ (extreal\<circ>joint_distribution X Y)"
proof -
- interpret X: finite_prob_space "S\<lparr>measure := distribution X\<rparr>"
+ interpret X: finite_prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>"
using X by (rule distribution_finite_prob_space)
- interpret Y: finite_prob_space "T\<lparr>measure := distribution Y\<rparr>"
+ interpret Y: finite_prob_space "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
using Y by (rule distribution_finite_prob_space)
interpret XY: pair_finite_prob_space
- "S\<lparr>measure := distribution X\<rparr>" "T\<lparr> measure := distribution Y\<rparr>" by default
- interpret P: finite_prob_space "XY.P\<lparr> measure := joint_distribution X Y\<rparr>"
+ "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" "T\<lparr> measure := extreal\<circ>distribution Y\<rparr>" by default
+ interpret P: finite_prob_space "XY.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>"
using assms by (auto intro!: joint_distribution_finite_prob_space)
note rv = assms[THEN finite_random_variableD]
- show "XY.absolutely_continuous (joint_distribution X Y)"
+ show "XY.absolutely_continuous (extreal\<circ>joint_distribution X Y)"
proof (rule XY.absolutely_continuousI)
- show "finite_measure_space (XY.P\<lparr> measure := joint_distribution X Y\<rparr>)" by default
+ show "finite_measure_space (XY.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)" by default
fix x assume "x \<in> space XY.P" and "XY.\<mu> {x} = 0"
- then obtain a b where "(a, b) = x" and "a \<in> space S" "b \<in> space T"
- and distr: "distribution X {a} * distribution Y {b} = 0"
+ then obtain a b where "x = (a, b)"
+ and "distribution X {a} = 0 \<or> distribution Y {b} = 0"
by (cases x) (auto simp: space_pair_measure)
- with X.sets_eq_Pow Y.sets_eq_Pow
- joint_distribution_Times_le_fst[OF rv, of "{a}" "{b}"]
- joint_distribution_Times_le_snd[OF rv, of "{a}" "{b}"]
- have "joint_distribution X Y {x} \<le> distribution Y {b}"
- "joint_distribution X Y {x} \<le> distribution X {a}"
- by (auto simp del: X.sets_eq_Pow Y.sets_eq_Pow)
- with distr show "joint_distribution X Y {x} = 0" by auto
+ with finite_distribution_order(5,6)[OF X Y]
+ show "(extreal \<circ> joint_distribution X Y) {x} = 0" by auto
qed
qed
@@ -320,28 +321,28 @@
assumes MY: "finite_random_variable MY Y"
shows mutual_information_generic_eq:
"mutual_information b MX MY X Y = (\<Sum> (x,y) \<in> space MX \<times> space MY.
- real (joint_distribution X Y {(x,y)}) *
- log b (real (joint_distribution X Y {(x,y)}) /
- (real (distribution X {x}) * real (distribution Y {y}))))"
+ joint_distribution X Y {(x,y)} *
+ log b (joint_distribution X Y {(x,y)} /
+ (distribution X {x} * distribution Y {y})))"
(is ?sum)
and mutual_information_positive_generic:
"0 \<le> mutual_information b MX MY X Y" (is ?positive)
proof -
- interpret X: finite_prob_space "MX\<lparr>measure := distribution X\<rparr>"
+ interpret X: finite_prob_space "MX\<lparr>measure := extreal\<circ>distribution X\<rparr>"
using MX by (rule distribution_finite_prob_space)
- interpret Y: finite_prob_space "MY\<lparr>measure := distribution Y\<rparr>"
+ interpret Y: finite_prob_space "MY\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
using MY by (rule distribution_finite_prob_space)
- interpret XY: pair_finite_prob_space "MX\<lparr>measure := distribution X\<rparr>" "MY\<lparr>measure := distribution Y\<rparr>" by default
- interpret P: finite_prob_space "XY.P\<lparr>measure := joint_distribution X Y\<rparr>"
+ interpret XY: pair_finite_prob_space "MX\<lparr>measure := extreal\<circ>distribution X\<rparr>" "MY\<lparr>measure := extreal\<circ>distribution Y\<rparr>" by default
+ interpret P: finite_prob_space "XY.P\<lparr>measure := extreal\<circ>joint_distribution X Y\<rparr>"
using assms by (auto intro!: joint_distribution_finite_prob_space)
- have P_ms: "finite_measure_space (XY.P\<lparr>measure :=joint_distribution X Y\<rparr>)" by default
- have P_ps: "finite_prob_space (XY.P\<lparr>measure := joint_distribution X Y\<rparr>)" by default
+ have P_ms: "finite_measure_space (XY.P\<lparr>measure := extreal\<circ>joint_distribution X Y\<rparr>)" by default
+ have P_ps: "finite_prob_space (XY.P\<lparr>measure := extreal\<circ>joint_distribution X Y\<rparr>)" by default
show ?sum
unfolding Let_def mutual_information_def
by (subst XY.KL_divergence_eq_finite[OF P_ms finite_variables_absolutely_continuous[OF MX MY]])
- (auto simp add: space_pair_measure setsum_cartesian_product' real_of_pextreal_mult[symmetric])
+ (auto simp add: space_pair_measure setsum_cartesian_product')
show ?positive
using XY.KL_divergence_positive_finite[OF P_ps finite_variables_absolutely_continuous[OF MX MY] b_gt_1]
@@ -351,10 +352,10 @@
lemma (in information_space) mutual_information_commute_generic:
assumes X: "random_variable S X" and Y: "random_variable T Y"
assumes ac: "measure_space.absolutely_continuous
- (S\<lparr>measure := distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := distribution Y\<rparr>) (joint_distribution X Y)"
+ (S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>) (extreal\<circ>joint_distribution X Y)"
shows "mutual_information b S T X Y = mutual_information b T S Y X"
proof -
- let ?S = "S\<lparr>measure := distribution X\<rparr>" and ?T = "T\<lparr>measure := distribution Y\<rparr>"
+ let ?S = "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" and ?T = "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
interpret S: prob_space ?S using X by (rule distribution_prob_space)
interpret T: prob_space ?T using Y by (rule distribution_prob_space)
interpret P: pair_prob_space ?S ?T ..
@@ -363,13 +364,13 @@
unfolding mutual_information_def
proof (intro Q.KL_divergence_vimage[OF Q.measure_preserving_swap _ _ _ _ ac b_gt_1])
show "(\<lambda>(x,y). (y,x)) \<in> measure_preserving
- (P.P \<lparr> measure := joint_distribution X Y\<rparr>) (Q.P \<lparr> measure := joint_distribution Y X\<rparr>)"
+ (P.P \<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>) (Q.P \<lparr> measure := extreal\<circ>joint_distribution Y X\<rparr>)"
using X Y unfolding measurable_def
unfolding measure_preserving_def using P.pair_sigma_algebra_swap_measurable
- by (auto simp add: space_pair_measure distribution_def intro!: arg_cong[where f=\<mu>])
- have "prob_space (P.P\<lparr> measure := joint_distribution X Y\<rparr>)"
+ by (auto simp add: space_pair_measure distribution_def intro!: arg_cong[where f=\<mu>'])
+ have "prob_space (P.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)"
using X Y by (auto intro!: distribution_prob_space random_variable_pairI)
- then show "measure_space (P.P\<lparr> measure := joint_distribution X Y\<rparr>)"
+ then show "measure_space (P.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)"
unfolding prob_space_def by simp
qed auto
qed
@@ -389,8 +390,8 @@
lemma (in information_space) mutual_information_eq:
assumes "simple_function M X" "simple_function M Y"
shows "\<I>(X;Y) = (\<Sum> (x,y) \<in> X ` space M \<times> Y ` space M.
- real (distribution (\<lambda>x. (X x, Y x)) {(x,y)}) * log b (real (distribution (\<lambda>x. (X x, Y x)) {(x,y)}) /
- (real (distribution X {x}) * real (distribution Y {y}))))"
+ distribution (\<lambda>x. (X x, Y x)) {(x,y)} * log b (distribution (\<lambda>x. (X x, Y x)) {(x,y)} /
+ (distribution X {x} * distribution Y {y})))"
using assms by (simp add: mutual_information_generic_eq)
lemma (in information_space) mutual_information_generic_cong:
@@ -416,22 +417,27 @@
abbreviation (in information_space)
entropy_Pow ("\<H>'(_')") where
- "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M), measure = distribution X \<rparr> X"
+ "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr> X"
lemma (in information_space) entropy_generic_eq:
+ fixes X :: "'a \<Rightarrow> 'c"
assumes MX: "finite_random_variable MX X"
- shows "entropy b MX X = -(\<Sum> x \<in> space MX. real (distribution X {x}) * log b (real (distribution X {x})))"
+ shows "entropy b MX X = -(\<Sum> x \<in> space MX. distribution X {x} * log b (distribution X {x}))"
proof -
- interpret MX: finite_prob_space "MX\<lparr>measure := distribution X\<rparr>"
+ interpret MX: finite_prob_space "MX\<lparr>measure := extreal\<circ>distribution X\<rparr>"
using MX by (rule distribution_finite_prob_space)
- let "?X x" = "real (distribution X {x})"
- let "?XX x y" = "real (joint_distribution X X {(x, y)})"
- { fix x y
- have "(\<lambda>x. (X x, X x)) -` {(x, y)} = (if x = y then X -` {x} else {})" by auto
+ let "?X x" = "distribution X {x}"
+ let "?XX x y" = "joint_distribution X X {(x, y)}"
+
+ { fix x y :: 'c
+ { assume "x \<noteq> y"
+ then have "(\<lambda>x. (X x, X x)) -` {(x,y)} \<inter> space M = {}" by auto
+ then have "joint_distribution X X {(x, y)} = 0" by (simp add: distribution_def) }
then have "?XX x y * log b (?XX x y / (?X x * ?X y)) =
(if x = y then - ?X y * log b (?X y) else 0)"
- unfolding distribution_def by (auto simp: log_simps zero_less_mult_iff) }
+ by (auto simp: log_simps zero_less_mult_iff) }
note remove_XX = this
+
show ?thesis
unfolding entropy_def mutual_information_generic_eq[OF MX MX]
unfolding setsum_cartesian_product[symmetric] setsum_negf[symmetric] remove_XX
@@ -440,7 +446,7 @@
lemma (in information_space) entropy_eq:
assumes "simple_function M X"
- shows "\<H>(X) = -(\<Sum> x \<in> X ` space M. real (distribution X {x}) * log b (real (distribution X {x})))"
+ shows "\<H>(X) = -(\<Sum> x \<in> X ` space M. distribution X {x} * log b (distribution X {x}))"
using assms by (simp add: entropy_generic_eq)
lemma (in information_space) entropy_positive:
@@ -448,63 +454,77 @@
unfolding entropy_def by (simp add: mutual_information_positive)
lemma (in information_space) entropy_certainty_eq_0:
- assumes "simple_function M X" and "x \<in> X ` space M" and "distribution X {x} = 1"
+ assumes X: "simple_function M X" and "x \<in> X ` space M" and "distribution X {x} = 1"
shows "\<H>(X) = 0"
proof -
- let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = distribution X\<rparr>"
+ let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = extreal\<circ>distribution X\<rparr>"
note simple_function_imp_finite_random_variable[OF `simple_function M X`]
- from distribution_finite_prob_space[OF this, of "\<lparr> measure = distribution X \<rparr>"]
+ from distribution_finite_prob_space[OF this, of "\<lparr> measure = extreal\<circ>distribution X \<rparr>"]
interpret X: finite_prob_space ?X by simp
have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}"
using X.measure_compl[of "{x}"] assms by auto
also have "\<dots> = 0" using X.prob_space assms by auto
finally have X0: "distribution X (X ` space M - {x}) = 0" by auto
- { fix y assume asm: "y \<noteq> x" "y \<in> X ` space M"
- hence "{y} \<subseteq> X ` space M - {x}" by auto
- from X.measure_mono[OF this] X0 asm
- have "distribution X {y} = 0" by auto }
- hence fi: "\<And> y. y \<in> X ` space M \<Longrightarrow> real (distribution X {y}) = (if x = y then 1 else 0)"
- using assms by auto
+ { fix y assume *: "y \<in> X ` space M"
+ { assume asm: "y \<noteq> x"
+ with * have "{y} \<subseteq> X ` space M - {x}" by auto
+ from X.measure_mono[OF this] X0 asm *
+ have "distribution X {y} = 0" by (auto intro: antisym) }
+ then have "distribution X {y} = (if x = y then 1 else 0)"
+ using assms by auto }
+ note fi = this
have y: "\<And>y. (if x = y then 1 else 0) * log b (if x = y then 1 else 0) = 0" by simp
show ?thesis unfolding entropy_eq[OF `simple_function M X`] by (auto simp: y fi)
qed
lemma (in information_space) entropy_le_card_not_0:
- assumes "simple_function M X"
- shows "\<H>(X) \<le> log b (real (card (X ` space M \<inter> {x . distribution X {x} \<noteq> 0})))"
+ assumes X: "simple_function M X"
+ shows "\<H>(X) \<le> log b (card (X ` space M \<inter> {x. distribution X {x} \<noteq> 0}))"
proof -
- let "?d x" = "distribution X {x}"
- let "?p x" = "real (?d x)"
+ let "?p x" = "distribution X {x}"
have "\<H>(X) = (\<Sum>x\<in>X`space M. ?p x * log b (1 / ?p x))"
- by (auto intro!: setsum_cong simp: entropy_eq[OF `simple_function M X`] setsum_negf[symmetric] log_simps not_less)
+ unfolding entropy_eq[OF X] setsum_negf[symmetric]
+ by (auto intro!: setsum_cong simp: log_simps)
also have "\<dots> \<le> log b (\<Sum>x\<in>X`space M. ?p x * (1 / ?p x))"
- apply (rule log_setsum')
- using not_empty b_gt_1 `simple_function M X` sum_over_space_real_distribution
- by (auto simp: simple_function_def)
- also have "\<dots> = log b (\<Sum>x\<in>X`space M. if ?d x \<noteq> 0 then 1 else 0)"
- using distribution_finite[OF `simple_function M X`[THEN simple_function_imp_random_variable], simplified]
- by (intro arg_cong[where f="\<lambda>X. log b X"] setsum_cong) (auto simp: real_of_pextreal_eq_0)
+ using not_empty b_gt_1 `simple_function M X` sum_over_space_real_distribution[OF X]
+ by (intro log_setsum') (auto simp: simple_function_def)
+ also have "\<dots> = log b (\<Sum>x\<in>X`space M. if ?p x \<noteq> 0 then 1 else 0)"
+ by (intro arg_cong[where f="\<lambda>X. log b X"] setsum_cong) auto
finally show ?thesis
using `simple_function M X` by (auto simp: setsum_cases real_eq_of_nat simple_function_def)
qed
+lemma (in prob_space) measure'_translate:
+ assumes X: "random_variable S X" and A: "A \<in> sets S"
+ shows "finite_measure.\<mu>' (S\<lparr> measure := extreal\<circ>distribution X \<rparr>) A = distribution X A"
+proof -
+ interpret S: prob_space "S\<lparr> measure := extreal\<circ>distribution X \<rparr>"
+ using distribution_prob_space[OF X] .
+ from A show "S.\<mu>' A = distribution X A"
+ unfolding S.\<mu>'_def by (simp add: distribution_def_raw \<mu>'_def)
+qed
+
lemma (in information_space) entropy_uniform_max:
- assumes "simple_function M X"
+ assumes X: "simple_function M X"
assumes "\<And>x y. \<lbrakk> x \<in> X ` space M ; y \<in> X ` space M \<rbrakk> \<Longrightarrow> distribution X {x} = distribution X {y}"
shows "\<H>(X) = log b (real (card (X ` space M)))"
proof -
- let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = distribution X\<rparr>"
- note simple_function_imp_finite_random_variable[OF `simple_function M X`]
- from distribution_finite_prob_space[OF this, of "\<lparr> measure = distribution X \<rparr>"]
+ let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = undefined\<rparr>\<lparr> measure := extreal\<circ>distribution X\<rparr>"
+ note frv = simple_function_imp_finite_random_variable[OF X]
+ from distribution_finite_prob_space[OF this, of "\<lparr> measure = extreal\<circ>distribution X \<rparr>"]
interpret X: finite_prob_space ?X by simp
+ note rv = finite_random_variableD[OF frv]
have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff
using `simple_function M X` not_empty by (auto simp: simple_function_def)
- { fix x assume "x \<in> X ` space M"
- hence "real (distribution X {x}) = 1 / real (card (X ` space M))"
- proof (rule X.uniform_prob[simplified])
- fix x y assume "x \<in> X`space M" "y \<in> X`space M"
- from assms(2)[OF this] show "real (distribution X {x}) = real (distribution X {y})" by simp
- qed }
+ { fix x assume "x \<in> space ?X"
+ moreover then have "X.\<mu>' {x} = 1 / card (space ?X)"
+ proof (rule X.uniform_prob)
+ fix x y assume "x \<in> space ?X" "y \<in> space ?X"
+ with assms(2)[of x y] show "X.\<mu>' {x} = X.\<mu>' {y}"
+ by (subst (1 2) measure'_translate[OF rv]) auto
+ qed
+ ultimately have "distribution X {x} = 1 / card (space ?X)"
+ by (subst (asm) measure'_translate[OF rv]) auto }
thus ?thesis
using not_empty X.finite_space b_gt_1 card_gt0
by (simp add: entropy_eq[OF `simple_function M X`] real_eq_of_nat[symmetric] log_simps)
@@ -552,8 +572,7 @@
lemma (in information_space) entropy_eq_cartesian_product:
assumes "simple_function M X" "simple_function M Y"
shows "\<H>(\<lambda>x. (X x, Y x)) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M.
- real (joint_distribution X Y {(x,y)}) *
- log b (real (joint_distribution X Y {(x,y)})))"
+ joint_distribution X Y {(x,y)} * log b (joint_distribution X Y {(x,y)}))"
proof -
have sf: "simple_function M (\<lambda>x. (X x, Y x))"
using assms by (auto intro: simple_function_Pair)
@@ -576,9 +595,9 @@
abbreviation (in information_space)
conditional_mutual_information_Pow ("\<I>'( _ ; _ | _ ')") where
"\<I>(X ; Y | Z) \<equiv> conditional_mutual_information b
- \<lparr> space = X`space M, sets = Pow (X`space M), measure = distribution X \<rparr>
- \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = distribution Y \<rparr>
- \<lparr> space = Z`space M, sets = Pow (Z`space M), measure = distribution Z \<rparr>
+ \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr>
+ \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = extreal\<circ>distribution Y \<rparr>
+ \<lparr> space = Z`space M, sets = Pow (Z`space M), measure = extreal\<circ>distribution Z \<rparr>
X Y Z"
lemma (in information_space) conditional_mutual_information_generic_eq:
@@ -586,58 +605,44 @@
and MY: "finite_random_variable MY Y"
and MZ: "finite_random_variable MZ Z"
shows "conditional_mutual_information b MX MY MZ X Y Z = (\<Sum>(x, y, z) \<in> space MX \<times> space MY \<times> space MZ.
- real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) *
- log b (real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) /
- (real (joint_distribution X Z {(x, z)}) * real (joint_distribution Y Z {(y,z)} / distribution Z {z}))))"
- (is "_ = (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * ?YZdZ y z)))")
+ distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} *
+ log b (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} /
+ (joint_distribution X Z {(x, z)} * (joint_distribution Y Z {(y,z)} / distribution Z {z}))))"
+ (is "_ = (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * (?YZ y z / ?Z z))))")
proof -
- let ?YZ = "\<lambda>y z. real (joint_distribution Y Z {(y, z)})"
- let ?X = "\<lambda>x. real (distribution X {x})"
- let ?Z = "\<lambda>z. real (distribution Z {z})"
-
- txt {* This proof is actually quiet easy, however we need to show that the
- distributions are finite and the joint distributions are zero when one of
- the variables distribution is also zero. *}
-
+ let ?X = "\<lambda>x. distribution X {x}"
note finite_var = MX MY MZ
- note random_var = finite_var[THEN finite_random_variableD]
-
- note space_simps = space_pair_measure space_sigma algebra.simps
-
note YZ = finite_random_variable_pairI[OF finite_var(2,3)]
+ note XYZ = finite_random_variable_pairI[OF MX YZ]
note XZ = finite_random_variable_pairI[OF finite_var(1,3)]
note ZX = finite_random_variable_pairI[OF finite_var(3,1)]
note YZX = finite_random_variable_pairI[OF finite_var(2) ZX]
note order1 =
- finite_distribution_order(5,6)[OF finite_var(1) YZ, simplified space_simps]
- finite_distribution_order(5,6)[OF finite_var(1,3), simplified space_simps]
+ finite_distribution_order(5,6)[OF finite_var(1) YZ]
+ finite_distribution_order(5,6)[OF finite_var(1,3)]
+ note random_var = finite_var[THEN finite_random_variableD]
note finite = finite_var(1) YZ finite_var(3) XZ YZX
- note finite[THEN finite_distribution_finite, simplified space_simps, simp]
have order2: "\<And>x y z. \<lbrakk>x \<in> space MX; y \<in> space MY; z \<in> space MZ; joint_distribution X Z {(x, z)} = 0\<rbrakk>
\<Longrightarrow> joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} = 0"
unfolding joint_distribution_commute_singleton[of X]
unfolding joint_distribution_assoc_singleton[symmetric]
using finite_distribution_order(6)[OF finite_var(2) ZX]
- by (auto simp: space_simps)
+ by auto
- have "(\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * ?YZdZ y z))) =
+ have "(\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * (?YZ y z / ?Z z)))) =
(\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * (log b (?XYZ x y z / (?X x * ?YZ y z)) - log b (?XZ x z / (?X x * ?Z z))))"
(is "(\<Sum>(x, y, z)\<in>?S. ?L x y z) = (\<Sum>(x, y, z)\<in>?S. ?R x y z)")
proof (safe intro!: setsum_cong)
fix x y z assume space: "x \<in> space MX" "y \<in> space MY" "z \<in> space MZ"
- then have *: "?XYZ x y z / (?XZ x z * ?YZdZ y z) =
- (?XYZ x y z / (?X x * ?YZ y z)) / (?XZ x z / (?X x * ?Z z))"
- using order1(3)
- by (auto simp: real_of_pextreal_mult[symmetric] real_of_pextreal_eq_0)
show "?L x y z = ?R x y z"
proof cases
assume "?XYZ x y z \<noteq> 0"
- with space b_gt_1 order1 order2 show ?thesis unfolding *
- by (subst log_divide)
- (auto simp: zero_less_divide_iff zero_less_real_of_pextreal
- real_of_pextreal_eq_0 zero_less_mult_iff)
+ with space have "0 < ?X x" "0 < ?Z z" "0 < ?XZ x z" "0 < ?YZ y z" "0 < ?XYZ x y z"
+ using order1 order2 by (auto simp: less_le)
+ with b_gt_1 show ?thesis
+ by (simp add: log_mult log_divide zero_less_mult_iff zero_less_divide_iff)
qed simp
qed
also have "\<dots> = (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?X x * ?YZ y z))) -
@@ -649,8 +654,8 @@
setsum_left_distrib[symmetric]
unfolding joint_distribution_commute_singleton[of X]
unfolding joint_distribution_assoc_singleton[symmetric]
- using setsum_real_joint_distribution_singleton[OF finite_var(2) ZX, unfolded space_simps]
- by (intro setsum_cong refl) simp
+ using setsum_joint_distribution_singleton[OF finite_var(2) ZX]
+ by (intro setsum_cong refl) (simp add: space_pair_measure)
also have "(\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?X x * ?YZ y z))) -
(\<Sum>(x, z)\<in>space MX \<times> space MZ. ?XZ x z * log b (?XZ x z / (?X x * ?Z z))) =
conditional_mutual_information b MX MY MZ X Y Z"
@@ -664,11 +669,11 @@
lemma (in information_space) conditional_mutual_information_eq:
assumes "simple_function M X" "simple_function M Y" "simple_function M Z"
shows "\<I>(X;Y|Z) = (\<Sum>(x, y, z) \<in> X`space M \<times> Y`space M \<times> Z`space M.
- real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) *
- log b (real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) /
- (real (joint_distribution X Z {(x, z)}) * real (joint_distribution Y Z {(y,z)} / distribution Z {z}))))"
- using conditional_mutual_information_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]]
- by simp
+ distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} *
+ log b (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} /
+ (joint_distribution X Z {(x, z)} * joint_distribution Y Z {(y,z)} / distribution Z {z})))"
+ by (subst conditional_mutual_information_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]])
+ simp
lemma (in information_space) conditional_mutual_information_eq_mutual_information:
assumes X: "simple_function M X" and Y: "simple_function M Y"
@@ -683,10 +688,10 @@
qed
lemma (in prob_space) distribution_unit[simp]: "distribution (\<lambda>x. ()) {()} = 1"
- unfolding distribution_def using measure_space_1 by auto
+ unfolding distribution_def using prob_space by auto
lemma (in prob_space) joint_distribution_unit[simp]: "distribution (\<lambda>x. (X x, ())) {(a, ())} = distribution X {a}"
- unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>])
+ unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
lemma (in prob_space) setsum_distribution:
assumes X: "finite_random_variable MX X" shows "(\<Sum>a\<in>space MX. distribution X {a}) = 1"
@@ -695,12 +700,13 @@
lemma (in prob_space) setsum_real_distribution:
fixes MX :: "('c, 'd) measure_space_scheme"
- assumes X: "finite_random_variable MX X" shows "(\<Sum>a\<in>space MX. real (distribution X {a})) = 1"
- using setsum_real_joint_distribution[OF assms, of "\<lparr> space = UNIV, sets = Pow UNIV, measure = undefined \<rparr>" "\<lambda>x. ()" "{()}"]
- using sigma_algebra_Pow[of "UNIV::unit set" "\<lparr> measure = undefined \<rparr>"] by simp
+ assumes X: "finite_random_variable MX X" shows "(\<Sum>a\<in>space MX. distribution X {a}) = 1"
+ using setsum_joint_distribution[OF assms, of "\<lparr> space = UNIV, sets = Pow UNIV, measure = undefined \<rparr>" "\<lambda>x. ()" "{()}"]
+ using sigma_algebra_Pow[of "UNIV::unit set" "\<lparr> measure = undefined \<rparr>"]
+ by auto
lemma (in information_space) conditional_mutual_information_generic_positive:
- assumes "finite_random_variable MX X" and "finite_random_variable MY Y" and "finite_random_variable MZ Z"
+ assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" and Z: "finite_random_variable MZ Z"
shows "0 \<le> conditional_mutual_information b MX MY MZ X Y Z"
proof (cases "space MX \<times> space MY \<times> space MZ = {}")
case True show ?thesis
@@ -708,43 +714,35 @@
by simp
next
case False
- let "?dXYZ A" = "real (distribution (\<lambda>x. (X x, Y x, Z x)) A)"
- let "?dXZ A" = "real (joint_distribution X Z A)"
- let "?dYZ A" = "real (joint_distribution Y Z A)"
- let "?dX A" = "real (distribution X A)"
- let "?dZ A" = "real (distribution Z A)"
+ let ?dXYZ = "distribution (\<lambda>x. (X x, Y x, Z x))"
+ let ?dXZ = "joint_distribution X Z"
+ let ?dYZ = "joint_distribution Y Z"
+ let ?dX = "distribution X"
+ let ?dZ = "distribution Z"
let ?M = "space MX \<times> space MY \<times> space MZ"
- have split_beta: "\<And>f. split f = (\<lambda>x. f (fst x) (snd x))" by (simp add: fun_eq_iff)
-
- note space_simps = space_pair_measure space_sigma algebra.simps
-
- note finite_var = assms
- note YZ = finite_random_variable_pairI[OF finite_var(2,3)]
- note XZ = finite_random_variable_pairI[OF finite_var(1,3)]
- note ZX = finite_random_variable_pairI[OF finite_var(3,1)]
- note YZ = finite_random_variable_pairI[OF finite_var(2,3)]
- note XYZ = finite_random_variable_pairI[OF finite_var(1) YZ]
- note finite = finite_var(3) YZ XZ XYZ
- note finite = finite[THEN finite_distribution_finite, simplified space_simps]
-
+ note YZ = finite_random_variable_pairI[OF Y Z]
+ note XZ = finite_random_variable_pairI[OF X Z]
+ note ZX = finite_random_variable_pairI[OF Z X]
+ note YZ = finite_random_variable_pairI[OF Y Z]
+ note XYZ = finite_random_variable_pairI[OF X YZ]
+ note finite = Z YZ XZ XYZ
have order: "\<And>x y z. \<lbrakk>x \<in> space MX; y \<in> space MY; z \<in> space MZ; joint_distribution X Z {(x, z)} = 0\<rbrakk>
\<Longrightarrow> joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} = 0"
unfolding joint_distribution_commute_singleton[of X]
unfolding joint_distribution_assoc_singleton[symmetric]
- using finite_distribution_order(6)[OF finite_var(2) ZX]
- by (auto simp: space_simps)
+ using finite_distribution_order(6)[OF Y ZX]
+ by auto
note order = order
- finite_distribution_order(5,6)[OF finite_var(1) YZ, simplified space_simps]
- finite_distribution_order(5,6)[OF finite_var(2,3), simplified space_simps]
+ finite_distribution_order(5,6)[OF X YZ]
+ finite_distribution_order(5,6)[OF Y Z]
have "- conditional_mutual_information b MX MY MZ X Y Z = - (\<Sum>(x, y, z) \<in> ?M. ?dXYZ {(x, y, z)} *
log b (?dXYZ {(x, y, z)} / (?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})))"
- unfolding conditional_mutual_information_generic_eq[OF assms] neg_equal_iff_equal
- by (intro setsum_cong) (auto intro!: arg_cong[where f="log b"] simp: real_of_pextreal_mult[symmetric])
+ unfolding conditional_mutual_information_generic_eq[OF assms] neg_equal_iff_equal by auto
also have "\<dots> \<le> log b (\<Sum>(x, y, z) \<in> ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})"
- unfolding split_beta
+ unfolding split_beta'
proof (rule log_setsum_divide)
show "?M \<noteq> {}" using False by simp
show "1 < b" using b_gt_1 .
@@ -757,33 +755,31 @@
unfolding setsum_commute[of _ "space MY"]
unfolding setsum_commute[of _ "space MZ"]
by (simp_all add: space_pair_measure
- setsum_real_joint_distribution_singleton[OF `finite_random_variable MX X` YZ]
- setsum_real_joint_distribution_singleton[OF `finite_random_variable MY Y` finite_var(3)]
- setsum_real_distribution[OF `finite_random_variable MZ Z`])
+ setsum_joint_distribution_singleton[OF X YZ]
+ setsum_joint_distribution_singleton[OF Y Z]
+ setsum_distribution[OF Z])
fix x assume "x \<in> ?M"
let ?x = "(fst x, fst (snd x), snd (snd x))"
- show "0 \<le> ?dXYZ {?x}" using real_pextreal_nonneg .
- show "0 \<le> ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
- by (simp add: real_pextreal_nonneg mult_nonneg_nonneg divide_nonneg_nonneg)
+ show "0 \<le> ?dXYZ {?x}"
+ "0 \<le> ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
+ by (simp_all add: mult_nonneg_nonneg divide_nonneg_nonneg)
assume *: "0 < ?dXYZ {?x}"
- with `x \<in> ?M` show "0 < ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
- using finite order
- by (cases x)
- (auto simp add: zero_less_real_of_pextreal zero_less_mult_iff zero_less_divide_iff)
+ with `x \<in> ?M` finite order show "0 < ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
+ by (cases x) (auto simp add: zero_le_mult_iff zero_le_divide_iff less_le)
qed
also have "(\<Sum>(x, y, z) \<in> ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z}) = (\<Sum>z\<in>space MZ. ?dZ {z})"
apply (simp add: setsum_cartesian_product')
apply (subst setsum_commute)
apply (subst (2) setsum_commute)
by (auto simp: setsum_divide_distrib[symmetric] setsum_product[symmetric]
- setsum_real_joint_distribution_singleton[OF finite_var(1,3)]
- setsum_real_joint_distribution_singleton[OF finite_var(2,3)]
+ setsum_joint_distribution_singleton[OF X Z]
+ setsum_joint_distribution_singleton[OF Y Z]
intro!: setsum_cong)
also have "log b (\<Sum>z\<in>space MZ. ?dZ {z}) = 0"
- unfolding setsum_real_distribution[OF finite_var(3)] by simp
+ unfolding setsum_real_distribution[OF Z] by simp
finally show ?thesis by simp
qed
@@ -800,57 +796,52 @@
abbreviation (in information_space)
conditional_entropy_Pow ("\<H>'(_ | _')") where
"\<H>(X | Y) \<equiv> conditional_entropy b
- \<lparr> space = X`space M, sets = Pow (X`space M), measure = distribution X \<rparr>
- \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = distribution Y \<rparr> X Y"
+ \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr>
+ \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = extreal\<circ>distribution Y \<rparr> X Y"
lemma (in information_space) conditional_entropy_positive:
"simple_function M X \<Longrightarrow> simple_function M Y \<Longrightarrow> 0 \<le> \<H>(X | Y)"
unfolding conditional_entropy_def by (auto intro!: conditional_mutual_information_positive)
-lemma (in measure_space) empty_measureI: "A = {} \<Longrightarrow> \<mu> A = 0" by simp
-
lemma (in information_space) conditional_entropy_generic_eq:
fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme"
assumes MX: "finite_random_variable MX X"
assumes MZ: "finite_random_variable MZ Z"
shows "conditional_entropy b MX MZ X Z =
- (\<Sum>(x, z)\<in>space MX \<times> space MZ.
- real (joint_distribution X Z {(x, z)}) *
- log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))"
+ joint_distribution X Z {(x, z)} * log b (joint_distribution X Z {(x, z)} / distribution Z {z}))"
proof -
interpret MX: finite_sigma_algebra MX using MX by simp
interpret MZ: finite_sigma_algebra MZ using MZ by simp
let "?XXZ x y z" = "joint_distribution X (\<lambda>x. (X x, Z x)) {(x, y, z)}"
let "?XZ x z" = "joint_distribution X Z {(x, z)}"
let "?Z z" = "distribution Z {z}"
- let "?f x y z" = "log b (real (?XXZ x y z) / (real (?XZ x z) * real (?XZ y z / ?Z z)))"
+ let "?f x y z" = "log b (?XXZ x y z * ?Z z / (?XZ x z * ?XZ y z))"
{ fix x z have "?XXZ x x z = ?XZ x z"
- unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>]) }
+ unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>']) }
note this[simp]
{ fix x x' :: 'c and z assume "x' \<noteq> x"
then have "?XXZ x x' z = 0"
- by (auto simp: distribution_def intro!: arg_cong[where f=\<mu>] empty_measureI) }
+ by (auto simp: distribution_def empty_measure'[symmetric]
+ simp del: empty_measure' intro!: arg_cong[where f=\<mu>']) }
note this[simp]
{ fix x x' z assume *: "x \<in> space MX" "z \<in> space MZ"
- then have "(\<Sum>x'\<in>space MX. real (?XXZ x x' z) * ?f x x' z)
- = (\<Sum>x'\<in>space MX. if x = x' then real (?XZ x z) * ?f x x z else 0)"
+ then have "(\<Sum>x'\<in>space MX. ?XXZ x x' z * ?f x x' z)
+ = (\<Sum>x'\<in>space MX. if x = x' then ?XZ x z * ?f x x z else 0)"
by (auto intro!: setsum_cong)
- also have "\<dots> = real (?XZ x z) * ?f x x z"
+ also have "\<dots> = ?XZ x z * ?f x x z"
using `x \<in> space MX` by (simp add: setsum_cases[OF MX.finite_space])
- also have "\<dots> = real (?XZ x z) * log b (real (?Z z) / real (?XZ x z))"
- by (auto simp: real_of_pextreal_mult[symmetric])
- also have "\<dots> = - real (?XZ x z) * log b (real (?XZ x z) / real (?Z z))"
- using assms[THEN finite_distribution_finite]
+ also have "\<dots> = ?XZ x z * log b (?Z z / ?XZ x z)" by auto
+ also have "\<dots> = - ?XZ x z * log b (?XZ x z / ?Z z)"
using finite_distribution_order(6)[OF MX MZ]
- by (auto simp: log_simps field_simps zero_less_mult_iff zero_less_real_of_pextreal real_of_pextreal_eq_0)
- finally have "(\<Sum>x'\<in>space MX. real (?XXZ x x' z) * ?f x x' z) =
- - real (?XZ x z) * log b (real (?XZ x z) / real (?Z z))" . }
+ by (auto simp: log_simps field_simps zero_less_mult_iff)
+ finally have "(\<Sum>x'\<in>space MX. ?XXZ x x' z * ?f x x' z) = - ?XZ x z * log b (?XZ x z / ?Z z)" . }
note * = this
show ?thesis
unfolding conditional_entropy_def
unfolding conditional_mutual_information_generic_eq[OF MX MX MZ]
by (auto simp: setsum_cartesian_product' setsum_negf[symmetric]
- setsum_commute[of _ "space MZ"] * simp del: divide_pextreal_def
+ setsum_commute[of _ "space MZ"] *
intro!: setsum_cong)
qed
@@ -858,29 +849,27 @@
assumes "simple_function M X" "simple_function M Z"
shows "\<H>(X | Z) =
- (\<Sum>(x, z)\<in>X ` space M \<times> Z ` space M.
- real (joint_distribution X Z {(x, z)}) *
- log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))"
- using conditional_entropy_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]]
- by simp
+ joint_distribution X Z {(x, z)} *
+ log b (joint_distribution X Z {(x, z)} / distribution Z {z}))"
+ by (subst conditional_entropy_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]])
+ simp
lemma (in information_space) conditional_entropy_eq_ce_with_hypothesis:
assumes X: "simple_function M X" and Y: "simple_function M Y"
shows "\<H>(X | Y) =
- -(\<Sum>y\<in>Y`space M. real (distribution Y {y}) *
- (\<Sum>x\<in>X`space M. real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}) *
- log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}))))"
+ -(\<Sum>y\<in>Y`space M. distribution Y {y} *
+ (\<Sum>x\<in>X`space M. joint_distribution X Y {(x,y)} / distribution Y {(y)} *
+ log b (joint_distribution X Y {(x,y)} / distribution Y {(y)})))"
unfolding conditional_entropy_eq[OF assms]
- using finite_distribution_finite[OF finite_random_variable_pairI[OF assms[THEN simple_function_imp_finite_random_variable]]]
using finite_distribution_order(5,6)[OF assms[THEN simple_function_imp_finite_random_variable]]
- using finite_distribution_finite[OF Y[THEN simple_function_imp_finite_random_variable]]
- by (auto simp: setsum_cartesian_product' setsum_commute[of _ "Y`space M"] setsum_right_distrib real_of_pextreal_eq_0
+ by (auto simp: setsum_cartesian_product' setsum_commute[of _ "Y`space M"] setsum_right_distrib
intro!: setsum_cong)
lemma (in information_space) conditional_entropy_eq_cartesian_product:
assumes "simple_function M X" "simple_function M Y"
shows "\<H>(X | Y) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M.
- real (joint_distribution X Y {(x,y)}) *
- log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {y})))"
+ joint_distribution X Y {(x,y)} *
+ log b (joint_distribution X Y {(x,y)} / distribution Y {y}))"
unfolding conditional_entropy_eq[OF assms]
by (auto intro!: setsum_cong simp: setsum_cartesian_product')
@@ -890,24 +879,22 @@
assumes X: "simple_function M X" and Z: "simple_function M Z"
shows "\<I>(X ; Z) = \<H>(X) - \<H>(X | Z)"
proof -
- let "?XZ x z" = "real (joint_distribution X Z {(x, z)})"
- let "?Z z" = "real (distribution Z {z})"
- let "?X x" = "real (distribution X {x})"
+ let "?XZ x z" = "joint_distribution X Z {(x, z)}"
+ let "?Z z" = "distribution Z {z}"
+ let "?X x" = "distribution X {x}"
note fX = X[THEN simple_function_imp_finite_random_variable]
note fZ = Z[THEN simple_function_imp_finite_random_variable]
- note fX[THEN finite_distribution_finite, simp] and fZ[THEN finite_distribution_finite, simp]
note finite_distribution_order[OF fX fZ, simp]
{ fix x z assume "x \<in> X`space M" "z \<in> Z`space M"
have "?XZ x z * log b (?XZ x z / (?X x * ?Z z)) =
?XZ x z * log b (?XZ x z / ?Z z) - ?XZ x z * log b (?X x)"
- by (auto simp: log_simps real_of_pextreal_mult[symmetric] zero_less_mult_iff
- zero_less_real_of_pextreal field_simps real_of_pextreal_eq_0 abs_mult) }
+ by (auto simp: log_simps zero_le_mult_iff field_simps less_le) }
note * = this
show ?thesis
unfolding entropy_eq[OF X] conditional_entropy_eq[OF X Z] mutual_information_eq[OF X Z]
- using setsum_real_joint_distribution_singleton[OF fZ fX, unfolded joint_distribution_commute_singleton[of Z X]]
+ using setsum_joint_distribution_singleton[OF fZ fX, unfolded joint_distribution_commute_singleton[of Z X]]
by (simp add: * setsum_cartesian_product' setsum_subtractf setsum_left_distrib[symmetric]
- setsum_real_distribution)
+ setsum_distribution)
qed
lemma (in information_space) conditional_entropy_less_eq_entropy:
@@ -923,21 +910,19 @@
assumes X: "simple_function M X" and Y: "simple_function M Y"
shows "\<H>(\<lambda>x. (X x, Y x)) = \<H>(X) + \<H>(Y|X)"
proof -
- let "?XY x y" = "real (joint_distribution X Y {(x, y)})"
- let "?Y y" = "real (distribution Y {y})"
- let "?X x" = "real (distribution X {x})"
+ let "?XY x y" = "joint_distribution X Y {(x, y)}"
+ let "?Y y" = "distribution Y {y}"
+ let "?X x" = "distribution X {x}"
note fX = X[THEN simple_function_imp_finite_random_variable]
note fY = Y[THEN simple_function_imp_finite_random_variable]
- note fX[THEN finite_distribution_finite, simp] and fY[THEN finite_distribution_finite, simp]
note finite_distribution_order[OF fX fY, simp]
{ fix x y assume "x \<in> X`space M" "y \<in> Y`space M"
have "?XY x y * log b (?XY x y / ?X x) =
?XY x y * log b (?XY x y) - ?XY x y * log b (?X x)"
- by (auto simp: log_simps real_of_pextreal_mult[symmetric] zero_less_mult_iff
- zero_less_real_of_pextreal field_simps real_of_pextreal_eq_0 abs_mult) }
+ by (auto simp: log_simps zero_le_mult_iff field_simps less_le) }
note * = this
show ?thesis
- using setsum_real_joint_distribution_singleton[OF fY fX]
+ using setsum_joint_distribution_singleton[OF fY fX]
unfolding entropy_eq[OF X] conditional_entropy_eq_cartesian_product[OF Y X] entropy_eq_cartesian_product[OF X Y]
unfolding joint_distribution_commute_singleton[of Y X] setsum_commute[of _ "X`space M"]
by (simp add: * setsum_subtractf setsum_left_distrib[symmetric])
@@ -1063,23 +1048,21 @@
assumes svi: "subvimage (space M) X P"
shows "\<H>(X) = \<H>(P) + \<H>(X|P)"
proof -
- let "?XP x p" = "real (joint_distribution X P {(x, p)})"
- let "?X x" = "real (distribution X {x})"
- let "?P p" = "real (distribution P {p})"
+ let "?XP x p" = "joint_distribution X P {(x, p)}"
+ let "?X x" = "distribution X {x}"
+ let "?P p" = "distribution P {p}"
note fX = sf(1)[THEN simple_function_imp_finite_random_variable]
note fP = sf(2)[THEN simple_function_imp_finite_random_variable]
- note fX[THEN finite_distribution_finite, simp] and fP[THEN finite_distribution_finite, simp]
note finite_distribution_order[OF fX fP, simp]
- have "(\<Sum>x\<in>X ` space M. real (distribution X {x}) * log b (real (distribution X {x}))) =
- (\<Sum>y\<in>P `space M. \<Sum>x\<in>X ` space M.
- real (joint_distribution X P {(x, y)}) * log b (real (joint_distribution X P {(x, y)})))"
+ have "(\<Sum>x\<in>X ` space M. ?X x * log b (?X x)) =
+ (\<Sum>y\<in>P `space M. \<Sum>x\<in>X ` space M. ?XP x y * log b (?XP x y))"
proof (subst setsum_image_split[OF svi],
safe intro!: setsum_mono_zero_cong_left imageI)
show "finite (X ` space M)" "finite (X ` space M)" "finite (P ` space M)"
using sf unfolding simple_function_def by auto
next
fix p x assume in_space: "p \<in> space M" "x \<in> space M"
- assume "real (joint_distribution X P {(X x, P p)}) * log b (real (joint_distribution X P {(X x, P p)})) \<noteq> 0"
+ assume "?XP (X x) (P p) * log b (?XP (X x) (P p)) \<noteq> 0"
hence "(\<lambda>x. (X x, P x)) -` {(X x, P p)} \<inter> space M \<noteq> {}" by (auto simp: distribution_def)
with svi[unfolded subvimage_def, rule_format, OF `x \<in> space M`]
show "x \<in> P -` {P p}" by auto
@@ -1091,20 +1074,16 @@
by auto
hence "(\<lambda>x. (X x, P x)) -` {(X x, P p)} \<inter> space M = X -` {X x} \<inter> space M"
by auto
- thus "real (distribution X {X x}) * log b (real (distribution X {X x})) =
- real (joint_distribution X P {(X x, P p)}) *
- log b (real (joint_distribution X P {(X x, P p)}))"
+ thus "?X (X x) * log b (?X (X x)) = ?XP (X x) (P p) * log b (?XP (X x) (P p))"
by (auto simp: distribution_def)
qed
- moreover have "\<And>x y. real (joint_distribution X P {(x, y)}) *
- log b (real (joint_distribution X P {(x, y)}) / real (distribution P {y})) =
- real (joint_distribution X P {(x, y)}) * log b (real (joint_distribution X P {(x, y)})) -
- real (joint_distribution X P {(x, y)}) * log b (real (distribution P {y}))"
+ moreover have "\<And>x y. ?XP x y * log b (?XP x y / ?P y) =
+ ?XP x y * log b (?XP x y) - ?XP x y * log b (?P y)"
by (auto simp add: log_simps zero_less_mult_iff field_simps)
ultimately show ?thesis
unfolding sf[THEN entropy_eq] conditional_entropy_eq[OF sf]
- using setsum_real_joint_distribution_singleton[OF fX fP]
- by (simp add: setsum_cartesian_product' setsum_subtractf setsum_real_distribution
+ using setsum_joint_distribution_singleton[OF fX fP]
+ by (simp add: setsum_cartesian_product' setsum_subtractf setsum_distribution
setsum_left_distrib[symmetric] setsum_commute[where B="P`space M"])
qed
--- a/src/HOL/Probability/Lebesgue_Integration.thy Mon Mar 14 14:37:47 2011 +0100
+++ b/src/HOL/Probability/Lebesgue_Integration.thy Mon Mar 14 14:37:49 2011 +0100
@@ -6,6 +6,88 @@
imports Measure Borel_Space
begin
+lemma extreal_indicator_pos[simp,intro]: "0 \<le> (indicator A x ::extreal)"
+ unfolding indicator_def by auto
+
+lemma tendsto_real_max:
+ fixes x y :: real
+ assumes "(X ---> x) net"
+ assumes "(Y ---> y) net"
+ shows "((\<lambda>x. max (X x) (Y x)) ---> max x y) net"
+proof -
+ have *: "\<And>x y :: real. max x y = y + ((x - y) + norm (x - y)) / 2"
+ by (auto split: split_max simp: field_simps)
+ show ?thesis
+ unfolding *
+ by (intro tendsto_add assms tendsto_divide tendsto_norm tendsto_diff) auto
+qed
+
+lemma (in measure_space) measure_Union:
+ assumes "finite S" "S \<subseteq> sets M" "\<And>A B. A \<in> S \<Longrightarrow> B \<in> S \<Longrightarrow> A \<noteq> B \<Longrightarrow> A \<inter> B = {}"
+ shows "setsum \<mu> S = \<mu> (\<Union>S)"
+proof -
+ have "setsum \<mu> S = \<mu> (\<Union>i\<in>S. i)"
+ using assms by (intro measure_setsum[OF `finite S`]) (auto simp: disjoint_family_on_def)
+ also have "\<dots> = \<mu> (\<Union>S)" by (auto intro!: arg_cong[where f=\<mu>])
+ finally show ?thesis .
+qed
+
+lemma (in sigma_algebra) measurable_sets2[intro]:
+ assumes "f \<in> measurable M M'" "g \<in> measurable M M''"
+ and "A \<in> sets M'" "B \<in> sets M''"
+ shows "f -` A \<inter> g -` B \<inter> space M \<in> sets M"
+proof -
+ have "f -` A \<inter> g -` B \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)"
+ by auto
+ then show ?thesis using assms by (auto intro: measurable_sets)
+qed
+
+lemma incseq_extreal: "incseq f \<Longrightarrow> incseq (\<lambda>x. extreal (f x))"
+ unfolding incseq_def by auto
+
+lemma incseq_Suc_iff: "incseq f \<longleftrightarrow> (\<forall>n. f n \<le> f (Suc n))"
+proof
+ assume "\<forall>n. f n \<le> f (Suc n)" then show "incseq f" by (auto intro!: incseq_SucI)
+qed (auto simp: incseq_def)
+
+lemma borel_measurable_real_floor:
+ "(\<lambda>x::real. real \<lfloor>x\<rfloor>) \<in> borel_measurable borel"
+ unfolding borel.borel_measurable_iff_ge
+proof (intro allI)
+ fix a :: real
+ { fix x have "a \<le> real \<lfloor>x\<rfloor> \<longleftrightarrow> real \<lceil>a\<rceil> \<le> x"
+ using le_floor_eq[of "\<lceil>a\<rceil>" x] ceiling_le_iff[of a "\<lfloor>x\<rfloor>"]
+ unfolding real_eq_of_int by simp }
+ then have "{w::real \<in> space borel. a \<le> real \<lfloor>w\<rfloor>} = {real \<lceil>a\<rceil>..}" by auto
+ then show "{w::real \<in> space borel. a \<le> real \<lfloor>w\<rfloor>} \<in> sets borel" by auto
+qed
+
+lemma measure_preservingD2:
+ "f \<in> measure_preserving A B \<Longrightarrow> f \<in> measurable A B"
+ unfolding measure_preserving_def by auto
+
+lemma measure_preservingD3:
+ "f \<in> measure_preserving A B \<Longrightarrow> f \<in> space A \<rightarrow> space B"
+ unfolding measure_preserving_def measurable_def by auto
+
+lemma measure_preservingD:
+ "T \<in> measure_preserving A B \<Longrightarrow> X \<in> sets B \<Longrightarrow> measure A (T -` X \<inter> space A) = measure B X"
+ unfolding measure_preserving_def by auto
+
+lemma (in sigma_algebra) borel_measurable_real_natfloor[intro, simp]:
+ assumes "f \<in> borel_measurable M"
+ shows "(\<lambda>x. real (natfloor (f x))) \<in> borel_measurable M"
+proof -
+ have "\<And>x. real (natfloor (f x)) = max 0 (real \<lfloor>f x\<rfloor>)"
+ by (auto simp: max_def natfloor_def)
+ with borel_measurable_max[OF measurable_comp[OF assms borel_measurable_real_floor] borel_measurable_const]
+ show ?thesis by (simp add: comp_def)
+qed
+
+lemma (in measure_space) AE_not_in:
+ assumes N: "N \<in> null_sets" shows "AE x. x \<notin> N"
+ using N by (rule AE_I') auto
+
lemma sums_If_finite:
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
assumes finite: "finite {r. P r}"
@@ -55,8 +137,17 @@
by (auto intro!: finite_UN simp del: UN_simps simp: simple_function_def)
qed
+lemma (in sigma_algebra) simple_function_measurable2[intro]:
+ assumes "simple_function M f" "simple_function M g"
+ shows "f -` A \<inter> g -` B \<inter> space M \<in> sets M"
+proof -
+ have "f -` A \<inter> g -` B \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)"
+ by auto
+ then show ?thesis using assms[THEN simple_functionD(2)] by auto
+qed
+
lemma (in sigma_algebra) simple_function_indicator_representation:
- fixes f ::"'a \<Rightarrow> pextreal"
+ fixes f ::"'a \<Rightarrow> extreal"
assumes f: "simple_function M f" and x: "x \<in> space M"
shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
(is "?l = ?r")
@@ -71,7 +162,7 @@
qed
lemma (in measure_space) simple_function_notspace:
- "simple_function M (\<lambda>x. h x * indicator (- space M) x::pextreal)" (is "simple_function M ?h")
+ "simple_function M (\<lambda>x. h x * indicator (- space M) x::extreal)" (is "simple_function M ?h")
proof -
have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset)
@@ -111,16 +202,22 @@
qed
lemma (in sigma_algebra) simple_function_borel_measurable:
- fixes f :: "'a \<Rightarrow> 'x::t2_space"
+ fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
assumes "f \<in> borel_measurable M" and "finite (f ` space M)"
shows "simple_function M f"
using assms unfolding simple_function_def
by (auto intro: borel_measurable_vimage)
+lemma (in sigma_algebra) simple_function_eq_borel_measurable:
+ fixes f :: "'a \<Rightarrow> extreal"
+ shows "simple_function M f \<longleftrightarrow> finite (f`space M) \<and> f \<in> borel_measurable M"
+ using simple_function_borel_measurable[of f]
+ borel_measurable_simple_function[of f]
+ by (fastsimp simp: simple_function_def)
+
lemma (in sigma_algebra) simple_function_const[intro, simp]:
"simple_function M (\<lambda>x. c)"
by (auto intro: finite_subset simp: simple_function_def)
-
lemma (in sigma_algebra) simple_function_compose[intro, simp]:
assumes "simple_function M f"
shows "simple_function M (g \<circ> f)"
@@ -189,6 +286,7 @@
and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"]
and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"]
and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"]
+ and simple_function_max[intro, simp] = simple_function_compose2[where h=max]
lemma (in sigma_algebra) simple_function_setsum[intro, simp]:
assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
@@ -197,247 +295,168 @@
assume "finite P" from this assms show ?thesis by induct auto
qed auto
-lemma (in sigma_algebra) simple_function_le_measurable:
- assumes "simple_function M f" "simple_function M g"
- shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
-proof -
- have *: "{x \<in> space M. f x \<le> g x} =
- (\<Union>(F, G)\<in>f`space M \<times> g`space M.
- if F \<le> G then (f -` {F} \<inter> space M) \<inter> (g -` {G} \<inter> space M) else {})"
- apply (auto split: split_if_asm)
- apply (rule_tac x=x in bexI)
- apply (rule_tac x=x in bexI)
- by simp_all
- have **: "\<And>x y. x \<in> space M \<Longrightarrow> y \<in> space M \<Longrightarrow>
- (f -` {f x} \<inter> space M) \<inter> (g -` {g y} \<inter> space M) \<in> sets M"
- using assms unfolding simple_function_def by auto
- have "finite (f`space M \<times> g`space M)"
- using assms unfolding simple_function_def by auto
- thus ?thesis unfolding *
- apply (rule finite_UN)
- using assms unfolding simple_function_def
- by (auto intro!: **)
-qed
+lemma (in sigma_algebra)
+ fixes f g :: "'a \<Rightarrow> real" assumes sf: "simple_function M f"
+ shows simple_function_extreal[intro, simp]: "simple_function M (\<lambda>x. extreal (f x))"
+ by (auto intro!: simple_function_compose1[OF sf])
+
+lemma (in sigma_algebra)
+ fixes f g :: "'a \<Rightarrow> nat" assumes sf: "simple_function M f"
+ shows simple_function_real_of_nat[intro, simp]: "simple_function M (\<lambda>x. real (f x))"
+ by (auto intro!: simple_function_compose1[OF sf])
lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence:
- fixes u :: "'a \<Rightarrow> pextreal"
+ fixes u :: "'a \<Rightarrow> extreal"
assumes u: "u \<in> borel_measurable M"
- shows "\<exists>f. (\<forall>i. simple_function M (f i) \<and> (\<forall>x\<in>space M. f i x \<noteq> \<omega>)) \<and> f \<up> u"
+ shows "\<exists>f. incseq f \<and> (\<forall>i. \<infinity> \<notin> range (f i) \<and> simple_function M (f i)) \<and>
+ (\<forall>x. (SUP i. f i x) = max 0 (u x)) \<and> (\<forall>i x. 0 \<le> f i x)"
proof -
- have "\<exists>f. \<forall>x j. (of_nat j \<le> u x \<longrightarrow> f x j = j*2^j) \<and>
- (u x < of_nat j \<longrightarrow> of_nat (f x j) \<le> u x * 2^j \<and> u x * 2^j < of_nat (Suc (f x j)))"
- (is "\<exists>f. \<forall>x j. ?P x j (f x j)")
- proof(rule choice, rule, rule choice, rule)
- fix x j show "\<exists>n. ?P x j n"
- proof cases
- assume *: "u x < of_nat j"
- then obtain r where r: "u x = Real r" "0 \<le> r" by (cases "u x") auto
- from reals_Archimedean6a[of "r * 2^j"]
- obtain n :: nat where "real n \<le> r * 2 ^ j" "r * 2 ^ j < real (Suc n)"
- using `0 \<le> r` by (auto simp: zero_le_power zero_le_mult_iff)
- thus ?thesis using r * by (auto intro!: exI[of _ n])
- qed auto
- qed
- then obtain f where top: "\<And>j x. of_nat j \<le> u x \<Longrightarrow> f x j = j*2^j" and
- upper: "\<And>j x. u x < of_nat j \<Longrightarrow> of_nat (f x j) \<le> u x * 2^j" and
- lower: "\<And>j x. u x < of_nat j \<Longrightarrow> u x * 2^j < of_nat (Suc (f x j))" by blast
-
- { fix j x P
- assume 1: "of_nat j \<le> u x \<Longrightarrow> P (j * 2^j)"
- assume 2: "\<And>k. \<lbrakk> u x < of_nat j ; of_nat k \<le> u x * 2^j ; u x * 2^j < of_nat (Suc k) \<rbrakk> \<Longrightarrow> P k"
- have "P (f x j)"
- proof cases
- assume "of_nat j \<le> u x" thus "P (f x j)"
- using top[of j x] 1 by auto
- next
- assume "\<not> of_nat j \<le> u x"
- hence "u x < of_nat j" "of_nat (f x j) \<le> u x * 2^j" "u x * 2^j < of_nat (Suc (f x j))"
- using upper lower by auto
- from 2[OF this] show "P (f x j)" .
- qed }
- note fI = this
-
- { fix j x
- have "f x j = j * 2 ^ j \<longleftrightarrow> of_nat j \<le> u x"
- by (rule fI, simp, cases "u x") (auto split: split_if_asm) }
- note f_eq = this
-
- { fix j x
- have "f x j \<le> j * 2 ^ j"
- proof (rule fI)
- fix k assume *: "u x < of_nat j"
- assume "of_nat k \<le> u x * 2 ^ j"
- also have "\<dots> \<le> of_nat (j * 2^j)"
- using * by (cases "u x") (auto simp: zero_le_mult_iff)
- finally show "k \<le> j*2^j" by (auto simp del: real_of_nat_mult)
- qed simp }
+ def f \<equiv> "\<lambda>x i. if real i \<le> u x then i * 2 ^ i else natfloor (real (u x) * 2 ^ i)"
+ { fix x j have "f x j \<le> j * 2 ^ j" unfolding f_def
+ proof (split split_if, intro conjI impI)
+ assume "\<not> real j \<le> u x"
+ then have "natfloor (real (u x) * 2 ^ j) \<le> natfloor (j * 2 ^ j)"
+ by (cases "u x") (auto intro!: natfloor_mono simp: mult_nonneg_nonneg)
+ moreover have "real (natfloor (j * 2 ^ j)) \<le> j * 2^j"
+ by (intro real_natfloor_le) (auto simp: mult_nonneg_nonneg)
+ ultimately show "natfloor (real (u x) * 2 ^ j) \<le> j * 2 ^ j"
+ unfolding real_of_nat_le_iff by auto
+ qed auto }
note f_upper = this
- let "?g j x" = "of_nat (f x j) / 2^j :: pextreal"
- show ?thesis unfolding simple_function_def isoton_fun_expand unfolding isoton_def
- proof (safe intro!: exI[of _ ?g])
- fix j
- have *: "?g j ` space M \<subseteq> (\<lambda>x. of_nat x / 2^j) ` {..j * 2^j}"
- using f_upper by auto
- thus "finite (?g j ` space M)" by (rule finite_subset) auto
- next
- fix j t assume "t \<in> space M"
- have **: "?g j -` {?g j t} \<inter> space M = {x \<in> space M. f t j = f x j}"
- by (auto simp: power_le_zero_eq Real_eq_Real mult_le_0_iff)
+ have real_f:
+ "\<And>i x. real (f x i) = (if real i \<le> u x then i * 2 ^ i else real (natfloor (real (u x) * 2 ^ i)))"
+ unfolding f_def by auto
- show "?g j -` {?g j t} \<inter> space M \<in> sets M"
- proof cases
- assume "of_nat j \<le> u t"
- hence "?g j -` {?g j t} \<inter> space M = {x\<in>space M. of_nat j \<le> u x}"
- unfolding ** f_eq[symmetric] by auto
- thus "?g j -` {?g j t} \<inter> space M \<in> sets M"
- using u by auto
- next
- assume not_t: "\<not> of_nat j \<le> u t"
- hence less: "f t j < j*2^j" using f_eq[symmetric] `f t j \<le> j*2^j` by auto
- have split_vimage: "?g j -` {?g j t} \<inter> space M =
- {x\<in>space M. of_nat (f t j)/2^j \<le> u x} \<inter> {x\<in>space M. u x < of_nat (Suc (f t j))/2^j}"
- unfolding **
- proof safe
- fix x assume [simp]: "f t j = f x j"
- have *: "\<not> of_nat j \<le> u x" using not_t f_eq[symmetric] by simp
- hence "of_nat (f t j) \<le> u x * 2^j \<and> u x * 2^j < of_nat (Suc (f t j))"
- using upper lower by auto
- hence "of_nat (f t j) / 2^j \<le> u x \<and> u x < of_nat (Suc (f t j))/2^j" using *
- by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps)
- thus "of_nat (f t j) / 2^j \<le> u x" "u x < of_nat (Suc (f t j))/2^j" by auto
+ let "?g j x" = "real (f x j) / 2^j :: extreal"
+ show ?thesis
+ proof (intro exI[of _ ?g] conjI allI ballI)
+ fix i
+ have "simple_function M (\<lambda>x. real (f x i))"
+ proof (intro simple_function_borel_measurable)
+ show "(\<lambda>x. real (f x i)) \<in> borel_measurable M"
+ using u by (auto intro!: measurable_If simp: real_f)
+ have "(\<lambda>x. real (f x i))`space M \<subseteq> real`{..i*2^i}"
+ using f_upper[of _ i] by auto
+ then show "finite ((\<lambda>x. real (f x i))`space M)"
+ by (rule finite_subset) auto
+ qed
+ then show "simple_function M (?g i)"
+ by (auto intro: simple_function_extreal simple_function_div)
+ next
+ show "incseq ?g"
+ proof (intro incseq_extreal incseq_SucI le_funI)
+ fix x and i :: nat
+ have "f x i * 2 \<le> f x (Suc i)" unfolding f_def
+ proof ((split split_if)+, intro conjI impI)
+ assume "extreal (real i) \<le> u x" "\<not> extreal (real (Suc i)) \<le> u x"
+ then show "i * 2 ^ i * 2 \<le> natfloor (real (u x) * 2 ^ Suc i)"
+ by (cases "u x") (auto intro!: le_natfloor)
next
- fix x
- assume "of_nat (f t j) / 2^j \<le> u x" "u x < of_nat (Suc (f t j))/2^j"
- hence "of_nat (f t j) \<le> u x * 2 ^ j \<and> u x * 2 ^ j < of_nat (Suc (f t j))"
- by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps)
- hence 1: "of_nat (f t j) \<le> u x * 2 ^ j" and 2: "u x * 2 ^ j < of_nat (Suc (f t j))" by auto
- note 2
- also have "\<dots> \<le> of_nat (j*2^j)"
- using less by (auto simp: zero_le_mult_iff simp del: real_of_nat_mult)
- finally have bound_ux: "u x < of_nat j"
- by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq)
- show "f t j = f x j"
- proof (rule antisym)
- from 1 lower[OF bound_ux]
- show "f t j \<le> f x j" by (cases "u x") (auto split: split_if_asm)
- from upper[OF bound_ux] 2
- show "f x j \<le> f t j" by (cases "u x") (auto split: split_if_asm)
+ assume "\<not> extreal (real i) \<le> u x" "extreal (real (Suc i)) \<le> u x"
+ then show "natfloor (real (u x) * 2 ^ i) * 2 \<le> Suc i * 2 ^ Suc i"
+ by (cases "u x") auto
+ next
+ assume "\<not> extreal (real i) \<le> u x" "\<not> extreal (real (Suc i)) \<le> u x"
+ have "natfloor (real (u x) * 2 ^ i) * 2 = natfloor (real (u x) * 2 ^ i) * natfloor 2"
+ by simp
+ also have "\<dots> \<le> natfloor (real (u x) * 2 ^ i * 2)"
+ proof cases
+ assume "0 \<le> u x" then show ?thesis
+ by (intro le_mult_natfloor) (cases "u x", auto intro!: mult_nonneg_nonneg)
+ next
+ assume "\<not> 0 \<le> u x" then show ?thesis
+ by (cases "u x") (auto simp: natfloor_neg mult_nonpos_nonneg)
qed
- qed
- show ?thesis unfolding split_vimage using u by auto
+ also have "\<dots> = natfloor (real (u x) * 2 ^ Suc i)"
+ by (simp add: ac_simps)
+ finally show "natfloor (real (u x) * 2 ^ i) * 2 \<le> natfloor (real (u x) * 2 ^ Suc i)" .
+ qed simp
+ then show "?g i x \<le> ?g (Suc i) x"
+ by (auto simp: field_simps)
qed
next
- fix j t assume "?g j t = \<omega>" thus False by (auto simp: power_le_zero_eq)
- next
- fix t
- { fix i
- have "f t i * 2 \<le> f t (Suc i)"
- proof (rule fI)
- assume "of_nat (Suc i) \<le> u t"
- hence "of_nat i \<le> u t" by (cases "u t") auto
- thus "f t i * 2 \<le> Suc i * 2 ^ Suc i" unfolding f_eq[symmetric] by simp
- next
- fix k
- assume *: "u t * 2 ^ Suc i < of_nat (Suc k)"
- show "f t i * 2 \<le> k"
- proof (rule fI)
- assume "of_nat i \<le> u t"
- hence "of_nat (i * 2^Suc i) \<le> u t * 2^Suc i"
- by (cases "u t") (auto simp: zero_le_mult_iff power_le_zero_eq)
- also have "\<dots> < of_nat (Suc k)" using * by auto
- finally show "i * 2 ^ i * 2 \<le> k"
- by (auto simp del: real_of_nat_mult)
- next
- fix j assume "of_nat j \<le> u t * 2 ^ i"
- with * show "j * 2 \<le> k" by (cases "u t") (auto simp: zero_le_mult_iff power_le_zero_eq)
+ fix x show "(SUP i. ?g i x) = max 0 (u x)"
+ proof (rule extreal_SUPI)
+ fix i show "?g i x \<le> max 0 (u x)" unfolding max_def f_def
+ by (cases "u x") (auto simp: field_simps real_natfloor_le natfloor_neg
+ mult_nonpos_nonneg mult_nonneg_nonneg)
+ next
+ fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> ?g i x \<le> y"
+ have "\<And>i. 0 \<le> ?g i x" by (auto simp: divide_nonneg_pos)
+ from order_trans[OF this *] have "0 \<le> y" by simp
+ show "max 0 (u x) \<le> y"
+ proof (cases y)
+ case (real r)
+ with * have *: "\<And>i. f x i \<le> r * 2^i" by (auto simp: divide_le_eq)
+ from real_arch_lt[of r] * have "u x \<noteq> \<infinity>" by (auto simp: f_def) (metis less_le_not_le)
+ then have "\<exists>p. max 0 (u x) = extreal p \<and> 0 \<le> p" by (cases "u x") (auto simp: max_def)
+ then guess p .. note ux = this
+ obtain m :: nat where m: "p < real m" using real_arch_lt ..
+ have "p \<le> r"
+ proof (rule ccontr)
+ assume "\<not> p \<le> r"
+ with LIMSEQ_inverse_realpow_zero[unfolded LIMSEQ_iff, rule_format, of 2 "p - r"]
+ obtain N where "\<forall>n\<ge>N. r * 2^n < p * 2^n - 1" by (auto simp: inverse_eq_divide field_simps)
+ then have "r * 2^max N m < p * 2^max N m - 1" by simp
+ moreover
+ have "real (natfloor (p * 2 ^ max N m)) \<le> r * 2 ^ max N m"
+ using *[of "max N m"] m unfolding real_f using ux
+ by (cases "0 \<le> u x") (simp_all add: max_def mult_nonneg_nonneg split: split_if_asm)
+ then have "p * 2 ^ max N m - 1 < r * 2 ^ max N m"
+ by (metis real_natfloor_gt_diff_one less_le_trans)
+ ultimately show False by auto
qed
- qed
- thus "?g i t \<le> ?g (Suc i) t"
- by (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps) }
- hence mono: "mono (\<lambda>i. ?g i t)" unfolding mono_iff_le_Suc by auto
+ then show "max 0 (u x) \<le> y" using real ux by simp
+ qed (insert `0 \<le> y`, auto)
+ qed
+ qed (auto simp: divide_nonneg_pos)
+qed
- show "(SUP j. of_nat (f t j) / 2 ^ j) = u t"
- proof (rule pextreal_SUPI)
- fix j show "of_nat (f t j) / 2 ^ j \<le> u t"
- proof (rule fI)
- assume "of_nat j \<le> u t" thus "of_nat (j * 2 ^ j) / 2 ^ j \<le> u t"
- by (cases "u t") (auto simp: power_le_zero_eq divide_real_def[symmetric] field_simps)
- next
- fix k assume "of_nat k \<le> u t * 2 ^ j"
- thus "of_nat k / 2 ^ j \<le> u t"
- by (cases "u t")
- (auto simp: power_le_zero_eq divide_real_def[symmetric] field_simps zero_le_mult_iff)
- qed
- next
- fix y :: pextreal assume *: "\<And>j. j \<in> UNIV \<Longrightarrow> of_nat (f t j) / 2 ^ j \<le> y"
- show "u t \<le> y"
- proof (cases "u t")
- case (preal r)
- show ?thesis
- proof (rule ccontr)
- assume "\<not> u t \<le> y"
- then obtain p where p: "y = Real p" "0 \<le> p" "p < r" using preal by (cases y) auto
- with LIMSEQ_inverse_realpow_zero[of 2, unfolded LIMSEQ_iff, rule_format, of "r - p"]
- obtain n where n: "\<And>N. n \<le> N \<Longrightarrow> inverse (2^N) < r - p" by auto
- let ?N = "max n (natfloor r + 1)"
- have "u t < of_nat ?N" "n \<le> ?N"
- using ge_natfloor_plus_one_imp_gt[of r n] preal
- using real_natfloor_add_one_gt
- by (auto simp: max_def real_of_nat_Suc)
- from lower[OF this(1)]
- have "r < (real (f t ?N) + 1) / 2 ^ ?N" unfolding less_divide_eq
- using preal by (auto simp add: power_le_zero_eq zero_le_mult_iff real_of_nat_Suc split: split_if_asm)
- hence "u t < of_nat (f t ?N) / 2 ^ ?N + 1 / 2 ^ ?N"
- using preal by (auto simp: field_simps divide_real_def[symmetric])
- with n[OF `n \<le> ?N`] p preal *[of ?N]
- show False
- by (cases "f t ?N") (auto simp: power_le_zero_eq split: split_if_asm)
- qed
- next
- case infinite
- { fix j have "f t j = j*2^j" using top[of j t] infinite by simp
- hence "of_nat j \<le> y" using *[of j]
- by (cases y) (auto simp: power_le_zero_eq zero_le_mult_iff field_simps) }
- note all_less_y = this
- show ?thesis unfolding infinite
- proof (rule ccontr)
- assume "\<not> \<omega> \<le> y" then obtain r where r: "y = Real r" "0 \<le> r" by (cases y) auto
- moreover obtain n ::nat where "r < real n" using ex_less_of_nat by (auto simp: real_eq_of_nat)
- with all_less_y[of n] r show False by auto
- qed
- qed
- qed
+lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence':
+ fixes u :: "'a \<Rightarrow> extreal"
+ assumes u: "u \<in> borel_measurable M"
+ obtains f where "\<And>i. simple_function M (f i)" "incseq f" "\<And>i. \<infinity> \<notin> range (f i)"
+ "\<And>x. (SUP i. f i x) = max 0 (u x)" "\<And>i x. 0 \<le> f i x"
+ using borel_measurable_implies_simple_function_sequence[OF u] by auto
+
+lemma (in sigma_algebra) simple_function_If_set:
+ assumes sf: "simple_function M f" "simple_function M g" and A: "A \<inter> space M \<in> sets M"
+ shows "simple_function M (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function M ?IF")
+proof -
+ def F \<equiv> "\<lambda>x. f -` {x} \<inter> space M" and G \<equiv> "\<lambda>x. g -` {x} \<inter> space M"
+ show ?thesis unfolding simple_function_def
+ proof safe
+ have "?IF ` space M \<subseteq> f ` space M \<union> g ` space M" by auto
+ from finite_subset[OF this] assms
+ show "finite (?IF ` space M)" unfolding simple_function_def by auto
+ next
+ fix x assume "x \<in> space M"
+ then have *: "?IF -` {?IF x} \<inter> space M = (if x \<in> A
+ then ((F (f x) \<inter> (A \<inter> space M)) \<union> (G (f x) - (G (f x) \<inter> (A \<inter> space M))))
+ else ((F (g x) \<inter> (A \<inter> space M)) \<union> (G (g x) - (G (g x) \<inter> (A \<inter> space M)))))"
+ using sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def)
+ have [intro]: "\<And>x. F x \<in> sets M" "\<And>x. G x \<in> sets M"
+ unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto
+ show "?IF -` {?IF x} \<inter> space M \<in> sets M" unfolding * using A by auto
qed
qed
-lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence':
- fixes u :: "'a \<Rightarrow> pextreal"
- assumes "u \<in> borel_measurable M"
- obtains (x) f where "f \<up> u" "\<And>i. simple_function M (f i)" "\<And>i. \<omega>\<notin>f i`space M"
+lemma (in sigma_algebra) simple_function_If:
+ assumes sf: "simple_function M f" "simple_function M g" and P: "{x\<in>space M. P x} \<in> sets M"
+ shows "simple_function M (\<lambda>x. if P x then f x else g x)"
proof -
- from borel_measurable_implies_simple_function_sequence[OF assms]
- obtain f where x: "\<And>i. simple_function M (f i)" "f \<up> u"
- and fin: "\<And>i. \<And>x. x\<in>space M \<Longrightarrow> f i x \<noteq> \<omega>" by auto
- { fix i from fin[of _ i] have "\<omega> \<notin> f i`space M" by fastsimp }
- with x show thesis by (auto intro!: that[of f])
+ have "{x\<in>space M. P x} = {x. P x} \<inter> space M" by auto
+ with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp
qed
-lemma (in sigma_algebra) simple_function_eq_borel_measurable:
- fixes f :: "'a \<Rightarrow> pextreal"
- shows "simple_function M f \<longleftrightarrow>
- finite (f`space M) \<and> f \<in> borel_measurable M"
- using simple_function_borel_measurable[of f]
- borel_measurable_simple_function[of f]
- by (fastsimp simp: simple_function_def)
-
lemma (in measure_space) simple_function_restricted:
- fixes f :: "'a \<Rightarrow> pextreal" assumes "A \<in> sets M"
+ fixes f :: "'a \<Rightarrow> extreal" assumes "A \<in> sets M"
shows "simple_function (restricted_space A) f \<longleftrightarrow> simple_function M (\<lambda>x. f x * indicator A x)"
(is "simple_function ?R f \<longleftrightarrow> simple_function M ?f")
proof -
interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
- have "finite (f`A) \<longleftrightarrow> finite (?f`space M)"
+ have f: "finite (f`A) \<longleftrightarrow> finite (?f`space M)"
proof cases
assume "A = space M"
then have "f`A = ?f`space M" by (fastsimp simp: image_iff)
@@ -456,7 +475,7 @@
using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x])
next
fix x
- assume "indicator A x \<noteq> (0::pextreal)"
+ assume "indicator A x \<noteq> (0::extreal)"
then have "x \<in> A" by (auto simp: indicator_def split: split_if_asm)
moreover assume "x \<in> space M" "\<forall>y\<in>A. ?f x \<noteq> f y"
ultimately show "f x = 0" by auto
@@ -467,7 +486,8 @@
unfolding simple_function_eq_borel_measurable
R.simple_function_eq_borel_measurable
unfolding borel_measurable_restricted[OF `A \<in> sets M`]
- by auto
+ using assms(1)[THEN sets_into_space]
+ by (auto simp: indicator_def)
qed
lemma (in sigma_algebra) simple_function_subalgebra:
@@ -504,7 +524,7 @@
"integral\<^isup>S M f = (\<Sum>x \<in> f ` space M. x * measure M (f -` {x} \<inter> space M))"
syntax
- "_simple_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> pextreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> pextreal" ("\<integral>\<^isup>S _. _ \<partial>_" [60,61] 110)
+ "_simple_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> extreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> extreal" ("\<integral>\<^isup>S _. _ \<partial>_" [60,61] 110)
translations
"\<integral>\<^isup>S x. f \<partial>M" == "CONST integral\<^isup>S M (%x. f)"
@@ -540,7 +560,7 @@
qed
lemma (in measure_space) simple_function_partition:
- assumes "simple_function M f" and "simple_function M g"
+ assumes f: "simple_function M f" and g: "simple_function M g"
shows "integral\<^isup>S M f = (\<Sum>A\<in>(\<lambda>x. f -` {f x} \<inter> g -` {g x} \<inter> space M) ` space M. the_elem (f`A) * \<mu> A)"
(is "_ = setsum _ (?p ` space M)")
proof-
@@ -559,23 +579,16 @@
hence "finite (?p ` (A \<inter> space M))"
by (rule finite_subset) auto }
note this[intro, simp]
+ note sets = simple_function_measurable2[OF f g]
{ fix x assume "x \<in> space M"
have "\<Union>(?sub (f x)) = (f -` {f x} \<inter> space M)" by auto
- moreover {
- fix x y
- have *: "f -` {f x} \<inter> g -` {g x} \<inter> space M
- = (f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)" by auto
- assume "x \<in> space M" "y \<in> space M"
- hence "f -` {f x} \<inter> g -` {g x} \<inter> space M \<in> sets M"
- using assms unfolding simple_function_def * by auto }
- ultimately
- have "\<mu> (f -` {f x} \<inter> space M) = setsum (\<mu>) (?sub (f x))"
- by (subst measure_finitely_additive) auto }
+ with sets have "\<mu> (f -` {f x} \<inter> space M) = setsum \<mu> (?sub (f x))"
+ by (subst measure_Union) auto }
hence "integral\<^isup>S M f = (\<Sum>(x,A)\<in>?SIGMA. x * \<mu> A)"
- unfolding simple_integral_def
- by (subst setsum_Sigma[symmetric],
- auto intro!: setsum_cong simp: setsum_right_distrib[symmetric])
+ unfolding simple_integral_def using f sets
+ by (subst setsum_Sigma[symmetric])
+ (auto intro!: setsum_cong setsum_extreal_right_distrib)
also have "\<dots> = (\<Sum>A\<in>?p ` space M. the_elem (f`A) * \<mu> A)"
proof -
have [simp]: "\<And>x. x \<in> space M \<Longrightarrow> f ` ?p x = {f x}" by (auto intro!: imageI)
@@ -595,7 +608,7 @@
qed
lemma (in measure_space) simple_integral_add[simp]:
- assumes "simple_function M f" and "simple_function M g"
+ assumes f: "simple_function M f" and "\<And>x. 0 \<le> f x" and g: "simple_function M g" and "\<And>x. 0 \<le> g x"
shows "(\<integral>\<^isup>Sx. f x + g x \<partial>M) = integral\<^isup>S M f + integral\<^isup>S M g"
proof -
{ fix x let ?S = "g -` {g x} \<inter> f -` {f x} \<inter> space M"
@@ -603,63 +616,43 @@
hence "(\<lambda>a. f a + g a) ` ?S = {f x + g x}" "f ` ?S = {f x}" "g ` ?S = {g x}"
"(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M = ?S"
by auto }
- thus ?thesis
+ with assms show ?thesis
unfolding
- simple_function_partition[OF simple_function_add[OF assms] simple_function_Pair[OF assms]]
- simple_function_partition[OF `simple_function M f` `simple_function M g`]
- simple_function_partition[OF `simple_function M g` `simple_function M f`]
- apply (subst (3) Int_commute)
- by (auto simp add: field_simps setsum_addf[symmetric] intro!: setsum_cong)
+ simple_function_partition[OF simple_function_add[OF f g] simple_function_Pair[OF f g]]
+ simple_function_partition[OF f g]
+ simple_function_partition[OF g f]
+ by (subst (3) Int_commute)
+ (auto simp add: extreal_left_distrib setsum_addf[symmetric] intro!: setsum_cong)
qed
lemma (in measure_space) simple_integral_setsum[simp]:
+ assumes "\<And>i x. i \<in> P \<Longrightarrow> 0 \<le> f i x"
assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
shows "(\<integral>\<^isup>Sx. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^isup>S M (f i))"
proof cases
assume "finite P"
from this assms show ?thesis
- by induct (auto simp: simple_function_setsum simple_integral_add)
+ by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg)
qed auto
lemma (in measure_space) simple_integral_mult[simp]:
- assumes "simple_function M f"
+ assumes f: "simple_function M f" "\<And>x. 0 \<le> f x"
shows "(\<integral>\<^isup>Sx. c * f x \<partial>M) = c * integral\<^isup>S M f"
proof -
- note mult = simple_function_mult[OF simple_function_const[of c] assms]
+ note mult = simple_function_mult[OF simple_function_const[of c] f(1)]
{ fix x let ?S = "f -` {f x} \<inter> (\<lambda>x. c * f x) -` {c * f x} \<inter> space M"
assume "x \<in> space M"
hence "(\<lambda>x. c * f x) ` ?S = {c * f x}" "f ` ?S = {f x}"
by auto }
- thus ?thesis
- unfolding simple_function_partition[OF mult assms]
- simple_function_partition[OF assms mult]
- by (auto simp: setsum_right_distrib field_simps intro!: setsum_cong)
-qed
-
-lemma (in sigma_algebra) simple_function_If:
- assumes sf: "simple_function M f" "simple_function M g" and A: "A \<in> sets M"
- shows "simple_function M (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function M ?IF")
-proof -
- def F \<equiv> "\<lambda>x. f -` {x} \<inter> space M" and G \<equiv> "\<lambda>x. g -` {x} \<inter> space M"
- show ?thesis unfolding simple_function_def
- proof safe
- have "?IF ` space M \<subseteq> f ` space M \<union> g ` space M" by auto
- from finite_subset[OF this] assms
- show "finite (?IF ` space M)" unfolding simple_function_def by auto
- next
- fix x assume "x \<in> space M"
- then have *: "?IF -` {?IF x} \<inter> space M = (if x \<in> A
- then ((F (f x) \<inter> A) \<union> (G (f x) - (G (f x) \<inter> A)))
- else ((F (g x) \<inter> A) \<union> (G (g x) - (G (g x) \<inter> A))))"
- using sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def)
- have [intro]: "\<And>x. F x \<in> sets M" "\<And>x. G x \<in> sets M"
- unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto
- show "?IF -` {?IF x} \<inter> space M \<in> sets M" unfolding * using A by auto
- qed
+ with assms show ?thesis
+ unfolding simple_function_partition[OF mult f(1)]
+ simple_function_partition[OF f(1) mult]
+ by (subst setsum_extreal_right_distrib)
+ (auto intro!: extreal_0_le_mult setsum_cong simp: mult_assoc)
qed
lemma (in measure_space) simple_integral_mono_AE:
- assumes "simple_function M f" and "simple_function M g"
+ assumes f: "simple_function M f" and g: "simple_function M g"
and mono: "AE x. f x \<le> g x"
shows "integral\<^isup>S M f \<le> integral\<^isup>S M g"
proof -
@@ -668,14 +661,16 @@
"\<And>x. f -` {f x} \<inter> g -` {g x} \<inter> space M = ?S x" by auto
show ?thesis
unfolding *
- simple_function_partition[OF `simple_function M f` `simple_function M g`]
- simple_function_partition[OF `simple_function M g` `simple_function M f`]
+ simple_function_partition[OF f g]
+ simple_function_partition[OF g f]
proof (safe intro!: setsum_mono)
fix x assume "x \<in> space M"
then have *: "f ` ?S x = {f x}" "g ` ?S x = {g x}" by auto
show "the_elem (f`?S x) * \<mu> (?S x) \<le> the_elem (g`?S x) * \<mu> (?S x)"
proof (cases "f x \<le> g x")
- case True then show ?thesis using * by (auto intro!: mult_right_mono)
+ case True then show ?thesis
+ using * assms(1,2)[THEN simple_functionD(2)]
+ by (auto intro!: extreal_mult_right_mono)
next
case False
obtain N where N: "{x\<in>space M. \<not> f x \<le> g x} \<subseteq> N" "N \<in> sets M" "\<mu> N = 0"
@@ -685,7 +680,10 @@
by (rule_tac Int) (auto intro!: simple_functionD)
ultimately have "\<mu> (?S x) \<le> \<mu> N"
using `N \<in> sets M` by (auto intro!: measure_mono)
- then show ?thesis using `\<mu> N = 0` by auto
+ moreover have "0 \<le> \<mu> (?S x)"
+ using assms(1,2)[THEN simple_functionD(2)] by auto
+ ultimately have "\<mu> (?S x) = 0" using `\<mu> N = 0` by auto
+ then show ?thesis by simp
qed
qed
qed
@@ -697,7 +695,8 @@
using assms by (intro simple_integral_mono_AE) auto
lemma (in measure_space) simple_integral_cong_AE:
- assumes "simple_function M f" "simple_function M g" and "AE x. f x = g x"
+ assumes "simple_function M f" and "simple_function M g"
+ and "AE x. f x = g x"
shows "integral\<^isup>S M f = integral\<^isup>S M g"
using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE)
@@ -765,7 +764,7 @@
assume "space M = {}" hence "A = {}" using sets_into_space[OF assms] by auto
thus ?thesis unfolding simple_integral_def using `space M = {}` by auto
next
- assume "space M \<noteq> {}" hence "(\<lambda>x. 1) ` space M = {1::pextreal}" by auto
+ assume "space M \<noteq> {}" hence "(\<lambda>x. 1) ` space M = {1::extreal}" by auto
thus ?thesis
using simple_integral_indicator[OF assms simple_function_const[of 1]]
using sets_into_space[OF assms]
@@ -773,13 +772,13 @@
qed
lemma (in measure_space) simple_integral_null_set:
- assumes "simple_function M u" "N \<in> null_sets"
+ assumes "simple_function M u" "\<And>x. 0 \<le> u x" and "N \<in> null_sets"
shows "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = 0"
proof -
- have "AE x. indicator N x = (0 :: pextreal)"
+ have "AE x. indicator N x = (0 :: extreal)"
using `N \<in> null_sets` by (auto simp: indicator_def intro!: AE_I[of _ N])
then have "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^isup>Sx. 0 \<partial>M)"
- using assms by (intro simple_integral_cong_AE) auto
+ using assms apply (intro simple_integral_cong_AE) by auto
then show ?thesis by simp
qed
@@ -813,7 +812,7 @@
by (auto simp: indicator_def split: split_if_asm)
then show "f x * \<mu> (f -` {f x} \<inter> A) =
f x * \<mu> (?f -` {f x} \<inter> space M)"
- unfolding pextreal_mult_cancel_left by auto
+ unfolding extreal_mult_cancel_left by auto
qed
lemma (in measure_space) simple_integral_subalgebra:
@@ -821,10 +820,6 @@
shows "integral\<^isup>S N = integral\<^isup>S M"
unfolding simple_integral_def_raw by simp
-lemma measure_preservingD:
- "T \<in> measure_preserving A B \<Longrightarrow> X \<in> sets B \<Longrightarrow> measure A (T -` X \<inter> space A) = measure B X"
- unfolding measure_preserving_def by auto
-
lemma (in measure_space) simple_integral_vimage:
assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'"
and f: "simple_function M' f"
@@ -853,196 +848,164 @@
qed
qed
+lemma (in measure_space) simple_integral_cmult_indicator:
+ assumes A: "A \<in> sets M"
+ shows "(\<integral>\<^isup>Sx. c * indicator A x \<partial>M) = c * \<mu> A"
+ using simple_integral_mult[OF simple_function_indicator[OF A]]
+ unfolding simple_integral_indicator_only[OF A] by simp
+
+lemma (in measure_space) simple_integral_positive:
+ assumes f: "simple_function M f" and ae: "AE x. 0 \<le> f x"
+ shows "0 \<le> integral\<^isup>S M f"
+proof -
+ have "integral\<^isup>S M (\<lambda>x. 0) \<le> integral\<^isup>S M f"
+ using simple_integral_mono_AE[OF _ f ae] by auto
+ then show ?thesis by simp
+qed
+
section "Continuous positive integration"
definition positive_integral_def:
- "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> f}. integral\<^isup>S M g)"
+ "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f}. integral\<^isup>S M g)"
syntax
- "_positive_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> pextreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> pextreal" ("\<integral>\<^isup>+ _. _ \<partial>_" [60,61] 110)
+ "_positive_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> extreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> extreal" ("\<integral>\<^isup>+ _. _ \<partial>_" [60,61] 110)
translations
"\<integral>\<^isup>+ x. f \<partial>M" == "CONST integral\<^isup>P M (%x. f)"
-lemma (in measure_space) positive_integral_alt: "integral\<^isup>P M f =
- (SUP g : {g. simple_function M g \<and> g \<le> f \<and> \<omega> \<notin> g`space M}. integral\<^isup>S M g)"
- (is "_ = ?alt")
-proof (rule antisym SUP_leI)
- show "integral\<^isup>P M f \<le> ?alt" unfolding positive_integral_def
- proof (safe intro!: SUP_leI)
- fix g assume g: "simple_function M g" "g \<le> f"
- let ?G = "g -` {\<omega>} \<inter> space M"
- show "integral\<^isup>S M g \<le>
- (SUP h : {i. simple_function M i \<and> i \<le> f \<and> \<omega> \<notin> i ` space M}. integral\<^isup>S M h)"
- (is "integral\<^isup>S M g \<le> SUPR ?A _")
- proof cases
- let ?g = "\<lambda>x. indicator (space M - ?G) x * g x"
- have g': "simple_function M ?g"
- using g by (auto intro: simple_functionD)
- moreover
- assume "\<mu> ?G = 0"
- then have "AE x. g x = ?g x" using g
- by (intro AE_I[where N="?G"])
- (auto intro: simple_functionD simp: indicator_def)
- with g(1) g' have "integral\<^isup>S M g = integral\<^isup>S M ?g"
- by (rule simple_integral_cong_AE)
- moreover have "?g \<le> g" by (auto simp: le_fun_def indicator_def)
- from this `g \<le> f` have "?g \<le> f" by (rule order_trans)
- moreover have "\<omega> \<notin> ?g ` space M"
- by (auto simp: indicator_def split: split_if_asm)
- ultimately show ?thesis by (auto intro!: le_SUPI)
- next
- assume "\<mu> ?G \<noteq> 0"
- then have "?G \<noteq> {}" by auto
- then have "\<omega> \<in> g`space M" by force
- then have "space M \<noteq> {}" by auto
- have "SUPR ?A (integral\<^isup>S M) = \<omega>"
- proof (intro SUP_\<omega>[THEN iffD2] allI impI)
- fix x assume "x < \<omega>"
- then guess n unfolding less_\<omega>_Ex_of_nat .. note n = this
- then have "0 < n" by (intro neq0_conv[THEN iffD1] notI) simp
- let ?g = "\<lambda>x. (of_nat n / (if \<mu> ?G = \<omega> then 1 else \<mu> ?G)) * indicator ?G x"
- show "\<exists>i\<in>?A. x < integral\<^isup>S M i"
- proof (intro bexI impI CollectI conjI)
- show "simple_function M ?g" using g
- by (auto intro!: simple_functionD simple_function_add)
- have "?g \<le> g" by (auto simp: le_fun_def indicator_def)
- from this g(2) show "?g \<le> f" by (rule order_trans)
- show "\<omega> \<notin> ?g ` space M"
- using `\<mu> ?G \<noteq> 0` by (auto simp: indicator_def split: split_if_asm)
- have "x < (of_nat n / (if \<mu> ?G = \<omega> then 1 else \<mu> ?G)) * \<mu> ?G"
- using n `\<mu> ?G \<noteq> 0` `0 < n`
- by (auto simp: pextreal_noteq_omega_Ex field_simps)
- also have "\<dots> = integral\<^isup>S M ?g" using g `space M \<noteq> {}`
- by (subst simple_integral_indicator)
- (auto simp: image_constant ac_simps dest: simple_functionD)
- finally show "x < integral\<^isup>S M ?g" .
- qed
- qed
- then show ?thesis by simp
- qed
- qed
-qed (auto intro!: SUP_subset simp: positive_integral_def)
-
lemma (in measure_space) positive_integral_cong_measure:
assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
shows "integral\<^isup>P N f = integral\<^isup>P M f"
-proof -
- interpret v: measure_space N
- by (rule measure_space_cong) fact+
- with assms show ?thesis
- unfolding positive_integral_def SUPR_def
- by (auto intro!: arg_cong[where f=Sup] image_cong
- simp: simple_integral_cong_measure[OF assms]
- simple_function_cong_algebra[OF assms(2,3)])
-qed
+ unfolding positive_integral_def
+ unfolding simple_function_cong_algebra[OF assms(2,3), symmetric]
+ using AE_cong_measure[OF assms]
+ using simple_integral_cong_measure[OF assms]
+ by (auto intro!: SUP_cong)
+
+lemma (in measure_space) positive_integral_positive:
+ "0 \<le> integral\<^isup>P M f"
+ by (auto intro!: le_SUPI2[of "\<lambda>x. 0"] simp: positive_integral_def le_fun_def)
-lemma (in measure_space) positive_integral_alt1:
- "integral\<^isup>P M f =
- (SUP g : {g. simple_function M g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}. integral\<^isup>S M g)"
- unfolding positive_integral_alt SUPR_def
-proof (safe intro!: arg_cong[where f=Sup])
- fix g let ?g = "\<lambda>x. if x \<in> space M then g x else f x"
- assume "simple_function M g" "\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>"
- hence "?g \<le> f" "simple_function M ?g" "integral\<^isup>S M g = integral\<^isup>S M ?g"
- "\<omega> \<notin> g`space M"
- unfolding le_fun_def by (auto cong: simple_function_cong simple_integral_cong)
- thus "integral\<^isup>S M g \<in> integral\<^isup>S M ` {g. simple_function M g \<and> g \<le> f \<and> \<omega> \<notin> g`space M}"
- by auto
-next
- fix g assume "simple_function M g" "g \<le> f" "\<omega> \<notin> g`space M"
- hence "simple_function M g" "\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>"
- by (auto simp add: le_fun_def image_iff)
- thus "integral\<^isup>S M g \<in> integral\<^isup>S M ` {g. simple_function M g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}"
- by auto
-qed
+lemma (in measure_space) positive_integral_def_finite:
+ "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f \<and> range g \<subseteq> {0 ..< \<infinity>}}. integral\<^isup>S M g)"
+ (is "_ = SUPR ?A ?f")
+ unfolding positive_integral_def
+proof (safe intro!: antisym SUP_leI)
+ fix g assume g: "simple_function M g" "g \<le> max 0 \<circ> f"
+ let ?G = "{x \<in> space M. \<not> g x \<noteq> \<infinity>}"
+ note gM = g(1)[THEN borel_measurable_simple_function]
+ have \<mu>G_pos: "0 \<le> \<mu> ?G" using gM by auto
+ let "?g y x" = "if g x = \<infinity> then y else max 0 (g x)"
+ from g gM have g_in_A: "\<And>y. 0 \<le> y \<Longrightarrow> y \<noteq> \<infinity> \<Longrightarrow> ?g y \<in> ?A"
+ apply (safe intro!: simple_function_max simple_function_If)
+ apply (force simp: max_def le_fun_def split: split_if_asm)+
+ done
+ show "integral\<^isup>S M g \<le> SUPR ?A ?f"
+ proof cases
+ have g0: "?g 0 \<in> ?A" by (intro g_in_A) auto
+ assume "\<mu> ?G = 0"
+ with gM have "AE x. x \<notin> ?G" by (simp add: AE_iff_null_set)
+ with gM g show ?thesis
+ by (intro le_SUPI2[OF g0] simple_integral_mono_AE)
+ (auto simp: max_def intro!: simple_function_If)
+ next
+ assume \<mu>G: "\<mu> ?G \<noteq> 0"
+ have "SUPR ?A (integral\<^isup>S M) = \<infinity>"
+ proof (intro SUP_PInfty)
+ fix n :: nat
+ let ?y = "extreal (real n) / (if \<mu> ?G = \<infinity> then 1 else \<mu> ?G)"
+ have "0 \<le> ?y" "?y \<noteq> \<infinity>" using \<mu>G \<mu>G_pos by (auto simp: extreal_divide_eq)
+ then have "?g ?y \<in> ?A" by (rule g_in_A)
+ have "real n \<le> ?y * \<mu> ?G"
+ using \<mu>G \<mu>G_pos by (cases "\<mu> ?G") (auto simp: field_simps)
+ also have "\<dots> = (\<integral>\<^isup>Sx. ?y * indicator ?G x \<partial>M)"
+ using `0 \<le> ?y` `?g ?y \<in> ?A` gM
+ by (subst simple_integral_cmult_indicator) auto
+ also have "\<dots> \<le> integral\<^isup>S M (?g ?y)" using `?g ?y \<in> ?A` gM
+ by (intro simple_integral_mono) auto
+ finally show "\<exists>i\<in>?A. real n \<le> integral\<^isup>S M i"
+ using `?g ?y \<in> ?A` by blast
+ qed
+ then show ?thesis by simp
+ qed
+qed (auto intro: le_SUPI)
-lemma (in measure_space) positive_integral_cong:
- assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
- shows "integral\<^isup>P M f = integral\<^isup>P M g"
-proof -
- have "\<And>h. (\<forall>x\<in>space M. h x \<le> f x \<and> h x \<noteq> \<omega>) = (\<forall>x\<in>space M. h x \<le> g x \<and> h x \<noteq> \<omega>)"
- using assms by auto
- thus ?thesis unfolding positive_integral_alt1 by auto
+lemma (in measure_space) positive_integral_mono_AE:
+ assumes ae: "AE x. u x \<le> v x" shows "integral\<^isup>P M u \<le> integral\<^isup>P M v"
+ unfolding positive_integral_def
+proof (safe intro!: SUP_mono)
+ fix n assume n: "simple_function M n" "n \<le> max 0 \<circ> u"
+ from ae[THEN AE_E] guess N . note N = this
+ then have ae_N: "AE x. x \<notin> N" by (auto intro: AE_not_in)
+ let "?n x" = "n x * indicator (space M - N) x"
+ have "AE x. n x \<le> ?n x" "simple_function M ?n"
+ using n N ae_N by auto
+ moreover
+ { fix x have "?n x \<le> max 0 (v x)"
+ proof cases
+ assume x: "x \<in> space M - N"
+ with N have "u x \<le> v x" by auto
+ with n(2)[THEN le_funD, of x] x show ?thesis
+ by (auto simp: max_def split: split_if_asm)
+ qed simp }
+ then have "?n \<le> max 0 \<circ> v" by (auto simp: le_funI)
+ moreover have "integral\<^isup>S M n \<le> integral\<^isup>S M ?n"
+ using ae_N N n by (auto intro!: simple_integral_mono_AE)
+ ultimately show "\<exists>m\<in>{g. simple_function M g \<and> g \<le> max 0 \<circ> v}. integral\<^isup>S M n \<le> integral\<^isup>S M m"
+ by force
qed
-lemma (in measure_space) positive_integral_eq_simple_integral:
- assumes "simple_function M f"
- shows "integral\<^isup>P M f = integral\<^isup>S M f"
- unfolding positive_integral_def
-proof (safe intro!: pextreal_SUPI)
- fix g assume "simple_function M g" "g \<le> f"
- with assms show "integral\<^isup>S M g \<le> integral\<^isup>S M f"
- by (auto intro!: simple_integral_mono simp: le_fun_def)
-next
- fix y assume "\<forall>x. x\<in>{g. simple_function M g \<and> g \<le> f} \<longrightarrow> integral\<^isup>S M x \<le> y"
- with assms show "integral\<^isup>S M f \<le> y" by auto
-qed
-
-lemma (in measure_space) positive_integral_mono_AE:
- assumes ae: "AE x. u x \<le> v x"
- shows "integral\<^isup>P M u \<le> integral\<^isup>P M v"
- unfolding positive_integral_alt1
-proof (safe intro!: SUPR_mono)
- fix a assume a: "simple_function M a" and mono: "\<forall>x\<in>space M. a x \<le> u x \<and> a x \<noteq> \<omega>"
- from ae obtain N where N: "{x\<in>space M. \<not> u x \<le> v x} \<subseteq> N" "N \<in> sets M" "\<mu> N = 0"
- by (auto elim!: AE_E)
- have "simple_function M (\<lambda>x. a x * indicator (space M - N) x)"
- using `N \<in> sets M` a by auto
- with a show "\<exists>b\<in>{g. simple_function M g \<and> (\<forall>x\<in>space M. g x \<le> v x \<and> g x \<noteq> \<omega>)}.
- integral\<^isup>S M a \<le> integral\<^isup>S M b"
- proof (safe intro!: bexI[of _ "\<lambda>x. a x * indicator (space M - N) x"]
- simple_integral_mono_AE)
- show "AE x. a x \<le> a x * indicator (space M - N) x"
- proof (rule AE_I, rule subset_refl)
- have *: "{x \<in> space M. \<not> a x \<le> a x * indicator (space M - N) x} =
- N \<inter> {x \<in> space M. a x \<noteq> 0}" (is "?N = _")
- using `N \<in> sets M`[THEN sets_into_space] by (auto simp: indicator_def)
- then show "?N \<in> sets M"
- using `N \<in> sets M` `simple_function M a`[THEN borel_measurable_simple_function]
- by (auto intro!: measure_mono Int)
- then have "\<mu> ?N \<le> \<mu> N"
- unfolding * using `N \<in> sets M` by (auto intro!: measure_mono)
- then show "\<mu> ?N = 0" using `\<mu> N = 0` by auto
- qed
- next
- fix x assume "x \<in> space M"
- show "a x * indicator (space M - N) x \<le> v x"
- proof (cases "x \<in> N")
- case True then show ?thesis by simp
- next
- case False
- with N mono have "a x \<le> u x" "u x \<le> v x" using `x \<in> space M` by auto
- with False `x \<in> space M` show "a x * indicator (space M - N) x \<le> v x" by auto
- qed
- assume "a x * indicator (space M - N) x = \<omega>"
- with mono `x \<in> space M` show False
- by (simp split: split_if_asm add: indicator_def)
- qed
-qed
+lemma (in measure_space) positive_integral_mono:
+ "(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^isup>P M u \<le> integral\<^isup>P M v"
+ by (auto intro: positive_integral_mono_AE)
lemma (in measure_space) positive_integral_cong_AE:
"AE x. u x = v x \<Longrightarrow> integral\<^isup>P M u = integral\<^isup>P M v"
by (auto simp: eq_iff intro!: positive_integral_mono_AE)
-lemma (in measure_space) positive_integral_mono:
- "(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^isup>P M u \<le> integral\<^isup>P M v"
- by (auto intro: positive_integral_mono_AE)
+lemma (in measure_space) positive_integral_cong:
+ "(\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^isup>P M u = integral\<^isup>P M v"
+ by (auto intro: positive_integral_cong_AE)
-lemma image_set_cong:
- assumes A: "\<And>x. x \<in> A \<Longrightarrow> \<exists>y\<in>B. f x = g y"
- assumes B: "\<And>y. y \<in> B \<Longrightarrow> \<exists>x\<in>A. g y = f x"
- shows "f ` A = g ` B"
- using assms by blast
+lemma (in measure_space) positive_integral_eq_simple_integral:
+ assumes f: "simple_function M f" "\<And>x. 0 \<le> f x" shows "integral\<^isup>P M f = integral\<^isup>S M f"
+proof -
+ let "?f x" = "f x * indicator (space M) x"
+ have f': "simple_function M ?f" using f by auto
+ with f(2) have [simp]: "max 0 \<circ> ?f = ?f"
+ by (auto simp: fun_eq_iff max_def split: split_indicator)
+ have "integral\<^isup>P M ?f \<le> integral\<^isup>S M ?f" using f'
+ by (force intro!: SUP_leI simple_integral_mono simp: le_fun_def positive_integral_def)
+ moreover have "integral\<^isup>S M ?f \<le> integral\<^isup>P M ?f"
+ unfolding positive_integral_def
+ using f' by (auto intro!: le_SUPI)
+ ultimately show ?thesis
+ by (simp cong: positive_integral_cong simple_integral_cong)
+qed
+
+lemma (in measure_space) positive_integral_eq_simple_integral_AE:
+ assumes f: "simple_function M f" "AE x. 0 \<le> f x" shows "integral\<^isup>P M f = integral\<^isup>S M f"
+proof -
+ have "AE x. f x = max 0 (f x)" using f by (auto split: split_max)
+ with f have "integral\<^isup>P M f = integral\<^isup>S M (\<lambda>x. max 0 (f x))"
+ by (simp cong: positive_integral_cong_AE simple_integral_cong_AE
+ add: positive_integral_eq_simple_integral)
+ with assms show ?thesis
+ by (auto intro!: simple_integral_cong_AE split: split_max)
+qed
lemma (in measure_space) positive_integral_SUP_approx:
- assumes "f \<up> s"
- and f: "\<And>i. f i \<in> borel_measurable M"
- and "simple_function M u"
- and le: "u \<le> s" and real: "\<omega> \<notin> u`space M"
+ assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
+ and u: "simple_function M u" "u \<le> (SUP i. f i)" "u`space M \<subseteq> {0..<\<infinity>}"
shows "integral\<^isup>S M u \<le> (SUP i. integral\<^isup>P M (f i))" (is "_ \<le> ?S")
-proof (rule pextreal_le_mult_one_interval)
- fix a :: pextreal assume "0 < a" "a < 1"
+proof (rule extreal_le_mult_one_interval)
+ have "0 \<le> (SUP i. integral\<^isup>P M (f i))"
+ using f(3) by (auto intro!: le_SUPI2 positive_integral_positive)
+ then show "(SUP i. integral\<^isup>P M (f i)) \<noteq> -\<infinity>" by auto
+ have u_range: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> u x \<and> u x \<noteq> \<infinity>"
+ using u(3) by auto
+ fix a :: extreal assume "0 < a" "a < 1"
hence "a \<noteq> 0" by auto
let "?B i" = "{x \<in> space M. a * u x \<le> f i x}"
have B: "\<And>i. ?B i \<in> sets M"
@@ -1054,203 +1017,269 @@
proof safe
fix i x assume "a * u x \<le> f i x"
also have "\<dots> \<le> f (Suc i) x"
- using `f \<up> s` unfolding isoton_def le_fun_def by auto
+ using `incseq f`[THEN incseq_SucD] unfolding le_fun_def by auto
finally show "a * u x \<le> f (Suc i) x" .
qed }
note B_mono = this
- have u: "\<And>x. x \<in> space M \<Longrightarrow> u -` {u x} \<inter> space M \<in> sets M"
- using `simple_function M u` by (auto simp add: simple_function_def)
+ note B_u = Int[OF u(1)[THEN simple_functionD(2)] B]
- have "\<And>i. (\<lambda>n. (u -` {i} \<inter> space M) \<inter> ?B n) \<up> (u -` {i} \<inter> space M)" using B_mono unfolding isoton_def
- proof safe
- fix x i assume "x \<in> space M"
- show "x \<in> (\<Union>i. (u -` {u x} \<inter> space M) \<inter> ?B i)"
- proof cases
- assume "u x = 0" thus ?thesis using `x \<in> space M` by simp
- next
- assume "u x \<noteq> 0"
- with `a < 1` real `x \<in> space M`
- have "a * u x < 1 * u x" by (rule_tac pextreal_mult_strict_right_mono) (auto simp: image_iff)
- also have "\<dots> \<le> (SUP i. f i x)" using le `f \<up> s`
- unfolding isoton_fun_expand by (auto simp: isoton_def le_fun_def)
- finally obtain i where "a * u x < f i x" unfolding SUPR_def
- by (auto simp add: less_Sup_iff)
- hence "a * u x \<le> f i x" by auto
- thus ?thesis using `x \<in> space M` by auto
+ let "?B' i n" = "(u -` {i} \<inter> space M) \<inter> ?B n"
+ have measure_conv: "\<And>i. \<mu> (u -` {i} \<inter> space M) = (SUP n. \<mu> (?B' i n))"
+ proof -
+ fix i
+ have 1: "range (?B' i) \<subseteq> sets M" using B_u by auto
+ have 2: "incseq (?B' i)" using B_mono by (auto intro!: incseq_SucI)
+ have "(\<Union>n. ?B' i n) = u -` {i} \<inter> space M"
+ proof safe
+ fix x i assume x: "x \<in> space M"
+ show "x \<in> (\<Union>i. ?B' (u x) i)"
+ proof cases
+ assume "u x = 0" thus ?thesis using `x \<in> space M` f(3) by simp
+ next
+ assume "u x \<noteq> 0"
+ with `a < 1` u_range[OF `x \<in> space M`]
+ have "a * u x < 1 * u x"
+ by (intro extreal_mult_strict_right_mono) (auto simp: image_iff)
+ also have "\<dots> \<le> (SUP i. f i x)" using u(2) by (auto simp: le_fun_def SUPR_apply)
+ finally obtain i where "a * u x < f i x" unfolding SUPR_def
+ by (auto simp add: less_Sup_iff)
+ hence "a * u x \<le> f i x" by auto
+ thus ?thesis using `x \<in> space M` by auto
+ qed
qed
- qed auto
- note measure_conv = measure_up[OF Int[OF u B] this]
+ then show "?thesis i" using continuity_from_below[OF 1 2] by simp
+ qed
have "integral\<^isup>S M u = (SUP i. integral\<^isup>S M (?uB i))"
unfolding simple_integral_indicator[OF B `simple_function M u`]
- proof (subst SUPR_pextreal_setsum, safe)
+ proof (subst SUPR_extreal_setsum, safe)
fix x n assume "x \<in> space M"
- have "\<mu> (u -` {u x} \<inter> space M \<inter> {x \<in> space M. a * u x \<le> f n x})
- \<le> \<mu> (u -` {u x} \<inter> space M \<inter> {x \<in> space M. a * u x \<le> f (Suc n) x})"
- using B_mono Int[OF u[OF `x \<in> space M`] B] by (auto intro!: measure_mono)
- thus "u x * \<mu> (u -` {u x} \<inter> space M \<inter> ?B n)
- \<le> u x * \<mu> (u -` {u x} \<inter> space M \<inter> ?B (Suc n))"
- by (auto intro: mult_left_mono)
+ with u_range show "incseq (\<lambda>i. u x * \<mu> (?B' (u x) i))" "\<And>i. 0 \<le> u x * \<mu> (?B' (u x) i)"
+ using B_mono B_u by (auto intro!: measure_mono extreal_mult_left_mono incseq_SucI simp: extreal_zero_le_0_iff)
next
- show "integral\<^isup>S M u =
- (\<Sum>i\<in>u ` space M. SUP n. i * \<mu> (u -` {i} \<inter> space M \<inter> ?B n))"
- using measure_conv unfolding simple_integral_def isoton_def
- by (auto intro!: setsum_cong simp: pextreal_SUP_cmult)
+ show "integral\<^isup>S M u = (\<Sum>i\<in>u ` space M. SUP n. i * \<mu> (?B' i n))"
+ using measure_conv u_range B_u unfolding simple_integral_def
+ by (auto intro!: setsum_cong SUPR_extreal_cmult[symmetric])
qed
moreover
have "a * (SUP i. integral\<^isup>S M (?uB i)) \<le> ?S"
- unfolding pextreal_SUP_cmult[symmetric]
+ apply (subst SUPR_extreal_cmult[symmetric])
proof (safe intro!: SUP_mono bexI)
fix i
have "a * integral\<^isup>S M (?uB i) = (\<integral>\<^isup>Sx. a * ?uB i x \<partial>M)"
- using B `simple_function M u`
- by (subst simple_integral_mult[symmetric]) (auto simp: field_simps)
+ using B `simple_function M u` u_range
+ by (subst simple_integral_mult) (auto split: split_indicator)
also have "\<dots> \<le> integral\<^isup>P M (f i)"
proof -
- have "\<And>x. a * ?uB i x \<le> f i x" unfolding indicator_def by auto
- hence *: "simple_function M (\<lambda>x. a * ?uB i x)" using B assms(3)
- by (auto intro!: simple_integral_mono)
- show ?thesis unfolding positive_integral_eq_simple_integral[OF *, symmetric]
- by (auto intro!: positive_integral_mono simp: indicator_def)
+ have *: "simple_function M (\<lambda>x. a * ?uB i x)" using B `0 < a` u(1) by auto
+ show ?thesis using f(3) * u_range `0 < a`
+ by (subst positive_integral_eq_simple_integral[symmetric])
+ (auto intro!: positive_integral_mono split: split_indicator)
qed
finally show "a * integral\<^isup>S M (?uB i) \<le> integral\<^isup>P M (f i)"
by auto
- qed simp
+ next
+ fix i show "0 \<le> \<integral>\<^isup>S x. ?uB i x \<partial>M" using B `0 < a` u(1) u_range
+ by (intro simple_integral_positive) (auto split: split_indicator)
+ qed (insert `0 < a`, auto)
ultimately show "a * integral\<^isup>S M u \<le> ?S" by simp
qed
+lemma (in measure_space) incseq_positive_integral:
+ assumes "incseq f" shows "incseq (\<lambda>i. integral\<^isup>P M (f i))"
+proof -
+ have "\<And>i x. f i x \<le> f (Suc i) x"
+ using assms by (auto dest!: incseq_SucD simp: le_fun_def)
+ then show ?thesis
+ by (auto intro!: incseq_SucI positive_integral_mono)
+qed
+
text {* Beppo-Levi monotone convergence theorem *}
-lemma (in measure_space) positive_integral_isoton:
- assumes "f \<up> u" "\<And>i. f i \<in> borel_measurable M"
- shows "(\<lambda>i. integral\<^isup>P M (f i)) \<up> integral\<^isup>P M u"
- unfolding isoton_def
-proof safe
- fix i show "integral\<^isup>P M (f i) \<le> integral\<^isup>P M (f (Suc i))"
- apply (rule positive_integral_mono)
- using `f \<up> u` unfolding isoton_def le_fun_def by auto
+lemma (in measure_space) positive_integral_monotone_convergence_SUP:
+ assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
+ shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
+proof (rule antisym)
+ show "(SUP j. integral\<^isup>P M (f j)) \<le> (\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M)"
+ by (auto intro!: SUP_leI le_SUPI positive_integral_mono)
next
- have u: "u = (SUP i. f i)" using `f \<up> u` unfolding isoton_def by auto
- show "(SUP i. integral\<^isup>P M (f i)) = integral\<^isup>P M u"
- proof (rule antisym)
- from `f \<up> u`[THEN isoton_Sup, unfolded le_fun_def]
- show "(SUP j. integral\<^isup>P M (f j)) \<le> integral\<^isup>P M u"
- by (auto intro!: SUP_leI positive_integral_mono)
- next
- show "integral\<^isup>P M u \<le> (SUP i. integral\<^isup>P M (f i))"
- unfolding positive_integral_alt[of u]
- by (auto intro!: SUP_leI positive_integral_SUP_approx[OF assms])
+ show "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) \<le> (SUP j. integral\<^isup>P M (f j))"
+ unfolding positive_integral_def_finite[of "\<lambda>x. SUP i. f i x"]
+ proof (safe intro!: SUP_leI)
+ fix g assume g: "simple_function M g"
+ and "g \<le> max 0 \<circ> (\<lambda>x. SUP i. f i x)" "range g \<subseteq> {0..<\<infinity>}"
+ moreover then have "\<And>x. 0 \<le> (SUP i. f i x)" and g': "g`space M \<subseteq> {0..<\<infinity>}"
+ using f by (auto intro!: le_SUPI2)
+ ultimately show "integral\<^isup>S M g \<le> (SUP j. integral\<^isup>P M (f j))"
+ by (intro positive_integral_SUP_approx[OF f g _ g'])
+ (auto simp: le_fun_def max_def SUPR_apply)
qed
qed
-lemma (in measure_space) positive_integral_monotone_convergence_SUP:
- assumes "\<And>i x. x \<in> space M \<Longrightarrow> f i x \<le> f (Suc i) x"
- assumes "\<And>i. f i \<in> borel_measurable M"
- shows "(SUP i. integral\<^isup>P M (f i)) = (\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M)"
- (is "_ = integral\<^isup>P M ?u")
+lemma (in measure_space) positive_integral_monotone_convergence_SUP_AE:
+ assumes f: "\<And>i. AE x. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x" "\<And>i. f i \<in> borel_measurable M"
+ shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
proof -
- show ?thesis
- proof (rule antisym)
- show "(SUP j. integral\<^isup>P M (f j)) \<le> integral\<^isup>P M ?u"
- by (auto intro!: SUP_leI positive_integral_mono le_SUPI)
- next
- def rf \<equiv> "\<lambda>i. \<lambda>x\<in>space M. f i x" and ru \<equiv> "\<lambda>x\<in>space M. ?u x"
- have "\<And>i. rf i \<in> borel_measurable M" unfolding rf_def
- using assms by (simp cong: measurable_cong)
- moreover have iso: "rf \<up> ru" using assms unfolding rf_def ru_def
- unfolding isoton_def le_fun_def fun_eq_iff SUPR_apply
- using SUP_const[OF UNIV_not_empty]
- by (auto simp: restrict_def le_fun_def fun_eq_iff)
- ultimately have "integral\<^isup>P M ru \<le> (SUP i. integral\<^isup>P M (rf i))"
- unfolding positive_integral_alt[of ru]
- by (auto simp: le_fun_def intro!: SUP_leI positive_integral_SUP_approx)
- then show "integral\<^isup>P M ?u \<le> (SUP i. integral\<^isup>P M (f i))"
- unfolding ru_def rf_def by (simp cong: positive_integral_cong)
+ from f have "AE x. \<forall>i. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x"
+ by (simp add: AE_all_countable)
+ from this[THEN AE_E] guess N . note N = this
+ let "?f i x" = "if x \<in> space M - N then f i x else 0"
+ have f_eq: "AE x. \<forall>i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ N])
+ then have "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (\<integral>\<^isup>+ x. (SUP i. ?f i x) \<partial>M)"
+ by (auto intro!: positive_integral_cong_AE)
+ also have "\<dots> = (SUP i. (\<integral>\<^isup>+ x. ?f i x \<partial>M))"
+ proof (rule positive_integral_monotone_convergence_SUP)
+ show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI)
+ { fix i show "(\<lambda>x. if x \<in> space M - N then f i x else 0) \<in> borel_measurable M"
+ using f N(3) by (intro measurable_If_set) auto
+ fix x show "0 \<le> ?f i x"
+ using N(1) by auto }
qed
+ also have "\<dots> = (SUP i. (\<integral>\<^isup>+ x. f i x \<partial>M))"
+ using f_eq by (force intro!: arg_cong[where f="SUPR UNIV"] positive_integral_cong_AE ext)
+ finally show ?thesis .
+qed
+
+lemma (in measure_space) positive_integral_monotone_convergence_SUP_AE_incseq:
+ assumes f: "incseq f" "\<And>i. AE x. 0 \<le> f i x" and borel: "\<And>i. f i \<in> borel_measurable M"
+ shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
+ using f[unfolded incseq_Suc_iff le_fun_def]
+ by (intro positive_integral_monotone_convergence_SUP_AE[OF _ borel])
+ auto
+
+lemma (in measure_space) positive_integral_monotone_convergence_simple:
+ assumes f: "incseq f" "\<And>i x. 0 \<le> f i x" "\<And>i. simple_function M (f i)"
+ shows "(SUP i. integral\<^isup>S M (f i)) = (\<integral>\<^isup>+x. (SUP i. f i x) \<partial>M)"
+ using assms unfolding positive_integral_monotone_convergence_SUP[OF f(1)
+ f(3)[THEN borel_measurable_simple_function] f(2)]
+ by (auto intro!: positive_integral_eq_simple_integral[symmetric] arg_cong[where f="SUPR UNIV"] ext)
+
+lemma positive_integral_max_0:
+ "(\<integral>\<^isup>+x. max 0 (f x) \<partial>M) = integral\<^isup>P M f"
+ by (simp add: le_fun_def positive_integral_def)
+
+lemma (in measure_space) positive_integral_cong_pos:
+ assumes "\<And>x. x \<in> space M \<Longrightarrow> f x \<le> 0 \<and> g x \<le> 0 \<or> f x = g x"
+ shows "integral\<^isup>P M f = integral\<^isup>P M g"
+proof -
+ have "integral\<^isup>P M (\<lambda>x. max 0 (f x)) = integral\<^isup>P M (\<lambda>x. max 0 (g x))"
+ proof (intro positive_integral_cong)
+ fix x assume "x \<in> space M"
+ from assms[OF this] show "max 0 (f x) = max 0 (g x)"
+ by (auto split: split_max)
+ qed
+ then show ?thesis by (simp add: positive_integral_max_0)
qed
lemma (in measure_space) SUP_simple_integral_sequences:
- assumes f: "f \<up> u" "\<And>i. simple_function M (f i)"
- and g: "g \<up> u" "\<And>i. simple_function M (g i)"
+ assumes f: "incseq f" "\<And>i x. 0 \<le> f i x" "\<And>i. simple_function M (f i)"
+ and g: "incseq g" "\<And>i x. 0 \<le> g i x" "\<And>i. simple_function M (g i)"
+ and eq: "AE x. (SUP i. f i x) = (SUP i. g i x)"
shows "(SUP i. integral\<^isup>S M (f i)) = (SUP i. integral\<^isup>S M (g i))"
(is "SUPR _ ?F = SUPR _ ?G")
proof -
- have "(SUP i. ?F i) = (SUP i. integral\<^isup>P M (f i))"
- using assms by (simp add: positive_integral_eq_simple_integral)
- also have "\<dots> = integral\<^isup>P M u"
- using positive_integral_isoton[OF f(1) borel_measurable_simple_function[OF f(2)]]
- unfolding isoton_def by simp
- also have "\<dots> = (SUP i. integral\<^isup>P M (g i))"
- using positive_integral_isoton[OF g(1) borel_measurable_simple_function[OF g(2)]]
- unfolding isoton_def by simp
+ have "(SUP i. integral\<^isup>S M (f i)) = (\<integral>\<^isup>+x. (SUP i. f i x) \<partial>M)"
+ using f by (rule positive_integral_monotone_convergence_simple)
+ also have "\<dots> = (\<integral>\<^isup>+x. (SUP i. g i x) \<partial>M)"
+ unfolding eq[THEN positive_integral_cong_AE] ..
also have "\<dots> = (SUP i. ?G i)"
- using assms by (simp add: positive_integral_eq_simple_integral)
- finally show ?thesis .
+ using g by (rule positive_integral_monotone_convergence_simple[symmetric])
+ finally show ?thesis by simp
qed
lemma (in measure_space) positive_integral_const[simp]:
- "(\<integral>\<^isup>+ x. c \<partial>M) = c * \<mu> (space M)"
+ "0 \<le> c \<Longrightarrow> (\<integral>\<^isup>+ x. c \<partial>M) = c * \<mu> (space M)"
by (subst positive_integral_eq_simple_integral) auto
-lemma (in measure_space) positive_integral_isoton_simple:
- assumes "f \<up> u" and e: "\<And>i. simple_function M (f i)"
- shows "(\<lambda>i. integral\<^isup>S M (f i)) \<up> integral\<^isup>P M u"
- using positive_integral_isoton[OF `f \<up> u` e(1)[THEN borel_measurable_simple_function]]
- unfolding positive_integral_eq_simple_integral[OF e] .
-
-lemma measure_preservingD2:
- "f \<in> measure_preserving A B \<Longrightarrow> f \<in> measurable A B"
- unfolding measure_preserving_def by auto
-
lemma (in measure_space) positive_integral_vimage:
- assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'" and f: "f \<in> borel_measurable M'"
+ assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'"
+ and f: "f \<in> borel_measurable M'"
shows "integral\<^isup>P M' f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
proof -
interpret T: measure_space M' by (rule measure_space_vimage[OF T])
- obtain f' where f': "f' \<up> f" "\<And>i. simple_function M' (f' i)"
- using T.borel_measurable_implies_simple_function_sequence[OF f] by blast
- then have f: "(\<lambda>i x. f' i (T x)) \<up> (\<lambda>x. f (T x))" "\<And>i. simple_function M (\<lambda>x. f' i (T x))"
- using simple_function_vimage[OF T(1) measure_preservingD2[OF T(2)]] unfolding isoton_fun_expand by auto
+ from T.borel_measurable_implies_simple_function_sequence'[OF f]
+ guess f' . note f' = this
+ let "?f i x" = "f' i (T x)"
+ have inc: "incseq ?f" using f' by (force simp: le_fun_def incseq_def)
+ have sup: "\<And>x. (SUP i. ?f i x) = max 0 (f (T x))"
+ using f'(4) .
+ have sf: "\<And>i. simple_function M (\<lambda>x. f' i (T x))"
+ using simple_function_vimage[OF T(1) measure_preservingD2[OF T(2)] f'(1)] .
show "integral\<^isup>P M' f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
- using positive_integral_isoton_simple[OF f]
- using T.positive_integral_isoton_simple[OF f']
- by (simp add: simple_integral_vimage[OF T f'(2)] isoton_def)
+ using
+ T.positive_integral_monotone_convergence_simple[OF f'(2,5,1)]
+ positive_integral_monotone_convergence_simple[OF inc f'(5) sf]
+ by (simp add: positive_integral_max_0 simple_integral_vimage[OF T f'(1)] f')
qed
lemma (in measure_space) positive_integral_linear:
- assumes f: "f \<in> borel_measurable M"
- and g: "g \<in> borel_measurable M"
+ assumes f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" and "0 \<le> a"
+ and g: "g \<in> borel_measurable M" "\<And>x. 0 \<le> g x"
shows "(\<integral>\<^isup>+ x. a * f x + g x \<partial>M) = a * integral\<^isup>P M f + integral\<^isup>P M g"
(is "integral\<^isup>P M ?L = _")
proof -
- from borel_measurable_implies_simple_function_sequence'[OF f] guess u .
- note u = this positive_integral_isoton_simple[OF this(1-2)]
- from borel_measurable_implies_simple_function_sequence'[OF g] guess v .
- note v = this positive_integral_isoton_simple[OF this(1-2)]
+ from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u .
+ note u = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this
+ from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v .
+ note v = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this
let "?L' i x" = "a * u i x + v i x"
- have "?L \<in> borel_measurable M"
- using assms by simp
+ have "?L \<in> borel_measurable M" using assms by auto
from borel_measurable_implies_simple_function_sequence'[OF this] guess l .
- note positive_integral_isoton_simple[OF this(1-2)] and l = this
- moreover have "(SUP i. integral\<^isup>S M (l i)) = (SUP i. integral\<^isup>S M (?L' i))"
- proof (rule SUP_simple_integral_sequences[OF l(1-2)])
- show "?L' \<up> ?L" "\<And>i. simple_function M (?L' i)"
- using u v by (auto simp: isoton_fun_expand isoton_add isoton_cmult_right)
+ note l = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this
+
+ have inc: "incseq (\<lambda>i. a * integral\<^isup>S M (u i))" "incseq (\<lambda>i. integral\<^isup>S M (v i))"
+ using u v `0 \<le> a`
+ by (auto simp: incseq_Suc_iff le_fun_def
+ intro!: add_mono extreal_mult_left_mono simple_integral_mono)
+ have pos: "\<And>i. 0 \<le> integral\<^isup>S M (u i)" "\<And>i. 0 \<le> integral\<^isup>S M (v i)" "\<And>i. 0 \<le> a * integral\<^isup>S M (u i)"
+ using u v `0 \<le> a` by (auto simp: simple_integral_positive)
+ { fix i from pos[of i] have "a * integral\<^isup>S M (u i) \<noteq> -\<infinity>" "integral\<^isup>S M (v i) \<noteq> -\<infinity>"
+ by (auto split: split_if_asm) }
+ note not_MInf = this
+
+ have l': "(SUP i. integral\<^isup>S M (l i)) = (SUP i. integral\<^isup>S M (?L' i))"
+ proof (rule SUP_simple_integral_sequences[OF l(3,6,2)])
+ show "incseq ?L'" "\<And>i x. 0 \<le> ?L' i x" "\<And>i. simple_function M (?L' i)"
+ using u v `0 \<le> a` unfolding incseq_Suc_iff le_fun_def
+ by (auto intro!: add_mono extreal_mult_left_mono extreal_add_nonneg_nonneg)
+ { fix x
+ { fix i have "a * u i x \<noteq> -\<infinity>" "v i x \<noteq> -\<infinity>" "u i x \<noteq> -\<infinity>" using `0 \<le> a` u(6)[of i x] v(6)[of i x]
+ by auto }
+ then have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)"
+ using `0 \<le> a` u(3) v(3) u(6)[of _ x] v(6)[of _ x]
+ by (subst SUPR_extreal_cmult[symmetric, OF u(6) `0 \<le> a`])
+ (auto intro!: SUPR_extreal_add
+ simp: incseq_Suc_iff le_fun_def add_mono extreal_mult_left_mono extreal_add_nonneg_nonneg) }
+ then show "AE x. (SUP i. l i x) = (SUP i. ?L' i x)"
+ unfolding l(5) using `0 \<le> a` u(5) v(5) l(5) f(2) g(2)
+ by (intro AE_I2) (auto split: split_max simp add: extreal_add_nonneg_nonneg)
qed
- moreover from u v have L'_isoton:
- "(\<lambda>i. integral\<^isup>S M (?L' i)) \<up> a * integral\<^isup>P M f + integral\<^isup>P M g"
- by (simp add: isoton_add isoton_cmult_right)
- ultimately show ?thesis by (simp add: isoton_def)
+ also have "\<dots> = (SUP i. a * integral\<^isup>S M (u i) + integral\<^isup>S M (v i))"
+ using u(2, 6) v(2, 6) `0 \<le> a` by (auto intro!: arg_cong[where f="SUPR UNIV"] ext)
+ finally have "(\<integral>\<^isup>+ x. max 0 (a * f x + g x) \<partial>M) = a * (\<integral>\<^isup>+x. max 0 (f x) \<partial>M) + (\<integral>\<^isup>+x. max 0 (g x) \<partial>M)"
+ unfolding l(5)[symmetric] u(5)[symmetric] v(5)[symmetric]
+ unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric]
+ apply (subst SUPR_extreal_cmult[symmetric, OF pos(1) `0 \<le> a`])
+ apply (subst SUPR_extreal_add[symmetric, OF inc not_MInf]) .
+ then show ?thesis by (simp add: positive_integral_max_0)
qed
lemma (in measure_space) positive_integral_cmult:
- assumes "f \<in> borel_measurable M"
+ assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x" "0 \<le> c"
shows "(\<integral>\<^isup>+ x. c * f x \<partial>M) = c * integral\<^isup>P M f"
- using positive_integral_linear[OF assms, of "\<lambda>x. 0" c] by auto
+proof -
+ have [simp]: "\<And>x. c * max 0 (f x) = max 0 (c * f x)" using `0 \<le> c`
+ by (auto split: split_max simp: extreal_zero_le_0_iff)
+ have "(\<integral>\<^isup>+ x. c * f x \<partial>M) = (\<integral>\<^isup>+ x. c * max 0 (f x) \<partial>M)"
+ by (simp add: positive_integral_max_0)
+ then show ?thesis
+ using positive_integral_linear[OF _ _ `0 \<le> c`, of "\<lambda>x. max 0 (f x)" "\<lambda>x. 0"] f
+ by (auto simp: positive_integral_max_0)
+qed
lemma (in measure_space) positive_integral_multc:
- assumes "f \<in> borel_measurable M"
+ assumes "f \<in> borel_measurable M" "AE x. 0 \<le> f x" "0 \<le> c"
shows "(\<integral>\<^isup>+ x. f x * c \<partial>M) = integral\<^isup>P M f * c"
unfolding mult_commute[of _ c] positive_integral_cmult[OF assms] by simp
@@ -1260,143 +1289,172 @@
(auto simp: simple_function_indicator simple_integral_indicator)
lemma (in measure_space) positive_integral_cmult_indicator:
- "A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. c * indicator A x \<partial>M) = c * \<mu> A"
+ "0 \<le> c \<Longrightarrow> A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. c * indicator A x \<partial>M) = c * \<mu> A"
by (subst positive_integral_eq_simple_integral)
(auto simp: simple_function_indicator simple_integral_indicator)
lemma (in measure_space) positive_integral_add:
- assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+ assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
+ and g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
shows "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = integral\<^isup>P M f + integral\<^isup>P M g"
- using positive_integral_linear[OF assms, of 1] by simp
+proof -
+ have ae: "AE x. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)"
+ using assms by (auto split: split_max simp: extreal_add_nonneg_nonneg)
+ have "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = (\<integral>\<^isup>+ x. max 0 (f x + g x) \<partial>M)"
+ by (simp add: positive_integral_max_0)
+ also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) + max 0 (g x) \<partial>M)"
+ unfolding ae[THEN positive_integral_cong_AE] ..
+ also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) \<partial>M) + (\<integral>\<^isup>+ x. max 0 (g x) \<partial>M)"
+ using positive_integral_linear[of "\<lambda>x. max 0 (f x)" 1 "\<lambda>x. max 0 (g x)"] f g
+ by auto
+ finally show ?thesis
+ by (simp add: positive_integral_max_0)
+qed
lemma (in measure_space) positive_integral_setsum:
- assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M"
+ assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M" "\<And>i. i\<in>P \<Longrightarrow> AE x. 0 \<le> f i x"
shows "(\<integral>\<^isup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^isup>P M (f i))"
proof cases
- assume "finite P"
- from this assms show ?thesis
+ assume f: "finite P"
+ from assms have "AE x. \<forall>i\<in>P. 0 \<le> f i x" unfolding AE_finite_all[OF f] by auto
+ from f this assms(1) show ?thesis
proof induct
case (insert i P)
- have "f i \<in> borel_measurable M"
- "(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M"
- using insert by (auto intro!: borel_measurable_pextreal_setsum)
+ then have "f i \<in> borel_measurable M" "AE x. 0 \<le> f i x"
+ "(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M" "AE x. 0 \<le> (\<Sum>i\<in>P. f i x)"
+ by (auto intro!: borel_measurable_extreal_setsum setsum_nonneg)
from positive_integral_add[OF this]
show ?case using insert by auto
qed simp
qed simp
+lemma (in measure_space) positive_integral_Markov_inequality:
+ assumes u: "u \<in> borel_measurable M" "AE x. 0 \<le> u x" and "A \<in> sets M" and c: "0 \<le> c" "c \<noteq> \<infinity>"
+ shows "\<mu> ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
+ (is "\<mu> ?A \<le> _ * ?PI")
+proof -
+ have "?A \<in> sets M"
+ using `A \<in> sets M` u by auto
+ hence "\<mu> ?A = (\<integral>\<^isup>+ x. indicator ?A x \<partial>M)"
+ using positive_integral_indicator by simp
+ also have "\<dots> \<le> (\<integral>\<^isup>+ x. c * (u x * indicator A x) \<partial>M)" using u c
+ by (auto intro!: positive_integral_mono_AE
+ simp: indicator_def extreal_zero_le_0_iff)
+ also have "\<dots> = c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
+ using assms
+ by (auto intro!: positive_integral_cmult borel_measurable_indicator simp: extreal_zero_le_0_iff)
+ finally show ?thesis .
+qed
+
+lemma (in measure_space) positive_integral_noteq_infinite:
+ assumes g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
+ and "integral\<^isup>P M g \<noteq> \<infinity>"
+ shows "AE x. g x \<noteq> \<infinity>"
+proof (rule ccontr)
+ assume c: "\<not> (AE x. g x \<noteq> \<infinity>)"
+ have "\<mu> {x\<in>space M. g x = \<infinity>} \<noteq> 0"
+ using c g by (simp add: AE_iff_null_set)
+ moreover have "0 \<le> \<mu> {x\<in>space M. g x = \<infinity>}" using g by (auto intro: measurable_sets)
+ ultimately have "0 < \<mu> {x\<in>space M. g x = \<infinity>}" by auto
+ then have "\<infinity> = \<infinity> * \<mu> {x\<in>space M. g x = \<infinity>}" by auto
+ also have "\<dots> \<le> (\<integral>\<^isup>+x. \<infinity> * indicator {x\<in>space M. g x = \<infinity>} x \<partial>M)"
+ using g by (subst positive_integral_cmult_indicator) auto
+ also have "\<dots> \<le> integral\<^isup>P M g"
+ using assms by (auto intro!: positive_integral_mono_AE simp: indicator_def)
+ finally show False using `integral\<^isup>P M g \<noteq> \<infinity>` by auto
+qed
+
lemma (in measure_space) positive_integral_diff:
- assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M"
- and fin: "integral\<^isup>P M g \<noteq> \<omega>"
- and mono: "\<And>x. x \<in> space M \<Longrightarrow> g x \<le> f x"
+ assumes f: "f \<in> borel_measurable M"
+ and g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
+ and fin: "integral\<^isup>P M g \<noteq> \<infinity>"
+ and mono: "AE x. g x \<le> f x"
shows "(\<integral>\<^isup>+ x. f x - g x \<partial>M) = integral\<^isup>P M f - integral\<^isup>P M g"
proof -
- have borel: "(\<lambda>x. f x - g x) \<in> borel_measurable M"
- using f g by (rule borel_measurable_pextreal_diff)
- have "(\<integral>\<^isup>+x. f x - g x \<partial>M) + integral\<^isup>P M g = integral\<^isup>P M f"
- unfolding positive_integral_add[OF borel g, symmetric]
- proof (rule positive_integral_cong)
- fix x assume "x \<in> space M"
- from mono[OF this] show "f x - g x + g x = f x"
- by (cases "f x", cases "g x", simp, simp, cases "g x", auto)
- qed
- with mono show ?thesis
- by (subst minus_pextreal_eq2[OF _ fin]) (auto intro!: positive_integral_mono)
+ have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M" "AE x. 0 \<le> f x - g x"
+ using assms by (auto intro: extreal_diff_positive)
+ have pos_f: "AE x. 0 \<le> f x" using mono g by auto
+ { fix a b :: extreal assume "0 \<le> a" "a \<noteq> \<infinity>" "0 \<le> b" "a \<le> b" then have "b - a + a = b"
+ by (cases rule: extreal2_cases[of a b]) auto }
+ note * = this
+ then have "AE x. f x = f x - g x + g x"
+ using mono positive_integral_noteq_infinite[OF g fin] assms by auto
+ then have **: "integral\<^isup>P M f = (\<integral>\<^isup>+x. f x - g x \<partial>M) + integral\<^isup>P M g"
+ unfolding positive_integral_add[OF diff g, symmetric]
+ by (rule positive_integral_cong_AE)
+ show ?thesis unfolding **
+ using fin positive_integral_positive[of g]
+ by (cases rule: extreal2_cases[of "\<integral>\<^isup>+ x. f x - g x \<partial>M" "integral\<^isup>P M g"]) auto
qed
-lemma (in measure_space) positive_integral_psuminf:
- assumes "\<And>i. f i \<in> borel_measurable M"
- shows "(\<integral>\<^isup>+ x. (\<Sum>\<^isub>\<infinity> i. f i x) \<partial>M) = (\<Sum>\<^isub>\<infinity> i. integral\<^isup>P M (f i))"
+lemma (in measure_space) positive_integral_suminf:
+ assumes f: "\<And>i. f i \<in> borel_measurable M" "\<And>i. AE x. 0 \<le> f i x"
+ shows "(\<integral>\<^isup>+ x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^isup>P M (f i))"
proof -
- have "(\<lambda>i. (\<integral>\<^isup>+x. (\<Sum>i<i. f i x) \<partial>M)) \<up> (\<integral>\<^isup>+x. (\<Sum>\<^isub>\<infinity>i. f i x) \<partial>M)"
- by (rule positive_integral_isoton)
- (auto intro!: borel_measurable_pextreal_setsum assms positive_integral_mono
- arg_cong[where f=Sup]
- simp: isoton_def le_fun_def psuminf_def fun_eq_iff SUPR_def Sup_fun_def)
- thus ?thesis
- by (auto simp: isoton_def psuminf_def positive_integral_setsum[OF assms])
+ have all_pos: "AE x. \<forall>i. 0 \<le> f i x"
+ using assms by (auto simp: AE_all_countable)
+ have "(\<Sum>i. integral\<^isup>P M (f i)) = (SUP n. \<Sum>i<n. integral\<^isup>P M (f i))"
+ using positive_integral_positive by (rule suminf_extreal_eq_SUPR)
+ also have "\<dots> = (SUP n. \<integral>\<^isup>+x. (\<Sum>i<n. f i x) \<partial>M)"
+ unfolding positive_integral_setsum[OF f] ..
+ also have "\<dots> = \<integral>\<^isup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos
+ by (intro positive_integral_monotone_convergence_SUP_AE[symmetric])
+ (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3)
+ also have "\<dots> = \<integral>\<^isup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos
+ by (intro positive_integral_cong_AE) (auto simp: suminf_extreal_eq_SUPR)
+ finally show ?thesis by simp
qed
text {* Fatou's lemma: convergence theorem on limes inferior *}
lemma (in measure_space) positive_integral_lim_INF:
- fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> pextreal"
- assumes "\<And>i. u i \<in> borel_measurable M"
- shows "(\<integral>\<^isup>+ x. (SUP n. INF m. u (m + n) x) \<partial>M) \<le>
- (SUP n. INF m. integral\<^isup>P M (u (m + n)))"
+ fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
+ assumes u: "\<And>i. u i \<in> borel_measurable M" "\<And>i. AE x. 0 \<le> u i x"
+ shows "(\<integral>\<^isup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^isup>P M (u n))"
proof -
- have "(\<integral>\<^isup>+x. (SUP n. INF m. u (m + n) x) \<partial>M)
- = (SUP n. (\<integral>\<^isup>+x. (INF m. u (m + n) x) \<partial>M))"
- using assms
- by (intro positive_integral_monotone_convergence_SUP[symmetric] INF_mono)
- (auto simp del: add_Suc simp add: add_Suc[symmetric])
- also have "\<dots> \<le> (SUP n. INF m. integral\<^isup>P M (u (m + n)))"
- by (auto intro!: SUP_mono bexI le_INFI positive_integral_mono INF_leI)
+ have pos: "AE x. \<forall>i. 0 \<le> u i x" using u by (auto simp: AE_all_countable)
+ have "(\<integral>\<^isup>+ x. liminf (\<lambda>n. u n x) \<partial>M) =
+ (SUP n. \<integral>\<^isup>+ x. (INF i:{n..}. u i x) \<partial>M)"
+ unfolding liminf_SUPR_INFI using pos u
+ by (intro positive_integral_monotone_convergence_SUP_AE)
+ (elim AE_mp, auto intro!: AE_I2 intro: le_INFI INF_subset)
+ also have "\<dots> \<le> liminf (\<lambda>n. integral\<^isup>P M (u n))"
+ unfolding liminf_SUPR_INFI
+ by (auto intro!: SUP_mono exI le_INFI positive_integral_mono INF_leI)
finally show ?thesis .
qed
lemma (in measure_space) measure_space_density:
- assumes borel: "u \<in> borel_measurable M"
+ assumes u: "u \<in> borel_measurable M" "AE x. 0 \<le> u x"
and M'[simp]: "M' = (M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)\<rparr>)"
shows "measure_space M'"
proof -
interpret M': sigma_algebra M' by (intro sigma_algebra_cong) auto
show ?thesis
proof
- show "measure M' {} = 0" unfolding M' by simp
+ have pos: "\<And>A. AE x. 0 \<le> u x * indicator A x"
+ using u by (auto simp: extreal_zero_le_0_iff)
+ then show "positive M' (measure M')" unfolding M'
+ using u(1) by (auto simp: positive_def intro!: positive_integral_positive)
show "countably_additive M' (measure M')"
proof (intro countably_additiveI)
fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets M'"
- then have "\<And>i. (\<lambda>x. u x * indicator (A i) x) \<in> borel_measurable M"
- using borel by (auto intro: borel_measurable_indicator)
- moreover assume "disjoint_family A"
- note psuminf_indicator[OF this]
- ultimately show "(\<Sum>\<^isub>\<infinity>n. measure M' (A n)) = measure M' (\<Union>x. A x)"
- by (simp add: positive_integral_psuminf[symmetric])
+ then have *: "\<And>i. (\<lambda>x. u x * indicator (A i) x) \<in> borel_measurable M"
+ using u by (auto intro: borel_measurable_indicator)
+ assume disj: "disjoint_family A"
+ have "(\<Sum>n. measure M' (A n)) = (\<integral>\<^isup>+ x. (\<Sum>n. u x * indicator (A n) x) \<partial>M)"
+ unfolding M' using u(1) *
+ by (simp add: positive_integral_suminf[OF _ pos, symmetric])
+ also have "\<dots> = (\<integral>\<^isup>+ x. u x * (\<Sum>n. indicator (A n) x) \<partial>M)" using u
+ by (intro positive_integral_cong_AE)
+ (elim AE_mp, auto intro!: AE_I2 suminf_cmult_extreal)
+ also have "\<dots> = (\<integral>\<^isup>+ x. u x * indicator (\<Union>n. A n) x \<partial>M)"
+ unfolding suminf_indicator[OF disj] ..
+ finally show "(\<Sum>n. measure M' (A n)) = measure M' (\<Union>x. A x)"
+ unfolding M' by simp
qed
qed
qed
-lemma (in measure_space) positive_integral_translated_density:
- assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
- and M': "M' = (M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)\<rparr>)"
- shows "integral\<^isup>P M' g = (\<integral>\<^isup>+ x. f x * g x \<partial>M)"
-proof -
- from measure_space_density[OF assms(1) M']
- interpret T: measure_space M' .
- have borel[simp]:
- "borel_measurable M' = borel_measurable M"
- "simple_function M' = simple_function M"
- unfolding measurable_def simple_function_def_raw by (auto simp: M')
- from borel_measurable_implies_simple_function_sequence[OF assms(2)]
- obtain G where G: "\<And>i. simple_function M (G i)" "G \<up> g" by blast
- note G_borel = borel_measurable_simple_function[OF this(1)]
- from T.positive_integral_isoton[unfolded borel, OF `G \<up> g` G_borel]
- have *: "(\<lambda>i. integral\<^isup>P M' (G i)) \<up> integral\<^isup>P M' g" .
- { fix i
- have [simp]: "finite (G i ` space M)"
- using G(1) unfolding simple_function_def by auto
- have "integral\<^isup>P M' (G i) = integral\<^isup>S M' (G i)"
- using G T.positive_integral_eq_simple_integral by simp
- also have "\<dots> = (\<integral>\<^isup>+x. f x * (\<Sum>y\<in>G i`space M. y * indicator (G i -` {y} \<inter> space M) x) \<partial>M)"
- apply (simp add: simple_integral_def M')
- apply (subst positive_integral_cmult[symmetric])
- using G_borel assms(1) apply (fastsimp intro: borel_measurable_vimage)
- apply (subst positive_integral_setsum[symmetric])
- using G_borel assms(1) apply (fastsimp intro: borel_measurable_vimage)
- by (simp add: setsum_right_distrib field_simps)
- also have "\<dots> = (\<integral>\<^isup>+x. f x * G i x \<partial>M)"
- by (auto intro!: positive_integral_cong
- simp: indicator_def if_distrib setsum_cases)
- finally have "integral\<^isup>P M' (G i) = (\<integral>\<^isup>+x. f x * G i x \<partial>M)" . }
- with * have eq_Tg: "(\<lambda>i. (\<integral>\<^isup>+x. f x * G i x \<partial>M)) \<up> integral\<^isup>P M' g" by simp
- from G(2) have "(\<lambda>i x. f x * G i x) \<up> (\<lambda>x. f x * g x)"
- unfolding isoton_fun_expand by (auto intro!: isoton_cmult_right)
- then have "(\<lambda>i. (\<integral>\<^isup>+x. f x * G i x \<partial>M)) \<up> (\<integral>\<^isup>+x. f x * g x \<partial>M)"
- using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pextreal_times)
- with eq_Tg show "integral\<^isup>P M' g = (\<integral>\<^isup>+x. f x * g x \<partial>M)"
- unfolding isoton_def by simp
-qed
-
lemma (in measure_space) positive_integral_null_set:
assumes "N \<in> null_sets" shows "(\<integral>\<^isup>+ x. u x * indicator N x \<partial>M) = 0"
proof -
@@ -1410,144 +1468,199 @@
then show ?thesis by simp
qed
-lemma (in measure_space) positive_integral_Markov_inequality:
- assumes borel: "u \<in> borel_measurable M" and "A \<in> sets M" and c: "c \<noteq> \<omega>"
- shows "\<mu> ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
- (is "\<mu> ?A \<le> _ * ?PI")
+lemma (in measure_space) positive_integral_translated_density:
+ assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
+ assumes g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
+ and M': "M' = (M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)\<rparr>)"
+ shows "integral\<^isup>P M' g = (\<integral>\<^isup>+ x. f x * g x \<partial>M)"
proof -
- have "?A \<in> sets M"
- using `A \<in> sets M` borel by auto
- hence "\<mu> ?A = (\<integral>\<^isup>+ x. indicator ?A x \<partial>M)"
- using positive_integral_indicator by simp
- also have "\<dots> \<le> (\<integral>\<^isup>+ x. c * (u x * indicator A x) \<partial>M)"
- proof (rule positive_integral_mono)
- fix x assume "x \<in> space M"
- show "indicator ?A x \<le> c * (u x * indicator A x)"
- by (cases "x \<in> ?A") auto
- qed
- also have "\<dots> = c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
- using assms
- by (auto intro!: positive_integral_cmult borel_measurable_indicator)
- finally show ?thesis .
+ from measure_space_density[OF f M']
+ interpret T: measure_space M' .
+ have borel[simp]:
+ "borel_measurable M' = borel_measurable M"
+ "simple_function M' = simple_function M"
+ unfolding measurable_def simple_function_def_raw by (auto simp: M')
+ from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess G . note G = this
+ note G' = borel_measurable_simple_function[OF this(1)] simple_functionD[OF G(1)]
+ note G'(2)[simp]
+ { fix P have "AE x. P x \<Longrightarrow> AE x in M'. P x"
+ using positive_integral_null_set[of _ f]
+ unfolding T.almost_everywhere_def almost_everywhere_def
+ by (auto simp: M') }
+ note ac = this
+ from G(4) g(2) have G_M': "AE x in M'. (SUP i. G i x) = g x"
+ by (auto intro!: ac split: split_max)
+ { fix i
+ let "?I y x" = "indicator (G i -` {y} \<inter> space M) x"
+ { fix x assume *: "x \<in> space M" "0 \<le> f x" "0 \<le> g x"
+ then have [simp]: "G i ` space M \<inter> {y. G i x = y \<and> x \<in> space M} = {G i x}" by auto
+ from * G' G have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * (\<Sum>y\<in>G i`space M. (y * ?I y x))"
+ by (subst setsum_extreal_right_distrib) (auto simp: ac_simps)
+ also have "\<dots> = f x * G i x"
+ by (simp add: indicator_def if_distrib setsum_cases)
+ finally have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * G i x" . }
+ note to_singleton = this
+ have "integral\<^isup>P M' (G i) = integral\<^isup>S M' (G i)"
+ using G T.positive_integral_eq_simple_integral by simp
+ also have "\<dots> = (\<Sum>y\<in>G i`space M. y * (\<integral>\<^isup>+x. f x * ?I y x \<partial>M))"
+ unfolding simple_integral_def M' by simp
+ also have "\<dots> = (\<Sum>y\<in>G i`space M. (\<integral>\<^isup>+x. y * (f x * ?I y x) \<partial>M))"
+ using f G' G by (auto intro!: setsum_cong positive_integral_cmult[symmetric])
+ also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) \<partial>M)"
+ using f G' G by (auto intro!: positive_integral_setsum[symmetric])
+ finally have "integral\<^isup>P M' (G i) = (\<integral>\<^isup>+x. f x * G i x \<partial>M)"
+ using f g G' to_singleton by (auto intro!: positive_integral_cong_AE) }
+ note [simp] = this
+ have "integral\<^isup>P M' g = (SUP i. integral\<^isup>P M' (G i))" using G'(1) G_M'(1) G
+ using T.positive_integral_monotone_convergence_SUP[symmetric, OF `incseq G`]
+ by (simp cong: T.positive_integral_cong_AE)
+ also have "\<dots> = (SUP i. (\<integral>\<^isup>+x. f x * G i x \<partial>M))" by simp
+ also have "\<dots> = (\<integral>\<^isup>+x. (SUP i. f x * G i x) \<partial>M)"
+ using f G' G(2)[THEN incseq_SucD] G
+ by (intro positive_integral_monotone_convergence_SUP_AE[symmetric])
+ (auto simp: extreal_mult_left_mono le_fun_def extreal_zero_le_0_iff)
+ also have "\<dots> = (\<integral>\<^isup>+x. f x * g x \<partial>M)" using f G' G g
+ by (intro positive_integral_cong_AE)
+ (auto simp add: SUPR_extreal_cmult split: split_max)
+ finally show "integral\<^isup>P M' g = (\<integral>\<^isup>+x. f x * g x \<partial>M)" .
qed
lemma (in measure_space) positive_integral_0_iff:
- assumes borel: "u \<in> borel_measurable M"
+ assumes u: "u \<in> borel_measurable M" and pos: "AE x. 0 \<le> u x"
shows "integral\<^isup>P M u = 0 \<longleftrightarrow> \<mu> {x\<in>space M. u x \<noteq> 0} = 0"
(is "_ \<longleftrightarrow> \<mu> ?A = 0")
proof -
- have A: "?A \<in> sets M" using borel by auto
- have u: "(\<integral>\<^isup>+ x. u x * indicator ?A x \<partial>M) = integral\<^isup>P M u"
+ have u_eq: "(\<integral>\<^isup>+ x. u x * indicator ?A x \<partial>M) = integral\<^isup>P M u"
by (auto intro!: positive_integral_cong simp: indicator_def)
-
show ?thesis
proof
assume "\<mu> ?A = 0"
- hence "?A \<in> null_sets" using `?A \<in> sets M` by auto
- from positive_integral_null_set[OF this]
- have "0 = (\<integral>\<^isup>+ x. u x * indicator ?A x \<partial>M)" by simp
- thus "integral\<^isup>P M u = 0" unfolding u by simp
+ with positive_integral_null_set[of ?A u] u
+ show "integral\<^isup>P M u = 0" by (simp add: u_eq)
next
+ { fix r :: extreal and n :: nat assume gt_1: "1 \<le> real n * r"
+ then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_extreal_def)
+ then have "0 \<le> r" by (auto simp add: extreal_zero_less_0_iff) }
+ note gt_1 = this
assume *: "integral\<^isup>P M u = 0"
- let "?M n" = "{x \<in> space M. 1 \<le> of_nat n * u x}"
+ let "?M n" = "{x \<in> space M. 1 \<le> real (n::nat) * u x}"
have "0 = (SUP n. \<mu> (?M n \<inter> ?A))"
proof -
- { fix n
- from positive_integral_Markov_inequality[OF borel `?A \<in> sets M`, of "of_nat n"]
- have "\<mu> (?M n \<inter> ?A) = 0" unfolding * u by simp }
+ { fix n :: nat
+ from positive_integral_Markov_inequality[OF u pos, of ?A "extreal (real n)"]
+ have "\<mu> (?M n \<inter> ?A) \<le> 0" unfolding u_eq * using u by simp
+ moreover have "0 \<le> \<mu> (?M n \<inter> ?A)" using u by auto
+ ultimately have "\<mu> (?M n \<inter> ?A) = 0" by auto }
thus ?thesis by simp
qed
also have "\<dots> = \<mu> (\<Union>n. ?M n \<inter> ?A)"
proof (safe intro!: continuity_from_below)
fix n show "?M n \<inter> ?A \<in> sets M"
- using borel by (auto intro!: Int)
+ using u by (auto intro!: Int)
next
- fix n x assume "1 \<le> of_nat n * u x"
- also have "\<dots> \<le> of_nat (Suc n) * u x"
- by (subst (1 2) mult_commute) (auto intro!: pextreal_mult_cancel)
- finally show "1 \<le> of_nat (Suc n) * u x" .
+ show "incseq (\<lambda>n. {x \<in> space M. 1 \<le> real n * u x} \<inter> {x \<in> space M. u x \<noteq> 0})"
+ proof (safe intro!: incseq_SucI)
+ fix n :: nat and x
+ assume *: "1 \<le> real n * u x"
+ also from gt_1[OF this] have "real n * u x \<le> real (Suc n) * u x"
+ using `0 \<le> u x` by (auto intro!: extreal_mult_right_mono)
+ finally show "1 \<le> real (Suc n) * u x" by auto
+ qed
qed
- also have "\<dots> = \<mu> ?A"
- proof (safe intro!: arg_cong[where f="\<mu>"])
- fix x assume "u x \<noteq> 0" and [simp, intro]: "x \<in> space M"
+ also have "\<dots> = \<mu> {x\<in>space M. 0 < u x}"
+ proof (safe intro!: arg_cong[where f="\<mu>"] dest!: gt_1)
+ fix x assume "0 < u x" and [simp, intro]: "x \<in> space M"
show "x \<in> (\<Union>n. ?M n \<inter> ?A)"
proof (cases "u x")
- case (preal r)
- obtain j where "1 / r \<le> of_nat j" using ex_le_of_nat ..
- hence "1 / r * r \<le> of_nat j * r" using preal unfolding mult_le_cancel_right by auto
- hence "1 \<le> of_nat j * r" using preal `u x \<noteq> 0` by auto
- thus ?thesis using `u x \<noteq> 0` preal by (auto simp: real_of_nat_def[symmetric])
- qed auto
- qed
- finally show "\<mu> ?A = 0" by simp
+ case (real r) with `0 < u x` have "0 < r" by auto
+ obtain j :: nat where "1 / r \<le> real j" using real_arch_simple ..
+ hence "1 / r * r \<le> real j * r" unfolding mult_le_cancel_right using `0 < r` by auto
+ hence "1 \<le> real j * r" using real `0 < r` by auto
+ thus ?thesis using `0 < r` real by (auto simp: one_extreal_def)
+ qed (insert `0 < u x`, auto)
+ qed auto
+ finally have "\<mu> {x\<in>space M. 0 < u x} = 0" by simp
+ moreover
+ from pos have "AE x. \<not> (u x < 0)" by auto
+ then have "\<mu> {x\<in>space M. u x < 0} = 0"
+ using AE_iff_null_set u by auto
+ moreover have "\<mu> {x\<in>space M. u x \<noteq> 0} = \<mu> {x\<in>space M. u x < 0} + \<mu> {x\<in>space M. 0 < u x}"
+ using u by (subst measure_additive) (auto intro!: arg_cong[where f=\<mu>])
+ ultimately show "\<mu> ?A = 0" by simp
qed
qed
lemma (in measure_space) positive_integral_0_iff_AE:
assumes u: "u \<in> borel_measurable M"
- shows "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x. u x = 0)"
+ shows "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x. u x \<le> 0)"
proof -
- have sets: "{x\<in>space M. u x \<noteq> 0} \<in> sets M"
+ have sets: "{x\<in>space M. max 0 (u x) \<noteq> 0} \<in> sets M"
using u by auto
- then show ?thesis
- using positive_integral_0_iff[OF u] AE_iff_null_set[OF sets] by auto
+ from positive_integral_0_iff[of "\<lambda>x. max 0 (u x)"]
+ have "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x. max 0 (u x) = 0)"
+ unfolding positive_integral_max_0
+ using AE_iff_null_set[OF sets] u by auto
+ also have "\<dots> \<longleftrightarrow> (AE x. u x \<le> 0)" by (auto split: split_max)
+ finally show ?thesis .
qed
lemma (in measure_space) positive_integral_restricted:
- assumes "A \<in> sets M"
+ assumes A: "A \<in> sets M"
shows "integral\<^isup>P (restricted_space A) f = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)"
(is "integral\<^isup>P ?R f = integral\<^isup>P M ?f")
proof -
- have msR: "measure_space ?R" by (rule restricted_measure_space[OF `A \<in> sets M`])
- then interpret R: measure_space ?R .
- have saR: "sigma_algebra ?R" by fact
- have *: "integral\<^isup>P ?R f = integral\<^isup>P ?R ?f"
- by (intro R.positive_integral_cong) auto
+ interpret R: measure_space ?R
+ by (rule restricted_measure_space) fact
+ let "?I g x" = "g x * indicator A x :: extreal"
show ?thesis
- unfolding * positive_integral_def
- unfolding simple_function_restricted[OF `A \<in> sets M`]
- apply (simp add: SUPR_def)
- apply (rule arg_cong[where f=Sup])
- proof (auto simp add: image_iff simple_integral_restricted[OF `A \<in> sets M`])
- fix g assume "simple_function M (\<lambda>x. g x * indicator A x)"
- "g \<le> f"
- then show "\<exists>x. simple_function M x \<and> x \<le> (\<lambda>x. f x * indicator A x) \<and>
- (\<integral>\<^isup>Sx. g x * indicator A x \<partial>M) = integral\<^isup>S M x"
- apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"])
- by (auto simp: indicator_def le_fun_def)
+ unfolding positive_integral_def
+ unfolding simple_function_restricted[OF A]
+ unfolding AE_restricted[OF A]
+ proof (safe intro!: SUPR_eq)
+ fix g assume g: "simple_function M (?I g)" and le: "g \<le> max 0 \<circ> f"
+ show "\<exists>j\<in>{g. simple_function M g \<and> g \<le> max 0 \<circ> ?I f}.
+ integral\<^isup>S (restricted_space A) g \<le> integral\<^isup>S M j"
+ proof (safe intro!: bexI[of _ "?I g"])
+ show "integral\<^isup>S (restricted_space A) g \<le> integral\<^isup>S M (?I g)"
+ using g A by (simp add: simple_integral_restricted)
+ show "?I g \<le> max 0 \<circ> ?I f"
+ using le by (auto simp: le_fun_def max_def indicator_def split: split_if_asm)
+ qed fact
next
- fix g assume g: "simple_function M g" "g \<le> (\<lambda>x. f x * indicator A x)"
- then have *: "(\<lambda>x. g x * indicator A x) = g"
- "\<And>x. g x * indicator A x = g x"
- "\<And>x. g x \<le> f x"
- by (auto simp: le_fun_def fun_eq_iff indicator_def split: split_if_asm)
- from g show "\<exists>x. simple_function M (\<lambda>xa. x xa * indicator A xa) \<and> x \<le> f \<and>
- integral\<^isup>S M g = integral\<^isup>S M (\<lambda>xa. x xa * indicator A xa)"
- using `A \<in> sets M`[THEN sets_into_space]
- apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"])
- by (fastsimp simp: le_fun_def *)
+ fix g assume g: "simple_function M g" and le: "g \<le> max 0 \<circ> ?I f"
+ show "\<exists>i\<in>{g. simple_function M (?I g) \<and> g \<le> max 0 \<circ> f}.
+ integral\<^isup>S M g \<le> integral\<^isup>S (restricted_space A) i"
+ proof (safe intro!: bexI[of _ "?I g"])
+ show "?I g \<le> max 0 \<circ> f"
+ using le by (auto simp: le_fun_def max_def indicator_def split: split_if_asm)
+ from le have "\<And>x. g x \<le> ?I (?I g) x"
+ by (auto simp: le_fun_def max_def indicator_def split: split_if_asm)
+ then show "integral\<^isup>S M g \<le> integral\<^isup>S (restricted_space A) (?I g)"
+ using A g by (auto intro!: simple_integral_mono simp: simple_integral_restricted)
+ show "simple_function M (?I (?I g))" using g A by auto
+ qed
qed
qed
lemma (in measure_space) positive_integral_subalgebra:
- assumes borel: "f \<in> borel_measurable N"
+ assumes f: "f \<in> borel_measurable N" "AE x in N. 0 \<le> f x"
and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A"
and sa: "sigma_algebra N"
shows "integral\<^isup>P N f = integral\<^isup>P M f"
proof -
interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
- from N.borel_measurable_implies_simple_function_sequence[OF borel]
- obtain fs where Nsf: "\<And>i. simple_function N (fs i)" and "fs \<up> f" by blast
- then have sf: "\<And>i. simple_function M (fs i)"
- using simple_function_subalgebra[OF _ N(1,2)] by blast
- from N.positive_integral_isoton_simple[OF `fs \<up> f` Nsf]
+ from N.borel_measurable_implies_simple_function_sequence'[OF f(1)] guess fs . note fs = this
+ note sf = simple_function_subalgebra[OF fs(1) N(1,2)]
+ from N.positive_integral_monotone_convergence_simple[OF fs(2,5,1), symmetric]
have "integral\<^isup>P N f = (SUP i. \<Sum>x\<in>fs i ` space M. x * N.\<mu> (fs i -` {x} \<inter> space M))"
- unfolding isoton_def simple_integral_def `space N = space M` by simp
+ unfolding fs(4) positive_integral_max_0
+ unfolding simple_integral_def `space N = space M` by simp
also have "\<dots> = (SUP i. \<Sum>x\<in>fs i ` space M. x * \<mu> (fs i -` {x} \<inter> space M))"
- using N N.simple_functionD(2)[OF Nsf] unfolding `space N = space M` by auto
+ using N N.simple_functionD(2)[OF fs(1)] unfolding `space N = space M` by auto
also have "\<dots> = integral\<^isup>P M f"
- using positive_integral_isoton_simple[OF `fs \<up> f` sf]
- unfolding isoton_def simple_integral_def `space N = space M` by simp
+ using positive_integral_monotone_convergence_simple[OF fs(2,5) sf, symmetric]
+ unfolding fs(4) positive_integral_max_0
+ unfolding simple_integral_def `space N = space M` by simp
finally show ?thesis .
qed
@@ -1555,16 +1668,15 @@
definition integrable where
"integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and>
- (\<integral>\<^isup>+ x. Real (f x) \<partial>M) \<noteq> \<omega> \<and>
- (\<integral>\<^isup>+ x. Real (- f x) \<partial>M) \<noteq> \<omega>"
+ (\<integral>\<^isup>+ x. extreal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^isup>+ x. extreal (- f x) \<partial>M) \<noteq> \<infinity>"
lemma integrableD[dest]:
assumes "integrable M f"
- shows "f \<in> borel_measurable M" "(\<integral>\<^isup>+ x. Real (f x) \<partial>M) \<noteq> \<omega>" "(\<integral>\<^isup>+ x. Real (- f x) \<partial>M) \<noteq> \<omega>"
+ shows "f \<in> borel_measurable M" "(\<integral>\<^isup>+ x. extreal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^isup>+ x. extreal (- f x) \<partial>M) \<noteq> \<infinity>"
using assms unfolding integrable_def by auto
definition lebesgue_integral_def:
- "integral\<^isup>L M f = real ((\<integral>\<^isup>+ x. Real (f x) \<partial>M)) - real ((\<integral>\<^isup>+ x. Real (- f x) \<partial>M))"
+ "integral\<^isup>L M f = real ((\<integral>\<^isup>+ x. extreal (f x) \<partial>M)) - real ((\<integral>\<^isup>+ x. extreal (- f x) \<partial>M))"
syntax
"_lebesgue_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> real" ("\<integral> _. _ \<partial>_" [60,61] 110)
@@ -1572,6 +1684,17 @@
translations
"\<integral> x. f \<partial>M" == "CONST integral\<^isup>L M (%x. f)"
+lemma (in measure_space) integrableE:
+ assumes "integrable M f"
+ obtains r q where
+ "(\<integral>\<^isup>+x. extreal (f x)\<partial>M) = extreal r"
+ "(\<integral>\<^isup>+x. extreal (-f x)\<partial>M) = extreal q"
+ "f \<in> borel_measurable M" "integral\<^isup>L M f = r - q"
+ using assms unfolding integrable_def lebesgue_integral_def
+ using positive_integral_positive[of "\<lambda>x. extreal (f x)"]
+ using positive_integral_positive[of "\<lambda>x. extreal (-f x)"]
+ by (cases rule: extreal2_cases[of "(\<integral>\<^isup>+x. extreal (-f x)\<partial>M)" "(\<integral>\<^isup>+x. extreal (f x)\<partial>M)"]) auto
+
lemma (in measure_space) integral_cong:
assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
shows "integral\<^isup>L M f = integral\<^isup>L M g"
@@ -1580,21 +1703,16 @@
lemma (in measure_space) integral_cong_measure:
assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
shows "integral\<^isup>L N f = integral\<^isup>L M f"
-proof -
- interpret v: measure_space N
- by (rule measure_space_cong) fact+
- show ?thesis
- by (simp add: positive_integral_cong_measure[OF assms] lebesgue_integral_def)
-qed
+ by (simp add: positive_integral_cong_measure[OF assms] lebesgue_integral_def)
lemma (in measure_space) integral_cong_AE:
assumes cong: "AE x. f x = g x"
shows "integral\<^isup>L M f = integral\<^isup>L M g"
proof -
- have "AE x. Real (f x) = Real (g x)" using cong by auto
- moreover have "AE x. Real (- f x) = Real (- g x)" using cong by auto
- ultimately show ?thesis
- by (simp cong: positive_integral_cong_AE add: lebesgue_integral_def)
+ have *: "AE x. extreal (f x) = extreal (g x)"
+ "AE x. extreal (- f x) = extreal (- g x)" using cong by auto
+ show ?thesis
+ unfolding *[THEN positive_integral_cong_AE] lebesgue_integral_def ..
qed
lemma (in measure_space) integrable_cong:
@@ -1602,11 +1720,14 @@
by (simp cong: positive_integral_cong measurable_cong add: integrable_def)
lemma (in measure_space) integral_eq_positive_integral:
- assumes "\<And>x. 0 \<le> f x"
- shows "integral\<^isup>L M f = real (\<integral>\<^isup>+ x. Real (f x) \<partial>M)"
+ assumes f: "\<And>x. 0 \<le> f x"
+ shows "integral\<^isup>L M f = real (\<integral>\<^isup>+ x. extreal (f x) \<partial>M)"
proof -
- have "\<And>x. Real (- f x) = 0" using assms by simp
- thus ?thesis by (simp del: Real_eq_0 add: lebesgue_integral_def)
+ { fix x have "max 0 (extreal (- f x)) = 0" using f[of x] by (simp split: split_max) }
+ then have "0 = (\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>M)" by simp
+ also have "\<dots> = (\<integral>\<^isup>+ x. extreal (- f x) \<partial>M)" unfolding positive_integral_max_0 ..
+ finally show ?thesis
+ unfolding lebesgue_integral_def by simp
qed
lemma (in measure_space) integral_vimage:
@@ -1616,7 +1737,7 @@
proof -
interpret T: measure_space M' by (rule measure_space_vimage[OF T])
from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel]
- have borel: "(\<lambda>x. Real (f x)) \<in> borel_measurable M'" "(\<lambda>x. Real (- f x)) \<in> borel_measurable M'"
+ have borel: "(\<lambda>x. extreal (f x)) \<in> borel_measurable M'" "(\<lambda>x. extreal (- f x)) \<in> borel_measurable M'"
and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
using f by (auto simp: comp_def)
then show ?thesis
@@ -1631,7 +1752,7 @@
proof -
interpret T: measure_space M' by (rule measure_space_vimage[OF T])
from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel]
- have borel: "(\<lambda>x. Real (f x)) \<in> borel_measurable M'" "(\<lambda>x. Real (- f x)) \<in> borel_measurable M'"
+ have borel: "(\<lambda>x. extreal (f x)) \<in> borel_measurable M'" "(\<lambda>x. extreal (- f x)) \<in> borel_measurable M'"
and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
using f by (auto simp: comp_def)
then show ?thesis
@@ -1649,10 +1770,10 @@
and f_def: "\<And>x. f x = u x - v x" and pos: "\<And>x. 0 \<le> u x" "\<And>x. 0 \<le> v x"
shows "integrable M f" and "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v"
proof -
- let "?f x" = "Real (f x)"
- let "?mf x" = "Real (- f x)"
- let "?u x" = "Real (u x)"
- let "?v x" = "Real (v x)"
+ let "?f x" = "max 0 (extreal (f x))"
+ let "?mf x" = "max 0 (extreal (- f x))"
+ let "?u x" = "max 0 (extreal (u x))"
+ let "?v x" = "max 0 (extreal (v x))"
from borel_measurable_diff[of u v] integrable
have f_borel: "?f \<in> borel_measurable M" and
@@ -1662,73 +1783,62 @@
"f \<in> borel_measurable M"
by (auto simp: f_def[symmetric] integrable_def)
- have "(\<integral>\<^isup>+ x. Real (u x - v x) \<partial>M) \<le> integral\<^isup>P M ?u"
- using pos by (auto intro!: positive_integral_mono)
- moreover have "(\<integral>\<^isup>+ x. Real (v x - u x) \<partial>M) \<le> integral\<^isup>P M ?v"
- using pos by (auto intro!: positive_integral_mono)
+ have "(\<integral>\<^isup>+ x. extreal (u x - v x) \<partial>M) \<le> integral\<^isup>P M ?u"
+ using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0)
+ moreover have "(\<integral>\<^isup>+ x. extreal (v x - u x) \<partial>M) \<le> integral\<^isup>P M ?v"
+ using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0)
ultimately show f: "integrable M f"
using `integrable M u` `integrable M v` `f \<in> borel_measurable M`
- by (auto simp: integrable_def f_def)
- hence mf: "integrable M (\<lambda>x. - f x)" ..
-
- have *: "\<And>x. Real (- v x) = 0" "\<And>x. Real (- u x) = 0"
- using pos by auto
+ by (auto simp: integrable_def f_def positive_integral_max_0)
have "\<And>x. ?u x + ?mf x = ?v x + ?f x"
- unfolding f_def using pos by simp
- hence "(\<integral>\<^isup>+ x. ?u x + ?mf x \<partial>M) = (\<integral>\<^isup>+ x. ?v x + ?f x \<partial>M)" by simp
- hence "real (integral\<^isup>P M ?u + integral\<^isup>P M ?mf) =
+ unfolding f_def using pos by (simp split: split_max)
+ then have "(\<integral>\<^isup>+ x. ?u x + ?mf x \<partial>M) = (\<integral>\<^isup>+ x. ?v x + ?f x \<partial>M)" by simp
+ then have "real (integral\<^isup>P M ?u + integral\<^isup>P M ?mf) =
real (integral\<^isup>P M ?v + integral\<^isup>P M ?f)"
- using positive_integral_add[OF u_borel mf_borel]
- using positive_integral_add[OF v_borel f_borel]
+ using positive_integral_add[OF u_borel _ mf_borel]
+ using positive_integral_add[OF v_borel _ f_borel]
by auto
then show "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v"
- using f mf `integrable M u` `integrable M v`
- unfolding lebesgue_integral_def integrable_def *
- by (cases "integral\<^isup>P M ?f", cases "integral\<^isup>P M ?mf", cases "integral\<^isup>P M ?v", cases "integral\<^isup>P M ?u")
- (auto simp add: field_simps)
+ unfolding positive_integral_max_0
+ unfolding pos[THEN integral_eq_positive_integral]
+ using integrable f by (auto elim!: integrableE)
qed
lemma (in measure_space) integral_linear:
assumes "integrable M f" "integrable M g" and "0 \<le> a"
shows "integrable M (\<lambda>t. a * f t + g t)"
- and "(\<integral> t. a * f t + g t \<partial>M) = a * integral\<^isup>L M f + integral\<^isup>L M g"
+ and "(\<integral> t. a * f t + g t \<partial>M) = a * integral\<^isup>L M f + integral\<^isup>L M g" (is ?EQ)
proof -
- let ?PI = "integral\<^isup>P M"
- let "?f x" = "Real (f x)"
- let "?g x" = "Real (g x)"
- let "?mf x" = "Real (- f x)"
- let "?mg x" = "Real (- g x)"
+ let "?f x" = "max 0 (extreal (f x))"
+ let "?g x" = "max 0 (extreal (g x))"
+ let "?mf x" = "max 0 (extreal (- f x))"
+ let "?mg x" = "max 0 (extreal (- g x))"
let "?p t" = "max 0 (a * f t) + max 0 (g t)"
let "?n t" = "max 0 (- (a * f t)) + max 0 (- g t)"
- have pos: "?f \<in> borel_measurable M" "?g \<in> borel_measurable M"
- and neg: "?mf \<in> borel_measurable M" "?mg \<in> borel_measurable M"
- and p: "?p \<in> borel_measurable M"
- and n: "?n \<in> borel_measurable M"
- using assms by (simp_all add: integrable_def)
+ from assms have linear:
+ "(\<integral>\<^isup>+ x. extreal a * ?f x + ?g x \<partial>M) = extreal a * integral\<^isup>P M ?f + integral\<^isup>P M ?g"
+ "(\<integral>\<^isup>+ x. extreal a * ?mf x + ?mg x \<partial>M) = extreal a * integral\<^isup>P M ?mf + integral\<^isup>P M ?mg"
+ by (auto intro!: positive_integral_linear simp: integrable_def)
- have *: "\<And>x. Real (?p x) = Real a * ?f x + ?g x"
- "\<And>x. Real (?n x) = Real a * ?mf x + ?mg x"
- "\<And>x. Real (- ?p x) = 0"
- "\<And>x. Real (- ?n x) = 0"
- using `0 \<le> a` by (auto simp: max_def min_def zero_le_mult_iff mult_le_0_iff add_nonpos_nonpos)
-
- note linear =
- positive_integral_linear[OF pos]
- positive_integral_linear[OF neg]
+ have *: "(\<integral>\<^isup>+x. extreal (- ?p x) \<partial>M) = 0" "(\<integral>\<^isup>+x. extreal (- ?n x) \<partial>M) = 0"
+ using `0 \<le> a` assms by (auto simp: positive_integral_0_iff_AE integrable_def)
+ have **: "\<And>x. extreal a * ?f x + ?g x = max 0 (extreal (?p x))"
+ "\<And>x. extreal a * ?mf x + ?mg x = max 0 (extreal (?n x))"
+ using `0 \<le> a` by (auto split: split_max simp: zero_le_mult_iff mult_le_0_iff)
have "integrable M ?p" "integrable M ?n"
"\<And>t. a * f t + g t = ?p t - ?n t" "\<And>t. 0 \<le> ?p t" "\<And>t. 0 \<le> ?n t"
- using assms p n unfolding integrable_def * linear by auto
+ using linear assms unfolding integrable_def ** *
+ by (auto simp: positive_integral_max_0)
note diff = integral_of_positive_diff[OF this]
show "integrable M (\<lambda>t. a * f t + g t)" by (rule diff)
-
- from assms show "(\<integral> t. a * f t + g t \<partial>M) = a * integral\<^isup>L M f + integral\<^isup>L M g"
- unfolding diff(2) unfolding lebesgue_integral_def * linear integrable_def
- by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?g", cases "?PI ?mg")
- (auto simp add: field_simps zero_le_mult_iff)
+ from assms linear show ?EQ
+ unfolding diff(2) ** positive_integral_max_0
+ unfolding lebesgue_integral_def *
+ by (auto elim!: integrableE simp: field_simps)
qed
lemma (in measure_space) integral_add[simp, intro]:
@@ -1772,13 +1882,13 @@
and mono: "AE t. f t \<le> g t"
shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
proof -
- have "AE x. Real (f x) \<le> Real (g x)"
+ have "AE x. extreal (f x) \<le> extreal (g x)"
using mono by auto
- moreover have "AE x. Real (- g x) \<le> Real (- f x)"
+ moreover have "AE x. extreal (- g x) \<le> extreal (- f x)"
using mono by auto
ultimately show ?thesis using fg
- by (auto simp: lebesgue_integral_def integrable_def diff_minus
- intro!: add_mono real_of_pextreal_mono positive_integral_mono_AE)
+ by (auto intro!: add_mono positive_integral_mono_AE real_of_extreal_positive_mono
+ simp: positive_integral_positive lebesgue_integral_def diff_minus)
qed
lemma (in measure_space) integral_mono:
@@ -1795,20 +1905,21 @@
by auto
lemma (in measure_space) integral_indicator[simp, intro]:
- assumes "a \<in> sets M" and "\<mu> a \<noteq> \<omega>"
- shows "integral\<^isup>L M (indicator a) = real (\<mu> a)" (is ?int)
- and "integrable M (indicator a)" (is ?able)
+ assumes "A \<in> sets M" and "\<mu> A \<noteq> \<infinity>"
+ shows "integral\<^isup>L M (indicator A) = real (\<mu> A)" (is ?int)
+ and "integrable M (indicator A)" (is ?able)
proof -
- have *:
- "\<And>A x. Real (indicator A x) = indicator A x"
- "\<And>A x. Real (- indicator A x) = 0" unfolding indicator_def by auto
+ from `A \<in> sets M` have *:
+ "\<And>x. extreal (indicator A x) = indicator A x"
+ "(\<integral>\<^isup>+x. extreal (- indicator A x) \<partial>M) = 0"
+ by (auto split: split_indicator simp: positive_integral_0_iff_AE one_extreal_def)
show ?int ?able
using assms unfolding lebesgue_integral_def integrable_def
by (auto simp: * positive_integral_indicator borel_measurable_indicator)
qed
lemma (in measure_space) integral_cmul_indicator:
- assumes "A \<in> sets M" and "c \<noteq> 0 \<Longrightarrow> \<mu> A \<noteq> \<omega>"
+ assumes "A \<in> sets M" and "c \<noteq> 0 \<Longrightarrow> \<mu> A \<noteq> \<infinity>"
shows "integrable M (\<lambda>x. c * indicator A x)" (is ?P)
and "(\<integral>x. c * indicator A x \<partial>M) = c * real (\<mu> A)" (is ?I)
proof -
@@ -1840,15 +1951,11 @@
assumes "integrable M f"
shows "integrable M (\<lambda> x. \<bar>f x\<bar>)"
proof -
- have *:
- "\<And>x. Real \<bar>f x\<bar> = Real (f x) + Real (- f x)"
- "\<And>x. Real (- \<bar>f x\<bar>) = 0" by auto
- have abs: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M" and
- f: "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
- "(\<lambda>x. Real (- f x)) \<in> borel_measurable M"
- using assms unfolding integrable_def by auto
- from abs assms show ?thesis unfolding integrable_def *
- using positive_integral_linear[OF f, of 1] by simp
+ from assms have *: "(\<integral>\<^isup>+x. extreal (- \<bar>f x\<bar>)\<partial>M) = 0"
+ "\<And>x. extreal \<bar>f x\<bar> = max 0 (extreal (f x)) + max 0 (extreal (- f x))"
+ by (auto simp: integrable_def positive_integral_0_iff_AE split: split_max)
+ with assms show ?thesis
+ by (simp add: positive_integral_add positive_integral_max_0 integrable_def)
qed
lemma (in measure_space) integral_subalgebra:
@@ -1858,12 +1965,13 @@
and "integral\<^isup>L N f = integral\<^isup>L M f" (is ?I)
proof -
interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
- have "(\<lambda>x. Real (f x)) \<in> borel_measurable N" "(\<lambda>x. Real (- f x)) \<in> borel_measurable N"
- using borel by auto
- note * = this[THEN positive_integral_subalgebra[OF _ N sa]]
- have "f \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable N"
+ have "(\<integral>\<^isup>+ x. max 0 (extreal (f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (extreal (f x)) \<partial>M)"
+ "(\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>M)"
+ using borel by (auto intro!: positive_integral_subalgebra N sa)
+ moreover have "f \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable N"
using assms unfolding measurable_def by auto
- then show ?P ?I by (auto simp: * integrable_def lebesgue_integral_def)
+ ultimately show ?P ?I
+ by (auto simp: integrable_def lebesgue_integral_def positive_integral_max_0)
qed
lemma (in measure_space) integrable_bound:
@@ -1873,21 +1981,21 @@
assumes borel: "g \<in> borel_measurable M"
shows "integrable M g"
proof -
- have "(\<integral>\<^isup>+ x. Real (g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. Real \<bar>g x\<bar> \<partial>M)"
+ have "(\<integral>\<^isup>+ x. extreal (g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. extreal \<bar>g x\<bar> \<partial>M)"
by (auto intro!: positive_integral_mono)
- also have "\<dots> \<le> (\<integral>\<^isup>+ x. Real (f x) \<partial>M)"
+ also have "\<dots> \<le> (\<integral>\<^isup>+ x. extreal (f x) \<partial>M)"
using f by (auto intro!: positive_integral_mono)
- also have "\<dots> < \<omega>"
- using `integrable M f` unfolding integrable_def by (auto simp: pextreal_less_\<omega>)
- finally have pos: "(\<integral>\<^isup>+ x. Real (g x) \<partial>M) < \<omega>" .
+ also have "\<dots> < \<infinity>"
+ using `integrable M f` unfolding integrable_def by auto
+ finally have pos: "(\<integral>\<^isup>+ x. extreal (g x) \<partial>M) < \<infinity>" .
- have "(\<integral>\<^isup>+ x. Real (- g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. Real (\<bar>g x\<bar>) \<partial>M)"
+ have "(\<integral>\<^isup>+ x. extreal (- g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. extreal (\<bar>g x\<bar>) \<partial>M)"
by (auto intro!: positive_integral_mono)
- also have "\<dots> \<le> (\<integral>\<^isup>+ x. Real (f x) \<partial>M)"
+ also have "\<dots> \<le> (\<integral>\<^isup>+ x. extreal (f x) \<partial>M)"
using f by (auto intro!: positive_integral_mono)
- also have "\<dots> < \<omega>"
- using `integrable M f` unfolding integrable_def by (auto simp: pextreal_less_\<omega>)
- finally have neg: "(\<integral>\<^isup>+ x. Real (- g x) \<partial>M) < \<omega>" .
+ also have "\<dots> < \<infinity>"
+ using `integrable M f` unfolding integrable_def by auto
+ finally have neg: "(\<integral>\<^isup>+ x. extreal (- g x) \<partial>M) < \<infinity>" .
from neg pos borel show ?thesis
unfolding integrable_def by auto
@@ -1959,41 +2067,34 @@
by (simp add: mono_def incseq_def) }
note pos_u = this
- hence [simp]: "\<And>i x. Real (- f i x) = 0" "\<And>x. Real (- u x) = 0"
- using pos by auto
+ have SUP_F: "\<And>x. (SUP n. extreal (f n x)) = extreal (u x)"
+ unfolding SUP_eq_LIMSEQ[OF mono] by (rule lim)
- have SUP_F: "\<And>x. (SUP n. Real (f n x)) = Real (u x)"
- using mono pos pos_u lim by (rule SUP_eq_LIMSEQ[THEN iffD2])
-
- have borel_f: "\<And>i. (\<lambda>x. Real (f i x)) \<in> borel_measurable M"
+ have borel_f: "\<And>i. (\<lambda>x. extreal (f i x)) \<in> borel_measurable M"
using i unfolding integrable_def by auto
- hence "(\<lambda>x. SUP i. Real (f i x)) \<in> borel_measurable M"
+ hence "(\<lambda>x. SUP i. extreal (f i x)) \<in> borel_measurable M"
by auto
hence borel_u: "u \<in> borel_measurable M"
- using pos_u by (auto simp: borel_measurable_Real_eq SUP_F)
+ by (auto simp: borel_measurable_extreal_iff SUP_F)
- have integral_eq: "\<And>n. (\<integral>\<^isup>+ x. Real (f n x) \<partial>M) = Real (integral\<^isup>L M (f n))"
- using i unfolding lebesgue_integral_def integrable_def by (auto simp: Real_real)
+ hence [simp]: "\<And>i. (\<integral>\<^isup>+x. extreal (- f i x) \<partial>M) = 0" "(\<integral>\<^isup>+x. extreal (- u x) \<partial>M) = 0"
+ using i borel_u pos pos_u by (auto simp: positive_integral_0_iff_AE integrable_def)
+
+ have integral_eq: "\<And>n. (\<integral>\<^isup>+ x. extreal (f n x) \<partial>M) = extreal (integral\<^isup>L M (f n))"
+ using i positive_integral_positive by (auto simp: extreal_real lebesgue_integral_def integrable_def)
have pos_integral: "\<And>n. 0 \<le> integral\<^isup>L M (f n)"
using pos i by (auto simp: integral_positive)
hence "0 \<le> x"
using LIMSEQ_le_const[OF ilim, of 0] by auto
- have "(\<lambda>i. (\<integral>\<^isup>+ x. Real (f i x) \<partial>M)) \<up> (\<integral>\<^isup>+ x. Real (u x) \<partial>M)"
- proof (rule positive_integral_isoton)
- from SUP_F mono pos
- show "(\<lambda>i x. Real (f i x)) \<up> (\<lambda>x. Real (u x))"
- unfolding isoton_fun_expand by (auto simp: isoton_def mono_def)
- qed (rule borel_f)
- hence pI: "(\<integral>\<^isup>+ x. Real (u x) \<partial>M) = (SUP n. (\<integral>\<^isup>+ x. Real (f n x) \<partial>M))"
- unfolding isoton_def by simp
- also have "\<dots> = Real x" unfolding integral_eq
+ from mono pos i have pI: "(\<integral>\<^isup>+ x. extreal (u x) \<partial>M) = (SUP n. (\<integral>\<^isup>+ x. extreal (f n x) \<partial>M))"
+ by (auto intro!: positive_integral_monotone_convergence_SUP
+ simp: integrable_def incseq_mono incseq_Suc_iff le_fun_def SUP_F[symmetric])
+ also have "\<dots> = extreal x" unfolding integral_eq
proof (rule SUP_eq_LIMSEQ[THEN iffD2])
show "mono (\<lambda>n. integral\<^isup>L M (f n))"
using mono i by (auto simp: mono_def intro!: integral_mono)
- show "\<And>n. 0 \<le> integral\<^isup>L M (f n)" using pos_integral .
- show "0 \<le> x" using `0 \<le> x` .
show "(\<lambda>n. integral\<^isup>L M (f n)) ----> x" using ilim .
qed
finally show "integrable M u" "integral\<^isup>L M u = x" using borel_u `0 \<le> x`
@@ -2028,61 +2129,72 @@
assumes "integrable M f"
shows "(\<integral>x. \<bar>f x\<bar> \<partial>M) = 0 \<longleftrightarrow> \<mu> {x\<in>space M. f x \<noteq> 0} = 0"
proof -
- have *: "\<And>x. Real (- \<bar>f x\<bar>) = 0" by auto
+ have *: "(\<integral>\<^isup>+x. extreal (- \<bar>f x\<bar>) \<partial>M) = 0"
+ using assms by (auto simp: positive_integral_0_iff_AE integrable_def)
have "integrable M (\<lambda>x. \<bar>f x\<bar>)" using assms by (rule integrable_abs)
- hence "(\<lambda>x. Real (\<bar>f x\<bar>)) \<in> borel_measurable M"
- "(\<integral>\<^isup>+ x. Real \<bar>f x\<bar> \<partial>M) \<noteq> \<omega>" unfolding integrable_def by auto
+ hence "(\<lambda>x. extreal (\<bar>f x\<bar>)) \<in> borel_measurable M"
+ "(\<integral>\<^isup>+ x. extreal \<bar>f x\<bar> \<partial>M) \<noteq> \<infinity>" unfolding integrable_def by auto
from positive_integral_0_iff[OF this(1)] this(2)
show ?thesis unfolding lebesgue_integral_def *
- by (simp add: real_of_pextreal_eq_0)
+ using positive_integral_positive[of "\<lambda>x. extreal \<bar>f x\<bar>"]
+ by (auto simp add: real_of_extreal_eq_0)
qed
-lemma (in measure_space) positive_integral_omega:
- assumes "f \<in> borel_measurable M"
- and "integral\<^isup>P M f \<noteq> \<omega>"
- shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
+lemma (in measure_space) positive_integral_PInf:
+ assumes f: "f \<in> borel_measurable M"
+ and not_Inf: "integral\<^isup>P M f \<noteq> \<infinity>"
+ shows "\<mu> (f -` {\<infinity>} \<inter> space M) = 0"
proof -
- have "\<omega> * \<mu> (f -` {\<omega>} \<inter> space M) = (\<integral>\<^isup>+ x. \<omega> * indicator (f -` {\<omega>} \<inter> space M) x \<partial>M)"
- using positive_integral_cmult_indicator[OF borel_measurable_vimage, OF assms(1), of \<omega> \<omega>] by simp
- also have "\<dots> \<le> integral\<^isup>P M f"
- by (auto intro!: positive_integral_mono simp: indicator_def)
- finally show ?thesis
- using assms(2) by (cases ?thesis) auto
+ have "\<infinity> * \<mu> (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^isup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
+ using f by (subst positive_integral_cmult_indicator) (auto simp: measurable_sets)
+ also have "\<dots> \<le> integral\<^isup>P M (\<lambda>x. max 0 (f x))"
+ by (auto intro!: positive_integral_mono simp: indicator_def max_def)
+ finally have "\<infinity> * \<mu> (f -` {\<infinity>} \<inter> space M) \<le> integral\<^isup>P M f"
+ by (simp add: positive_integral_max_0)
+ moreover have "0 \<le> \<mu> (f -` {\<infinity>} \<inter> space M)"
+ using f by (simp add: measurable_sets)
+ ultimately show ?thesis
+ using assms by (auto split: split_if_asm)
qed
-lemma (in measure_space) positive_integral_omega_AE:
- assumes "f \<in> borel_measurable M" "integral\<^isup>P M f \<noteq> \<omega>" shows "AE x. f x \<noteq> \<omega>"
+lemma (in measure_space) positive_integral_PInf_AE:
+ assumes "f \<in> borel_measurable M" "integral\<^isup>P M f \<noteq> \<infinity>" shows "AE x. f x \<noteq> \<infinity>"
proof (rule AE_I)
- show "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
- by (rule positive_integral_omega[OF assms])
- show "f -` {\<omega>} \<inter> space M \<in> sets M"
+ show "\<mu> (f -` {\<infinity>} \<inter> space M) = 0"
+ by (rule positive_integral_PInf[OF assms])
+ show "f -` {\<infinity>} \<inter> space M \<in> sets M"
using assms by (auto intro: borel_measurable_vimage)
qed auto
-lemma (in measure_space) simple_integral_omega:
- assumes "simple_function M f"
- and "integral\<^isup>S M f \<noteq> \<omega>"
- shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
-proof (rule positive_integral_omega)
+lemma (in measure_space) simple_integral_PInf:
+ assumes "simple_function M f" "\<And>x. 0 \<le> f x"
+ and "integral\<^isup>S M f \<noteq> \<infinity>"
+ shows "\<mu> (f -` {\<infinity>} \<inter> space M) = 0"
+proof (rule positive_integral_PInf)
show "f \<in> borel_measurable M" using assms by (auto intro: borel_measurable_simple_function)
- show "integral\<^isup>P M f \<noteq> \<omega>"
+ show "integral\<^isup>P M f \<noteq> \<infinity>"
using assms by (simp add: positive_integral_eq_simple_integral)
qed
lemma (in measure_space) integral_real:
- fixes f :: "'a \<Rightarrow> pextreal"
- assumes [simp]: "AE x. f x \<noteq> \<omega>"
- shows "(\<integral>x. real (f x) \<partial>M) = real (integral\<^isup>P M f)" (is ?plus)
- and "(\<integral>x. - real (f x) \<partial>M) = - real (integral\<^isup>P M f)" (is ?minus)
-proof -
- have "(\<integral>\<^isup>+ x. Real (real (f x)) \<partial>M) = integral\<^isup>P M f"
- by (auto intro!: positive_integral_cong_AE simp: Real_real)
- moreover
- have "(\<integral>\<^isup>+ x. Real (- real (f x)) \<partial>M) = (\<integral>\<^isup>+ x. 0 \<partial>M)"
- by (intro positive_integral_cong) auto
- ultimately show ?plus ?minus
- by (auto simp: lebesgue_integral_def integrable_def)
-qed
+ "AE x. \<bar>f x\<bar> \<noteq> \<infinity> \<Longrightarrow> (\<integral>x. real (f x) \<partial>M) = real (integral\<^isup>P M f) - real (integral\<^isup>P M (\<lambda>x. - f x))"
+ using assms unfolding lebesgue_integral_def
+ by (subst (1 2) positive_integral_cong_AE) (auto simp add: extreal_real)
+
+lemma liminf_extreal_cminus:
+ fixes f :: "nat \<Rightarrow> extreal" assumes "c \<noteq> -\<infinity>"
+ shows "liminf (\<lambda>x. c - f x) = c - limsup f"
+proof (cases c)
+ case PInf then show ?thesis by (simp add: Liminf_const)
+next
+ case (real r) then show ?thesis
+ unfolding liminf_SUPR_INFI limsup_INFI_SUPR
+ apply (subst INFI_extreal_cminus)
+ apply auto
+ apply (subst SUPR_extreal_cminus)
+ apply auto
+ done
+qed (insert `c \<noteq> -\<infinity>`, simp)
lemma (in measure_space) integral_dominated_convergence:
assumes u: "\<And>i. integrable M (u i)" and bound: "\<And>x j. x\<in>space M \<Longrightarrow> \<bar>u j x\<bar> \<le> w x"
@@ -2129,61 +2241,76 @@
finally have "\<bar>u j x - u' x\<bar> \<le> 2 * w x" by simp }
note diff_less_2w = this
- have PI_diff: "\<And>m n. (\<integral>\<^isup>+ x. Real (?diff (m + n) x) \<partial>M) =
- (\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) - (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M)"
+ have PI_diff: "\<And>n. (\<integral>\<^isup>+ x. extreal (?diff n x) \<partial>M) =
+ (\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M) - (\<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"
using diff w diff_less_2w w_pos
by (subst positive_integral_diff[symmetric])
(auto simp: integrable_def intro!: positive_integral_cong)
have "integrable M (\<lambda>x. 2 * w x)"
using w by (auto intro: integral_cmult)
- hence I2w_fin: "(\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) \<noteq> \<omega>" and
- borel_2w: "(\<lambda>x. Real (2 * w x)) \<in> borel_measurable M"
+ hence I2w_fin: "(\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M) \<noteq> \<infinity>" and
+ borel_2w: "(\<lambda>x. extreal (2 * w x)) \<in> borel_measurable M"
unfolding integrable_def by auto
- have "(INF n. SUP m. (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M)) = 0" (is "?lim_SUP = 0")
+ have "limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M) = 0" (is "limsup ?f = 0")
proof cases
- assume eq_0: "(\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) = 0"
- have "\<And>i. (\<integral>\<^isup>+ x. Real \<bar>u i x - u' x\<bar> \<partial>M) \<le> (\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M)"
- proof (rule positive_integral_mono)
- fix i x assume "x \<in> space M" from diff_less_2w[OF this, of i]
- show "Real \<bar>u i x - u' x\<bar> \<le> Real (2 * w x)" by auto
- qed
- hence "\<And>i. (\<integral>\<^isup>+ x. Real \<bar>u i x - u' x\<bar> \<partial>M) = 0" using eq_0 by auto
- thus ?thesis by simp
+ assume eq_0: "(\<integral>\<^isup>+ x. max 0 (extreal (2 * w x)) \<partial>M) = 0" (is "?wx = 0")
+ { fix n
+ have "?f n \<le> ?wx" (is "integral\<^isup>P M ?f' \<le> _")
+ using diff_less_2w[of _ n] unfolding positive_integral_max_0
+ by (intro positive_integral_mono) auto
+ then have "?f n = 0"
+ using positive_integral_positive[of ?f'] eq_0 by auto }
+ then show ?thesis by (simp add: Limsup_const)
next
- assume neq_0: "(\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) \<noteq> 0"
- have "(\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) = (\<integral>\<^isup>+ x. (SUP n. INF m. Real (?diff (m + n) x)) \<partial>M)"
- proof (rule positive_integral_cong, subst add_commute)
+ assume neq_0: "(\<integral>\<^isup>+ x. max 0 (extreal (2 * w x)) \<partial>M) \<noteq> 0" (is "?wx \<noteq> 0")
+ have "0 = limsup (\<lambda>n. 0 :: extreal)" by (simp add: Limsup_const)
+ also have "\<dots> \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"
+ by (intro limsup_mono positive_integral_positive)
+ finally have pos: "0 \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)" .
+ have "?wx = (\<integral>\<^isup>+ x. liminf (\<lambda>n. max 0 (extreal (?diff n x))) \<partial>M)"
+ proof (rule positive_integral_cong)
fix x assume x: "x \<in> space M"
- show "Real (2 * w x) = (SUP n. INF m. Real (?diff (n + m) x))"
- proof (rule LIMSEQ_imp_lim_INF[symmetric])
- fix j show "0 \<le> ?diff j x" using diff_less_2w[OF x, of j] by simp
- next
+ show "max 0 (extreal (2 * w x)) = liminf (\<lambda>n. max 0 (extreal (?diff n x)))"
+ unfolding extreal_max_0
+ proof (rule lim_imp_Liminf[symmetric], unfold lim_extreal)
have "(\<lambda>i. ?diff i x) ----> 2 * w x - \<bar>u' x - u' x\<bar>"
using u'[OF x] by (safe intro!: LIMSEQ_diff LIMSEQ_const LIMSEQ_imp_rabs)
- thus "(\<lambda>i. ?diff i x) ----> 2 * w x" by simp
- qed
+ then show "(\<lambda>i. max 0 (?diff i x)) ----> max 0 (2 * w x)"
+ by (auto intro!: tendsto_real_max simp add: lim_extreal)
+ qed (rule trivial_limit_sequentially)
qed
- also have "\<dots> \<le> (SUP n. INF m. (\<integral>\<^isup>+ x. Real (?diff (m + n) x) \<partial>M))"
+ also have "\<dots> \<le> liminf (\<lambda>n. \<integral>\<^isup>+ x. max 0 (extreal (?diff n x)) \<partial>M)"
using u'_borel w u unfolding integrable_def
- by (auto intro!: positive_integral_lim_INF)
- also have "\<dots> = (\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) -
- (INF n. SUP m. \<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M)"
- unfolding PI_diff pextreal_INF_minus[OF I2w_fin] pextreal_SUP_minus ..
- finally show ?thesis using neq_0 I2w_fin by (rule pextreal_le_minus_imp_0)
+ by (intro positive_integral_lim_INF) (auto intro!: positive_integral_lim_INF)
+ also have "\<dots> = (\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M) -
+ limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"
+ unfolding PI_diff positive_integral_max_0
+ using positive_integral_positive[of "\<lambda>x. extreal (2 * w x)"]
+ by (subst liminf_extreal_cminus) auto
+ finally show ?thesis
+ using neq_0 I2w_fin positive_integral_positive[of "\<lambda>x. extreal (2 * w x)"] pos
+ unfolding positive_integral_max_0
+ by (cases rule: extreal2_cases[of "\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M" "limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"])
+ auto
qed
- have [simp]: "\<And>n m x. Real (- \<bar>u (m + n) x - u' x\<bar>) = 0" by auto
-
- have [simp]: "\<And>n m. (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M) =
- Real ((\<integral>x. \<bar>u (n + m) x - u' x\<bar> \<partial>M))"
- using diff by (subst add_commute) (simp add: lebesgue_integral_def integrable_def Real_real)
-
- have "(SUP n. INF m. (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M)) \<le> ?lim_SUP"
- (is "?lim_INF \<le> _") by (subst (1 2) add_commute) (rule lim_INF_le_lim_SUP)
- hence "?lim_INF = Real 0" "?lim_SUP = Real 0" using `?lim_SUP = 0` by auto
- thus ?lim_diff using diff by (auto intro!: integral_positive lim_INF_eq_lim_SUP)
+ have "liminf ?f \<le> limsup ?f"
+ by (intro extreal_Liminf_le_Limsup trivial_limit_sequentially)
+ moreover
+ { have "0 = liminf (\<lambda>n. 0 :: extreal)" by (simp add: Liminf_const)
+ also have "\<dots> \<le> liminf ?f"
+ by (intro liminf_mono positive_integral_positive)
+ finally have "0 \<le> liminf ?f" . }
+ ultimately have liminf_limsup_eq: "liminf ?f = extreal 0" "limsup ?f = extreal 0"
+ using `limsup ?f = 0` by auto
+ have "\<And>n. (\<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M) = extreal (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M)"
+ using diff positive_integral_positive
+ by (subst integral_eq_positive_integral) (auto simp: extreal_real integrable_def)
+ then show ?lim_diff
+ using extreal_Liminf_eq_Limsup[OF trivial_limit_sequentially liminf_limsup_eq]
+ by (simp add: lim_extreal)
show ?lim
proof (rule LIMSEQ_I)
@@ -2266,7 +2393,7 @@
assumes f: "f \<in> borel_measurable M"
and bij: "bij_betw enum S (f ` space M)"
and enum_zero: "enum ` (-S) \<subseteq> {0}"
- and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<omega>"
+ and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<infinity>"
and abs_summable: "summable (\<lambda>r. \<bar>enum r * real (\<mu> (f -` {enum r} \<inter> space M))\<bar>)"
shows "integrable M f"
and "(\<lambda>r. enum r * real (\<mu> (f -` {enum r} \<inter> space M))) sums integral\<^isup>L M f" (is ?sums)
@@ -2303,7 +2430,7 @@
also have "\<dots> = \<bar>enum r\<bar> * real (\<mu> (?A r))"
using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
finally have "(\<integral>x. \<bar>?F r x\<bar> \<partial>M) = \<bar>enum r * real (\<mu> (?A r))\<bar>"
- by (simp add: abs_mult_pos real_pextreal_pos) }
+ using f by (subst (2) abs_mult_pos[symmetric]) (auto intro!: real_of_extreal_pos measurable_sets) }
note int_abs_F = this
have 1: "\<And>i. integrable M (\<lambda>x. ?F i x)"
@@ -2323,7 +2450,7 @@
lemma (in measure_space) integral_on_finite:
assumes f: "f \<in> borel_measurable M" and finite: "finite (f`space M)"
- and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<omega>"
+ and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<infinity>"
shows "integrable M f"
and "(\<integral>x. f x \<partial>M) =
(\<Sum> r \<in> f`space M. r * real (\<mu> (f -` {r} \<inter> space M)))" (is "?integral")
@@ -2353,11 +2480,12 @@
by (auto intro: borel_measurable_simple_function)
lemma (in finite_measure_space) positive_integral_finite_eq_setsum:
- "integral\<^isup>P M f = (\<Sum>x \<in> space M. f x * \<mu> {x})"
+ assumes pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
+ shows "integral\<^isup>P M f = (\<Sum>x \<in> space M. f x * \<mu> {x})"
proof -
have *: "integral\<^isup>P M f = (\<integral>\<^isup>+ x. (\<Sum>y\<in>space M. f y * indicator {y} x) \<partial>M)"
by (auto intro!: positive_integral_cong simp add: indicator_def if_distrib setsum_cases[OF finite_space])
- show ?thesis unfolding * using borel_measurable_finite[of f]
+ show ?thesis unfolding * using borel_measurable_finite[of f] pos
by (simp add: positive_integral_setsum positive_integral_cmult_indicator)
qed
@@ -2365,16 +2493,20 @@
shows "integrable M f"
and "integral\<^isup>L M f = (\<Sum>x \<in> space M. f x * real (\<mu> {x}))" (is ?I)
proof -
- have [simp]:
- "(\<integral>\<^isup>+ x. Real (f x) \<partial>M) = (\<Sum>x \<in> space M. Real (f x) * \<mu> {x})"
- "(\<integral>\<^isup>+ x. Real (- f x) \<partial>M) = (\<Sum>x \<in> space M. Real (- f x) * \<mu> {x})"
- unfolding positive_integral_finite_eq_setsum by auto
- show "integrable M f" using finite_space finite_measure
- by (simp add: setsum_\<omega> integrable_def)
- show ?I using finite_measure
- apply (simp add: lebesgue_integral_def real_of_pextreal_setsum[symmetric]
- real_of_pextreal_mult[symmetric] setsum_subtractf[symmetric])
- by (rule setsum_cong) (simp_all split: split_if)
+ have *:
+ "(\<integral>\<^isup>+ x. max 0 (extreal (f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (extreal (f x)) * \<mu> {x})"
+ "(\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (extreal (- f x)) * \<mu> {x})"
+ by (simp_all add: positive_integral_finite_eq_setsum)
+ then show "integrable M f" using finite_space finite_measure
+ by (simp add: setsum_Pinfty integrable_def positive_integral_max_0
+ split: split_max)
+ show ?I using finite_measure *
+ apply (simp add: positive_integral_max_0 lebesgue_integral_def)
+ apply (subst (1 2) setsum_real_of_extreal[symmetric])
+ apply (simp_all split: split_max add: setsum_subtractf[symmetric])
+ apply (intro setsum_cong[OF refl])
+ apply (simp split: split_max)
+ done
qed
end
--- a/src/HOL/Probability/Lebesgue_Measure.thy Mon Mar 14 14:37:47 2011 +0100
+++ b/src/HOL/Probability/Lebesgue_Measure.thy Mon Mar 14 14:37:49 2011 +0100
@@ -48,12 +48,12 @@
lemma Pi_iff: "f \<in> Pi I X \<longleftrightarrow> (\<forall>i\<in>I. f i \<in> X i)"
unfolding Pi_def by auto
-subsection {* Lebesgue measure *}
+subsection {* Lebesgue measure *}
definition lebesgue :: "'a::ordered_euclidean_space measure_space" where
"lebesgue = \<lparr> space = UNIV,
sets = {A. \<forall>n. (indicator A :: 'a \<Rightarrow> real) integrable_on cube n},
- measure = \<lambda>A. SUP n. Real (integral (cube n) (indicator A)) \<rparr>"
+ measure = \<lambda>A. SUP n. extreal (integral (cube n) (indicator A)) \<rparr>"
lemma space_lebesgue[simp]: "space lebesgue = UNIV"
unfolding lebesgue_def by simp
@@ -114,10 +114,33 @@
qed (auto intro: LIMSEQ_indicator_UN simp: cube_def)
qed simp
+lemma suminf_SUP_eq:
+ fixes f :: "nat \<Rightarrow> nat \<Rightarrow> extreal"
+ assumes "\<And>i. incseq (\<lambda>n. f n i)" "\<And>n i. 0 \<le> f n i"
+ shows "(\<Sum>i. SUP n. f n i) = (SUP n. \<Sum>i. f n i)"
+proof -
+ { fix n :: nat
+ have "(\<Sum>i<n. SUP k. f k i) = (SUP k. \<Sum>i<n. f k i)"
+ using assms by (auto intro!: SUPR_extreal_setsum[symmetric]) }
+ note * = this
+ show ?thesis using assms
+ apply (subst (1 2) suminf_extreal_eq_SUPR)
+ unfolding *
+ apply (auto intro!: le_SUPI2)
+ apply (subst SUP_commute) ..
+qed
+
interpretation lebesgue: measure_space lebesgue
proof
have *: "indicator {} = (\<lambda>x. 0 :: real)" by (simp add: fun_eq_iff)
- show "measure lebesgue {} = 0" by (simp add: integral_0 * lebesgue_def)
+ show "positive lebesgue (measure lebesgue)"
+ proof (unfold positive_def, safe)
+ show "measure lebesgue {} = 0" by (simp add: integral_0 * lebesgue_def)
+ fix A assume "A \<in> sets lebesgue"
+ then show "0 \<le> measure lebesgue A"
+ unfolding lebesgue_def
+ by (auto intro!: le_SUPI2 integral_nonneg)
+ qed
next
show "countably_additive lebesgue (measure lebesgue)"
proof (intro countably_additive_def[THEN iffD2] allI impI)
@@ -130,23 +153,17 @@
assume "(\<Union>i. A i) \<in> sets lebesgue"
then have UN_A[simp, intro]: "\<And>i n. (indicator (\<Union>i. A i) :: _ \<Rightarrow> real) integrable_on cube n"
by (auto dest: lebesgueD)
- show "(\<Sum>\<^isub>\<infinity>n. measure lebesgue (A n)) = measure lebesgue (\<Union>i. A i)"
- proof (simp add: lebesgue_def, subst psuminf_SUP_eq)
- fix n i show "Real (?m n i) \<le> Real (?m (Suc n) i)"
- using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le)
+ show "(\<Sum>n. measure lebesgue (A n)) = measure lebesgue (\<Union>i. A i)"
+ proof (simp add: lebesgue_def, subst suminf_SUP_eq, safe intro!: incseq_SucI)
+ fix i n show "extreal (?m n i) \<le> extreal (?m (Suc n) i)"
+ using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le incseq_SucI)
next
- show "(SUP n. \<Sum>\<^isub>\<infinity>i. Real (?m n i)) = (SUP n. Real (?M n UNIV))"
- unfolding psuminf_def
- proof (subst setsum_Real, (intro arg_cong[where f="SUPR UNIV"] ext ballI nn SUP_eq_LIMSEQ[THEN iffD2])+)
- fix n :: nat show "mono (\<lambda>m. \<Sum>x<m. ?m n x)"
- proof (intro mono_iff_le_Suc[THEN iffD2] allI)
- fix m show "(\<Sum>x<m. ?m n x) \<le> (\<Sum>x<Suc m. ?m n x)"
- using nn[of n m] by auto
- qed
- show "0 \<le> ?M n UNIV"
- using UN_A by (auto intro!: integral_nonneg)
- fix m show "0 \<le> (\<Sum>x<m. ?m n x)" by (auto intro!: setsum_nonneg)
- next
+ fix i n show "0 \<le> extreal (?m n i)"
+ using rA unfolding lebesgue_def
+ by (auto intro!: le_SUPI2 integral_nonneg)
+ next
+ show "(SUP n. \<Sum>i. extreal (?m n i)) = (SUP n. extreal (?M n UNIV))"
+ proof (intro arg_cong[where f="SUPR UNIV"] ext sums_unique[symmetric] sums_extreal[THEN iffD2] sums_def[THEN iffD2])
fix n
have "\<And>m. (UNION {..<m} A) \<in> sets lebesgue" using rA by auto
from lebesgueD[OF this]
@@ -171,8 +188,8 @@
ultimately show ?case
using Suc A by (simp add: integral_add[symmetric])
qed auto }
- ultimately show "(\<lambda>m. \<Sum>x<m. ?m n x) ----> ?M n UNIV"
- by simp
+ ultimately show "(\<lambda>m. \<Sum>x = 0..<m. ?m n x) ----> ?M n UNIV"
+ by (simp add: atLeast0LessThan)
qed
qed
qed
@@ -232,13 +249,11 @@
lemma lmeasure_iff_LIMSEQ:
assumes "A \<in> sets lebesgue" "0 \<le> m"
- shows "lebesgue.\<mu> A = Real m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m"
+ shows "lebesgue.\<mu> A = extreal m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m"
proof (simp add: lebesgue_def, intro SUP_eq_LIMSEQ)
show "mono (\<lambda>n. integral (cube n) (indicator A::_=>real))"
using cube_subset assms by (intro monoI integral_subset_le) (auto dest!: lebesgueD)
- fix n show "0 \<le> integral (cube n) (indicator A::_=>real)"
- using assms by (auto dest!: lebesgueD intro!: integral_nonneg)
-qed fact
+qed
lemma has_integral_indicator_UNIV:
fixes s A :: "'a::ordered_euclidean_space set" and x :: real
@@ -260,7 +275,7 @@
lemma lmeasure_finite_has_integral:
fixes s :: "'a::ordered_euclidean_space set"
- assumes "s \<in> sets lebesgue" "lebesgue.\<mu> s = Real m" "0 \<le> m"
+ assumes "s \<in> sets lebesgue" "lebesgue.\<mu> s = extreal m" "0 \<le> m"
shows "(indicator s has_integral m) UNIV"
proof -
let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
@@ -302,12 +317,14 @@
unfolding m by (intro integrable_integral **)
qed
-lemma lmeasure_finite_integrable: assumes "s \<in> sets lebesgue" "lebesgue.\<mu> s \<noteq> \<omega>"
+lemma lmeasure_finite_integrable: assumes s: "s \<in> sets lebesgue" and "lebesgue.\<mu> s \<noteq> \<infinity>"
shows "(indicator s :: _ \<Rightarrow> real) integrable_on UNIV"
proof (cases "lebesgue.\<mu> s")
- case (preal m) from lmeasure_finite_has_integral[OF `s \<in> sets lebesgue` this]
+ case (real m)
+ with lmeasure_finite_has_integral[OF `s \<in> sets lebesgue` this]
+ lebesgue.positive_measure[OF s]
show ?thesis unfolding integrable_on_def by auto
-qed (insert assms, auto)
+qed (insert assms lebesgue.positive_measure[OF s], auto)
lemma has_integral_lebesgue: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
shows "s \<in> sets lebesgue"
@@ -321,7 +338,7 @@
qed
lemma has_integral_lmeasure: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
- shows "lebesgue.\<mu> s = Real m"
+ shows "lebesgue.\<mu> s = extreal m"
proof (intro lmeasure_iff_LIMSEQ[THEN iffD2])
let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
show "s \<in> sets lebesgue" using has_integral_lebesgue[OF assms] .
@@ -346,28 +363,28 @@
qed
lemma has_integral_iff_lmeasure:
- "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = Real m)"
+ "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = extreal m)"
proof
assume "(indicator A has_integral m) UNIV"
with has_integral_lmeasure[OF this] has_integral_lebesgue[OF this]
- show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = Real m"
+ show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = extreal m"
by (auto intro: has_integral_nonneg)
next
- assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = Real m"
+ assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = extreal m"
then show "(indicator A has_integral m) UNIV" by (intro lmeasure_finite_has_integral) auto
qed
lemma lmeasure_eq_integral: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV"
- shows "lebesgue.\<mu> s = Real (integral UNIV (indicator s))"
+ shows "lebesgue.\<mu> s = extreal (integral UNIV (indicator s))"
using assms unfolding integrable_on_def
proof safe
fix y :: real assume "(indicator s has_integral y) UNIV"
from this[unfolded has_integral_iff_lmeasure] integral_unique[OF this]
- show "lebesgue.\<mu> s = Real (integral UNIV (indicator s))" by simp
+ show "lebesgue.\<mu> s = extreal (integral UNIV (indicator s))" by simp
qed
lemma lebesgue_simple_function_indicator:
- fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
+ fixes f::"'a::ordered_euclidean_space \<Rightarrow> extreal"
assumes f:"simple_function lebesgue f"
shows "f = (\<lambda>x. (\<Sum>y \<in> f ` UNIV. y * indicator (f -` {y}) x))"
by (rule, subst lebesgue.simple_function_indicator_representation[OF f]) auto
@@ -376,7 +393,7 @@
"(indicator s::_\<Rightarrow>real) integrable_on UNIV \<Longrightarrow> integral UNIV (indicator s) = real (lebesgue.\<mu> s)"
by (subst lmeasure_eq_integral) (auto intro!: integral_nonneg)
-lemma lmeasure_finite: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV" shows "lebesgue.\<mu> s \<noteq> \<omega>"
+lemma lmeasure_finite: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV" shows "lebesgue.\<mu> s \<noteq> \<infinity>"
using lmeasure_eq_integral[OF assms] by auto
lemma negligible_iff_lebesgue_null_sets:
@@ -409,37 +426,29 @@
shows "integral {a .. b} (\<lambda>x. c) = content {a .. b} *\<^sub>R c"
by (rule integral_unique) (rule has_integral_const)
-lemma lmeasure_UNIV[intro]: "lebesgue.\<mu> (UNIV::'a::ordered_euclidean_space set) = \<omega>"
-proof (simp add: lebesgue_def SUP_\<omega>, intro allI impI)
- fix x assume "x < \<omega>"
- then obtain r where r: "x = Real r" "0 \<le> r" by (cases x) auto
- then obtain n where n: "r < of_nat n" using ex_less_of_nat by auto
- show "\<exists>i. x < Real (integral (cube i) (indicator UNIV::'a\<Rightarrow>real))"
- proof (intro exI[of _ n])
- have [simp]: "indicator UNIV = (\<lambda>x. 1)" by (auto simp: fun_eq_iff)
- { fix m :: nat assume "0 < m" then have "real n \<le> (\<Prod>x<m. 2 * real n)"
- proof (induct m)
- case (Suc m)
- show ?case
- proof cases
- assume "m = 0" then show ?thesis by (simp add: lessThan_Suc)
- next
- assume "m \<noteq> 0" then have "real n \<le> (\<Prod>x<m. 2 * real n)" using Suc by auto
- then show ?thesis
- by (auto simp: lessThan_Suc field_simps mult_le_cancel_left1)
- qed
- qed auto } note this[OF DIM_positive[where 'a='a], simp]
- then have [simp]: "0 \<le> (\<Prod>x<DIM('a). 2 * real n)" using real_of_nat_ge_zero by arith
- have "x < Real (of_nat n)" using n r by auto
- also have "Real (of_nat n) \<le> Real (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))"
- by (auto simp: real_eq_of_nat[symmetric] cube_def content_closed_interval_cases)
- finally show "x < Real (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))" .
- qed
-qed
+lemma lmeasure_UNIV[intro]: "lebesgue.\<mu> (UNIV::'a::ordered_euclidean_space set) = \<infinity>"
+proof (simp add: lebesgue_def, intro SUP_PInfty bexI)
+ fix n :: nat
+ have "indicator UNIV = (\<lambda>x::'a. 1 :: real)" by auto
+ moreover
+ { have "real n \<le> (2 * real n) ^ DIM('a)"
+ proof (cases n)
+ case 0 then show ?thesis by auto
+ next
+ case (Suc n')
+ have "real n \<le> (2 * real n)^1" by auto
+ also have "(2 * real n)^1 \<le> (2 * real n) ^ DIM('a)"
+ using Suc DIM_positive[where 'a='a] by (intro power_increasing) (auto simp: real_of_nat_Suc)
+ finally show ?thesis .
+ qed }
+ ultimately show "extreal (real n) \<le> extreal (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))"
+ using integral_const DIM_positive[where 'a='a]
+ by (auto simp: cube_def content_closed_interval_cases setprod_constant)
+qed simp
lemma
fixes a b ::"'a::ordered_euclidean_space"
- shows lmeasure_atLeastAtMost[simp]: "lebesgue.\<mu> {a..b} = Real (content {a..b})"
+ shows lmeasure_atLeastAtMost[simp]: "lebesgue.\<mu> {a..b} = extreal (content {a..b})"
proof -
have "(indicator (UNIV \<inter> {a..b})::_\<Rightarrow>real) integrable_on UNIV"
unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def_raw)
@@ -467,7 +476,7 @@
lemma
fixes a b :: real
shows lmeasure_real_greaterThanAtMost[simp]:
- "lebesgue.\<mu> {a <.. b} = Real (if a \<le> b then b - a else 0)"
+ "lebesgue.\<mu> {a <.. b} = extreal (if a \<le> b then b - a else 0)"
proof cases
assume "a < b"
then have "lebesgue.\<mu> {a <.. b} = lebesgue.\<mu> {a .. b} - lebesgue.\<mu> {a}"
@@ -479,7 +488,7 @@
lemma
fixes a b :: real
shows lmeasure_real_atLeastLessThan[simp]:
- "lebesgue.\<mu> {a ..< b} = Real (if a \<le> b then b - a else 0)"
+ "lebesgue.\<mu> {a ..< b} = extreal (if a \<le> b then b - a else 0)"
proof cases
assume "a < b"
then have "lebesgue.\<mu> {a ..< b} = lebesgue.\<mu> {a .. b} - lebesgue.\<mu> {b}"
@@ -491,7 +500,7 @@
lemma
fixes a b :: real
shows lmeasure_real_greaterThanLessThan[simp]:
- "lebesgue.\<mu> {a <..< b} = Real (if a \<le> b then b - a else 0)"
+ "lebesgue.\<mu> {a <..< b} = extreal (if a \<le> b then b - a else 0)"
proof cases
assume "a < b"
then have "lebesgue.\<mu> {a <..< b} = lebesgue.\<mu> {a <.. b} - lebesgue.\<mu> {b}"
@@ -511,19 +520,16 @@
and measurable_lborel[simp]: "measurable lborel = measurable borel"
by (simp_all add: measurable_def_raw lborel_def)
-interpretation lborel: measure_space lborel
+interpretation lborel: measure_space "lborel :: ('a::ordered_euclidean_space) measure_space"
where "space lborel = UNIV"
and "sets lborel = sets borel"
and "measure lborel = lebesgue.\<mu>"
and "measurable lborel = measurable borel"
-proof -
- show "measure_space lborel"
- proof
- show "countably_additive lborel (measure lborel)"
- using lebesgue.ca unfolding countably_additive_def lborel_def
- apply safe apply (erule_tac x=A in allE) by auto
- qed (auto simp: lborel_def)
-qed simp_all
+proof (rule lebesgue.measure_space_subalgebra)
+ have "sigma_algebra (lborel::'a measure_space) \<longleftrightarrow> sigma_algebra (borel::'a algebra)"
+ unfolding sigma_algebra_iff2 lborel_def by simp
+ then show "sigma_algebra (lborel::'a measure_space)" by simp default
+qed auto
interpretation lborel: sigma_finite_measure lborel
where "space lborel = UNIV"
@@ -536,7 +542,7 @@
show "range cube \<subseteq> sets lborel" by (auto intro: borel_closed)
{ fix x have "\<exists>n. x\<in>cube n" using mem_big_cube by auto }
thus "(\<Union>i. cube i) = space lborel" by auto
- show "\<forall>i. measure lborel (cube i) \<noteq> \<omega>" by (simp add: cube_def)
+ show "\<forall>i. measure lborel (cube i) \<noteq> \<infinity>" by (simp add: cube_def)
qed
qed simp_all
@@ -544,171 +550,221 @@
proof
from lborel.sigma_finite guess A ..
moreover then have "range A \<subseteq> sets lebesgue" using lebesgueI_borel by blast
- ultimately show "\<exists>A::nat \<Rightarrow> 'b set. range A \<subseteq> sets lebesgue \<and> (\<Union>i. A i) = space lebesgue \<and> (\<forall>i. lebesgue.\<mu> (A i) \<noteq> \<omega>)"
+ ultimately show "\<exists>A::nat \<Rightarrow> 'b set. range A \<subseteq> sets lebesgue \<and> (\<Union>i. A i) = space lebesgue \<and> (\<forall>i. lebesgue.\<mu> (A i) \<noteq> \<infinity>)"
by auto
qed
subsection {* Lebesgue integrable implies Gauge integrable *}
+lemma positive_not_Inf:
+ "0 \<le> x \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> \<bar>x\<bar> \<noteq> \<infinity>"
+ by (cases x) auto
+
+lemma has_integral_cmult_real:
+ fixes c :: real
+ assumes "c \<noteq> 0 \<Longrightarrow> (f has_integral x) A"
+ shows "((\<lambda>x. c * f x) has_integral c * x) A"
+proof cases
+ assume "c \<noteq> 0"
+ from has_integral_cmul[OF assms[OF this], of c] show ?thesis
+ unfolding real_scaleR_def .
+qed simp
+
lemma simple_function_has_integral:
- fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
+ fixes f::"'a::ordered_euclidean_space \<Rightarrow> extreal"
assumes f:"simple_function lebesgue f"
- and f':"\<forall>x. f x \<noteq> \<omega>"
- and om:"\<forall>x\<in>range f. lebesgue.\<mu> (f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
+ and f':"range f \<subseteq> {0..<\<infinity>}"
+ and om:"\<And>x. x \<in> range f \<Longrightarrow> lebesgue.\<mu> (f -` {x} \<inter> UNIV) = \<infinity> \<Longrightarrow> x = 0"
shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV"
- unfolding simple_integral_def
- apply(subst lebesgue_simple_function_indicator[OF f])
-proof -
- case goal1
- have *:"\<And>x. \<forall>y\<in>range f. y * indicator (f -` {y}) x \<noteq> \<omega>"
- "\<forall>x\<in>range f. x * lebesgue.\<mu> (f -` {x} \<inter> UNIV) \<noteq> \<omega>"
- using f' om unfolding indicator_def by auto
- show ?case unfolding space_lebesgue real_of_pextreal_setsum'[OF *(1),THEN sym]
- unfolding real_of_pextreal_setsum'[OF *(2),THEN sym]
- unfolding real_of_pextreal_setsum space_lebesgue
- apply(rule has_integral_setsum)
- proof safe show "finite (range f)" using f by (auto dest: lebesgue.simple_functionD)
- fix y::'a show "((\<lambda>x. real (f y * indicator (f -` {f y}) x)) has_integral
- real (f y * lebesgue.\<mu> (f -` {f y} \<inter> UNIV))) UNIV"
- proof(cases "f y = 0") case False
- have mea:"(indicator (f -` {f y}) ::_\<Rightarrow>real) integrable_on UNIV"
- apply(rule lmeasure_finite_integrable)
- using assms unfolding simple_function_def using False by auto
- have *:"\<And>x. real (indicator (f -` {f y}) x::pextreal) = (indicator (f -` {f y}) x)"
- by (auto simp: indicator_def)
- show ?thesis unfolding real_of_pextreal_mult[THEN sym]
- apply(rule has_integral_cmul[where 'b=real, unfolded real_scaleR_def])
- unfolding Int_UNIV_right lmeasure_eq_integral[OF mea,THEN sym]
- unfolding integral_eq_lmeasure[OF mea, symmetric] *
- apply(rule integrable_integral) using mea .
- qed auto
+ unfolding simple_integral_def space_lebesgue
+proof (subst lebesgue_simple_function_indicator)
+ let "?M x" = "lebesgue.\<mu> (f -` {x} \<inter> UNIV)"
+ let "?F x" = "indicator (f -` {x})"
+ { fix x y assume "y \<in> range f"
+ from subsetD[OF f' this] have "y * ?F y x = extreal (real y * ?F y x)"
+ by (cases rule: extreal2_cases[of y "?F y x"])
+ (auto simp: indicator_def one_extreal_def split: split_if_asm) }
+ moreover
+ { fix x assume x: "x\<in>range f"
+ have "x * ?M x = real x * real (?M x)"
+ proof cases
+ assume "x \<noteq> 0" with om[OF x] have "?M x \<noteq> \<infinity>" by auto
+ with subsetD[OF f' x] f[THEN lebesgue.simple_functionD(2)] show ?thesis
+ by (cases rule: extreal2_cases[of x "?M x"]) auto
+ qed simp }
+ ultimately
+ have "((\<lambda>x. real (\<Sum>y\<in>range f. y * ?F y x)) has_integral real (\<Sum>x\<in>range f. x * ?M x)) UNIV \<longleftrightarrow>
+ ((\<lambda>x. \<Sum>y\<in>range f. real y * ?F y x) has_integral (\<Sum>x\<in>range f. real x * real (?M x))) UNIV"
+ by simp
+ also have \<dots>
+ proof (intro has_integral_setsum has_integral_cmult_real lmeasure_finite_has_integral
+ real_of_extreal_pos lebesgue.positive_measure ballI)
+ show *: "finite (range f)" "\<And>y. f -` {y} \<in> sets lebesgue" "\<And>y. f -` {y} \<inter> UNIV \<in> sets lebesgue"
+ using lebesgue.simple_functionD[OF f] by auto
+ fix y assume "real y \<noteq> 0" "y \<in> range f"
+ with * om[OF this(2)] show "lebesgue.\<mu> (f -` {y}) = extreal (real (?M y))"
+ by (auto simp: extreal_real)
qed
-qed
+ finally show "((\<lambda>x. real (\<Sum>y\<in>range f. y * ?F y x)) has_integral real (\<Sum>x\<in>range f. x * ?M x)) UNIV" .
+qed fact
lemma bounded_realI: assumes "\<forall>x\<in>s. abs (x::real) \<le> B" shows "bounded s"
unfolding bounded_def dist_real_def apply(rule_tac x=0 in exI)
using assms by auto
lemma simple_function_has_integral':
- fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
- assumes f:"simple_function lebesgue f"
- and i: "integral\<^isup>S lebesgue f \<noteq> \<omega>"
+ fixes f::"'a::ordered_euclidean_space \<Rightarrow> extreal"
+ assumes f: "simple_function lebesgue f" "\<And>x. 0 \<le> f x"
+ and i: "integral\<^isup>S lebesgue f \<noteq> \<infinity>"
shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV"
-proof- let ?f = "\<lambda>x. if f x = \<omega> then 0 else f x"
- { fix x have "real (f x) = real (?f x)" by (cases "f x") auto } note * = this
- have **:"{x. f x \<noteq> ?f x} = f -` {\<omega>}" by auto
- have **:"lebesgue.\<mu> {x\<in>space lebesgue. f x \<noteq> ?f x} = 0"
- using lebesgue.simple_integral_omega[OF assms] by(auto simp add:**)
- show ?thesis apply(subst lebesgue.simple_integral_cong'[OF f _ **])
- apply(rule lebesgue.simple_function_compose1[OF f])
- unfolding * defer apply(rule simple_function_has_integral)
- proof-
- show "simple_function lebesgue ?f"
- using lebesgue.simple_function_compose1[OF f] .
- show "\<forall>x. ?f x \<noteq> \<omega>" by auto
- show "\<forall>x\<in>range ?f. lebesgue.\<mu> (?f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
- proof (safe, simp, safe, rule ccontr)
- fix y assume "f y \<noteq> \<omega>" "f y \<noteq> 0"
- hence "(\<lambda>x. if f x = \<omega> then 0 else f x) -` {if f y = \<omega> then 0 else f y} = f -` {f y}"
- by (auto split: split_if_asm)
- moreover assume "lebesgue.\<mu> ((\<lambda>x. if f x = \<omega> then 0 else f x) -` {if f y = \<omega> then 0 else f y}) = \<omega>"
- ultimately have "lebesgue.\<mu> (f -` {f y}) = \<omega>" by simp
- moreover
- have "f y * lebesgue.\<mu> (f -` {f y}) \<noteq> \<omega>" using i f
- unfolding simple_integral_def setsum_\<omega> simple_function_def
- by auto
- ultimately have "f y = 0" by (auto split: split_if_asm)
- then show False using `f y \<noteq> 0` by simp
- qed
+proof -
+ let ?f = "\<lambda>x. if x \<in> f -` {\<infinity>} then 0 else f x"
+ note f(1)[THEN lebesgue.simple_functionD(2)]
+ then have [simp, intro]: "\<And>X. f -` X \<in> sets lebesgue" by auto
+ have f': "simple_function lebesgue ?f"
+ using f by (intro lebesgue.simple_function_If_set) auto
+ have rng: "range ?f \<subseteq> {0..<\<infinity>}" using f by auto
+ have "AE x in lebesgue. f x = ?f x"
+ using lebesgue.simple_integral_PInf[OF f i]
+ by (intro lebesgue.AE_I[where N="f -` {\<infinity>} \<inter> space lebesgue"]) auto
+ from f(1) f' this have eq: "integral\<^isup>S lebesgue f = integral\<^isup>S lebesgue ?f"
+ by (rule lebesgue.simple_integral_cong_AE)
+ have real_eq: "\<And>x. real (f x) = real (?f x)" by auto
+
+ show ?thesis
+ unfolding eq real_eq
+ proof (rule simple_function_has_integral[OF f' rng])
+ fix x assume x: "x \<in> range ?f" and inf: "lebesgue.\<mu> (?f -` {x} \<inter> UNIV) = \<infinity>"
+ have "x * lebesgue.\<mu> (?f -` {x} \<inter> UNIV) = (\<integral>\<^isup>S y. x * indicator (?f -` {x}) y \<partial>lebesgue)"
+ using f'[THEN lebesgue.simple_functionD(2)]
+ by (simp add: lebesgue.simple_integral_cmult_indicator)
+ also have "\<dots> \<le> integral\<^isup>S lebesgue f"
+ using f'[THEN lebesgue.simple_functionD(2)] f
+ by (intro lebesgue.simple_integral_mono lebesgue.simple_function_mult lebesgue.simple_function_indicator)
+ (auto split: split_indicator)
+ finally show "x = 0" unfolding inf using i subsetD[OF rng x] by (auto split: split_if_asm)
qed
qed
-lemma (in measure_space) positive_integral_monotone_convergence:
- fixes f::"nat \<Rightarrow> 'a \<Rightarrow> pextreal"
- assumes i: "\<And>i. f i \<in> borel_measurable M" and mono: "\<And>x. mono (\<lambda>n. f n x)"
- and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
- shows "u \<in> borel_measurable M"
- and "(\<lambda>i. integral\<^isup>P M (f i)) ----> integral\<^isup>P M u" (is ?ilim)
-proof -
- from positive_integral_isoton[unfolded isoton_fun_expand isoton_iff_Lim_mono, of f u]
- show ?ilim using mono lim i by auto
- have "\<And>x. (SUP i. f i x) = u x" using mono lim SUP_Lim_pextreal
- unfolding fun_eq_iff mono_def by auto
- moreover have "(\<lambda>x. SUP i. f i x) \<in> borel_measurable M"
- using i by auto
- ultimately show "u \<in> borel_measurable M" by simp
-qed
+lemma real_of_extreal_positive_mono:
+ "\<lbrakk>0 \<le> x; x \<le> y; y \<noteq> \<infinity>\<rbrakk> \<Longrightarrow> real x \<le> real y"
+ by (cases rule: extreal2_cases[of x y]) auto
lemma positive_integral_has_integral:
- fixes f::"'a::ordered_euclidean_space => pextreal"
- assumes f:"f \<in> borel_measurable lebesgue"
- and int_om:"integral\<^isup>P lebesgue f \<noteq> \<omega>"
- and f_om:"\<forall>x. f x \<noteq> \<omega>" (* TODO: remove this *)
+ fixes f :: "'a::ordered_euclidean_space \<Rightarrow> extreal"
+ assumes f: "f \<in> borel_measurable lebesgue" "range f \<subseteq> {0..<\<infinity>}" "integral\<^isup>P lebesgue f \<noteq> \<infinity>"
shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>P lebesgue f))) UNIV"
-proof- let ?i = "integral\<^isup>P lebesgue f"
- from lebesgue.borel_measurable_implies_simple_function_sequence[OF f]
- guess u .. note conjunctD2[OF this,rule_format] note u = conjunctD2[OF this(1)] this(2)
- let ?u = "\<lambda>i x. real (u i x)" and ?f = "\<lambda>x. real (f x)"
- have u_simple:"\<And>k. integral\<^isup>S lebesgue (u k) = integral\<^isup>P lebesgue (u k)"
- apply(subst lebesgue.positive_integral_eq_simple_integral[THEN sym,OF u(1)]) ..
- have int_u_le:"\<And>k. integral\<^isup>S lebesgue (u k) \<le> integral\<^isup>P lebesgue f"
- unfolding u_simple apply(rule lebesgue.positive_integral_mono)
- using isoton_Sup[OF u(3)] unfolding le_fun_def by auto
- have u_int_om:"\<And>i. integral\<^isup>S lebesgue (u i) \<noteq> \<omega>"
- proof- case goal1 thus ?case using int_u_le[of i] int_om by auto qed
+proof -
+ from lebesgue.borel_measurable_implies_simple_function_sequence'[OF f(1)]
+ guess u . note u = this
+ have SUP_eq: "\<And>x. (SUP i. u i x) = f x"
+ using u(4) f(2)[THEN subsetD] by (auto split: split_max)
+ let "?u i x" = "real (u i x)"
+ note u_eq = lebesgue.positive_integral_eq_simple_integral[OF u(1,5), symmetric]
+ { fix i
+ note u_eq
+ also have "integral\<^isup>P lebesgue (u i) \<le> (\<integral>\<^isup>+x. max 0 (f x) \<partial>lebesgue)"
+ by (intro lebesgue.positive_integral_mono) (auto intro: le_SUPI simp: u(4)[symmetric])
+ finally have "integral\<^isup>S lebesgue (u i) \<noteq> \<infinity>"
+ unfolding positive_integral_max_0 using f by auto }
+ note u_fin = this
+ then have u_int: "\<And>i. (?u i has_integral real (integral\<^isup>S lebesgue (u i))) UNIV"
+ by (rule simple_function_has_integral'[OF u(1,5)])
+ have "\<forall>x. \<exists>r\<ge>0. f x = extreal r"
+ proof
+ fix x from f(2) have "0 \<le> f x" "f x \<noteq> \<infinity>" by (auto simp: subset_eq)
+ then show "\<exists>r\<ge>0. f x = extreal r" by (cases "f x") auto
+ qed
+ from choice[OF this] obtain f' where f': "f = (\<lambda>x. extreal (f' x))" "\<And>x. 0 \<le> f' x" by auto
+
+ have "\<forall>i. \<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = extreal (r x)"
+ proof
+ fix i show "\<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = extreal (r x)"
+ proof (intro choice allI)
+ fix i x have "u i x \<noteq> \<infinity>" using u(3)[of i] by (auto simp: image_iff) metis
+ then show "\<exists>r\<ge>0. u i x = extreal r" using u(5)[of i x] by (cases "u i x") auto
+ qed
+ qed
+ from choice[OF this] obtain u' where
+ u': "u = (\<lambda>i x. extreal (u' i x))" "\<And>i x. 0 \<le> u' i x" by (auto simp: fun_eq_iff)
- note u_int = simple_function_has_integral'[OF u(1) this]
- have "(\<lambda>x. real (f x)) integrable_on UNIV \<and>
- (\<lambda>k. Integration.integral UNIV (\<lambda>x. real (u k x))) ----> Integration.integral UNIV (\<lambda>x. real (f x))"
- apply(rule monotone_convergence_increasing) apply(rule,rule,rule u_int)
- proof safe case goal1 show ?case apply(rule real_of_pextreal_mono) using u(2,3) by auto
- next case goal2 show ?case using u(3) apply(subst lim_Real[THEN sym])
- prefer 3 apply(subst Real_real') defer apply(subst Real_real')
- using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] using f_om u by auto
- next case goal3
- show ?case apply(rule bounded_realI[where B="real (integral\<^isup>P lebesgue f)"])
- apply safe apply(subst abs_of_nonneg) apply(rule integral_nonneg,rule) apply(rule u_int)
- unfolding integral_unique[OF u_int] defer apply(rule real_of_pextreal_mono[OF _ int_u_le])
- using u int_om by auto
- qed note int = conjunctD2[OF this]
+ have convergent: "f' integrable_on UNIV \<and>
+ (\<lambda>k. integral UNIV (u' k)) ----> integral UNIV f'"
+ proof (intro monotone_convergence_increasing allI ballI)
+ show int: "\<And>k. (u' k) integrable_on UNIV"
+ using u_int unfolding integrable_on_def u' by auto
+ show "\<And>k x. u' k x \<le> u' (Suc k) x" using u(2,3,5)
+ by (auto simp: incseq_Suc_iff le_fun_def image_iff u' intro!: real_of_extreal_positive_mono)
+ show "\<And>x. (\<lambda>k. u' k x) ----> f' x"
+ using SUP_eq u(2)
+ by (intro SUP_eq_LIMSEQ[THEN iffD1]) (auto simp: u' f' incseq_mono incseq_Suc_iff le_fun_def)
+ show "bounded {integral UNIV (u' k)|k. True}"
+ proof (safe intro!: bounded_realI)
+ fix k
+ have "\<bar>integral UNIV (u' k)\<bar> = integral UNIV (u' k)"
+ by (intro abs_of_nonneg integral_nonneg int ballI u')
+ also have "\<dots> = real (integral\<^isup>S lebesgue (u k))"
+ using u_int[THEN integral_unique] by (simp add: u')
+ also have "\<dots> = real (integral\<^isup>P lebesgue (u k))"
+ using lebesgue.positive_integral_eq_simple_integral[OF u(1,5)] by simp
+ also have "\<dots> \<le> real (integral\<^isup>P lebesgue f)" using f
+ by (auto intro!: real_of_extreal_positive_mono lebesgue.positive_integral_positive
+ lebesgue.positive_integral_mono le_SUPI simp: SUP_eq[symmetric])
+ finally show "\<bar>integral UNIV (u' k)\<bar> \<le> real (integral\<^isup>P lebesgue f)" .
+ qed
+ qed
- have "(\<lambda>i. integral\<^isup>S lebesgue (u i)) ----> ?i" unfolding u_simple
- apply(rule lebesgue.positive_integral_monotone_convergence(2))
- apply(rule lebesgue.borel_measurable_simple_function[OF u(1)])
- using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] by auto
- hence "(\<lambda>i. real (integral\<^isup>S lebesgue (u i))) ----> real ?i" apply-
- apply(subst lim_Real[THEN sym]) prefer 3
- apply(subst Real_real') defer apply(subst Real_real')
- using u f_om int_om u_int_om by auto
- note * = LIMSEQ_unique[OF this int(2)[unfolded integral_unique[OF u_int]]]
- show ?thesis unfolding * by(rule integrable_integral[OF int(1)])
+ have "integral\<^isup>P lebesgue f = extreal (integral UNIV f')"
+ proof (rule tendsto_unique[OF trivial_limit_sequentially])
+ have "(\<lambda>i. integral\<^isup>S lebesgue (u i)) ----> (SUP i. integral\<^isup>P lebesgue (u i))"
+ unfolding u_eq by (intro LIMSEQ_extreal_SUPR lebesgue.incseq_positive_integral u)
+ also note lebesgue.positive_integral_monotone_convergence_SUP
+ [OF u(2) lebesgue.borel_measurable_simple_function[OF u(1)] u(5), symmetric]
+ finally show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> integral\<^isup>P lebesgue f"
+ unfolding SUP_eq .
+
+ { fix k
+ have "0 \<le> integral\<^isup>S lebesgue (u k)"
+ using u by (auto intro!: lebesgue.simple_integral_positive)
+ then have "integral\<^isup>S lebesgue (u k) = extreal (real (integral\<^isup>S lebesgue (u k)))"
+ using u_fin by (auto simp: extreal_real) }
+ note * = this
+ show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> extreal (integral UNIV f')"
+ using convergent using u_int[THEN integral_unique, symmetric]
+ by (subst *) (simp add: lim_extreal u')
+ qed
+ then show ?thesis using convergent by (simp add: f' integrable_integral)
qed
lemma lebesgue_integral_has_integral:
- fixes f::"'a::ordered_euclidean_space => real"
- assumes f:"integrable lebesgue f"
+ fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
+ assumes f: "integrable lebesgue f"
shows "(f has_integral (integral\<^isup>L lebesgue f)) UNIV"
-proof- let ?n = "\<lambda>x. - min (f x) 0" and ?p = "\<lambda>x. max (f x) 0"
- have *:"f = (\<lambda>x. ?p x - ?n x)" apply rule by auto
- note f = integrableD[OF f]
- show ?thesis unfolding lebesgue_integral_def apply(subst *)
- proof(rule has_integral_sub) case goal1
- have *:"\<forall>x. Real (f x) \<noteq> \<omega>" by auto
- note lebesgue.borel_measurable_Real[OF f(1)]
- from positive_integral_has_integral[OF this f(2) *]
- show ?case unfolding real_Real_max .
- next case goal2
- have *:"\<forall>x. Real (- f x) \<noteq> \<omega>" by auto
- note lebesgue.borel_measurable_uminus[OF f(1)]
- note lebesgue.borel_measurable_Real[OF this]
- from positive_integral_has_integral[OF this f(3) *]
- show ?case unfolding real_Real_max minus_min_eq_max by auto
- qed
+proof -
+ let ?n = "\<lambda>x. real (extreal (max 0 (- f x)))" and ?p = "\<lambda>x. real (extreal (max 0 (f x)))"
+ have *: "f = (\<lambda>x. ?p x - ?n x)" by (auto simp del: extreal_max)
+ { fix f have "(\<integral>\<^isup>+ x. extreal (f x) \<partial>lebesgue) = (\<integral>\<^isup>+ x. extreal (max 0 (f x)) \<partial>lebesgue)"
+ by (intro lebesgue.positive_integral_cong_pos) (auto split: split_max) }
+ note eq = this
+ show ?thesis
+ unfolding lebesgue_integral_def
+ apply (subst *)
+ apply (rule has_integral_sub)
+ unfolding eq[of f] eq[of "\<lambda>x. - f x"]
+ apply (safe intro!: positive_integral_has_integral)
+ using integrableD[OF f]
+ by (auto simp: zero_extreal_def[symmetric] positive_integral_max_0 split: split_max
+ intro!: lebesgue.measurable_If lebesgue.borel_measurable_extreal)
qed
lemma lebesgue_positive_integral_eq_borel:
- "f \<in> borel_measurable borel \<Longrightarrow> integral\<^isup>P lebesgue f = integral\<^isup>P lborel f"
- by (auto intro!: lebesgue.positive_integral_subalgebra[symmetric]) default
+ assumes f: "f \<in> borel_measurable borel"
+ shows "integral\<^isup>P lebesgue f = integral\<^isup>P lborel f"
+proof -
+ from f have "integral\<^isup>P lebesgue (\<lambda>x. max 0 (f x)) = integral\<^isup>P lborel (\<lambda>x. max 0 (f x))"
+ by (auto intro!: lebesgue.positive_integral_subalgebra[symmetric]) default
+ then show ?thesis unfolding positive_integral_max_0 .
+qed
lemma lebesgue_integral_eq_borel:
assumes "f \<in> borel_measurable borel"
@@ -771,7 +827,7 @@
have "sets ?G = sets (\<Pi>\<^isub>M i\<in>I.
sigma \<lparr> space = UNIV::real set, sets = range lessThan, measure = lebesgue.\<mu> \<rparr>)"
by (subst sigma_product_algebra_sigma_eq[of I "\<lambda>_ i. {..<real i}" ])
- (auto intro!: measurable_sigma_sigma isotoneI real_arch_lt
+ (auto intro!: measurable_sigma_sigma incseq_SucI real_arch_lt
simp: product_algebra_def)
then show ?thesis
unfolding lborel_def borel_eq_lessThan lebesgue_def sigma_def by simp
@@ -838,9 +894,10 @@
let ?E = "\<lparr> space = UNIV :: 'a set, sets = range (\<lambda>(a,b). {a..b}) \<rparr>"
show "Int_stable ?E" using Int_stable_cuboids .
show "range cube \<subseteq> sets ?E" unfolding cube_def_raw by auto
+ show "incseq cube" using cube_subset_Suc by (auto intro!: incseq_SucI)
{ fix x have "\<exists>n. x \<in> cube n" using mem_big_cube[of x] by fastsimp }
- then show "cube \<up> space ?E" by (intro isotoneI cube_subset_Suc) auto
- { fix i show "lborel.\<mu> (cube i) \<noteq> \<omega>" unfolding cube_def by auto }
+ then show "(\<Union>i. cube i) = space ?E" by auto
+ { fix i show "lborel.\<mu> (cube i) \<noteq> \<infinity>" unfolding cube_def by auto }
show "A \<in> sets (sigma ?E)" "sets (sigma ?E) = sets lborel" "space ?E = space lborel"
using assms by (simp_all add: borel_eq_atLeastAtMost)
@@ -857,7 +914,7 @@
by (simp add: interval_ne_empty eucl_le[where 'a='a])
then have "lborel.\<mu> X = (\<Prod>x<DIM('a). lborel.\<mu> {a $$ x .. b $$ x})"
by (auto simp: content_closed_interval eucl_le[where 'a='a]
- intro!: Real_setprod )
+ intro!: setprod_extreal[symmetric])
also have "\<dots> = measure ?P (?T X)"
unfolding * by (subst lborel_space.measure_times) auto
finally show ?thesis .
@@ -882,7 +939,7 @@
using lborel_eq_lborel_space[OF A] by simp
lemma borel_fubini_positiv_integral:
- fixes f :: "'a::ordered_euclidean_space \<Rightarrow> pextreal"
+ fixes f :: "'a::ordered_euclidean_space \<Rightarrow> extreal"
assumes f: "f \<in> borel_measurable borel"
shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(lborel_space.P DIM('a))"
proof (rule lborel_space.positive_integral_vimage[OF _ measure_preserving_p2e _])
--- a/src/HOL/Probability/Measure.thy Mon Mar 14 14:37:47 2011 +0100
+++ b/src/HOL/Probability/Measure.thy Mon Mar 14 14:37:49 2011 +0100
@@ -76,7 +76,7 @@
lemma (in measure_space) measure_countably_additive:
assumes "range A \<subseteq> sets M" "disjoint_family A"
- shows "psuminf (\<lambda>i. \<mu> (A i)) = \<mu> (\<Union>i. A i)"
+ shows "(\<Sum>i. \<mu> (A i)) = \<mu> (\<Union>i. A i)"
proof -
have "(\<Union> i. A i) \<in> sets M" using assms(1) by (rule countable_UN)
with ca assms show ?thesis by (simp add: countably_additive_def)
@@ -94,13 +94,13 @@
interpret N: sigma_algebra N by (intro sigma_algebra_cong assms)
show ?thesis
proof
- show "measure N {} = 0" using assms by auto
+ show "positive N (measure N)" using assms by (auto simp: positive_def)
show "countably_additive N (measure N)" unfolding countably_additive_def
proof safe
fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets N" "disjoint_family A"
then have "\<And>i. A i \<in> sets M" "(UNION UNIV A) \<in> sets M" unfolding assms by auto
from measure_countably_additive[of A] A this[THEN assms(1)]
- show "(\<Sum>\<^isub>\<infinity>n. measure N (A n)) = measure N (UNION UNIV A)"
+ show "(\<Sum>n. measure N (A n)) = measure N (UNION UNIV A)"
unfolding assms by simp
qed
qed
@@ -124,51 +124,51 @@
have "b = a \<union> (b - a)" using assms by auto
moreover have "{} = a \<inter> (b - a)" by auto
ultimately have "\<mu> b = \<mu> a + \<mu> (b - a)"
- using measure_additive[of a "b - a"] local.Diff[of b a] assms by auto
- moreover have "\<mu> (b - a) \<ge> 0" using assms by auto
+ using measure_additive[of a "b - a"] Diff[of b a] assms by auto
+ moreover have "\<mu> a + 0 \<le> \<mu> a + \<mu> (b - a)" using assms by (intro add_mono) auto
ultimately show "\<mu> a \<le> \<mu> b" by auto
qed
lemma (in measure_space) measure_compl:
- assumes s: "s \<in> sets M" and fin: "\<mu> s \<noteq> \<omega>"
+ assumes s: "s \<in> sets M" and fin: "\<mu> s \<noteq> \<infinity>"
shows "\<mu> (space M - s) = \<mu> (space M) - \<mu> s"
proof -
have s_less_space: "\<mu> s \<le> \<mu> (space M)"
using s by (auto intro!: measure_mono sets_into_space)
-
+ from s have "0 \<le> \<mu> s" by auto
have "\<mu> (space M) = \<mu> (s \<union> (space M - s))" using s
by (metis Un_Diff_cancel Un_absorb1 s sets_into_space)
also have "... = \<mu> s + \<mu> (space M - s)"
by (rule additiveD [OF additive]) (auto simp add: s)
finally have "\<mu> (space M) = \<mu> s + \<mu> (space M - s)" .
- thus ?thesis
- unfolding minus_pextreal_eq2[OF s_less_space fin]
- by (simp add: ac_simps)
+ then show ?thesis
+ using fin `0 \<le> \<mu> s`
+ unfolding extreal_eq_minus_iff by (auto simp: ac_simps)
qed
lemma (in measure_space) measure_Diff:
- assumes finite: "\<mu> B \<noteq> \<omega>"
+ assumes finite: "\<mu> B \<noteq> \<infinity>"
and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
shows "\<mu> (A - B) = \<mu> A - \<mu> B"
proof -
- have *: "(A - B) \<union> B = A" using `B \<subseteq> A` by auto
-
- have "\<mu> ((A - B) \<union> B) = \<mu> (A - B) + \<mu> B"
- using measurable by (rule_tac measure_additive[symmetric]) auto
- thus ?thesis unfolding * using `\<mu> B \<noteq> \<omega>`
- by (simp add: pextreal_cancel_plus_minus)
+ have "0 \<le> \<mu> B" using assms by auto
+ have "(A - B) \<union> B = A" using `B \<subseteq> A` by auto
+ then have "\<mu> A = \<mu> ((A - B) \<union> B)" by simp
+ also have "\<dots> = \<mu> (A - B) + \<mu> B"
+ using measurable by (subst measure_additive[symmetric]) auto
+ finally show "\<mu> (A - B) = \<mu> A - \<mu> B"
+ unfolding extreal_eq_minus_iff
+ using finite `0 \<le> \<mu> B` by auto
qed
lemma (in measure_space) measure_countable_increasing:
assumes A: "range A \<subseteq> sets M"
and A0: "A 0 = {}"
- and ASuc: "\<And>n. A n \<subseteq> A (Suc n)"
+ and ASuc: "\<And>n. A n \<subseteq> A (Suc n)"
shows "(SUP n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
proof -
- {
- fix n
- have "\<mu> (A n) =
- setsum (\<mu> \<circ> (\<lambda>i. A (Suc i) - A i)) {..<n}"
+ { fix n
+ have "\<mu> (A n) = (\<Sum>i<n. \<mu> (A (Suc i) - A i))"
proof (induct n)
case 0 thus ?case by (auto simp add: A0)
next
@@ -199,92 +199,83 @@
by (metis A Diff range_subsetD)
have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M"
by (blast intro: range_subsetD [OF A])
- have "psuminf ( (\<lambda>i. \<mu> (A (Suc i) - A i))) = \<mu> (\<Union>i. A (Suc i) - A i)"
+ have "(SUP n. \<Sum>i<n. \<mu> (A (Suc i) - A i)) = (\<Sum>i. \<mu> (A (Suc i) - A i))"
+ using A by (auto intro!: suminf_extreal_eq_SUPR[symmetric])
+ also have "\<dots> = \<mu> (\<Union>i. A (Suc i) - A i)"
by (rule measure_countably_additive)
(auto simp add: disjoint_family_Suc ASuc A1 A2)
also have "... = \<mu> (\<Union>i. A i)"
by (simp add: Aeq)
- finally have "psuminf (\<lambda>i. \<mu> (A (Suc i) - A i)) = \<mu> (\<Union>i. A i)" .
- thus ?thesis
- by (auto simp add: Meq psuminf_def)
+ finally have "(SUP n. \<Sum>i<n. \<mu> (A (Suc i) - A i)) = \<mu> (\<Union>i. A i)" .
+ then show ?thesis by (auto simp add: Meq)
qed
lemma (in measure_space) continuity_from_below:
- assumes A: "range A \<subseteq> sets M"
- and ASuc: "!!n. A n \<subseteq> A (Suc n)"
+ assumes A: "range A \<subseteq> sets M" and "incseq A"
shows "(SUP n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
proof -
have *: "(SUP n. \<mu> (nat_case {} A (Suc n))) = (SUP n. \<mu> (nat_case {} A n))"
- apply (rule Sup_mono_offset_Suc)
- apply (rule measure_mono)
- using assms by (auto split: nat.split)
-
+ using A by (auto intro!: SUPR_eq exI split: nat.split)
have ueq: "(\<Union>i. nat_case {} A i) = (\<Union>i. A i)"
by (auto simp add: split: nat.splits)
have meq: "\<And>n. \<mu> (A n) = (\<mu> \<circ> nat_case {} A) (Suc n)"
by simp
have "(SUP n. \<mu> (nat_case {} A n)) = \<mu> (\<Union>i. nat_case {} A i)"
- by (rule measure_countable_increasing)
- (auto simp add: range_subsetD [OF A] subsetD [OF ASuc] split: nat.splits)
+ using range_subsetD[OF A] incseq_SucD[OF `incseq A`]
+ by (force split: nat.splits intro!: measure_countable_increasing)
also have "\<mu> (\<Union>i. nat_case {} A i) = \<mu> (\<Union>i. A i)"
by (simp add: ueq)
finally have "(SUP n. \<mu> (nat_case {} A n)) = \<mu> (\<Union>i. A i)" .
thus ?thesis unfolding meq * comp_def .
qed
-lemma (in measure_space) measure_up:
- assumes "\<And>i. B i \<in> sets M" "B \<up> P"
- shows "(\<lambda>i. \<mu> (B i)) \<up> \<mu> P"
- using assms unfolding isoton_def
- by (auto intro!: measure_mono continuity_from_below)
+lemma (in measure_space) measure_incseq:
+ assumes "range B \<subseteq> sets M" "incseq B"
+ shows "incseq (\<lambda>i. \<mu> (B i))"
+ using assms by (auto simp: incseq_def intro!: measure_mono)
-lemma (in measure_space) continuity_from_below':
- assumes A: "range A \<subseteq> sets M" "\<And>i. A i \<subseteq> A (Suc i)"
- shows "(\<lambda>i. (\<mu> (A i))) ----> (\<mu> (\<Union>i. A i))"
-proof- let ?A = "\<Union>i. A i"
- have " (\<lambda>i. \<mu> (A i)) \<up> \<mu> ?A" apply(rule measure_up)
- using assms unfolding complete_lattice_class.isoton_def by auto
- thus ?thesis by(rule isotone_Lim(1))
-qed
+lemma (in measure_space) continuity_from_below_Lim:
+ assumes A: "range A \<subseteq> sets M" "incseq A"
+ shows "(\<lambda>i. (\<mu> (A i))) ----> \<mu> (\<Union>i. A i)"
+ using LIMSEQ_extreal_SUPR[OF measure_incseq, OF A]
+ continuity_from_below[OF A] by simp
+
+lemma (in measure_space) measure_decseq:
+ assumes "range B \<subseteq> sets M" "decseq B"
+ shows "decseq (\<lambda>i. \<mu> (B i))"
+ using assms by (auto simp: decseq_def intro!: measure_mono)
lemma (in measure_space) continuity_from_above:
- assumes A: "range A \<subseteq> sets M"
- and mono_Suc: "\<And>n. A (Suc n) \<subseteq> A n"
- and finite: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+ assumes A: "range A \<subseteq> sets M" and "decseq A"
+ and finite: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
shows "(INF n. \<mu> (A n)) = \<mu> (\<Inter>i. A i)"
proof -
- { fix n have "A n \<subseteq> A 0" using mono_Suc by (induct n) auto }
- note mono = this
-
have le_MI: "\<mu> (\<Inter>i. A i) \<le> \<mu> (A 0)"
using A by (auto intro!: measure_mono)
- hence *: "\<mu> (\<Inter>i. A i) \<noteq> \<omega>" using finite[of 0] by auto
+ hence *: "\<mu> (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by auto
+
+ have A0: "0 \<le> \<mu> (A 0)" using A by auto
- have le_IM: "(INF n. \<mu> (A n)) \<le> \<mu> (A 0)"
- by (rule INF_leI) simp
-
- have "\<mu> (A 0) - (INF n. \<mu> (A n)) = (SUP n. \<mu> (A 0 - A n))"
- unfolding pextreal_SUP_minus[symmetric]
- using mono A finite by (subst measure_Diff) auto
+ have "\<mu> (A 0) - (INF n. \<mu> (A n)) = \<mu> (A 0) + (SUP n. - \<mu> (A n))"
+ by (simp add: extreal_SUPR_uminus minus_extreal_def)
+ also have "\<dots> = (SUP n. \<mu> (A 0) - \<mu> (A n))"
+ unfolding minus_extreal_def using A0 assms
+ by (subst SUPR_extreal_add) (auto simp add: measure_decseq)
+ also have "\<dots> = (SUP n. \<mu> (A 0 - A n))"
+ using A finite `decseq A`[unfolded decseq_def] by (subst measure_Diff) auto
also have "\<dots> = \<mu> (\<Union>i. A 0 - A i)"
proof (rule continuity_from_below)
show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M"
using A by auto
- show "\<And>n. A 0 - A n \<subseteq> A 0 - A (Suc n)"
- using mono_Suc by auto
+ show "incseq (\<lambda>n. A 0 - A n)"
+ using `decseq A` by (auto simp add: incseq_def decseq_def)
qed
also have "\<dots> = \<mu> (A 0) - \<mu> (\<Inter>i. A i)"
- using mono A finite * by (simp, subst measure_Diff) auto
+ using A finite * by (simp, subst measure_Diff) auto
finally show ?thesis
- by (rule pextreal_diff_eq_diff_imp_eq[OF finite[of 0] le_IM le_MI])
+ unfolding extreal_minus_eq_minus_iff using finite A0 by auto
qed
-lemma (in measure_space) measure_down:
- assumes "\<And>i. B i \<in> sets M" "B \<down> P"
- and finite: "\<And>i. \<mu> (B i) \<noteq> \<omega>"
- shows "(\<lambda>i. \<mu> (B i)) \<down> \<mu> P"
- using assms unfolding antiton_def
- by (auto intro!: measure_mono continuity_from_above)
lemma (in measure_space) measure_insert:
assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
shows "\<mu> (insert x A) = \<mu> {x} + \<mu> A"
@@ -293,109 +284,26 @@
from measure_additive[OF sets this] show ?thesis by simp
qed
-lemma (in measure_space) measure_finite_singleton:
- assumes fin: "finite S"
- and ssets: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
- shows "\<mu> S = (\<Sum>x\<in>S. \<mu> {x})"
+lemma (in measure_space) measure_setsum:
+ assumes "finite S" and "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M"
+ assumes disj: "disjoint_family_on A S"
+ shows "(\<Sum>i\<in>S. \<mu> (A i)) = \<mu> (\<Union>i\<in>S. A i)"
using assms proof induct
- case (insert x S)
- have *: "\<mu> S = (\<Sum>x\<in>S. \<mu> {x})" "{x} \<in> sets M"
- using insert.prems by (blast intro!: insert.hyps(3))+
-
- have "(\<Union>x\<in>S. {x}) \<in> sets M"
- using insert.prems `finite S` by (blast intro!: finite_UN)
- hence "S \<in> sets M" by auto
- from measure_insert[OF `{x} \<in> sets M` this `x \<notin> S`]
- show ?case using `x \<notin> S` `finite S` * by simp
+ case (insert i S)
+ then have "(\<Sum>i\<in>S. \<mu> (A i)) = \<mu> (\<Union>a\<in>S. A a)"
+ by (auto intro: disjoint_family_on_mono)
+ moreover have "A i \<inter> (\<Union>a\<in>S. A a) = {}"
+ using `disjoint_family_on A (insert i S)` `i \<notin> S`
+ by (auto simp: disjoint_family_on_def)
+ ultimately show ?case using insert
+ by (auto simp: measure_additive finite_UN)
qed simp
-lemma (in measure_space) measure_finitely_additive':
- assumes "f \<in> ({..< n :: nat} \<rightarrow> sets M)"
- assumes "\<And> a b. \<lbrakk>a < n ; b < n ; a \<noteq> b\<rbrakk> \<Longrightarrow> f a \<inter> f b = {}"
- assumes "s = \<Union> (f ` {..< n})"
- shows "(\<Sum>i<n. (\<mu> \<circ> f) i) = \<mu> s"
-proof -
- def f' == "\<lambda> i. (if i < n then f i else {})"
- have rf: "range f' \<subseteq> sets M" unfolding f'_def
- using assms empty_sets by auto
- have df: "disjoint_family f'" unfolding f'_def disjoint_family_on_def
- using assms by simp
- have "\<Union> range f' = (\<Union> i \<in> {..< n}. f i)"
- unfolding f'_def by auto
- then have "\<mu> s = \<mu> (\<Union> range f')"
- using assms by simp
- then have part1: "(\<Sum>\<^isub>\<infinity> i. \<mu> (f' i)) = \<mu> s"
- using df rf ca[unfolded countably_additive_def, rule_format, of f']
- by auto
-
- have "(\<Sum>\<^isub>\<infinity> i. \<mu> (f' i)) = (\<Sum> i< n. \<mu> (f' i))"
- by (rule psuminf_finite) (simp add: f'_def)
- also have "\<dots> = (\<Sum>i<n. \<mu> (f i))"
- unfolding f'_def by auto
- finally have part2: "(\<Sum>\<^isub>\<infinity> i. \<mu> (f' i)) = (\<Sum>i<n. \<mu> (f i))" by simp
- show ?thesis using part1 part2 by auto
-qed
-
-
-lemma (in measure_space) measure_finitely_additive:
- assumes "finite r"
- assumes "r \<subseteq> sets M"
- assumes d: "\<And> a b. \<lbrakk>a \<in> r ; b \<in> r ; a \<noteq> b\<rbrakk> \<Longrightarrow> a \<inter> b = {}"
- shows "(\<Sum> i \<in> r. \<mu> i) = \<mu> (\<Union> r)"
-using assms
-proof -
- (* counting the elements in r is enough *)
- let ?R = "{..<card r}"
- obtain f where f: "f ` ?R = r" "inj_on f ?R"
- using ex_bij_betw_nat_finite[unfolded bij_betw_def, OF `finite r`]
- unfolding atLeast0LessThan by auto
- hence f_into_sets: "f \<in> ?R \<rightarrow> sets M" using assms by auto
- have disj: "\<And> a b. \<lbrakk>a \<in> ?R ; b \<in> ?R ; a \<noteq> b\<rbrakk> \<Longrightarrow> f a \<inter> f b = {}"
- proof -
- fix a b assume asm: "a \<in> ?R" "b \<in> ?R" "a \<noteq> b"
- hence neq: "f a \<noteq> f b" using f[unfolded inj_on_def, rule_format] by blast
- from asm have "f a \<in> r" "f b \<in> r" using f by auto
- thus "f a \<inter> f b = {}" using d[of "f a" "f b"] f using neq by auto
- qed
- have "(\<Union> r) = (\<Union> i \<in> ?R. f i)"
- using f by auto
- hence "\<mu> (\<Union> r)= \<mu> (\<Union> i \<in> ?R. f i)" by simp
- also have "\<dots> = (\<Sum> i \<in> ?R. \<mu> (f i))"
- using measure_finitely_additive'[OF f_into_sets disj] by simp
- also have "\<dots> = (\<Sum> a \<in> r. \<mu> a)"
- using f[rule_format] setsum_reindex[of f ?R "\<lambda> a. \<mu> a"] by auto
- finally show ?thesis by simp
-qed
-
-lemma (in measure_space) measure_finitely_additive'':
- assumes "finite s"
- assumes "\<And> i. i \<in> s \<Longrightarrow> a i \<in> sets M"
- assumes d: "disjoint_family_on a s"
- shows "(\<Sum> i \<in> s. \<mu> (a i)) = \<mu> (\<Union> i \<in> s. a i)"
-using assms
-proof -
- (* counting the elements in s is enough *)
- let ?R = "{..< card s}"
- obtain f where f: "f ` ?R = s" "inj_on f ?R"
- using ex_bij_betw_nat_finite[unfolded bij_betw_def, OF `finite s`]
- unfolding atLeast0LessThan by auto
- hence f_into_sets: "a \<circ> f \<in> ?R \<rightarrow> sets M" using assms unfolding o_def by auto
- have disj: "\<And> i j. \<lbrakk>i \<in> ?R ; j \<in> ?R ; i \<noteq> j\<rbrakk> \<Longrightarrow> (a \<circ> f) i \<inter> (a \<circ> f) j = {}"
- proof -
- fix i j assume asm: "i \<in> ?R" "j \<in> ?R" "i \<noteq> j"
- hence neq: "f i \<noteq> f j" using f[unfolded inj_on_def, rule_format] by blast
- from asm have "f i \<in> s" "f j \<in> s" using f by auto
- thus "(a \<circ> f) i \<inter> (a \<circ> f) j = {}"
- using d f neq unfolding disjoint_family_on_def by auto
- qed
- have "(\<Union> i \<in> s. a i) = (\<Union> i \<in> f ` ?R. a i)" using f by auto
- hence "(\<Union> i \<in> s. a i) = (\<Union> i \<in> ?R. a (f i))" by auto
- hence "\<mu> (\<Union> i \<in> s. a i) = (\<Sum> i \<in> ?R. \<mu> (a (f i)))"
- using measure_finitely_additive'[OF f_into_sets disj] by simp
- also have "\<dots> = (\<Sum> i \<in> s. \<mu> (a i))"
- using f[rule_format] setsum_reindex[of f ?R "\<lambda> i. \<mu> (a i)"] by auto
- finally show ?thesis by simp
-qed
+lemma (in measure_space) measure_finite_singleton:
+ assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
+ shows "\<mu> S = (\<Sum>x\<in>S. \<mu> {x})"
+ using measure_setsum[of S "\<lambda>x. {x}", OF assms]
+ by (auto simp: disjoint_family_on_def)
lemma finite_additivity_sufficient:
assumes "sigma_algebra M"
@@ -405,7 +313,7 @@
interpret sigma_algebra M by fact
show ?thesis
proof
- show [simp]: "measure M {} = 0" using pos by (simp add: positive_def)
+ show [simp]: "positive M (measure M)" using pos by (simp add: positive_def)
show "countably_additive M (measure M)"
proof (auto simp add: countably_additive_def)
fix A :: "nat \<Rightarrow> 'a set"
@@ -434,12 +342,12 @@
by blast
qed
then obtain N where N: "\<forall>m\<ge>N. A m = {}" by blast
- then have "\<forall>m\<ge>N. measure M (A m) = 0" by simp
- then have "(\<Sum>\<^isub>\<infinity> n. measure M (A n)) = setsum (\<lambda>m. measure M (A m)) {..<N}"
- by (simp add: psuminf_finite)
+ then have "\<forall>m\<ge>N. measure M (A m) = 0" using pos[unfolded positive_def] by simp
+ then have "(\<Sum>n. measure M (A n)) = (\<Sum>m<N. measure M (A m))"
+ by (simp add: suminf_finite)
also have "... = measure M (\<Union>i<N. A i)"
proof (induct N)
- case 0 thus ?case by simp
+ case 0 thus ?case using pos[unfolded positive_def] by simp
next
case (Suc n)
have "measure M (A n \<union> (\<Union> x<n. A x)) = measure M (A n) + measure M (\<Union> i<n. A i)"
@@ -465,30 +373,25 @@
by auto (metis Int_absorb N disjoint_iff_not_equal lessThan_iff not_leE)
thus ?thesis by simp
qed
- finally show "(\<Sum>\<^isub>\<infinity> n. measure M (A n)) = measure M (\<Union>i. A i)" .
+ finally show "(\<Sum>n. measure M (A n)) = measure M (\<Union>i. A i)" .
qed
qed
qed
lemma (in measure_space) measure_setsum_split:
- assumes "finite r" and "a \<in> sets M" and br_in_M: "b ` r \<subseteq> sets M"
- assumes "(\<Union>i \<in> r. b i) = space M"
- assumes "disjoint_family_on b r"
- shows "\<mu> a = (\<Sum> i \<in> r. \<mu> (a \<inter> (b i)))"
+ assumes "finite S" and "A \<in> sets M" and br_in_M: "B ` S \<subseteq> sets M"
+ assumes "(\<Union>i\<in>S. B i) = space M"
+ assumes "disjoint_family_on B S"
+ shows "\<mu> A = (\<Sum>i\<in>S. \<mu> (A \<inter> (B i)))"
proof -
- have *: "\<mu> a = \<mu> (\<Union>i \<in> r. a \<inter> b i)"
+ have *: "\<mu> A = \<mu> (\<Union>i\<in>S. A \<inter> B i)"
using assms by auto
show ?thesis unfolding *
- proof (rule measure_finitely_additive''[symmetric])
- show "finite r" using `finite r` by auto
- { fix i assume "i \<in> r"
- hence "b i \<in> sets M" using br_in_M by auto
- thus "a \<inter> b i \<in> sets M" using `a \<in> sets M` by auto
- }
- show "disjoint_family_on (\<lambda>i. a \<inter> b i) r"
- using `disjoint_family_on b r`
+ proof (rule measure_setsum[symmetric])
+ show "disjoint_family_on (\<lambda>i. A \<inter> B i) S"
+ using `disjoint_family_on B S`
unfolding disjoint_family_on_def by auto
- qed
+ qed (insert assms, auto)
qed
lemma (in measure_space) measure_subadditive:
@@ -506,7 +409,7 @@
lemma (in measure_space) measure_eq_0:
assumes "N \<in> sets M" and "\<mu> N = 0" and "K \<subseteq> N" and "K \<in> sets M"
shows "\<mu> K = 0"
-using measure_mono[OF assms(3,4,1)] assms(2) by auto
+ using measure_mono[OF assms(3,4,1)] assms(2) positive_measure[OF assms(4)] by auto
lemma (in measure_space) measure_finitely_subadditive:
assumes "finite I" "A ` I \<subseteq> sets M"
@@ -523,35 +426,38 @@
lemma (in measure_space) measure_countably_subadditive:
assumes "range f \<subseteq> sets M"
- shows "\<mu> (\<Union>i. f i) \<le> (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
+ shows "\<mu> (\<Union>i. f i) \<le> (\<Sum>i. \<mu> (f i))"
proof -
have "\<mu> (\<Union>i. f i) = \<mu> (\<Union>i. disjointed f i)"
unfolding UN_disjointed_eq ..
- also have "\<dots> = (\<Sum>\<^isub>\<infinity> i. \<mu> (disjointed f i))"
+ also have "\<dots> = (\<Sum>i. \<mu> (disjointed f i))"
using range_disjointed_sets[OF assms] measure_countably_additive
by (simp add: disjoint_family_disjointed comp_def)
- also have "\<dots> \<le> (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
- proof (rule psuminf_le, rule measure_mono)
- fix i show "disjointed f i \<subseteq> f i" by (rule disjointed_subset)
- show "f i \<in> sets M" "disjointed f i \<in> sets M"
- using assms range_disjointed_sets[OF assms] by auto
- qed
+ also have "\<dots> \<le> (\<Sum>i. \<mu> (f i))"
+ using range_disjointed_sets[OF assms] assms
+ by (auto intro!: suminf_le_pos measure_mono positive_measure disjointed_subset)
finally show ?thesis .
qed
lemma (in measure_space) measure_UN_eq_0:
- assumes "\<And> i :: nat. \<mu> (N i) = 0" and "range N \<subseteq> sets M"
+ assumes "\<And>i::nat. \<mu> (N i) = 0" and "range N \<subseteq> sets M"
shows "\<mu> (\<Union> i. N i) = 0"
-using measure_countably_subadditive[OF assms(2)] assms(1) by auto
+proof -
+ have "0 \<le> \<mu> (\<Union> i. N i)" using assms by auto
+ moreover have "\<mu> (\<Union> i. N i) \<le> 0"
+ using measure_countably_subadditive[OF assms(2)] assms(1) by simp
+ ultimately show ?thesis by simp
+qed
lemma (in measure_space) measure_inter_full_set:
- assumes "S \<in> sets M" "T \<in> sets M" and not_\<omega>: "\<mu> (T - S) \<noteq> \<omega>"
+ assumes "S \<in> sets M" "T \<in> sets M" and fin: "\<mu> (T - S) \<noteq> \<infinity>"
assumes T: "\<mu> T = \<mu> (space M)"
shows "\<mu> (S \<inter> T) = \<mu> S"
proof (rule antisym)
show " \<mu> (S \<inter> T) \<le> \<mu> S"
using assms by (auto intro!: measure_mono)
+ have pos: "0 \<le> \<mu> (T - S)" using assms by auto
show "\<mu> S \<le> \<mu> (S \<inter> T)"
proof (rule ccontr)
assume contr: "\<not> ?thesis"
@@ -560,7 +466,7 @@
also have "\<dots> \<le> \<mu> (T - S) + \<mu> (S \<inter> T)"
using assms by (auto intro!: measure_subadditive)
also have "\<dots> < \<mu> (T - S) + \<mu> S"
- by (rule pextreal_less_add[OF not_\<omega>]) (insert contr, auto)
+ using fin contr pos by (intro extreal_less_add) auto
also have "\<dots> = \<mu> (T \<union> S)"
using assms by (subst measure_additive) auto
also have "\<dots> \<le> \<mu> (space M)"
@@ -572,11 +478,11 @@
lemma measure_unique_Int_stable:
fixes E :: "('a, 'b) algebra_scheme" and A :: "nat \<Rightarrow> 'a set"
assumes "Int_stable E"
- and A: "range A \<subseteq> sets E" "A \<up> space E"
+ and A: "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E"
and M: "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = \<mu>\<rparr>" (is "measure_space ?M")
and N: "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = \<nu>\<rparr>" (is "measure_space ?N")
and eq: "\<And>X. X \<in> sets E \<Longrightarrow> \<mu> X = \<nu> X"
- and finite: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+ and finite: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
assumes "X \<in> sets (sigma E)"
shows "\<mu> X = \<nu> X"
proof -
@@ -585,9 +491,9 @@
where "space ?M = space E" and "sets ?M = sets (sigma E)" and "measure ?M = \<mu>" by (simp_all add: M)
interpret N: measure_space ?N
where "space ?N = space E" and "sets ?N = sets (sigma E)" and "measure ?N = \<nu>" by (simp_all add: N)
- { fix F assume "F \<in> sets E" and "\<mu> F \<noteq> \<omega>"
+ { fix F assume "F \<in> sets E" and "\<mu> F \<noteq> \<infinity>"
then have [intro]: "F \<in> sets (sigma E)" by auto
- have "\<nu> F \<noteq> \<omega>" using `\<mu> F \<noteq> \<omega>` `F \<in> sets E` eq by simp
+ have "\<nu> F \<noteq> \<infinity>" using `\<mu> F \<noteq> \<infinity>` `F \<in> sets E` eq by simp
interpret D: dynkin_system "\<lparr>space=space E, sets=?D F\<rparr>"
proof (rule dynkin_systemI, simp_all)
fix A assume "A \<in> sets (sigma E) \<and> \<mu> (F \<inter> A) = \<nu> (F \<inter> A)"
@@ -602,14 +508,14 @@
and [intro]: "F \<inter> A \<in> sets (sigma E)"
using `F \<in> sets E` M.sets_into_space by auto
have "\<nu> (F \<inter> A) \<le> \<nu> F" by (auto intro!: N.measure_mono)
- then have "\<nu> (F \<inter> A) \<noteq> \<omega>" using `\<nu> F \<noteq> \<omega>` by auto
+ then have "\<nu> (F \<inter> A) \<noteq> \<infinity>" using `\<nu> F \<noteq> \<infinity>` by auto
have "\<mu> (F \<inter> A) \<le> \<mu> F" by (auto intro!: M.measure_mono)
- then have "\<mu> (F \<inter> A) \<noteq> \<omega>" using `\<mu> F \<noteq> \<omega>` by auto
+ then have "\<mu> (F \<inter> A) \<noteq> \<infinity>" using `\<mu> F \<noteq> \<infinity>` by auto
then have "\<mu> (F \<inter> (space (sigma E) - A)) = \<mu> F - \<mu> (F \<inter> A)" unfolding **
using `F \<inter> A \<in> sets (sigma E)` by (auto intro!: M.measure_Diff)
also have "\<dots> = \<nu> F - \<nu> (F \<inter> A)" using eq `F \<in> sets E` * by simp
also have "\<dots> = \<nu> (F \<inter> (space (sigma E) - A))" unfolding **
- using `F \<inter> A \<in> sets (sigma E)` `\<nu> (F \<inter> A) \<noteq> \<omega>` by (auto intro!: N.measure_Diff[symmetric])
+ using `F \<inter> A \<in> sets (sigma E)` `\<nu> (F \<inter> A) \<noteq> \<infinity>` by (auto intro!: N.measure_Diff[symmetric])
finally show "space E - A \<in> sets (sigma E) \<and> \<mu> (F \<inter> (space E - A)) = \<nu> (F \<inter> (space E - A))"
using * by auto
next
@@ -630,15 +536,13 @@
have "\<And>D. D \<in> sets (sigma E) \<Longrightarrow> \<mu> (F \<inter> D) = \<nu> (F \<inter> D)"
by (subst (asm) *) auto }
note * = this
- { fix i have "\<mu> (A i \<inter> X) = \<nu> (A i \<inter> X)"
+ let "?A i" = "A i \<inter> X"
+ have A': "range ?A \<subseteq> sets (sigma E)" "incseq ?A"
+ using A(1,2) `X \<in> sets (sigma E)` by (auto simp: incseq_def)
+ { fix i have "\<mu> (?A i) = \<nu> (?A i)"
using *[of "A i" X] `X \<in> sets (sigma E)` A finite by auto }
- moreover
- have "(\<lambda>i. A i \<inter> X) \<up> X"
- using `X \<in> sets (sigma E)` M.sets_into_space A
- by (auto simp: isoton_def)
- then have "(\<lambda>i. \<mu> (A i \<inter> X)) \<up> \<mu> X" "(\<lambda>i. \<nu> (A i \<inter> X)) \<up> \<nu> X"
- using `X \<in> sets (sigma E)` A by (auto intro!: M.measure_up N.measure_up M.Int simp: subset_eq)
- ultimately show ?thesis by (simp add: isoton_def)
+ with M.continuity_from_below[OF A'] N.continuity_from_below[OF A']
+ show ?thesis using A(3) `X \<in> sets (sigma E)` by auto
qed
section "@{text \<mu>}-null sets"
@@ -650,10 +554,10 @@
shows "N \<union> N' \<in> null_sets"
proof (intro conjI CollectI)
show "N \<union> N' \<in> sets M" using assms by auto
- have "\<mu> (N \<union> N') \<le> \<mu> N + \<mu> N'"
+ then have "0 \<le> \<mu> (N \<union> N')" by simp
+ moreover have "\<mu> (N \<union> N') \<le> \<mu> N + \<mu> N'"
using assms by (intro measure_subadditive) auto
- then show "\<mu> (N \<union> N') = 0"
- using assms by auto
+ ultimately show "\<mu> (N \<union> N') = 0" using assms by auto
qed
lemma UN_from_nat: "(\<Union>i. N i) = (\<Union>i. N (Countable.from_nat i))"
@@ -669,11 +573,11 @@
shows "(\<Union>i. N i) \<in> null_sets"
proof (intro conjI CollectI)
show "(\<Union>i. N i) \<in> sets M" using assms by auto
- have "\<mu> (\<Union>i. N i) \<le> (\<Sum>\<^isub>\<infinity> n. \<mu> (N (Countable.from_nat n)))"
+ then have "0 \<le> \<mu> (\<Union>i. N i)" by simp
+ moreover have "\<mu> (\<Union>i. N i) \<le> (\<Sum>n. \<mu> (N (Countable.from_nat n)))"
unfolding UN_from_nat[of N]
using assms by (intro measure_countably_subadditive) auto
- then show "\<mu> (\<Union>i. N i) = 0"
- using assms by auto
+ ultimately show "\<mu> (\<Union>i. N i) = 0" using assms by auto
qed
lemma (in measure_space) null_sets_finite_UN:
@@ -681,10 +585,10 @@
shows "(\<Union>i\<in>S. A i) \<in> null_sets"
proof (intro CollectI conjI)
show "(\<Union>i\<in>S. A i) \<in> sets M" using assms by (intro finite_UN) auto
- have "\<mu> (\<Union>i\<in>S. A i) \<le> (\<Sum>i\<in>S. \<mu> (A i))"
+ then have "0 \<le> \<mu> (\<Union>i\<in>S. A i)" by simp
+ moreover have "\<mu> (\<Union>i\<in>S. A i) \<le> (\<Sum>i\<in>S. \<mu> (A i))"
using assms by (intro measure_finitely_subadditive) auto
- then show "\<mu> (\<Union>i\<in>S. A i) = 0"
- using assms by auto
+ ultimately show "\<mu> (\<Union>i\<in>S. A i) = 0" using assms by auto
qed
lemma (in measure_space) null_set_Int1:
@@ -731,6 +635,23 @@
almost_everywhere :: "('a \<Rightarrow> bool) \<Rightarrow> bool" (binder "AE " 10) where
"almost_everywhere P \<longleftrightarrow> (\<exists>N\<in>null_sets. { x \<in> space M. \<not> P x } \<subseteq> N)"
+syntax
+ "_almost_everywhere" :: "'a \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
+
+translations
+ "AE x in M. P" == "CONST measure_space.almost_everywhere M (%x. P)"
+
+lemma (in measure_space) AE_cong_measure:
+ assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
+ shows "(AE x in N. P x) \<longleftrightarrow> (AE x. P x)"
+proof -
+ interpret N: measure_space N
+ by (rule measure_space_cong) fact+
+ show ?thesis
+ unfolding N.almost_everywhere_def almost_everywhere_def
+ by (auto simp: assms)
+qed
+
lemma (in measure_space) AE_I':
"N \<in> null_sets \<Longrightarrow> {x\<in>space M. \<not> P x} \<subseteq> N \<Longrightarrow> (AE x. P x)"
unfolding almost_everywhere_def by auto
@@ -741,13 +662,19 @@
proof
assume "AE x. P x" then obtain N where N: "N \<in> sets M" "?P \<subseteq> N" "\<mu> N = 0"
unfolding almost_everywhere_def by auto
+ have "0 \<le> \<mu> ?P" using assms by simp
moreover have "\<mu> ?P \<le> \<mu> N"
using assms N(1,2) by (auto intro: measure_mono)
- ultimately show "?P \<in> null_sets" using assms by auto
+ ultimately have "\<mu> ?P = 0" unfolding `\<mu> N = 0` by auto
+ then show "?P \<in> null_sets" using assms by simp
next
assume "?P \<in> null_sets" with assms show "AE x. P x" by (auto intro: AE_I')
qed
+lemma (in measure_space) AE_iff_measurable:
+ "N \<in> sets M \<Longrightarrow> {x\<in>space M. \<not> P x} = N \<Longrightarrow> (AE x. P x) \<longleftrightarrow> \<mu> N = 0"
+ using AE_iff_null_set[of P] by simp
+
lemma (in measure_space) AE_True[intro, simp]: "AE x. True"
unfolding almost_everywhere_def by auto
@@ -760,13 +687,9 @@
assumes "AE x. P x" "{x\<in>space M. P x} \<in> sets M"
shows "\<mu> {x\<in>space M. \<not> P x} = 0" (is "\<mu> ?P = 0")
proof -
- obtain A where A: "?P \<subseteq> A" "A \<in> sets M" "\<mu> A = 0"
- using assms by (auto elim!: AE_E)
- have "?P = space M - {x\<in>space M. P x}" by auto
- then have "?P \<in> sets M" using assms by auto
- with assms `?P \<subseteq> A` `A \<in> sets M` have "\<mu> ?P \<le> \<mu> A"
- by (auto intro!: measure_mono)
- then show "\<mu> ?P = 0" using A by simp
+ have "{x\<in>space M. \<not> P x} = space M - {x\<in>space M. P x}"
+ by auto
+ with AE_iff_null_set[of P] assms show ?thesis by auto
qed
lemma (in measure_space) AE_I:
@@ -788,8 +711,10 @@
show ?thesis
proof (intro AE_I)
- show "A \<union> B \<in> sets M" "\<mu> (A \<union> B) = 0" using A B
- using measure_subadditive[of A B] by auto
+ have "0 \<le> \<mu> (A \<union> B)" using A B by auto
+ moreover have "\<mu> (A \<union> B) \<le> 0"
+ using measure_subadditive[of A B] A B by auto
+ ultimately show "A \<union> B \<in> sets M" "\<mu> (A \<union> B) = 0" using A B by auto
show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B"
using P imp by auto
qed
@@ -818,8 +743,8 @@
"(\<And>x. x \<in> space M \<Longrightarrow> P x \<longleftrightarrow> Q x) \<Longrightarrow> (AE x. P x) \<longleftrightarrow> (AE x. Q x)"
by auto
-lemma (in measure_space) all_AE_countable:
- "(\<forall>i::'i::countable. AE x. P i x) \<longleftrightarrow> (AE x. \<forall>i. P i x)"
+lemma (in measure_space) AE_all_countable:
+ "(AE x. \<forall>i. P i x) \<longleftrightarrow> (\<forall>i::'i::countable. AE x. P i x)"
proof
assume "\<forall>i. AE x. P i x"
from this[unfolded almost_everywhere_def Bex_def, THEN choice]
@@ -833,6 +758,10 @@
unfolding almost_everywhere_def by auto
qed auto
+lemma (in measure_space) AE_finite_all:
+ assumes f: "finite S" shows "(AE x. \<forall>i\<in>S. P i x) \<longleftrightarrow> (\<forall>i\<in>S. AE x. P i x)"
+ using f by induct auto
+
lemma (in measure_space) restricted_measure_space:
assumes "S \<in> sets M"
shows "measure_space (restricted_space S)"
@@ -840,7 +769,7 @@
unfolding measure_space_def measure_space_axioms_def
proof safe
show "sigma_algebra ?r" using restricted_sigma_algebra[OF assms] .
- show "measure ?r {} = 0" by simp
+ show "positive ?r (measure ?r)" using `S \<in> sets M` by (auto simp: positive_def)
show "countably_additive ?r (measure ?r)"
unfolding countably_additive_def
@@ -848,13 +777,38 @@
fix A :: "nat \<Rightarrow> 'a set"
assume *: "range A \<subseteq> sets ?r" and **: "disjoint_family A"
from restriction_in_sets[OF assms *[simplified]] **
- show "(\<Sum>\<^isub>\<infinity> n. measure ?r (A n)) = measure ?r (\<Union>i. A i)"
+ show "(\<Sum>n. measure ?r (A n)) = measure ?r (\<Union>i. A i)"
using measure_countably_additive by simp
qed
qed
+lemma (in measure_space) AE_restricted:
+ assumes "A \<in> sets M"
+ shows "(AE x in restricted_space A. P x) \<longleftrightarrow> (AE x. x \<in> A \<longrightarrow> P x)"
+proof -
+ interpret R: measure_space "restricted_space A"
+ by (rule restricted_measure_space[OF `A \<in> sets M`])
+ show ?thesis
+ proof
+ assume "AE x in restricted_space A. P x"
+ from this[THEN R.AE_E] guess N' .
+ then obtain N where "{x \<in> A. \<not> P x} \<subseteq> A \<inter> N" "\<mu> (A \<inter> N) = 0" "N \<in> sets M"
+ by auto
+ moreover then have "{x \<in> space M. \<not> (x \<in> A \<longrightarrow> P x)} \<subseteq> A \<inter> N"
+ using `A \<in> sets M` sets_into_space by auto
+ ultimately show "AE x. x \<in> A \<longrightarrow> P x"
+ using `A \<in> sets M` by (auto intro!: AE_I[where N="A \<inter> N"])
+ next
+ assume "AE x. x \<in> A \<longrightarrow> P x"
+ from this[THEN AE_E] guess N .
+ then show "AE x in restricted_space A. P x"
+ using null_set_Int1[OF _ `A \<in> sets M`] `A \<in> sets M`[THEN sets_into_space]
+ by (auto intro!: R.AE_I[where N="A \<inter> N"] simp: subset_eq)
+ qed
+qed
+
lemma (in measure_space) measure_space_subalgebra:
- assumes "sigma_algebra N" and [simp]: "sets N \<subseteq> sets M" "space N = space M"
+ assumes "sigma_algebra N" and "sets N \<subseteq> sets M" "space N = space M"
and measure[simp]: "\<And>X. X \<in> sets N \<Longrightarrow> measure N X = measure M X"
shows "measure_space N"
proof -
@@ -864,13 +818,26 @@
from `sets N \<subseteq> sets M` have "\<And>A. range A \<subseteq> sets N \<Longrightarrow> range A \<subseteq> sets M" by blast
then show "countably_additive N (measure N)"
by (auto intro!: measure_countably_additive simp: countably_additive_def subset_eq)
- qed simp
+ show "positive N (measure_space.measure N)"
+ using assms(2) by (auto simp add: positive_def)
+ qed
+qed
+
+lemma (in measure_space) AE_subalgebra:
+ assumes ae: "AE x in N. P x"
+ and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A"
+ and sa: "sigma_algebra N"
+ shows "AE x. P x"
+proof -
+ interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
+ from ae[THEN N.AE_E] guess N .
+ with N show ?thesis unfolding almost_everywhere_def by auto
qed
section "@{text \<sigma>}-finite Measures"
locale sigma_finite_measure = measure_space +
- assumes sigma_finite: "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<omega>)"
+ assumes sigma_finite: "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<infinity>)"
lemma (in sigma_finite_measure) restricted_sigma_finite_measure:
assumes "S \<in> sets M"
@@ -881,9 +848,9 @@
show "measure_space ?r" using restricted_measure_space[OF assms] .
next
obtain A :: "nat \<Rightarrow> 'a set" where
- "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+ "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
using sigma_finite by auto
- show "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets ?r \<and> (\<Union>i. A i) = space ?r \<and> (\<forall>i. measure ?r (A i) \<noteq> \<omega>)"
+ show "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets ?r \<and> (\<Union>i. A i) = space ?r \<and> (\<forall>i. measure ?r (A i) \<noteq> \<infinity>)"
proof (safe intro!: exI[of _ "\<lambda>i. A i \<inter> S"] del: notI)
fix i
show "A i \<inter> S \<in> sets ?r"
@@ -897,8 +864,7 @@
fix i
have "\<mu> (A i \<inter> S) \<le> \<mu> (A i)"
using `range A \<subseteq> sets M` `S \<in> sets M` by (auto intro!: measure_mono)
- also have "\<dots> < \<omega>" using `\<mu> (A i) \<noteq> \<omega>` by (auto simp: pextreal_less_\<omega>)
- finally show "measure ?r (A i \<inter> S) \<noteq> \<omega>" by (auto simp: pextreal_less_\<omega>)
+ then show "measure ?r (A i \<inter> S) \<noteq> \<infinity>" using `\<mu> (A i) \<noteq> \<infinity>` by auto
qed
qed
@@ -909,7 +875,7 @@
interpret M': measure_space M' by (intro measure_space_cong cong)
from sigma_finite guess A .. note A = this
then have "\<And>i. A i \<in> sets M" by auto
- with A have fin: "(\<forall>i. measure M' (A i) \<noteq> \<omega>)" using cong by auto
+ with A have fin: "\<forall>i. measure M' (A i) \<noteq> \<infinity>" using cong by auto
show ?thesis
apply default
using A fin cong by auto
@@ -917,30 +883,30 @@
lemma (in sigma_finite_measure) disjoint_sigma_finite:
"\<exists>A::nat\<Rightarrow>'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and>
- (\<forall>i. \<mu> (A i) \<noteq> \<omega>) \<and> disjoint_family A"
+ (\<forall>i. \<mu> (A i) \<noteq> \<infinity>) \<and> disjoint_family A"
proof -
obtain A :: "nat \<Rightarrow> 'a set" where
range: "range A \<subseteq> sets M" and
space: "(\<Union>i. A i) = space M" and
- measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+ measure: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
using sigma_finite by auto
note range' = range_disjointed_sets[OF range] range
{ fix i
have "\<mu> (disjointed A i) \<le> \<mu> (A i)"
using range' disjointed_subset[of A i] by (auto intro!: measure_mono)
- then have "\<mu> (disjointed A i) \<noteq> \<omega>"
+ then have "\<mu> (disjointed A i) \<noteq> \<infinity>"
using measure[of i] by auto }
with disjoint_family_disjointed UN_disjointed_eq[of A] space range'
show ?thesis by (auto intro!: exI[of _ "disjointed A"])
qed
lemma (in sigma_finite_measure) sigma_finite_up:
- "\<exists>F. range F \<subseteq> sets M \<and> F \<up> space M \<and> (\<forall>i. \<mu> (F i) \<noteq> \<omega>)"
+ "\<exists>F. range F \<subseteq> sets M \<and> incseq F \<and> (\<Union>i. F i) = space M \<and> (\<forall>i. \<mu> (F i) \<noteq> \<infinity>)"
proof -
obtain F :: "nat \<Rightarrow> 'a set" where
- F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. \<mu> (F i) \<noteq> \<omega>"
+ F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. \<mu> (F i) \<noteq> \<infinity>"
using sigma_finite by auto
- then show ?thesis unfolding isoton_def
+ then show ?thesis
proof (intro exI[of _ "\<lambda>n. \<Union>i\<le>n. F i"] conjI allI)
from F have "\<And>x. x \<in> space M \<Longrightarrow> \<exists>i. x \<in> F i" by auto
then show "(\<Union>n. \<Union> i\<le>n. F i) = space M"
@@ -949,16 +915,16 @@
fix n
have "\<mu> (\<Union> i\<le>n. F i) \<le> (\<Sum>i\<le>n. \<mu> (F i))" using F
by (auto intro!: measure_finitely_subadditive)
- also have "\<dots> < \<omega>"
- using F by (auto simp: setsum_\<omega>)
- finally show "\<mu> (\<Union> i\<le>n. F i) \<noteq> \<omega>" by simp
- qed force+
+ also have "\<dots> < \<infinity>"
+ using F by (auto simp: setsum_Pinfty)
+ finally show "\<mu> (\<Union> i\<le>n. F i) \<noteq> \<infinity>" by simp
+ qed (force simp: incseq_def)+
qed
section {* Measure preserving *}
definition "measure_preserving A B =
- {f \<in> measurable A B. (\<forall>y \<in> sets B. measure A (f -` y \<inter> space A) = measure B y)}"
+ {f \<in> measurable A B. (\<forall>y \<in> sets B. measure B y = measure A (f -` y \<inter> space A))}"
lemma measure_preservingI[intro?]:
assumes "f \<in> measurable A B"
@@ -974,7 +940,8 @@
interpret M': sigma_algebra M' by fact
show ?thesis
proof
- show "measure M' {} = 0" using T by (force simp: measure_preserving_def)
+ show "positive M' (measure M')" using T
+ by (auto simp: measure_preserving_def positive_def measurable_sets)
show "countably_additive M' (measure M')"
proof (intro countably_additiveI)
@@ -986,7 +953,7 @@
using * by blast
moreover have **: "disjoint_family (\<lambda>i. T -` A i \<inter> space M)"
using `disjoint_family A` by (auto simp: disjoint_family_on_def)
- ultimately show "(\<Sum>\<^isub>\<infinity> i. measure M' (A i)) = measure M' (\<Union>i. A i)"
+ ultimately show "(\<Sum>i. measure M' (A i)) = measure M' (\<Union>i. A i)"
using measure_countably_additive[OF _ **] A T
by (auto simp: comp_def vimage_UN measure_preserving_def)
qed
@@ -1009,13 +976,13 @@
lemma measure_unique_Int_stable_vimage:
fixes A :: "nat \<Rightarrow> 'a set"
assumes E: "Int_stable E"
- and A: "range A \<subseteq> sets E" "A \<up> space E" "\<And>i. measure M (A i) \<noteq> \<omega>"
+ and A: "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E" "\<And>i. measure M (A i) \<noteq> \<infinity>"
and N: "measure_space N" "T \<in> measurable N M"
and M: "measure_space M" "sets (sigma E) = sets M" "space E = space M"
and eq: "\<And>X. X \<in> sets E \<Longrightarrow> measure M X = measure N (T -` X \<inter> space N)"
assumes X: "X \<in> sets (sigma E)"
shows "measure M X = measure N (T -` X \<inter> space N)"
-proof (rule measure_unique_Int_stable[OF E A(1,2) _ _ eq _ X])
+proof (rule measure_unique_Int_stable[OF E A(1,2,3) _ _ eq _ X])
interpret M: measure_space M by fact
interpret N: measure_space N by fact
let "?T X" = "T -` X \<inter> space N"
@@ -1028,12 +995,12 @@
show "T \<in> measure_preserving N ?E"
using `T \<in> measurable N M` by (auto simp: M measurable_def measure_preserving_def)
qed
- show "\<And>i. M.\<mu> (A i) \<noteq> \<omega>" by fact
+ show "\<And>i. M.\<mu> (A i) \<noteq> \<infinity>" by fact
qed
lemma (in measure_space) measure_preserving_Int_stable:
fixes A :: "nat \<Rightarrow> 'a set"
- assumes E: "Int_stable E" "range A \<subseteq> sets E" "A \<up> space E" "\<And>i. measure E (A i) \<noteq> \<omega>"
+ assumes E: "Int_stable E" "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E" "\<And>i. measure E (A i) \<noteq> \<infinity>"
and N: "measure_space (sigma E)"
and T: "T \<in> measure_preserving M E"
shows "T \<in> measure_preserving M (sigma E)"
@@ -1046,7 +1013,7 @@
show "\<mu> (T -` X \<inter> space M) = E.\<mu> X"
proof (rule measure_unique_Int_stable_vimage[symmetric])
show "sets (sigma E) = sets (sigma E)" "space E = space (sigma E)"
- "\<And>i. E.\<mu> (A i) \<noteq> \<omega>" using E by auto
+ "\<And>i. E.\<mu> (A i) \<noteq> \<infinity>" using E by auto
show "measure_space M" by default
next
fix X assume "X \<in> sets E" then show "E.\<mu> X = \<mu> (T -` X \<inter> space M)"
@@ -1057,28 +1024,28 @@
section "Real measure values"
lemma (in measure_space) real_measure_Union:
- assumes finite: "\<mu> A \<noteq> \<omega>" "\<mu> B \<noteq> \<omega>"
+ assumes finite: "\<mu> A \<noteq> \<infinity>" "\<mu> B \<noteq> \<infinity>"
and measurable: "A \<in> sets M" "B \<in> sets M" "A \<inter> B = {}"
shows "real (\<mu> (A \<union> B)) = real (\<mu> A) + real (\<mu> B)"
unfolding measure_additive[symmetric, OF measurable]
- using finite by (auto simp: real_of_pextreal_add)
+ using measurable(1,2)[THEN positive_measure]
+ using finite by (cases rule: extreal2_cases[of "\<mu> A" "\<mu> B"]) auto
lemma (in measure_space) real_measure_finite_Union:
assumes measurable:
"finite S" "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M" "disjoint_family_on A S"
- assumes finite: "\<And>i. i \<in> S \<Longrightarrow> \<mu> (A i) \<noteq> \<omega>"
+ assumes finite: "\<And>i. i \<in> S \<Longrightarrow> \<mu> (A i) \<noteq> \<infinity>"
shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
- using real_of_pextreal_setsum[of S, OF finite]
- measure_finitely_additive''[symmetric, OF measurable]
- by simp
+ using finite measurable(2)[THEN positive_measure]
+ by (force intro!: setsum_real_of_extreal[symmetric]
+ simp: measure_setsum[OF measurable, symmetric])
lemma (in measure_space) real_measure_Diff:
- assumes finite: "\<mu> A \<noteq> \<omega>"
+ assumes finite: "\<mu> A \<noteq> \<infinity>"
and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
shows "real (\<mu> (A - B)) = real (\<mu> A) - real (\<mu> B)"
proof -
- have "\<mu> (A - B) \<le> \<mu> A"
- "\<mu> B \<le> \<mu> A"
+ have "\<mu> (A - B) \<le> \<mu> A" "\<mu> B \<le> \<mu> A"
using measurable by (auto intro!: measure_mono)
hence "real (\<mu> ((A - B) \<union> B)) = real (\<mu> (A - B)) + real (\<mu> B)"
using measurable finite by (rule_tac real_measure_Union) auto
@@ -1087,117 +1054,155 @@
lemma (in measure_space) real_measure_UNION:
assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
- assumes finite: "\<mu> (\<Union>i. A i) \<noteq> \<omega>"
+ assumes finite: "\<mu> (\<Union>i. A i) \<noteq> \<infinity>"
shows "(\<lambda>i. real (\<mu> (A i))) sums (real (\<mu> (\<Union>i. A i)))"
proof -
- have *: "(\<Sum>\<^isub>\<infinity> i. \<mu> (A i)) = \<mu> (\<Union>i. A i)"
- using measure_countably_additive[OF measurable] by (simp add: comp_def)
- hence "(\<Sum>\<^isub>\<infinity> i. \<mu> (A i)) \<noteq> \<omega>" using finite by simp
- from psuminf_imp_suminf[OF this]
- show ?thesis using * by simp
+ have "\<And>i. 0 \<le> \<mu> (A i)" using measurable by auto
+ with summable_sums[OF summable_extreal_pos, of "\<lambda>i. \<mu> (A i)"]
+ measure_countably_additive[OF measurable]
+ have "(\<lambda>i. \<mu> (A i)) sums (\<mu> (\<Union>i. A i))" by simp
+ moreover
+ { fix i
+ have "\<mu> (A i) \<le> \<mu> (\<Union>i. A i)"
+ using measurable by (auto intro!: measure_mono)
+ moreover have "0 \<le> \<mu> (A i)" using measurable by auto
+ ultimately have "\<mu> (A i) = extreal (real (\<mu> (A i)))"
+ using finite by (cases "\<mu> (A i)") auto }
+ moreover
+ have "0 \<le> \<mu> (\<Union>i. A i)" using measurable by auto
+ then have "\<mu> (\<Union>i. A i) = extreal (real (\<mu> (\<Union>i. A i)))"
+ using finite by (cases "\<mu> (\<Union>i. A i)") auto
+ ultimately show ?thesis
+ unfolding sums_extreal[symmetric] by simp
qed
lemma (in measure_space) real_measure_subadditive:
assumes measurable: "A \<in> sets M" "B \<in> sets M"
- and fin: "\<mu> A \<noteq> \<omega>" "\<mu> B \<noteq> \<omega>"
+ and fin: "\<mu> A \<noteq> \<infinity>" "\<mu> B \<noteq> \<infinity>"
shows "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
proof -
- have "real (\<mu> (A \<union> B)) \<le> real (\<mu> A + \<mu> B)"
- using measure_subadditive[OF measurable] fin by (auto intro!: real_of_pextreal_mono)
- also have "\<dots> = real (\<mu> A) + real (\<mu> B)"
- using fin by (simp add: real_of_pextreal_add)
- finally show ?thesis .
-qed
-
-lemma (in measure_space) real_measure_countably_subadditive:
- assumes "range f \<subseteq> sets M" and "(\<Sum>\<^isub>\<infinity> i. \<mu> (f i)) \<noteq> \<omega>"
- shows "real (\<mu> (\<Union>i. f i)) \<le> (\<Sum> i. real (\<mu> (f i)))"
-proof -
- have "real (\<mu> (\<Union>i. f i)) \<le> real (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
- using assms by (auto intro!: real_of_pextreal_mono measure_countably_subadditive)
- also have "\<dots> = (\<Sum> i. real (\<mu> (f i)))"
- using assms by (auto intro!: sums_unique psuminf_imp_suminf)
- finally show ?thesis .
+ have "0 \<le> \<mu> (A \<union> B)" using measurable by auto
+ then show "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
+ using measure_subadditive[OF measurable] fin
+ by (cases rule: extreal3_cases[of "\<mu> (A \<union> B)" "\<mu> A" "\<mu> B"]) auto
qed
lemma (in measure_space) real_measure_setsum_singleton:
assumes S: "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
- and fin: "\<And>x. x \<in> S \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
+ and fin: "\<And>x. x \<in> S \<Longrightarrow> \<mu> {x} \<noteq> \<infinity>"
shows "real (\<mu> S) = (\<Sum>x\<in>S. real (\<mu> {x}))"
- using measure_finite_singleton[OF S] fin by (simp add: real_of_pextreal_setsum)
+ using measure_finite_singleton[OF S] fin
+ using positive_measure[OF S(2)]
+ by (force intro!: setsum_real_of_extreal[symmetric])
lemma (in measure_space) real_continuity_from_below:
- assumes A: "range A \<subseteq> sets M" "\<And>i. A i \<subseteq> A (Suc i)" and "\<mu> (\<Union>i. A i) \<noteq> \<omega>"
+ assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "\<mu> (\<Union>i. A i) \<noteq> \<infinity>"
shows "(\<lambda>i. real (\<mu> (A i))) ----> real (\<mu> (\<Union>i. A i))"
-proof (rule SUP_eq_LIMSEQ[THEN iffD1])
- { fix n have "\<mu> (A n) \<le> \<mu> (\<Union>i. A i)"
- using A by (auto intro!: measure_mono)
- hence "\<mu> (A n) \<noteq> \<omega>" using assms by auto }
- note this[simp]
+proof -
+ have "0 \<le> \<mu> (\<Union>i. A i)" using A by auto
+ then have "extreal (real (\<mu> (\<Union>i. A i))) = \<mu> (\<Union>i. A i)"
+ using fin by (auto intro: extreal_real')
+ then show ?thesis
+ using continuity_from_below_Lim[OF A]
+ by (intro lim_real_of_extreal) simp
+qed
- show "mono (\<lambda>i. real (\<mu> (A i)))" unfolding mono_iff_le_Suc using A
- by (auto intro!: real_of_pextreal_mono measure_mono)
-
- show "(SUP n. Real (real (\<mu> (A n)))) = Real (real (\<mu> (\<Union>i. A i)))"
- using continuity_from_below[OF A(1), OF A(2)]
- using assms by (simp add: Real_real)
-qed simp_all
+lemma (in measure_space) continuity_from_above_Lim:
+ assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
+ shows "(\<lambda>i. (\<mu> (A i))) ----> \<mu> (\<Inter>i. A i)"
+ using LIMSEQ_extreal_INFI[OF measure_decseq, OF A]
+ using continuity_from_above[OF A fin] by simp
lemma (in measure_space) real_continuity_from_above:
- assumes A: "range A \<subseteq> sets M"
- and mono_Suc: "\<And>n. A (Suc n) \<subseteq> A n"
- and finite: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+ assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
shows "(\<lambda>n. real (\<mu> (A n))) ----> real (\<mu> (\<Inter>i. A i))"
-proof (rule INF_eq_LIMSEQ[THEN iffD1])
- { fix n have "\<mu> (\<Inter>i. A i) \<le> \<mu> (A n)"
- using A by (auto intro!: measure_mono)
- hence "\<mu> (\<Inter>i. A i) \<noteq> \<omega>" using assms by auto }
- note this[simp]
+proof -
+ have "0 \<le> \<mu> (\<Inter>i. A i)" using A by auto
+ moreover
+ have "\<mu> (\<Inter>i. A i) \<le> \<mu> (A 0)"
+ using A by (auto intro!: measure_mono)
+ ultimately have "extreal (real (\<mu> (\<Inter>i. A i))) = \<mu> (\<Inter>i. A i)"
+ using fin by (auto intro: extreal_real')
+ then show ?thesis
+ using continuity_from_above_Lim[OF A fin]
+ by (intro lim_real_of_extreal) simp
+qed
- show "mono (\<lambda>i. - real (\<mu> (A i)))" unfolding mono_iff_le_Suc using assms
- by (auto intro!: real_of_pextreal_mono measure_mono)
-
- show "(INF n. Real (real (\<mu> (A n)))) =
- Real (real (\<mu> (\<Inter>i. A i)))"
- using continuity_from_above[OF A, OF mono_Suc finite]
- using assms by (simp add: Real_real)
-qed simp_all
+lemma (in measure_space) real_measure_countably_subadditive:
+ assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. \<mu> (A i)) \<noteq> \<infinity>"
+ shows "real (\<mu> (\<Union>i. A i)) \<le> (\<Sum>i. real (\<mu> (A i)))"
+proof -
+ { fix i
+ have "0 \<le> \<mu> (A i)" using A by auto
+ moreover have "\<mu> (A i) \<noteq> \<infinity>" using A by (intro suminf_PInfty[OF _ fin]) auto
+ ultimately have "\<bar>\<mu> (A i)\<bar> \<noteq> \<infinity>" by auto }
+ moreover have "0 \<le> \<mu> (\<Union>i. A i)" using A by auto
+ ultimately have "extreal (real (\<mu> (\<Union>i. A i))) \<le> (\<Sum>i. extreal (real (\<mu> (A i))))"
+ using measure_countably_subadditive[OF A] by (auto simp: extreal_real)
+ also have "\<dots> = extreal (\<Sum>i. real (\<mu> (A i)))"
+ using A
+ by (auto intro!: sums_unique[symmetric] sums_extreal[THEN iffD2] summable_sums summable_real_of_extreal fin)
+ finally show ?thesis by simp
+qed
locale finite_measure = measure_space +
- assumes finite_measure_of_space: "\<mu> (space M) \<noteq> \<omega>"
+ assumes finite_measure_of_space: "\<mu> (space M) \<noteq> \<infinity>"
sublocale finite_measure < sigma_finite_measure
proof
- show "\<exists>A. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<omega>)"
+ show "\<exists>A. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<infinity>)"
using finite_measure_of_space by (auto intro!: exI[of _ "\<lambda>x. space M"])
qed
lemma (in finite_measure) finite_measure[simp, intro]:
assumes "A \<in> sets M"
- shows "\<mu> A \<noteq> \<omega>"
+ shows "\<mu> A \<noteq> \<infinity>"
proof -
from `A \<in> sets M` have "A \<subseteq> space M"
using sets_into_space by blast
- hence "\<mu> A \<le> \<mu> (space M)"
+ then have "\<mu> A \<le> \<mu> (space M)"
using assms top by (rule measure_mono)
- also have "\<dots> < \<omega>"
- using finite_measure_of_space unfolding pextreal_less_\<omega> .
- finally show ?thesis unfolding pextreal_less_\<omega> .
+ then show ?thesis
+ using finite_measure_of_space by auto
qed
+lemma (in finite_measure) measure_not_inf:
+ assumes A: "A \<in> sets M"
+ shows "\<bar>\<mu> A\<bar> \<noteq> \<infinity>"
+ using finite_measure[OF A] positive_measure[OF A] by auto
+
+definition (in finite_measure)
+ "\<mu>' A = (if A \<in> sets M then real (\<mu> A) else 0)"
+
+lemma (in finite_measure) finite_measure_eq: "A \<in> sets M \<Longrightarrow> \<mu> A = extreal (\<mu>' A)"
+ using measure_not_inf[of A] by (auto simp: \<mu>'_def)
+
+lemma (in finite_measure) positive_measure': "0 \<le> \<mu>' A"
+ unfolding \<mu>'_def by (auto simp: real_of_extreal_pos)
+
+lemma (in finite_measure) bounded_measure: "\<mu>' A \<le> \<mu>' (space M)"
+proof cases
+ assume "A \<in> sets M"
+ moreover then have "\<mu> A \<le> \<mu> (space M)"
+ using sets_into_space by (auto intro!: measure_mono)
+ ultimately show ?thesis
+ using measure_not_inf[of A] measure_not_inf[of "space M"]
+ by (auto simp: \<mu>'_def)
+qed (simp add: \<mu>'_def real_of_extreal_pos)
+
lemma (in finite_measure) restricted_finite_measure:
assumes "S \<in> sets M"
shows "finite_measure (restricted_space S)"
(is "finite_measure ?r")
unfolding finite_measure_def finite_measure_axioms_def
-proof (safe del: notI)
+proof (intro conjI)
show "measure_space ?r" using restricted_measure_space[OF assms] .
next
- show "measure ?r (space ?r) \<noteq> \<omega>" using finite_measure[OF `S \<in> sets M`] by auto
+ show "measure ?r (space ?r) \<noteq> \<infinity>" using finite_measure[OF `S \<in> sets M`] by auto
qed
lemma (in measure_space) restricted_to_finite_measure:
- assumes "S \<in> sets M" "\<mu> S \<noteq> \<omega>"
+ assumes "S \<in> sets M" "\<mu> S \<noteq> \<infinity>"
shows "finite_measure (restricted_space S)"
proof -
have "measure_space (restricted_space S)"
@@ -1207,202 +1212,128 @@
using assms by auto
qed
-lemma (in finite_measure) real_finite_measure_Diff:
- assumes "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
- shows "real (\<mu> (A - B)) = real (\<mu> A) - real (\<mu> B)"
- using finite_measure[OF `A \<in> sets M`] by (rule real_measure_Diff[OF _ assms])
-
-lemma (in finite_measure) real_finite_measure_Union:
- assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
- shows "real (\<mu> (A \<union> B)) = real (\<mu> A) + real (\<mu> B)"
- using sets by (auto intro!: real_measure_Union[OF _ _ assms] finite_measure)
+lemma (in finite_measure) finite_measure_Diff:
+ assumes sets: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
+ shows "\<mu>' (A - B) = \<mu>' A - \<mu>' B"
+ using sets[THEN finite_measure_eq]
+ using Diff[OF sets, THEN finite_measure_eq]
+ using measure_Diff[OF _ assms] by simp
-lemma (in finite_measure) real_finite_measure_finite_Union:
- assumes "finite S" and dis: "disjoint_family_on A S"
- and *: "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M"
- shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
-proof (rule real_measure_finite_Union[OF `finite S` _ dis])
- fix i assume "i \<in> S" from *[OF this] show "A i \<in> sets M" .
- from finite_measure[OF this] show "\<mu> (A i) \<noteq> \<omega>" .
-qed
-
-lemma (in finite_measure) real_finite_measure_UNION:
- assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
- shows "(\<lambda>i. real (\<mu> (A i))) sums (real (\<mu> (\<Union>i. A i)))"
-proof (rule real_measure_UNION[OF assms])
- have "(\<Union>i. A i) \<in> sets M" using measurable(1) by (rule countable_UN)
- thus "\<mu> (\<Union>i. A i) \<noteq> \<omega>" by (rule finite_measure)
-qed
-
-lemma (in finite_measure) real_measure_mono:
- "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> A \<subseteq> B \<Longrightarrow> real (\<mu> A) \<le> real (\<mu> B)"
- by (auto intro!: measure_mono real_of_pextreal_mono finite_measure)
-
-lemma (in finite_measure) real_finite_measure_subadditive:
- assumes measurable: "A \<in> sets M" "B \<in> sets M"
- shows "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
- using measurable measurable[THEN finite_measure] by (rule real_measure_subadditive)
+lemma (in finite_measure) finite_measure_Union:
+ assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
+ shows "\<mu>' (A \<union> B) = \<mu>' A + \<mu>' B"
+ using measure_additive[OF assms]
+ using sets[THEN finite_measure_eq]
+ using Un[OF sets, THEN finite_measure_eq]
+ by simp
-lemma (in finite_measure) real_finite_measure_countably_subadditive:
- assumes "range f \<subseteq> sets M" and "summable (\<lambda>i. real (\<mu> (f i)))"
- shows "real (\<mu> (\<Union>i. f i)) \<le> (\<Sum> i. real (\<mu> (f i)))"
-proof (rule real_measure_countably_subadditive[OF assms(1)])
- have *: "\<And>i. f i \<in> sets M" using assms by auto
- have "(\<lambda>i. real (\<mu> (f i))) sums (\<Sum>i. real (\<mu> (f i)))"
- using assms(2) by (rule summable_sums)
- from suminf_imp_psuminf[OF this]
- have "(\<Sum>\<^isub>\<infinity>i. \<mu> (f i)) = Real (\<Sum>i. real (\<mu> (f i)))"
- using * by (simp add: Real_real finite_measure)
- thus "(\<Sum>\<^isub>\<infinity>i. \<mu> (f i)) \<noteq> \<omega>" by simp
-qed
-
-lemma (in finite_measure) real_finite_measure_finite_singelton:
- assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
- shows "real (\<mu> S) = (\<Sum>x\<in>S. real (\<mu> {x}))"
-proof (rule real_measure_setsum_singleton[OF `finite S`])
- fix x assume "x \<in> S" thus "{x} \<in> sets M" by (rule *)
- with finite_measure show "\<mu> {x} \<noteq> \<omega>" .
-qed
+lemma (in finite_measure) finite_measure_finite_Union:
+ assumes S: "finite S" "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M"
+ and dis: "disjoint_family_on A S"
+ shows "\<mu>' (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. \<mu>' (A i))"
+ using measure_setsum[OF assms]
+ using finite_UN[of S A, OF S, THEN finite_measure_eq]
+ using S(2)[THEN finite_measure_eq]
+ by simp
-lemma (in finite_measure) real_finite_continuity_from_below:
- assumes "range A \<subseteq> sets M" "\<And>i. A i \<subseteq> A (Suc i)"
- shows "(\<lambda>i. real (\<mu> (A i))) ----> real (\<mu> (\<Union>i. A i))"
- using real_continuity_from_below[OF assms(1), OF assms(2) finite_measure] assms by auto
-
-lemma (in finite_measure) real_finite_continuity_from_above:
- assumes A: "range A \<subseteq> sets M"
- and mono_Suc: "\<And>n. A (Suc n) \<subseteq> A n"
- shows "(\<lambda>n. real (\<mu> (A n))) ----> real (\<mu> (\<Inter>i. A i))"
- using real_continuity_from_above[OF A, OF mono_Suc finite_measure] A
- by auto
-
-lemma (in finite_measure) real_finite_measure_finite_Union':
- assumes "finite S" "A`S \<subseteq> sets M" "disjoint_family_on A S"
- shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
- using assms real_finite_measure_finite_Union[of S A] by auto
-
-lemma (in finite_measure) finite_measure_compl:
- assumes s: "s \<in> sets M"
- shows "\<mu> (space M - s) = \<mu> (space M) - \<mu> s"
- using measure_compl[OF s, OF finite_measure, OF s] .
-
-lemma (in finite_measure) finite_measure_inter_full_set:
- assumes "S \<in> sets M" "T \<in> sets M"
- assumes T: "\<mu> T = \<mu> (space M)"
- shows "\<mu> (S \<inter> T) = \<mu> S"
- using measure_inter_full_set[OF assms(1,2) finite_measure assms(3)] assms
+lemma (in finite_measure) finite_measure_UNION:
+ assumes A: "range A \<subseteq> sets M" "disjoint_family A"
+ shows "(\<lambda>i. \<mu>' (A i)) sums (\<mu>' (\<Union>i. A i))"
+ using real_measure_UNION[OF A]
+ using countable_UN[OF A(1), THEN finite_measure_eq]
+ using A(1)[THEN subsetD, THEN finite_measure_eq]
by auto
-lemma (in finite_measure) measure_preserving_lift:
- fixes f :: "'a \<Rightarrow> 'c" and A :: "('c, 'd) measure_space_scheme"
- assumes "algebra A" "finite_measure (sigma A)" (is "finite_measure ?sA")
- assumes fmp: "f \<in> measure_preserving M A"
- shows "f \<in> measure_preserving M (sigma A)"
+lemma (in finite_measure) finite_measure_mono:
+ assumes B: "B \<in> sets M" and "A \<subseteq> B" shows "\<mu>' A \<le> \<mu>' B"
+proof cases
+ assume "A \<in> sets M"
+ from this[THEN finite_measure_eq] B[THEN finite_measure_eq]
+ show ?thesis using measure_mono[OF `A \<subseteq> B` `A \<in> sets M` `B \<in> sets M`] by simp
+next
+ assume "A \<notin> sets M" then show ?thesis
+ using positive_measure'[of B] unfolding \<mu>'_def by auto
+qed
+
+lemma (in finite_measure) finite_measure_subadditive:
+ assumes m: "A \<in> sets M" "B \<in> sets M"
+ shows "\<mu>' (A \<union> B) \<le> \<mu>' A + \<mu>' B"
+ using measure_subadditive[OF m]
+ using m[THEN finite_measure_eq] Un[OF m, THEN finite_measure_eq] by simp
+
+lemma (in finite_measure) finite_measure_countably_subadditive:
+ assumes A: "range A \<subseteq> sets M" and sum: "summable (\<lambda>i. \<mu>' (A i))"
+ shows "\<mu>' (\<Union>i. A i) \<le> (\<Sum>i. \<mu>' (A i))"
proof -
- interpret sA: finite_measure ?sA by fact
- interpret A: algebra A by fact
- show ?thesis using fmp
- proof (clarsimp simp add: measure_preserving_def)
- assume fm: "f \<in> measurable M A"
- and "\<forall>y\<in>sets A. \<mu> (f -` y \<inter> space M) = measure A y"
- then have meq: "\<forall>y\<in>sets A. \<mu> (f -` y \<inter> space M) = sA.\<mu> y"
- by simp
- have f12: "f \<in> measurable M ?sA"
- using measurable_subset[OF A.sets_into_space] fm by auto
- hence ffn: "f \<in> space M \<rightarrow> space A"
- by (simp add: measurable_def)
- have "space M \<subseteq> f -` (space A)"
- by auto (metis PiE ffn)
- hence fveq [simp]: "(f -` (space A)) \<inter> space M = space M"
- by blast
- {
- fix y
- assume y: "y \<in> sets ?sA"
- have "sets ?sA = sigma_sets (space A) (sets A)" (is "_ = ?A") by (auto simp: sigma_def)
- also have "\<dots> \<subseteq> {s . \<mu> ((f -` s) \<inter> space M) = sA.\<mu> s}"
- proof (rule A.sigma_property_disjoint, safe)
- fix x assume "x \<in> sets A" then show "\<mu> (f -` x \<inter> space M) = sA.\<mu> x" by (simp add: meq)
- next
- fix s
- assume eq: "\<mu> (f -` s \<inter> space M) = sA.\<mu> s" and s: "s \<in> ?A"
- then have s': "s \<in> sets ?sA" by (simp add: sigma_def)
- show "\<mu> (f -` (space A - s) \<inter> space M) = measure ?sA (space A - s)"
- using sA.finite_measure_compl[OF s']
- using measurable_sets[OF f12 s'] meq[THEN bspec, OF A.top]
- by (simp add: vimage_Diff Diff_Int_distrib2 finite_measure_compl eq)
- next
- fix S
- assume S0: "S 0 = {}"
- and SSuc: "\<And>n. S n \<subseteq> S (Suc n)"
- and rS1: "range S \<subseteq> {s. \<mu> (f -` s \<inter> space M) = sA.\<mu> s} \<inter> ?A"
- hence rS2: "range S \<subseteq> sets ?sA" by (simp add: sigma_def)
- have eq1: "\<And>i. \<mu> (f -` S i \<inter> space M) = sA.\<mu> (S i)"
- using rS1 by blast
- have *: "(\<lambda>n. sA.\<mu> (S n)) = (\<lambda>n. \<mu> (f -` S n \<inter> space M))"
- by (simp add: eq1)
- have "(SUP n. ... n) = \<mu> (\<Union>i. f -` S i \<inter> space M)"
- proof (rule measure_countable_increasing)
- show "range (\<lambda>i. f -` S i \<inter> space M) \<subseteq> sets M"
- using f12 rS2 by (auto simp add: measurable_def)
- show "f -` S 0 \<inter> space M = {}" using S0
- by blast
- show "\<And>n. f -` S n \<inter> space M \<subseteq> f -` S (Suc n) \<inter> space M"
- using SSuc by auto
- qed
- also have "\<mu> (\<Union>i. f -` S i \<inter> space M) = \<mu> (f -` (\<Union>i. S i) \<inter> space M)"
- by (simp add: vimage_UN)
- finally have "(SUP n. sA.\<mu> (S n)) = \<mu> (f -` (\<Union>i. S i) \<inter> space M)" unfolding * .
- moreover
- have "(SUP n. sA.\<mu> (S n)) = sA.\<mu> (\<Union>i. S i)"
- by (rule sA.measure_countable_increasing[OF rS2, OF S0 SSuc])
- ultimately
- show "\<mu> (f -` (\<Union>i. S i) \<inter> space M) = sA.\<mu> (\<Union>i. S i)" by simp
- next
- fix S :: "nat \<Rightarrow> 'c set"
- assume dS: "disjoint_family S"
- and rS1: "range S \<subseteq> {s. \<mu> (f -` s \<inter> space M) = sA.\<mu> s} \<inter> ?A"
- hence rS2: "range S \<subseteq> sets ?sA" by (simp add: sigma_def)
- have "\<And>i. \<mu> (f -` S i \<inter> space M) = sA.\<mu> (S i)"
- using rS1 by blast
- hence *: "(\<lambda>i. sA.\<mu> (S i)) = (\<lambda>n. \<mu> (f -` S n \<inter> space M))"
- by simp
- have "psuminf ... = \<mu> (\<Union>i. f -` S i \<inter> space M)"
- proof (rule measure_countably_additive)
- show "range (\<lambda>i. f -` S i \<inter> space M) \<subseteq> sets M"
- using f12 rS2 by (auto simp add: measurable_def)
- show "disjoint_family (\<lambda>i. f -` S i \<inter> space M)" using dS
- by (auto simp add: disjoint_family_on_def)
- qed
- hence "(\<Sum>\<^isub>\<infinity> i. sA.\<mu> (S i)) = \<mu> (\<Union>i. f -` S i \<inter> space M)" unfolding * .
- with sA.measure_countably_additive [OF rS2 dS]
- have "\<mu> (\<Union>i. f -` S i \<inter> space M) = sA.\<mu> (\<Union>i. S i)"
- by simp
- moreover have "\<mu> (f -` (\<Union>i. S i) \<inter> space M) = \<mu> (\<Union>i. f -` S i \<inter> space M)"
- by (simp add: vimage_UN)
- ultimately show "\<mu> (f -` (\<Union>i. S i) \<inter> space M) = sA.\<mu> (\<Union>i. S i)"
- by metis
- qed
- finally have "sets ?sA \<subseteq> {s . \<mu> ((f -` s) \<inter> space M) = sA.\<mu> s}" .
- hence "\<mu> (f -` y \<inter> space M) = sA.\<mu> y" using y by force
- }
- thus "f \<in> measurable M ?sA \<and> (\<forall>y\<in>sets ?sA. \<mu> (f -` y \<inter> space M) = measure A y)"
- by simp_all (blast intro: f12)
- qed
+ note A[THEN subsetD, THEN finite_measure_eq, simp]
+ note countable_UN[OF A, THEN finite_measure_eq, simp]
+ from `summable (\<lambda>i. \<mu>' (A i))`
+ have "(\<lambda>i. extreal (\<mu>' (A i))) sums extreal (\<Sum>i. \<mu>' (A i))"
+ by (simp add: sums_extreal) (rule summable_sums)
+ from sums_unique[OF this, symmetric]
+ measure_countably_subadditive[OF A]
+ show ?thesis by simp
qed
+lemma (in finite_measure) finite_measure_finite_singleton:
+ assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
+ shows "\<mu>' S = (\<Sum>x\<in>S. \<mu>' {x})"
+ using real_measure_setsum_singleton[OF assms]
+ using *[THEN finite_measure_eq]
+ using finite_UN[of S "\<lambda>x. {x}", OF assms, THEN finite_measure_eq]
+ by simp
+
+lemma (in finite_measure) finite_continuity_from_below:
+ assumes A: "range A \<subseteq> sets M" and "incseq A"
+ shows "(\<lambda>i. \<mu>' (A i)) ----> \<mu>' (\<Union>i. A i)"
+ using real_continuity_from_below[OF A, OF `incseq A` finite_measure] assms
+ using A[THEN subsetD, THEN finite_measure_eq]
+ using countable_UN[OF A, THEN finite_measure_eq]
+ by auto
+
+lemma (in finite_measure) finite_continuity_from_above:
+ assumes A: "range A \<subseteq> sets M" and "decseq A"
+ shows "(\<lambda>n. \<mu>' (A n)) ----> \<mu>' (\<Inter>i. A i)"
+ using real_continuity_from_above[OF A, OF `decseq A` finite_measure] assms
+ using A[THEN subsetD, THEN finite_measure_eq]
+ using countable_INT[OF A, THEN finite_measure_eq]
+ by auto
+
+lemma (in finite_measure) finite_measure_compl:
+ assumes S: "S \<in> sets M"
+ shows "\<mu>' (space M - S) = \<mu>' (space M) - \<mu>' S"
+ using measure_compl[OF S, OF finite_measure, OF S]
+ using S[THEN finite_measure_eq]
+ using compl_sets[OF S, THEN finite_measure_eq]
+ using top[THEN finite_measure_eq]
+ by simp
+
+lemma (in finite_measure) finite_measure_inter_full_set:
+ assumes S: "S \<in> sets M" "T \<in> sets M"
+ assumes T: "\<mu>' T = \<mu>' (space M)"
+ shows "\<mu>' (S \<inter> T) = \<mu>' S"
+ using measure_inter_full_set[OF S finite_measure]
+ using T Diff[OF S(2,1)] Diff[OF S, THEN finite_measure_eq]
+ using Int[OF S, THEN finite_measure_eq]
+ using S[THEN finite_measure_eq] top[THEN finite_measure_eq]
+ by simp
+
+lemma (in finite_measure) empty_measure'[simp]: "\<mu>' {} = 0"
+ unfolding \<mu>'_def by simp
+
section "Finite spaces"
locale finite_measure_space = measure_space + finite_sigma_algebra +
- assumes finite_single_measure[simp]: "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
+ assumes finite_single_measure[simp]: "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<infinity>"
lemma (in finite_measure_space) sum_over_space: "(\<Sum>x\<in>space M. \<mu> {x}) = \<mu> (space M)"
- using measure_finitely_additive''[of "space M" "\<lambda>i. {i}"]
+ using measure_setsum[of "space M" "\<lambda>i. {i}"]
by (simp add: sets_eq_Pow disjoint_family_on_def finite_space)
lemma finite_measure_spaceI:
- assumes "finite (space M)" "sets M = Pow(space M)" and space: "measure M (space M) \<noteq> \<omega>"
+ assumes "finite (space M)" "sets M = Pow(space M)" and space: "measure M (space M) \<noteq> \<infinity>"
and add: "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M B"
- and "measure M {} = 0"
+ and "measure M {} = 0" "\<And>A. A \<subseteq> space M \<Longrightarrow> 0 \<le> measure M A"
shows "finite_measure_space M"
unfolding finite_measure_space_def finite_measure_space_axioms_def
proof (intro allI impI conjI)
@@ -1414,7 +1345,7 @@
using sigma_algebra_Pow[of "space M" "algebra.more M"]
unfolding * .
show "finite (space M)" by fact
- show "positive M (measure M)" unfolding positive_def by fact
+ show "positive M (measure M)" unfolding positive_def using assms by auto
show "additive M (measure M)" unfolding additive_def using assms by simp
qed
then interpret measure_space M .
@@ -1425,19 +1356,20 @@
qed
{ fix x assume *: "x \<in> space M"
with add[of "{x}" "space M - {x}"] space
- show "\<mu> {x} \<noteq> \<omega>" by (auto simp: insert_absorb[OF *] Diff_subset) }
+ show "\<mu> {x} \<noteq> \<infinity>" by (auto simp: insert_absorb[OF *] Diff_subset) }
qed
sublocale finite_measure_space \<subseteq> finite_measure
proof
- show "\<mu> (space M) \<noteq> \<omega>"
- unfolding sum_over_space[symmetric] setsum_\<omega>
+ show "\<mu> (space M) \<noteq> \<infinity>"
+ unfolding sum_over_space[symmetric] setsum_Pinfty
using finite_space finite_single_measure by auto
qed
lemma finite_measure_space_iff:
"finite_measure_space M \<longleftrightarrow>
- finite (space M) \<and> sets M = Pow(space M) \<and> measure M (space M) \<noteq> \<omega> \<and> measure M {} = 0 \<and>
+ finite (space M) \<and> sets M = Pow(space M) \<and> measure M (space M) \<noteq> \<infinity> \<and>
+ measure M {} = 0 \<and> (\<forall>A\<subseteq>space M. 0 \<le> measure M A) \<and>
(\<forall>A\<subseteq>space M. \<forall>B\<subseteq>space M. A \<inter> B = {} \<longrightarrow> measure M (A \<union> B) = measure M A + measure M B)"
(is "_ = ?rhs")
proof (intro iffI)
@@ -1451,29 +1383,31 @@
by (auto intro!: finite_measure_spaceI)
qed
-lemma psuminf_cmult_indicator:
- assumes "disjoint_family A" "x \<in> A i"
- shows "(\<Sum>\<^isub>\<infinity> n. f n * indicator (A n) x) = f i"
+lemma suminf_cmult_indicator:
+ fixes f :: "nat \<Rightarrow> extreal"
+ assumes "disjoint_family A" "x \<in> A i" "\<And>i. 0 \<le> f i"
+ shows "(\<Sum>n. f n * indicator (A n) x) = f i"
proof -
- have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: pextreal)"
+ have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: extreal)"
using `x \<in> A i` assms unfolding disjoint_family_on_def indicator_def by auto
- then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: pextreal)"
+ then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: extreal)"
by (auto simp: setsum_cases)
- moreover have "(SUP n. if i < n then f i else 0) = (f i :: pextreal)"
- proof (rule pextreal_SUPI)
- fix y :: pextreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
+ moreover have "(SUP n. if i < n then f i else 0) = (f i :: extreal)"
+ proof (rule extreal_SUPI)
+ fix y :: extreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
from this[of "Suc i"] show "f i \<le> y" by auto
- qed simp
- ultimately show ?thesis using `x \<in> A i` unfolding psuminf_def by auto
+ qed (insert assms, simp)
+ ultimately show ?thesis using assms
+ by (subst suminf_extreal_eq_SUPR) (auto simp: indicator_def)
qed
-lemma psuminf_indicator:
+lemma suminf_indicator:
assumes "disjoint_family A"
- shows "(\<Sum>\<^isub>\<infinity> n. indicator (A n) x) = indicator (\<Union>i. A i) x"
+ shows "(\<Sum>n. indicator (A n) x :: extreal) = indicator (\<Union>i. A i) x"
proof cases
assume *: "x \<in> (\<Union>i. A i)"
then obtain i where "x \<in> A i" by auto
- from psuminf_cmult_indicator[OF assms, OF this, of "\<lambda>i. 1"]
+ from suminf_cmult_indicator[OF assms(1), OF `x \<in> A i`, of "\<lambda>k. 1"]
show ?thesis using * by simp
qed simp
--- a/src/HOL/Probability/Positive_Extended_Real.thy Mon Mar 14 14:37:47 2011 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2908 +0,0 @@
-(* Author: Johannes Hoelzl, TU Muenchen *)
-
-header {* A type for positive real numbers with infinity *}
-
-theory Positive_Extended_Real
- imports Complex_Main "~~/src/HOL/Library/Nat_Bijection" Multivariate_Analysis
-begin
-
-lemma (in complete_lattice) Sup_start:
- assumes *: "\<And>x. f x \<le> f 0"
- shows "(SUP n. f n) = f 0"
-proof (rule antisym)
- show "f 0 \<le> (SUP n. f n)" by (rule le_SUPI) auto
- show "(SUP n. f n) \<le> f 0" by (rule SUP_leI[OF *])
-qed
-
-lemma (in complete_lattice) Inf_start:
- assumes *: "\<And>x. f 0 \<le> f x"
- shows "(INF n. f n) = f 0"
-proof (rule antisym)
- show "(INF n. f n) \<le> f 0" by (rule INF_leI) simp
- show "f 0 \<le> (INF n. f n)" by (rule le_INFI[OF *])
-qed
-
-lemma (in complete_lattice) Sup_mono_offset:
- fixes f :: "'b :: {ordered_ab_semigroup_add,monoid_add} \<Rightarrow> 'a"
- assumes *: "\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y" and "0 \<le> k"
- shows "(SUP n . f (k + n)) = (SUP n. f n)"
-proof (rule antisym)
- show "(SUP n. f (k + n)) \<le> (SUP n. f n)"
- by (auto intro!: Sup_mono simp: SUPR_def)
- { fix n :: 'b
- have "0 + n \<le> k + n" using `0 \<le> k` by (rule add_right_mono)
- with * have "f n \<le> f (k + n)" by simp }
- thus "(SUP n. f n) \<le> (SUP n. f (k + n))"
- by (auto intro!: Sup_mono exI simp: SUPR_def)
-qed
-
-lemma (in complete_lattice) Sup_mono_offset_Suc:
- assumes *: "\<And>x. f x \<le> f (Suc x)"
- shows "(SUP n . f (Suc n)) = (SUP n. f n)"
- unfolding Suc_eq_plus1
- apply (subst add_commute)
- apply (rule Sup_mono_offset)
- by (auto intro!: order.lift_Suc_mono_le[of "op \<le>" "op <" f, OF _ *]) default
-
-lemma (in complete_lattice) Inf_mono_offset:
- fixes f :: "'b :: {ordered_ab_semigroup_add,monoid_add} \<Rightarrow> 'a"
- assumes *: "\<And>x y. x \<le> y \<Longrightarrow> f y \<le> f x" and "0 \<le> k"
- shows "(INF n . f (k + n)) = (INF n. f n)"
-proof (rule antisym)
- show "(INF n. f n) \<le> (INF n. f (k + n))"
- by (auto intro!: Inf_mono simp: INFI_def)
- { fix n :: 'b
- have "0 + n \<le> k + n" using `0 \<le> k` by (rule add_right_mono)
- with * have "f (k + n) \<le> f n" by simp }
- thus "(INF n. f (k + n)) \<le> (INF n. f n)"
- by (auto intro!: Inf_mono exI simp: INFI_def)
-qed
-
-lemma (in complete_lattice) isotone_converge:
- fixes f :: "nat \<Rightarrow> 'a" assumes "\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y "
- shows "(INF n. SUP m. f (n + m)) = (SUP n. INF m. f (n + m))"
-proof -
- have "\<And>n. (SUP m. f (n + m)) = (SUP n. f n)"
- apply (rule Sup_mono_offset)
- apply (rule assms)
- by simp_all
- moreover
- { fix n have "(INF m. f (n + m)) = f n"
- using Inf_start[of "\<lambda>m. f (n + m)"] assms by simp }
- ultimately show ?thesis by simp
-qed
-
-lemma (in complete_lattice) antitone_converges:
- fixes f :: "nat \<Rightarrow> 'a" assumes "\<And>x y. x \<le> y \<Longrightarrow> f y \<le> f x"
- shows "(INF n. SUP m. f (n + m)) = (SUP n. INF m. f (n + m))"
-proof -
- have "\<And>n. (INF m. f (n + m)) = (INF n. f n)"
- apply (rule Inf_mono_offset)
- apply (rule assms)
- by simp_all
- moreover
- { fix n have "(SUP m. f (n + m)) = f n"
- using Sup_start[of "\<lambda>m. f (n + m)"] assms by simp }
- ultimately show ?thesis by simp
-qed
-
-lemma (in complete_lattice) lim_INF_le_lim_SUP:
- fixes f :: "nat \<Rightarrow> 'a"
- shows "(SUP n. INF m. f (n + m)) \<le> (INF n. SUP m. f (n + m))"
-proof (rule SUP_leI, rule le_INFI)
- fix i j show "(INF m. f (i + m)) \<le> (SUP m. f (j + m))"
- proof (cases rule: le_cases)
- assume "i \<le> j"
- have "(INF m. f (i + m)) \<le> f (i + (j - i))" by (rule INF_leI) simp
- also have "\<dots> = f (j + 0)" using `i \<le> j` by auto
- also have "\<dots> \<le> (SUP m. f (j + m))" by (rule le_SUPI) simp
- finally show ?thesis .
- next
- assume "j \<le> i"
- have "(INF m. f (i + m)) \<le> f (i + 0)" by (rule INF_leI) simp
- also have "\<dots> = f (j + (i - j))" using `j \<le> i` by auto
- also have "\<dots> \<le> (SUP m. f (j + m))" by (rule le_SUPI) simp
- finally show ?thesis .
- qed
-qed
-
-text {*
-
-We introduce the the positive real numbers as needed for measure theory.
-
-*}
-
-typedef pextreal = "(Some ` {0::real..}) \<union> {None}"
- by (rule exI[of _ None]) simp
-
-subsection "Introduce @{typ pextreal} similar to a datatype"
-
-definition "Real x = Abs_pextreal (Some (sup 0 x))"
-definition "\<omega> = Abs_pextreal None"
-
-definition "pextreal_case f i x = (if x = \<omega> then i else f (THE r. 0 \<le> r \<and> x = Real r))"
-
-definition "of_pextreal = pextreal_case (\<lambda>x. x) 0"
-
-defs (overloaded)
- real_of_pextreal_def [code_unfold]: "real == of_pextreal"
-
-lemma pextreal_Some[simp]: "0 \<le> x \<Longrightarrow> Some x \<in> pextreal"
- unfolding pextreal_def by simp
-
-lemma pextreal_Some_sup[simp]: "Some (sup 0 x) \<in> pextreal"
- by (simp add: sup_ge1)
-
-lemma pextreal_None[simp]: "None \<in> pextreal"
- unfolding pextreal_def by simp
-
-lemma Real_inj[simp]:
- assumes "0 \<le> x" and "0 \<le> y"
- shows "Real x = Real y \<longleftrightarrow> x = y"
- unfolding Real_def assms[THEN sup_absorb2]
- using assms by (simp add: Abs_pextreal_inject)
-
-lemma Real_neq_\<omega>[simp]:
- "Real x = \<omega> \<longleftrightarrow> False"
- "\<omega> = Real x \<longleftrightarrow> False"
- by (simp_all add: Abs_pextreal_inject \<omega>_def Real_def)
-
-lemma Real_neg: "x < 0 \<Longrightarrow> Real x = Real 0"
- unfolding Real_def by (auto simp add: Abs_pextreal_inject intro!: sup_absorb1)
-
-lemma pextreal_cases[case_names preal infinite, cases type: pextreal]:
- assumes preal: "\<And>r. x = Real r \<Longrightarrow> 0 \<le> r \<Longrightarrow> P" and inf: "x = \<omega> \<Longrightarrow> P"
- shows P
-proof (cases x rule: pextreal.Abs_pextreal_cases)
- case (Abs_pextreal y)
- hence "y = None \<or> (\<exists>x \<ge> 0. y = Some x)"
- unfolding pextreal_def by auto
- thus P
- proof (rule disjE)
- assume "\<exists>x\<ge>0. y = Some x" then guess x ..
- thus P by (simp add: preal[of x] Real_def Abs_pextreal(1) sup_absorb2)
- qed (simp add: \<omega>_def Abs_pextreal(1) inf)
-qed
-
-lemma pextreal_case_\<omega>[simp]: "pextreal_case f i \<omega> = i"
- unfolding pextreal_case_def by simp
-
-lemma pextreal_case_Real[simp]: "pextreal_case f i (Real x) = (if 0 \<le> x then f x else f 0)"
-proof (cases "0 \<le> x")
- case True thus ?thesis unfolding pextreal_case_def by (auto intro: theI2)
-next
- case False
- moreover have "(THE r. 0 \<le> r \<and> Real 0 = Real r) = 0"
- by (auto intro!: the_equality)
- ultimately show ?thesis unfolding pextreal_case_def by (simp add: Real_neg)
-qed
-
-lemma pextreal_case_cancel[simp]: "pextreal_case (\<lambda>c. i) i x = i"
- by (cases x) simp_all
-
-lemma pextreal_case_split:
- "P (pextreal_case f i x) = ((x = \<omega> \<longrightarrow> P i) \<and> (\<forall>r\<ge>0. x = Real r \<longrightarrow> P (f r)))"
- by (cases x) simp_all
-
-lemma pextreal_case_split_asm:
- "P (pextreal_case f i x) = (\<not> (x = \<omega> \<and> \<not> P i \<or> (\<exists>r. r \<ge> 0 \<and> x = Real r \<and> \<not> P (f r))))"
- by (cases x) auto
-
-lemma pextreal_case_cong[cong]:
- assumes eq: "x = x'" "i = i'" and cong: "\<And>r. 0 \<le> r \<Longrightarrow> f r = f' r"
- shows "pextreal_case f i x = pextreal_case f' i' x'"
- unfolding eq using cong by (cases x') simp_all
-
-lemma real_Real[simp]: "real (Real x) = (if 0 \<le> x then x else 0)"
- unfolding real_of_pextreal_def of_pextreal_def by simp
-
-lemma Real_real_image:
- assumes "\<omega> \<notin> A" shows "Real ` real ` A = A"
-proof safe
- fix x assume "x \<in> A"
- hence *: "x = Real (real x)"
- using `\<omega> \<notin> A` by (cases x) auto
- show "x \<in> Real ` real ` A"
- using `x \<in> A` by (subst *) (auto intro!: imageI)
-next
- fix x assume "x \<in> A"
- thus "Real (real x) \<in> A"
- using `\<omega> \<notin> A` by (cases x) auto
-qed
-
-lemma real_pextreal_nonneg[simp, intro]: "0 \<le> real (x :: pextreal)"
- unfolding real_of_pextreal_def of_pextreal_def
- by (cases x) auto
-
-lemma real_\<omega>[simp]: "real \<omega> = 0"
- unfolding real_of_pextreal_def of_pextreal_def by simp
-
-lemma pextreal_noteq_omega_Ex: "X \<noteq> \<omega> \<longleftrightarrow> (\<exists>r\<ge>0. X = Real r)" by (cases X) auto
-
-subsection "@{typ pextreal} is a monoid for addition"
-
-instantiation pextreal :: comm_monoid_add
-begin
-
-definition "0 = Real 0"
-definition "x + y = pextreal_case (\<lambda>r. pextreal_case (\<lambda>p. Real (r + p)) \<omega> y) \<omega> x"
-
-lemma pextreal_plus[simp]:
- "Real r + Real p = (if 0 \<le> r then if 0 \<le> p then Real (r + p) else Real r else Real p)"
- "x + 0 = x"
- "0 + x = x"
- "x + \<omega> = \<omega>"
- "\<omega> + x = \<omega>"
- by (simp_all add: plus_pextreal_def Real_neg zero_pextreal_def split: pextreal_case_split)
-
-lemma \<omega>_neq_0[simp]:
- "\<omega> = 0 \<longleftrightarrow> False"
- "0 = \<omega> \<longleftrightarrow> False"
- by (simp_all add: zero_pextreal_def)
-
-lemma Real_eq_0[simp]:
- "Real r = 0 \<longleftrightarrow> r \<le> 0"
- "0 = Real r \<longleftrightarrow> r \<le> 0"
- by (auto simp add: Abs_pextreal_inject zero_pextreal_def Real_def sup_real_def)
-
-lemma Real_0[simp]: "Real 0 = 0" by (simp add: zero_pextreal_def)
-
-instance
-proof
- fix a :: pextreal
- show "0 + a = a" by (cases a) simp_all
-
- fix b show "a + b = b + a"
- by (cases a, cases b) simp_all
-
- fix c show "a + b + c = a + (b + c)"
- by (cases a, cases b, cases c) simp_all
-qed
-end
-
-lemma Real_minus_abs[simp]: "Real (- \<bar>x\<bar>) = 0"
- by simp
-
-lemma pextreal_plus_eq_\<omega>[simp]: "(a :: pextreal) + b = \<omega> \<longleftrightarrow> a = \<omega> \<or> b = \<omega>"
- by (cases a, cases b) auto
-
-lemma pextreal_add_cancel_left:
- "a + b = a + c \<longleftrightarrow> (a = \<omega> \<or> b = c)"
- by (cases a, cases b, cases c, simp_all, cases c, simp_all)
-
-lemma pextreal_add_cancel_right:
- "b + a = c + a \<longleftrightarrow> (a = \<omega> \<or> b = c)"
- by (cases a, cases b, cases c, simp_all, cases c, simp_all)
-
-lemma Real_eq_Real:
- "Real a = Real b \<longleftrightarrow> (a = b \<or> (a \<le> 0 \<and> b \<le> 0))"
-proof (cases "a \<le> 0 \<or> b \<le> 0")
- case False with Real_inj[of a b] show ?thesis by auto
-next
- case True
- thus ?thesis
- proof
- assume "a \<le> 0"
- hence *: "Real a = 0" by simp
- show ?thesis using `a \<le> 0` unfolding * by auto
- next
- assume "b \<le> 0"
- hence *: "Real b = 0" by simp
- show ?thesis using `b \<le> 0` unfolding * by auto
- qed
-qed
-
-lemma real_pextreal_0[simp]: "real (0 :: pextreal) = 0"
- unfolding zero_pextreal_def real_Real by simp
-
-lemma real_of_pextreal_eq_0: "real X = 0 \<longleftrightarrow> (X = 0 \<or> X = \<omega>)"
- by (cases X) auto
-
-lemma real_of_pextreal_eq: "real X = real Y \<longleftrightarrow>
- (X = Y \<or> (X = 0 \<and> Y = \<omega>) \<or> (Y = 0 \<and> X = \<omega>))"
- by (cases X, cases Y) (auto simp add: real_of_pextreal_eq_0)
-
-lemma real_of_pextreal_add: "real X + real Y =
- (if X = \<omega> then real Y else if Y = \<omega> then real X else real (X + Y))"
- by (auto simp: pextreal_noteq_omega_Ex)
-
-subsection "@{typ pextreal} is a monoid for multiplication"
-
-instantiation pextreal :: comm_monoid_mult
-begin
-
-definition "1 = Real 1"
-definition "x * y = (if x = 0 \<or> y = 0 then 0 else
- pextreal_case (\<lambda>r. pextreal_case (\<lambda>p. Real (r * p)) \<omega> y) \<omega> x)"
-
-lemma pextreal_times[simp]:
- "Real r * Real p = (if 0 \<le> r \<and> 0 \<le> p then Real (r * p) else 0)"
- "\<omega> * x = (if x = 0 then 0 else \<omega>)"
- "x * \<omega> = (if x = 0 then 0 else \<omega>)"
- "0 * x = 0"
- "x * 0 = 0"
- "1 = \<omega> \<longleftrightarrow> False"
- "\<omega> = 1 \<longleftrightarrow> False"
- by (auto simp add: times_pextreal_def one_pextreal_def)
-
-lemma pextreal_one_mult[simp]:
- "Real x + 1 = (if 0 \<le> x then Real (x + 1) else 1)"
- "1 + Real x = (if 0 \<le> x then Real (1 + x) else 1)"
- unfolding one_pextreal_def by simp_all
-
-instance
-proof
- fix a :: pextreal show "1 * a = a"
- by (cases a) (simp_all add: one_pextreal_def)
-
- fix b show "a * b = b * a"
- by (cases a, cases b) (simp_all add: mult_nonneg_nonneg)
-
- fix c show "a * b * c = a * (b * c)"
- apply (cases a, cases b, cases c)
- apply (simp_all add: mult_nonneg_nonneg not_le mult_pos_pos)
- apply (cases b, cases c)
- apply (simp_all add: mult_nonneg_nonneg not_le mult_pos_pos)
- done
-qed
-end
-
-lemma pextreal_mult_cancel_left:
- "a * b = a * c \<longleftrightarrow> (a = 0 \<or> b = c \<or> (a = \<omega> \<and> b \<noteq> 0 \<and> c \<noteq> 0))"
- by (cases a, cases b, cases c, auto simp: Real_eq_Real mult_le_0_iff, cases c, auto)
-
-lemma pextreal_mult_cancel_right:
- "b * a = c * a \<longleftrightarrow> (a = 0 \<or> b = c \<or> (a = \<omega> \<and> b \<noteq> 0 \<and> c \<noteq> 0))"
- by (cases a, cases b, cases c, auto simp: Real_eq_Real mult_le_0_iff, cases c, auto)
-
-lemma Real_1[simp]: "Real 1 = 1" by (simp add: one_pextreal_def)
-
-lemma real_pextreal_1[simp]: "real (1 :: pextreal) = 1"
- unfolding one_pextreal_def real_Real by simp
-
-lemma real_of_pextreal_mult: "real X * real Y = real (X * Y :: pextreal)"
- by (cases X, cases Y) (auto simp: zero_le_mult_iff)
-
-lemma Real_mult_nonneg: assumes "x \<ge> 0" "y \<ge> 0"
- shows "Real (x * y) = Real x * Real y" using assms by auto
-
-lemma Real_setprod: assumes "\<forall>x\<in>A. f x \<ge> 0" shows "Real (setprod f A) = setprod (\<lambda>x. Real (f x)) A"
-proof(cases "finite A")
- case True thus ?thesis using assms
- proof(induct A) case (insert x A)
- have "0 \<le> setprod f A" apply(rule setprod_nonneg) using insert by auto
- thus ?case unfolding setprod_insert[OF insert(1-2)] apply-
- apply(subst Real_mult_nonneg) prefer 3 apply(subst insert(3)[THEN sym])
- using insert by auto
- qed auto
-qed auto
-
-subsection "@{typ pextreal} is a linear order"
-
-instantiation pextreal :: linorder
-begin
-
-definition "x < y \<longleftrightarrow> pextreal_case (\<lambda>i. pextreal_case (\<lambda>j. i < j) True y) False x"
-definition "x \<le> y \<longleftrightarrow> pextreal_case (\<lambda>j. pextreal_case (\<lambda>i. i \<le> j) False x) True y"
-
-lemma pextreal_less[simp]:
- "Real r < \<omega>"
- "Real r < Real p \<longleftrightarrow> (if 0 \<le> r \<and> 0 \<le> p then r < p else 0 < p)"
- "\<omega> < x \<longleftrightarrow> False"
- "0 < \<omega>"
- "0 < Real r \<longleftrightarrow> 0 < r"
- "x < 0 \<longleftrightarrow> False"
- "0 < (1::pextreal)"
- by (simp_all add: less_pextreal_def zero_pextreal_def one_pextreal_def del: Real_0 Real_1)
-
-lemma pextreal_less_eq[simp]:
- "x \<le> \<omega>"
- "Real r \<le> Real p \<longleftrightarrow> (if 0 \<le> r \<and> 0 \<le> p then r \<le> p else r \<le> 0)"
- "0 \<le> x"
- by (simp_all add: less_eq_pextreal_def zero_pextreal_def del: Real_0)
-
-lemma pextreal_\<omega>_less_eq[simp]:
- "\<omega> \<le> x \<longleftrightarrow> x = \<omega>"
- by (cases x) (simp_all add: not_le less_eq_pextreal_def)
-
-lemma pextreal_less_eq_zero[simp]:
- "(x::pextreal) \<le> 0 \<longleftrightarrow> x = 0"
- by (cases x) (simp_all add: zero_pextreal_def del: Real_0)
-
-instance
-proof
- fix x :: pextreal
- show "x \<le> x" by (cases x) simp_all
- fix y
- show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)"
- by (cases x, cases y) auto
- show "x \<le> y \<or> y \<le> x "
- by (cases x, cases y) auto
- { assume "x \<le> y" "y \<le> x" thus "x = y"
- by (cases x, cases y) auto }
- { fix z assume "x \<le> y" "y \<le> z"
- thus "x \<le> z" by (cases x, cases y, cases z) auto }
-qed
-end
-
-lemma pextreal_zero_lessI[intro]:
- "(a :: pextreal) \<noteq> 0 \<Longrightarrow> 0 < a"
- by (cases a) auto
-
-lemma pextreal_less_omegaI[intro, simp]:
- "a \<noteq> \<omega> \<Longrightarrow> a < \<omega>"
- by (cases a) auto
-
-lemma pextreal_plus_eq_0[simp]: "(a :: pextreal) + b = 0 \<longleftrightarrow> a = 0 \<and> b = 0"
- by (cases a, cases b) auto
-
-lemma pextreal_le_add1[simp, intro]: "n \<le> n + (m::pextreal)"
- by (cases n, cases m) simp_all
-
-lemma pextreal_le_add2: "(n::pextreal) + m \<le> k \<Longrightarrow> m \<le> k"
- by (cases n, cases m, cases k) simp_all
-
-lemma pextreal_le_add3: "(n::pextreal) + m \<le> k \<Longrightarrow> n \<le> k"
- by (cases n, cases m, cases k) simp_all
-
-lemma pextreal_less_\<omega>: "x < \<omega> \<longleftrightarrow> x \<noteq> \<omega>"
- by (cases x) auto
-
-lemma pextreal_0_less_mult_iff[simp]:
- fixes x y :: pextreal shows "0 < x * y \<longleftrightarrow> 0 < x \<and> 0 < y"
- by (cases x, cases y) (auto simp: zero_less_mult_iff)
-
-lemma pextreal_ord_one[simp]:
- "Real p < 1 \<longleftrightarrow> p < 1"
- "Real p \<le> 1 \<longleftrightarrow> p \<le> 1"
- "1 < Real p \<longleftrightarrow> 1 < p"
- "1 \<le> Real p \<longleftrightarrow> 1 \<le> p"
- by (simp_all add: one_pextreal_def del: Real_1)
-
-subsection {* @{text "x - y"} on @{typ pextreal} *}
-
-instantiation pextreal :: minus
-begin
-definition "x - y = (if y < x then THE d. x = y + d else 0 :: pextreal)"
-
-lemma minus_pextreal_eq:
- "(x - y = (z :: pextreal)) \<longleftrightarrow> (if y < x then x = y + z else z = 0)"
- (is "?diff \<longleftrightarrow> ?if")
-proof
- assume ?diff
- thus ?if
- proof (cases "y < x")
- case True
- then obtain p where p: "y = Real p" "0 \<le> p" by (cases y) auto
-
- show ?thesis unfolding `?diff`[symmetric] if_P[OF True] minus_pextreal_def
- proof (rule theI2[where Q="\<lambda>d. x = y + d"])
- show "x = y + pextreal_case (\<lambda>r. Real (r - real y)) \<omega> x" (is "x = y + ?d")
- using `y < x` p by (cases x) simp_all
-
- fix d assume "x = y + d"
- thus "d = ?d" using `y < x` p by (cases d, cases x) simp_all
- qed simp
- qed (simp add: minus_pextreal_def)
-next
- assume ?if
- thus ?diff
- proof (cases "y < x")
- case True
- then obtain p where p: "y = Real p" "0 \<le> p" by (cases y) auto
-
- from True `?if` have "x = y + z" by simp
-
- show ?thesis unfolding minus_pextreal_def if_P[OF True] unfolding `x = y + z`
- proof (rule the_equality)
- fix d :: pextreal assume "y + z = y + d"
- thus "d = z" using `y < x` p
- by (cases d, cases z) simp_all
- qed simp
- qed (simp add: minus_pextreal_def)
-qed
-
-instance ..
-end
-
-lemma pextreal_minus[simp]:
- "Real r - Real p = (if 0 \<le> r \<and> p < r then if 0 \<le> p then Real (r - p) else Real r else 0)"
- "(A::pextreal) - A = 0"
- "\<omega> - Real r = \<omega>"
- "Real r - \<omega> = 0"
- "A - 0 = A"
- "0 - A = 0"
- by (auto simp: minus_pextreal_eq not_less)
-
-lemma pextreal_le_epsilon:
- fixes x y :: pextreal
- assumes "\<And>e. 0 < e \<Longrightarrow> x \<le> y + e"
- shows "x \<le> y"
-proof (cases y)
- case (preal r)
- then obtain p where x: "x = Real p" "0 \<le> p"
- using assms[of 1] by (cases x) auto
- { fix e have "0 < e \<Longrightarrow> p \<le> r + e"
- using assms[of "Real e"] preal x by auto }
- hence "p \<le> r" by (rule field_le_epsilon)
- thus ?thesis using preal x by auto
-qed simp
-
-instance pextreal :: "{ordered_comm_semiring, comm_semiring_1}"
-proof
- show "0 \<noteq> (1::pextreal)" unfolding zero_pextreal_def one_pextreal_def
- by (simp del: Real_1 Real_0)
-
- fix a :: pextreal
- show "0 * a = 0" "a * 0 = 0" by simp_all
-
- fix b c :: pextreal
- show "(a + b) * c = a * c + b * c"
- by (cases c, cases a, cases b)
- (auto intro!: arg_cong[where f=Real] simp: field_simps not_le mult_le_0_iff mult_less_0_iff)
-
- { assume "a \<le> b" thus "c + a \<le> c + b"
- by (cases c, cases a, cases b) auto }
-
- assume "a \<le> b" "0 \<le> c"
- thus "c * a \<le> c * b"
- apply (cases c, cases a, cases b)
- by (auto simp: mult_left_mono mult_le_0_iff mult_less_0_iff not_le)
-qed
-
-lemma mult_\<omega>[simp]: "x * y = \<omega> \<longleftrightarrow> (x = \<omega> \<or> y = \<omega>) \<and> x \<noteq> 0 \<and> y \<noteq> 0"
- by (cases x, cases y) auto
-
-lemma \<omega>_mult[simp]: "(\<omega> = x * y) = ((x = \<omega> \<or> y = \<omega>) \<and> x \<noteq> 0 \<and> y \<noteq> 0)"
- by (cases x, cases y) auto
-
-lemma pextreal_mult_0[simp]: "x * y = 0 \<longleftrightarrow> x = 0 \<or> (y::pextreal) = 0"
- by (cases x, cases y) (auto simp: mult_le_0_iff)
-
-lemma pextreal_mult_cancel:
- fixes x y z :: pextreal
- assumes "y \<le> z"
- shows "x * y \<le> x * z"
- using assms
- by (cases x, cases y, cases z)
- (auto simp: mult_le_cancel_left mult_le_0_iff mult_less_0_iff not_le)
-
-lemma Real_power[simp]:
- "Real x ^ n = (if x \<le> 0 then (if n = 0 then 1 else 0) else Real (x ^ n))"
- by (induct n) auto
-
-lemma Real_power_\<omega>[simp]:
- "\<omega> ^ n = (if n = 0 then 1 else \<omega>)"
- by (induct n) auto
-
-lemma pextreal_of_nat[simp]: "of_nat m = Real (real m)"
- by (induct m) (auto simp: real_of_nat_Suc one_pextreal_def simp del: Real_1)
-
-lemma less_\<omega>_Ex_of_nat: "x < \<omega> \<longleftrightarrow> (\<exists>n. x < of_nat n)"
-proof safe
- assume "x < \<omega>"
- then obtain r where "0 \<le> r" "x = Real r" by (cases x) auto
- moreover obtain n where "r < of_nat n" using ex_less_of_nat by auto
- ultimately show "\<exists>n. x < of_nat n" by (auto simp: real_eq_of_nat)
-qed auto
-
-lemma real_of_pextreal_mono:
- fixes a b :: pextreal
- assumes "b \<noteq> \<omega>" "a \<le> b"
- shows "real a \<le> real b"
-using assms by (cases b, cases a) auto
-
-lemma setprod_pextreal_0:
- "(\<Prod>i\<in>I. f i) = (0::pextreal) \<longleftrightarrow> finite I \<and> (\<exists>i\<in>I. f i = 0)"
-proof cases
- assume "finite I" then show ?thesis
- proof (induct I)
- case (insert i I)
- then show ?case by simp
- qed simp
-qed simp
-
-lemma setprod_\<omega>:
- "(\<Prod>i\<in>I. f i) = \<omega> \<longleftrightarrow> finite I \<and> (\<exists>i\<in>I. f i = \<omega>) \<and> (\<forall>i\<in>I. f i \<noteq> 0)"
-proof cases
- assume "finite I" then show ?thesis
- proof (induct I)
- case (insert i I) then show ?case
- by (auto simp: setprod_pextreal_0)
- qed simp
-qed simp
-
-instance pextreal :: "semiring_char_0"
-proof
- fix m n
- show "inj (of_nat::nat\<Rightarrow>pextreal)" by (auto intro!: inj_onI)
-qed
-
-subsection "@{typ pextreal} is a complete lattice"
-
-instantiation pextreal :: lattice
-begin
-definition [simp]: "sup x y = (max x y :: pextreal)"
-definition [simp]: "inf x y = (min x y :: pextreal)"
-instance proof qed simp_all
-end
-
-instantiation pextreal :: complete_lattice
-begin
-
-definition "bot = Real 0"
-definition "top = \<omega>"
-
-definition "Sup S = (LEAST z. \<forall>x\<in>S. x \<le> z :: pextreal)"
-definition "Inf S = (GREATEST z. \<forall>x\<in>S. z \<le> x :: pextreal)"
-
-lemma pextreal_complete_Sup:
- fixes S :: "pextreal set" assumes "S \<noteq> {}"
- shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>S. y \<le> z) \<longrightarrow> x \<le> z)"
-proof (cases "\<exists>x\<ge>0. \<forall>a\<in>S. a \<le> Real x")
- case False
- hence *: "\<And>x. x\<ge>0 \<Longrightarrow> \<exists>a\<in>S. \<not>a \<le> Real x" by simp
- show ?thesis
- proof (safe intro!: exI[of _ \<omega>])
- fix y assume **: "\<forall>z\<in>S. z \<le> y"
- show "\<omega> \<le> y" unfolding pextreal_\<omega>_less_eq
- proof (rule ccontr)
- assume "y \<noteq> \<omega>"
- then obtain x where [simp]: "y = Real x" and "0 \<le> x" by (cases y) auto
- from *[OF `0 \<le> x`] show False using ** by auto
- qed
- qed simp
-next
- case True then obtain y where y: "\<And>a. a\<in>S \<Longrightarrow> a \<le> Real y" and "0 \<le> y" by auto
- from y[of \<omega>] have "\<omega> \<notin> S" by auto
-
- with `S \<noteq> {}` obtain x where "x \<in> S" and "x \<noteq> \<omega>" by auto
-
- have bound: "\<forall>x\<in>real ` S. x \<le> y"
- proof
- fix z ass