the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
changed syntax for simple_function, simple_integral, positive_integral, integral and RN_deriv.
introduced binder variants for simple_integral, positive_integral and integral.
--- a/src/HOL/Probability/Caratheodory.thy Wed Feb 02 10:35:41 2011 +0100
+++ b/src/HOL/Probability/Caratheodory.thy Wed Feb 02 12:34:45 2011 +0100
@@ -8,77 +8,72 @@
subsection {* Measure Spaces *}
-definition "positive f \<longleftrightarrow> f {} = (0::pextreal)" -- "Positive is enforced by the type"
+record 'a measure_space = "'a algebra" +
+ measure :: "'a set \<Rightarrow> pextreal"
-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 positive where "positive M f \<longleftrightarrow> f {} = (0::pextreal)"
+ -- "Positive is enforced by the type"
-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 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
- 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 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
- 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)"
+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
- 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> psuminf (\<lambda>n. f (A n)))"
+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)"
-definition
- lambda_system where
- "lambda_system M f =
- {l. l \<in> sets M & (\<forall>x \<in> sets M. f (l \<inter> x) + f ((space M - l) \<inter> x) = f x)}"
+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)))"
+
+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}"
-definition
- outer_measure_space where
- "outer_measure_space M f \<longleftrightarrow>
- positive f \<and> increasing M f \<and> countably_subadditive M f"
+definition outer_measure_space where "outer_measure_space M f \<longleftrightarrow>
+ 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}"
-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}"
+locale measure_space = sigma_algebra M for M :: "('a, 'b) measure_space_scheme" +
+ assumes empty_measure [simp]: "measure M {} = 0"
+ and ca: "countably_additive M (measure M)"
-locale measure_space = sigma_algebra +
- fixes \<mu> :: "'a set \<Rightarrow> pextreal"
- assumes empty_measure [simp]: "\<mu> {} = 0"
- and ca: "countably_additive M \<mu>"
+abbreviation (in measure_space) "\<mu> \<equiv> measure M"
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"
+ "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)
lemma subadditiveD:
- "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x\<in>sets M \<Longrightarrow> y\<in>sets M
- \<Longrightarrow> f (x \<union> y) \<le> f x + f y"
+ "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> sets M \<Longrightarrow> y \<in> sets M
+ \<Longrightarrow> f (x \<union> y) \<le> f x + f y"
by (auto simp add: subadditive_def)
lemma additiveD:
- "additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x\<in>sets M \<Longrightarrow> y\<in>sets M
- \<Longrightarrow> f (x \<union> y) = f x + f y"
+ "additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> sets M \<Longrightarrow> y \<in> sets M
+ \<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)"
+ \<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)
section "Extend binary sets"
@@ -109,7 +104,8 @@
by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f])
lemma suminf_binaryset_eq:
- "f {} = 0 \<Longrightarrow> suminf (\<lambda>n. f (binaryset A B n)) = f A + f B"
+ fixes f :: "'a set \<Rightarrow> real"
+ shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
by (metis binaryset_sums sums_unique)
lemma binaryset_psuminf:
@@ -130,8 +126,8 @@
subsection {* Lambda Systems *}
lemma (in algebra) lambda_system_eq:
- "lambda_system M f =
- {l. l \<in> sets M & (\<forall>x \<in> sets M. f (x \<inter> l) + f (x - l) = f x)}"
+ shows "lambda_system M f = {l \<in> sets M.
+ \<forall>x \<in> sets M. f (x \<inter> l) + f (x - l) = f x}"
proof -
have [simp]: "!!l x. l \<in> sets M \<Longrightarrow> x \<in> sets M \<Longrightarrow> (space M - l) \<inter> x = x - l"
by (metis Int_Diff Int_absorb1 Int_commute sets_into_space)
@@ -140,60 +136,59 @@
qed
lemma (in algebra) lambda_system_empty:
- "positive f \<Longrightarrow> {} \<in> lambda_system M f"
- by (auto simp add: positive_def lambda_system_eq)
+ "positive M f \<Longrightarrow> {} \<in> lambda_system M f"
+ by (auto simp add: positive_def lambda_system_eq algebra_def)
lemma lambda_system_sets:
- "x \<in> lambda_system M f \<Longrightarrow> x \<in> sets M"
- by (simp add: lambda_system_def)
+ "x \<in> lambda_system M f \<Longrightarrow> x \<in> sets M"
+ by (simp add: lambda_system_def)
lemma (in algebra) lambda_system_Compl:
fixes f:: "'a set \<Rightarrow> pextreal"
assumes x: "x \<in> lambda_system M f"
shows "space M - x \<in> lambda_system M f"
- proof -
- have "x \<subseteq> space M"
- by (metis sets_into_space lambda_system_sets x)
- hence "space M - (space M - x) = x"
- by (metis double_diff equalityE)
- with x show ?thesis
- by (force simp add: lambda_system_def ac_simps)
- qed
+proof -
+ have "x \<subseteq> space M"
+ by (metis sets_into_space lambda_system_sets x)
+ hence "space M - (space M - x) = x"
+ by (metis double_diff equalityE)
+ with x show ?thesis
+ by (force simp add: lambda_system_def ac_simps)
+qed
lemma (in algebra) lambda_system_Int:
fixes f:: "'a set \<Rightarrow> pextreal"
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 -
- from xl yl show ?thesis
- proof (auto simp add: positive_def lambda_system_eq Int)
- fix u
- assume x: "x \<in> sets M" and y: "y \<in> sets M" and u: "u \<in> sets M"
- and fx: "\<forall>z\<in>sets M. f (z \<inter> x) + f (z - x) = f z"
- and fy: "\<forall>z\<in>sets M. f (z \<inter> y) + f (z - y) = f z"
- have "u - x \<inter> y \<in> sets M"
- by (metis Diff Diff_Int Un u x y)
- moreover
- have "(u - (x \<inter> y)) \<inter> y = u \<inter> y - x" by blast
- moreover
- have "u - x \<inter> y - y = u - y" by blast
- ultimately
- have ey: "f (u - x \<inter> y) = f (u \<inter> y - x) + f (u - y)" using fy
- by force
- have "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y)
- = (f (u \<inter> (x \<inter> y)) + f (u \<inter> y - x)) + f (u - y)"
- by (simp add: ey ac_simps)
- also have "... = (f ((u \<inter> y) \<inter> x) + f (u \<inter> y - x)) + f (u - y)"
- by (simp add: Int_ac)
- also have "... = f (u \<inter> y) + f (u - y)"
- using fx [THEN bspec, of "u \<inter> y"] Int y u
- by force
- also have "... = f u"
- by (metis fy u)
- finally show "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y) = f u" .
- qed
+proof -
+ from xl yl show ?thesis
+ proof (auto simp add: positive_def lambda_system_eq Int)
+ fix u
+ assume x: "x \<in> sets M" and y: "y \<in> sets M" and u: "u \<in> sets M"
+ and fx: "\<forall>z\<in>sets M. f (z \<inter> x) + f (z - x) = f z"
+ and fy: "\<forall>z\<in>sets M. f (z \<inter> y) + f (z - y) = f z"
+ have "u - x \<inter> y \<in> sets M"
+ by (metis Diff Diff_Int Un u x y)
+ moreover
+ have "(u - (x \<inter> y)) \<inter> y = u \<inter> y - x" by blast
+ moreover
+ have "u - x \<inter> y - y = u - y" by blast
+ ultimately
+ have ey: "f (u - x \<inter> y) = f (u \<inter> y - x) + f (u - y)" using fy
+ by force
+ have "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y)
+ = (f (u \<inter> (x \<inter> y)) + f (u \<inter> y - x)) + f (u - y)"
+ by (simp add: ey ac_simps)
+ also have "... = (f ((u \<inter> y) \<inter> x) + f (u \<inter> y - x)) + f (u - y)"
+ by (simp add: Int_ac)
+ also have "... = f (u \<inter> y) + f (u - y)"
+ using fx [THEN bspec, of "u \<inter> y"] Int y u
+ by force
+ also have "... = f u"
+ by (metis fy u)
+ finally show "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y) = f u" .
qed
-
+qed
lemma (in algebra) lambda_system_Un:
fixes f:: "'a set \<Rightarrow> pextreal"
@@ -210,7 +205,7 @@
qed
lemma (in algebra) lambda_system_algebra:
- "positive f \<Longrightarrow> algebra (M (|sets := lambda_system M f|))"
+ "positive M f \<Longrightarrow> algebra (M\<lparr>sets := lambda_system M f\<rparr>)"
apply (auto simp add: algebra_def)
apply (metis lambda_system_sets set_mp sets_into_space)
apply (metis lambda_system_empty)
@@ -222,32 +217,31 @@
assumes z: "z \<in> sets M" and disj: "x \<inter> y = {}"
and xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
shows "f (z \<inter> (x \<union> y)) = f (z \<inter> x) + f (z \<inter> y)"
- proof -
- have "z \<inter> x = (z \<inter> (x \<union> y)) \<inter> x" using disj by blast
- moreover
- have "z \<inter> y = (z \<inter> (x \<union> y)) - x" using disj by blast
- moreover
- have "(z \<inter> (x \<union> y)) \<in> sets M"
- by (metis Int Un lambda_system_sets xl yl z)
- ultimately show ?thesis using xl yl
- by (simp add: lambda_system_eq)
- qed
+proof -
+ have "z \<inter> x = (z \<inter> (x \<union> y)) \<inter> x" using disj by blast
+ moreover
+ have "z \<inter> y = (z \<inter> (x \<union> y)) - x" using disj by blast
+ moreover
+ have "(z \<inter> (x \<union> y)) \<in> sets M"
+ by (metis Int Un lambda_system_sets xl yl z)
+ ultimately show ?thesis using xl yl
+ by (simp add: lambda_system_eq)
+qed
lemma (in algebra) lambda_system_additive:
"additive (M (|sets := lambda_system M f|)) f"
- proof (auto simp add: additive_def)
- fix x and y
- assume disj: "x \<inter> y = {}"
- and xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
- hence "x \<in> sets M" "y \<in> sets M" by (blast intro: lambda_system_sets)+
- thus "f (x \<union> y) = f x + f y"
- using lambda_system_strong_additive [OF top disj xl yl]
- by (simp add: Un)
- qed
-
+proof (auto simp add: additive_def)
+ fix x and y
+ assume disj: "x \<inter> y = {}"
+ and xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
+ hence "x \<in> sets M" "y \<in> sets M" by (blast intro: lambda_system_sets)+
+ thus "f (x \<union> y) = f x + f y"
+ using lambda_system_strong_additive [OF top disj xl yl]
+ by (simp add: Un)
+qed
lemma (in algebra) countably_subadditive_subadditive:
- assumes f: "positive f" and cs: "countably_subadditive M f"
+ assumes f: "positive M f" and cs: "countably_subadditive M f"
shows "subadditive M f"
proof (auto simp add: subadditive_def)
fix x y
@@ -267,7 +261,7 @@
lemma (in algebra) additive_sum:
fixes A:: "nat \<Rightarrow> 'a set"
- assumes f: "positive f" and ad: "additive M f"
+ assumes f: "positive M f" and ad: "additive M f"
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)"
@@ -288,15 +282,9 @@
by (auto simp add: atLeastLessThanSuc additive_def)
qed
-
-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 (in algebra) increasing_additive_bound:
fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> pextreal"
- assumes f: "positive f" and ad: "additive M f"
+ 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"
@@ -311,12 +299,16 @@
qed
lemma lambda_system_increasing:
- "increasing M f \<Longrightarrow> increasing (M (|sets := lambda_system M f|)) f"
+ "increasing M f \<Longrightarrow> increasing (M (|sets := lambda_system M f|)) f"
by (simp add: increasing_def lambda_system_def)
+lemma lambda_system_positive:
+ "positive M f \<Longrightarrow> positive (M (|sets := lambda_system M f|)) f"
+ 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"
- assumes f: "positive f" and a: "a \<in> sets M"
+ 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"
shows "(\<Sum>i = 0..<n. f (a \<inter>A i)) = f (a \<inter> (\<Union>i\<in>{0..<n}. A i))"
@@ -335,14 +327,13 @@
by (simp add: atLeastLessThanSuc lambda_system_strong_additive [OF a 2 3 4])
qed
-
lemma (in sigma_algebra) lambda_system_caratheodory:
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)"
proof -
- have pos: "positive f" and inc: "increasing M f"
+ have pos: "positive M f" and inc: "increasing M f"
and csa: "countably_subadditive M f"
by (metis oms outer_measure_space_def)+
have sa: "subadditive M f"
@@ -357,15 +348,15 @@
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)"
- 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_additive subset_Un_eq increasingD [OF inc]
- A' A'' UNION_in_sets U_in)
- qed
+ 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)
+ qed
{
fix a
assume a [iff]: "a \<in> sets M"
@@ -424,19 +415,20 @@
lemma (in sigma_algebra) caratheodory_lemma:
assumes oms: "outer_measure_space M f"
- shows "measure_space (|space = space M, sets = lambda_system M f|) f"
+ shows "measure_space \<lparr> space = space M, sets = lambda_system M f, measure = f \<rparr>"
+ (is "measure_space ?M")
proof -
- have pos: "positive f"
+ have pos: "positive M f"
by (metis oms outer_measure_space_def)
- have alg: "algebra (|space = space M, sets = lambda_system M f|)"
+ have alg: "algebra ?M"
using lambda_system_algebra [of f, OF pos]
by (simp add: algebra_def)
then moreover
- have "sigma_algebra (|space = space M, sets = lambda_system M f|)"
+ have "sigma_algebra ?M"
using lambda_system_caratheodory [OF oms]
by (simp add: sigma_algebra_disjoint_iff)
moreover
- have "measure_space_axioms (|space = space M, sets = lambda_system M f|) f"
+ have "measure_space_axioms ?M"
using pos lambda_system_caratheodory [OF oms]
by (simp add: measure_space_axioms_def positive_def lambda_system_sets
countably_additive_def o_def)
@@ -446,7 +438,7 @@
qed
lemma (in algebra) additive_increasing:
- assumes posf: "positive f" and addf: "additive M f"
+ assumes posf: "positive M f" and addf: "additive M f"
shows "increasing M f"
proof (auto simp add: increasing_def)
fix x y
@@ -460,7 +452,7 @@
qed
lemma (in algebra) countably_additive_additive:
- assumes posf: "positive f" and ca: "countably_additive M f"
+ assumes posf: "positive M f" and ca: "countably_additive M f"
shows "additive M f"
proof (auto simp add: additive_def)
fix x y
@@ -480,7 +472,7 @@
qed
lemma inf_measure_nonempty:
- assumes f: "positive f" and b: "b \<in> sets M" and a: "a \<subseteq> b" "{} \<in> sets M"
+ 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}"
@@ -494,7 +486,7 @@
qed
lemma (in algebra) inf_measure_agrees:
- assumes posf: "positive f" and ca: "countably_additive M f"
+ 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
@@ -530,25 +522,25 @@
fix y
assume y: "\<forall>u \<in> measure_set M f s. y \<le> u"
thus "y \<le> f s"
- by (blast intro: inf_measure_nonempty [of f, OF posf s subset_refl])
+ by (blast intro: inf_measure_nonempty [of _ f, OF posf s subset_refl])
qed
-lemma (in algebra) inf_measure_empty:
- assumes posf: "positive f" "{} \<in> sets M"
+lemma inf_measure_empty:
+ assumes posf: "positive M f" "{} \<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])
+ by (metis complete_lattice_class.Inf_lower `{} \<in> sets M`
+ inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def])
qed simp
lemma (in algebra) inf_measure_positive:
- "positive f \<Longrightarrow>
- positive (\<lambda>x. Inf (measure_set M f x))"
- by (simp add: positive_def inf_measure_empty)
+ "positive M f \<Longrightarrow> positive M (\<lambda>x. Inf (measure_set M f x))"
+ by (simp add: positive_def inf_measure_empty)
lemma (in algebra) inf_measure_increasing:
- assumes posf: "positive f"
- shows "increasing (| space = space M, sets = Pow (space M) |)
+ assumes posf: "positive M f"
+ shows "increasing \<lparr> space = space M, sets = Pow (space M) \<rparr>
(\<lambda>x. Inf (measure_set M f x))"
apply (auto simp add: increasing_def)
apply (rule complete_lattice_class.Inf_greatest)
@@ -558,7 +550,7 @@
lemma (in algebra) inf_measure_le:
- assumes posf: "positive f" and inc: "increasing M f"
+ 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}"
shows "Inf (measure_set M f s) \<le> x"
proof -
@@ -584,7 +576,7 @@
qed
lemma (in algebra) inf_measure_close:
- assumes posf: "positive f" and e: "0 < e" and ss: "s \<subseteq> (space M)"
+ 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>")
@@ -596,7 +588,7 @@
next
case True
have "measure_set M f s \<noteq> {}"
- by (metis emptyE ss inf_measure_nonempty [of f, OF posf top _ empty_sets])
+ by (metis emptyE ss inf_measure_nonempty [of _ f, OF posf top _ empty_sets])
then obtain l where "l \<in> measure_set M f s" by auto
moreover from True have "l \<le> Inf (measure_set M f s) + e" by simp
ultimately show ?thesis
@@ -604,7 +596,7 @@
qed
lemma (in algebra) inf_measure_countably_subadditive:
- assumes posf: "positive f" and inc: "increasing M f"
+ assumes posf: "positive M f" and inc: "increasing M f"
shows "countably_subadditive (| space = space M, sets = Pow (space M) |)
(\<lambda>x. Inf (measure_set M f x))"
unfolding countably_subadditive_def o_def
@@ -666,8 +658,8 @@
qed
lemma (in algebra) inf_measure_outer:
- "\<lbrakk> positive f ; increasing M f \<rbrakk>
- \<Longrightarrow> outer_measure_space (| space = space M, sets = Pow (space M) |)
+ "\<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))"
by (simp add: outer_measure_space_def inf_measure_empty
inf_measure_increasing inf_measure_countably_subadditive positive_def)
@@ -675,7 +667,7 @@
(*MOVE UP*)
lemma (in algebra) algebra_subset_lambda_system:
- assumes posf: "positive f" and inc: "increasing M f"
+ assumes posf: "positive M f" and inc: "increasing M f"
and add: "additive M f"
shows "sets M \<subseteq> lambda_system (| space = space M, sets = Pow (space M) |)
(\<lambda>x. Inf (measure_set M f x))"
@@ -739,10 +731,10 @@
by (metis Un_Diff_Int Un_commute)
also have "... \<le> Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))"
apply (rule subadditiveD)
- apply (iprover intro: algebra.countably_subadditive_subadditive algebra_Pow
- inf_measure_positive inf_measure_countably_subadditive posf inc)
- apply (auto simp add: subsetD [OF s])
- done
+ apply (rule algebra.countably_subadditive_subadditive[OF algebra_Pow])
+ apply (simp add: positive_def inf_measure_empty[OF posf])
+ apply (rule inf_measure_countably_subadditive)
+ using s by (auto intro!: posf inc)
finally show ?thesis .
qed
ultimately
@@ -752,37 +744,38 @@
qed
lemma measure_down:
- "measure_space N \<mu> \<Longrightarrow> sigma_algebra M \<Longrightarrow> sets M \<subseteq> sets N \<Longrightarrow>
- (\<nu> = \<mu>) \<Longrightarrow> measure_space M \<nu>"
+ "measure_space N \<Longrightarrow> sigma_algebra M \<Longrightarrow> sets M \<subseteq> sets N \<Longrightarrow> measure N = measure M \<Longrightarrow> measure_space M"
by (simp add: measure_space_def measure_space_axioms_def positive_def
countably_additive_def)
blast
theorem (in algebra) caratheodory:
- assumes posf: "positive f" and ca: "countably_additive M f"
- shows "\<exists>\<mu> :: 'a set \<Rightarrow> pextreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and> measure_space (sigma M) \<mu>"
- proof -
- have inc: "increasing M f"
- by (metis additive_increasing ca countably_additive_additive posf)
- let ?infm = "(\<lambda>x. Inf (measure_set M f x))"
- def ls \<equiv> "lambda_system (|space = space M, sets = Pow (space M)|) ?infm"
- have mls: "measure_space \<lparr>space = space M, sets = ls\<rparr> ?infm"
- using sigma_algebra.caratheodory_lemma
- [OF sigma_algebra_Pow inf_measure_outer [OF posf inc]]
- by (simp add: ls_def)
- hence sls: "sigma_algebra (|space = space M, sets = ls|)"
- by (simp add: measure_space_def)
- have "sets M \<subseteq> ls"
- by (simp add: ls_def)
- (metis ca posf inc countably_additive_additive algebra_subset_lambda_system)
- hence sgs_sb: "sigma_sets (space M) (sets M) \<subseteq> ls"
- using sigma_algebra.sigma_sets_subset [OF sls, of "sets M"]
- by simp
- have "measure_space (sigma M) ?infm"
- unfolding sigma_def
- by (rule measure_down [OF mls], rule sigma_algebra_sigma_sets)
- (simp_all add: sgs_sb space_closed)
- thus ?thesis using inf_measure_agrees [OF posf ca] by (auto intro!: exI[of _ ?infm])
- qed
+ 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>
+ measure_space \<lparr> space = space M, sets = sets (sigma M), measure = \<mu> \<rparr>"
+proof -
+ have inc: "increasing M f"
+ by (metis additive_increasing ca countably_additive_additive posf)
+ let ?infm = "(\<lambda>x. Inf (measure_set M f x))"
+ def ls \<equiv> "lambda_system (|space = space M, sets = Pow (space M)|) ?infm"
+ have mls: "measure_space \<lparr>space = space M, sets = ls, measure = ?infm\<rparr>"
+ using sigma_algebra.caratheodory_lemma
+ [OF sigma_algebra_Pow inf_measure_outer [OF posf inc]]
+ by (simp add: ls_def)
+ hence sls: "sigma_algebra (|space = space M, sets = ls, measure = ?infm|)"
+ by (simp add: measure_space_def)
+ have "sets M \<subseteq> ls"
+ by (simp add: ls_def)
+ (metis ca posf inc countably_additive_additive algebra_subset_lambda_system)
+ hence sgs_sb: "sigma_sets (space M) (sets M) \<subseteq> ls"
+ using sigma_algebra.sigma_sets_subset [OF sls, of "sets M"]
+ by simp
+ have "measure_space \<lparr> space = space M, sets = sets (sigma M), measure = ?infm \<rparr>"
+ unfolding sigma_def
+ by (rule measure_down [OF mls], rule sigma_algebra_sigma_sets)
+ (simp_all add: sgs_sb space_closed)
+ thus ?thesis using inf_measure_agrees [OF posf ca]
+ by (intro exI[of _ ?infm]) auto
+qed
end
--- a/src/HOL/Probability/Complete_Measure.thy Wed Feb 02 10:35:41 2011 +0100
+++ b/src/HOL/Probability/Complete_Measure.thy Wed Feb 02 12:34:45 2011 +0100
@@ -7,9 +7,24 @@
locale completeable_measure_space = measure_space
-definition (in completeable_measure_space) completion :: "'a algebra" where
+definition (in completeable_measure_space)
+ "split_completion A p = (\<exists>N'. A = fst p \<union> snd p \<and> fst p \<inter> snd p = {} \<and>
+ fst p \<in> sets M \<and> snd p \<subseteq> N' \<and> N' \<in> null_sets)"
+
+definition (in completeable_measure_space)
+ "main_part A = fst (Eps (split_completion A))"
+
+definition (in completeable_measure_space)
+ "null_part A = snd (Eps (split_completion A))"
+
+abbreviation (in completeable_measure_space) "\<mu>' A \<equiv> \<mu> (main_part A)"
+
+definition (in completeable_measure_space) completion :: "('a, 'b) measure_space_scheme" where
"completion = \<lparr> space = space M,
- sets = { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets \<and> N \<subseteq> N' } \<rparr>"
+ sets = { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets \<and> N \<subseteq> N' },
+ measure = \<mu>',
+ \<dots> = more M \<rparr>"
+
lemma (in completeable_measure_space) space_completion[simp]:
"space completion = space M" unfolding completion_def by simp
@@ -58,16 +73,6 @@
auto
qed auto
-definition (in completeable_measure_space)
- "split_completion A p = (\<exists>N'. A = fst p \<union> snd p \<and> fst p \<inter> snd p = {} \<and>
- fst p \<in> sets M \<and> snd p \<subseteq> N' \<and> N' \<in> null_sets)"
-
-definition (in completeable_measure_space)
- "main_part A = fst (Eps (split_completion A))"
-
-definition (in completeable_measure_space)
- "null_part A = snd (Eps (split_completion A))"
-
lemma (in completeable_measure_space) split_completion:
assumes "A \<in> sets completion"
shows "split_completion A (main_part A, null_part A)"
@@ -108,17 +113,15 @@
show "\<mu> (null_part S) = 0" by auto
qed
-definition (in completeable_measure_space) "\<mu>' A = \<mu> (main_part A)"
-
lemma (in completeable_measure_space) \<mu>'_set[simp]:
assumes "S \<in> sets M" shows "\<mu>' S = \<mu> S"
proof -
have S: "S \<in> sets completion" using assms by auto
then have "\<mu> S = \<mu> (main_part S \<union> null_part S)" by simp
- also have "\<dots> = \<mu> (main_part S)"
+ also have "\<dots> = \<mu>' S"
using S assms measure_additive[of "main_part S" "null_part S"]
by (auto simp: measure_additive)
- finally show ?thesis unfolding \<mu>'_def by simp
+ finally show ?thesis by simp
qed
lemma (in completeable_measure_space) sets_completionI_sub:
@@ -154,7 +157,7 @@
unfolding * ..
also have "\<dots> = \<mu> (\<Union>i. main_part (S i))"
using null_set S by (intro measure_Un_null_set) auto
- finally show ?thesis unfolding \<mu>'_def .
+ finally show ?thesis .
qed
lemma (in completeable_measure_space) \<mu>_main_part_Un:
@@ -168,30 +171,35 @@
unfolding range_binary_eq Un_range_binary UN by auto
qed
-sublocale completeable_measure_space \<subseteq> completion!: measure_space completion \<mu>'
-proof
- show "\<mu>' {} = 0" by auto
-next
- show "countably_additive completion \<mu>'"
- proof (unfold countably_additive_def, intro allI conjI impI)
- fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets completion" "disjoint_family A"
- have "disjoint_family (\<lambda>i. main_part (A i))"
- proof (intro disjoint_family_on_bisimulation[OF A(2)])
- fix n m assume "A n \<inter> A m = {}"
- then have "(main_part (A n) \<union> null_part (A n)) \<inter> (main_part (A m) \<union> null_part (A m)) = {}"
- 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
+sublocale completeable_measure_space \<subseteq> completion!: measure_space completion
+ where "measure completion = \<mu>'"
+proof -
+ show "measure_space completion"
+ proof
+ show "measure completion {} = 0" by (auto simp: completion_def)
+ next
+ show "countably_additive completion (measure completion)"
+ proof (intro countably_additiveI)
+ fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets completion" "disjoint_family A"
+ have "disjoint_family (\<lambda>i. main_part (A i))"
+ proof (intro disjoint_family_on_bisimulation[OF A(2)])
+ fix n m assume "A n \<inter> A m = {}"
+ then have "(main_part (A n) \<union> null_part (A n)) \<inter> (main_part (A m) \<union> null_part (A m)) = {}"
+ 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))"
+ 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)"
+ by (simp add: completion_def \<mu>_main_part_UN[OF A(1)])
qed
- then have "(\<Sum>\<^isub>\<infinity>n. \<mu>' (A n)) = \<mu> (\<Union>i. main_part (A i))"
- unfolding \<mu>'_def using A by (intro measure_countably_additive) auto
- then show "(\<Sum>\<^isub>\<infinity>n. \<mu>' (A n)) = \<mu>' (UNION UNIV A)"
- unfolding \<mu>_main_part_UN[OF A(1)] .
qed
+ show "measure completion = \<mu>'" unfolding completion_def by simp
qed
lemma (in completeable_measure_space) completion_ex_simple_function:
- assumes f: "completion.simple_function f"
- shows "\<exists>f'. simple_function f' \<and> (AE x. f x = f' x)"
+ assumes f: "simple_function completion f"
+ shows "\<exists>f'. simple_function M f' \<and> (AE x. f x = f' x)"
proof -
let "?F x" = "f -` {x} \<inter> space M"
have F: "\<And>x. ?F x \<in> sets completion" and fin: "finite (f`space M)"
@@ -248,11 +256,11 @@
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. completion.simple_function (f i)" "f \<up> g" by auto
- then have "\<forall>i. \<exists>f'. simple_function f' \<and> (AE x. f i x = f' x)"
+ 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 this[THEN choice] obtain f' where
- sf: "\<And>i. simple_function (f' i)" and
+ 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)
--- a/src/HOL/Probability/Information.thy Wed Feb 02 10:35:41 2011 +0100
+++ b/src/HOL/Probability/Information.thy Wed Feb 02 12:34:45 2011 +0100
@@ -165,43 +165,45 @@
Kullback$-$Leibler distance. *}
definition
- "KL_divergence b M \<mu> \<nu> =
- measure_space.integral M \<mu> (\<lambda>x. log b (real (sigma_finite_measure.RN_deriv M \<nu> \<mu> x)))"
+ "KL_divergence b M \<nu> = \<integral>x. log b (real (RN_deriv M \<nu> x)) \<partial>M\<lparr>measure := \<nu>\<rparr>"
lemma (in sigma_finite_measure) KL_divergence_cong:
- assumes "measure_space M \<nu>"
- and cong: "\<And>A. A \<in> sets M \<Longrightarrow> \<mu>' A = \<mu> A" "\<And>A. A \<in> sets M \<Longrightarrow> \<nu>' A = \<nu> A"
- shows "KL_divergence b M \<nu>' \<mu>' = KL_divergence b M \<nu> \<mu>"
+ assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?\<nu>")
+ assumes [simp]: "sets N = sets M" "space N = space M"
+ "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A"
+ "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = \<nu>' A"
+ shows "KL_divergence b M \<nu> = KL_divergence b N \<nu>'"
proof -
- interpret \<nu>: measure_space M \<nu> by fact
- show ?thesis
- unfolding KL_divergence_def
- using RN_deriv_cong[OF cong, of "\<lambda>A. A"]
- by (simp add: cong \<nu>.integral_cong_measure[OF cong(2)])
+ interpret \<nu>: measure_space ?\<nu> by fact
+ have "KL_divergence b M \<nu> = \<integral>x. log b (real (RN_deriv N \<nu>' x)) \<partial>?\<nu>"
+ by (simp cong: RN_deriv_cong \<nu>.integral_cong add: KL_divergence_def)
+ also have "\<dots> = KL_divergence b N \<nu>'"
+ by (auto intro!: \<nu>.integral_cong_measure[symmetric] simp: KL_divergence_def)
+ finally show ?thesis .
qed
lemma (in finite_measure_space) KL_divergence_eq_finite:
- assumes v: "finite_measure_space M \<nu>"
+ assumes v: "finite_measure_space (M\<lparr>measure := \<nu>\<rparr>)"
assumes ac: "absolutely_continuous \<nu>"
- shows "KL_divergence b M \<nu> \<mu> = (\<Sum>x\<in>space M. real (\<nu> {x}) * log b (real (\<nu> {x}) / real (\<mu> {x})))" (is "_ = ?sum")
+ shows "KL_divergence b M \<nu> = (\<Sum>x\<in>space M. real (\<nu> {x}) * log b (real (\<nu> {x}) / real (\<mu> {x})))" (is "_ = ?sum")
proof (simp add: KL_divergence_def finite_measure_space.integral_finite_singleton[OF v])
- interpret v: finite_measure_space M \<nu> by fact
- have ms: "measure_space M \<nu>" by fact
- show "(\<Sum>x \<in> space M. log b (real (RN_deriv \<nu> x)) * real (\<nu> {x})) = ?sum"
+ interpret v: finite_measure_space "M\<lparr>measure := \<nu>\<rparr>" by fact
+ 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])
qed
lemma (in finite_prob_space) KL_divergence_positive_finite:
- assumes v: "finite_prob_space M \<nu>"
+ assumes v: "finite_prob_space (M\<lparr>measure := \<nu>\<rparr>)"
assumes ac: "absolutely_continuous \<nu>"
and "1 < b"
- shows "0 \<le> KL_divergence b M \<nu> \<mu>"
+ shows "0 \<le> KL_divergence b M \<nu>"
proof -
- interpret v: finite_prob_space M \<nu> using v .
- have ms: "finite_measure_space M \<nu>" by default
+ interpret v: finite_prob_space "M\<lparr>measure := \<nu>\<rparr>" by fact
+ have ms: "finite_measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
- have "- (KL_divergence b M \<nu> \<mu>) \<le> log b (\<Sum>x\<in>space M. real (\<mu> {x}))"
+ have "- (KL_divergence b M \<nu>) \<le> log b (\<Sum>x\<in>space M. real (\<mu> {x}))"
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
@@ -215,16 +217,15 @@
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> \<mu>" using finite_sum_over_space_eq_1 by simp
+ thus "0 \<le> KL_divergence b M \<nu>" using finite_sum_over_space_eq_1 by simp
qed
subsection {* Mutual Information *}
definition (in prob_space)
"mutual_information b S T X Y =
- KL_divergence b (sigma (pair_algebra S T))
- (joint_distribution X Y)
- (pair_sigma_finite.pair_measure S (distribution X) T (distribution Y))"
+ KL_divergence b (S\<lparr>measure := distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := distribution Y\<rparr>)
+ (joint_distribution X Y)"
definition (in prob_space)
"entropy b s X = mutual_information b s s X X"
@@ -232,32 +233,49 @@
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) \<rparr>
- \<lparr> space = Y`space M, sets = Pow (Y`space M) \<rparr> X Y"
+ \<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"
+
+lemma algebra_measure_update[simp]:
+ "algebra (M'\<lparr>measure := m\<rparr>) \<longleftrightarrow> algebra M'"
+ unfolding algebra_def by simp
+
+lemma sigma_algebra_measure_update[simp]:
+ "sigma_algebra (M'\<lparr>measure := m\<rparr>) \<longleftrightarrow> sigma_algebra M'"
+ unfolding sigma_algebra_def sigma_algebra_axioms_def by simp
+
+lemma finite_sigma_algebra_measure_update[simp]:
+ "finite_sigma_algebra (M'\<lparr>measure := m\<rparr>) \<longleftrightarrow> finite_sigma_algebra M'"
+ unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp
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 (sigma (pair_algebra S T))
- (pair_sigma_finite.pair_measure S (distribution X) T (distribution Y)) (joint_distribution X 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)"
proof -
- interpret X: finite_prob_space S "distribution X" using X by (rule distribution_finite_prob_space)
- interpret Y: finite_prob_space T "distribution Y" using Y by (rule distribution_finite_prob_space)
- interpret XY: pair_finite_prob_space S "distribution X" T "distribution Y" by default
- interpret P: finite_prob_space XY.P "joint_distribution X Y"
- using assms by (intro joint_distribution_finite_prob_space)
+ interpret X: finite_prob_space "S\<lparr>measure := distribution X\<rparr>"
+ using X by (rule distribution_finite_prob_space)
+ interpret Y: finite_prob_space "T\<lparr>measure := 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>"
+ 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)"
proof (rule XY.absolutely_continuousI)
- show "finite_measure_space XY.P (joint_distribution X Y)" by default
- fix x assume "x \<in> space XY.P" and "XY.pair_measure {x} = 0"
+ show "finite_measure_space (XY.P\<lparr> measure := 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"
- by (cases x) (auto simp: pair_algebra_def)
- with assms[THEN finite_random_variableD]
- joint_distribution_Times_le_fst[of S X T Y "{a}" "{b}"]
- joint_distribution_Times_le_snd[of S X T Y "{a}" "{b}"]
+ 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
+ by (auto simp del: X.sets_eq_Pow Y.sets_eq_Pow)
with distr show "joint_distribution X Y {x} = 0" by auto
qed
qed
@@ -274,19 +292,21 @@
and mutual_information_positive_generic:
"0 \<le> mutual_information b MX MY X Y" (is ?positive)
proof -
- interpret X: finite_prob_space MX "distribution X" using MX by (rule distribution_finite_prob_space)
- interpret Y: finite_prob_space MY "distribution Y" using MY by (rule distribution_finite_prob_space)
- interpret XY: pair_finite_prob_space MX "distribution X" MY "distribution Y" by default
- interpret P: finite_prob_space XY.P "joint_distribution X Y"
- using assms by (intro joint_distribution_finite_prob_space)
+ interpret X: finite_prob_space "MX\<lparr>measure := distribution X\<rparr>"
+ using MX by (rule distribution_finite_prob_space)
+ interpret Y: finite_prob_space "MY\<lparr>measure := 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>"
+ using assms by (auto intro!: joint_distribution_finite_prob_space)
- have P_ms: "finite_measure_space XY.P (joint_distribution X Y)" by default
- have P_ps: "finite_prob_space XY.P (joint_distribution X Y)" by default
+ 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
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: pair_algebra_def setsum_cartesian_product' real_of_pextreal_mult[symmetric])
+ (auto simp add: space_pair_measure setsum_cartesian_product' real_of_pextreal_mult[symmetric])
show ?positive
using XY.KL_divergence_positive_finite[OF P_ps finite_variables_absolutely_continuous[OF MX MY] b_gt_1]
@@ -301,12 +321,12 @@
by (auto simp add: ac_simps intro!: setsum_reindex_cong[OF swap_inj_on])
lemma (in information_space) mutual_information_commute_simple:
- assumes X: "simple_function X" and Y: "simple_function Y"
+ assumes X: "simple_function M X" and Y: "simple_function M Y"
shows "\<I>(X;Y) = \<I>(Y;X)"
by (intro X Y simple_function_imp_finite_random_variable mutual_information_commute)
lemma (in information_space) mutual_information_eq:
- assumes "simple_function X" "simple_function Y"
+ 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}))))"
@@ -327,7 +347,7 @@
by (simp cong: distribution_cong image_cong)
lemma (in information_space) mutual_information_positive:
- assumes "simple_function X" "simple_function Y"
+ assumes "simple_function M X" "simple_function M Y"
shows "0 \<le> \<I>(X;Y)"
using assms by (simp add: mutual_information_positive_generic)
@@ -335,13 +355,14 @@
abbreviation (in information_space)
entropy_Pow ("\<H>'(_')") where
- "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M) \<rparr> X"
+ "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M), measure = distribution X \<rparr> X"
lemma (in information_space) entropy_generic_eq:
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})))"
proof -
- interpret MX: finite_prob_space MX "distribution X" using MX by (rule distribution_finite_prob_space)
+ interpret MX: finite_prob_space "MX\<lparr>measure := 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
@@ -353,25 +374,26 @@
show ?thesis
unfolding entropy_def mutual_information_generic_eq[OF MX MX]
unfolding setsum_cartesian_product[symmetric] setsum_negf[symmetric] remove_XX
- by (auto simp: setsum_cases MX.finite_space)
+ using MX.finite_space by (auto simp: setsum_cases)
qed
lemma (in information_space) entropy_eq:
- assumes "simple_function X"
+ assumes "simple_function M X"
shows "\<H>(X) = -(\<Sum> x \<in> X ` space M. real (distribution X {x}) * log b (real (distribution X {x})))"
using assms by (simp add: entropy_generic_eq)
lemma (in information_space) entropy_positive:
- "simple_function X \<Longrightarrow> 0 \<le> \<H>(X)"
+ "simple_function M X \<Longrightarrow> 0 \<le> \<H>(X)"
unfolding entropy_def by (simp add: mutual_information_positive)
lemma (in information_space) entropy_certainty_eq_0:
- assumes "simple_function X" and "x \<in> X ` space M" and "distribution X {x} = 1"
+ assumes "simple_function M X" and "x \<in> X ` space M" and "distribution X {x} = 1"
shows "\<H>(X) = 0"
proof -
- interpret X: finite_prob_space "\<lparr> space = X ` space M, sets = Pow (X ` space M) \<rparr>" "distribution X"
- using simple_function_imp_finite_random_variable[OF `simple_function X`]
- by (rule distribution_finite_prob_space)
+ 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>"]
+ 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
@@ -383,38 +405,39 @@
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
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 X`] by (auto simp: y fi)
+ 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 X"
+ assumes "simple_function M X"
shows "\<H>(X) \<le> log b (real (card (X ` space M \<inter> {x . distribution X {x} \<noteq> 0})))"
proof -
let "?d x" = "distribution X {x}"
let "?p x" = "real (?d 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 X`] setsum_negf[symmetric] log_simps not_less)
+ by (auto intro!: setsum_cong simp: entropy_eq[OF `simple_function M X`] setsum_negf[symmetric] log_simps not_less)
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 X` sum_over_space_real_distribution
+ 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 X`[THEN simple_function_imp_random_variable], simplified]
+ 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)
finally show ?thesis
- using `simple_function X` by (auto simp: setsum_cases real_eq_of_nat simple_function_def)
+ using `simple_function M X` by (auto simp: setsum_cases real_eq_of_nat simple_function_def)
qed
lemma (in information_space) entropy_uniform_max:
- assumes "simple_function X"
+ assumes "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 -
- interpret X: finite_prob_space "\<lparr> space = X ` space M, sets = Pow (X ` space M) \<rparr>" "distribution X"
- using simple_function_imp_finite_random_variable[OF `simple_function X`]
- by (rule distribution_finite_prob_space)
+ 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>"]
+ interpret X: finite_prob_space ?X by simp
have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff
- using `simple_function X` not_empty by (auto simp: simple_function_def)
+ 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])
@@ -423,18 +446,18 @@
qed }
thus ?thesis
using not_empty X.finite_space b_gt_1 card_gt0
- by (simp add: entropy_eq[OF `simple_function X`] real_eq_of_nat[symmetric] log_simps)
+ by (simp add: entropy_eq[OF `simple_function M X`] real_eq_of_nat[symmetric] log_simps)
qed
lemma (in information_space) entropy_le_card:
- assumes "simple_function X"
+ assumes "simple_function M X"
shows "\<H>(X) \<le> log b (real (card (X ` space M)))"
proof cases
assume "X ` space M \<inter> {x. distribution X {x} \<noteq> 0} = {}"
then have "\<And>x. x\<in>X`space M \<Longrightarrow> distribution X {x} = 0" by auto
moreover
have "0 < card (X`space M)"
- using `simple_function X` not_empty
+ using `simple_function M X` not_empty
by (auto simp: card_gt_0_iff simple_function_def)
then have "log b 1 \<le> log b (real (card (X`space M)))"
using b_gt_1 by (intro log_le) auto
@@ -451,10 +474,10 @@
qed
lemma (in information_space) entropy_commute:
- assumes "simple_function X" "simple_function Y"
+ assumes "simple_function M X" "simple_function M Y"
shows "\<H>(\<lambda>x. (X x, Y x)) = \<H>(\<lambda>x. (Y x, X x))"
proof -
- have sf: "simple_function (\<lambda>x. (X x, Y x))" "simple_function (\<lambda>x. (Y x, X x))"
+ have sf: "simple_function M (\<lambda>x. (X x, Y x))" "simple_function M (\<lambda>x. (Y x, X x))"
using assms by (auto intro: simple_function_Pair)
have *: "(\<lambda>x. (Y x, X x))`space M = (\<lambda>(a,b). (b,a))`(\<lambda>x. (X x, Y x))`space M"
by auto
@@ -466,12 +489,12 @@
qed
lemma (in information_space) entropy_eq_cartesian_product:
- assumes "simple_function X" "simple_function Y"
+ 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)})))"
proof -
- have sf: "simple_function (\<lambda>x. (X x, Y x))"
+ have sf: "simple_function M (\<lambda>x. (X x, Y x))"
using assms by (auto intro: simple_function_Pair)
{ fix x assume "x\<notin>(\<lambda>x. (X x, Y x))`space M"
then have "(\<lambda>x. (X x, Y x)) -` {x} \<inter> space M = {}" by auto
@@ -485,19 +508,18 @@
subsection {* Conditional Mutual Information *}
definition (in prob_space)
- "conditional_mutual_information b M1 M2 M3 X Y Z \<equiv>
- mutual_information b M1 (sigma (pair_algebra M2 M3)) X (\<lambda>x. (Y x, Z x)) -
- mutual_information b M1 M3 X Z"
+ "conditional_mutual_information b MX MY MZ X Y Z \<equiv>
+ mutual_information b MX (MY \<Otimes>\<^isub>M MZ) X (\<lambda>x. (Y x, Z x)) -
+ mutual_information b MX MZ X Z"
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) \<rparr>
- \<lparr> space = Y`space M, sets = Pow (Y`space M) \<rparr>
- \<lparr> space = Z`space M, sets = Pow (Z`space M) \<rparr>
+ \<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>
X Y Z"
-
lemma (in information_space) conditional_mutual_information_generic_eq:
assumes MX: "finite_random_variable MX X"
and MY: "finite_random_variable MY Y"
@@ -519,7 +541,7 @@
note finite_var = MX MY MZ
note random_var = finite_var[THEN finite_random_variableD]
- note space_simps = space_pair_algebra space_sigma algebra.simps
+ note space_simps = space_pair_measure space_sigma algebra.simps
note YZ = finite_random_variable_pairI[OF finite_var(2,3)]
note XZ = finite_random_variable_pairI[OF finite_var(1,3)]
@@ -574,12 +596,12 @@
unfolding conditional_mutual_information_def
unfolding mutual_information_generic_eq[OF finite_var(1,3)]
unfolding mutual_information_generic_eq[OF finite_var(1) YZ]
- by (simp add: space_sigma space_pair_algebra setsum_cartesian_product')
+ by (simp add: space_sigma space_pair_measure setsum_cartesian_product')
finally show ?thesis by simp
qed
lemma (in information_space) conditional_mutual_information_eq:
- assumes "simple_function X" "simple_function Y" "simple_function Z"
+ 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)}) /
@@ -588,11 +610,11 @@
by simp
lemma (in information_space) conditional_mutual_information_eq_mutual_information:
- assumes X: "simple_function X" and Y: "simple_function Y"
+ assumes X: "simple_function M X" and Y: "simple_function M Y"
shows "\<I>(X ; Y) = \<I>(X ; Y | (\<lambda>x. ()))"
proof -
have [simp]: "(\<lambda>x. ()) ` space M = {()}" using not_empty by auto
- have C: "simple_function (\<lambda>x. ())" by auto
+ have C: "simple_function M (\<lambda>x. ())" by auto
show ?thesis
unfolding conditional_mutual_information_eq[OF X Y C]
unfolding mutual_information_eq[OF X Y]
@@ -608,12 +630,13 @@
lemma (in prob_space) setsum_distribution:
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 \<rparr>" "\<lambda>x. ()" "{()}"]
- using sigma_algebra_Pow[of "UNIV::unit set"] by simp
+ using sigma_algebra_Pow[of "UNIV::unit set" "()"] by simp
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 \<rparr>" "\<lambda>x. ()" "{()}"]
- using sigma_algebra_Pow[of "UNIV::unit set"] by simp
+ 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
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"
@@ -633,7 +656,7 @@
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_algebra space_sigma algebra.simps
+ 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)]
@@ -672,7 +695,7 @@
unfolding setsum_cartesian_product'
unfolding setsum_commute[of _ "space MY"]
unfolding setsum_commute[of _ "space MZ"]
- by (simp_all add: space_pair_algebra
+ 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`])
@@ -704,10 +727,9 @@
qed
lemma (in information_space) conditional_mutual_information_positive:
- assumes "simple_function X" and "simple_function Y" and "simple_function Z"
+ assumes "simple_function M X" and "simple_function M Y" and "simple_function M Z"
shows "0 \<le> \<I>(X;Y|Z)"
- using conditional_mutual_information_generic_positive[OF assms[THEN simple_function_imp_finite_random_variable]]
- by simp
+ by (rule conditional_mutual_information_generic_positive[OF assms[THEN simple_function_imp_finite_random_variable]])
subsection {* Conditional Entropy *}
@@ -717,16 +739,17 @@
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) \<rparr>
- \<lparr> space = Y`space M, sets = Pow (Y`space M) \<rparr> X Y"
+ \<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"
lemma (in information_space) conditional_entropy_positive:
- "simple_function X \<Longrightarrow> simple_function Y \<Longrightarrow> 0 \<le> \<H>(X | Y)"
+ "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 =
@@ -743,7 +766,7 @@
{ fix x z have "?XXZ x x z = ?XZ x z"
unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>]) }
note this[simp]
- { fix x x' :: 'b and z assume "x' \<noteq> x"
+ { 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) }
note this[simp]
@@ -762,7 +785,6 @@
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))" . }
note * = this
-
show ?thesis
unfolding conditional_entropy_def
unfolding conditional_mutual_information_generic_eq[OF MX MX MZ]
@@ -772,7 +794,7 @@
qed
lemma (in information_space) conditional_entropy_eq:
- assumes "simple_function X" "simple_function Z"
+ 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)}) *
@@ -781,7 +803,7 @@
by simp
lemma (in information_space) conditional_entropy_eq_ce_with_hypothesis:
- assumes X: "simple_function X" and Y: "simple_function Y"
+ 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)}) *
@@ -794,7 +816,7 @@
intro!: setsum_cong)
lemma (in information_space) conditional_entropy_eq_cartesian_product:
- assumes "simple_function X" "simple_function Y"
+ 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})))"
@@ -804,7 +826,7 @@
subsection {* Equalities *}
lemma (in information_space) mutual_information_eq_entropy_conditional_entropy:
- assumes X: "simple_function X" and Z: "simple_function Z"
+ 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)})"
@@ -828,7 +850,7 @@
qed
lemma (in information_space) conditional_entropy_less_eq_entropy:
- assumes X: "simple_function X" and Z: "simple_function Z"
+ assumes X: "simple_function M X" and Z: "simple_function M Z"
shows "\<H>(X | Z) \<le> \<H>(X)"
proof -
have "\<I>(X ; Z) = \<H>(X) - \<H>(X | Z)" using mutual_information_eq_entropy_conditional_entropy[OF assms] .
@@ -837,7 +859,7 @@
qed
lemma (in information_space) entropy_chain_rule:
- assumes X: "simple_function X" and Y: "simple_function Y"
+ 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)})"
@@ -976,7 +998,7 @@
qed
lemma (in information_space) entropy_partition:
- assumes sf: "simple_function X" "simple_function P"
+ assumes sf: "simple_function M X" "simple_function M P"
assumes svi: "subvimage (space M) X P"
shows "\<H>(X) = \<H>(P) + \<H>(X|P)"
proof -
@@ -1026,10 +1048,10 @@
qed
corollary (in information_space) entropy_data_processing:
- assumes X: "simple_function X" shows "\<H>(f \<circ> X) \<le> \<H>(X)"
+ assumes X: "simple_function M X" shows "\<H>(f \<circ> X) \<le> \<H>(X)"
proof -
note X
- moreover have fX: "simple_function (f \<circ> X)" using X by auto
+ moreover have fX: "simple_function M (f \<circ> X)" using X by auto
moreover have "subvimage (space M) X (f \<circ> X)" by auto
ultimately have "\<H>(X) = \<H>(f\<circ>X) + \<H>(X|f\<circ>X)" by (rule entropy_partition)
then show "\<H>(f \<circ> X) \<le> \<H>(X)"
@@ -1037,12 +1059,12 @@
qed
corollary (in information_space) entropy_of_inj:
- assumes X: "simple_function X" and inj: "inj_on f (X`space M)"
+ assumes X: "simple_function M X" and inj: "inj_on f (X`space M)"
shows "\<H>(f \<circ> X) = \<H>(X)"
proof (rule antisym)
show "\<H>(f \<circ> X) \<le> \<H>(X)" using entropy_data_processing[OF X] .
next
- have sf: "simple_function (f \<circ> X)"
+ have sf: "simple_function M (f \<circ> X)"
using X by auto
have "\<H>(X) = \<H>(the_inv_into (X`space M) f \<circ> (f \<circ> X))"
by (auto intro!: mutual_information_cong simp: entropy_def the_inv_into_f_f[OF inj])
--- a/src/HOL/Probability/Lebesgue_Integration.thy Wed Feb 02 10:35:41 2011 +0100
+++ b/src/HOL/Probability/Lebesgue_Integration.thy Wed Feb 02 12:34:45 2011 +0100
@@ -7,6 +7,7 @@
begin
lemma sums_If_finite:
+ fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
assumes finite: "finite {r. P r}"
shows "(\<lambda>r. if P r then f r else 0) sums (\<Sum>r\<in>{r. P r}. f r)" (is "?F sums _")
proof cases
@@ -24,7 +25,8 @@
qed
lemma sums_single:
- "(\<lambda>r. if r = i then f r else 0) sums f i"
+ fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
+ shows "(\<lambda>r. if r = i then f r else 0) sums f i"
using sums_If_finite[of "\<lambda>r. r = i" f] by simp
section "Simple function"
@@ -37,12 +39,12 @@
*}
-definition (in sigma_algebra) "simple_function g \<longleftrightarrow>
+definition "simple_function M g \<longleftrightarrow>
finite (g ` space M) \<and>
(\<forall>x \<in> g ` space M. g -` {x} \<inter> space M \<in> sets M)"
lemma (in sigma_algebra) simple_functionD:
- assumes "simple_function g"
+ assumes "simple_function M g"
shows "finite (g ` space M)" and "g -` X \<inter> space M \<in> sets M"
proof -
show "finite (g ` space M)"
@@ -55,7 +57,7 @@
lemma (in sigma_algebra) simple_function_indicator_representation:
fixes f ::"'a \<Rightarrow> pextreal"
- assumes f: "simple_function f" and x: "x \<in> space M"
+ 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")
proof -
@@ -69,7 +71,7 @@
qed
lemma (in measure_space) simple_function_notspace:
- "simple_function (\<lambda>x. h x * indicator (- space M) x::pextreal)" (is "simple_function ?h")
+ "simple_function M (\<lambda>x. h x * indicator (- space M) x::pextreal)" (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)
@@ -79,7 +81,7 @@
lemma (in sigma_algebra) simple_function_cong:
assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
- shows "simple_function f \<longleftrightarrow> simple_function g"
+ shows "simple_function M f \<longleftrightarrow> simple_function M g"
proof -
have "f ` space M = g ` space M"
"\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
@@ -87,15 +89,21 @@
thus ?thesis unfolding simple_function_def using assms by simp
qed
+lemma (in sigma_algebra) simple_function_cong_algebra:
+ assumes "sets N = sets M" "space N = space M"
+ shows "simple_function M f \<longleftrightarrow> simple_function N f"
+ unfolding simple_function_def assms ..
+
lemma (in sigma_algebra) borel_measurable_simple_function:
- assumes "simple_function f"
+ assumes "simple_function M f"
shows "f \<in> borel_measurable M"
proof (rule borel_measurableI)
fix S
let ?I = "f ` (f -` S \<inter> space M)"
have *: "(\<Union>x\<in>?I. f -` {x} \<inter> space M) = f -` S \<inter> space M" (is "?U = _") by auto
have "finite ?I"
- using assms unfolding simple_function_def by (auto intro: finite_subset)
+ using assms unfolding simple_function_def
+ using finite_subset[of "f ` (f -` S \<inter> space M)" "f ` space M"] by auto
hence "?U \<in> sets M"
apply (rule finite_UN)
using assms unfolding simple_function_def by auto
@@ -105,17 +113,17 @@
lemma (in sigma_algebra) simple_function_borel_measurable:
fixes f :: "'a \<Rightarrow> 'x::t2_space"
assumes "f \<in> borel_measurable M" and "finite (f ` space M)"
- shows "simple_function f"
+ shows "simple_function M f"
using assms unfolding simple_function_def
by (auto intro: borel_measurable_vimage)
lemma (in sigma_algebra) simple_function_const[intro, simp]:
- "simple_function (\<lambda>x. c)"
+ "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 f"
- shows "simple_function (g \<circ> f)"
+ assumes "simple_function M f"
+ shows "simple_function M (g \<circ> f)"
unfolding simple_function_def
proof safe
show "finite ((g \<circ> f) ` space M)"
@@ -132,7 +140,7 @@
lemma (in sigma_algebra) simple_function_indicator[intro, simp]:
assumes "A \<in> sets M"
- shows "simple_function (indicator A)"
+ shows "simple_function M (indicator A)"
proof -
have "indicator A ` space M \<subseteq> {0, 1}" (is "?S \<subseteq> _")
by (auto simp: indicator_def)
@@ -143,9 +151,9 @@
qed
lemma (in sigma_algebra) simple_function_Pair[intro, simp]:
- assumes "simple_function f"
- assumes "simple_function g"
- shows "simple_function (\<lambda>x. (f x, g x))" (is "simple_function ?p")
+ assumes "simple_function M f"
+ assumes "simple_function M g"
+ shows "simple_function M (\<lambda>x. (f x, g x))" (is "simple_function M ?p")
unfolding simple_function_def
proof safe
show "finite (?p ` space M)"
@@ -161,16 +169,16 @@
qed
lemma (in sigma_algebra) simple_function_compose1:
- assumes "simple_function f"
- shows "simple_function (\<lambda>x. g (f x))"
+ assumes "simple_function M f"
+ shows "simple_function M (\<lambda>x. g (f x))"
using simple_function_compose[OF assms, of g]
by (simp add: comp_def)
lemma (in sigma_algebra) simple_function_compose2:
- assumes "simple_function f" and "simple_function g"
- shows "simple_function (\<lambda>x. h (f x) (g x))"
+ assumes "simple_function M f" and "simple_function M g"
+ shows "simple_function M (\<lambda>x. h (f x) (g x))"
proof -
- have "simple_function ((\<lambda>(x, y). h x y) \<circ> (\<lambda>x. (f x, g x)))"
+ have "simple_function M ((\<lambda>(x, y). h x y) \<circ> (\<lambda>x. (f x, g x)))"
using assms by auto
thus ?thesis by (simp_all add: comp_def)
qed
@@ -183,14 +191,14 @@
and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"]
lemma (in sigma_algebra) simple_function_setsum[intro, simp]:
- assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function (f i)"
- shows "simple_function (\<lambda>x. \<Sum>i\<in>P. f i x)"
+ assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
+ shows "simple_function M (\<lambda>x. \<Sum>i\<in>P. f i x)"
proof cases
assume "finite P" from this assms show ?thesis by induct auto
qed auto
lemma (in sigma_algebra) simple_function_le_measurable:
- assumes "simple_function f" "simple_function g"
+ 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} =
@@ -214,7 +222,7 @@
lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence:
fixes u :: "'a \<Rightarrow> pextreal"
assumes u: "u \<in> borel_measurable M"
- shows "\<exists>f. (\<forall>i. simple_function (f i) \<and> (\<forall>x\<in>space M. f i x \<noteq> \<omega>)) \<and> f \<up> u"
+ 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"
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)))"
@@ -406,10 +414,10 @@
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 (f i)" "\<And>i. \<omega>\<notin>f i`space M"
+ obtains (x) f where "f \<up> u" "\<And>i. simple_function M (f i)" "\<And>i. \<omega>\<notin>f i`space M"
proof -
from borel_measurable_implies_simple_function_sequence[OF assms]
- obtain f where x: "\<And>i. simple_function (f i)" "f \<up> u"
+ 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])
@@ -417,7 +425,7 @@
lemma (in sigma_algebra) simple_function_eq_borel_measurable:
fixes f :: "'a \<Rightarrow> pextreal"
- shows "simple_function f \<longleftrightarrow>
+ 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]
@@ -425,8 +433,8 @@
lemma (in measure_space) simple_function_restricted:
fixes f :: "'a \<Rightarrow> pextreal" assumes "A \<in> sets M"
- shows "sigma_algebra.simple_function (restricted_space A) f \<longleftrightarrow> simple_function (\<lambda>x. f x * indicator A x)"
- (is "sigma_algebra.simple_function ?R f \<longleftrightarrow> simple_function ?f")
+ 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)"
@@ -463,29 +471,26 @@
qed
lemma (in sigma_algebra) simple_function_subalgebra:
- assumes "sigma_algebra.simple_function N f"
- and N_subalgebra: "sets N \<subseteq> sets M" "space N = space M" "sigma_algebra N"
- shows "simple_function f"
- using assms
- unfolding simple_function_def
- unfolding sigma_algebra.simple_function_def[OF N_subalgebra(3)]
- by auto
+ assumes "simple_function N f"
+ and N_subalgebra: "sets N \<subseteq> sets M" "space N = space M"
+ shows "simple_function M f"
+ using assms unfolding simple_function_def by auto
lemma (in measure_space) simple_function_vimage:
assumes T: "sigma_algebra M'" "T \<in> measurable M M'"
- and f: "sigma_algebra.simple_function M' f"
- shows "simple_function (\<lambda>x. f (T x))"
+ and f: "simple_function M' f"
+ shows "simple_function M (\<lambda>x. f (T x))"
proof (intro simple_function_def[THEN iffD2] conjI ballI)
interpret T: sigma_algebra M' by fact
have "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space M'"
using T unfolding measurable_def by auto
then show "finite ((\<lambda>x. f (T x)) ` space M)"
- using f unfolding T.simple_function_def by (auto intro: finite_subset)
+ using f unfolding simple_function_def by (auto intro: finite_subset)
fix i assume i: "i \<in> (\<lambda>x. f (T x)) ` space M"
then have "i \<in> f ` space M'"
using T unfolding measurable_def by auto
then have "f -` {i} \<inter> space M' \<in> sets M'"
- using f unfolding T.simple_function_def by auto
+ using f unfolding simple_function_def by auto
then have "T -` (f -` {i} \<inter> space M') \<inter> space M \<in> sets M"
using T unfolding measurable_def by auto
also have "T -` (f -` {i} \<inter> space M') \<inter> space M = (\<lambda>x. f (T x)) -` {i} \<inter> space M"
@@ -495,12 +500,18 @@
section "Simple integral"
-definition (in measure_space) simple_integral (binder "\<integral>\<^isup>S " 10) where
- "simple_integral f = (\<Sum>x \<in> f ` space M. x * \<mu> (f -` {x} \<inter> space M))"
+definition simple_integral_def:
+ "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)
+
+translations
+ "\<integral>\<^isup>S x. f \<partial>M" == "CONST integral\<^isup>S M (%x. f)"
lemma (in measure_space) simple_integral_cong:
assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
- shows "simple_integral f = simple_integral g"
+ shows "integral\<^isup>S M f = integral\<^isup>S M g"
proof -
have "f ` space M = g ` space M"
"\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
@@ -509,18 +520,18 @@
qed
lemma (in measure_space) simple_integral_cong_measure:
- assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = \<mu> A" and "simple_function f"
- shows "measure_space.simple_integral M \<nu> f = simple_integral f"
+ assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
+ and "simple_function M f"
+ shows "integral\<^isup>S N f = integral\<^isup>S M f"
proof -
- interpret v: measure_space M \<nu>
- by (rule measure_space_cong) fact
- from simple_functionD[OF `simple_function f`] assms show ?thesis
- unfolding simple_integral_def v.simple_integral_def
- by (auto intro!: setsum_cong)
+ interpret v: measure_space N
+ by (rule measure_space_cong) fact+
+ from simple_functionD[OF `simple_function M f`] assms show ?thesis
+ by (auto intro!: setsum_cong simp: simple_integral_def)
qed
lemma (in measure_space) simple_integral_const[simp]:
- "(\<integral>\<^isup>Sx. c) = c * \<mu> (space M)"
+ "(\<integral>\<^isup>Sx. c \<partial>M) = c * \<mu> (space M)"
proof (cases "space M = {}")
case True thus ?thesis unfolding simple_integral_def by simp
next
@@ -529,8 +540,8 @@
qed
lemma (in measure_space) simple_function_partition:
- assumes "simple_function f" and "simple_function g"
- shows "simple_integral f = (\<Sum>A\<in>(\<lambda>x. f -` {f x} \<inter> g -` {g x} \<inter> space M) ` space M. the_elem (f`A) * \<mu> A)"
+ assumes "simple_function M f" and "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-
let "?sub x" = "?p ` (f -` {x} \<inter> space M)"
@@ -561,7 +572,7 @@
ultimately
have "\<mu> (f -` {f x} \<inter> space M) = setsum (\<mu>) (?sub (f x))"
by (subst measure_finitely_additive) auto }
- hence "simple_integral f = (\<Sum>(x,A)\<in>?SIGMA. x * \<mu> A)"
+ 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])
@@ -584,8 +595,8 @@
qed
lemma (in measure_space) simple_integral_add[simp]:
- assumes "simple_function f" and "simple_function g"
- shows "(\<integral>\<^isup>Sx. f x + g x) = simple_integral f + simple_integral g"
+ assumes "simple_function M f" and "simple_function M g"
+ 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"
assume "x \<in> space M"
@@ -595,15 +606,15 @@
thus ?thesis
unfolding
simple_function_partition[OF simple_function_add[OF assms] simple_function_Pair[OF assms]]
- simple_function_partition[OF `simple_function f` `simple_function g`]
- simple_function_partition[OF `simple_function g` `simple_function f`]
+ 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)
qed
lemma (in measure_space) simple_integral_setsum[simp]:
- assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function (f i)"
- shows "(\<integral>\<^isup>Sx. \<Sum>i\<in>P. f i x) = (\<Sum>i\<in>P. simple_integral (f i))"
+ 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
@@ -611,8 +622,8 @@
qed auto
lemma (in measure_space) simple_integral_mult[simp]:
- assumes "simple_function f"
- shows "(\<integral>\<^isup>Sx. c * f x) = c * simple_integral f"
+ assumes "simple_function M f"
+ 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]
{ fix x let ?S = "f -` {f x} \<inter> (\<lambda>x. c * f x) -` {c * f x} \<inter> space M"
@@ -626,8 +637,8 @@
qed
lemma (in sigma_algebra) simple_function_If:
- assumes sf: "simple_function f" "simple_function g" and A: "A \<in> sets M"
- shows "simple_function (\<lambda>x. if x \<in> A then f x else g x)" (is "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
@@ -648,17 +659,17 @@
qed
lemma (in measure_space) simple_integral_mono_AE:
- assumes "simple_function f" and "simple_function g"
+ assumes "simple_function M f" and "simple_function M g"
and mono: "AE x. f x \<le> g x"
- shows "simple_integral f \<le> simple_integral g"
+ shows "integral\<^isup>S M f \<le> integral\<^isup>S M g"
proof -
let "?S x" = "(g -` {g x} \<inter> space M) \<inter> (f -` {f x} \<inter> space M)"
have *: "\<And>x. g -` {g x} \<inter> f -` {f x} \<inter> space M = ?S x"
"\<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 f` `simple_function g`]
- simple_function_partition[OF `simple_function g` `simple_function f`]
+ simple_function_partition[OF `simple_function M f` `simple_function M g`]
+ simple_function_partition[OF `simple_function M g` `simple_function M 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
@@ -680,23 +691,23 @@
qed
lemma (in measure_space) simple_integral_mono:
- assumes "simple_function f" and "simple_function g"
+ assumes "simple_function M f" and "simple_function M g"
and mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x"
- shows "simple_integral f \<le> simple_integral g"
+ shows "integral\<^isup>S M f \<le> integral\<^isup>S M g"
proof (rule simple_integral_mono_AE[OF assms(1, 2)])
show "AE x. f x \<le> g x"
using mono by (rule AE_cong) auto
qed
lemma (in measure_space) simple_integral_cong_AE:
- assumes "simple_function f" "simple_function g" and "AE x. f x = g x"
- shows "simple_integral f = simple_integral g"
+ assumes "simple_function M f" "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)
lemma (in measure_space) simple_integral_cong':
- assumes sf: "simple_function f" "simple_function g"
+ assumes sf: "simple_function M f" "simple_function M g"
and mea: "\<mu> {x\<in>space M. f x \<noteq> g x} = 0"
- shows "simple_integral f = simple_integral g"
+ shows "integral\<^isup>S M f = integral\<^isup>S M g"
proof (intro simple_integral_cong_AE sf AE_I)
show "\<mu> {x\<in>space M. f x \<noteq> g x} = 0" by fact
show "{x \<in> space M. f x \<noteq> g x} \<in> sets M"
@@ -705,12 +716,12 @@
lemma (in measure_space) simple_integral_indicator:
assumes "A \<in> sets M"
- assumes "simple_function f"
- shows "(\<integral>\<^isup>Sx. f x * indicator A x) =
+ assumes "simple_function M f"
+ shows "(\<integral>\<^isup>Sx. f x * indicator A x \<partial>M) =
(\<Sum>x \<in> f ` space M. x * \<mu> (f -` {x} \<inter> space M \<inter> A))"
proof cases
assume "A = space M"
- moreover hence "(\<integral>\<^isup>Sx. f x * indicator A x) = simple_integral f"
+ moreover hence "(\<integral>\<^isup>Sx. f x * indicator A x \<partial>M) = integral\<^isup>S M f"
by (auto intro!: simple_integral_cong)
moreover have "\<And>X. X \<inter> space M \<inter> space M = X \<inter> space M" by auto
ultimately show ?thesis by (simp add: simple_integral_def)
@@ -726,7 +737,7 @@
next
show "0 \<in> ?I ` space M" using x by (auto intro!: image_eqI[of _ _ x])
qed
- have *: "(\<integral>\<^isup>Sx. f x * indicator A x) =
+ have *: "(\<integral>\<^isup>Sx. f x * indicator A x \<partial>M) =
(\<Sum>x \<in> f ` space M \<union> {0}. x * \<mu> (f -` {x} \<inter> space M \<inter> A))"
unfolding simple_integral_def I
proof (rule setsum_mono_zero_cong_left)
@@ -752,7 +763,7 @@
lemma (in measure_space) simple_integral_indicator_only[simp]:
assumes "A \<in> sets M"
- shows "simple_integral (indicator A) = \<mu> A"
+ shows "integral\<^isup>S M (indicator A) = \<mu> A"
proof cases
assume "space M = {}" hence "A = {}" using sets_into_space[OF assms] by auto
thus ?thesis unfolding simple_integral_def using `space M = {}` by auto
@@ -765,22 +776,22 @@
qed
lemma (in measure_space) simple_integral_null_set:
- assumes "simple_function u" "N \<in> null_sets"
- shows "(\<integral>\<^isup>Sx. u x * indicator N x) = 0"
+ assumes "simple_function M u" "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)"
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) = (\<integral>\<^isup>Sx. 0)"
+ 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 intro!: AE_disjI2)
then show ?thesis by simp
qed
lemma (in measure_space) simple_integral_cong_AE_mult_indicator:
- assumes sf: "simple_function f" and eq: "AE x. x \<in> S" and "S \<in> sets M"
- shows "simple_integral f = (\<integral>\<^isup>Sx. f x * indicator S x)"
+ assumes sf: "simple_function M f" and eq: "AE x. x \<in> S" and "S \<in> sets M"
+ shows "integral\<^isup>S M f = (\<integral>\<^isup>Sx. f x * indicator S x \<partial>M)"
proof (rule simple_integral_cong_AE)
- show "simple_function f" by fact
- show "simple_function (\<lambda>x. f x * indicator S x)"
+ show "simple_function M f" by fact
+ show "simple_function M (\<lambda>x. f x * indicator S x)"
using sf `S \<in> sets M` by auto
from eq show "AE x. f x = f x * indicator S x"
by (rule AE_mp) simp
@@ -788,10 +799,9 @@
lemma (in measure_space) simple_integral_restricted:
assumes "A \<in> sets M"
- assumes sf: "simple_function (\<lambda>x. f x * indicator A x)"
- shows "measure_space.simple_integral (restricted_space A) \<mu> f = (\<integral>\<^isup>Sx. f x * indicator A x)"
- (is "_ = simple_integral ?f")
- unfolding measure_space.simple_integral_def[OF restricted_measure_space[OF `A \<in> sets M`]]
+ assumes sf: "simple_function M (\<lambda>x. f x * indicator A x)"
+ shows "integral\<^isup>S (restricted_space A) f = (\<integral>\<^isup>Sx. f x * indicator A x \<partial>M)"
+ (is "_ = integral\<^isup>S M ?f")
unfolding simple_integral_def
proof (simp, safe intro!: setsum_mono_zero_cong_left)
from sf show "finite (?f ` space M)"
@@ -816,64 +826,71 @@
qed
lemma (in measure_space) simple_integral_subalgebra:
- assumes N: "measure_space N \<mu>" and [simp]: "space N = space M"
- shows "measure_space.simple_integral N \<mu> = simple_integral"
- unfolding simple_integral_def_raw
- unfolding measure_space.simple_integral_def_raw[OF N] by simp
+ assumes N: "measure_space N" and [simp]: "space N = space M" "measure N = measure M"
+ shows "integral\<^isup>S N = integral\<^isup>S M"
+ unfolding simple_integral_def_raw by simp
lemma (in measure_space) simple_integral_vimage:
assumes T: "sigma_algebra M'" "T \<in> measurable M M'"
- and f: "sigma_algebra.simple_function M' f"
- shows "measure_space.simple_integral M' (\<lambda>A. \<mu> (T -` A \<inter> space M)) f = (\<integral>\<^isup>S x. f (T x))"
- (is "measure_space.simple_integral M' ?nu f = _")
+ and f: "simple_function M' f"
+ and \<nu>: "\<And>A. A \<in> sets M' \<Longrightarrow> measure M' A = \<mu> (T -` A \<inter> space M)"
+ shows "integral\<^isup>S M' f = (\<integral>\<^isup>S x. f (T x) \<partial>M)"
proof -
- interpret T: measure_space M' ?nu using T by (rule measure_space_vimage) auto
- show "T.simple_integral f = (\<integral>\<^isup>S x. f (T x))"
- unfolding simple_integral_def T.simple_integral_def
+ interpret T: measure_space M' using \<nu> by (rule measure_space_vimage[OF T])
+ show "integral\<^isup>S M' f = (\<integral>\<^isup>S x. f (T x) \<partial>M)"
+ unfolding simple_integral_def
proof (intro setsum_mono_zero_cong_right ballI)
show "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space M'"
using T unfolding measurable_def by auto
show "finite (f ` space M')"
- using f unfolding T.simple_function_def by auto
+ using f unfolding simple_function_def by auto
next
fix i assume "i \<in> f ` space M' - (\<lambda>x. f (T x)) ` space M"
then have "T -` (f -` {i} \<inter> space M') \<inter> space M = {}" by (auto simp: image_iff)
- then show "i * \<mu> (T -` (f -` {i} \<inter> space M') \<inter> space M) = 0" by simp
+ with f[THEN T.simple_functionD(2), THEN \<nu>]
+ show "i * T.\<mu> (f -` {i} \<inter> space M') = 0" by simp
next
fix i assume "i \<in> (\<lambda>x. f (T x)) ` space M"
then have "T -` (f -` {i} \<inter> space M') \<inter> space M = (\<lambda>x. f (T x)) -` {i} \<inter> space M"
using T unfolding measurable_def by auto
- then show "i * \<mu> (T -` (f -` {i} \<inter> space M') \<inter> space M) = i * \<mu> ((\<lambda>x. f (T x)) -` {i} \<inter> space M)"
+ with f[THEN T.simple_functionD(2), THEN \<nu>]
+ show "i * T.\<mu> (f -` {i} \<inter> space M') = i * \<mu> ((\<lambda>x. f (T x)) -` {i} \<inter> space M)"
by auto
qed
qed
-section "Continuous posititve integration"
+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)"
-definition (in measure_space) positive_integral (binder "\<integral>\<^isup>+ " 10) where
- "positive_integral f = SUPR {g. simple_function g \<and> g \<le> f} simple_integral"
+syntax
+ "_positive_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> pextreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> pextreal" ("\<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:
- "positive_integral f =
- (SUPR {g. simple_function g \<and> g \<le> f \<and> \<omega> \<notin> g`space M} simple_integral)" (is "_ = ?alt")
+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 "positive_integral f \<le> ?alt" unfolding positive_integral_def
+ show "integral\<^isup>P M f \<le> ?alt" unfolding positive_integral_def
proof (safe intro!: SUP_leI)
- fix g assume g: "simple_function g" "g \<le> f"
+ fix g assume g: "simple_function M g" "g \<le> f"
let ?G = "g -` {\<omega>} \<inter> space M"
- show "simple_integral g \<le>
- SUPR {g. simple_function g \<and> g \<le> f \<and> \<omega> \<notin> g ` space M} simple_integral"
- (is "simple_integral g \<le> SUPR ?A simple_integral")
+ 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 ?g"
+ 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 "simple_integral g = simple_integral ?g"
+ 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)
@@ -885,15 +902,15 @@
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 simple_integral = \<omega>"
+ 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 < simple_integral i"
+ show "\<exists>i\<in>?A. x < integral\<^isup>S M i"
proof (intro bexI impI CollectI conjI)
- show "simple_function ?g" using g
+ 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)
@@ -902,10 +919,10 @@
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> = simple_integral ?g" using g `space M \<noteq> {}`
+ 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 < simple_integral ?g" .
+ finally show "x < integral\<^isup>S M ?g" .
qed
qed
then show ?thesis by simp
@@ -914,40 +931,41 @@
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> \<nu> A = \<mu> A"
- shows "measure_space.positive_integral M \<nu> f = positive_integral f"
+ 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 M \<nu>
- by (rule measure_space_cong) fact
+ interpret v: measure_space N
+ by (rule measure_space_cong) fact+
with assms show ?thesis
- unfolding positive_integral_def v.positive_integral_def SUPR_def
+ unfolding positive_integral_def SUPR_def
by (auto intro!: arg_cong[where f=Sup] image_cong
- simp: simple_integral_cong_measure[of \<nu>])
+ simp: simple_integral_cong_measure[OF assms]
+ simple_function_cong_algebra[OF assms(2,3)])
qed
lemma (in measure_space) positive_integral_alt1:
- "positive_integral f =
- (SUP g : {g. simple_function g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}. simple_integral g)"
+ "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 g" "\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>"
- hence "?g \<le> f" "simple_function ?g" "simple_integral g = simple_integral ?g"
+ 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 "simple_integral g \<in> simple_integral ` {g. simple_function g \<and> g \<le> f \<and> \<omega> \<notin> g`space M}"
+ 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 g" "g \<le> f" "\<omega> \<notin> g`space M"
- hence "simple_function g" "\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>"
+ 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 "simple_integral g \<in> simple_integral ` {g. simple_function g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}"
+ 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_cong:
assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
- shows "positive_integral f = positive_integral g"
+ 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
@@ -955,30 +973,30 @@
qed
lemma (in measure_space) positive_integral_eq_simple_integral:
- assumes "simple_function f"
- shows "positive_integral f = simple_integral f"
+ 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 g" "g \<le> f"
- with assms show "simple_integral g \<le> simple_integral f"
+ 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 g \<and> g \<le> f} \<longrightarrow> simple_integral x \<le> y"
- with assms show "simple_integral f \<le> y" by auto
+ 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 "positive_integral u \<le> positive_integral v"
+ 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 a" and mono: "\<forall>x\<in>space M. a x \<le> u x \<and> a x \<noteq> \<omega>"
+ 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 (\<lambda>x. a x * indicator (space M - N) x)"
+ 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 g \<and> (\<forall>x\<in>space M. g x \<le> v x \<and> g x \<noteq> \<omega>)}.
- simple_integral a \<le> simple_integral b"
+ 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"
@@ -987,7 +1005,7 @@
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 a`[THEN borel_measurable_simple_function]
+ 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)
@@ -1010,12 +1028,12 @@
qed
lemma (in measure_space) positive_integral_cong_AE:
- "AE x. u x = v x \<Longrightarrow> positive_integral u = positive_integral v"
+ "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:
assumes mono: "\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x"
- shows "positive_integral u \<le> positive_integral v"
+ shows "integral\<^isup>P M u \<le> integral\<^isup>P M v"
using mono by (auto intro!: AE_cong positive_integral_mono_AE)
lemma image_set_cong:
@@ -1027,15 +1045,15 @@
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 u"
+ and "simple_function M u"
and le: "u \<le> s" and real: "\<omega> \<notin> u`space M"
- shows "simple_integral u \<le> (SUP i. positive_integral (f i))" (is "_ \<le> ?S")
+ 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"
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"
- using f `simple_function u` by (auto simp: borel_measurable_simple_function)
+ using f `simple_function M u` by (auto simp: borel_measurable_simple_function)
let "?uB i x" = "u x * indicator (?B i) x"
@@ -1049,7 +1067,7 @@
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 u` by (auto simp add: simple_function_def)
+ using `simple_function M u` by (auto simp add: simple_function_def)
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
@@ -1071,8 +1089,8 @@
qed auto
note measure_conv = measure_up[OF Int[OF u B] this]
- have "simple_integral u = (SUP i. simple_integral (?uB i))"
- unfolding simple_integral_indicator[OF B `simple_function u`]
+ 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)
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})
@@ -1082,52 +1100,51 @@
\<le> u x * \<mu> (u -` {u x} \<inter> space M \<inter> ?B (Suc n))"
by (auto intro: mult_left_mono)
next
- show "simple_integral u =
+ 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)
qed
moreover
- have "a * (SUP i. simple_integral (?uB i)) \<le> ?S"
+ have "a * (SUP i. integral\<^isup>S M (?uB i)) \<le> ?S"
unfolding pextreal_SUP_cmult[symmetric]
proof (safe intro!: SUP_mono bexI)
fix i
- have "a * simple_integral (?uB i) = (\<integral>\<^isup>Sx. a * ?uB i x)"
- using B `simple_function u`
+ 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)
- also have "\<dots> \<le> positive_integral (f i)"
+ 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 (\<lambda>x. a * ?uB i x)" using B assms(3)
+ 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)
qed
- finally show "a * simple_integral (?uB i) \<le> positive_integral (f i)"
+ finally show "a * integral\<^isup>S M (?uB i) \<le> integral\<^isup>P M (f i)"
by auto
qed simp
- ultimately show "a * simple_integral u \<le> ?S" by simp
+ ultimately show "a * integral\<^isup>S M u \<le> ?S" by simp
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. positive_integral (f i)) \<up> positive_integral u"
+ shows "(\<lambda>i. integral\<^isup>P M (f i)) \<up> integral\<^isup>P M u"
unfolding isoton_def
proof safe
- fix i show "positive_integral (f i) \<le> positive_integral (f (Suc i))"
+ 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
next
have u: "u = (SUP i. f i)" using `f \<up> u` unfolding isoton_def by auto
-
- show "(SUP i. positive_integral (f i)) = positive_integral u"
+ 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. positive_integral (f j)) \<le> positive_integral u"
+ 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 "positive_integral u \<le> (SUP i. positive_integral (f i))"
+ 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])
qed
@@ -1136,12 +1153,12 @@
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. positive_integral (f i)) = (\<integral>\<^isup>+ x. SUP i. f i x)"
- (is "_ = positive_integral ?u")
+ 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")
proof -
show ?thesis
proof (rule antisym)
- show "(SUP j. positive_integral (f j)) \<le> positive_integral ?u"
+ 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"
@@ -1151,26 +1168,26 @@
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 "positive_integral ru \<le> (SUP i. positive_integral (rf i))"
+ 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 "positive_integral ?u \<le> (SUP i. positive_integral (f i))"
+ 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)
qed
qed
lemma (in measure_space) SUP_simple_integral_sequences:
- assumes f: "f \<up> u" "\<And>i. simple_function (f i)"
- and g: "g \<up> u" "\<And>i. simple_function (g i)"
- shows "(SUP i. simple_integral (f i)) = (SUP i. simple_integral (g i))"
+ assumes f: "f \<up> u" "\<And>i. simple_function M (f i)"
+ and g: "g \<up> u" "\<And>i. simple_function M (g i)"
+ 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. positive_integral (f i))"
+ 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> = positive_integral u"
+ 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. positive_integral (g i))"
+ 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
also have "\<dots> = (SUP i. ?G i)"
@@ -1179,38 +1196,36 @@
qed
lemma (in measure_space) positive_integral_const[simp]:
- "(\<integral>\<^isup>+ x. c) = c * \<mu> (space M)"
+ "(\<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 (f i)"
- shows "(\<lambda>i. simple_integral (f i)) \<up> positive_integral u"
+ 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 (in measure_space) positive_integral_vimage:
assumes T: "sigma_algebra M'" "T \<in> measurable M M'" and f: "f \<in> borel_measurable M'"
- shows "measure_space.positive_integral M' (\<lambda>A. \<mu> (T -` A \<inter> space M)) f = (\<integral>\<^isup>+ x. f (T x))"
- (is "measure_space.positive_integral M' ?nu f = _")
+ and \<nu>: "\<And>A. A \<in> sets M' \<Longrightarrow> measure M' A = \<mu> (T -` A \<inter> space M)"
+ shows "integral\<^isup>P M' f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
proof -
- interpret T: measure_space M' ?nu using T by (rule measure_space_vimage) auto
- obtain f' where f': "f' \<up> f" "\<And>i. T.simple_function (f' i)"
+ interpret T: measure_space M' using \<nu> 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 (\<lambda>x. f' i (T x))"
+ 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] unfolding isoton_fun_expand by auto
- show "T.positive_integral f = (\<integral>\<^isup>+ x. f (T x))"
+ 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']
- unfolding simple_integral_vimage[OF T f'(2)] isoton_def
- by simp
+ by (simp add: simple_integral_vimage[OF T f'(2) \<nu>] isoton_def)
qed
lemma (in measure_space) positive_integral_linear:
assumes f: "f \<in> borel_measurable M"
and g: "g \<in> borel_measurable M"
- shows "(\<integral>\<^isup>+ x. a * f x + g x) =
- a * positive_integral f + positive_integral g"
- (is "positive_integral ?L = _")
+ 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)]
@@ -1222,46 +1237,45 @@
using assms by simp
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. simple_integral (l i)) = (SUP i. simple_integral (?L' i))"
+ 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 (?L' i)"
+ 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)
qed
moreover from u v have L'_isoton:
- "(\<lambda>i. simple_integral (?L' i)) \<up> a * positive_integral f + positive_integral g"
+ "(\<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)
qed
lemma (in measure_space) positive_integral_cmult:
assumes "f \<in> borel_measurable M"
- shows "(\<integral>\<^isup>+ x. c * f x) = c * positive_integral f"
+ 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
lemma (in measure_space) positive_integral_multc:
assumes "f \<in> borel_measurable M"
- shows "(\<integral>\<^isup>+ x. f x * c) = positive_integral f * 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
lemma (in measure_space) positive_integral_indicator[simp]:
- "A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. indicator A x) = \<mu> A"
+ "A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. indicator A x\<partial>M) = \<mu> A"
by (subst positive_integral_eq_simple_integral)
(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) = c * \<mu> A"
+ "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"
- shows "(\<integral>\<^isup>+ x. f x + g x) = positive_integral f + positive_integral g"
+ 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
lemma (in measure_space) positive_integral_setsum:
assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M"
- shows "(\<integral>\<^isup>+ x. \<Sum>i\<in>P. f i x) = (\<Sum>i\<in>P. positive_integral (f i))"
+ 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
@@ -1277,14 +1291,13 @@
lemma (in measure_space) positive_integral_diff:
assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M"
- and fin: "positive_integral g \<noteq> \<omega>"
+ and fin: "integral\<^isup>P M g \<noteq> \<omega>"
and mono: "\<And>x. x \<in> space M \<Longrightarrow> g x \<le> f x"
- shows "(\<integral>\<^isup>+ x. f x - g x) = positive_integral f - positive_integral g"
+ 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) + positive_integral g =
- positive_integral f"
+ 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"
@@ -1297,9 +1310,9 @@
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) = (\<Sum>\<^isub>\<infinity> i. positive_integral (f i))"
+ shows "(\<integral>\<^isup>+ x. (\<Sum>\<^isub>\<infinity> i. f i x) \<partial>M) = (\<Sum>\<^isub>\<infinity> i. integral\<^isup>P M (f i))"
proof -
- have "(\<lambda>i. (\<integral>\<^isup>+x. \<Sum>i<i. f i x)) \<up> (\<integral>\<^isup>+x. \<Sum>\<^isub>\<infinity>i. f i x)"
+ 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]
@@ -1312,79 +1325,86 @@
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) \<le>
- (SUP n. INF m. positive_integral (u (m + n)))"
+ 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)))"
proof -
- have "(\<integral>\<^isup>+x. SUP n. INF m. u (m + n) x)
- = (SUP n. (\<integral>\<^isup>+x. INF m. u (m + n) x))"
+ 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. positive_integral (u (m + n)))"
+ 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)
finally show ?thesis .
qed
lemma (in measure_space) measure_space_density:
assumes borel: "u \<in> borel_measurable M"
- shows "measure_space M (\<lambda>A. (\<integral>\<^isup>+ x. u x * indicator A x))" (is "measure_space M ?v")
-proof
- show "?v {} = 0" by simp
- show "countably_additive M ?v"
- unfolding countably_additive_def
- proof safe
- fix A :: "nat \<Rightarrow> 'a set"
- assume "range A \<subseteq> sets M"
- hence "\<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. ?v (A n)) = ?v (\<Union>x. A x)"
- by (simp add: positive_integral_psuminf[symmetric])
+ 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
+ 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])
+ qed
qed
qed
lemma (in measure_space) positive_integral_translated_density:
assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
- shows "measure_space.positive_integral M (\<lambda>A. (\<integral>\<^isup>+ x. f x * indicator A x)) g =
- (\<integral>\<^isup>+ x. f x * g x)" (is "measure_space.positive_integral M ?T _ = _")
+ 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)]
- interpret T: measure_space M ?T .
+ 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 (G i)" "G \<up> g" by blast
+ 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[OF `G \<up> g` G_borel]
- have *: "(\<lambda>i. T.positive_integral (G i)) \<up> T.positive_integral g" .
+ 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 "T.positive_integral (G i) = T.simple_integral (G i)"
+ 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))"
- apply (simp add: T.simple_integral_def)
+ 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_indicator borel_measurable_vimage)
+ 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_indicator borel_measurable_vimage)
+ 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)"
+ 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 "T.positive_integral (G i) = (\<integral>\<^isup>+x. f x * G i x)" . }
- with * have eq_Tg: "(\<lambda>i. (\<integral>\<^isup>+x. f x * G i x)) \<up> T.positive_integral g" by simp
+ 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)) \<up> (\<integral>\<^isup>+x. f x * g x)"
+ 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 "T.positive_integral g = (\<integral>\<^isup>+x. f x * g x)"
+ 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) = 0"
+ assumes "N \<in> null_sets" shows "(\<integral>\<^isup>+ x. u x * indicator N x \<partial>M) = 0"
proof -
- have "(\<integral>\<^isup>+ x. u x * indicator N x) = (\<integral>\<^isup>+ x. 0)"
+ have "(\<integral>\<^isup>+ x. u x * indicator N x \<partial>M) = (\<integral>\<^isup>+ x. 0 \<partial>M)"
proof (intro positive_integral_cong_AE AE_I)
show "{x \<in> space M. u x * indicator N x \<noteq> 0} \<subseteq> N"
by (auto simp: indicator_def)
@@ -1396,20 +1416,20 @@
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)"
+ 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` borel by auto
- hence "\<mu> ?A = (\<integral>\<^isup>+ x. indicator ?A x)"
+ 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))"
+ 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)"
+ 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 .
@@ -1417,11 +1437,11 @@
lemma (in measure_space) positive_integral_0_iff:
assumes borel: "u \<in> borel_measurable M"
- shows "positive_integral u = 0 \<longleftrightarrow> \<mu> {x\<in>space M. u x \<noteq> 0} = 0"
+ 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) = positive_integral u"
+ have u: "(\<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
@@ -1429,10 +1449,10 @@
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)" by simp
- thus "positive_integral u = 0" unfolding u by simp
+ 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
next
- assume *: "positive_integral u = 0"
+ assume *: "integral\<^isup>P M u = 0"
let "?M n" = "{x \<in> space M. 1 \<le> of_nat n * u x}"
have "0 = (SUP n. \<mu> (?M n \<inter> ?A))"
proof -
@@ -1469,34 +1489,34 @@
lemma (in measure_space) positive_integral_restricted:
assumes "A \<in> sets M"
- shows "measure_space.positive_integral (restricted_space A) \<mu> f = (\<integral>\<^isup>+ x. f x * indicator A x)"
- (is "measure_space.positive_integral ?R \<mu> f = positive_integral ?f")
+ 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 \<mu>" by (rule restricted_measure_space[OF `A \<in> sets M`])
- then interpret R: measure_space ?R \<mu> .
+ 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 *: "R.positive_integral f = R.positive_integral ?f"
+ have *: "integral\<^isup>P ?R f = integral\<^isup>P ?R ?f"
by (intro R.positive_integral_cong) auto
show ?thesis
- unfolding * R.positive_integral_def positive_integral_def
+ 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 (\<lambda>x. g x * indicator A x)"
+ fix g assume "simple_function M (\<lambda>x. g x * indicator A x)"
"g \<le> f"
- then show "\<exists>x. simple_function x \<and> x \<le> (\<lambda>x. f x * indicator A x) \<and>
- (\<integral>\<^isup>Sx. g x * indicator A x) = simple_integral x"
+ 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)
next
- fix g assume g: "simple_function g" "g \<le> (\<lambda>x. f x * indicator A x)"
+ 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 (\<lambda>xa. x xa * indicator A xa) \<and> x \<le> f \<and>
- simple_integral g = simple_integral (\<lambda>xa. x xa * indicator A xa)"
+ 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 *)
@@ -1505,103 +1525,113 @@
lemma (in measure_space) positive_integral_subalgebra:
assumes borel: "f \<in> borel_measurable N"
- and N: "sets N \<subseteq> sets M" "space N = space M" and sa: "sigma_algebra N"
- shows "measure_space.positive_integral N \<mu> f = positive_integral f"
+ 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 \<mu> using measure_space_subalgebra[OF sa N] .
+ 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. N.simple_function (fs i)" and "fs \<up> f" by blast
- then have sf: "\<And>i. simple_function (fs i)"
- using simple_function_subalgebra[OF _ N sa] by blast
- from positive_integral_isoton_simple[OF `fs \<up> f` sf] N.positive_integral_isoton_simple[OF `fs \<up> f` Nsf]
- show ?thesis unfolding isoton_def simple_integral_def N.simple_integral_def `space N = space M` by simp
+ 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]
+ 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
+ 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
+ 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
+ finally show ?thesis .
qed
section "Lebesgue Integral"
-definition (in measure_space) integrable where
- "integrable f \<longleftrightarrow> f \<in> borel_measurable M \<and>
- (\<integral>\<^isup>+ x. Real (f x)) \<noteq> \<omega> \<and>
- (\<integral>\<^isup>+ x. Real (- f x)) \<noteq> \<omega>"
+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>"
-lemma (in measure_space) integrableD[dest]:
- assumes "integrable f"
- shows "f \<in> borel_measurable M"
- "(\<integral>\<^isup>+ x. Real (f x)) \<noteq> \<omega>"
- "(\<integral>\<^isup>+ x. Real (- f x)) \<noteq> \<omega>"
+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>"
using assms unfolding integrable_def by auto
-definition (in measure_space) integral (binder "\<integral> " 10) where
- "integral f = real ((\<integral>\<^isup>+ x. Real (f x))) - real ((\<integral>\<^isup>+ x. Real (- f x)))"
+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))"
+
+syntax
+ "_lebesgue_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> real" ("\<integral> _. _ \<partial>_" [60,61] 110)
+
+translations
+ "\<integral> x. f \<partial>M" == "CONST integral\<^isup>L M (%x. f)"
lemma (in measure_space) integral_cong:
- assumes cong: "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
- shows "integral f = integral g"
- using assms by (simp cong: positive_integral_cong add: integral_def)
+ assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
+ shows "integral\<^isup>L M f = integral\<^isup>L M g"
+ using assms by (simp cong: positive_integral_cong add: lebesgue_integral_def)
lemma (in measure_space) integral_cong_measure:
- assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = \<mu> A"
- shows "measure_space.integral M \<nu> f = integral f"
+ 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 M \<nu>
- by (rule measure_space_cong) fact
+ interpret v: measure_space N
+ by (rule measure_space_cong) fact+
show ?thesis
- unfolding integral_def v.integral_def
- by (simp add: positive_integral_cong_measure[OF assms])
+ by (simp add: positive_integral_cong_measure[OF assms] lebesgue_integral_def)
qed
lemma (in measure_space) integral_cong_AE:
assumes cong: "AE x. f x = g x"
- shows "integral f = integral g"
+ shows "integral\<^isup>L M f = integral\<^isup>L M g"
proof -
have "AE x. Real (f x) = Real (g x)"
using cong by (rule AE_mp) simp
moreover have "AE x. Real (- f x) = Real (- g x)"
using cong by (rule AE_mp) simp
ultimately show ?thesis
- by (simp cong: positive_integral_cong_AE add: integral_def)
+ by (simp cong: positive_integral_cong_AE add: lebesgue_integral_def)
qed
lemma (in measure_space) integrable_cong:
- "(\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> integrable f \<longleftrightarrow> integrable g"
+ "(\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> integrable M f \<longleftrightarrow> integrable M g"
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 f = real ((\<integral>\<^isup>+ x. Real (f x)))"
+ shows "integral\<^isup>L M f = real (\<integral>\<^isup>+ x. Real (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: integral_def)
+ thus ?thesis by (simp del: Real_eq_0 add: lebesgue_integral_def)
qed
lemma (in measure_space) integral_vimage:
assumes T: "sigma_algebra M'" "T \<in> measurable M M'"
- assumes f: "measure_space.integrable M' (\<lambda>A. \<mu> (T -` A \<inter> space M)) f"
- shows "integrable (\<lambda>x. f (T x))" (is ?P)
- and "measure_space.integral M' (\<lambda>A. \<mu> (T -` A \<inter> space M)) f = (\<integral>x. f (T x))" (is ?I)
+ and \<nu>: "\<And>A. A \<in> sets M' \<Longrightarrow> measure M' A = \<mu> (T -` A \<inter> space M)"
+ assumes f: "integrable M' f"
+ shows "integrable M (\<lambda>x. f (T x))" (is ?P)
+ and "integral\<^isup>L M' f = (\<integral>x. f (T x) \<partial>M)" (is ?I)
proof -
- interpret T: measure_space M' "\<lambda>A. \<mu> (T -` A \<inter> space M)"
- using T by (rule measure_space_vimage) auto
+ interpret T: measure_space M' using \<nu> by (rule measure_space_vimage[OF T])
from measurable_comp[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'"
and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
- using f unfolding T.integrable_def by (auto simp: comp_def)
+ using f unfolding integrable_def by (auto simp: comp_def)
then show ?P ?I
- using f unfolding T.integral_def integral_def T.integrable_def integrable_def
- unfolding borel[THEN positive_integral_vimage[OF T]] by auto
+ using f unfolding lebesgue_integral_def integrable_def
+ by (auto simp: borel[THEN positive_integral_vimage[OF T], OF \<nu>])
qed
lemma (in measure_space) integral_minus[intro, simp]:
- assumes "integrable f"
- shows "integrable (\<lambda>x. - f x)" "(\<integral>x. - f x) = - integral f"
- using assms by (auto simp: integrable_def integral_def)
+ assumes "integrable M f"
+ shows "integrable M (\<lambda>x. - f x)" "(\<integral>x. - f x \<partial>M) = - integral\<^isup>L M f"
+ using assms by (auto simp: integrable_def lebesgue_integral_def)
lemma (in measure_space) integral_of_positive_diff:
- assumes integrable: "integrable u" "integrable v"
+ assumes integrable: "integrable M u" "integrable M v"
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 f" and "integral f = integral u - integral v"
+ shows "integrable M f" and "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v"
proof -
- let ?PI = positive_integral
let "?f x" = "Real (f x)"
let "?mf x" = "Real (- f x)"
let "?u x" = "Real (u x)"
@@ -1615,38 +1645,39 @@
"f \<in> borel_measurable M"
by (auto simp: f_def[symmetric] integrable_def)
- have "?PI (\<lambda>x. Real (u x - v x)) \<le> ?PI ?u"
+ 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 "?PI (\<lambda>x. Real (v x - u x)) \<le> ?PI ?v"
+ 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)
- ultimately show f: "integrable f"
- using `integrable u` `integrable v` `f \<in> borel_measurable M`
+ 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 (\<lambda>x. - f x)" ..
+ 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
have "\<And>x. ?u x + ?mf x = ?v x + ?f x"
unfolding f_def using pos by simp
- hence "?PI (\<lambda>x. ?u x + ?mf x) = ?PI (\<lambda>x. ?v x + ?f x)" by simp
- hence "real (?PI ?u + ?PI ?mf) = real (?PI ?v + ?PI ?f)"
+ 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) =
+ 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]
by auto
- then show "integral f = integral u - integral v"
- using f mf `integrable u` `integrable v`
- unfolding integral_def integrable_def *
- by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?v", cases "?PI ?u")
+ 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)
qed
lemma (in measure_space) integral_linear:
- assumes "integrable f" "integrable g" and "0 \<le> a"
- shows "integrable (\<lambda>t. a * f t + g t)"
- and "(\<integral> t. a * f t + g t) = a * integral f + integral g"
+ 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"
proof -
- let ?PI = positive_integral
+ let ?PI = "integral\<^isup>P M"
let "?f x" = "Real (f x)"
let "?g x" = "Real (g x)"
let "?mf x" = "Real (- f x)"
@@ -1670,37 +1701,36 @@
positive_integral_linear[OF pos]
positive_integral_linear[OF neg]
- have "integrable ?p" "integrable ?n"
+ 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
note diff = integral_of_positive_diff[OF this]
- show "integrable (\<lambda>t. a * f t + g t)" by (rule diff)
+ show "integrable M (\<lambda>t. a * f t + g t)" by (rule diff)
- from assms show "(\<integral> t. a * f t + g t) = a * integral f + integral g"
- unfolding diff(2) unfolding integral_def * linear integrable_def
+ 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)
qed
lemma (in measure_space) integral_add[simp, intro]:
- assumes "integrable f" "integrable g"
- shows "integrable (\<lambda>t. f t + g t)"
- and "(\<integral> t. f t + g t) = integral f + integral g"
+ assumes "integrable M f" "integrable M g"
+ shows "integrable M (\<lambda>t. f t + g t)"
+ and "(\<integral> t. f t + g t \<partial>M) = integral\<^isup>L M f + integral\<^isup>L M g"
using assms integral_linear[where a=1] by auto
lemma (in measure_space) integral_zero[simp, intro]:
- shows "integrable (\<lambda>x. 0)"
- and "(\<integral> x.0) = 0"
- unfolding integrable_def integral_def
+ shows "integrable M (\<lambda>x. 0)" "(\<integral> x.0 \<partial>M) = 0"
+ unfolding integrable_def lebesgue_integral_def
by (auto simp add: borel_measurable_const)
lemma (in measure_space) integral_cmult[simp, intro]:
- assumes "integrable f"
- shows "integrable (\<lambda>t. a * f t)" (is ?P)
- and "(\<integral> t. a * f t) = a * integral f" (is ?I)
+ assumes "integrable M f"
+ shows "integrable M (\<lambda>t. a * f t)" (is ?P)
+ and "(\<integral> t. a * f t \<partial>M) = a * integral\<^isup>L M f" (is ?I)
proof -
- have "integrable (\<lambda>t. a * f t) \<and> (\<integral> t. a * f t) = a * integral f"
+ have "integrable M (\<lambda>t. a * f t) \<and> (\<integral> t. a * f t \<partial>M) = a * integral\<^isup>L M f"
proof (cases rule: le_cases)
assume "0 \<le> a" show ?thesis
using integral_linear[OF assms integral_zero(1) `0 \<le> a`]
@@ -1716,56 +1746,56 @@
qed
lemma (in measure_space) integral_multc:
- assumes "integrable f"
- shows "(\<integral> x. f x * c) = integral f * c"
+ assumes "integrable M f"
+ shows "(\<integral> x. f x * c \<partial>M) = integral\<^isup>L M f * c"
unfolding mult_commute[of _ c] integral_cmult[OF assms] ..
lemma (in measure_space) integral_mono_AE:
- assumes fg: "integrable f" "integrable g"
+ assumes fg: "integrable M f" "integrable M g"
and mono: "AE t. f t \<le> g t"
- shows "integral f \<le> integral g"
+ shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
proof -
have "AE x. Real (f x) \<le> Real (g x)"
using mono by (rule AE_mp) (auto intro!: AE_cong)
- moreover have "AE x. Real (- g x) \<le> Real (- f x)"
+ moreover have "AE x. Real (- g x) \<le> Real (- f x)"
using mono by (rule AE_mp) (auto intro!: AE_cong)
ultimately show ?thesis using fg
- by (auto simp: integral_def integrable_def diff_minus
+ by (auto simp: lebesgue_integral_def integrable_def diff_minus
intro!: add_mono real_of_pextreal_mono positive_integral_mono_AE)
qed
lemma (in measure_space) integral_mono:
- assumes fg: "integrable f" "integrable g"
+ assumes fg: "integrable M f" "integrable M g"
and mono: "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"
- shows "integral f \<le> integral g"
+ shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
apply (rule integral_mono_AE[OF fg])
using mono by (rule AE_cong) auto
lemma (in measure_space) integral_diff[simp, intro]:
- assumes f: "integrable f" and g: "integrable g"
- shows "integrable (\<lambda>t. f t - g t)"
- and "(\<integral> t. f t - g t) = integral f - integral g"
+ assumes f: "integrable M f" and g: "integrable M g"
+ shows "integrable M (\<lambda>t. f t - g t)"
+ and "(\<integral> t. f t - g t \<partial>M) = integral\<^isup>L M f - integral\<^isup>L M g"
using integral_add[OF f integral_minus(1)[OF g]]
unfolding diff_minus integral_minus(2)[OF g]
by auto
lemma (in measure_space) integral_indicator[simp, intro]:
assumes "a \<in> sets M" and "\<mu> a \<noteq> \<omega>"
- shows "integral (indicator a) = real (\<mu> a)" (is ?int)
- and "integrable (indicator a)" (is ?able)
+ 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
show ?int ?able
- using assms unfolding integral_def integrable_def
+ 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>"
- shows "integrable (\<lambda>x. c * indicator A x)" (is ?P)
- and "(\<integral>x. c * indicator A x) = c * real (\<mu> A)" (is ?I)
+ 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 -
show ?P
proof (cases "c = 0")
@@ -1779,9 +1809,9 @@
qed
lemma (in measure_space) integral_setsum[simp, intro]:
- assumes "\<And>n. n \<in> S \<Longrightarrow> integrable (f n)"
- shows "(\<integral>x. \<Sum> i \<in> S. f i x) = (\<Sum> i \<in> S. integral (f i))" (is "?int S")
- and "integrable (\<lambda>x. \<Sum> i \<in> S. f i x)" (is "?I S")
+ assumes "\<And>n. n \<in> S \<Longrightarrow> integrable M (f n)"
+ shows "(\<integral>x. (\<Sum> i \<in> S. f i x) \<partial>M) = (\<Sum> i \<in> S. integral\<^isup>L M (f i))" (is "?int S")
+ and "integrable M (\<lambda>x. \<Sum> i \<in> S. f i x)" (is "?I S")
proof -
have "?int S \<and> ?I S"
proof (cases "finite S")
@@ -1792,8 +1822,8 @@
qed
lemma (in measure_space) integrable_abs:
- assumes "integrable f"
- shows "integrable (\<lambda> x. \<bar>f x\<bar>)"
+ 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)"
@@ -1808,56 +1838,55 @@
lemma (in measure_space) integral_subalgebra:
assumes borel: "f \<in> borel_measurable N"
- and N: "sets N \<subseteq> sets M" "space N = space M" and sa: "sigma_algebra N"
- shows "measure_space.integrable N \<mu> f \<longleftrightarrow> integrable f" (is ?P)
- and "measure_space.integral N \<mu> f = integral f" (is ?I)
+ 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 "integrable N f \<longleftrightarrow> integrable M f" (is ?P)
+ and "integral\<^isup>L N f = integral\<^isup>L M f" (is ?I)
proof -
- interpret N: measure_space N \<mu> using measure_space_subalgebra[OF sa N] .
+ 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"
using assms unfolding measurable_def by auto
- then show ?P ?I unfolding integrable_def N.integrable_def integral_def N.integral_def
- unfolding * by auto
+ then show ?P ?I by (auto simp: * integrable_def lebesgue_integral_def)
qed
lemma (in measure_space) integrable_bound:
- assumes "integrable f"
+ assumes "integrable M f"
and f: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
"\<And>x. x \<in> space M \<Longrightarrow> \<bar>g x\<bar> \<le> f x"
assumes borel: "g \<in> borel_measurable M"
- shows "integrable g"
+ shows "integrable M g"
proof -
- have "(\<integral>\<^isup>+ x. Real (g x)) \<le> (\<integral>\<^isup>+ x. Real \<bar>g x\<bar>)"
+ have "(\<integral>\<^isup>+ x. Real (g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. Real \<bar>g x\<bar> \<partial>M)"
by (auto intro!: positive_integral_mono)
- also have "\<dots> \<le> (\<integral>\<^isup>+ x. Real (f x))"
+ also have "\<dots> \<le> (\<integral>\<^isup>+ x. Real (f x) \<partial>M)"
using f by (auto intro!: positive_integral_mono)
also have "\<dots> < \<omega>"
- using `integrable f` unfolding integrable_def by (auto simp: pextreal_less_\<omega>)
- finally have pos: "(\<integral>\<^isup>+ x. Real (g x)) < \<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>" .
- have "(\<integral>\<^isup>+ x. Real (- g x)) \<le> (\<integral>\<^isup>+ x. Real (\<bar>g x\<bar>))"
+ have "(\<integral>\<^isup>+ x. Real (- g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. Real (\<bar>g x\<bar>) \<partial>M)"
by (auto intro!: positive_integral_mono)
- also have "\<dots> \<le> (\<integral>\<^isup>+ x. Real (f x))"
+ also have "\<dots> \<le> (\<integral>\<^isup>+ x. Real (f x) \<partial>M)"
using f by (auto intro!: positive_integral_mono)
also have "\<dots> < \<omega>"
- using `integrable f` unfolding integrable_def by (auto simp: pextreal_less_\<omega>)
- finally have neg: "(\<integral>\<^isup>+ x. Real (- g x)) < \<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>" .
from neg pos borel show ?thesis
unfolding integrable_def by auto
qed
lemma (in measure_space) integrable_abs_iff:
- "f \<in> borel_measurable M \<Longrightarrow> integrable (\<lambda> x. \<bar>f x\<bar>) \<longleftrightarrow> integrable f"
+ "f \<in> borel_measurable M \<Longrightarrow> integrable M (\<lambda> x. \<bar>f x\<bar>) \<longleftrightarrow> integrable M f"
by (auto intro!: integrable_bound[where g=f] integrable_abs)
lemma (in measure_space) integrable_max:
- assumes int: "integrable f" "integrable g"
- shows "integrable (\<lambda> x. max (f x) (g x))"
+ assumes int: "integrable M f" "integrable M g"
+ shows "integrable M (\<lambda> x. max (f x) (g x))"
proof (rule integrable_bound)
- show "integrable (\<lambda>x. \<bar>f x\<bar> + \<bar>g x\<bar>)"
+ show "integrable M (\<lambda>x. \<bar>f x\<bar> + \<bar>g x\<bar>)"
using int by (simp add: integrable_abs)
show "(\<lambda>x. max (f x) (g x)) \<in> borel_measurable M"
using int unfolding integrable_def by auto
@@ -1868,10 +1897,10 @@
qed
lemma (in measure_space) integrable_min:
- assumes int: "integrable f" "integrable g"
- shows "integrable (\<lambda> x. min (f x) (g x))"
+ assumes int: "integrable M f" "integrable M g"
+ shows "integrable M (\<lambda> x. min (f x) (g x))"
proof (rule integrable_bound)
- show "integrable (\<lambda>x. \<bar>f x\<bar> + \<bar>g x\<bar>)"
+ show "integrable M (\<lambda>x. \<bar>f x\<bar> + \<bar>g x\<bar>)"
using int by (simp add: integrable_abs)
show "(\<lambda>x. min (f x) (g x)) \<in> borel_measurable M"
using int unfolding integrable_def by auto
@@ -1882,33 +1911,33 @@
qed
lemma (in measure_space) integral_triangle_inequality:
- assumes "integrable f"
- shows "\<bar>integral f\<bar> \<le> (\<integral>x. \<bar>f x\<bar>)"
+ assumes "integrable M f"
+ shows "\<bar>integral\<^isup>L M f\<bar> \<le> (\<integral>x. \<bar>f x\<bar> \<partial>M)"
proof -
- have "\<bar>integral f\<bar> = max (integral f) (- integral f)" by auto
- also have "\<dots> \<le> (\<integral>x. \<bar>f x\<bar>)"
+ have "\<bar>integral\<^isup>L M f\<bar> = max (integral\<^isup>L M f) (- integral\<^isup>L M f)" by auto
+ also have "\<dots> \<le> (\<integral>x. \<bar>f x\<bar> \<partial>M)"
using assms integral_minus(2)[of f, symmetric]
by (auto intro!: integral_mono integrable_abs simp del: integral_minus)
finally show ?thesis .
qed
lemma (in measure_space) integral_positive:
- assumes "integrable f" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
- shows "0 \<le> integral f"
+ assumes "integrable M f" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
+ shows "0 \<le> integral\<^isup>L M f"
proof -
- have "0 = (\<integral>x. 0)" by (auto simp: integral_zero)
- also have "\<dots> \<le> integral f"
+ have "0 = (\<integral>x. 0 \<partial>M)" by (auto simp: integral_zero)
+ also have "\<dots> \<le> integral\<^isup>L M f"
using assms by (rule integral_mono[OF integral_zero(1)])
finally show ?thesis .
qed
lemma (in measure_space) integral_monotone_convergence_pos:
- assumes i: "\<And>i. integrable (f i)" and mono: "\<And>x. mono (\<lambda>n. f n x)"
+ assumes i: "\<And>i. integrable M (f i)" and mono: "\<And>x. mono (\<lambda>n. f n x)"
and pos: "\<And>x i. 0 \<le> f i x"
and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
- and ilim: "(\<lambda>i. integral (f i)) ----> x"
- shows "integrable u"
- and "integral u = x"
+ and ilim: "(\<lambda>i. integral\<^isup>L M (f i)) ----> x"
+ shows "integrable M u"
+ and "integral\<^isup>L M u = x"
proof -
{ fix x have "0 \<le> u x"
using mono pos[of 0 x] incseq_le[OF _ lim, of x 0]
@@ -1928,43 +1957,42 @@
hence borel_u: "u \<in> borel_measurable M"
using pos_u by (auto simp: borel_measurable_Real_eq SUP_F)
- have integral_eq: "\<And>n. (\<integral>\<^isup>+ x. Real (f n x)) = Real (integral (f n))"
- using i unfolding integral_def integrable_def by (auto simp: Real_real)
+ 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)
- have pos_integral: "\<And>n. 0 \<le> integral (f n)"
+ 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))) \<up> (\<integral>\<^isup>+ x. Real (u x))"
+ 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)) =
- (SUP n. (\<integral>\<^isup>+ x. Real (f n x)))"
+ 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
proof (rule SUP_eq_LIMSEQ[THEN iffD2])
- show "mono (\<lambda>n. integral (f n))"
+ 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 (f n)" using pos_integral .
+ 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 (f n)) ----> x" using ilim .
+ show "(\<lambda>n. integral\<^isup>L M (f n)) ----> x" using ilim .
qed
- finally show "integrable u" "integral u = x" using borel_u `0 \<le> x`
- unfolding integrable_def integral_def by auto
+ finally show "integrable M u" "integral\<^isup>L M u = x" using borel_u `0 \<le> x`
+ unfolding integrable_def lebesgue_integral_def by auto
qed
lemma (in measure_space) integral_monotone_convergence:
- assumes f: "\<And>i. integrable (f i)" and "mono f"
+ assumes f: "\<And>i. integrable M (f i)" and "mono f"
and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
- and ilim: "(\<lambda>i. integral (f i)) ----> x"
- shows "integrable u"
- and "integral u = x"
+ and ilim: "(\<lambda>i. integral\<^isup>L M (f i)) ----> x"
+ shows "integrable M u"
+ and "integral\<^isup>L M u = x"
proof -
- have 1: "\<And>i. integrable (\<lambda>x. f i x - f 0 x)"
+ have 1: "\<And>i. integrable M (\<lambda>x. f i x - f 0 x)"
using f by (auto intro!: integral_diff)
have 2: "\<And>x. mono (\<lambda>n. f n x - f 0 x)" using `mono f`
unfolding mono_def le_fun_def by auto
@@ -1972,43 +2000,43 @@
unfolding mono_def le_fun_def by (auto simp: field_simps)
have 4: "\<And>x. (\<lambda>i. f i x - f 0 x) ----> u x - f 0 x"
using lim by (auto intro!: LIMSEQ_diff)
- have 5: "(\<lambda>i. (\<integral>x. f i x - f 0 x)) ----> x - integral (f 0)"
+ have 5: "(\<lambda>i. (\<integral>x. f i x - f 0 x \<partial>M)) ----> x - integral\<^isup>L M (f 0)"
using f ilim by (auto intro!: LIMSEQ_diff simp: integral_diff)
note diff = integral_monotone_convergence_pos[OF 1 2 3 4 5]
- have "integrable (\<lambda>x. (u x - f 0 x) + f 0 x)"
+ have "integrable M (\<lambda>x. (u x - f 0 x) + f 0 x)"
using diff(1) f by (rule integral_add(1))
- with diff(2) f show "integrable u" "integral u = x"
+ with diff(2) f show "integrable M u" "integral\<^isup>L M u = x"
by (auto simp: integral_diff)
qed
lemma (in measure_space) integral_0_iff:
- assumes "integrable f"
- shows "(\<integral>x. \<bar>f x\<bar>) = 0 \<longleftrightarrow> \<mu> {x\<in>space M. f x \<noteq> 0} = 0"
+ 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 "integrable (\<lambda>x. \<bar>f x\<bar>)" using assms by (rule integrable_abs)
+ 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>) \<noteq> \<omega>" unfolding integrable_def by auto
+ "(\<integral>\<^isup>+ x. Real \<bar>f x\<bar> \<partial>M) \<noteq> \<omega>" unfolding integrable_def by auto
from positive_integral_0_iff[OF this(1)] this(2)
- show ?thesis unfolding integral_def *
+ show ?thesis unfolding lebesgue_integral_def *
by (simp add: real_of_pextreal_eq_0)
qed
lemma (in measure_space) positive_integral_omega:
assumes "f \<in> borel_measurable M"
- and "positive_integral f \<noteq> \<omega>"
+ and "integral\<^isup>P M f \<noteq> \<omega>"
shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
proof -
- have "\<omega> * \<mu> (f -` {\<omega>} \<inter> space M) = (\<integral>\<^isup>+ x. \<omega> * indicator (f -` {\<omega>} \<inter> space M) x)"
+ 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> positive_integral f"
+ 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
qed
lemma (in measure_space) positive_integral_omega_AE:
- assumes "f \<in> borel_measurable M" "positive_integral f \<noteq> \<omega>" shows "AE x. f x \<noteq> \<omega>"
+ assumes "f \<in> borel_measurable M" "integral\<^isup>P M f \<noteq> \<omega>" shows "AE x. f x \<noteq> \<omega>"
proof (rule AE_I)
show "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
by (rule positive_integral_omega[OF assms])
@@ -2017,39 +2045,39 @@
qed auto
lemma (in measure_space) simple_integral_omega:
- assumes "simple_function f"
- and "simple_integral f \<noteq> \<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)
show "f \<in> borel_measurable M" using assms by (auto intro: borel_measurable_simple_function)
- show "positive_integral f \<noteq> \<omega>"
+ show "integral\<^isup>P M f \<noteq> \<omega>"
using assms by (simp add: positive_integral_eq_simple_integral)
qed
lemma (in measure_space) integral_real:
fixes f :: "'a \<Rightarrow> pextreal"
assumes "AE x. f x \<noteq> \<omega>"
- shows "(\<integral>x. real (f x)) = real (positive_integral f)" (is ?plus)
- and "(\<integral>x. - real (f x)) = - real (positive_integral f)" (is ?minus)
+ 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))) = positive_integral f"
+ have "(\<integral>\<^isup>+ x. Real (real (f x)) \<partial>M) = integral\<^isup>P M f"
apply (rule positive_integral_cong_AE)
apply (rule AE_mp[OF assms(1)])
by (auto intro!: AE_cong simp: Real_real)
moreover
- have "(\<integral>\<^isup>+ x. Real (- real (f x))) = (\<integral>\<^isup>+ x. 0)"
+ 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: integral_def integrable_def)
+ by (auto simp: lebesgue_integral_def integrable_def)
qed
lemma (in measure_space) integral_dominated_convergence:
- assumes u: "\<And>i. integrable (u i)" and bound: "\<And>x j. x\<in>space M \<Longrightarrow> \<bar>u j x\<bar> \<le> w x"
- and w: "integrable w" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> w x"
+ 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"
+ and w: "integrable M w" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> w x"
and u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
- shows "integrable u'"
- and "(\<lambda>i. (\<integral>x. \<bar>u i x - u' x\<bar>)) ----> 0" (is "?lim_diff")
- and "(\<lambda>i. integral (u i)) ----> integral u'" (is ?lim)
+ shows "integrable M u'"
+ and "(\<lambda>i. (\<integral>x. \<bar>u i x - u' x\<bar> \<partial>M)) ----> 0" (is "?lim_diff")
+ and "(\<lambda>i. integral\<^isup>L M (u i)) ----> integral\<^isup>L M u'" (is ?lim)
proof -
{ fix x j assume x: "x \<in> space M"
from u'[OF x] have "(\<lambda>i. \<bar>u i x\<bar>) ----> \<bar>u' x\<bar>" by (rule LIMSEQ_imp_rabs)
@@ -2061,9 +2089,9 @@
have u'_borel: "u' \<in> borel_measurable M"
using u' by (blast intro: borel_measurable_LIMSEQ[of u])
- show "integrable u'"
+ show "integrable M u'"
proof (rule integrable_bound)
- show "integrable w" by fact
+ show "integrable M w" by fact
show "u' \<in> borel_measurable M" by fact
next
fix x assume x: "x \<in> space M"
@@ -2072,8 +2100,8 @@
qed
let "?diff n x" = "2 * w x - \<bar>u n x - u' x\<bar>"
- have diff: "\<And>n. integrable (\<lambda>x. \<bar>u n x - u' x\<bar>)"
- using w u `integrable u'`
+ have diff: "\<And>n. integrable M (\<lambda>x. \<bar>u n x - u' x\<bar>)"
+ using w u `integrable M u'`
by (auto intro!: integral_add integral_diff integral_cmult integrable_abs)
{ fix j x assume x: "x \<in> space M"
@@ -2083,31 +2111,31 @@
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)) =
- (\<integral>\<^isup>+ x. Real (2 * w x)) - (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar>)"
+ 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)"
using diff w diff_less_2w
by (subst positive_integral_diff[symmetric])
(auto simp: integrable_def intro!: positive_integral_cong)
- have "integrable (\<lambda>x. 2 * w x)"
+ 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)) \<noteq> \<omega>" and
+ 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"
unfolding integrable_def by auto
- have "(INF n. SUP m. (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar>)) = 0" (is "?lim_SUP = 0")
+ have "(INF n. SUP m. (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M)) = 0" (is "?lim_SUP = 0")
proof cases
- assume eq_0: "(\<integral>\<^isup>+ x. Real (2 * w x)) = 0"
- have "\<And>i. (\<integral>\<^isup>+ x. Real \<bar>u i x - u' x\<bar>) \<le> (\<integral>\<^isup>+ x. Real (2 * w x))"
+ 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>) = 0" using eq_0 by auto
+ 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
next
- assume neq_0: "(\<integral>\<^isup>+ x. Real (2 * w x)) \<noteq> 0"
- have "(\<integral>\<^isup>+ x. Real (2 * w x)) = (\<integral>\<^isup>+ x. SUP n. INF m. Real (?diff (m + n) x))"
+ 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)
fix x assume x: "x \<in> space M"
show "Real (2 * w x) = (SUP n. INF m. Real (?diff (n + m) x))"
@@ -2119,22 +2147,22 @@
thus "(\<lambda>i. ?diff i x) ----> 2 * w x" by simp
qed
qed
- also have "\<dots> \<le> (SUP n. INF m. (\<integral>\<^isup>+ x. Real (?diff (m + n) x)))"
+ also have "\<dots> \<le> (SUP n. INF m. (\<integral>\<^isup>+ x. Real (?diff (m + 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)) -
- (INF n. SUP m. (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar>))"
+ 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)
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>) =
- Real ((\<integral>x. \<bar>u (n + m) x - u' x\<bar>))"
- using diff by (subst add_commute) (simp add: integral_def integrable_def Real_real)
+ 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>)) \<le> ?lim_SUP"
+ 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)
@@ -2143,28 +2171,28 @@
proof (rule LIMSEQ_I)
fix r :: real assume "0 < r"
from LIMSEQ_D[OF `?lim_diff` this]
- obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> (\<integral>x. \<bar>u n x - u' x\<bar>) < r"
+ obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M) < r"
using diff by (auto simp: integral_positive)
- show "\<exists>N. \<forall>n\<ge>N. norm (integral (u n) - integral u') < r"
+ show "\<exists>N. \<forall>n\<ge>N. norm (integral\<^isup>L M (u n) - integral\<^isup>L M u') < r"
proof (safe intro!: exI[of _ N])
fix n assume "N \<le> n"
- have "\<bar>integral (u n) - integral u'\<bar> = \<bar>(\<integral>x. u n x - u' x)\<bar>"
- using u `integrable u'` by (auto simp: integral_diff)
- also have "\<dots> \<le> (\<integral>x. \<bar>u n x - u' x\<bar>)" using u `integrable u'`
+ have "\<bar>integral\<^isup>L M (u n) - integral\<^isup>L M u'\<bar> = \<bar>(\<integral>x. u n x - u' x \<partial>M)\<bar>"
+ using u `integrable M u'` by (auto simp: integral_diff)
+ also have "\<dots> \<le> (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M)" using u `integrable M u'`
by (rule_tac integral_triangle_inequality) (auto intro!: integral_diff)
also note N[OF `N \<le> n`]
- finally show "norm (integral (u n) - integral u') < r" by simp
+ finally show "norm (integral\<^isup>L M (u n) - integral\<^isup>L M u') < r" by simp
qed
qed
qed
lemma (in measure_space) integral_sums:
- assumes borel: "\<And>i. integrable (f i)"
+ assumes borel: "\<And>i. integrable M (f i)"
and summable: "\<And>x. x \<in> space M \<Longrightarrow> summable (\<lambda>i. \<bar>f i x\<bar>)"
- and sums: "summable (\<lambda>i. (\<integral>x. \<bar>f i x\<bar>))"
- shows "integrable (\<lambda>x. (\<Sum>i. f i x))" (is "integrable ?S")
- and "(\<lambda>i. integral (f i)) sums (\<integral>x. (\<Sum>i. f i x))" (is ?integral)
+ and sums: "summable (\<lambda>i. (\<integral>x. \<bar>f i x\<bar> \<partial>M))"
+ shows "integrable M (\<lambda>x. (\<Sum>i. f i x))" (is "integrable M ?S")
+ and "(\<lambda>i. integral\<^isup>L M (f i)) sums (\<integral>x. (\<Sum>i. f i x) \<partial>M)" (is ?integral)
proof -
have "\<forall>x\<in>space M. \<exists>w. (\<lambda>i. \<bar>f i x\<bar>) sums w"
using summable unfolding summable_def by auto
@@ -2173,10 +2201,10 @@
let "?w y" = "if y \<in> space M then w y else 0"
- obtain x where abs_sum: "(\<lambda>i. (\<integral>x. \<bar>f i x\<bar>)) sums x"
+ obtain x where abs_sum: "(\<lambda>i. (\<integral>x. \<bar>f i x\<bar> \<partial>M)) sums x"
using sums unfolding summable_def ..
- have 1: "\<And>n. integrable (\<lambda>x. \<Sum>i = 0..<n. f i x)"
+ have 1: "\<And>n. integrable M (\<lambda>x. \<Sum>i = 0..<n. f i x)"
using borel by (auto intro!: integral_setsum)
{ fix j x assume [simp]: "x \<in> space M"
@@ -2185,21 +2213,21 @@
finally have "\<bar>\<Sum>i = 0..<j. f i x\<bar> \<le> ?w x" by simp }
note 2 = this
- have 3: "integrable ?w"
+ have 3: "integrable M ?w"
proof (rule integral_monotone_convergence(1))
let "?F n y" = "(\<Sum>i = 0..<n. \<bar>f i y\<bar>)"
let "?w' n y" = "if y \<in> space M then ?F n y else 0"
- have "\<And>n. integrable (?F n)"
+ have "\<And>n. integrable M (?F n)"
using borel by (auto intro!: integral_setsum integrable_abs)
- thus "\<And>n. integrable (?w' n)" by (simp cong: integrable_cong)
+ thus "\<And>n. integrable M (?w' n)" by (simp cong: integrable_cong)
show "mono ?w'"
by (auto simp: mono_def le_fun_def intro!: setsum_mono2)
{ fix x show "(\<lambda>n. ?w' n x) ----> ?w x"
using w by (cases "x \<in> space M") (simp_all add: LIMSEQ_const sums_def) }
- have *: "\<And>n. integral (?w' n) = (\<Sum>i = 0..< n. (\<integral>x. \<bar>f i x\<bar>))"
+ have *: "\<And>n. integral\<^isup>L M (?w' n) = (\<Sum>i = 0..< n. (\<integral>x. \<bar>f i x\<bar> \<partial>M))"
using borel by (simp add: integral_setsum integrable_abs cong: integral_cong)
from abs_sum
- show "(\<lambda>i. integral (?w' i)) ----> x" unfolding * sums_def .
+ show "(\<lambda>i. integral\<^isup>L M (?w' i)) ----> x" unfolding * sums_def .
qed
have 4: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> ?w x" using 2[of _ 0] by simp
@@ -2210,7 +2238,7 @@
note int = integral_dominated_convergence(1,3)[OF 1 2 3 4 5]
- from int show "integrable ?S" by simp
+ from int show "integrable M ?S" by simp
show ?integral unfolding sums_def integral_setsum(1)[symmetric, OF borel]
using int(2) by simp
@@ -2224,12 +2252,12 @@
and enum_zero: "enum ` (-S) \<subseteq> {0}"
and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<omega>"
and abs_summable: "summable (\<lambda>r. \<bar>enum r * real (\<mu> (f -` {enum r} \<inter> space M))\<bar>)"
- shows "integrable f"
- and "(\<lambda>r. enum r * real (\<mu> (f -` {enum r} \<inter> space M))) sums integral f" (is ?sums)
+ 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)
proof -
let "?A r" = "f -` {enum r} \<inter> space M"
let "?F r x" = "enum r * indicator (?A r) x"
- have enum_eq: "\<And>r. enum r * real (\<mu> (?A r)) = integral (?F r)"
+ have enum_eq: "\<And>r. enum r * real (\<mu> (?A r)) = integral\<^isup>L M (?F r)"
using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
{ fix x assume "x \<in> space M"
@@ -2250,19 +2278,19 @@
by (auto intro!: sums_single simp: F F_abs) }
note F_sums_f = this(1) and F_abs_sums_f = this(2)
- have int_f: "integral f = (\<integral>x. \<Sum>r. ?F r x)" "integrable f = integrable (\<lambda>x. \<Sum>r. ?F r x)"
+ have int_f: "integral\<^isup>L M f = (\<integral>x. (\<Sum>r. ?F r x) \<partial>M)" "integrable M f = integrable M (\<lambda>x. \<Sum>r. ?F r x)"
using F_sums_f by (auto intro!: integral_cong integrable_cong simp: sums_iff)
{ fix r
- have "(\<integral>x. \<bar>?F r x\<bar>) = (\<integral>x. \<bar>enum r\<bar> * indicator (?A r) x)"
+ have "(\<integral>x. \<bar>?F r x\<bar> \<partial>M) = (\<integral>x. \<bar>enum r\<bar> * indicator (?A r) x \<partial>M)"
by (auto simp: indicator_def intro!: integral_cong)
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>) = \<bar>enum r * real (\<mu> (?A r))\<bar>"
+ 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) }
note int_abs_F = this
- have 1: "\<And>i. integrable (\<lambda>x. ?F i x)"
+ have 1: "\<And>i. integrable M (\<lambda>x. ?F i x)"
using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
have 2: "\<And>x. x \<in> space M \<Longrightarrow> summable (\<lambda>i. \<bar>?F i x\<bar>)"
@@ -2272,7 +2300,7 @@
show ?sums unfolding enum_eq int_f by simp
from integral_sums(1)[OF 1 2, unfolded int_abs_F, OF _ abs_summable]
- show "integrable f" unfolding int_f by simp
+ show "integrable M f" unfolding int_f by simp
qed
section "Lebesgue integration on finite space"
@@ -2280,8 +2308,8 @@
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>"
- shows "integrable f"
- and "(\<integral>x. f x) =
+ 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")
proof -
let "?A r" = "f -` {r} \<inter> space M"
@@ -2295,40 +2323,40 @@
finally have "f x = ?S x" . }
note f_eq = this
- have f_eq_S: "integrable f \<longleftrightarrow> integrable ?S" "integral f = integral ?S"
+ have f_eq_S: "integrable M f \<longleftrightarrow> integrable M ?S" "integral\<^isup>L M f = integral\<^isup>L M ?S"
by (auto intro!: integrable_cong integral_cong simp only: f_eq)
- show "integrable f" ?integral using fin f f_eq_S
+ show "integrable M f" ?integral using fin f f_eq_S
by (simp_all add: integral_cmul_indicator borel_measurable_vimage)
qed
-lemma (in finite_measure_space) simple_function_finite[simp, intro]: "simple_function f"
+lemma (in finite_measure_space) simple_function_finite[simp, intro]: "simple_function M f"
unfolding simple_function_def using finite_space by auto
lemma (in finite_measure_space) borel_measurable_finite[intro, simp]: "f \<in> borel_measurable M"
by (auto intro: borel_measurable_simple_function)
lemma (in finite_measure_space) positive_integral_finite_eq_setsum:
- "positive_integral f = (\<Sum>x \<in> space M. f x * \<mu> {x})"
+ "integral\<^isup>P M f = (\<Sum>x \<in> space M. f x * \<mu> {x})"
proof -
- have *: "positive_integral f = (\<integral>\<^isup>+ x. \<Sum>y\<in>space M. f y * indicator {y} x)"
+ 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]
by (simp add: positive_integral_setsum positive_integral_cmult_indicator)
qed
lemma (in finite_measure_space) integral_finite_singleton:
- shows "integrable f"
- and "integral f = (\<Sum>x \<in> space M. f x * real (\<mu> {x}))" (is ?I)
+ 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)) = (\<Sum>x \<in> space M. Real (f x) * \<mu> {x})"
- "(\<integral>\<^isup>+ x. Real (- f x)) = (\<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})"
+ "(\<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 f" using finite_space finite_measure
+ show "integrable M f" using finite_space finite_measure
by (simp add: setsum_\<omega> integrable_def)
show ?I using finite_measure
- apply (simp add: integral_def real_of_pextreal_setsum[symmetric]
+ 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)
qed
--- a/src/HOL/Probability/Lebesgue_Measure.thy Wed Feb 02 10:35:41 2011 +0100
+++ b/src/HOL/Probability/Lebesgue_Measure.thy Wed Feb 02 12:34:45 2011 +0100
@@ -1,7 +1,7 @@
(* Author: Robert Himmelmann, TU Muenchen *)
header {* Lebsegue measure *}
theory Lebesgue_Measure
- imports Product_Measure Complete_Measure
+ imports Product_Measure
begin
subsection {* Standard Cubes *}
@@ -42,10 +42,16 @@
by (auto simp add:dist_norm)
qed
-definition lebesgue :: "'a::ordered_euclidean_space algebra" where
- "lebesgue = \<lparr> space = UNIV, sets = {A. \<forall>n. (indicator A :: 'a \<Rightarrow> real) integrable_on cube n} \<rparr>"
+lemma cube_subset_Suc[intro]: "cube n \<subseteq> cube (Suc n)"
+ unfolding cube_def_raw subset_eq apply safe unfolding mem_interval by auto
-definition "lmeasure A = (SUP n. Real (integral (cube n) (indicator A)))"
+lemma Pi_iff: "f \<in> Pi I X \<longleftrightarrow> (\<forall>i\<in>I. f i \<in> X i)"
+ unfolding Pi_def by auto
+
+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>"
lemma space_lebesgue[simp]: "space lebesgue = UNIV"
unfolding lebesgue_def by simp
@@ -106,12 +112,12 @@
qed (auto intro: LIMSEQ_indicator_UN simp: cube_def)
qed simp
-interpretation lebesgue: measure_space lebesgue lmeasure
+interpretation lebesgue: measure_space lebesgue
proof
have *: "indicator {} = (\<lambda>x. 0 :: real)" by (simp add: fun_eq_iff)
- show "lmeasure {} = 0" by (simp add: integral_0 * lmeasure_def)
+ show "measure lebesgue {} = 0" by (simp add: integral_0 * lebesgue_def)
next
- show "countably_additive lebesgue lmeasure"
+ show "countably_additive lebesgue (measure lebesgue)"
proof (intro countably_additive_def[THEN iffD2] allI impI)
fix A :: "nat \<Rightarrow> 'b set" assume rA: "range A \<subseteq> sets lebesgue" "disjoint_family A"
then have A[simp, intro]: "\<And>i n. (indicator (A i) :: _ \<Rightarrow> real) integrable_on cube n"
@@ -122,8 +128,8 @@
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. lmeasure (A n)) = lmeasure (\<Union>i. A i)" unfolding lmeasure_def
- proof (subst psuminf_SUP_eq)
+ 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)
next
@@ -213,20 +219,19 @@
using assms by (force simp: cube_def integrable_on_def negligible_def intro!: lebesgueI)
lemma lmeasure_eq_0:
- fixes S :: "'a::ordered_euclidean_space set" assumes "negligible S" shows "lmeasure S = 0"
+ fixes S :: "'a::ordered_euclidean_space set" assumes "negligible S" shows "lebesgue.\<mu> S = 0"
proof -
have "\<And>n. integral (cube n) (indicator S :: 'a\<Rightarrow>real) = 0"
- unfolding integral_def using assms
- by (intro some1_equality ex_ex1I has_integral_unique)
- (auto simp: cube_def negligible_def intro: )
- then show ?thesis unfolding lmeasure_def by auto
+ unfolding lebesgue_integral_def using assms
+ by (intro integral_unique some1_equality ex_ex1I)
+ (auto simp: cube_def negligible_def)
+ then show ?thesis by (auto simp: lebesgue_def)
qed
lemma lmeasure_iff_LIMSEQ:
assumes "A \<in> sets lebesgue" "0 \<le> m"
- shows "lmeasure A = Real m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m"
- unfolding lmeasure_def
-proof (intro SUP_eq_LIMSEQ)
+ shows "lebesgue.\<mu> A = Real 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)"
@@ -253,7 +258,7 @@
lemma lmeasure_finite_has_integral:
fixes s :: "'a::ordered_euclidean_space set"
- assumes "s \<in> sets lebesgue" "lmeasure s = Real m" "0 \<le> m"
+ assumes "s \<in> sets lebesgue" "lebesgue.\<mu> s = Real m" "0 \<le> m"
shows "(indicator s has_integral m) UNIV"
proof -
let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
@@ -295,9 +300,9 @@
unfolding m by (intro integrable_integral **)
qed
-lemma lmeasure_finite_integrable: assumes "s \<in> sets lebesgue" "lmeasure s \<noteq> \<omega>"
+lemma lmeasure_finite_integrable: assumes "s \<in> sets lebesgue" "lebesgue.\<mu> s \<noteq> \<omega>"
shows "(indicator s :: _ \<Rightarrow> real) integrable_on UNIV"
-proof (cases "lmeasure s")
+proof (cases "lebesgue.\<mu> s")
case (preal m) from lmeasure_finite_has_integral[OF `s \<in> sets lebesgue` this]
show ?thesis unfolding integrable_on_def by auto
qed (insert assms, auto)
@@ -314,7 +319,7 @@
qed
lemma has_integral_lmeasure: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
- shows "lmeasure s = Real m"
+ shows "lebesgue.\<mu> s = Real 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] .
@@ -339,37 +344,37 @@
qed
lemma has_integral_iff_lmeasure:
- "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> lmeasure A = Real m)"
+ "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = Real 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> lmeasure A = Real m"
+ show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = Real m"
by (auto intro: has_integral_nonneg)
next
- assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lmeasure A = Real m"
+ assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = Real 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 "lmeasure s = Real (integral UNIV (indicator s))"
+ shows "lebesgue.\<mu> s = Real (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 "lmeasure s = Real (integral UNIV (indicator s))" by simp
+ show "lebesgue.\<mu> s = Real (integral UNIV (indicator s))" by simp
qed
lemma lebesgue_simple_function_indicator:
fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
- assumes f:"lebesgue.simple_function f"
+ assumes f:"simple_function lebesgue f"
shows "f = (\<lambda>x. (\<Sum>y \<in> f ` UNIV. y * indicator (f -` {y}) x))"
- apply(rule,subst lebesgue.simple_function_indicator_representation[OF f]) by auto
+ by (rule, subst lebesgue.simple_function_indicator_representation[OF f]) auto
lemma integral_eq_lmeasure:
- "(indicator s::_\<Rightarrow>real) integrable_on UNIV \<Longrightarrow> integral UNIV (indicator s) = real (lmeasure s)"
+ "(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 "lmeasure s \<noteq> \<omega>"
+lemma lmeasure_finite: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV" shows "lebesgue.\<mu> s \<noteq> \<omega>"
using lmeasure_eq_integral[OF assms] by auto
lemma negligible_iff_lebesgue_null_sets:
@@ -402,14 +407,13 @@
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]: "lmeasure (UNIV::'a::ordered_euclidean_space set) = \<omega>"
- unfolding lmeasure_def SUP_\<omega>
-proof (intro allI impI)
+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\<in>UNIV. x < Real (integral (cube i) (indicator UNIV::'a\<Rightarrow>real))"
- proof (intro bexI[of _ n])
+ 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)
@@ -428,12 +432,12 @@
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 auto
+ qed
qed
lemma
fixes a b ::"'a::ordered_euclidean_space"
- shows lmeasure_atLeastAtMost[simp]: "lmeasure {a..b} = Real (content {a..b})"
+ shows lmeasure_atLeastAtMost[simp]: "lebesgue.\<mu> {a..b} = Real (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)
@@ -453,7 +457,7 @@
qed
lemma lmeasure_singleton[simp]:
- fixes a :: "'a::ordered_euclidean_space" shows "lmeasure {a} = 0"
+ fixes a :: "'a::ordered_euclidean_space" shows "lebesgue.\<mu> {a} = 0"
using lmeasure_atLeastAtMost[of a a] by simp
declare content_real[simp]
@@ -461,74 +465,97 @@
lemma
fixes a b :: real
shows lmeasure_real_greaterThanAtMost[simp]:
- "lmeasure {a <.. b} = Real (if a \<le> b then b - a else 0)"
+ "lebesgue.\<mu> {a <.. b} = Real (if a \<le> b then b - a else 0)"
proof cases
assume "a < b"
- then have "lmeasure {a <.. b} = lmeasure {a .. b} - lmeasure {a}"
+ then have "lebesgue.\<mu> {a <.. b} = lebesgue.\<mu> {a .. b} - lebesgue.\<mu> {a}"
by (subst lebesgue.measure_Diff[symmetric])
- (auto intro!: arg_cong[where f=lmeasure])
+ (auto intro!: arg_cong[where f=lebesgue.\<mu>])
then show ?thesis by auto
qed auto
lemma
fixes a b :: real
shows lmeasure_real_atLeastLessThan[simp]:
- "lmeasure {a ..< b} = Real (if a \<le> b then b - a else 0)"
+ "lebesgue.\<mu> {a ..< b} = Real (if a \<le> b then b - a else 0)"
proof cases
assume "a < b"
- then have "lmeasure {a ..< b} = lmeasure {a .. b} - lmeasure {b}"
+ then have "lebesgue.\<mu> {a ..< b} = lebesgue.\<mu> {a .. b} - lebesgue.\<mu> {b}"
by (subst lebesgue.measure_Diff[symmetric])
- (auto intro!: arg_cong[where f=lmeasure])
+ (auto intro!: arg_cong[where f=lebesgue.\<mu>])
then show ?thesis by auto
qed auto
lemma
fixes a b :: real
shows lmeasure_real_greaterThanLessThan[simp]:
- "lmeasure {a <..< b} = Real (if a \<le> b then b - a else 0)"
+ "lebesgue.\<mu> {a <..< b} = Real (if a \<le> b then b - a else 0)"
proof cases
assume "a < b"
- then have "lmeasure {a <..< b} = lmeasure {a <.. b} - lmeasure {b}"
+ then have "lebesgue.\<mu> {a <..< b} = lebesgue.\<mu> {a <.. b} - lebesgue.\<mu> {b}"
by (subst lebesgue.measure_Diff[symmetric])
- (auto intro!: arg_cong[where f=lmeasure])
+ (auto intro!: arg_cong[where f=lebesgue.\<mu>])
then show ?thesis by auto
qed auto
-interpretation borel: measure_space borel lmeasure
-proof
- show "countably_additive borel lmeasure"
- using lebesgue.ca unfolding countably_additive_def
- apply safe apply (erule_tac x=A in allE) by auto
-qed auto
+definition "lborel = lebesgue \<lparr> sets := sets borel \<rparr>"
+
+lemma
+ shows space_lborel[simp]: "space lborel = UNIV"
+ and sets_lborel[simp]: "sets lborel = sets borel"
+ and measure_lborel[simp]: "measure lborel = lebesgue.\<mu>"
+ and measurable_lborel[simp]: "measurable lborel = measurable borel"
+ by (simp_all add: measurable_def_raw lborel_def)
-interpretation borel: sigma_finite_measure borel lmeasure
-proof (default, intro conjI exI[of _ "\<lambda>n. cube n"])
- show "range cube \<subseteq> sets borel" 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 borel" by auto
- show "\<forall>i. lmeasure (cube i) \<noteq> \<omega>" unfolding cube_def by auto
-qed
+interpretation lborel: measure_space lborel
+ 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
-interpretation lebesgue: sigma_finite_measure lebesgue lmeasure
+interpretation lborel: sigma_finite_measure lborel
+ where "space lborel = UNIV"
+ and "sets lborel = sets borel"
+ and "measure lborel = lebesgue.\<mu>"
+ and "measurable lborel = measurable borel"
+proof -
+ show "sigma_finite_measure lborel"
+ proof (default, intro conjI exI[of _ "\<lambda>n. cube n"])
+ 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)
+ qed
+qed simp_all
+
+interpretation lebesgue: sigma_finite_measure lebesgue
proof
- from borel.sigma_finite guess A ..
+ 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. lmeasure (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> \<omega>)"
by auto
qed
lemma simple_function_has_integral:
fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
- assumes f:"lebesgue.simple_function f"
+ assumes f:"simple_function lebesgue f"
and f':"\<forall>x. f x \<noteq> \<omega>"
- and om:"\<forall>x\<in>range f. lmeasure (f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
- shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.simple_integral f))) UNIV"
- unfolding lebesgue.simple_integral_def
+ and om:"\<forall>x\<in>range f. lebesgue.\<mu> (f -` {x} \<inter> UNIV) = \<omega> \<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 * lmeasure (f -` {x} \<inter> UNIV) \<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]
@@ -536,11 +563,11 @@
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 * lmeasure (f -` {f y} \<inter> UNIV))) UNIV"
+ 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 lebesgue.simple_function_def using False by auto
+ 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]
@@ -558,31 +585,31 @@
lemma simple_function_has_integral':
fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
- assumes f:"lebesgue.simple_function f"
- and i: "lebesgue.simple_integral f \<noteq> \<omega>"
- shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.simple_integral f))) UNIV"
+ assumes f:"simple_function lebesgue f"
+ and i: "integral\<^isup>S lebesgue f \<noteq> \<omega>"
+ 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 **:"lmeasure {x\<in>space lebesgue. f x \<noteq> ?f x} = 0"
+ 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 "lebesgue.simple_function ?f"
+ 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. lmeasure (?f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
+ 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 "lmeasure ((\<lambda>x. if f x = \<omega> then 0 else f x) -` {if f y = \<omega> then 0 else f y}) = \<omega>"
- ultimately have "lmeasure (f -` {f y}) = \<omega>" by simp
+ 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 * lmeasure (f -` {f y}) \<noteq> \<omega>" using i f
- unfolding lebesgue.simple_integral_def setsum_\<omega> lebesgue.simple_function_def
+ 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
@@ -595,7 +622,7 @@
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. positive_integral (f i)) ----> positive_integral u" (is ?ilim)
+ 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
@@ -609,19 +636,19 @@
lemma positive_integral_has_integral:
fixes f::"'a::ordered_euclidean_space => pextreal"
assumes f:"f \<in> borel_measurable lebesgue"
- and int_om:"lebesgue.positive_integral f \<noteq> \<omega>"
+ and int_om:"integral\<^isup>P lebesgue f \<noteq> \<omega>"
and f_om:"\<forall>x. f x \<noteq> \<omega>" (* TODO: remove this *)
- shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.positive_integral f))) UNIV"
-proof- let ?i = "lebesgue.positive_integral f"
+ 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. lebesgue.simple_integral (u k) = lebesgue.positive_integral (u k)"
+ 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. lebesgue.simple_integral (u k) \<le> lebesgue.positive_integral f"
+ 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. lebesgue.simple_integral (u i) \<noteq> \<omega>"
+ 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
note u_int = simple_function_has_integral'[OF u(1) this]
@@ -633,17 +660,17 @@
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 (lebesgue.positive_integral f)"])
+ 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 "(\<lambda>i. lebesgue.simple_integral (u i)) ----> ?i" unfolding u_simple
+ 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 (lebesgue.simple_integral (u i))) ----> real ?i" apply-
+ 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
@@ -653,12 +680,12 @@
lemma lebesgue_integral_has_integral:
fixes f::"'a::ordered_euclidean_space => real"
- assumes f:"lebesgue.integrable f"
- shows "(f has_integral (lebesgue.integral f)) UNIV"
+ 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 = lebesgue.integrableD[OF f]
- show ?thesis unfolding lebesgue.integral_def apply(subst *)
+ 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)]
@@ -674,27 +701,27 @@
qed
lemma lebesgue_positive_integral_eq_borel:
- "f \<in> borel_measurable borel \<Longrightarrow> lebesgue.positive_integral f = borel.positive_integral f "
+ "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
lemma lebesgue_integral_eq_borel:
assumes "f \<in> borel_measurable borel"
- shows "lebesgue.integrable f = borel.integrable f" (is ?P)
- and "lebesgue.integral f = borel.integral f" (is ?I)
+ shows "integrable lebesgue f \<longleftrightarrow> integrable lborel f" (is ?P)
+ and "integral\<^isup>L lebesgue f = integral\<^isup>L lborel f" (is ?I)
proof -
- have *: "sigma_algebra borel" by default
- have "sets borel \<subseteq> sets lebesgue" by auto
- from lebesgue.integral_subalgebra[OF assms this _ *]
+ have *: "sigma_algebra lborel" by default
+ have "sets lborel \<subseteq> sets lebesgue" by auto
+ from lebesgue.integral_subalgebra[of f lborel, OF _ this _ _ *] assms
show ?P ?I by auto
qed
lemma borel_integral_has_integral:
fixes f::"'a::ordered_euclidean_space => real"
- assumes f:"borel.integrable f"
- shows "(f has_integral (borel.integral f)) UNIV"
+ assumes f:"integrable lborel f"
+ shows "(f has_integral (integral\<^isup>L lborel f)) UNIV"
proof -
have borel: "f \<in> borel_measurable borel"
- using f unfolding borel.integrable_def by auto
+ using f unfolding integrable_def by auto
from f show ?thesis
using lebesgue_integral_has_integral[of f]
unfolding lebesgue_integral_eq_borel[OF borel] by simp
@@ -708,11 +735,11 @@
using continuous_open_preimage[OF assms] unfolding vimage_def by auto
lemma (in measure_space) integral_monotone_convergence_pos':
- assumes i: "\<And>i. integrable (f i)" and mono: "\<And>x. mono (\<lambda>n. f n x)"
+ assumes i: "\<And>i. integrable M (f i)" and mono: "\<And>x. mono (\<lambda>n. f n x)"
and pos: "\<And>x i. 0 \<le> f i x"
and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
- and ilim: "(\<lambda>i. integral (f i)) ----> x"
- shows "integrable u \<and> integral u = x"
+ and ilim: "(\<lambda>i. integral\<^isup>L M (f i)) ----> x"
+ shows "integrable M u \<and> integral\<^isup>L M u = x"
using integral_monotone_convergence_pos[OF assms] by auto
definition e2p :: "'a::ordered_euclidean_space \<Rightarrow> (nat \<Rightarrow> real)" where
@@ -751,53 +778,68 @@
thus "x \<in> e2p ` A" unfolding image_iff apply(rule_tac x="p2e x" in bexI) apply(subst e2p_p2e) by auto
qed
-interpretation borel_product: product_sigma_finite "\<lambda>x. borel::real algebra" "\<lambda>x. lmeasure"
+interpretation lborel_product: product_sigma_finite "\<lambda>x. lborel::real measure_space"
by default
-lemma cube_subset_Suc[intro]: "cube n \<subseteq> cube (Suc n)"
- unfolding cube_def_raw subset_eq apply safe unfolding mem_interval by auto
-
-lemma Pi_iff: "f \<in> Pi I X \<longleftrightarrow> (\<forall>i\<in>I. f i \<in> X i)"
- unfolding Pi_def by auto
+interpretation lborel_space: finite_product_sigma_finite "\<lambda>x. lborel::real measure_space" "{..<DIM('a::ordered_euclidean_space)}"
+ where "space lborel = UNIV"
+ and "sets lborel = sets borel"
+ and "measure lborel = lebesgue.\<mu>"
+ and "measurable lborel = measurable borel"
+proof -
+ show "finite_product_sigma_finite (\<lambda>x. lborel::real measure_space) {..<DIM('a::ordered_euclidean_space)}"
+ by default simp
+qed simp_all
-lemma measurable_e2p_on_generator:
- "e2p \<in> measurable \<lparr> space = UNIV::'a set, sets = range lessThan \<rparr>
- (product_algebra
- (\<lambda>x. \<lparr> space = UNIV::real set, sets = range lessThan \<rparr>)
- {..<DIM('a::ordered_euclidean_space)})"
- (is "e2p \<in> measurable ?E ?P")
-proof (unfold measurable_def, intro CollectI conjI ballI)
- show "e2p \<in> space ?E \<rightarrow> space ?P" by (auto simp: e2p_def)
- fix A assume "A \<in> sets ?P"
- then obtain E where A: "A = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. E i)"
- and "E \<in> {..<DIM('a)} \<rightarrow> (range lessThan)"
- by (auto elim!: product_algebraE)
- then have "\<forall>i\<in>{..<DIM('a)}. \<exists>xs. E i = {..< xs}" by auto
- from this[THEN bchoice] guess xs ..
- then have A_eq: "A = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..< xs i})"
- using A by auto
- have "e2p -` A = {..< (\<chi>\<chi> i. xs i) :: 'a}"
- using DIM_positive by (auto simp add: Pi_iff set_eq_iff e2p_def A_eq
- euclidean_eq[where 'a='a] eucl_less[where 'a='a])
- then show "e2p -` A \<inter> space ?E \<in> sets ?E" by simp
+lemma sets_product_borel:
+ assumes [intro]: "finite I"
+ shows "sets (\<Pi>\<^isub>M i\<in>I.
+ \<lparr> space = UNIV::real set, sets = range lessThan, measure = lebesgue.\<mu> \<rparr>) =
+ sets (\<Pi>\<^isub>M i\<in>I. lborel)" (is "sets ?G = _")
+proof -
+ 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
+ simp: product_algebra_def)
+ then show ?thesis
+ unfolding lborel_def borel_eq_lessThan lebesgue_def sigma_def by simp
qed
lemma measurable_e2p:
- "e2p \<in> measurable (borel::'a algebra)
- (sigma (product_algebra (\<lambda>x. borel :: real algebra) {..<DIM('a::ordered_euclidean_space)}))"
- using measurable_e2p_on_generator[where 'a='a] unfolding borel_eq_lessThan
- by (subst sigma_product_algebra_sigma_eq[where S="\<lambda>_ i. {..<real i}"])
- (auto intro!: measurable_sigma_sigma isotoneI real_arch_lt
- simp: product_algebra_def)
+ "e2p \<in> measurable (borel::'a::ordered_euclidean_space algebra)
+ (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure_space))"
+ (is "_ \<in> measurable ?E ?P")
+proof -
+ let ?B = "\<lparr> space = UNIV::real set, sets = range lessThan, measure = lebesgue.\<mu> \<rparr>"
+ let ?G = "product_algebra_generator {..<DIM('a)} (\<lambda>_. ?B)"
+ have "e2p \<in> measurable ?E (sigma ?G)"
+ proof (rule borel.measurable_sigma)
+ show "e2p \<in> space ?E \<rightarrow> space ?G" by (auto simp: e2p_def)
+ fix A assume "A \<in> sets ?G"
+ then obtain E where A: "A = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. E i)"
+ and "E \<in> {..<DIM('a)} \<rightarrow> (range lessThan)"
+ by (auto elim!: product_algebraE simp: )
+ then have "\<forall>i\<in>{..<DIM('a)}. \<exists>xs. E i = {..< xs}" by auto
+ from this[THEN bchoice] guess xs ..
+ then have A_eq: "A = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..< xs i})"
+ using A by auto
+ have "e2p -` A = {..< (\<chi>\<chi> i. xs i) :: 'a}"
+ using DIM_positive by (auto simp add: Pi_iff set_eq_iff e2p_def A_eq
+ euclidean_eq[where 'a='a] eucl_less[where 'a='a])
+ then show "e2p -` A \<inter> space ?E \<in> sets ?E" by simp
+ qed (auto simp: product_algebra_generator_def)
+ with sets_product_borel[of "{..<DIM('a)}"] show ?thesis
+ unfolding measurable_def product_algebra_def by simp
+qed
-lemma measurable_p2e_on_generator:
- "p2e \<in> measurable
- (product_algebra
- (\<lambda>x. \<lparr> space = UNIV::real set, sets = range lessThan \<rparr>)
- {..<DIM('a::ordered_euclidean_space)})
- \<lparr> space = UNIV::'a set, sets = range lessThan \<rparr>"
- (is "p2e \<in> measurable ?P ?E")
-proof (unfold measurable_def, intro CollectI conjI ballI)
+lemma measurable_p2e:
+ "p2e \<in> measurable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure_space))
+ (borel :: 'a::ordered_euclidean_space algebra)"
+ (is "p2e \<in> measurable ?P _")
+ unfolding borel_eq_lessThan
+proof (intro lborel_space.measurable_sigma)
+ let ?E = "\<lparr> space = UNIV :: 'a set, sets = range lessThan \<rparr>"
show "p2e \<in> space ?P \<rightarrow> space ?E" by simp
fix A assume "A \<in> sets ?E"
then obtain x where "A = {..<x}" by auto
@@ -806,15 +848,7 @@
by (auto simp: Pi_iff set_eq_iff p2e_def
euclidean_eq[where 'a='a] eucl_less[where 'a='a])
then show "p2e -` A \<inter> space ?P \<in> sets ?P" by auto
-qed
-
-lemma measurable_p2e:
- "p2e \<in> measurable (sigma (product_algebra (\<lambda>x. borel :: real algebra) {..<DIM('a::ordered_euclidean_space)}))
- (borel::'a algebra)"
- using measurable_p2e_on_generator[where 'a='a] unfolding borel_eq_lessThan
- by (subst sigma_product_algebra_sigma_eq[where S="\<lambda>_ i. {..<real i}"])
- (auto intro!: measurable_sigma_sigma isotoneI real_arch_lt
- simp: product_algebra_def)
+qed simp
lemma e2p_Int:"e2p ` A \<inter> e2p ` B = e2p ` (A \<inter> B)" (is "?L = ?R")
apply(rule image_Int[THEN sym])
@@ -840,15 +874,13 @@
lemma lmeasure_measure_eq_borel_prod:
fixes A :: "('a::ordered_euclidean_space) set"
assumes "A \<in> sets borel"
- shows "lmeasure A = borel_product.product_measure {..<DIM('a)} (e2p ` A :: (nat \<Rightarrow> real) set)"
+ shows "lebesgue.\<mu> A = lborel_space.\<mu> TYPE('a) (e2p ` A)" (is "_ = ?m A")
proof (rule measure_unique_Int_stable[where X=A and A=cube])
- interpret fprod: finite_product_sigma_finite "\<lambda>x. borel :: real algebra" "\<lambda>x. lmeasure" "{..<DIM('a)}" by default auto
show "Int_stable \<lparr> space = UNIV :: 'a set, sets = range (\<lambda>(a,b). {a..b}) \<rparr>"
(is "Int_stable ?E" ) using Int_stable_cuboids' .
- show "borel = sigma ?E" using borel_eq_atLeastAtMost .
- show "\<And>i. lmeasure (cube i) \<noteq> \<omega>" unfolding cube_def by auto
- show "\<And>X. X \<in> sets ?E \<Longrightarrow>
- lmeasure X = borel_product.product_measure {..<DIM('a)} (e2p ` X :: (nat \<Rightarrow> real) set)"
+ have [simp]: "sigma ?E = borel" using borel_eq_atLeastAtMost ..
+ show "\<And>i. lebesgue.\<mu> (cube i) \<noteq> \<omega>" unfolding cube_def by auto
+ show "\<And>X. X \<in> sets ?E \<Longrightarrow> lebesgue.\<mu> X = ?m X"
proof- case goal1 then obtain a b where X:"X = {a..b}" by auto
{ presume *:"X \<noteq> {} \<Longrightarrow> ?case"
show ?case apply(cases,rule *,assumption) by auto }
@@ -861,12 +893,12 @@
show "x \<in> Pi\<^isub>E {..<DIM('a)} XX" unfolding y using y(1)
unfolding Pi_def extensional_def e2p_def restrict_def X mem_interval XX_def by auto
qed
- have "lmeasure X = (\<Prod>x<DIM('a). Real (b $$ x - a $$ x))" using X' apply- unfolding X
+ have "lebesgue.\<mu> X = (\<Prod>x<DIM('a). Real (b $$ x - a $$ x))" using X' apply- unfolding X
unfolding lmeasure_atLeastAtMost content_closed_interval apply(subst Real_setprod) by auto
- also have "... = (\<Prod>i<DIM('a). lmeasure (XX i))" apply(rule setprod_cong2)
+ also have "... = (\<Prod>i<DIM('a). lebesgue.\<mu> (XX i))" apply(rule setprod_cong2)
unfolding XX_def lmeasure_atLeastAtMost apply(subst content_real) using X' by auto
- also have "... = borel_product.product_measure {..<DIM('a)} (e2p ` X)" unfolding *[THEN sym]
- apply(rule fprod.measure_times[THEN sym]) unfolding XX_def by auto
+ also have "... = ?m X" unfolding *[THEN sym]
+ apply(rule lborel_space.measure_times[symmetric]) unfolding XX_def by auto
finally show ?case .
qed
@@ -875,18 +907,21 @@
have "\<And>x. \<exists>xa. x \<in> cube xa" apply(rule_tac x=x in mem_big_cube) by fastsimp
thus "cube \<up> space \<lparr>space = UNIV, sets = range (\<lambda>(a, b). {a..b})\<rparr>"
apply-apply(rule isotoneI) apply(rule cube_subset_Suc) by auto
- show "A \<in> sets borel " by fact
- show "measure_space borel lmeasure" by default
- show "measure_space borel
- (\<lambda>a::'a set. finite_product_sigma_finite.measure (\<lambda>x. borel) (\<lambda>x. lmeasure) {..<DIM('a)} (e2p ` a))"
- proof (rule fprod.measure_space_vimage)
- show "sigma_algebra borel" by default
- show "(p2e :: (nat \<Rightarrow> real) \<Rightarrow> 'a) \<in> measurable fprod.P borel" by (rule measurable_p2e)
- fix A :: "'a set" assume "A \<in> sets borel"
- show "fprod.measure (e2p ` A) = fprod.measure (p2e -` A \<inter> space fprod.P)"
+ show "A \<in> sets (sigma ?E)" using assms by simp
+ have "measure_space lborel" by default
+ then show "measure_space \<lparr> space = space ?E, sets = sets (sigma ?E), measure = measure lebesgue\<rparr>"
+ unfolding lebesgue_def lborel_def by simp
+ let ?M = "\<lparr> space = space ?E, sets = sets (sigma ?E), measure = ?m \<rparr>"
+ show "measure_space ?M"
+ proof (rule lborel_space.measure_space_vimage)
+ show "sigma_algebra ?M" by (rule lborel.sigma_algebra_cong) auto
+ show "p2e \<in> measurable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel) ?M"
+ using measurable_p2e unfolding measurable_def by auto
+ fix A :: "'a set" assume "A \<in> sets ?M"
+ show "measure ?M A = lborel_space.\<mu> TYPE('a) (p2e -` A \<inter> space (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel))"
by (simp add: e2p_image_vimage)
qed
-qed
+qed simp
lemma range_e2p:"range (e2p::'a::ordered_euclidean_space \<Rightarrow> _) = extensional {..<DIM('a)}"
unfolding e2p_def_raw
@@ -896,41 +931,30 @@
lemma borel_fubini_positiv_integral:
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> pextreal"
assumes f: "f \<in> borel_measurable borel"
- shows "borel.positive_integral f =
- borel_product.product_positive_integral {..<DIM('a)} (f \<circ> p2e)"
-proof- def U \<equiv> "extensional {..<DIM('a)} :: (nat \<Rightarrow> real) set"
- interpret fprod: finite_product_sigma_finite "\<lambda>x. borel" "\<lambda>x. lmeasure" "{..<DIM('a)}" by default auto
- show ?thesis
- proof (subst borel.positive_integral_vimage[symmetric, of _ "e2p :: 'a \<Rightarrow> _" "(\<lambda>x. f (p2e x))", unfolded p2e_e2p])
- show "(e2p :: 'a \<Rightarrow> _) \<in> measurable borel fprod.P" by (rule measurable_e2p)
- show "sigma_algebra fprod.P" by default
- from measurable_comp[OF measurable_p2e f]
- show "(\<lambda>x. f (p2e x)) \<in> borel_measurable fprod.P" by (simp add: comp_def)
- let "?L A" = "lmeasure ((e2p::'a \<Rightarrow> (nat \<Rightarrow> real)) -` A \<inter> space borel)"
- show "measure_space.positive_integral fprod.P ?L (\<lambda>x. f (p2e x)) =
- fprod.positive_integral (f \<circ> p2e)"
- unfolding comp_def
- proof (rule fprod.positive_integral_cong_measure)
- fix A :: "(nat \<Rightarrow> real) set" assume "A \<in> sets fprod.P"
- then have A: "(e2p::'a \<Rightarrow> (nat \<Rightarrow> real)) -` A \<inter> space borel \<in> sets borel"
- by (rule measurable_sets[OF measurable_e2p])
- have [simp]: "A \<inter> extensional {..<DIM('a)} = A"
- using `A \<in> sets fprod.P`[THEN fprod.sets_into_space] by auto
- show "?L A = fprod.measure A"
- unfolding lmeasure_measure_eq_borel_prod[OF A]
- by (simp add: range_e2p)
- qed
- qed
+ shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(lborel_space.P TYPE('a))"
+proof (rule lborel.positive_integral_vimage[symmetric, of _ "e2p :: 'a \<Rightarrow> _" "(\<lambda>x. f (p2e x))", unfolded p2e_e2p])
+ show "(e2p :: 'a \<Rightarrow> _) \<in> measurable borel (lborel_space.P TYPE('a))" by (rule measurable_e2p)
+ show "sigma_algebra (lborel_space.P TYPE('a))" by default
+ from measurable_comp[OF measurable_p2e f]
+ show "(\<lambda>x. f (p2e x)) \<in> borel_measurable (lborel_space.P TYPE('a))" by (simp add: comp_def)
+ let "?L A" = "lebesgue.\<mu> ((e2p::'a \<Rightarrow> (nat \<Rightarrow> real)) -` A \<inter> UNIV)"
+ fix A :: "(nat \<Rightarrow> real) set" assume "A \<in> sets (lborel_space.P TYPE('a))"
+ then have A: "(e2p::'a \<Rightarrow> (nat \<Rightarrow> real)) -` A \<inter> space borel \<in> sets borel"
+ by (rule measurable_sets[OF measurable_e2p])
+ have [simp]: "A \<inter> extensional {..<DIM('a)} = A"
+ using `A \<in> sets (lborel_space.P TYPE('a))`[THEN lborel_space.sets_into_space] by auto
+ show "lborel_space.\<mu> TYPE('a) A = ?L A"
+ using lmeasure_measure_eq_borel_prod[OF A] by (simp add: range_e2p)
qed
lemma borel_fubini:
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
assumes f: "f \<in> borel_measurable borel"
- shows "borel.integral f = borel_product.product_integral {..<DIM('a)} (f \<circ> p2e)"
-proof- interpret fprod: finite_product_sigma_finite "\<lambda>x. borel" "\<lambda>x. lmeasure" "{..<DIM('a)}" by default auto
+ shows "integral\<^isup>L lborel f = \<integral>x. f (p2e x) \<partial>(lborel_space.P TYPE('a))"
+proof -
have 1:"(\<lambda>x. Real (f x)) \<in> borel_measurable borel" using f by auto
have 2:"(\<lambda>x. Real (- f x)) \<in> borel_measurable borel" using f by auto
- show ?thesis unfolding fprod.integral_def borel.integral_def
+ show ?thesis unfolding lebesgue_integral_def
unfolding borel_fubini_positiv_integral[OF 1] borel_fubini_positiv_integral[OF 2]
unfolding o_def ..
qed
--- a/src/HOL/Probability/Measure.thy Wed Feb 02 10:35:41 2011 +0100
+++ b/src/HOL/Probability/Measure.thy Wed Feb 02 12:34:45 2011 +0100
@@ -4,6 +4,18 @@
imports Caratheodory
begin
+lemma measure_algebra_more[simp]:
+ "\<lparr> space = A, sets = B, \<dots> = algebra.more M \<rparr> \<lparr> measure := m \<rparr> =
+ \<lparr> space = A, sets = B, \<dots> = algebra.more (M \<lparr> measure := m \<rparr>) \<rparr>"
+ by (cases M) simp
+
+lemma measure_algebra_more_eq[simp]:
+ "\<And>X. measure \<lparr> space = T, sets = A, \<dots> = algebra.more X \<rparr> = measure X"
+ unfolding measure_space.splits by simp
+
+lemma measure_sigma[simp]: "measure (sigma A) = measure A"
+ unfolding sigma_def by simp
+
lemma inj_on_image_eq_iff:
assumes "inj_on f S"
assumes "A \<subseteq> S" "B \<subseteq> S"
@@ -53,24 +65,34 @@
with ca assms show ?thesis by (simp add: countably_additive_def)
qed
+lemma (in sigma_algebra) sigma_algebra_cong:
+ assumes "space N = space M" "sets N = sets M"
+ shows "sigma_algebra N"
+ by default (insert sets_into_space, auto simp: assms)
+
lemma (in measure_space) measure_space_cong:
- assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = \<mu> A"
- shows "measure_space M \<nu>"
-proof
- show "\<nu> {} = 0" using assms by auto
- show "countably_additive M \<nu>" unfolding countably_additive_def
- proof safe
- fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "disjoint_family A"
- then have "\<And>i. A i \<in> sets M" "(UNION UNIV A) \<in> sets M" by auto
- from this[THEN assms] measure_countably_additive[OF A]
- show "(\<Sum>\<^isub>\<infinity>n. \<nu> (A n)) = \<nu> (UNION UNIV A)" by simp
+ assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "space N = space M" "sets N = sets M"
+ shows "measure_space N"
+proof -
+ interpret N: sigma_algebra N by (intro sigma_algebra_cong assms)
+ show ?thesis
+ proof
+ show "measure N {} = 0" using assms by auto
+ 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)"
+ unfolding assms by simp
+ qed
qed
qed
lemma (in measure_space) additive: "additive M \<mu>"
proof (rule algebra.countably_additive_additive [OF _ _ ca])
show "algebra M" by default
- show "positive \<mu>" by (simp add: positive_def)
+ show "positive M \<mu>" by (simp add: positive_def)
qed
lemma (in measure_space) measure_additive:
@@ -358,12 +380,16 @@
finally show ?thesis by simp
qed
-lemma (in sigma_algebra) finite_additivity_sufficient:
- assumes fin: "finite (space M)" and pos: "positive \<mu>" and add: "additive M \<mu>"
- shows "measure_space M \<mu>"
-proof
- show [simp]: "\<mu> {} = 0" using pos by (simp add: positive_def)
- show "countably_additive M \<mu>"
+lemma finite_additivity_sufficient:
+ assumes "sigma_algebra M"
+ assumes fin: "finite (space M)" and pos: "positive M (measure M)" and add: "additive M (measure M)"
+ shows "measure_space M"
+proof -
+ interpret sigma_algebra M by fact
+ show ?thesis
+ proof
+ show [simp]: "measure M {} = 0" 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"
assume A: "range A \<subseteq> sets M"
@@ -391,15 +417,15 @@
by blast
qed
then obtain N where N: "\<forall>m\<ge>N. A m = {}" by blast
- then have "\<forall>m\<ge>N. \<mu> (A m) = 0" by simp
- then have "(\<Sum>\<^isub>\<infinity> n. \<mu> (A n)) = setsum (\<lambda>m. \<mu> (A m)) {..<N}"
+ 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)
- also have "... = \<mu> (\<Union>i<N. A i)"
+ also have "... = measure M (\<Union>i<N. A i)"
proof (induct N)
case 0 thus ?case by simp
next
case (Suc n)
- have "\<mu> (A n \<union> (\<Union> x<n. A x)) = \<mu> (A n) + \<mu> (\<Union> i<n. A i)"
+ have "measure M (A n \<union> (\<Union> x<n. A x)) = measure M (A n) + measure M (\<Union> i<n. A i)"
proof (rule Caratheodory.additiveD [OF add])
show "A n \<inter> (\<Union> x<n. A x) = {}" using disj
by (auto simp add: disjoint_family_on_def nat_less_le) blast
@@ -416,14 +442,15 @@
thus ?case using Suc
by (simp add: lessThan_Suc)
qed
- also have "... = \<mu> (\<Union>i. A i)"
+ also have "... = measure M (\<Union>i. A i)"
proof -
have "(\<Union> i<N. A i) = (\<Union>i. A i)" using N
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. \<mu> (A n)) = \<mu> (\<Union>i. A i)" .
+ finally show "(\<Sum>\<^isub>\<infinity> n. measure M (A n)) = measure M (\<Union>i. A i)" .
qed
+ qed
qed
lemma (in measure_space) measure_setsum_split:
@@ -525,95 +552,75 @@
qed
qed
-lemma True
-proof
- fix x a b :: nat
- have "\<And>x a b :: int. x dvd a \<Longrightarrow> x dvd (a + b) \<Longrightarrow> x dvd b"
- by (metis dvd_mult_div_cancel zadd_commute zdvd_reduce)
- then have "x dvd a \<Longrightarrow> x dvd (a + b) \<Longrightarrow> x dvd b"
- unfolding zdvd_int[of x] zadd_int[symmetric] .
-qed
-
lemma measure_unique_Int_stable:
- fixes M E :: "'a algebra" and A :: "nat \<Rightarrow> 'a set"
- assumes "Int_stable E" "M = sigma E"
- and A: "range A \<subseteq> sets E" "A \<up> space E"
- and ms: "measure_space M \<mu>" "measure_space M \<nu>"
+ 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 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>"
- assumes "X \<in> sets M"
+ assumes "X \<in> sets (sigma E)"
shows "\<mu> X = \<nu> X"
proof -
- let "?D F" = "{D. D \<in> sets M \<and> \<mu> (F \<inter> D) = \<nu> (F \<inter> D)}"
- interpret M: measure_space M \<mu> by fact
- interpret M': measure_space M \<nu> by fact
- have "space E = space M"
- using `M = sigma E` by simp
- have sets_E: "sets E \<subseteq> Pow (space E)"
- proof
- fix X assume "X \<in> sets E"
- then have "X \<in> sets M" unfolding `M = sigma E`
- unfolding sigma_def by (auto intro!: sigma_sets.Basic)
- with M.sets_into_space show "X \<in> Pow (space E)"
- unfolding `space E = space M` by auto
- qed
- have A': "range A \<subseteq> sets M" using `M = sigma E` A
- by (auto simp: sets_sigma intro!: sigma_sets.Basic)
+ let "?D F" = "{D. D \<in> sets (sigma E) \<and> \<mu> (F \<inter> D) = \<nu> (F \<inter> D)}"
+ interpret M: measure_space ?M
+ 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>"
- then have [intro]: "F \<in> sets M" unfolding `M = sigma E` sets_sigma
- by (intro sigma_sets.Basic)
+ 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
interpret D: dynkin_system "\<lparr>space=space E, sets=?D F\<rparr>"
proof (rule dynkin_systemI, simp_all)
- fix A assume "A \<in> sets M \<and> \<mu> (F \<inter> A) = \<nu> (F \<inter> A)"
- then show "A \<subseteq> space E"
- unfolding `space E = space M` using M.sets_into_space by auto
+ fix A assume "A \<in> sets (sigma E) \<and> \<mu> (F \<inter> A) = \<nu> (F \<inter> A)"
+ then show "A \<subseteq> space E" using M.sets_into_space by auto
next
- have "F \<inter> space E = F" using `F \<in> sets E` sets_E by auto
- then show "space E \<in> sets M \<and> \<mu> (F \<inter> space E) = \<nu> (F \<inter> space E)"
- unfolding `space E = space M` using `F \<in> sets E` eq by auto
+ have "F \<inter> space E = F" using `F \<in> sets E` by auto
+ then show "\<mu> (F \<inter> space E) = \<nu> (F \<inter> space E)"
+ using `F \<in> sets E` eq by auto
next
- fix A assume *: "A \<in> sets M \<and> \<mu> (F \<inter> A) = \<nu> (F \<inter> A)"
- then have **: "F \<inter> (space M - A) = F - (F \<inter> A)"
- and [intro]: "F \<inter> A \<in> sets M"
- using `F \<in> sets E` sets_E `space E = space M` by auto
- have "\<nu> (F \<inter> A) \<le> \<nu> F" by (auto intro!: M'.measure_mono)
+ fix A assume *: "A \<in> sets (sigma E) \<and> \<mu> (F \<inter> A) = \<nu> (F \<inter> A)"
+ then have **: "F \<inter> (space (sigma E) - A) = F - (F \<inter> A)"
+ 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
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> (space M - A)) = \<mu> F - \<mu> (F \<inter> A)" unfolding **
- using `F \<inter> A \<in> sets M` by (auto intro!: M.measure_Diff)
+ 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 M - A))" unfolding **
- using `F \<inter> A \<in> sets M` `\<nu> (F \<inter> A) \<noteq> \<omega>` by (auto intro!: M'.measure_Diff[symmetric])
- finally show "space E - A \<in> sets M \<and> \<mu> (F \<inter> (space E - A)) = \<nu> (F \<inter> (space E - A))"
- using `space E = space M` * by auto
+ 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])
+ 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
fix A :: "nat \<Rightarrow> 'a set"
- assume "disjoint_family A" "range A \<subseteq> {X \<in> sets M. \<mu> (F \<inter> X) = \<nu> (F \<inter> X)}"
- then have A: "range (\<lambda>i. F \<inter> A i) \<subseteq> sets M" "F \<inter> (\<Union>x. A x) = (\<Union>x. F \<inter> A x)"
- "disjoint_family (\<lambda>i. F \<inter> A i)" "\<And>i. \<mu> (F \<inter> A i) = \<nu> (F \<inter> A i)" "range A \<subseteq> sets M"
- by ((fastsimp simp: disjoint_family_on_def)+)
- then show "(\<Union>x. A x) \<in> sets M \<and> \<mu> (F \<inter> (\<Union>x. A x)) = \<nu> (F \<inter> (\<Union>x. A x))"
+ assume "disjoint_family A" "range A \<subseteq> {X \<in> sets (sigma E). \<mu> (F \<inter> X) = \<nu> (F \<inter> X)}"
+ then have A: "range (\<lambda>i. F \<inter> A i) \<subseteq> sets (sigma E)" "F \<inter> (\<Union>x. A x) = (\<Union>x. F \<inter> A x)"
+ "disjoint_family (\<lambda>i. F \<inter> A i)" "\<And>i. \<mu> (F \<inter> A i) = \<nu> (F \<inter> A i)" "range A \<subseteq> sets (sigma E)"
+ by (auto simp: disjoint_family_on_def subset_eq)
+ then show "(\<Union>x. A x) \<in> sets (sigma E) \<and> \<mu> (F \<inter> (\<Union>x. A x)) = \<nu> (F \<inter> (\<Union>x. A x))"
by (auto simp: M.measure_countably_additive[symmetric]
- M'.measure_countably_additive[symmetric]
+ N.measure_countably_additive[symmetric]
simp del: UN_simps)
qed
- have *: "sigma E = \<lparr>space = space E, sets = ?D F\<rparr>"
- using `M = sigma E` `F \<in> sets E` `Int_stable E`
+ have *: "sets (sigma E) = sets \<lparr>space = space E, sets = ?D F\<rparr>"
+ using `F \<in> sets E` `Int_stable E`
by (intro D.dynkin_lemma)
(auto simp add: sets_sigma Int_stable_def eq intro: sigma_sets.Basic)
- have "\<And>D. D \<in> sets M \<Longrightarrow> \<mu> (F \<inter> D) = \<nu> (F \<inter> D)"
- unfolding `M = sigma E` by (auto simp: *) }
+ 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)"
- using *[of "A i" X] `X \<in> sets M` A finite by auto }
+ 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 M` M.sets_into_space A `space E = space M`
+ 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 M` A' by (auto intro!: M.measure_up M'.measure_up M.Int)
+ 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)
qed
@@ -830,37 +837,38 @@
lemma (in measure_space) restricted_measure_space:
assumes "S \<in> sets M"
- shows "measure_space (restricted_space S) \<mu>"
- (is "measure_space ?r \<mu>")
+ shows "measure_space (restricted_space S)"
+ (is "measure_space ?r")
unfolding measure_space_def measure_space_axioms_def
proof safe
show "sigma_algebra ?r" using restricted_sigma_algebra[OF assms] .
- show "\<mu> {} = 0" by simp
- show "countably_additive ?r \<mu>"
+ show "measure ?r {} = 0" by simp
+
+ show "countably_additive ?r (measure ?r)"
unfolding countably_additive_def
proof safe
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. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
+ show "(\<Sum>\<^isub>\<infinity> n. measure ?r (A n)) = measure ?r (\<Union>i. A i)"
using measure_countably_additive by simp
qed
qed
lemma (in measure_space) measure_space_vimage:
- fixes M' :: "'b algebra"
+ fixes M' :: "('c, 'd) measure_space_scheme"
assumes T: "sigma_algebra M'" "T \<in> measurable M M'"
- and \<nu>: "\<And>A. A \<in> sets M' \<Longrightarrow> \<nu> A = \<mu> (T -` A \<inter> space M)"
- shows "measure_space M' \<nu>"
+ and \<nu>: "\<And>A. A \<in> sets M' \<Longrightarrow> measure M' A = \<mu> (T -` A \<inter> space M)"
+ shows "measure_space M'"
proof -
interpret M': sigma_algebra M' by fact
show ?thesis
proof
- show "\<nu> {} = 0" using \<nu>[of "{}"] by simp
+ show "measure M' {} = 0" using \<nu>[of "{}"] by simp
- show "countably_additive M' \<nu>"
- proof (intro countably_additive_def[THEN iffD2] allI impI)
- fix A :: "nat \<Rightarrow> 'b set" assume "range A \<subseteq> sets M'" "disjoint_family A"
+ show "countably_additive M' (measure M')"
+ proof (intro countably_additiveI)
+ fix A :: "nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets M'" "disjoint_family A"
then have A: "\<And>i. A i \<in> sets M'" "(\<Union>i. A i) \<in> sets M'" by auto
then have *: "range (\<lambda>i. T -` (A i) \<inter> space M) \<subseteq> sets M"
using `T \<in> measurable M M'` by (auto simp: measurable_def)
@@ -868,7 +876,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. \<nu> (A i)) = \<nu> (\<Union>i. A i)"
+ ultimately show "(\<Sum>\<^isub>\<infinity> i. measure M' (A i)) = measure M' (\<Union>i. A i)"
using measure_countably_additive[OF _ **] A
by (auto simp: comp_def vimage_UN \<nu>)
qed
@@ -877,14 +885,15 @@
lemma (in measure_space) measure_space_subalgebra:
assumes "sigma_algebra N" and [simp]: "sets N \<subseteq> sets M" "space N = space M"
- shows "measure_space N \<mu>"
+ and measure[simp]: "\<And>X. X \<in> sets N \<Longrightarrow> measure N X = measure M X"
+ shows "measure_space N"
proof -
interpret N: sigma_algebra N by fact
show ?thesis
proof
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 \<mu>"
- by (auto intro!: measure_countably_additive simp: countably_additive_def)
+ then show "countably_additive N (measure N)"
+ by (auto intro!: measure_countably_additive simp: countably_additive_def subset_eq)
qed simp
qed
@@ -895,16 +904,16 @@
lemma (in sigma_finite_measure) restricted_sigma_finite_measure:
assumes "S \<in> sets M"
- shows "sigma_finite_measure (restricted_space S) \<mu>"
- (is "sigma_finite_measure ?r _")
+ shows "sigma_finite_measure (restricted_space S)"
+ (is "sigma_finite_measure ?r")
unfolding sigma_finite_measure_def sigma_finite_measure_axioms_def
proof safe
- show "measure_space ?r \<mu>" using restricted_measure_space[OF assms] .
+ 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>"
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. \<mu> (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> \<omega>)"
proof (safe intro!: exI[of _ "\<lambda>i. A i \<inter> S"] del: notI)
fix i
show "A i \<inter> S \<in> sets ?r"
@@ -919,22 +928,21 @@
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 "\<mu> (A i \<inter> S) \<noteq> \<omega>" by (auto simp: pextreal_less_\<omega>)
+ finally show "measure ?r (A i \<inter> S) \<noteq> \<omega>" by (auto simp: pextreal_less_\<omega>)
qed
qed
lemma (in sigma_finite_measure) sigma_finite_measure_cong:
- assumes cong: "\<And>A. A \<in> sets M \<Longrightarrow> \<mu>' A = \<mu> A"
- shows "sigma_finite_measure M \<mu>'"
+ assumes cong: "\<And>A. A \<in> sets M \<Longrightarrow> measure M' A = \<mu> A" "sets M' = sets M" "space M' = space M"
+ shows "sigma_finite_measure M'"
proof -
- interpret \<mu>': measure_space M \<mu>'
- using cong by (rule measure_space_cong)
+ 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. \<mu>' (A i) \<noteq> \<omega>)" using cong by auto
+ with A have fin: "(\<forall>i. measure M' (A i) \<noteq> \<omega>)" using cong by auto
show ?thesis
apply default
- using A fin by auto
+ using A fin cong by auto
qed
lemma (in sigma_finite_measure) disjoint_sigma_finite:
@@ -1110,20 +1118,20 @@
lemma (in finite_measure) restricted_finite_measure:
assumes "S \<in> sets M"
- shows "finite_measure (restricted_space S) \<mu>"
- (is "finite_measure ?r _")
+ shows "finite_measure (restricted_space S)"
+ (is "finite_measure ?r")
unfolding finite_measure_def finite_measure_axioms_def
proof (safe del: notI)
- show "measure_space ?r \<mu>" using restricted_measure_space[OF assms] .
+ show "measure_space ?r" using restricted_measure_space[OF assms] .
next
- show "\<mu> (space ?r) \<noteq> \<omega>" using finite_measure[OF `S \<in> sets M`] by auto
+ show "measure ?r (space ?r) \<noteq> \<omega>" 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>"
- shows "finite_measure (restricted_space S) \<mu>"
+ shows "finite_measure (restricted_space S)"
proof -
- have "measure_space (restricted_space S) \<mu>"
+ have "measure_space (restricted_space S)"
using `S \<in> sets M` by (rule restricted_measure_space)
then show ?thesis
unfolding finite_measure_def finite_measure_axioms_def
@@ -1218,105 +1226,104 @@
section {* Measure preserving *}
-definition "measure_preserving A \<mu> B \<nu> =
- {f \<in> measurable A B. (\<forall>y \<in> sets B. \<mu> (f -` y \<inter> space A) = \<nu> y)}"
+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)}"
lemma (in finite_measure) measure_preserving_lift:
- fixes f :: "'a \<Rightarrow> 'a2" and A :: "'a2 algebra"
- assumes "algebra A"
- assumes "finite_measure (sigma A) \<nu>" (is "finite_measure ?sA _")
- assumes fmp: "f \<in> measure_preserving M \<mu> A \<nu>"
- shows "f \<in> measure_preserving M \<mu> (sigma A) \<nu>"
+ 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)"
proof -
- interpret sA: finite_measure ?sA \<nu> by fact
+ 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 meq: "\<forall>y\<in>sets A. \<mu> (f -` y \<inter> space M) = \<nu> y"
- 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) = \<nu> s}"
- proof (rule A.sigma_property_disjoint, auto)
- fix x assume "x \<in> sets A" then show "\<mu> (f -` x \<inter> space M) = \<nu> x" by (simp add: meq)
- next
- fix s
- assume eq: "\<mu> (f -` s \<inter> space M) = \<nu> 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) = \<nu> (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) = \<nu> s}"
- and "range S \<subseteq> ?A"
- hence rS2: "range S \<subseteq> sets ?sA" by (simp add: sigma_def)
- have eq1: "\<And>i. \<mu> (f -` S i \<inter> space M) = \<nu> (S i)"
- using rS1 by blast
- have *: "(\<lambda>n. \<nu> (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. \<nu> (S n)) = \<mu> (f -` (\<Union>i. S i) \<inter> space M)" unfolding * .
- moreover
- have "(SUP n. \<nu> (S n)) = \<nu> (\<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) = \<nu> (\<Union>i. S i)" by simp
- next
- fix S :: "nat => 'a2 set"
- assume dS: "disjoint_family S"
- and rS1: "range S \<subseteq> {s. \<mu> (f -` s \<inter> space M) = \<nu> s}"
- and "range S \<subseteq> ?A"
- hence rS2: "range S \<subseteq> sets ?sA" by (simp add: sigma_def)
- have "\<And>i. \<mu> (f -` S i \<inter> space M) = \<nu> (S i)"
- using rS1 by blast
- hence *: "(\<lambda>i. \<nu> (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. \<nu> (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) = \<nu> (\<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) = \<nu> (\<Union>i. S i)"
- by metis
+ 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
- finally have "sets ?sA \<subseteq> {s . \<mu> ((f -` s) \<inter> space M) = \<nu> s}" .
- hence "\<mu> (f -` y \<inter> space M) = \<nu> y" using y by force
- }
- thus "f \<in> measurable M ?sA \<and> (\<forall>y\<in>sets ?sA. \<mu> (f -` y \<inter> space M) = \<nu> y)"
- by (blast intro: f12)
- 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
qed
section "Finite spaces"
@@ -1329,22 +1336,24 @@
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: "\<mu> (space M) \<noteq> \<omega>"
- and add: "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> \<mu> (A \<union> B) = \<mu> A + \<mu> B"
- and "\<mu> {} = 0"
- shows "finite_measure_space M \<mu>"
+ assumes "finite (space M)" "sets M = Pow(space M)" and space: "measure M (space M) \<noteq> \<omega>"
+ 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"
+ shows "finite_measure_space M"
unfolding finite_measure_space_def finite_measure_space_axioms_def
proof (intro allI impI conjI)
- show "measure_space M \<mu>"
- proof (rule sigma_algebra.finite_additivity_sufficient)
- have *: "\<lparr>space = space M, sets = sets M\<rparr> = M" by auto
+ show "measure_space M"
+ proof (rule finite_additivity_sufficient)
+ have *: "\<lparr>space = space M, sets = Pow (space M), \<dots> = algebra.more M\<rparr> = M"
+ unfolding assms(2)[symmetric] by (auto intro!: algebra.equality)
show "sigma_algebra M"
- using sigma_algebra_Pow[of "space M" "more M"] assms(2)[symmetric] by (simp add: *)
+ using sigma_algebra_Pow[of "space M" "algebra.more M"]
+ unfolding * .
show "finite (space M)" by fact
- show "positive \<mu>" unfolding positive_def by fact
- show "additive M \<mu>" unfolding additive_def using assms by simp
+ show "positive M (measure M)" unfolding positive_def by fact
+ show "additive M (measure M)" unfolding additive_def using assms by simp
qed
- then interpret measure_space M \<mu> .
+ then interpret measure_space M .
show "finite_sigma_algebra M"
proof
show "finite (space M)" by fact
@@ -1363,18 +1372,18 @@
qed
lemma finite_measure_space_iff:
- "finite_measure_space M \<mu> \<longleftrightarrow>
- finite (space M) \<and> sets M = Pow(space M) \<and> \<mu> (space M) \<noteq> \<omega> \<and> \<mu> {} = 0 \<and>
- (\<forall>A\<subseteq>space M. \<forall>B\<subseteq>space M. A \<inter> B = {} \<longrightarrow> \<mu> (A \<union> B) = \<mu> A + \<mu> B)"
+ "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>
+ (\<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)
- assume "finite_measure_space M \<mu>"
- then interpret finite_measure_space M \<mu> .
+ assume "finite_measure_space M"
+ then interpret finite_measure_space M .
show ?rhs
using finite_space sets_eq_Pow measure_additive empty_measure finite_measure
by auto
next
- assume ?rhs then show "finite_measure_space M \<mu>"
+ assume ?rhs then show "finite_measure_space M"
by (auto intro!: finite_measure_spaceI)
qed
--- a/src/HOL/Probability/Probability.thy Wed Feb 02 10:35:41 2011 +0100
+++ b/src/HOL/Probability/Probability.thy Wed Feb 02 12:34:45 2011 +0100
@@ -1,6 +1,8 @@
theory Probability
imports
+ Complete_Measure
Information
"ex/Dining_Cryptographers"
+ "ex/Koepf_Duermuth_Countermeasure"
begin
end
--- a/src/HOL/Probability/Probability_Space.thy Wed Feb 02 10:35:41 2011 +0100
+++ b/src/HOL/Probability/Probability_Space.thy Wed Feb 02 12:34:45 2011 +0100
@@ -14,7 +14,7 @@
by (cases X) auto
locale prob_space = measure_space +
- assumes measure_space_1: "\<mu> (space M) = 1"
+ assumes measure_space_1: "measure M (space M) = 1"
lemma abs_real_of_pextreal[simp]: "\<bar>real (X :: pextreal)\<bar> = real X"
by simp
@@ -31,7 +31,7 @@
abbreviation (in prob_space) "prob \<equiv> \<lambda>A. real (\<mu> A)"
abbreviation (in prob_space) "prob_preserving \<equiv> measure_preserving"
abbreviation (in prob_space) "random_variable M' X \<equiv> sigma_algebra M' \<and> X \<in> measurable M M'"
-abbreviation (in prob_space) "expectation \<equiv> integral"
+abbreviation (in prob_space) "expectation \<equiv> integral\<^isup>L M"
definition (in prob_space)
"indep A B \<longleftrightarrow> A \<in> events \<and> B \<in> events \<and> prob (A \<inter> B) = prob A * prob B"
@@ -193,12 +193,14 @@
lemma (in prob_space) distribution_prob_space:
assumes "random_variable S X"
- shows "prob_space S (distribution X)"
+ shows "prob_space (S\<lparr>measure := distribution X\<rparr>)"
proof -
- interpret S: measure_space S "distribution X" unfolding distribution_def
- using assms by (intro measure_space_vimage) auto
+ interpret S: measure_space "S\<lparr>measure := distribution X\<rparr>"
+ unfolding distribution_def using assms
+ by (intro measure_space_vimage)
+ (auto intro!: sigma_algebra.sigma_algebra_cong[of S])
show ?thesis
- proof
+ proof (default, simp)
have "X -` space S \<inter> space M = space M"
using `random_variable S X` by (auto simp: measurable_def)
then show "distribution X (space S) = 1"
@@ -207,10 +209,10 @@
qed
lemma (in prob_space) AE_distribution:
- assumes X: "random_variable MX X" and "measure_space.almost_everywhere MX (distribution X) (\<lambda>x. Q x)"
+ assumes X: "random_variable MX X" and "measure_space.almost_everywhere (MX\<lparr>measure := distribution X\<rparr>) (\<lambda>x. Q x)"
shows "AE x. Q (X x)"
proof -
- interpret X: prob_space MX "distribution X" using X by (rule distribution_prob_space)
+ interpret X: prob_space "MX\<lparr>measure := distribution X\<rparr>" using X by (rule distribution_prob_space)
obtain N where N: "N \<in> sets MX" "distribution X N = 0" "{x\<in>space MX. \<not> Q x} \<subseteq> N"
using assms unfolding X.almost_everywhere_def by auto
show "AE x. Q (X x)"
@@ -228,12 +230,10 @@
lemma (in prob_space) distribution_lebesgue_thm2:
assumes "random_variable S X" and "A \<in> sets S"
- shows "distribution X A =
- measure_space.positive_integral S (distribution X) (indicator A)"
- (is "_ = measure_space.positive_integral _ ?D _")
+ shows "distribution X A = integral\<^isup>P (S\<lparr>measure := distribution X\<rparr>) (indicator A)"
proof -
- interpret S: prob_space S "distribution X" using assms(1) by (rule distribution_prob_space)
-
+ interpret S: prob_space "S\<lparr>measure := distribution X\<rparr>"
+ using assms(1) by (rule distribution_prob_space)
show ?thesis
using S.positive_integral_indicator(1)
using assms unfolding distribution_def by auto
@@ -249,7 +249,7 @@
qed
lemma (in prob_space) finite_expectation:
- assumes "finite (space M)" "random_variable borel X"
+ assumes "finite (X`space M)" "random_variable borel X"
shows "expectation X = (\<Sum> r \<in> X ` (space M). r * real (distribution X {r}))"
using assms unfolding distribution_def using finite_expectation1 by auto
@@ -290,27 +290,24 @@
assumes "distribution X {x} = 1"
assumes "y \<noteq> x"
shows "distribution X {y} = 0"
-proof -
- from distribution_prob_space[OF X]
- interpret S: prob_space ?S "distribution X" by simp
- have x: "{x} \<in> sets ?S"
- proof (rule ccontr)
- assume "{x} \<notin> sets ?S"
- hence "X -` {x} \<inter> space M = {}" by auto
- thus "False" using assms unfolding distribution_def by auto
- qed
- have [simp]: "{y} \<inter> {x} = {}" "{x} - {y} = {x}" using `y \<noteq> x` by auto
- show ?thesis
- proof cases
- assume "{y} \<in> sets ?S"
- with `{x} \<in> sets ?S` assms show "distribution X {y} = 0"
- using S.measure_inter_full_set[of "{y}" "{x}"]
- by simp
- next
- assume "{y} \<notin> sets ?S"
- hence "X -` {y} \<inter> space M = {}" by auto
- thus "distribution X {y} = 0" unfolding distribution_def by auto
- qed
+proof cases
+ { fix x have "X -` {x} \<inter> space M \<in> sets M"
+ proof cases
+ assume "x \<in> X`space M" with X show ?thesis
+ by (auto simp: measurable_def image_iff)
+ next
+ assume "x \<notin> X`space M" then have "X -` {x} \<inter> space M = {}" by auto
+ then show ?thesis by auto
+ qed } note single = this
+ have "X -` {x} \<inter> space M - X -` {y} \<inter> space M = X -` {x} \<inter> space M"
+ "X -` {y} \<inter> space M \<inter> (X -` {x} \<inter> space M) = {}"
+ using `y \<noteq> x` by auto
+ with measure_inter_full_set[OF single single, of x y] assms(2)
+ show ?thesis unfolding distribution_def measure_space_1 by auto
+next
+ assume "{y} \<notin> sets ?S"
+ then have "X -` {y} \<inter> space M = {}" by auto
+ thus "distribution X {y} = 0" unfolding distribution_def by auto
qed
lemma (in prob_space) joint_distribution_Times_le_fst:
@@ -344,14 +341,14 @@
lemma (in prob_space) random_variable_pairI:
assumes "random_variable MX X"
assumes "random_variable MY Y"
- shows "random_variable (sigma (pair_algebra MX MY)) (\<lambda>x. (X x, Y x))"
+ shows "random_variable (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))"
proof
interpret MX: sigma_algebra MX using assms by simp
interpret MY: sigma_algebra MY using assms by simp
interpret P: pair_sigma_algebra MX MY by default
- show "sigma_algebra (sigma (pair_algebra MX MY))" by default
+ show "sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default
have sa: "sigma_algebra M" by default
- show "(\<lambda>x. (X x, Y x)) \<in> measurable M (sigma (pair_algebra MX MY))"
+ show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)"
unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def)
qed
@@ -377,15 +374,12 @@
joint_distribution (\<lambda>x. (X x, Y x)) Z {((x, y), z)}"
unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>])
-locale pair_prob_space = M1: prob_space M1 p1 + M2: prob_space M2 p2 for M1 p1 M2 p2
-
-sublocale pair_prob_space \<subseteq> pair_sigma_finite M1 p1 M2 p2 by default
+locale pair_prob_space = M1: prob_space M1 + M2: prob_space M2 for M1 M2
-sublocale pair_prob_space \<subseteq> P: prob_space P pair_measure
-proof
- show "pair_measure (space P) = 1"
- by (simp add: pair_algebra_def pair_measure_times M1.measure_space_1 M2.measure_space_1)
-qed
+sublocale pair_prob_space \<subseteq> pair_sigma_finite M1 M2 by default
+
+sublocale pair_prob_space \<subseteq> P: prob_space P
+by default (simp add: pair_measure_times M1.measure_space_1 M2.measure_space_1 space_pair_measure)
lemma countably_additiveI[case_names countably]:
assumes "\<And>A. \<lbrakk> range A \<subseteq> sets M ; disjoint_family A ; (\<Union>i. A i) \<in> sets M\<rbrakk> \<Longrightarrow>
@@ -395,38 +389,8 @@
lemma (in prob_space) joint_distribution_prob_space:
assumes "random_variable MX X" "random_variable MY Y"
- shows "prob_space (sigma (pair_algebra MX MY)) (joint_distribution X Y)"
-proof -
- interpret X: prob_space MX "distribution X" by (intro distribution_prob_space assms)
- interpret Y: prob_space MY "distribution Y" by (intro distribution_prob_space assms)
- interpret XY: pair_sigma_finite MX "distribution X" MY "distribution Y" by default
- show ?thesis
- proof
- let "?X A" = "(\<lambda>x. (X x, Y x)) -` A \<inter> space M"
- show "joint_distribution X Y {} = 0" by (simp add: distribution_def)
- show "countably_additive XY.P (joint_distribution X Y)"
- proof (rule countably_additiveI)
- fix A :: "nat \<Rightarrow> ('b \<times> 'c) set"
- assume A: "range A \<subseteq> sets XY.P" and df: "disjoint_family A"
- have "(\<Sum>\<^isub>\<infinity>n. \<mu> (?X (A n))) = \<mu> (\<Union>x. ?X (A x))"
- proof (intro measure_countably_additive)
- have "sigma_algebra M" by default
- then have *: "(\<lambda>x. (X x, Y x)) \<in> measurable M XY.P"
- using assms by (simp add: XY.measurable_pair comp_def)
- show "range (\<lambda>n. ?X (A n)) \<subseteq> events"
- using measurable_sets[OF *] A by auto
- show "disjoint_family (\<lambda>n. ?X (A n))"
- by (intro disjoint_family_on_bisimulation[OF df]) auto
- qed
- then show "(\<Sum>\<^isub>\<infinity>n. joint_distribution X Y (A n)) = joint_distribution X Y (\<Union>i. A i)"
- by (simp add: distribution_def vimage_UN)
- qed
- have "?X (space MX \<times> space MY) = space M"
- using assms by (auto simp: measurable_def)
- then show "joint_distribution X Y (space XY.P) = 1"
- by (simp add: space_pair_algebra distribution_def measure_space_1)
- qed
-qed
+ shows "prob_space ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := joint_distribution X Y\<rparr>)"
+ using random_variable_pairI[OF assms] by (rule distribution_prob_space)
section "Probability spaces on finite sets"
@@ -443,32 +407,32 @@
lemma (in prob_space) distribution_finite_prob_space:
assumes "finite_random_variable MX X"
- shows "finite_prob_space MX (distribution X)"
+ shows "finite_prob_space (MX\<lparr>measure := distribution X\<rparr>)"
proof -
- interpret X: prob_space MX "distribution X"
+ interpret X: prob_space "MX\<lparr>measure := distribution X\<rparr>"
using assms[THEN finite_random_variableD] by (rule distribution_prob_space)
interpret MX: finite_sigma_algebra MX
- using assms by simp
+ using assms by auto
show ?thesis
- proof
+ proof (default, simp_all)
fix x assume "x \<in> space MX"
then have "X -` {x} \<inter> space M \<in> sets M"
using assms unfolding measurable_def by simp
then show "distribution X {x} \<noteq> \<omega>"
unfolding distribution_def by simp
- qed
+ qed (rule MX.finite_space)
qed
lemma (in prob_space) simple_function_imp_finite_random_variable[simp, intro]:
- assumes "simple_function X"
- shows "finite_random_variable \<lparr> space = X`space M, sets = Pow (X`space M) \<rparr> X"
+ assumes "simple_function M X"
+ shows "finite_random_variable \<lparr> space = X`space M, sets = Pow (X`space M), \<dots> = x \<rparr> X"
+ (is "finite_random_variable ?X _")
proof (intro conjI)
have [simp]: "finite (X ` space M)" using assms unfolding simple_function_def by simp
- interpret X: sigma_algebra "\<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr>"
- by (rule sigma_algebra_Pow)
- show "finite_sigma_algebra \<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr>"
+ interpret X: sigma_algebra ?X by (rule sigma_algebra_Pow)
+ show "finite_sigma_algebra ?X"
by default auto
- show "X \<in> measurable M \<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr>"
+ show "X \<in> measurable M ?X"
proof (unfold measurable_def, clarsimp)
fix A assume A: "A \<subseteq> X`space M"
then have "finite A" by (rule finite_subset) simp
@@ -481,13 +445,13 @@
qed
lemma (in prob_space) simple_function_imp_random_variable[simp, intro]:
- assumes "simple_function X"
- shows "random_variable \<lparr> space = X`space M, sets = Pow (X`space M) \<rparr> X"
- using simple_function_imp_finite_random_variable[OF assms]
+ assumes "simple_function M X"
+ shows "random_variable \<lparr> space = X`space M, sets = Pow (X`space M), \<dots> = ext \<rparr> X"
+ using simple_function_imp_finite_random_variable[OF assms, of ext]
by (auto dest!: finite_random_variableD)
lemma (in prob_space) sum_over_space_real_distribution:
- "simple_function X \<Longrightarrow> (\<Sum>x\<in>X`space M. real (distribution X {x})) = 1"
+ "simple_function M X \<Longrightarrow> (\<Sum>x\<in>X`space M. real (distribution X {x})) = 1"
unfolding distribution_def prob_space[symmetric]
by (subst real_finite_measure_finite_Union[symmetric])
(auto simp add: disjoint_family_on_def simple_function_def
@@ -496,14 +460,14 @@
lemma (in prob_space) finite_random_variable_pairI:
assumes "finite_random_variable MX X"
assumes "finite_random_variable MY Y"
- shows "finite_random_variable (sigma (pair_algebra MX MY)) (\<lambda>x. (X x, Y x))"
+ shows "finite_random_variable (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))"
proof
interpret MX: finite_sigma_algebra MX using assms by simp
interpret MY: finite_sigma_algebra MY using assms by simp
interpret P: pair_finite_sigma_algebra MX MY by default
- show "finite_sigma_algebra (sigma (pair_algebra MX MY))" by default
+ show "finite_sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default
have sa: "sigma_algebra M" by default
- show "(\<lambda>x. (X x, Y x)) \<in> measurable M (sigma (pair_algebra MX MY))"
+ show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)"
unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def)
qed
@@ -599,6 +563,7 @@
finite_random_variable_imp_sets[OF Y]] by simp
lemma (in prob_space) setsum_real_joint_distribution:
+ fixes MX :: "('c, 'x) measure_space_scheme" and MY :: "('d, 'y) measure_space_scheme"
assumes X: "finite_random_variable MX X"
assumes Y: "random_variable MY Y" "B \<in> sets MY"
shows "(\<Sum>a\<in>space MX. real (joint_distribution X Y ({a} \<times> B))) = real (distribution Y B)"
@@ -607,10 +572,11 @@
show ?thesis
unfolding setsum_joint_distribution[OF assms, symmetric]
using distribution_finite[OF random_variable_pairI[OF finite_random_variableD[OF X] Y(1)]] Y(2)
- by (simp add: space_pair_algebra in_sigma pair_algebraI MX.sets_eq_Pow real_of_pextreal_setsum)
+ by (simp add: space_pair_measure real_of_pextreal_setsum)
qed
lemma (in prob_space) setsum_real_joint_distribution_singleton:
+ fixes MX :: "('c, 'x) measure_space_scheme" and MY :: "('d, 'y) measure_space_scheme"
assumes X: "finite_random_variable MX X"
assumes Y: "finite_random_variable MY Y" "b \<in> space MY"
shows "(\<Sum>a\<in>space MX. real (joint_distribution X Y {(a,b)})) = real (distribution Y {b})"
@@ -618,39 +584,20 @@
finite_random_variableD[OF Y(1)]
finite_random_variable_imp_sets[OF Y]] by simp
-locale pair_finite_prob_space = M1: finite_prob_space M1 p1 + M2: finite_prob_space M2 p2 for M1 p1 M2 p2
+locale pair_finite_prob_space = M1: finite_prob_space M1 + M2: finite_prob_space M2 for M1 M2
-sublocale pair_finite_prob_space \<subseteq> pair_prob_space M1 p1 M2 p2 by default
-sublocale pair_finite_prob_space \<subseteq> pair_finite_space M1 p1 M2 p2 by default
-sublocale pair_finite_prob_space \<subseteq> finite_prob_space P pair_measure by default
+sublocale pair_finite_prob_space \<subseteq> pair_prob_space M1 M2 by default
+sublocale pair_finite_prob_space \<subseteq> pair_finite_space M1 M2 by default
+sublocale pair_finite_prob_space \<subseteq> finite_prob_space P by default
lemma (in prob_space) joint_distribution_finite_prob_space:
assumes X: "finite_random_variable MX X"
assumes Y: "finite_random_variable MY Y"
- shows "finite_prob_space (sigma (pair_algebra MX MY)) (joint_distribution X Y)"
-proof -
- interpret X: finite_prob_space MX "distribution X"
- using X by (rule distribution_finite_prob_space)
- interpret Y: finite_prob_space MY "distribution Y"
- using Y by (rule distribution_finite_prob_space)
- interpret P: prob_space "sigma (pair_algebra MX MY)" "joint_distribution X Y"
- using assms[THEN finite_random_variableD] by (rule joint_distribution_prob_space)
- interpret XY: pair_finite_prob_space MX "distribution X" MY "distribution Y"
- by default
- show ?thesis
- proof
- fix x assume "x \<in> space XY.P"
- moreover have "(\<lambda>x. (X x, Y x)) \<in> measurable M XY.P"
- using X Y by (intro XY.measurable_pair) (simp_all add: o_def, default)
- ultimately have "(\<lambda>x. (X x, Y x)) -` {x} \<inter> space M \<in> sets M"
- unfolding measurable_def by simp
- then show "joint_distribution X Y {x} \<noteq> \<omega>"
- unfolding distribution_def by simp
- qed
-qed
+ shows "finite_prob_space ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := joint_distribution X Y\<rparr>)"
+ by (intro distribution_finite_prob_space finite_random_variable_pairI X Y)
lemma finite_prob_space_eq:
- "finite_prob_space M \<mu> \<longleftrightarrow> finite_measure_space M \<mu> \<and> \<mu> (space M) = 1"
+ "finite_prob_space M \<longleftrightarrow> finite_measure_space M \<and> measure M (space M) = 1"
unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def
by auto
@@ -852,48 +799,51 @@
lemma (in prob_space) prob_space_subalgebra:
assumes "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M"
- shows "prob_space N \<mu>"
+ and "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A"
+ shows "prob_space N"
proof -
- interpret N: measure_space N \<mu>
- using measure_space_subalgebra[OF assms] .
+ interpret N: measure_space N
+ by (rule measure_space_subalgebra[OF assms])
show ?thesis
- proof qed (simp add: `space N = space M` measure_space_1)
+ proof qed (insert assms(4)[OF N.top], simp add: assms measure_space_1)
qed
lemma (in prob_space) prob_space_of_restricted_space:
assumes "\<mu> A \<noteq> 0" "\<mu> A \<noteq> \<omega>" "A \<in> sets M"
- shows "prob_space (restricted_space A) (\<lambda>S. \<mu> S / \<mu> A)"
- unfolding prob_space_def prob_space_axioms_def
-proof
- show "\<mu> (space (restricted_space A)) / \<mu> A = 1"
- using `\<mu> A \<noteq> 0` `\<mu> A \<noteq> \<omega>` by (auto simp: pextreal_noteq_omega_Ex)
- have *: "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S" by (simp add: mult_commute)
- interpret A: measure_space "restricted_space A" \<mu>
+ shows "prob_space (restricted_space A \<lparr>measure := \<lambda>S. \<mu> S / \<mu> A\<rparr>)"
+ (is "prob_space ?P")
+proof -
+ interpret A: measure_space "restricted_space A"
using `A \<in> sets M` by (rule restricted_measure_space)
- show "measure_space (restricted_space A) (\<lambda>S. \<mu> S / \<mu> A)"
+ interpret A': sigma_algebra ?P
+ by (rule A.sigma_algebra_cong) auto
+ show "prob_space ?P"
proof
- show "\<mu> {} / \<mu> A = 0" by auto
- show "countably_additive (restricted_space A) (\<lambda>S. \<mu> S / \<mu> A)"
- unfolding countably_additive_def psuminf_cmult_right *
+ show "measure ?P (space ?P) = 1"
+ using `\<mu> A \<noteq> 0` `\<mu> A \<noteq> \<omega>` by (auto simp: pextreal_noteq_omega_Ex)
+ show "measure ?P {} = 0" by auto
+ have "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S" by (simp add: mult_commute)
+ then show "countably_additive ?P (measure ?P)"
+ unfolding countably_additive_def psuminf_cmult_right
using A.measure_countably_additive by auto
qed
qed
lemma finite_prob_spaceI:
- assumes "finite (space M)" "sets M = Pow(space M)" "\<mu> (space M) = 1" "\<mu> {} = 0"
- and "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> \<mu> (A \<union> B) = \<mu> A + \<mu> B"
- shows "finite_prob_space M \<mu>"
+ assumes "finite (space M)" "sets M = Pow(space M)" "measure M (space M) = 1" "measure M {} = 0"
+ and "\<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"
+ shows "finite_prob_space M"
unfolding finite_prob_space_eq
proof
- show "finite_measure_space M \<mu>" using assms
+ show "finite_measure_space M" using assms
by (auto intro!: finite_measure_spaceI)
- show "\<mu> (space M) = 1" by fact
+ show "measure M (space M) = 1" by fact
qed
lemma (in finite_prob_space) finite_measure_space:
fixes X :: "'a \<Rightarrow> 'x"
- shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)"
- (is "finite_measure_space ?S _")
+ shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = distribution X\<rparr>"
+ (is "finite_measure_space ?S")
proof (rule finite_measure_spaceI, simp_all)
show "finite (X ` space M)" using finite_space by simp
next
@@ -905,7 +855,7 @@
qed
lemma (in finite_prob_space) finite_prob_space_of_images:
- "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)"
+ "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = distribution X \<rparr>"
by (simp add: finite_prob_space_eq finite_measure_space)
lemma (in finite_prob_space) real_distribution_order':
@@ -919,8 +869,8 @@
lemma (in finite_prob_space) finite_product_measure_space:
fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y"
assumes "finite s1" "finite s2"
- shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2)\<rparr> (joint_distribution X Y)"
- (is "finite_measure_space ?M ?D")
+ shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = joint_distribution X Y\<rparr>"
+ (is "finite_measure_space ?M")
proof (rule finite_measure_spaceI, simp_all)
show "finite (s1 \<times> s2)"
using assms by auto
@@ -936,14 +886,14 @@
lemma (in finite_prob_space) finite_product_measure_space_of_images:
shows "finite_measure_space \<lparr> space = X ` space M \<times> Y ` space M,
- sets = Pow (X ` space M \<times> Y ` space M) \<rparr>
- (joint_distribution X Y)"
+ sets = Pow (X ` space M \<times> Y ` space M),
+ measure = joint_distribution X Y \<rparr>"
using finite_space by (auto intro!: finite_product_measure_space)
lemma (in finite_prob_space) finite_product_prob_space_of_images:
- "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M)\<rparr>
- (joint_distribution X Y)"
- (is "finite_prob_space ?S _")
+ "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M),
+ measure = joint_distribution X Y \<rparr>"
+ (is "finite_prob_space ?S")
proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images)
have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto
thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1"
@@ -953,27 +903,29 @@
section "Conditional Expectation and Probability"
lemma (in prob_space) conditional_expectation_exists:
- fixes X :: "'a \<Rightarrow> pextreal"
+ fixes X :: "'a \<Rightarrow> pextreal" and N :: "('a, 'b) measure_space_scheme"
assumes borel: "X \<in> borel_measurable M"
- and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M"
+ and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
shows "\<exists>Y\<in>borel_measurable N. \<forall>C\<in>sets N.
- (\<integral>\<^isup>+x. Y x * indicator C x) = (\<integral>\<^isup>+x. X x * indicator C x)"
+ (\<integral>\<^isup>+x. Y x * indicator C x \<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x \<partial>M)"
proof -
- interpret P: prob_space N \<mu>
+ note N(4)[simp]
+ interpret P: prob_space N
using prob_space_subalgebra[OF N] .
let "?f A" = "\<lambda>x. X x * indicator A x"
- let "?Q A" = "positive_integral (?f A)"
+ let "?Q A" = "integral\<^isup>P M (?f A)"
from measure_space_density[OF borel]
- have Q: "measure_space N ?Q"
- by (rule measure_space.measure_space_subalgebra[OF _ N])
- then interpret Q: measure_space N ?Q .
+ have Q: "measure_space (N\<lparr> measure := ?Q \<rparr>)"
+ apply (rule measure_space.measure_space_subalgebra[of "M\<lparr> measure := ?Q \<rparr>"])
+ using N by (auto intro!: P.sigma_algebra_cong)
+ then interpret Q: measure_space "N\<lparr> measure := ?Q \<rparr>" .
have "P.absolutely_continuous ?Q"
unfolding P.absolutely_continuous_def
proof safe
- fix A assume "A \<in> sets N" "\<mu> A = 0"
+ fix A assume "A \<in> sets N" "P.\<mu> A = 0"
moreover then have f_borel: "?f A \<in> borel_measurable M"
using borel N by (auto intro: borel_measurable_indicator)
moreover have "{x\<in>space M. ?f A x \<noteq> 0} = (?f A -` {0<..} \<inter> space M) \<inter> A"
@@ -986,28 +938,28 @@
qed
from P.Radon_Nikodym[OF Q this]
obtain Y where Y: "Y \<in> borel_measurable N"
- "\<And>A. A \<in> sets N \<Longrightarrow> ?Q A = P.positive_integral (\<lambda>x. Y x * indicator A x)"
+ "\<And>A. A \<in> sets N \<Longrightarrow> ?Q A =(\<integral>\<^isup>+x. Y x * indicator A x \<partial>N)"
by blast
with N(2) show ?thesis
- by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ N(2,3,1)])
+ by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ N(2,3,4,1)])
qed
definition (in prob_space)
"conditional_expectation N X = (SOME Y. Y\<in>borel_measurable N
- \<and> (\<forall>C\<in>sets N. (\<integral>\<^isup>+x. Y x * indicator C x) = (\<integral>\<^isup>+x. X x * indicator C x)))"
+ \<and> (\<forall>C\<in>sets N. (\<integral>\<^isup>+x. Y x * indicator C x\<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x\<partial>M)))"
abbreviation (in prob_space)
"conditional_prob N A \<equiv> conditional_expectation N (indicator A)"
lemma (in prob_space)
- fixes X :: "'a \<Rightarrow> pextreal"
+ fixes X :: "'a \<Rightarrow> pextreal" and N :: "('a, 'b) measure_space_scheme"
assumes borel: "X \<in> borel_measurable M"
- and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M"
+ and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
shows borel_measurable_conditional_expectation:
"conditional_expectation N X \<in> borel_measurable N"
and conditional_expectation: "\<And>C. C \<in> sets N \<Longrightarrow>
- (\<integral>\<^isup>+x. conditional_expectation N X x * indicator C x) =
- (\<integral>\<^isup>+x. X x * indicator C x)"
+ (\<integral>\<^isup>+x. conditional_expectation N X x * indicator C x \<partial>M) =
+ (\<integral>\<^isup>+x. X x * indicator C x \<partial>M)"
(is "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C")
proof -
note CE = conditional_expectation_exists[OF assms, unfolded Bex_def]
@@ -1042,21 +994,21 @@
assume "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
from va.borel_measurable_implies_simple_function_sequence[OF this]
- obtain f where f: "\<And>i. va.simple_function (f i)" and "f \<up> Z" by blast
+ obtain f where f: "\<And>i. simple_function (M'.vimage_algebra (space M) Y) (f i)" and "f \<up> Z" by blast
- have "\<forall>i. \<exists>g. M'.simple_function g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
+ have "\<forall>i. \<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
proof
fix i
from f[of i] have "finite (f i`space M)" and B_ex:
"\<forall>z\<in>(f i)`space M. \<exists>B. B \<in> sets M' \<and> (f i) -` {z} \<inter> space M = Y -` B \<inter> space M"
- unfolding va.simple_function_def by auto
+ unfolding simple_function_def by auto
from B_ex[THEN bchoice] guess B .. note B = this
let ?g = "\<lambda>x. \<Sum>z\<in>f i`space M. z * indicator (B z) x"
- show "\<exists>g. M'.simple_function g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
+ show "\<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
proof (intro exI[of _ ?g] conjI ballI)
- show "M'.simple_function ?g" using B by auto
+ show "simple_function M' ?g" using B by auto
fix x assume "x \<in> space M"
then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::pextreal)"
--- a/src/HOL/Probability/Product_Measure.thy Wed Feb 02 10:35:41 2011 +0100
+++ b/src/HOL/Probability/Product_Measure.thy Wed Feb 02 12:34:45 2011 +0100
@@ -2,6 +2,20 @@
imports Lebesgue_Integration
begin
+lemma measurable_cancel_measure2[simp]:
+ "measurable M1 (M2\<lparr>measure := m\<rparr>) = measurable M1 M2"
+ unfolding measurable_def by auto
+
+lemma measurable_cancel_measure1[simp]:
+ "measurable (M1\<lparr>measure := m\<rparr>) M2 = measurable M1 M2"
+ unfolding measurable_def by auto
+
+lemma sigma_sets_subseteq: assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
+proof
+ fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
+ by induct (insert `A \<subseteq> B`, auto intro: sigma_sets.intros)
+qed
+
lemma times_Int_times: "A \<times> B \<inter> C \<times> D = (A \<inter> C) \<times> (B \<inter> D)"
by auto
@@ -20,6 +34,18 @@
abbreviation
"Pi\<^isub>E A B \<equiv> Pi A B \<inter> extensional A"
+syntax
+ "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3PIE _:_./ _)" 10)
+
+syntax (xsymbols)
+ "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\<Pi>\<^isub>E _\<in>_./ _)" 10)
+
+syntax (HTML output)
+ "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\<Pi>\<^isub>E _\<in>_./ _)" 10)
+
+translations
+ "PIE x:A. B" == "CONST Pi\<^isub>E A (%x. B)"
+
abbreviation
funcset_extensional :: "['a set, 'b set] => ('a => 'b) set"
(infixr "->\<^isub>E" 60) where
@@ -171,61 +197,121 @@
"i \<notin> I \<Longrightarrow> (restrict f I)(i := y) = restrict (f(i := y)) (insert i I)"
by (auto simp: fun_eq_iff)
+lemma Pi_eq_subset:
+ assumes ne: "\<And>i. i \<in> I \<Longrightarrow> F i \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> F' i \<noteq> {}"
+ assumes eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and "i \<in> I"
+ shows "F i \<subseteq> F' i"
+proof
+ fix x assume "x \<in> F i"
+ with ne have "\<forall>j. \<exists>y. ((j \<in> I \<longrightarrow> y \<in> F j \<and> (i = j \<longrightarrow> x = y)) \<and> (j \<notin> I \<longrightarrow> y = undefined))" by auto
+ from choice[OF this] guess f .. note f = this
+ then have "f \<in> Pi\<^isub>E I F" by (auto simp: extensional_def)
+ then have "f \<in> Pi\<^isub>E I F'" using assms by simp
+ then show "x \<in> F' i" using f `i \<in> I` by auto
+qed
+
+lemma Pi_eq_iff_not_empty:
+ assumes ne: "\<And>i. i \<in> I \<Longrightarrow> F i \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> F' i \<noteq> {}"
+ shows "Pi\<^isub>E I F = Pi\<^isub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i)"
+proof (intro iffI ballI)
+ fix i assume eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and i: "i \<in> I"
+ show "F i = F' i"
+ using Pi_eq_subset[of I F F', OF ne eq i]
+ using Pi_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i]
+ by auto
+qed auto
+
+lemma Pi_eq_empty_iff:
+ "Pi\<^isub>E I F = {} \<longleftrightarrow> (\<exists>i\<in>I. F i = {})"
+proof
+ assume "Pi\<^isub>E I F = {}"
+ show "\<exists>i\<in>I. F i = {}"
+ proof (rule ccontr)
+ assume "\<not> ?thesis"
+ then have "\<forall>i. \<exists>y. (i \<in> I \<longrightarrow> y \<in> F i) \<and> (i \<notin> I \<longrightarrow> y = undefined)" by auto
+ from choice[OF this] guess f ..
+ then have "f \<in> Pi\<^isub>E I F" by (auto simp: extensional_def)
+ with `Pi\<^isub>E I F = {}` show False by auto
+ qed
+qed auto
+
+lemma Pi_eq_iff:
+ "Pi\<^isub>E I F = Pi\<^isub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i) \<or> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))"
+proof (intro iffI disjCI)
+ assume eq[simp]: "Pi\<^isub>E I F = Pi\<^isub>E I F'"
+ assume "\<not> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))"
+ then have "(\<forall>i\<in>I. F i \<noteq> {}) \<and> (\<forall>i\<in>I. F' i \<noteq> {})"
+ using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto
+ with Pi_eq_iff_not_empty[of I F F'] show "\<forall>i\<in>I. F i = F' i" by auto
+next
+ assume "(\<forall>i\<in>I. F i = F' i) \<or> (\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {})"
+ then show "Pi\<^isub>E I F = Pi\<^isub>E I F'"
+ using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto
+qed
+
section "Binary products"
definition
- "pair_algebra A B = \<lparr> space = space A \<times> space B,
- sets = {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B} \<rparr>"
+ "pair_measure_generator A B =
+ \<lparr> space = space A \<times> space B,
+ sets = {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B},
+ measure = \<lambda>X. \<integral>\<^isup>+x. (\<integral>\<^isup>+y. indicator X (x,y) \<partial>B) \<partial>A \<rparr>"
+
+definition pair_measure (infixr "\<Otimes>\<^isub>M" 80) where
+ "A \<Otimes>\<^isub>M B = sigma (pair_measure_generator A B)"
locale pair_sigma_algebra = M1: sigma_algebra M1 + M2: sigma_algebra M2
- for M1 M2
+ for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
+
+abbreviation (in pair_sigma_algebra)
+ "E \<equiv> pair_measure_generator M1 M2"
abbreviation (in pair_sigma_algebra)
- "E \<equiv> pair_algebra M1 M2"
+ "P \<equiv> M1 \<Otimes>\<^isub>M M2"
-abbreviation (in pair_sigma_algebra)
- "P \<equiv> sigma E"
+lemma sigma_algebra_pair_measure:
+ "sets M1 \<subseteq> Pow (space M1) \<Longrightarrow> sets M2 \<subseteq> Pow (space M2) \<Longrightarrow> sigma_algebra (pair_measure M1 M2)"
+ by (force simp: pair_measure_def pair_measure_generator_def intro!: sigma_algebra_sigma)
sublocale pair_sigma_algebra \<subseteq> sigma_algebra P
- using M1.sets_into_space M2.sets_into_space
- by (force simp: pair_algebra_def intro!: sigma_algebra_sigma)
+ using M1.space_closed M2.space_closed
+ by (rule sigma_algebra_pair_measure)
-lemma pair_algebraI[intro, simp]:
- "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (pair_algebra A B)"
- by (auto simp add: pair_algebra_def)
+lemma pair_measure_generatorI[intro, simp]:
+ "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (pair_measure_generator A B)"
+ by (auto simp add: pair_measure_generator_def)
-lemma space_pair_algebra:
- "space (pair_algebra A B) = space A \<times> space B"
- by (simp add: pair_algebra_def)
+lemma pair_measureI[intro, simp]:
+ "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^isub>M B)"
+ by (auto simp add: pair_measure_def)
-lemma sets_pair_algebra: "sets (pair_algebra N M) = (\<lambda>(x, y). x \<times> y) ` (sets N \<times> sets M)"
- unfolding pair_algebra_def by auto
+lemma space_pair_measure:
+ "space (A \<Otimes>\<^isub>M B) = space A \<times> space B"
+ by (simp add: pair_measure_def pair_measure_generator_def)
-lemma pair_algebra_sets_into_space:
- assumes "sets M \<subseteq> Pow (space M)" "sets N \<subseteq> Pow (space N)"
- shows "sets (pair_algebra M N) \<subseteq> Pow (space (pair_algebra M N))"
- using assms by (auto simp: pair_algebra_def)
+lemma sets_pair_measure_generator:
+ "sets (pair_measure_generator N M) = (\<lambda>(x, y). x \<times> y) ` (sets N \<times> sets M)"
+ unfolding pair_measure_generator_def by auto
-lemma pair_algebra_Int_snd:
+lemma pair_measure_generator_sets_into_space:
+ assumes "sets M \<subseteq> Pow (space M)" "sets N \<subseteq> Pow (space N)"
+ shows "sets (pair_measure_generator M N) \<subseteq> Pow (space (pair_measure_generator M N))"
+ using assms by (auto simp: pair_measure_generator_def)
+
+lemma pair_measure_generator_Int_snd:
assumes "sets S1 \<subseteq> Pow (space S1)"
- shows "pair_algebra S1 (algebra.restricted_space S2 A) =
- algebra.restricted_space (pair_algebra S1 S2) (space S1 \<times> A)"
+ shows "sets (pair_measure_generator S1 (algebra.restricted_space S2 A)) =
+ sets (algebra.restricted_space (pair_measure_generator S1 S2) (space S1 \<times> A))"
(is "?L = ?R")
-proof (intro algebra.equality set_eqI iffI)
- fix X assume "X \<in> sets ?L"
- then obtain A1 A2 where X: "X = A1 \<times> (A \<inter> A2)" and "A1 \<in> sets S1" "A2 \<in> sets S2"
- by (auto simp: pair_algebra_def)
- then show "X \<in> sets ?R" unfolding pair_algebra_def
- using assms apply simp by (intro image_eqI[of _ _ "A1 \<times> A2"]) auto
-next
- fix X assume "X \<in> sets ?R"
- then obtain A1 A2 where "X = space S1 \<times> A \<inter> A1 \<times> A2" "A1 \<in> sets S1" "A2 \<in> sets S2"
- by (auto simp: pair_algebra_def)
- moreover then have "X = A1 \<times> (A \<inter> A2)"
- using assms by auto
- ultimately show "X \<in> sets ?L"
- unfolding pair_algebra_def by auto
-qed (auto simp add: pair_algebra_def)
+ apply (auto simp: pair_measure_generator_def image_iff)
+ using assms
+ apply (rule_tac x="a \<times> xa" in exI)
+ apply force
+ using assms
+ apply (rule_tac x="a" in exI)
+ apply (rule_tac x="b \<inter> A" in exI)
+ apply auto
+ done
lemma (in pair_sigma_algebra)
shows measurable_fst[intro!, simp]:
@@ -243,7 +329,7 @@
apply - apply (rule bexI[of _ "space M1"]) apply (rule bexI[of _ X])
using M2.sets_into_space by force+ }
ultimately have "?fst \<and> ?snd"
- by (fastsimp simp: measurable_def sets_sigma space_pair_algebra
+ by (fastsimp simp: measurable_def sets_sigma space_pair_measure
intro!: sigma_sets.Basic)
then show ?fst ?snd by auto
qed
@@ -257,15 +343,15 @@
from assms show ?thesis
proof (safe intro!: measurable_comp[where b=P])
assume f: "(fst \<circ> f) \<in> measurable M M1" and s: "(snd \<circ> f) \<in> measurable M M2"
- show "f \<in> measurable M P"
+ show "f \<in> measurable M P" unfolding pair_measure_def
proof (rule M.measurable_sigma)
- show "sets (pair_algebra M1 M2) \<subseteq> Pow (space E)"
- unfolding pair_algebra_def using M1.sets_into_space M2.sets_into_space by auto
+ show "sets (pair_measure_generator M1 M2) \<subseteq> Pow (space E)"
+ unfolding pair_measure_generator_def using M1.sets_into_space M2.sets_into_space by auto
show "f \<in> space M \<rightarrow> space E"
- using f s by (auto simp: mem_Times_iff measurable_def comp_def space_sigma space_pair_algebra)
+ using f s by (auto simp: mem_Times_iff measurable_def comp_def space_sigma pair_measure_generator_def)
fix A assume "A \<in> sets E"
then obtain B C where "B \<in> sets M1" "C \<in> sets M2" "A = B \<times> C"
- unfolding pair_algebra_def by auto
+ unfolding pair_measure_generator_def by auto
moreover have "(fst \<circ> f) -` B \<inter> space M \<in> sets M"
using f `B \<in> sets M1` unfolding measurable_def by auto
moreover have "(snd \<circ> f) -` C \<inter> space M \<in> sets M"
@@ -283,52 +369,58 @@
shows "f \<in> measurable M P"
unfolding measurable_pair_iff[OF assms(1)] using assms(2,3) by simp
-lemma pair_algebraE:
- assumes "X \<in> sets (pair_algebra M1 M2)"
+lemma pair_measure_generatorE:
+ assumes "X \<in> sets (pair_measure_generator M1 M2)"
obtains A B where "X = A \<times> B" "A \<in> sets M1" "B \<in> sets M2"
- using assms unfolding pair_algebra_def by auto
+ using assms unfolding pair_measure_generator_def by auto
-lemma (in pair_sigma_algebra) pair_algebra_swap:
- "(\<lambda>X. (\<lambda>(x,y). (y,x)) -` X \<inter> space M2 \<times> space M1) ` sets E = sets (pair_algebra M2 M1)"
-proof (safe elim!: pair_algebraE)
+lemma (in pair_sigma_algebra) pair_measure_generator_swap:
+ "(\<lambda>X. (\<lambda>(x,y). (y,x)) -` X \<inter> space M2 \<times> space M1) ` sets E = sets (pair_measure_generator M2 M1)"
+proof (safe elim!: pair_measure_generatorE)
fix A B assume "A \<in> sets M1" "B \<in> sets M2"
moreover then have "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space M2 \<times> space M1 = B \<times> A"
using M1.sets_into_space M2.sets_into_space by auto
- ultimately show "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space M2 \<times> space M1 \<in> sets (pair_algebra M2 M1)"
- by (auto intro: pair_algebraI)
+ ultimately show "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space M2 \<times> space M1 \<in> sets (pair_measure_generator M2 M1)"
+ by (auto intro: pair_measure_generatorI)
next
fix A B assume "A \<in> sets M1" "B \<in> sets M2"
then show "B \<times> A \<in> (\<lambda>X. (\<lambda>(x, y). (y, x)) -` X \<inter> space M2 \<times> space M1) ` sets E"
using M1.sets_into_space M2.sets_into_space
- by (auto intro!: image_eqI[where x="A \<times> B"] pair_algebraI)
+ by (auto intro!: image_eqI[where x="A \<times> B"] pair_measure_generatorI)
qed
lemma (in pair_sigma_algebra) sets_pair_sigma_algebra_swap:
assumes Q: "Q \<in> sets P"
- shows "(\<lambda>(x,y). (y, x)) ` Q \<in> sets (sigma (pair_algebra M2 M1))" (is "_ \<in> sets ?Q")
+ shows "(\<lambda>(x,y). (y, x)) -` Q \<in> sets (M2 \<Otimes>\<^isub>M M1)" (is "_ \<in> sets ?Q")
proof -
- have *: "(\<lambda>(x,y). (y, x)) \<in> space M2 \<times> space M1 \<rightarrow> (space M1 \<times> space M2)"
- "sets (pair_algebra M1 M2) \<subseteq> Pow (space M1 \<times> space M2)"
- using M1.sets_into_space M2.sets_into_space by (auto elim!: pair_algebraE)
- from Q sets_into_space show ?thesis
- by (auto intro!: image_eqI[where x=Q]
- simp: pair_algebra_swap[symmetric] sets_sigma
- sigma_sets_vimage[OF *] space_pair_algebra)
+ let "?f Q" = "(\<lambda>(x,y). (y, x)) -` Q \<inter> space M2 \<times> space M1"
+ have *: "(\<lambda>(x,y). (y, x)) -` Q = ?f Q"
+ using sets_into_space[OF Q] by (auto simp: space_pair_measure)
+ have "sets (M2 \<Otimes>\<^isub>M M1) = sets (sigma (pair_measure_generator M2 M1))"
+ unfolding pair_measure_def ..
+ also have "\<dots> = sigma_sets (space M2 \<times> space M1) (?f ` sets E)"
+ unfolding sigma_def pair_measure_generator_swap[symmetric]
+ by (simp add: pair_measure_generator_def)
+ also have "\<dots> = ?f ` sigma_sets (space M1 \<times> space M2) (sets E)"
+ using M1.sets_into_space M2.sets_into_space
+ by (intro sigma_sets_vimage) (auto simp: pair_measure_generator_def)
+ also have "\<dots> = ?f ` sets P"
+ unfolding pair_measure_def pair_measure_generator_def sigma_def by simp
+ finally show ?thesis
+ using Q by (subst *) auto
qed
lemma (in pair_sigma_algebra) pair_sigma_algebra_swap_measurable:
- shows "(\<lambda>(x,y). (y, x)) \<in> measurable P (sigma (pair_algebra M2 M1))"
+ shows "(\<lambda>(x,y). (y, x)) \<in> measurable P (M2 \<Otimes>\<^isub>M M1)"
(is "?f \<in> measurable ?P ?Q")
unfolding measurable_def
proof (intro CollectI conjI Pi_I ballI)
fix x assume "x \<in> space ?P" then show "(case x of (x, y) \<Rightarrow> (y, x)) \<in> space ?Q"
- unfolding pair_algebra_def by auto
+ unfolding pair_measure_generator_def pair_measure_def by auto
next
- fix A assume "A \<in> sets ?Q"
+ fix A assume "A \<in> sets (M2 \<Otimes>\<^isub>M M1)"
interpret Q: pair_sigma_algebra M2 M1 by default
- have "?f -` A \<inter> space ?P = (\<lambda>(x,y). (y, x)) ` A"
- using Q.sets_into_space `A \<in> sets ?Q` by (auto simp: pair_algebra_def)
- with Q.sets_pair_sigma_algebra_swap[OF `A \<in> sets ?Q`]
+ with Q.sets_pair_sigma_algebra_swap[OF `A \<in> sets (M2 \<Otimes>\<^isub>M M1)`]
show "?f -` A \<inter> space ?P \<in> sets ?P" by simp
qed
@@ -338,13 +430,13 @@
let ?Q' = "{Q. Q \<subseteq> space P \<and> Pair x -` Q \<in> sets M2}"
let ?Q = "\<lparr> space = space P, sets = ?Q' \<rparr>"
interpret Q: sigma_algebra ?Q
- proof qed (auto simp: vimage_UN vimage_Diff space_pair_algebra)
+ proof qed (auto simp: vimage_UN vimage_Diff space_pair_measure)
have "sets E \<subseteq> sets ?Q"
using M1.sets_into_space M2.sets_into_space
- by (auto simp: pair_algebra_def space_pair_algebra)
+ by (auto simp: pair_measure_generator_def space_pair_measure)
then have "sets P \<subseteq> sets ?Q"
- by (subst pair_algebra_def, intro Q.sets_sigma_subset)
- (simp_all add: pair_algebra_def)
+ apply (subst pair_measure_def, intro Q.sets_sigma_subset)
+ by (simp add: pair_measure_def)
with assms show ?thesis by auto
qed
@@ -352,9 +444,8 @@
assumes Q: "Q \<in> sets P" shows "(\<lambda>x. (x, y)) -` Q \<in> sets M1" (is "?cut Q \<in> sets M1")
proof -
interpret Q: pair_sigma_algebra M2 M1 by default
- have "Pair y -` (\<lambda>(x, y). (y, x)) ` Q = (\<lambda>x. (x, y)) -` Q" by auto
with Q.measurable_cut_fst[OF sets_pair_sigma_algebra_swap[OF Q], of y]
- show ?thesis by simp
+ show ?thesis by (simp add: vimage_compose[symmetric] comp_def)
qed
lemma (in pair_sigma_algebra) measurable_pair_image_snd:
@@ -363,14 +454,15 @@
unfolding measurable_def
proof (intro CollectI conjI Pi_I ballI)
fix y assume "y \<in> space M2" with `f \<in> measurable P M` `x \<in> space M1`
- show "f (x, y) \<in> space M" unfolding measurable_def pair_algebra_def by auto
+ show "f (x, y) \<in> space M"
+ unfolding measurable_def pair_measure_generator_def pair_measure_def by auto
next
fix A assume "A \<in> sets M"
then have "Pair x -` (f -` A \<inter> space P) \<in> sets M2" (is "?C \<in> _")
using `f \<in> measurable P M`
by (intro measurable_cut_fst) (auto simp: measurable_def)
also have "?C = (\<lambda>y. f (x, y)) -` A \<inter> space M2"
- using `x \<in> space M1` by (auto simp: pair_algebra_def)
+ using `x \<in> space M1` by (auto simp: pair_measure_generator_def pair_measure_def)
finally show "(\<lambda>y. f (x, y)) -` A \<inter> space M2 \<in> sets M2" .
qed
@@ -384,31 +476,31 @@
show ?thesis by simp
qed
-lemma (in pair_sigma_algebra) Int_stable_pair_algebra: "Int_stable E"
+lemma (in pair_sigma_algebra) Int_stable_pair_measure_generator: "Int_stable E"
unfolding Int_stable_def
proof (intro ballI)
fix A B assume "A \<in> sets E" "B \<in> sets E"
then obtain A1 A2 B1 B2 where "A = A1 \<times> A2" "B = B1 \<times> B2"
"A1 \<in> sets M1" "A2 \<in> sets M2" "B1 \<in> sets M1" "B2 \<in> sets M2"
- unfolding pair_algebra_def by auto
+ unfolding pair_measure_generator_def by auto
then show "A \<inter> B \<in> sets E"
- by (auto simp add: times_Int_times pair_algebra_def)
+ by (auto simp add: times_Int_times pair_measure_generator_def)
qed
lemma finite_measure_cut_measurable:
- fixes M1 :: "'a algebra" and M2 :: "'b algebra"
- assumes "sigma_finite_measure M1 \<mu>1" "finite_measure M2 \<mu>2"
- assumes "Q \<in> sets (sigma (pair_algebra M1 M2))"
- shows "(\<lambda>x. \<mu>2 (Pair x -` Q)) \<in> borel_measurable M1"
+ fixes M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
+ assumes "sigma_finite_measure M1" "finite_measure M2"
+ assumes "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)"
+ shows "(\<lambda>x. measure M2 (Pair x -` Q)) \<in> borel_measurable M1"
(is "?s Q \<in> _")
proof -
- interpret M1: sigma_finite_measure M1 \<mu>1 by fact
- interpret M2: finite_measure M2 \<mu>2 by fact
+ interpret M1: sigma_finite_measure M1 by fact
+ interpret M2: finite_measure M2 by fact
interpret pair_sigma_algebra M1 M2 by default
have [intro]: "sigma_algebra M1" by fact
have [intro]: "sigma_algebra M2" by fact
let ?D = "\<lparr> space = space P, sets = {A\<in>sets P. ?s A \<in> borel_measurable M1} \<rparr>"
- note space_pair_algebra[simp]
+ note space_pair_measure[simp]
interpret dynkin_system ?D
proof (intro dynkin_systemI)
fix A assume "A \<in> sets ?D" then show "A \<subseteq> space ?D"
@@ -418,73 +510,83 @@
by (auto simp add: if_distrib intro!: M1.measurable_If)
next
fix A assume "A \<in> sets ?D"
- with sets_into_space have "\<And>x. \<mu>2 (Pair x -` (space M1 \<times> space M2 - A)) =
- (if x \<in> space M1 then \<mu>2 (space M2) - ?s A x else 0)"
+ with sets_into_space have "\<And>x. measure M2 (Pair x -` (space M1 \<times> space M2 - A)) =
+ (if x \<in> space M1 then measure M2 (space M2) - ?s A x else 0)"
by (auto intro!: M2.finite_measure_compl measurable_cut_fst
simp: vimage_Diff)
with `A \<in> sets ?D` top show "space ?D - A \<in> sets ?D"
by (auto intro!: Diff M1.measurable_If M1.borel_measurable_pextreal_diff)
next
fix F :: "nat \<Rightarrow> ('a\<times>'b) set" assume "disjoint_family F" "range F \<subseteq> sets ?D"
- moreover then have "\<And>x. \<mu>2 (\<Union>i. Pair x -` F i) = (\<Sum>\<^isub>\<infinity> i. ?s (F i) x)"
+ moreover then have "\<And>x. measure M2 (\<Union>i. Pair x -` F i) = (\<Sum>\<^isub>\<infinity> i. ?s (F i) x)"
by (intro M2.measure_countably_additive[symmetric])
(auto intro!: measurable_cut_fst simp: disjoint_family_on_def)
ultimately show "(\<Union>i. F i) \<in> sets ?D"
by (auto simp: vimage_UN intro!: M1.borel_measurable_psuminf)
qed
- have "P = ?D"
+ have "sets P = sets ?D" apply (subst pair_measure_def)
proof (intro dynkin_lemma)
- show "Int_stable E" by (rule Int_stable_pair_algebra)
+ show "Int_stable E" by (rule Int_stable_pair_measure_generator)
from M1.sets_into_space have "\<And>A. A \<in> sets M1 \<Longrightarrow> {x \<in> space M1. x \<in> A} = A"
by auto
then show "sets E \<subseteq> sets ?D"
- by (auto simp: pair_algebra_def sets_sigma if_distrib
+ by (auto simp: pair_measure_generator_def sets_sigma if_distrib
intro: sigma_sets.Basic intro!: M1.measurable_If)
- qed auto
+ qed (auto simp: pair_measure_def)
with `Q \<in> sets P` have "Q \<in> sets ?D" by simp
then show "?s Q \<in> borel_measurable M1" by simp
qed
subsection {* Binary products of $\sigma$-finite measure spaces *}
-locale pair_sigma_finite = M1: sigma_finite_measure M1 \<mu>1 + M2: sigma_finite_measure M2 \<mu>2
- for M1 \<mu>1 M2 \<mu>2
+locale pair_sigma_finite = M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2
+ for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
sublocale pair_sigma_finite \<subseteq> pair_sigma_algebra M1 M2
by default
+lemma times_eq_iff: "A \<times> B = C \<times> D \<longleftrightarrow> A = C \<and> B = D \<or> ((A = {} \<or> B = {}) \<and> (C = {} \<or> D = {}))"
+ by auto
+
lemma (in pair_sigma_finite) measure_cut_measurable_fst:
- assumes "Q \<in> sets P" shows "(\<lambda>x. \<mu>2 (Pair x -` Q)) \<in> borel_measurable M1" (is "?s Q \<in> _")
+ assumes "Q \<in> sets P" shows "(\<lambda>x. measure M2 (Pair x -` Q)) \<in> borel_measurable M1" (is "?s Q \<in> _")
proof -
have [intro]: "sigma_algebra M1" and [intro]: "sigma_algebra M2" by default+
- have M1: "sigma_finite_measure M1 \<mu>1" by default
-
+ have M1: "sigma_finite_measure M1" by default
from M2.disjoint_sigma_finite guess F .. note F = this
+ then have "\<And>i. F i \<in> sets M2" by auto
let "?C x i" = "F i \<inter> Pair x -` Q"
{ fix i
let ?R = "M2.restricted_space (F i)"
have [simp]: "space M1 \<times> F i \<inter> space M1 \<times> space M2 = space M1 \<times> F i"
using F M2.sets_into_space by auto
- have "(\<lambda>x. \<mu>2 (Pair x -` (space M1 \<times> F i \<inter> Q))) \<in> borel_measurable M1"
+ let ?R2 = "M2.restricted_space (F i)"
+ have "(\<lambda>x. measure ?R2 (Pair x -` (space M1 \<times> space ?R2 \<inter> Q))) \<in> borel_measurable M1"
proof (intro finite_measure_cut_measurable[OF M1])
- show "finite_measure (M2.restricted_space (F i)) \<mu>2"
+ show "finite_measure ?R2"
using F by (intro M2.restricted_to_finite_measure) auto
- have "space M1 \<times> F i \<in> sets P"
- using M1.top F by blast
- from sigma_sets_Int[symmetric,
- OF this[unfolded pair_sigma_algebra_def sets_sigma]]
- show "(space M1 \<times> F i) \<inter> Q \<in> sets (sigma (pair_algebra M1 ?R))"
- using `Q \<in> sets P`
- using pair_algebra_Int_snd[OF M1.space_closed, of "F i" M2]
- by (auto simp: pair_algebra_def sets_sigma)
+ have "(space M1 \<times> space ?R2) \<inter> Q \<in> (op \<inter> (space M1 \<times> F i)) ` sets P"
+ using `Q \<in> sets P` by (auto simp: image_iff)
+ also have "\<dots> = sigma_sets (space M1 \<times> F i) ((op \<inter> (space M1 \<times> F i)) ` sets E)"
+ unfolding pair_measure_def pair_measure_generator_def sigma_def
+ using `F i \<in> sets M2` M2.sets_into_space
+ by (auto intro!: sigma_sets_Int sigma_sets.Basic)
+ also have "\<dots> \<subseteq> sets (M1 \<Otimes>\<^isub>M ?R2)"
+ using M1.sets_into_space
+ apply (auto simp: times_Int_times pair_measure_def pair_measure_generator_def sigma_def
+ intro!: sigma_sets_subseteq)
+ apply (rule_tac x="a" in exI)
+ apply (rule_tac x="b \<inter> F i" in exI)
+ by auto
+ finally show "(space M1 \<times> space ?R2) \<inter> Q \<in> sets (M1 \<Otimes>\<^isub>M ?R2)" .
qed
moreover have "\<And>x. Pair x -` (space M1 \<times> F i \<inter> Q) = ?C x i"
- using `Q \<in> sets P` sets_into_space by (auto simp: space_pair_algebra)
- ultimately have "(\<lambda>x. \<mu>2 (?C x i)) \<in> borel_measurable M1"
+ using `Q \<in> sets P` sets_into_space by (auto simp: space_pair_measure)
+ ultimately have "(\<lambda>x. measure M2 (?C x i)) \<in> borel_measurable M1"
by simp }
moreover
{ fix x
- have "(\<Sum>\<^isub>\<infinity>i. \<mu>2 (?C x i)) = \<mu>2 (\<Union>i. ?C x i)"
+ have "(\<Sum>\<^isub>\<infinity>i. measure M2 (?C x i)) = measure M2 (\<Union>i. ?C x i)"
proof (intro M2.measure_countably_additive)
show "range (?C x) \<subseteq> sets M2"
using F `Q \<in> sets P` by (auto intro!: M2.Int measurable_cut_fst)
@@ -494,27 +596,24 @@
qed
also have "(\<Union>i. ?C x i) = Pair x -` Q"
using F sets_into_space `Q \<in> sets P`
- by (auto simp: space_pair_algebra)
- finally have "\<mu>2 (Pair x -` Q) = (\<Sum>\<^isub>\<infinity>i. \<mu>2 (?C x i))"
+ by (auto simp: space_pair_measure)
+ finally have "measure M2 (Pair x -` Q) = (\<Sum>\<^isub>\<infinity>i. measure M2 (?C x i))"
by simp }
ultimately show ?thesis
by (auto intro!: M1.borel_measurable_psuminf)
qed
lemma (in pair_sigma_finite) measure_cut_measurable_snd:
- assumes "Q \<in> sets P" shows "(\<lambda>y. \<mu>1 ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"
+ assumes "Q \<in> sets P" shows "(\<lambda>y. M1.\<mu> ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"
proof -
- interpret Q: pair_sigma_finite M2 \<mu>2 M1 \<mu>1 by default
- have [simp]: "\<And>y. (Pair y -` (\<lambda>(x, y). (y, x)) ` Q) = (\<lambda>x. (x, y)) -` Q"
- by auto
+ interpret Q: pair_sigma_finite M2 M1 by default
note sets_pair_sigma_algebra_swap[OF assms]
from Q.measure_cut_measurable_fst[OF this]
- show ?thesis by simp
+ show ?thesis by (simp add: vimage_compose[symmetric] comp_def)
qed
lemma (in pair_sigma_algebra) pair_sigma_algebra_measurable:
- assumes "f \<in> measurable P M"
- shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (sigma (pair_algebra M2 M1)) M"
+ assumes "f \<in> measurable P M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^isub>M M1) M"
proof -
interpret Q: pair_sigma_algebra M2 M1 by default
have *: "(\<lambda>(x,y). f (y, x)) = f \<circ> (\<lambda>(x,y). (y, x))" by (simp add: fun_eq_iff)
@@ -523,19 +622,15 @@
unfolding * by (rule measurable_comp)
qed
-definition (in pair_sigma_finite)
- "pair_measure A = M1.positive_integral (\<lambda>x.
- M2.positive_integral (\<lambda>y. indicator A (x, y)))"
-
lemma (in pair_sigma_finite) pair_measure_alt:
assumes "A \<in> sets P"
- shows "pair_measure A = M1.positive_integral (\<lambda>x. \<mu>2 (Pair x -` A))"
- unfolding pair_measure_def
+ shows "measure (M1 \<Otimes>\<^isub>M M2) A = (\<integral>\<^isup>+ x. measure M2 (Pair x -` A) \<partial>M1)"
+ apply (simp add: pair_measure_def pair_measure_generator_def)
proof (rule M1.positive_integral_cong)
fix x assume "x \<in> space M1"
have *: "\<And>y. indicator A (x, y) = (indicator (Pair x -` A) y :: pextreal)"
unfolding indicator_def by auto
- show "M2.positive_integral (\<lambda>y. indicator A (x, y)) = \<mu>2 (Pair x -` A)"
+ show "(\<integral>\<^isup>+ y. indicator A (x, y) \<partial>M2) = measure M2 (Pair x -` A)"
unfolding *
apply (subst M2.positive_integral_indicator)
apply (rule measurable_cut_fst[OF assms])
@@ -544,30 +639,29 @@
lemma (in pair_sigma_finite) pair_measure_times:
assumes A: "A \<in> sets M1" and "B \<in> sets M2"
- shows "pair_measure (A \<times> B) = \<mu>1 A * \<mu>2 B"
+ shows "measure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = M1.\<mu> A * measure M2 B"
proof -
- from assms have "pair_measure (A \<times> B) =
- M1.positive_integral (\<lambda>x. \<mu>2 B * indicator A x)"
- by (auto intro!: M1.positive_integral_cong simp: pair_measure_alt)
+ have "measure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = (\<integral>\<^isup>+ x. measure M2 B * indicator A x \<partial>M1)"
+ using assms by (auto intro!: M1.positive_integral_cong simp: pair_measure_alt)
with assms show ?thesis
by (simp add: M1.positive_integral_cmult_indicator ac_simps)
qed
-lemma (in pair_sigma_finite) sigma_finite_up_in_pair_algebra:
+lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:
"\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets E \<and> F \<up> space E \<and>
- (\<forall>i. pair_measure (F i) \<noteq> \<omega>)"
+ (\<forall>i. measure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<omega>)"
proof -
obtain F1 :: "nat \<Rightarrow> 'a set" and F2 :: "nat \<Rightarrow> 'b set" where
- F1: "range F1 \<subseteq> sets M1" "F1 \<up> space M1" "\<And>i. \<mu>1 (F1 i) \<noteq> \<omega>" and
- F2: "range F2 \<subseteq> sets M2" "F2 \<up> space M2" "\<And>i. \<mu>2 (F2 i) \<noteq> \<omega>"
+ F1: "range F1 \<subseteq> sets M1" "F1 \<up> space M1" "\<And>i. M1.\<mu> (F1 i) \<noteq> \<omega>" and
+ F2: "range F2 \<subseteq> sets M2" "F2 \<up> space M2" "\<And>i. M2.\<mu> (F2 i) \<noteq> \<omega>"
using M1.sigma_finite_up M2.sigma_finite_up by auto
then have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)"
unfolding isoton_def by auto
let ?F = "\<lambda>i. F1 i \<times> F2 i"
- show ?thesis unfolding isoton_def space_pair_algebra
+ show ?thesis unfolding isoton_def space_pair_measure
proof (intro exI[of _ ?F] conjI allI)
show "range ?F \<subseteq> sets E" using F1 F2
- by (fastsimp intro!: pair_algebraI)
+ by (fastsimp intro!: pair_measure_generatorI)
next
have "space M1 \<times> space M2 \<subseteq> (\<Union>i. ?F i)"
proof (intro subsetI)
@@ -581,8 +675,8 @@
by (intro SigmaI) (auto simp add: min_max.sup_commute)
then show "x \<in> (\<Union>i. ?F i)" by auto
qed
- then show "(\<Union>i. ?F i) = space M1 \<times> space M2"
- using space by (auto simp: space)
+ then show "(\<Union>i. ?F i) = space E"
+ using space by (auto simp: space pair_measure_generator_def)
next
fix i show "F1 i \<times> F2 i \<subseteq> F1 (Suc i) \<times> F2 (Suc i)"
using `F1 \<up> space M1` `F2 \<up> space M2` unfolding isoton_def
@@ -590,103 +684,163 @@
next
fix i
from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto
- with F1 F2 show "pair_measure (F1 i \<times> F2 i) \<noteq> \<omega>"
+ with F1 F2 show "measure P (F1 i \<times> F2 i) \<noteq> \<omega>"
by (simp add: pair_measure_times)
qed
qed
-sublocale pair_sigma_finite \<subseteq> sigma_finite_measure P pair_measure
+sublocale pair_sigma_finite \<subseteq> sigma_finite_measure P
proof
- show "pair_measure {} = 0"
- unfolding pair_measure_def by auto
+ show "measure P {} = 0"
+ unfolding pair_measure_def pair_measure_generator_def sigma_def by auto
- show "countably_additive P pair_measure"
+ show "countably_additive P (measure P)"
unfolding countably_additive_def
proof (intro allI impI)
fix F :: "nat \<Rightarrow> ('a \<times> 'b) set"
assume F: "range F \<subseteq> sets P" "disjoint_family F"
from F have *: "\<And>i. F i \<in> sets P" "(\<Union>i. F i) \<in> sets P" by auto
- moreover from F have "\<And>i. (\<lambda>x. \<mu>2 (Pair x -` F i)) \<in> borel_measurable M1"
+ moreover from F have "\<And>i. (\<lambda>x. measure M2 (Pair x -` F i)) \<in> borel_measurable M1"
by (intro measure_cut_measurable_fst) auto
moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)"
by (intro disjoint_family_on_bisimulation[OF F(2)]) auto
moreover have "\<And>x. x \<in> space M1 \<Longrightarrow> range (\<lambda>i. Pair x -` F i) \<subseteq> sets M2"
using F by (auto intro!: measurable_cut_fst)
- ultimately show "(\<Sum>\<^isub>\<infinity>n. pair_measure (F n)) = pair_measure (\<Union>i. F i)"
+ ultimately show "(\<Sum>\<^isub>\<infinity>n. measure P (F n)) = measure P (\<Union>i. F i)"
by (simp add: pair_measure_alt vimage_UN M1.positive_integral_psuminf[symmetric]
M2.measure_countably_additive
cong: M1.positive_integral_cong)
qed
- from sigma_finite_up_in_pair_algebra guess F :: "nat \<Rightarrow> ('a \<times> 'c) set" .. note F = this
- show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets P \<and> (\<Union>i. F i) = space P \<and> (\<forall>i. pair_measure (F i) \<noteq> \<omega>)"
+ from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
+ show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets P \<and> (\<Union>i. F i) = space P \<and> (\<forall>i. measure P (F i) \<noteq> \<omega>)"
proof (rule exI[of _ F], intro conjI)
- show "range F \<subseteq> sets P" using F by auto
+ show "range F \<subseteq> sets P" using F by (auto simp: pair_measure_def)
show "(\<Union>i. F i) = space P"
- using F by (auto simp: space_pair_algebra isoton_def)
- show "\<forall>i. pair_measure (F i) \<noteq> \<omega>" using F by auto
+ using F by (auto simp: pair_measure_def pair_measure_generator_def isoton_def)
+ show "\<forall>i. measure P (F i) \<noteq> \<omega>" using F by auto
qed
qed
lemma (in pair_sigma_algebra) sets_swap:
assumes "A \<in> sets P"
- shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (sigma (pair_algebra M2 M1)) \<in> sets (sigma (pair_algebra M2 M1))"
+ shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1) \<in> sets (M2 \<Otimes>\<^isub>M M1)"
(is "_ -` A \<inter> space ?Q \<in> sets ?Q")
proof -
- have *: "(\<lambda>(x, y). (y, x)) -` A \<inter> space ?Q = (\<lambda>(x, y). (y, x)) ` A"
- using `A \<in> sets P` sets_into_space by (auto simp: space_pair_algebra)
+ have *: "(\<lambda>(x, y). (y, x)) -` A \<inter> space ?Q = (\<lambda>(x, y). (y, x)) -` A"
+ using `A \<in> sets P` sets_into_space by (auto simp: space_pair_measure)
show ?thesis
unfolding * using assms by (rule sets_pair_sigma_algebra_swap)
qed
lemma (in pair_sigma_finite) pair_measure_alt2:
assumes "A \<in> sets P"
- shows "pair_measure A = M2.positive_integral (\<lambda>y. \<mu>1 ((\<lambda>x. (x, y)) -` A))"
+ shows "\<mu> A = (\<integral>\<^isup>+y. M1.\<mu> ((\<lambda>x. (x, y)) -` A) \<partial>M2)"
(is "_ = ?\<nu> A")
proof -
- from sigma_finite_up_in_pair_algebra guess F :: "nat \<Rightarrow> ('a \<times> 'c) set" .. note F = this
+ from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
+ have [simp]: "\<And>m. \<lparr> space = space E, sets = sets (sigma E), measure = m \<rparr> = P\<lparr> measure := m \<rparr>"
+ unfolding pair_measure_def by simp
show ?thesis
- proof (rule measure_unique_Int_stable[where \<nu>="?\<nu>", OF Int_stable_pair_algebra],
- simp_all add: pair_sigma_algebra_def[symmetric])
- show "range F \<subseteq> sets E" "F \<up> space E" "\<And>i. pair_measure (F i) \<noteq> \<omega>"
- using F by auto
- show "measure_space P pair_measure" by default
- interpret Q: pair_sigma_finite M2 \<mu>2 M1 \<mu>1 by default
- have P: "sigma_algebra P" by default
- show "measure_space P ?\<nu>"
+ proof (rule measure_unique_Int_stable[where \<nu>="?\<nu>", OF Int_stable_pair_measure_generator], simp_all)
+ show "range F \<subseteq> sets E" "F \<up> space E" "\<And>i. \<mu> (F i) \<noteq> \<omega>" "A \<in> sets (sigma E)"
+ using F `A \<in> sets P` by (auto simp: pair_measure_def)
+ show "measure_space P" by default
+ interpret Q: pair_sigma_finite M2 M1 by default
+ have P: "sigma_algebra (P\<lparr> measure := ?\<nu>\<rparr>)"
+ by (intro sigma_algebra_cong) auto
+ show "measure_space (P\<lparr> measure := ?\<nu>\<rparr>)"
apply (rule Q.measure_space_vimage[OF P])
+ apply (simp_all)
apply (rule Q.pair_sigma_algebra_swap_measurable)
proof -
fix A assume "A \<in> sets P"
from sets_swap[OF this]
- show "M2.positive_integral (\<lambda>y. \<mu>1 ((\<lambda>x. (x, y)) -` A)) =
- Q.pair_measure ((\<lambda>(x, y). (y, x)) -` A \<inter> space Q.P)"
+ show "(\<integral>\<^isup>+ y. M1.\<mu> ((\<lambda>x. (x, y)) -` A) \<partial>M2) = Q.\<mu> ((\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1))"
using sets_into_space[OF `A \<in> sets P`]
- by (auto simp add: Q.pair_measure_alt space_pair_algebra
- intro!: M2.positive_integral_cong arg_cong[where f=\<mu>1])
+ by (auto simp add: Q.pair_measure_alt space_pair_measure
+ intro!: M2.positive_integral_cong arg_cong[where f="M1.\<mu>"])
qed
fix X assume "X \<in> sets E"
then obtain A B where X: "X = A \<times> B" and AB: "A \<in> sets M1" "B \<in> sets M2"
- unfolding pair_algebra_def by auto
- show "pair_measure X = ?\<nu> X"
+ unfolding pair_measure_def pair_measure_generator_def by auto
+ show "\<mu> X = ?\<nu> X"
proof -
- from AB have "?\<nu> (A \<times> B) =
- M2.positive_integral (\<lambda>y. \<mu>1 A * indicator B y)"
+ from AB have "?\<nu> (A \<times> B) = (\<integral>\<^isup>+y. M1.\<mu> A * indicator B y \<partial>M2)"
by (auto intro!: M2.positive_integral_cong)
with AB show ?thesis
unfolding pair_measure_times[OF AB] X
by (simp add: M2.positive_integral_cmult_indicator ac_simps)
qed
- qed fact
+ qed
+qed
+
+
+lemma pair_sigma_algebra_sigma:
+ assumes 1: "S1 \<up> (space E1)" "range S1 \<subseteq> sets E1" and E1: "sets E1 \<subseteq> Pow (space E1)"
+ assumes 2: "S2 \<up> (space E2)" "range S2 \<subseteq> sets E2" and E2: "sets E2 \<subseteq> Pow (space E2)"
+ shows "sets (sigma (pair_measure_generator (sigma E1) (sigma E2))) = sets (sigma (pair_measure_generator E1 E2))"
+ (is "sets ?S = sets ?E")
+proof -
+ interpret M1: sigma_algebra "sigma E1" using E1 by (rule sigma_algebra_sigma)
+ interpret M2: sigma_algebra "sigma E2" using E2 by (rule sigma_algebra_sigma)
+ have P: "sets (pair_measure_generator E1 E2) \<subseteq> Pow (space E1 \<times> space E2)"
+ using E1 E2 by (auto simp add: pair_measure_generator_def)
+ interpret E: sigma_algebra ?E unfolding pair_measure_generator_def
+ using E1 E2 by (intro sigma_algebra_sigma) auto
+ { fix A assume "A \<in> sets E1"
+ then have "fst -` A \<inter> space ?E = A \<times> (\<Union>i. S2 i)"
+ using E1 2 unfolding isoton_def pair_measure_generator_def by auto
+ also have "\<dots> = (\<Union>i. A \<times> S2 i)" by auto
+ also have "\<dots> \<in> sets ?E" unfolding pair_measure_generator_def sets_sigma
+ using 2 `A \<in> sets E1`
+ by (intro sigma_sets.Union)
+ (auto simp: image_subset_iff intro!: sigma_sets.Basic)
+ finally have "fst -` A \<inter> space ?E \<in> sets ?E" . }
+ moreover
+ { fix B assume "B \<in> sets E2"
+ then have "snd -` B \<inter> space ?E = (\<Union>i. S1 i) \<times> B"
+ using E2 1 unfolding isoton_def pair_measure_generator_def by auto
+ also have "\<dots> = (\<Union>i. S1 i \<times> B)" by auto
+ also have "\<dots> \<in> sets ?E"
+ using 1 `B \<in> sets E2` unfolding pair_measure_generator_def sets_sigma
+ by (intro sigma_sets.Union)
+ (auto simp: image_subset_iff intro!: sigma_sets.Basic)
+ finally have "snd -` B \<inter> space ?E \<in> sets ?E" . }
+ ultimately have proj:
+ "fst \<in> measurable ?E (sigma E1) \<and> snd \<in> measurable ?E (sigma E2)"
+ using E1 E2 by (subst (1 2) E.measurable_iff_sigma)
+ (auto simp: pair_measure_generator_def sets_sigma)
+ { fix A B assume A: "A \<in> sets (sigma E1)" and B: "B \<in> sets (sigma E2)"
+ with proj have "fst -` A \<inter> space ?E \<in> sets ?E" "snd -` B \<inter> space ?E \<in> sets ?E"
+ unfolding measurable_def by simp_all
+ moreover have "A \<times> B = (fst -` A \<inter> space ?E) \<inter> (snd -` B \<inter> space ?E)"
+ using A B M1.sets_into_space M2.sets_into_space
+ by (auto simp: pair_measure_generator_def)
+ ultimately have "A \<times> B \<in> sets ?E" by auto }
+ then have "sigma_sets (space ?E) (sets (pair_measure_generator (sigma E1) (sigma E2))) \<subseteq> sets ?E"
+ by (intro E.sigma_sets_subset) (auto simp add: pair_measure_generator_def sets_sigma)
+ then have subset: "sets ?S \<subseteq> sets ?E"
+ by (simp add: sets_sigma pair_measure_generator_def)
+ show "sets ?S = sets ?E"
+ proof (intro set_eqI iffI)
+ fix A assume "A \<in> sets ?E" then show "A \<in> sets ?S"
+ unfolding sets_sigma
+ proof induct
+ case (Basic A) then show ?case
+ by (auto simp: pair_measure_generator_def sets_sigma intro: sigma_sets.Basic)
+ qed (auto intro: sigma_sets.intros simp: pair_measure_generator_def)
+ next
+ fix A assume "A \<in> sets ?S" then show "A \<in> sets ?E" using subset by auto
+ qed
qed
section "Fubinis theorem"
lemma (in pair_sigma_finite) simple_function_cut:
- assumes f: "simple_function f"
- shows "(\<lambda>x. M2.positive_integral (\<lambda> y. f (x, y))) \<in> borel_measurable M1"
- and "M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>y. f (x, y)))
- = positive_integral f"
+ assumes f: "simple_function P f"
+ shows "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
+ and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
proof -
have f_borel: "f \<in> borel_measurable P"
using f by (rule borel_measurable_simple_function)
@@ -696,42 +850,40 @@
have [simp]: "\<And>z y. indicator (?F z) (x, y) = indicator (?F' x z) y"
by (auto simp: indicator_def)
have "\<And>y. y \<in> space M2 \<Longrightarrow> (x, y) \<in> space P" using `x \<in> space M1`
- by (simp add: space_pair_algebra)
+ by (simp add: space_pair_measure)
moreover have "\<And>x z. ?F' x z \<in> sets M2" using f_borel
by (intro borel_measurable_vimage measurable_cut_fst)
- ultimately have "M2.simple_function (\<lambda> y. f (x, y))"
+ ultimately have "simple_function M2 (\<lambda> y. f (x, y))"
apply (rule_tac M2.simple_function_cong[THEN iffD2, OF _])
apply (rule simple_function_indicator_representation[OF f])
using `x \<in> space M1` by (auto simp del: space_sigma) }
note M2_sf = this
{ fix x assume x: "x \<in> space M1"
- then have "M2.positive_integral (\<lambda> y. f (x, y)) =
- (\<Sum>z\<in>f ` space P. z * \<mu>2 (?F' x z))"
+ then have "(\<integral>\<^isup>+y. f (x, y) \<partial>M2) = (\<Sum>z\<in>f ` space P. z * M2.\<mu> (?F' x z))"
unfolding M2.positive_integral_eq_simple_integral[OF M2_sf[OF x]]
- unfolding M2.simple_integral_def
+ unfolding simple_integral_def
proof (safe intro!: setsum_mono_zero_cong_left)
from f show "finite (f ` space P)" by (rule simple_functionD)
next
fix y assume "y \<in> space M2" then show "f (x, y) \<in> f ` space P"
- using `x \<in> space M1` by (auto simp: space_pair_algebra)
+ using `x \<in> space M1` by (auto simp: space_pair_measure)
next
fix x' y assume "(x', y) \<in> space P"
"f (x', y) \<notin> (\<lambda>y. f (x, y)) ` space M2"
then have *: "?F' x (f (x', y)) = {}"
- by (force simp: space_pair_algebra)
- show "f (x', y) * \<mu>2 (?F' x (f (x', y))) = 0"
+ by (force simp: space_pair_measure)
+ show "f (x', y) * M2.\<mu> (?F' x (f (x', y))) = 0"
unfolding * by simp
qed (simp add: vimage_compose[symmetric] comp_def
- space_pair_algebra) }
+ space_pair_measure) }
note eq = this
moreover have "\<And>z. ?F z \<in> sets P"
by (auto intro!: f_borel borel_measurable_vimage simp del: space_sigma)
- moreover then have "\<And>z. (\<lambda>x. \<mu>2 (?F' x z)) \<in> borel_measurable M1"
+ moreover then have "\<And>z. (\<lambda>x. M2.\<mu> (?F' x z)) \<in> borel_measurable M1"
by (auto intro!: measure_cut_measurable_fst simp del: vimage_Int)
ultimately
- show "(\<lambda> x. M2.positive_integral (\<lambda> y. f (x, y))) \<in> borel_measurable M1"
- and "M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>y. f (x, y)))
- = positive_integral f"
+ show "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
+ and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
by (auto simp del: vimage_Int cong: measurable_cong
intro!: M1.borel_measurable_pextreal_setsum
simp add: M1.positive_integral_setsum simple_integral_def
@@ -743,13 +895,12 @@
lemma (in pair_sigma_finite) positive_integral_fst_measurable:
assumes f: "f \<in> borel_measurable P"
- shows "(\<lambda> x. M2.positive_integral (\<lambda> y. f (x, y))) \<in> borel_measurable M1"
+ shows "(\<lambda>x. \<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
(is "?C f \<in> borel_measurable M1")
- and "M1.positive_integral (\<lambda> x. M2.positive_integral (\<lambda> y. f (x, y))) =
- positive_integral f"
+ and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
proof -
from borel_measurable_implies_simple_function_sequence[OF f]
- obtain F where F: "\<And>i. simple_function (F i)" "F \<up> f" by auto
+ obtain F where F: "\<And>i. simple_function P (F i)" "F \<up> f" by auto
then have F_borel: "\<And>i. F i \<in> borel_measurable P"
and F_mono: "\<And>i x. F i x \<le> F (Suc i) x"
and F_SUPR: "\<And>x. (SUP i. F i x) = f x"
@@ -770,56 +921,48 @@
note SUPR_C = this
ultimately show "?C f \<in> borel_measurable M1"
by (simp cong: measurable_cong)
- have "positive_integral (\<lambda>x. (SUP i. F i x)) = (SUP i. positive_integral (F i))"
+ have "(\<integral>\<^isup>+x. (SUP i. F i x) \<partial>P) = (SUP i. integral\<^isup>P P (F i))"
using F_borel F_mono
by (auto intro!: positive_integral_monotone_convergence_SUP[symmetric])
- also have "(SUP i. positive_integral (F i)) =
- (SUP i. M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>y. F i (x, y))))"
+ also have "(SUP i. integral\<^isup>P P (F i)) = (SUP i. \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1)"
unfolding sf(2) by simp
- also have "\<dots> = M1.positive_integral (\<lambda>x. SUP i. M2.positive_integral (\<lambda>y. F i (x, y)))"
+ also have "\<dots> = \<integral>\<^isup>+ x. (SUP i. \<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1"
by (auto intro!: M1.positive_integral_monotone_convergence_SUP[OF _ sf(1)]
M2.positive_integral_mono F_mono)
- also have "\<dots> = M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>y. SUP i. F i (x, y)))"
+ also have "\<dots> = \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. (SUP i. F i (x, y)) \<partial>M2) \<partial>M1"
using F_borel F_mono
by (auto intro!: M2.positive_integral_monotone_convergence_SUP
M1.positive_integral_cong measurable_pair_image_snd)
- finally show "M1.positive_integral (\<lambda> x. M2.positive_integral (\<lambda> y. f (x, y))) =
- positive_integral f"
+ finally show "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
unfolding F_SUPR by simp
qed
lemma (in pair_sigma_finite) positive_integral_product_swap:
assumes f: "f \<in> borel_measurable P"
- shows "measure_space.positive_integral
- (sigma (pair_algebra M2 M1)) (pair_sigma_finite.pair_measure M2 \<mu>2 M1 \<mu>1) (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x))) =
- positive_integral f"
+ shows "(\<integral>\<^isup>+x. f (case x of (x,y)\<Rightarrow>(y,x)) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>P P f"
proof -
- interpret Q: pair_sigma_finite M2 \<mu>2 M1 \<mu>1 by default
- have P: "sigma_algebra P" by default
+ interpret Q: pair_sigma_finite M2 M1 by default
+ have "sigma_algebra P" by default
show ?thesis
- unfolding Q.positive_integral_vimage[OF P Q.pair_sigma_algebra_swap_measurable f, symmetric]
- proof (rule positive_integral_cong_measure)
- fix A
- assume A: "A \<in> sets P"
+ proof (intro Q.positive_integral_vimage[symmetric] Q.pair_sigma_algebra_swap_measurable)
+ fix A assume "A \<in> sets P"
from Q.pair_sigma_algebra_swap_measurable[THEN measurable_sets, OF this] this sets_into_space[OF this]
- show "Q.pair_measure ((\<lambda>(x, y). (y, x)) -` A \<inter> space Q.P) = pair_measure A"
- by (auto intro!: M1.positive_integral_cong arg_cong[where f=\<mu>2]
- simp: pair_measure_alt Q.pair_measure_alt2 space_pair_algebra)
- qed
+ show "\<mu> A = Q.\<mu> ((\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1))"
+ by (auto intro!: M1.positive_integral_cong arg_cong[where f="M2.\<mu>"]
+ simp: pair_measure_alt Q.pair_measure_alt2 space_pair_measure)
+ qed fact+
qed
lemma (in pair_sigma_finite) positive_integral_snd_measurable:
assumes f: "f \<in> borel_measurable P"
- shows "M2.positive_integral (\<lambda>y. M1.positive_integral (\<lambda>x. f (x, y))) =
- positive_integral f"
+ shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>P P f"
proof -
- interpret Q: pair_sigma_finite M2 \<mu>2 M1 \<mu>1 by default
+ interpret Q: pair_sigma_finite M2 M1 by default
note pair_sigma_algebra_measurable[OF f]
from Q.positive_integral_fst_measurable[OF this]
- have "M2.positive_integral (\<lambda>y. M1.positive_integral (\<lambda>x. f (x, y))) =
- Q.positive_integral (\<lambda>(x, y). f (y, x))"
+ have "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ (x, y). f (y, x) \<partial>Q.P)"
by simp
- also have "Q.positive_integral (\<lambda>(x, y). f (y, x)) = positive_integral f"
+ also have "(\<integral>\<^isup>+ (x, y). f (y, x) \<partial>Q.P) = integral\<^isup>P P f"
unfolding positive_integral_product_swap[OF f, symmetric]
by (auto intro!: Q.positive_integral_cong)
finally show ?thesis .
@@ -827,8 +970,7 @@
lemma (in pair_sigma_finite) Fubini:
assumes f: "f \<in> borel_measurable P"
- shows "M2.positive_integral (\<lambda>y. M1.positive_integral (\<lambda>x. f (x, y))) =
- M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>y. f (x, y)))"
+ shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1)"
unfolding positive_integral_snd_measurable[OF assms]
unfolding positive_integral_fst_measurable[OF assms] ..
@@ -836,30 +978,30 @@
assumes "almost_everywhere (\<lambda>x. Q x)"
shows "M1.almost_everywhere (\<lambda>x. M2.almost_everywhere (\<lambda>y. Q (x, y)))"
proof -
- obtain N where N: "N \<in> sets P" "pair_measure N = 0" "{x\<in>space P. \<not> Q x} \<subseteq> N"
+ obtain N where N: "N \<in> sets P" "\<mu> N = 0" "{x\<in>space P. \<not> Q x} \<subseteq> N"
using assms unfolding almost_everywhere_def by auto
show ?thesis
proof (rule M1.AE_I)
from N measure_cut_measurable_fst[OF `N \<in> sets P`]
- show "\<mu>1 {x\<in>space M1. \<mu>2 (Pair x -` N) \<noteq> 0} = 0"
+ show "M1.\<mu> {x\<in>space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} = 0"
by (simp add: M1.positive_integral_0_iff pair_measure_alt vimage_def)
- show "{x \<in> space M1. \<mu>2 (Pair x -` N) \<noteq> 0} \<in> sets M1"
+ show "{x \<in> space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} \<in> sets M1"
by (intro M1.borel_measurable_pextreal_neq_const measure_cut_measurable_fst N)
- { fix x assume "x \<in> space M1" "\<mu>2 (Pair x -` N) = 0"
+ { fix x assume "x \<in> space M1" "M2.\<mu> (Pair x -` N) = 0"
have "M2.almost_everywhere (\<lambda>y. Q (x, y))"
proof (rule M2.AE_I)
- show "\<mu>2 (Pair x -` N) = 0" by fact
+ show "M2.\<mu> (Pair x -` N) = 0" by fact
show "Pair x -` N \<in> sets M2" by (intro measurable_cut_fst N)
show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N"
- using N `x \<in> space M1` unfolding space_sigma space_pair_algebra by auto
+ using N `x \<in> space M1` unfolding space_sigma space_pair_measure by auto
qed }
- then show "{x \<in> space M1. \<not> M2.almost_everywhere (\<lambda>y. Q (x, y))} \<subseteq> {x \<in> space M1. \<mu>2 (Pair x -` N) \<noteq> 0}"
+ then show "{x \<in> space M1. \<not> M2.almost_everywhere (\<lambda>y. Q (x, y))} \<subseteq> {x \<in> space M1. M2.\<mu> (Pair x -` N) \<noteq> 0}"
by auto
qed
qed
lemma (in pair_sigma_algebra) measurable_product_swap:
- "f \<in> measurable (sigma (pair_algebra M2 M1)) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable P M"
+ "f \<in> measurable (M2 \<Otimes>\<^isub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable P M"
proof -
interpret Q: pair_sigma_algebra M2 M1 by default
show ?thesis
@@ -868,114 +1010,106 @@
qed
lemma (in pair_sigma_finite) integrable_product_swap:
- assumes "integrable f"
- shows "measure_space.integrable
- (sigma (pair_algebra M2 M1)) (pair_sigma_finite.pair_measure M2 \<mu>2 M1 \<mu>1) (\<lambda>(x,y). f (y,x))"
+ assumes "integrable P f"
+ shows "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x))"
proof -
- interpret Q: pair_sigma_finite M2 \<mu>2 M1 \<mu>1 by default
+ interpret Q: pair_sigma_finite M2 M1 by default
have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
show ?thesis unfolding *
- using assms unfolding Q.integrable_def integrable_def
+ using assms unfolding integrable_def
apply (subst (1 2) positive_integral_product_swap)
- using `integrable f` unfolding integrable_def
+ using `integrable P f` unfolding integrable_def
by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric])
qed
lemma (in pair_sigma_finite) integrable_product_swap_iff:
- "measure_space.integrable
- (sigma (pair_algebra M2 M1)) (pair_sigma_finite.pair_measure M2 \<mu>2 M1 \<mu>1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow>
- integrable f"
+ "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable P f"
proof -
- interpret Q: pair_sigma_finite M2 \<mu>2 M1 \<mu>1 by default
+ interpret Q: pair_sigma_finite M2 M1 by default
from Q.integrable_product_swap[of "\<lambda>(x,y). f (y,x)"] integrable_product_swap[of f]
show ?thesis by auto
qed
lemma (in pair_sigma_finite) integral_product_swap:
- assumes "integrable f"
- shows "measure_space.integral
- (sigma (pair_algebra M2 M1)) (pair_sigma_finite.pair_measure M2 \<mu>2 M1 \<mu>1) (\<lambda>(x,y). f (y,x)) =
- integral f"
+ assumes "integrable P f"
+ shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>L P f"
proof -
- interpret Q: pair_sigma_finite M2 \<mu>2 M1 \<mu>1 by default
+ interpret Q: pair_sigma_finite M2 M1 by default
have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
show ?thesis
- unfolding integral_def Q.integral_def *
+ unfolding lebesgue_integral_def *
apply (subst (1 2) positive_integral_product_swap)
- using `integrable f` unfolding integrable_def
+ using `integrable P f` unfolding integrable_def
by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric])
qed
lemma (in pair_sigma_finite) integrable_fst_measurable:
- assumes f: "integrable f"
- shows "M1.almost_everywhere (\<lambda>x. M2.integrable (\<lambda> y. f (x, y)))" (is "?AE")
- and "M1.integral (\<lambda> x. M2.integral (\<lambda> y. f (x, y))) = integral f" (is "?INT")
+ assumes f: "integrable P f"
+ shows "M1.almost_everywhere (\<lambda>x. integrable M2 (\<lambda> y. f (x, y)))" (is "?AE")
+ and "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>L P f" (is "?INT")
proof -
let "?pf x" = "Real (f x)" and "?nf x" = "Real (- f x)"
have
borel: "?nf \<in> borel_measurable P""?pf \<in> borel_measurable P" and
- int: "positive_integral ?nf \<noteq> \<omega>" "positive_integral ?pf \<noteq> \<omega>"
+ int: "integral\<^isup>P P ?nf \<noteq> \<omega>" "integral\<^isup>P P ?pf \<noteq> \<omega>"
using assms by auto
- have "M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>y. Real (f (x, y)))) \<noteq> \<omega>"
- "M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>y. Real (- f (x, y)))) \<noteq> \<omega>"
+ have "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. Real (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<omega>"
+ "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. Real (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<omega>"
using borel[THEN positive_integral_fst_measurable(1)] int
unfolding borel[THEN positive_integral_fst_measurable(2)] by simp_all
with borel[THEN positive_integral_fst_measurable(1)]
- have AE: "M1.almost_everywhere (\<lambda>x. M2.positive_integral (\<lambda>y. Real (f (x, y))) \<noteq> \<omega>)"
- "M1.almost_everywhere (\<lambda>x. M2.positive_integral (\<lambda>y. Real (- f (x, y))) \<noteq> \<omega>)"
+ have AE: "M1.almost_everywhere (\<lambda>x. (\<integral>\<^isup>+y. Real (f (x, y)) \<partial>M2) \<noteq> \<omega>)"
+ "M1.almost_everywhere (\<lambda>x. (\<integral>\<^isup>+y. Real (- f (x, y)) \<partial>M2) \<noteq> \<omega>)"
by (auto intro!: M1.positive_integral_omega_AE)
then show ?AE
apply (rule M1.AE_mp[OF _ M1.AE_mp])
apply (rule M1.AE_cong)
- using assms unfolding M2.integrable_def
+ using assms unfolding integrable_def
by (auto intro!: measurable_pair_image_snd)
- have "M1.integrable
- (\<lambda>x. real (M2.positive_integral (\<lambda>xa. Real (f (x, xa)))))" (is "M1.integrable ?f")
- proof (unfold M1.integrable_def, intro conjI)
+ have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+y. Real (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")
+ proof (intro integrable_def[THEN iffD2] conjI)
show "?f \<in> borel_measurable M1"
using borel by (auto intro!: M1.borel_measurable_real positive_integral_fst_measurable)
- have "M1.positive_integral (\<lambda>x. Real (?f x)) =
- M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>xa. Real (f (x, xa))))"
+ have "(\<integral>\<^isup>+x. Real (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. Real (f (x, y)) \<partial>M2) \<partial>M1)"
apply (rule M1.positive_integral_cong_AE)
apply (rule M1.AE_mp[OF AE(1)])
apply (rule M1.AE_cong)
by (auto simp: Real_real)
- then show "M1.positive_integral (\<lambda>x. Real (?f x)) \<noteq> \<omega>"
+ then show "(\<integral>\<^isup>+x. Real (?f x) \<partial>M1) \<noteq> \<omega>"
using positive_integral_fst_measurable[OF borel(2)] int by simp
- have "M1.positive_integral (\<lambda>x. Real (- ?f x)) = M1.positive_integral (\<lambda>x. 0)"
+ have "(\<integral>\<^isup>+x. Real (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
by (intro M1.positive_integral_cong) simp
- then show "M1.positive_integral (\<lambda>x. Real (- ?f x)) \<noteq> \<omega>" by simp
+ then show "(\<integral>\<^isup>+x. Real (- ?f x) \<partial>M1) \<noteq> \<omega>" by simp
qed
- moreover have "M1.integrable
- (\<lambda>x. real (M2.positive_integral (\<lambda>xa. Real (- f (x, xa)))))" (is "M1.integrable ?f")
- proof (unfold M1.integrable_def, intro conjI)
+ moreover have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+ y. Real (- f (x, y)) \<partial>M2))"
+ (is "integrable M1 ?f")
+ proof (intro integrable_def[THEN iffD2] conjI)
show "?f \<in> borel_measurable M1"
using borel by (auto intro!: M1.borel_measurable_real positive_integral_fst_measurable)
- have "M1.positive_integral (\<lambda>x. Real (?f x)) =
- M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>xa. Real (- f (x, xa))))"
+ have "(\<integral>\<^isup>+x. Real (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. Real (- f (x, y)) \<partial>M2) \<partial>M1)"
apply (rule M1.positive_integral_cong_AE)
apply (rule M1.AE_mp[OF AE(2)])
apply (rule M1.AE_cong)
by (auto simp: Real_real)
- then show "M1.positive_integral (\<lambda>x. Real (?f x)) \<noteq> \<omega>"
+ then show "(\<integral>\<^isup>+x. Real (?f x) \<partial>M1) \<noteq> \<omega>"
using positive_integral_fst_measurable[OF borel(1)] int by simp
- have "M1.positive_integral (\<lambda>x. Real (- ?f x)) = M1.positive_integral (\<lambda>x. 0)"
+ have "(\<integral>\<^isup>+x. Real (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
by (intro M1.positive_integral_cong) simp
- then show "M1.positive_integral (\<lambda>x. Real (- ?f x)) \<noteq> \<omega>" by simp
+ then show "(\<integral>\<^isup>+x. Real (- ?f x) \<partial>M1) \<noteq> \<omega>" by simp
qed
ultimately show ?INT
- unfolding M2.integral_def integral_def
+ unfolding lebesgue_integral_def[of P] lebesgue_integral_def[of M2]
borel[THEN positive_integral_fst_measurable(2), symmetric]
by (simp add: M1.integral_real[OF AE(1)] M1.integral_real[OF AE(2)])
qed
lemma (in pair_sigma_finite) integrable_snd_measurable:
- assumes f: "integrable f"
- shows "M2.almost_everywhere (\<lambda>y. M1.integrable (\<lambda>x. f (x, y)))" (is "?AE")
- and "M2.integral (\<lambda>y. M1.integral (\<lambda>x. f (x, y))) = integral f" (is "?INT")
+ assumes f: "integrable P f"
+ shows "M2.almost_everywhere (\<lambda>y. integrable M1 (\<lambda>x. f (x, y)))" (is "?AE")
+ and "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>L P f" (is "?INT")
proof -
- interpret Q: pair_sigma_finite M2 \<mu>2 M1 \<mu>1 by default
- have Q_int: "Q.integrable (\<lambda>(x, y). f (y, x))"
+ interpret Q: pair_sigma_finite M2 M1 by default
+ have Q_int: "integrable Q.P (\<lambda>(x, y). f (y, x))"
using f unfolding integrable_product_swap_iff .
show ?INT
using Q.integrable_fst_measurable(2)[OF Q_int]
@@ -986,9 +1120,8 @@
qed
lemma (in pair_sigma_finite) Fubini_integral:
- assumes f: "integrable f"
- shows "M2.integral (\<lambda>y. M1.integral (\<lambda>x. f (x, y))) =
- M1.integral (\<lambda>x. M2.integral (\<lambda>y. f (x, y)))"
+ assumes f: "integrable P f"
+ shows "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1)"
unfolding integrable_snd_measurable[OF assms]
unfolding integrable_fst_measurable[OF assms] ..
@@ -997,291 +1130,140 @@
section "Products"
locale product_sigma_algebra =
- fixes M :: "'i \<Rightarrow> 'a algebra"
+ fixes M :: "'i \<Rightarrow> ('a, 'b) measure_space_scheme"
assumes sigma_algebras: "\<And>i. sigma_algebra (M i)"
-locale finite_product_sigma_algebra = product_sigma_algebra M for M :: "'i \<Rightarrow> 'a algebra" +
+locale finite_product_sigma_algebra = product_sigma_algebra M
+ for M :: "'i \<Rightarrow> ('a, 'b) measure_space_scheme" +
fixes I :: "'i set"
assumes finite_index: "finite I"
+definition
+ "product_algebra_generator I M = \<lparr> space = (\<Pi>\<^isub>E i \<in> I. space (M i)),
+ sets = Pi\<^isub>E I ` (\<Pi> i \<in> I. sets (M i)),
+ measure = \<lambda>A. (\<Prod>i\<in>I. measure (M i) ((SOME F. A = Pi\<^isub>E I F) i)) \<rparr>"
+
+definition product_algebra_def:
+ "Pi\<^isub>M I M = sigma (product_algebra_generator I M)
+ \<lparr> measure := (SOME \<mu>. sigma_finite_measure (sigma (product_algebra_generator I M) \<lparr> measure := \<mu> \<rparr>) \<and>
+ (\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<mu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))))\<rparr>"
+
syntax
- "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3PIE _:_./ _)" 10)
+ "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
+ ('i => 'a, 'b) measure_space_scheme" ("(3PIM _:_./ _)" 10)
syntax (xsymbols)
- "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\<Pi>\<^isub>E _\<in>_./ _)" 10)
+ "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
+ ('i => 'a, 'b) measure_space_scheme" ("(3\<Pi>\<^isub>M _\<in>_./ _)" 10)
syntax (HTML output)
- "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\<Pi>\<^isub>E _\<in>_./ _)" 10)
+ "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
+ ('i => 'a, 'b) measure_space_scheme" ("(3\<Pi>\<^isub>M _\<in>_./ _)" 10)
translations
- "PIE x:A. B" == "CONST Pi\<^isub>E A (%x. B)"
+ "PIM x:I. M" == "CONST Pi\<^isub>M I (%x. M)"
-definition
- "product_algebra M I = \<lparr> space = (\<Pi>\<^isub>E i\<in>I. space (M i)), sets = Pi\<^isub>E I ` (\<Pi> i \<in> I. sets (M i)) \<rparr>"
-
-abbreviation (in finite_product_sigma_algebra) "G \<equiv> product_algebra M I"
-abbreviation (in finite_product_sigma_algebra) "P \<equiv> sigma G"
+abbreviation (in finite_product_sigma_algebra) "G \<equiv> product_algebra_generator I M"
+abbreviation (in finite_product_sigma_algebra) "P \<equiv> Pi\<^isub>M I M"
sublocale product_sigma_algebra \<subseteq> M: sigma_algebra "M i" for i by (rule sigma_algebras)
-lemma (in finite_product_sigma_algebra) product_algebra_into_space:
- "sets G \<subseteq> Pow (space G)"
- using M.sets_into_space unfolding product_algebra_def
+lemma sigma_into_space:
+ assumes "sets M \<subseteq> Pow (space M)"
+ shows "sets (sigma M) \<subseteq> Pow (space M)"
+ using sigma_sets_into_sp[OF assms] unfolding sigma_def by auto
+
+lemma (in product_sigma_algebra) product_algebra_generator_into_space:
+ "sets (product_algebra_generator I M) \<subseteq> Pow (space (product_algebra_generator I M))"
+ using M.sets_into_space unfolding product_algebra_generator_def
by auto blast
+lemma (in product_sigma_algebra) product_algebra_into_space:
+ "sets (Pi\<^isub>M I M) \<subseteq> Pow (space (Pi\<^isub>M I M))"
+ using product_algebra_generator_into_space
+ by (auto intro!: sigma_into_space simp add: product_algebra_def)
+
+lemma (in product_sigma_algebra) sigma_algebra_product_algebra: "sigma_algebra (Pi\<^isub>M I M)"
+ using product_algebra_generator_into_space unfolding product_algebra_def
+ by (rule sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) simp_all
+
sublocale finite_product_sigma_algebra \<subseteq> sigma_algebra P
- using product_algebra_into_space by (rule sigma_algebra_sigma)
+ using sigma_algebra_product_algebra .
lemma product_algebraE:
- assumes "A \<in> sets (product_algebra M I)"
+ assumes "A \<in> sets (product_algebra_generator I M)"
obtains E where "A = Pi\<^isub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))"
- using assms unfolding product_algebra_def by auto
+ using assms unfolding product_algebra_generator_def by auto
-lemma product_algebraI[intro]:
+lemma product_algebra_generatorI[intro]:
assumes "E \<in> (\<Pi> i\<in>I. sets (M i))"
- shows "Pi\<^isub>E I E \<in> sets (product_algebra M I)"
- using assms unfolding product_algebra_def by auto
+ shows "Pi\<^isub>E I E \<in> sets (product_algebra_generator I M)"
+ using assms unfolding product_algebra_generator_def by auto
+
+lemma space_product_algebra_generator[simp]:
+ "space (product_algebra_generator I M) = Pi\<^isub>E I (\<lambda>i. space (M i))"
+ unfolding product_algebra_generator_def by simp
lemma space_product_algebra[simp]:
- "space (product_algebra M I) = Pi\<^isub>E I (\<lambda>i. space (M i))"
- unfolding product_algebra_def by simp
+ "space (Pi\<^isub>M I M) = (\<Pi>\<^isub>E i\<in>I. space (M i))"
+ unfolding product_algebra_def product_algebra_generator_def by simp
-lemma product_algebra_sets_into_space:
+lemma sets_product_algebra:
+ "sets (Pi\<^isub>M I M) = sets (sigma (product_algebra_generator I M))"
+ unfolding product_algebra_def sigma_def by simp
+
+lemma product_algebra_generator_sets_into_space:
assumes "\<And>i. i\<in>I \<Longrightarrow> sets (M i) \<subseteq> Pow (space (M i))"
- shows "sets (product_algebra M I) \<subseteq> Pow (space (product_algebra M I))"
- using assms by (auto simp: product_algebra_def) blast
-
-lemma (in finite_product_sigma_algebra) P_empty:
- "I = {} \<Longrightarrow> P = \<lparr> space = {\<lambda>k. undefined}, sets = { {}, {\<lambda>k. undefined} }\<rparr>"
- unfolding product_algebra_def by (simp add: sigma_def image_constant)
+ shows "sets (product_algebra_generator I M) \<subseteq> Pow (space (product_algebra_generator I M))"
+ using assms by (auto simp: product_algebra_generator_def) blast
lemma (in finite_product_sigma_algebra) in_P[simp, intro]:
"\<lbrakk> \<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i) \<rbrakk> \<Longrightarrow> Pi\<^isub>E I A \<in> sets P"
- by (auto simp: product_algebra_def sets_sigma intro!: sigma_sets.Basic)
-
-lemma (in product_sigma_algebra) bij_inv_restrict_merge:
- assumes [simp]: "I \<inter> J = {}"
- shows "bij_inv
- (space (sigma (product_algebra M (I \<union> J))))
- (space (sigma (pair_algebra (product_algebra M I) (product_algebra M J))))
- (\<lambda>x. (restrict x I, restrict x J)) (\<lambda>(x, y). merge I x J y)"
- by (rule bij_invI) (auto simp: space_pair_algebra extensional_restrict)
-
-lemma (in product_sigma_algebra) bij_inv_singleton:
- "bij_inv (space (sigma (product_algebra M {i}))) (space (M i))
- (\<lambda>x. x i) (\<lambda>x. (\<lambda>j\<in>{i}. x))"
- by (rule bij_invI) (auto simp: restrict_def extensional_def fun_eq_iff)
-
-lemma (in product_sigma_algebra) bij_inv_restrict_insert:
- assumes [simp]: "i \<notin> I"
- shows "bij_inv
- (space (sigma (product_algebra M (insert i I))))
- (space (sigma (pair_algebra (product_algebra M I) (M i))))
- (\<lambda>x. (restrict x I, x i)) (\<lambda>(x, y). x(i := y))"
- by (rule bij_invI) (auto simp: space_pair_algebra extensional_restrict)
-
-lemma (in product_sigma_algebra) measurable_restrict_on_generating:
- assumes [simp]: "I \<inter> J = {}"
- shows "(\<lambda>x. (restrict x I, restrict x J)) \<in> measurable
- (product_algebra M (I \<union> J))
- (pair_algebra (product_algebra M I) (product_algebra M J))"
- (is "?R \<in> measurable ?IJ ?P")
-proof (unfold measurable_def, intro CollectI conjI ballI)
- show "?R \<in> space ?IJ \<rightarrow> space ?P" by (auto simp: space_pair_algebra)
- { fix F E assume "E \<in> (\<Pi> i\<in>I. sets (M i))" "F \<in> (\<Pi> i\<in>J. sets (M i))"
- then have "Pi (I \<union> J) (merge I E J F) \<inter> (\<Pi>\<^isub>E i\<in>I \<union> J. space (M i)) =
- Pi\<^isub>E (I \<union> J) (merge I E J F)"
- using M.sets_into_space by auto blast+ }
- note this[simp]
- show "\<And>A. A \<in> sets ?P \<Longrightarrow> ?R -` A \<inter> space ?IJ \<in> sets ?IJ"
- by (force elim!: pair_algebraE product_algebraE
- simp del: vimage_Int simp: restrict_vimage merge_vimage space_pair_algebra)
- qed
-
-lemma (in product_sigma_algebra) measurable_merge_on_generating:
- assumes [simp]: "I \<inter> J = {}"
- shows "(\<lambda>(x, y). merge I x J y) \<in> measurable
- (pair_algebra (product_algebra M I) (product_algebra M J))
- (product_algebra M (I \<union> J))"
- (is "?M \<in> measurable ?P ?IJ")
-proof (unfold measurable_def, intro CollectI conjI ballI)
- show "?M \<in> space ?P \<rightarrow> space ?IJ" by (auto simp: space_pair_algebra)
- { fix E assume "E \<in> (\<Pi> i\<in>I. sets (M i))" "E \<in> (\<Pi> i\<in>J. sets (M i))"
- then have "Pi I E \<times> Pi J E \<inter> (\<Pi>\<^isub>E i\<in>I. space (M i)) \<times> (\<Pi>\<^isub>E i\<in>J. space (M i)) =
- Pi\<^isub>E I E \<times> Pi\<^isub>E J E"
- using M.sets_into_space by auto blast+ }
- note this[simp]
- show "\<And>A. A \<in> sets ?IJ \<Longrightarrow> ?M -` A \<inter> space ?P \<in> sets ?P"
- by (force elim!: pair_algebraE product_algebraE
- simp del: vimage_Int simp: restrict_vimage merge_vimage space_pair_algebra)
- qed
-
-lemma (in product_sigma_algebra) measurable_singleton_on_generator:
- "(\<lambda>x. \<lambda>j\<in>{i}. x) \<in> measurable (M i) (product_algebra M {i})"
- (is "?f \<in> measurable _ ?P")
-proof (unfold measurable_def, intro CollectI conjI)
- show "?f \<in> space (M i) \<rightarrow> space ?P" by auto
- have "\<And>E. E i \<in> sets (M i) \<Longrightarrow> ?f -` Pi\<^isub>E {i} E \<inter> space (M i) = E i"
- using M.sets_into_space by auto
- then show "\<forall>A \<in> sets ?P. ?f -` A \<inter> space (M i) \<in> sets (M i)"
- by (auto elim!: product_algebraE)
-qed
-
-lemma (in product_sigma_algebra) measurable_component_on_generator:
- assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (product_algebra M I) (M i)"
- (is "?f \<in> measurable ?P _")
-proof (unfold measurable_def, intro CollectI conjI ballI)
- show "?f \<in> space ?P \<rightarrow> space (M i)" using `i \<in> I` by auto
- fix A assume "A \<in> sets (M i)"
- moreover then have "(\<lambda>x. x i) -` A \<inter> (\<Pi>\<^isub>E i\<in>I. space (M i)) =
- (\<Pi>\<^isub>E j\<in>I. if i = j then A else space (M j))"
- using M.sets_into_space `i \<in> I`
- by (fastsimp dest: Pi_mem split: split_if_asm)
- ultimately show "?f -` A \<inter> space ?P \<in> sets ?P"
- by (auto intro!: product_algebraI)
-qed
-
-lemma (in product_sigma_algebra) measurable_restrict_singleton_on_generating:
- assumes [simp]: "i \<notin> I"
- shows "(\<lambda>x. (restrict x I, x i)) \<in> measurable
- (product_algebra M (insert i I))
- (pair_algebra (product_algebra M I) (M i))"
- (is "?R \<in> measurable ?I ?P")
-proof (unfold measurable_def, intro CollectI conjI ballI)
- show "?R \<in> space ?I \<rightarrow> space ?P" by (auto simp: space_pair_algebra)
- { fix E F assume "E \<in> (\<Pi> i\<in>I. sets (M i))" "F \<in> sets (M i)"
- then have "(\<lambda>x. (restrict x I, x i)) -` (Pi\<^isub>E I E \<times> F) \<inter> (\<Pi>\<^isub>E i\<in>insert i I. space (M i)) =
- Pi\<^isub>E (insert i I) (E(i := F))"
- using M.sets_into_space using `i\<notin>I` by (auto simp: restrict_Pi_cancel) blast+ }
- note this[simp]
- show "\<And>A. A \<in> sets ?P \<Longrightarrow> ?R -` A \<inter> space ?I \<in> sets ?I"
- by (force elim!: pair_algebraE product_algebraE
- simp del: vimage_Int simp: space_pair_algebra)
-qed
-
-lemma (in product_sigma_algebra) measurable_merge_singleton_on_generating:
- assumes [simp]: "i \<notin> I"
- shows "(\<lambda>(x, y). x(i := y)) \<in> measurable
- (pair_algebra (product_algebra M I) (M i))
- (product_algebra M (insert i I))"
- (is "?M \<in> measurable ?P ?IJ")
-proof (unfold measurable_def, intro CollectI conjI ballI)
- show "?M \<in> space ?P \<rightarrow> space ?IJ" by (auto simp: space_pair_algebra)
- { fix E assume "E \<in> (\<Pi> i\<in>I. sets (M i))" "E i \<in> sets (M i)"
- then have "(\<lambda>(x, y). x(i := y)) -` Pi\<^isub>E (insert i I) E \<inter> (\<Pi>\<^isub>E i\<in>I. space (M i)) \<times> space (M i) =
- Pi\<^isub>E I E \<times> E i"
- using M.sets_into_space by auto blast+ }
- note this[simp]
- show "\<And>A. A \<in> sets ?IJ \<Longrightarrow> ?M -` A \<inter> space ?P \<in> sets ?P"
- by (force elim!: pair_algebraE product_algebraE
- simp del: vimage_Int simp: space_pair_algebra)
-qed
+ by (auto simp: sets_product_algebra)
section "Generating set generates also product algebra"
-lemma pair_sigma_algebra_sigma:
- assumes 1: "S1 \<up> (space E1)" "range S1 \<subseteq> sets E1" and E1: "sets E1 \<subseteq> Pow (space E1)"
- assumes 2: "S2 \<up> (space E2)" "range S2 \<subseteq> sets E2" and E2: "sets E2 \<subseteq> Pow (space E2)"
- shows "sigma (pair_algebra (sigma E1) (sigma E2)) = sigma (pair_algebra E1 E2)"
- (is "?S = ?E")
-proof -
- interpret M1: sigma_algebra "sigma E1" using E1 by (rule sigma_algebra_sigma)
- interpret M2: sigma_algebra "sigma E2" using E2 by (rule sigma_algebra_sigma)
- have P: "sets (pair_algebra E1 E2) \<subseteq> Pow (space E1 \<times> space E2)"
- using E1 E2 by (auto simp add: pair_algebra_def)
- interpret E: sigma_algebra ?E unfolding pair_algebra_def
- using E1 E2 by (intro sigma_algebra_sigma) auto
- { fix A assume "A \<in> sets E1"
- then have "fst -` A \<inter> space ?E = A \<times> (\<Union>i. S2 i)"
- using E1 2 unfolding isoton_def pair_algebra_def by auto
- also have "\<dots> = (\<Union>i. A \<times> S2 i)" by auto
- also have "\<dots> \<in> sets ?E" unfolding pair_algebra_def sets_sigma
- using 2 `A \<in> sets E1`
- by (intro sigma_sets.Union)
- (auto simp: image_subset_iff intro!: sigma_sets.Basic)
- finally have "fst -` A \<inter> space ?E \<in> sets ?E" . }
- moreover
- { fix B assume "B \<in> sets E2"
- then have "snd -` B \<inter> space ?E = (\<Union>i. S1 i) \<times> B"
- using E2 1 unfolding isoton_def pair_algebra_def by auto
- also have "\<dots> = (\<Union>i. S1 i \<times> B)" by auto
- also have "\<dots> \<in> sets ?E"
- using 1 `B \<in> sets E2` unfolding pair_algebra_def sets_sigma
- by (intro sigma_sets.Union)
- (auto simp: image_subset_iff intro!: sigma_sets.Basic)
- finally have "snd -` B \<inter> space ?E \<in> sets ?E" . }
- ultimately have proj:
- "fst \<in> measurable ?E (sigma E1) \<and> snd \<in> measurable ?E (sigma E2)"
- using E1 E2 by (subst (1 2) E.measurable_iff_sigma)
- (auto simp: pair_algebra_def sets_sigma)
- { fix A B assume A: "A \<in> sets (sigma E1)" and B: "B \<in> sets (sigma E2)"
- with proj have "fst -` A \<inter> space ?E \<in> sets ?E" "snd -` B \<inter> space ?E \<in> sets ?E"
- unfolding measurable_def by simp_all
- moreover have "A \<times> B = (fst -` A \<inter> space ?E) \<inter> (snd -` B \<inter> space ?E)"
- using A B M1.sets_into_space M2.sets_into_space
- by (auto simp: pair_algebra_def)
- ultimately have "A \<times> B \<in> sets ?E" by auto }
- then have "sigma_sets (space ?E) (sets (pair_algebra (sigma E1) (sigma E2))) \<subseteq> sets ?E"
- by (intro E.sigma_sets_subset) (auto simp add: pair_algebra_def sets_sigma)
- then have subset: "sets ?S \<subseteq> sets ?E"
- by (simp add: sets_sigma pair_algebra_def)
- have "sets ?S = sets ?E"
- proof (intro set_eqI iffI)
- fix A assume "A \<in> sets ?E" then show "A \<in> sets ?S"
- unfolding sets_sigma
- proof induct
- case (Basic A) then show ?case
- by (auto simp: pair_algebra_def sets_sigma intro: sigma_sets.Basic)
- qed (auto intro: sigma_sets.intros simp: pair_algebra_def)
- next
- fix A assume "A \<in> sets ?S" then show "A \<in> sets ?E" using subset by auto
- qed
- then show ?thesis
- by (simp add: pair_algebra_def sigma_def)
-qed
-
lemma sigma_product_algebra_sigma_eq:
assumes "finite I"
assumes isotone: "\<And>i. i \<in> I \<Longrightarrow> (S i) \<up> (space (E i))"
assumes sets_into: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> sets (E i)"
and E: "\<And>i. sets (E i) \<subseteq> Pow (space (E i))"
- shows "sigma (product_algebra (\<lambda>i. sigma (E i)) I) = sigma (product_algebra E I)"
- (is "?S = ?E")
+ shows "sets (\<Pi>\<^isub>M i\<in>I. sigma (E i)) = sets (\<Pi>\<^isub>M i\<in>I. E i)"
+ (is "sets ?S = sets ?E")
proof cases
- assume "I = {}" then show ?thesis by (simp add: product_algebra_def)
+ assume "I = {}" then show ?thesis
+ by (simp add: product_algebra_def product_algebra_generator_def)
next
assume "I \<noteq> {}"
interpret E: sigma_algebra "sigma (E i)" for i
using E by (rule sigma_algebra_sigma)
-
have into_space[intro]: "\<And>i x A. A \<in> sets (E i) \<Longrightarrow> x i \<in> A \<Longrightarrow> x i \<in> space (E i)"
using E by auto
-
interpret G: sigma_algebra ?E
- unfolding product_algebra_def using E
- by (intro sigma_algebra_sigma) (auto dest: Pi_mem)
-
+ unfolding product_algebra_def product_algebra_generator_def using E
+ by (intro sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) (auto dest: Pi_mem)
{ fix A i assume "i \<in> I" and A: "A \<in> sets (E i)"
then have "(\<lambda>x. x i) -` A \<inter> space ?E = (\<Pi>\<^isub>E j\<in>I. if j = i then A else \<Union>n. S j n) \<inter> space ?E"
- using isotone unfolding isoton_def product_algebra_def by (auto dest: Pi_mem)
+ using isotone unfolding isoton_def space_product_algebra
+ by (auto dest: Pi_mem)
also have "\<dots> = (\<Union>n. (\<Pi>\<^isub>E j\<in>I. if j = i then A else S j n))"
- unfolding product_algebra_def
+ unfolding space_product_algebra
apply simp
apply (subst Pi_UN[OF `finite I`])
using isotone[THEN isoton_mono_le] apply simp
apply (simp add: PiE_Int)
apply (intro PiE_cong)
using A sets_into by (auto intro!: into_space)
- also have "\<dots> \<in> sets ?E" unfolding product_algebra_def sets_sigma
+ also have "\<dots> \<in> sets ?E"
using sets_into `A \<in> sets (E i)`
+ unfolding sets_product_algebra sets_sigma
by (intro sigma_sets.Union)
(auto simp: image_subset_iff intro!: sigma_sets.Basic)
finally have "(\<lambda>x. x i) -` A \<inter> space ?E \<in> sets ?E" . }
then have proj:
"\<And>i. i\<in>I \<Longrightarrow> (\<lambda>x. x i) \<in> measurable ?E (sigma (E i))"
using E by (subst G.measurable_iff_sigma)
- (auto simp: product_algebra_def sets_sigma)
-
+ (auto simp: sets_product_algebra sets_sigma)
{ fix A assume A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (sigma (E i))"
with proj have basic: "\<And>i. i \<in> I \<Longrightarrow> (\<lambda>x. x i) -` (A i) \<inter> space ?E \<in> sets ?E"
unfolding measurable_def by simp
@@ -1289,218 +1271,117 @@
using A E.sets_into_space `I \<noteq> {}` unfolding product_algebra_def by auto blast
then have "Pi\<^isub>E I A \<in> sets ?E"
using G.finite_INT[OF `finite I` `I \<noteq> {}` basic, of "\<lambda>i. i"] by simp }
- then have "sigma_sets (space ?E) (sets (product_algebra (\<lambda>i. sigma (E i)) I)) \<subseteq> sets ?E"
- by (intro G.sigma_sets_subset) (auto simp add: sets_sigma product_algebra_def)
+ then have "sigma_sets (space ?E) (sets (product_algebra_generator I (\<lambda>i. sigma (E i)))) \<subseteq> sets ?E"
+ by (intro G.sigma_sets_subset) (auto simp add: product_algebra_generator_def)
then have subset: "sets ?S \<subseteq> sets ?E"
- by (simp add: sets_sigma product_algebra_def)
-
- have "sets ?S = sets ?E"
+ by (simp add: sets_sigma sets_product_algebra)
+ show "sets ?S = sets ?E"
proof (intro set_eqI iffI)
fix A assume "A \<in> sets ?E" then show "A \<in> sets ?S"
- unfolding sets_sigma
+ unfolding sets_sigma sets_product_algebra
proof induct
case (Basic A) then show ?case
- by (auto simp: sets_sigma product_algebra_def intro: sigma_sets.Basic)
- qed (auto intro: sigma_sets.intros simp: product_algebra_def)
+ by (auto simp: sets_sigma product_algebra_generator_def intro: sigma_sets.Basic)
+ qed (auto intro: sigma_sets.intros simp: product_algebra_generator_def)
next
fix A assume "A \<in> sets ?S" then show "A \<in> sets ?E" using subset by auto
qed
- then show ?thesis
- by (simp add: product_algebra_def sigma_def)
+qed
+
+lemma product_algebraI[intro]:
+ "E \<in> (\<Pi> i\<in>I. sets (M i)) \<Longrightarrow> Pi\<^isub>E I E \<in> sets (Pi\<^isub>M I M)"
+ using assms unfolding product_algebra_def by (auto intro: product_algebra_generatorI)
+
+lemma (in product_sigma_algebra) measurable_component_update:
+ assumes "x \<in> space (Pi\<^isub>M I M)" and "i \<notin> I"
+ shows "(\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^isub>M (insert i I) M)" (is "?f \<in> _")
+ unfolding product_algebra_def apply simp
+proof (intro measurable_sigma)
+ let ?G = "product_algebra_generator (insert i I) M"
+ show "sets ?G \<subseteq> Pow (space ?G)" using product_algebra_generator_into_space .
+ show "?f \<in> space (M i) \<rightarrow> space ?G"
+ using M.sets_into_space assms by auto
+ fix A assume "A \<in> sets ?G"
+ from product_algebraE[OF this] guess E . note E = this
+ then have "?f -` A \<inter> space (M i) = E i \<or> ?f -` A \<inter> space (M i) = {}"
+ using M.sets_into_space assms by auto
+ then show "?f -` A \<inter> space (M i) \<in> sets (M i)"
+ using E by (auto intro!: product_algebraI)
qed
-lemma (in product_sigma_algebra) sigma_pair_algebra_sigma_eq:
- "sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J))) =
- sigma (pair_algebra (product_algebra M I) (product_algebra M J))"
- using M.sets_into_space
- by (intro pair_sigma_algebra_sigma[of "\<lambda>_. \<Pi>\<^isub>E i\<in>I. space (M i)", of _ "\<lambda>_. \<Pi>\<^isub>E i\<in>J. space (M i)"])
- (auto simp: isoton_const product_algebra_def, blast+)
-
-lemma (in product_sigma_algebra) sigma_pair_algebra_product_singleton:
- "sigma (pair_algebra (sigma (product_algebra M I)) (M i)) =
- sigma (pair_algebra (product_algebra M I) (M i))"
- using M.sets_into_space apply (subst M.sigma_eq[symmetric])
- by (intro pair_sigma_algebra_sigma[of "\<lambda>_. \<Pi>\<^isub>E i\<in>I. space (M i)" _ "\<lambda>_. space (M i)"])
- (auto simp: isoton_const product_algebra_def, blast+)
-
-lemma (in product_sigma_algebra) measurable_restrict:
- assumes [simp]: "I \<inter> J = {}"
- shows "(\<lambda>x. (restrict x I, restrict x J)) \<in> measurable
- (sigma (product_algebra M (I \<union> J)))
- (sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J))))"
- unfolding sigma_pair_algebra_sigma_eq using M.sets_into_space
- by (intro measurable_sigma_sigma measurable_restrict_on_generating
- pair_algebra_sets_into_space product_algebra_sets_into_space)
- auto
+lemma (in product_sigma_algebra) measurable_add_dim:
+ assumes "i \<notin> I"
+ shows "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>M (insert i I) M)"
+proof -
+ let ?f = "(\<lambda>(f, y). f(i := y))" and ?G = "product_algebra_generator (insert i I) M"
+ interpret Ii: pair_sigma_algebra "Pi\<^isub>M I M" "M i"
+ unfolding pair_sigma_algebra_def
+ by (intro sigma_algebra_product_algebra sigma_algebras conjI)
+ have "?f \<in> measurable Ii.P (sigma ?G)"
+ proof (rule Ii.measurable_sigma)
+ show "sets ?G \<subseteq> Pow (space ?G)"
+ using product_algebra_generator_into_space .
+ show "?f \<in> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) \<rightarrow> space ?G"
+ by (auto simp: space_pair_measure)
+ next
+ fix A assume "A \<in> sets ?G"
+ then obtain F where "A = Pi\<^isub>E (insert i I) F"
+ and F: "\<And>j. j \<in> I \<Longrightarrow> F j \<in> sets (M j)" "F i \<in> sets (M i)"
+ by (auto elim!: product_algebraE)
+ then have "?f -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) = Pi\<^isub>E I F \<times> (F i)"
+ using sets_into_space `i \<notin> I`
+ by (auto simp add: space_pair_measure) blast+
+ then show "?f -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) \<in> sets (Pi\<^isub>M I M \<Otimes>\<^isub>M M i)"
+ using F by (auto intro!: pair_measureI)
+ qed
+ then show ?thesis
+ by (simp add: product_algebra_def)
+qed
lemma (in product_sigma_algebra) measurable_merge:
assumes [simp]: "I \<inter> J = {}"
- shows "(\<lambda>(x, y). merge I x J y) \<in> measurable
- (sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J))))
- (sigma (product_algebra M (I \<union> J)))"
- unfolding sigma_pair_algebra_sigma_eq using M.sets_into_space
- by (intro measurable_sigma_sigma measurable_merge_on_generating
- pair_algebra_sets_into_space product_algebra_sets_into_space)
- auto
-
-lemma (in product_sigma_algebra) pair_product_product_vimage_algebra:
- assumes [simp]: "I \<inter> J = {}"
- shows "sigma_algebra.vimage_algebra (sigma (product_algebra M (I \<union> J)))
- (space (sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J))))) (\<lambda>(x,y). merge I x J y) =
- (sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J))))"
- unfolding sigma_pair_algebra_sigma_eq using sets_into_space
- by (intro vimage_algebra_sigma[OF bij_inv_restrict_merge[symmetric]]
- pair_algebra_sets_into_space product_algebra_sets_into_space
- measurable_merge_on_generating measurable_restrict_on_generating)
- auto
-
-lemma (in product_sigma_algebra) measurable_restrict_iff:
- assumes IJ[simp]: "I \<inter> J = {}"
- shows "f \<in> measurable (sigma (pair_algebra
- (sigma (product_algebra M I)) (sigma (product_algebra M J)))) M' \<longleftrightarrow>
- (\<lambda>x. f (restrict x I, restrict x J)) \<in> measurable (sigma (product_algebra M (I \<union> J))) M'"
- using M.sets_into_space
- apply (subst pair_product_product_vimage_algebra[OF IJ, symmetric])
- apply (subst sigma_pair_algebra_sigma_eq)
- apply (subst sigma_algebra.measurable_vimage_iff_inv[OF _
- bij_inv_restrict_merge[symmetric]])
- apply (intro sigma_algebra_sigma product_algebra_sets_into_space)
- by auto
-
-lemma (in product_sigma_algebra) measurable_merge_iff:
- assumes IJ: "I \<inter> J = {}"
- shows "f \<in> measurable (sigma (product_algebra M (I \<union> J))) M' \<longleftrightarrow>
- (\<lambda>(x, y). f (merge I x J y)) \<in>
- measurable (sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J)))) M'"
- unfolding measurable_restrict_iff[OF IJ]
- by (rule measurable_cong) (auto intro!: arg_cong[where f=f] simp: extensional_restrict)
-
-lemma (in product_sigma_algebra) measurable_component:
- assumes "i \<in> I" and f: "f \<in> measurable (M i) M'"
- shows "(\<lambda>x. f (x i)) \<in> measurable (sigma (product_algebra M I)) M'"
- (is "?f \<in> measurable ?P M'")
+ shows "(\<lambda>(x, y). merge I x J y) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)"
proof -
- have "f \<circ> (\<lambda>x. x i) \<in> measurable ?P M'"
- apply (rule measurable_comp[OF _ f])
- using measurable_up_sigma[of "product_algebra M I" "M i"]
- using measurable_component_on_generator[OF `i \<in> I`]
- by auto
- then show "?f \<in> measurable ?P M'" by (simp add: comp_def)
+ let ?I = "Pi\<^isub>M I M" and ?J = "Pi\<^isub>M J M"
+ interpret P: sigma_algebra "?I \<Otimes>\<^isub>M ?J"
+ by (intro sigma_algebra_pair_measure product_algebra_into_space)
+ let ?f = "\<lambda>(x, y). merge I x J y"
+ let ?G = "product_algebra_generator (I \<union> J) M"
+ have "?f \<in> measurable (?I \<Otimes>\<^isub>M ?J) (sigma ?G)"
+ proof (rule P.measurable_sigma)
+ fix A assume "A \<in> sets ?G"
+ from product_algebraE[OF this]
+ obtain E where E: "A = Pi\<^isub>E (I \<union> J) E" "E \<in> (\<Pi> i\<in>I \<union> J. sets (M i))" .
+ then have *: "?f -` A \<inter> space (?I \<Otimes>\<^isub>M ?J) = Pi\<^isub>E I E \<times> Pi\<^isub>E J E"
+ using sets_into_space `I \<inter> J = {}`
+ by (auto simp add: space_pair_measure) blast+
+ then show "?f -` A \<inter> space (?I \<Otimes>\<^isub>M ?J) \<in> sets (?I \<Otimes>\<^isub>M ?J)"
+ using E unfolding * by (auto intro!: pair_measureI in_sigma product_algebra_generatorI
+ simp: product_algebra_def)
+ qed (insert product_algebra_generator_into_space, auto simp: space_pair_measure)
+ then show "?f \<in> measurable (?I \<Otimes>\<^isub>M ?J) (Pi\<^isub>M (I \<union> J) M)"
+ unfolding product_algebra_def[of "I \<union> J"] by simp
qed
-lemma (in product_sigma_algebra) singleton_vimage_algebra:
- "sigma_algebra.vimage_algebra (sigma (product_algebra M {i})) (space (M i)) (\<lambda>x. \<lambda>j\<in>{i}. x) = M i"
- using sets_into_space
- by (intro vimage_algebra_sigma[of "M i", unfolded M.sigma_eq, OF bij_inv_singleton[symmetric]]
- product_algebra_sets_into_space measurable_singleton_on_generator measurable_component_on_generator)
- auto
-
lemma (in product_sigma_algebra) measurable_component_singleton:
- "(\<lambda>x. f (x i)) \<in> measurable (sigma (product_algebra M {i})) M' \<longleftrightarrow>
- f \<in> measurable (M i) M'"
- using sets_into_space
- apply (subst singleton_vimage_algebra[symmetric])
- apply (subst sigma_algebra.measurable_vimage_iff_inv[OF _ bij_inv_singleton[symmetric]])
- by (auto intro!: sigma_algebra_sigma product_algebra_sets_into_space)
-
-lemma (in product_sigma_algebra) measurable_component_iff:
- assumes "i \<in> I" and not_empty: "\<forall>i\<in>I. space (M i) \<noteq> {}"
- shows "(\<lambda>x. f (x i)) \<in> measurable (sigma (product_algebra M I)) M' \<longleftrightarrow>
- f \<in> measurable (M i) M'"
- (is "?f \<in> measurable ?P M' \<longleftrightarrow> _")
-proof
- assume "f \<in> measurable (M i) M'" then show "?f \<in> measurable ?P M'"
- by (rule measurable_component[OF `i \<in> I`])
-next
- assume f: "?f \<in> measurable ?P M'"
- def t \<equiv> "\<lambda>i. SOME x. x \<in> space (M i)"
- have t: "\<And>i. i\<in>I \<Longrightarrow> t i \<in> space (M i)"
- unfolding t_def using not_empty by (rule_tac someI_ex) auto
- have "?f \<circ> (\<lambda>x. (\<lambda>j\<in>I. if j = i then x else t j)) \<in> measurable (M i) M'"
- (is "?f \<circ> ?t \<in> measurable _ _")
- proof (rule measurable_comp[OF _ f])
- have "?t \<in> measurable (M i) (product_algebra M I)"
- proof (unfold measurable_def, intro CollectI conjI ballI)
- from t show "?t \<in> space (M i) \<rightarrow> (space (product_algebra M I))" by auto
- next
- fix A assume A: "A \<in> sets (product_algebra M I)"
- { fix E assume "E \<in> (\<Pi> i\<in>I. sets (M i))"
- then have "?t -` Pi\<^isub>E I E \<inter> space (M i) = (if (\<forall>j\<in>I-{i}. t j \<in> E j) then E i else {})"
- using `i \<in> I` sets_into_space by (auto dest: Pi_mem[where B=E]) }
- note * = this
- with A `i \<in> I` show "?t -` A \<inter> space (M i) \<in> sets (M i)"
- by (auto elim!: product_algebraE simp del: vimage_Int)
- qed
- also have "\<dots> \<subseteq> measurable (M i) (sigma (product_algebra M I))"
- using M.sets_into_space by (intro measurable_subset) (auto simp: product_algebra_def, blast)
- finally show "?t \<in> measurable (M i) (sigma (product_algebra M I))" .
- qed
- then show "f \<in> measurable (M i) M'" unfolding comp_def using `i \<in> I` by simp
-qed
-
-lemma (in product_sigma_algebra) measurable_singleton:
- shows "f \<in> measurable (sigma (product_algebra M {i})) M' \<longleftrightarrow>
- (\<lambda>x. f (\<lambda>j\<in>{i}. x)) \<in> measurable (M i) M'"
- using sets_into_space unfolding measurable_component_singleton[symmetric]
- by (auto intro!: measurable_cong arg_cong[where f=f] simp: fun_eq_iff extensional_def)
-
-
-lemma (in pair_sigma_algebra) measurable_pair_split:
- assumes "sigma_algebra M1'" "sigma_algebra M2'"
- assumes f: "f \<in> measurable M1 M1'" and g: "g \<in> measurable M2 M2'"
- shows "(\<lambda>(x, y). (f x, g y)) \<in> measurable P (sigma (pair_algebra M1' M2'))"
-proof (rule measurable_sigma)
- interpret M1': sigma_algebra M1' by fact
- interpret M2': sigma_algebra M2' by fact
- interpret Q: pair_sigma_algebra M1' M2' by default
- show "sets Q.E \<subseteq> Pow (space Q.E)"
- using M1'.sets_into_space M2'.sets_into_space by (auto simp: pair_algebra_def)
- show "(\<lambda>(x, y). (f x, g y)) \<in> space P \<rightarrow> space Q.E"
- using f g unfolding measurable_def pair_algebra_def by auto
- fix A assume "A \<in> sets Q.E"
- then obtain X Y where A: "A = X \<times> Y" "X \<in> sets M1'" "Y \<in> sets M2'"
- unfolding pair_algebra_def by auto
- then have *: "(\<lambda>(x, y). (f x, g y)) -` A \<inter> space P =
- (f -` X \<inter> space M1) \<times> (g -` Y \<inter> space M2)"
- by (auto simp: pair_algebra_def)
- then show "(\<lambda>(x, y). (f x, g y)) -` A \<inter> space P \<in> sets P"
- using f g A unfolding measurable_def *
- by (auto intro!: pair_algebraI in_sigma)
-qed
-
-lemma (in product_sigma_algebra) measurable_add_dim:
- assumes "i \<notin> I" "finite I"
- shows "(\<lambda>(f, y). f(i := y)) \<in> measurable (sigma (pair_algebra (sigma (product_algebra M I)) (M i)))
- (sigma (product_algebra M (insert i I)))"
-proof (subst measurable_cong)
- interpret I: finite_product_sigma_algebra M I by default fact
- interpret i: finite_product_sigma_algebra M "{i}" by default auto
- interpret Ii: pair_sigma_algebra I.P "M i" by default
- interpret Ii': pair_sigma_algebra I.P i.P by default
- have disj: "I \<inter> {i} = {}" using `i \<notin> I` by auto
- have "(\<lambda>(x, y). (id x, \<lambda>_\<in>{i}. y)) \<in> measurable Ii.P Ii'.P"
- proof (intro Ii.measurable_pair_split I.measurable_ident)
- show "(\<lambda>y. \<lambda>_\<in>{i}. y) \<in> measurable (M i) i.P"
- apply (rule measurable_singleton[THEN iffD1])
- using i.measurable_ident unfolding id_def .
- qed default
- from measurable_comp[OF this measurable_merge[OF disj]]
- show "(\<lambda>(x, y). merge I x {i} y) \<circ> (\<lambda>(x, y). (id x, \<lambda>_\<in>{i}. y))
- \<in> measurable (sigma (pair_algebra I.P (M i))) (sigma (product_algebra M (insert i I)))"
- (is "?f \<in> _") by simp
- fix x assume "x \<in> space Ii.P"
- with assms show "(\<lambda>(f, y). f(i := y)) x = ?f x"
- by (cases x) (simp add: merge_def fun_eq_iff pair_algebra_def extensional_def)
-qed
+ assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^isub>M I M) (M i)"
+proof (unfold measurable_def, intro CollectI conjI ballI)
+ fix A assume "A \<in> sets (M i)"
+ then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) = (\<Pi>\<^isub>E j\<in>I. if i = j then A else space (M j))"
+ using M.sets_into_space `i \<in> I` by (fastsimp dest: Pi_mem split: split_if_asm)
+ then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) \<in> sets (Pi\<^isub>M I M)"
+ using `A \<in> sets (M i)` by (auto intro!: product_algebraI)
+qed (insert `i \<in> I`, auto)
locale product_sigma_finite =
- fixes M :: "'i \<Rightarrow> 'a algebra" and \<mu> :: "'i \<Rightarrow> 'a set \<Rightarrow> pextreal"
- assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i) (\<mu> i)"
+ fixes M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme"
+ assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i)"
-locale finite_product_sigma_finite = product_sigma_finite M \<mu> for M :: "'i \<Rightarrow> 'a algebra" and \<mu> +
- fixes I :: "'i set" assumes finite_index': "finite I"
+locale finite_product_sigma_finite = product_sigma_finite M
+ for M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme" +
+ fixes I :: "'i set" assumes finite_index'[intro]: "finite I"
-sublocale product_sigma_finite \<subseteq> M: sigma_finite_measure "M i" "\<mu> i" for i
+sublocale product_sigma_finite \<subseteq> M: sigma_finite_measure "M i" for i
by (rule sigma_finite_measures)
sublocale product_sigma_finite \<subseteq> product_sigma_algebra
@@ -1509,6 +1390,30 @@
sublocale finite_product_sigma_finite \<subseteq> finite_product_sigma_algebra
by default (fact finite_index')
+lemma (in finite_product_sigma_finite) product_algebra_generator_measure:
+ assumes "Pi\<^isub>E I F \<in> sets G"
+ shows "measure G (Pi\<^isub>E I F) = (\<Prod>i\<in>I. M.\<mu> i (F i))"
+proof cases
+ assume ne: "\<forall>i\<in>I. F i \<noteq> {}"
+ have "\<forall>i\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') i = F i"
+ by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"])
+ (insert ne, auto simp: Pi_eq_iff)
+ then show ?thesis
+ unfolding product_algebra_generator_def by simp
+next
+ assume empty: "\<not> (\<forall>j\<in>I. F j \<noteq> {})"
+ then have "(\<Prod>j\<in>I. M.\<mu> j (F j)) = 0"
+ by (auto simp: setprod_pextreal_0 intro!: bexI)
+ moreover
+ have "\<exists>j\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j = {}"
+ by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"])
+ (insert empty, auto simp: Pi_eq_empty_iff[symmetric])
+ then have "(\<Prod>j\<in>I. M.\<mu> j ((SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j)) = 0"
+ by (auto simp: setprod_pextreal_0 intro!: bexI)
+ ultimately show ?thesis
+ unfolding product_algebra_generator_def by simp
+qed
+
lemma (in finite_product_sigma_finite) sigma_finite_pairs:
"\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set.
(\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and>
@@ -1544,46 +1449,60 @@
lemma (in product_sigma_finite) product_measure_exists:
assumes "finite I"
- shows "\<exists>\<nu>. (\<forall>A\<in>(\<Pi> i\<in>I. sets (M i)). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. \<mu> i (A i))) \<and>
- sigma_finite_measure (sigma (product_algebra M I)) \<nu>"
+ shows "\<exists>\<nu>. sigma_finite_measure (sigma (product_algebra_generator I M) \<lparr> measure := \<nu> \<rparr>) \<and>
+ (\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i)))"
using `finite I` proof induct
- case empty then show ?case unfolding product_algebra_def
- by (auto intro!: exI[of _ "\<lambda>A. if A = {} then 0 else 1"] sigma_algebra_sigma
- sigma_algebra.finite_additivity_sufficient
- simp add: positive_def additive_def sets_sigma sigma_finite_measure_def
- sigma_finite_measure_axioms_def image_constant)
+ case empty
+ interpret finite_product_sigma_finite M "{}" by default simp
+ let ?\<nu> = "(\<lambda>A. if A = {} then 0 else 1) :: 'd set \<Rightarrow> pextreal"
+ show ?case
+ proof (intro exI conjI ballI)
+ have "sigma_algebra (sigma G \<lparr>measure := ?\<nu>\<rparr>)"
+ by (rule sigma_algebra_cong) (simp_all add: product_algebra_def)
+ then have "measure_space (sigma G\<lparr>measure := ?\<nu>\<rparr>)"
+ by (rule finite_additivity_sufficient)
+ (simp_all add: positive_def additive_def sets_sigma
+ product_algebra_generator_def image_constant)
+ then show "sigma_finite_measure (sigma G\<lparr>measure := ?\<nu>\<rparr>)"
+ by (auto intro!: exI[of _ "\<lambda>i. {\<lambda>_. undefined}"]
+ simp: sigma_finite_measure_def sigma_finite_measure_axioms_def
+ product_algebra_generator_def)
+ qed auto
next
case (insert i I)
- interpret finite_product_sigma_finite M \<mu> I by default fact
+ interpret finite_product_sigma_finite M I by default fact
have "finite (insert i I)" using `finite I` by auto
- interpret I': finite_product_sigma_finite M \<mu> "insert i I" by default fact
+ interpret I': finite_product_sigma_finite M "insert i I" by default fact
from insert obtain \<nu> where
- prod: "\<And>A. A \<in> (\<Pi> i\<in>I. sets (M i)) \<Longrightarrow> \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. \<mu> i (A i))" and
- "sigma_finite_measure P \<nu>" by auto
- interpret I: sigma_finite_measure P \<nu> by fact
- interpret P: pair_sigma_finite P \<nu> "M i" "\<mu> i" ..
-
+ prod: "\<And>A. A \<in> (\<Pi> i\<in>I. sets (M i)) \<Longrightarrow> \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))" and
+ "sigma_finite_measure (sigma G\<lparr> measure := \<nu> \<rparr>)" by auto
+ then interpret I: sigma_finite_measure "P\<lparr> measure := \<nu>\<rparr>" unfolding product_algebra_def by simp
+ interpret P: pair_sigma_finite "P\<lparr> measure := \<nu>\<rparr>" "M i" ..
let ?h = "(\<lambda>(f, y). f(i := y))"
- let ?\<nu> = "\<lambda>A. P.pair_measure (?h -` A \<inter> space P.P)"
- have I': "sigma_algebra I'.P" by default
- interpret I': measure_space "sigma (product_algebra M (insert i I))" ?\<nu>
- apply (rule P.measure_space_vimage[OF I'])
- apply (rule measurable_add_dim[OF `i \<notin> I` `finite I`])
- by simp
+ let ?\<nu> = "\<lambda>A. P.\<mu> (?h -` A \<inter> space P.P)"
+ have I': "sigma_algebra (I'.P\<lparr> measure := ?\<nu> \<rparr>)"
+ by (rule I'.sigma_algebra_cong) simp_all
+ interpret I'': measure_space "I'.P\<lparr> measure := ?\<nu> \<rparr>"
+ using measurable_add_dim[OF `i \<notin> I`]
+ by (intro P.measure_space_vimage[OF I'])
+ (simp_all add: measurable_def pair_measure_def pair_measure_generator_def sets_sigma)
show ?case
proof (intro exI[of _ ?\<nu>] conjI ballI)
+ let "?m A" = "measure (Pi\<^isub>M I M\<lparr>measure := \<nu>\<rparr> \<Otimes>\<^isub>M M i) (?h -` A \<inter> space P.P)"
{ fix A assume A: "A \<in> (\<Pi> i\<in>insert i I. sets (M i))"
then have *: "?h -` Pi\<^isub>E (insert i I) A \<inter> space P.P = Pi\<^isub>E I A \<times> A i"
- using `i \<notin> I` M.sets_into_space by (auto simp: pair_algebra_def) blast
- show "?\<nu> (Pi\<^isub>E (insert i I) A) = (\<Prod>i\<in>insert i I. \<mu> i (A i))"
+ using `i \<notin> I` M.sets_into_space by (auto simp: space_pair_measure space_product_algebra) blast
+ show "?m (Pi\<^isub>E (insert i I) A) = (\<Prod>i\<in>insert i I. M.\<mu> i (A i))"
unfolding * using A
apply (subst P.pair_measure_times)
using A apply fastsimp
using A apply fastsimp
using `i \<notin> I` `finite I` prod[of A] A by (auto simp: ac_simps) }
note product = this
- show "sigma_finite_measure I'.P ?\<nu>"
- proof
+ have *: "sigma I'.G\<lparr> measure := ?\<nu> \<rparr> = I'.P\<lparr> measure := ?\<nu> \<rparr>"
+ by (simp add: product_algebra_def)
+ show "sigma_finite_measure (sigma I'.G\<lparr> measure := ?\<nu> \<rparr>)"
+ proof (unfold *, default, simp)
from I'.sigma_finite_pairs guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
then have F: "\<And>j. j \<in> insert i I \<Longrightarrow> range (F j) \<subseteq> sets (M j)"
"(\<lambda>k. \<Pi>\<^isub>E j \<in> insert i I. F j k) \<up> space I'.G"
@@ -1591,12 +1510,12 @@
by blast+
let "?F k" = "\<Pi>\<^isub>E j \<in> insert i I. F j k"
show "\<exists>F::nat \<Rightarrow> ('i \<Rightarrow> 'a) set. range F \<subseteq> sets I'.P \<and>
- (\<Union>i. F i) = space I'.P \<and> (\<forall>i. ?\<nu> (F i) \<noteq> \<omega>)"
+ (\<Union>i. F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i)) \<and> (\<forall>i. ?m (F i) \<noteq> \<omega>)"
proof (intro exI[of _ ?F] conjI allI)
show "range ?F \<subseteq> sets I'.P" using F(1) by auto
next
from F(2)[THEN isotonD(2)]
- show "(\<Union>i. ?F i) = space I'.P" by simp
+ show "(\<Union>i. ?F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i))" by simp
next
fix j
show "?\<nu> (?F j) \<noteq> \<omega>"
@@ -1607,62 +1526,43 @@
qed
qed
-definition (in finite_product_sigma_finite)
- measure :: "('i \<Rightarrow> 'a) set \<Rightarrow> pextreal" where
- "measure = (SOME \<nu>.
- (\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. \<mu> i (A i))) \<and>
- sigma_finite_measure P \<nu>)"
-
-sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure P measure
-proof -
- show "sigma_finite_measure P measure"
- unfolding measure_def
- by (rule someI2_ex[OF product_measure_exists[OF finite_index]]) auto
-qed
+sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure P
+ unfolding product_algebra_def
+ using product_measure_exists[OF finite_index]
+ by (rule someI2_ex) auto
lemma (in finite_product_sigma_finite) measure_times:
assumes "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)"
- shows "measure (Pi\<^isub>E I A) = (\<Prod>i\<in>I. \<mu> i (A i))"
-proof -
- note ex = product_measure_exists[OF finite_index]
- show ?thesis
- unfolding measure_def
- proof (rule someI2_ex[OF ex], elim conjE)
- fix \<nu> assume *: "\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. \<mu> i (A i))"
+ shows "\<mu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))"
+ unfolding product_algebra_def
+ using product_measure_exists[OF finite_index]
+ proof (rule someI2_ex, elim conjE)
+ fix \<nu> assume *: "\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))"
have "Pi\<^isub>E I A = Pi\<^isub>E I (\<lambda>i\<in>I. A i)" by (auto dest: Pi_mem)
then have "\<nu> (Pi\<^isub>E I A) = \<nu> (Pi\<^isub>E I (\<lambda>i\<in>I. A i))" by simp
- also have "\<dots> = (\<Prod>i\<in>I. \<mu> i ((\<lambda>i\<in>I. A i) i))" using assms * by auto
- finally show "\<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. \<mu> i (A i))" by simp
+ also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i ((\<lambda>i\<in>I. A i) i))" using assms * by auto
+ finally show "measure (sigma G\<lparr>measure := \<nu>\<rparr>) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))"
+ by (simp add: sigma_def)
qed
-qed
-
-abbreviation (in product_sigma_finite)
- "product_measure I \<equiv> finite_product_sigma_finite.measure M \<mu> I"
-
-abbreviation (in product_sigma_finite)
- "product_positive_integral I \<equiv>
- measure_space.positive_integral (sigma (product_algebra M I)) (product_measure I)"
-
-abbreviation (in product_sigma_finite)
- "product_integral I \<equiv>
- measure_space.integral (sigma (product_algebra M I)) (product_measure I)"
-
-abbreviation (in product_sigma_finite)
- "product_integrable I \<equiv>
- measure_space.integrable (sigma (product_algebra M I)) (product_measure I)"
lemma (in product_sigma_finite) product_measure_empty[simp]:
- "product_measure {} {\<lambda>x. undefined} = 1"
+ "measure (Pi\<^isub>M {} M) {\<lambda>x. undefined} = 1"
proof -
- interpret finite_product_sigma_finite M \<mu> "{}" by default auto
+ interpret finite_product_sigma_finite M "{}" by default auto
from measure_times[of "\<lambda>x. {}"] show ?thesis by simp
qed
+lemma (in finite_product_sigma_algebra) P_empty:
+ assumes "I = {}"
+ shows "space P = {\<lambda>k. undefined}" "sets P = { {}, {\<lambda>k. undefined} }"
+ unfolding product_algebra_def product_algebra_generator_def `I = {}`
+ by (simp_all add: sigma_def image_constant)
+
lemma (in product_sigma_finite) positive_integral_empty:
- "product_positive_integral {} f = f (\<lambda>k. undefined)"
+ "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\<lambda>k. undefined)"
proof -
- interpret finite_product_sigma_finite M \<mu> "{}" by default (fact finite.emptyI)
- have "\<And>A. measure (Pi\<^isub>E {} A) = 1"
+ interpret finite_product_sigma_finite M "{}" by default (fact finite.emptyI)
+ have "\<And>A. measure (Pi\<^isub>M {} M) (Pi\<^isub>E {} A) = 1"
using assms by (subst measure_times) auto
then show ?thesis
unfolding positive_integral_def simple_function_def simple_integral_def_raw
@@ -1676,18 +1576,15 @@
lemma (in product_sigma_finite) measure_fold:
assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
- assumes A: "A \<in> sets (sigma (product_algebra M (I \<union> J)))"
- shows "pair_sigma_finite.pair_measure
- (sigma (product_algebra M I)) (product_measure I)
- (sigma (product_algebra M J)) (product_measure J)
- ((\<lambda>(x,y). merge I x J y) -` A \<inter> space (sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J))))) =
- product_measure (I \<union> J) A"
+ assumes A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)"
+ shows "measure (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) ((\<lambda>(x,y). merge I x J y) -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)) =
+ measure (Pi\<^isub>M (I \<union> J) M) A"
proof -
- interpret I: finite_product_sigma_finite M \<mu> I by default fact
- interpret J: finite_product_sigma_finite M \<mu> J by default fact
+ interpret I: finite_product_sigma_finite M I by default fact
+ interpret J: finite_product_sigma_finite M J by default fact
have "finite (I \<union> J)" using fin by auto
- interpret IJ: finite_product_sigma_finite M \<mu> "I \<union> J" by default fact
- interpret P: pair_sigma_finite I.P I.measure J.P J.measure by default
+ interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
+ interpret P: pair_sigma_finite I.P J.P by default
let ?g = "\<lambda>(x,y). merge I x J y"
let "?X S" = "?g -` S \<inter> space P.P"
from IJ.sigma_finite_pairs obtain F where
@@ -1696,111 +1593,125 @@
"\<And>k. \<forall>i\<in>I\<union>J. \<mu> i (F i k) \<noteq> \<omega>"
by auto
let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k"
- show "P.pair_measure (?X A) = IJ.measure A"
- proof (rule measure_unique_Int_stable[OF _ _ _ _ _ _ _ _ A])
- show "Int_stable IJ.G" by (simp add: PiE_Int Int_stable_def product_algebra_def) auto
- show "range ?F \<subseteq> sets IJ.G" using F by (simp add: image_subset_iff product_algebra_def)
+ show "P.\<mu> (?X A) = IJ.\<mu> A"
+ proof (rule measure_unique_Int_stable[where X=A])
+ show "A \<in> sets (sigma IJ.G)"
+ using A unfolding product_algebra_def by auto
+ show "Int_stable IJ.G"
+ by (simp add: PiE_Int Int_stable_def product_algebra_def
+ product_algebra_generator_def)
+ auto
+ show "range ?F \<subseteq> sets IJ.G" using F
+ by (simp add: image_subset_iff product_algebra_def
+ product_algebra_generator_def)
show "?F \<up> space IJ.G " using F(2) by simp
- have "sigma_algebra IJ.P" by default
- then show "measure_space IJ.P (\<lambda>A. P.pair_measure (?X A))"
- apply (rule P.measure_space_vimage)
- apply (rule measurable_merge[OF `I \<inter> J = {}`])
- by simp
- show "measure_space IJ.P IJ.measure" by fact
+ have "measure_space IJ.P" by fact
+ also have "IJ.P = \<lparr> space = space IJ.G, sets = sets (sigma IJ.G), measure = IJ.\<mu> \<rparr>"
+ by (simp add: product_algebra_def)
+ finally show "measure_space \<dots>" .
+ let ?P = "\<lparr> space = space IJ.G, sets = sets (sigma IJ.G),
+ measure = \<lambda>A. P.\<mu> (?X A)\<rparr>"
+ have *: "?P = (sigma IJ.G \<lparr> measure := \<lambda>A. P.\<mu> (?X A) \<rparr>)"
+ by auto
+ have "sigma_algebra (sigma IJ.G \<lparr> measure := \<lambda>A. P.\<mu> (?X A) \<rparr>)"
+ by (rule IJ.sigma_algebra_cong) (auto simp: product_algebra_def)
+ then show "measure_space ?P" unfolding *
+ using measurable_merge[OF `I \<inter> J = {}`]
+ by (auto intro!: P.measure_space_vimage simp add: product_algebra_def)
next
fix A assume "A \<in> sets IJ.G"
then obtain F where A[simp]: "A = Pi\<^isub>E (I \<union> J) F"
and F: "\<And>i. i \<in> I \<union> J \<Longrightarrow> F i \<in> sets (M i)"
- by (auto simp: product_algebra_def)
+ by (auto simp: product_algebra_generator_def)
then have "?X A = (Pi\<^isub>E I F \<times> Pi\<^isub>E J F)"
- using sets_into_space by (auto simp: space_pair_algebra) blast+
- then have "P.pair_measure (?X A) = (\<Prod>i\<in>I. \<mu> i (F i)) * (\<Prod>i\<in>J. \<mu> i (F i))"
+ using sets_into_space by (auto simp: space_pair_measure) blast+
+ then have "P.\<mu> (?X A) = (\<Prod>i\<in>I. \<mu> i (F i)) * (\<Prod>i\<in>J. \<mu> i (F i))"
using `finite J` `finite I` F
by (simp add: P.pair_measure_times I.measure_times J.measure_times)
also have "\<dots> = (\<Prod>i\<in>I \<union> J. \<mu> i (F i))"
using `finite J` `finite I` `I \<inter> J = {}` by (simp add: setprod_Un_one)
- also have "\<dots> = IJ.measure A"
+ also have "\<dots> = IJ.\<mu> A"
using `finite J` `finite I` F unfolding A
by (intro IJ.measure_times[symmetric]) auto
- finally show "P.pair_measure (?X A) = IJ.measure A" .
+ finally show "P.\<mu> (?X A) = IJ.\<mu> A" .
next
fix k
have k: "\<And>i. i \<in> I \<union> J \<Longrightarrow> F i k \<in> sets (M i)" using F by auto
then have "?X (?F k) = (\<Pi>\<^isub>E i\<in>I. F i k) \<times> (\<Pi>\<^isub>E i\<in>J. F i k)"
- using sets_into_space by (auto simp: space_pair_algebra) blast+
- with k have "P.pair_measure (?X (?F k)) = (\<Prod>i\<in>I. \<mu> i (F i k)) * (\<Prod>i\<in>J. \<mu> i (F i k))"
+ using sets_into_space by (auto simp: space_pair_measure) blast+
+ with k have "P.\<mu> (?X (?F k)) = (\<Prod>i\<in>I. \<mu> i (F i k)) * (\<Prod>i\<in>J. \<mu> i (F i k))"
by (simp add: P.pair_measure_times I.measure_times J.measure_times)
- then show "P.pair_measure (?X (?F k)) \<noteq> \<omega>"
+ then show "P.\<mu> (?X (?F k)) \<noteq> \<omega>"
using `finite I` F by (simp add: setprod_\<omega>)
- qed simp
+ qed
qed
lemma (in product_sigma_finite) product_positive_integral_fold:
assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
- and f: "f \<in> borel_measurable (sigma (product_algebra M (I \<union> J)))"
- shows "product_positive_integral (I \<union> J) f =
- product_positive_integral I (\<lambda>x. product_positive_integral J (\<lambda>y. f (merge I x J y)))"
+ and f: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)"
+ shows "integral\<^isup>P (Pi\<^isub>M (I \<union> J) M) f =
+ (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (merge I x J y) \<partial>(Pi\<^isub>M J M)) \<partial>(Pi\<^isub>M I M))"
proof -
- interpret I: finite_product_sigma_finite M \<mu> I by default fact
- interpret J: finite_product_sigma_finite M \<mu> J by default fact
+ interpret I: finite_product_sigma_finite M I by default fact
+ interpret J: finite_product_sigma_finite M J by default fact
have "finite (I \<union> J)" using fin by auto
- interpret IJ: finite_product_sigma_finite M \<mu> "I \<union> J" by default fact
- interpret P: pair_sigma_finite I.P I.measure J.P J.measure by default
+ interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
+ interpret P: pair_sigma_finite I.P J.P by default
have P_borel: "(\<lambda>x. f (case x of (x, y) \<Rightarrow> merge I x J y)) \<in> borel_measurable P.P"
- unfolding case_prod_distrib measurable_merge_iff[OF IJ, symmetric] using f .
+ using measurable_comp[OF measurable_merge[OF IJ] f] by (simp add: comp_def)
show ?thesis
unfolding P.positive_integral_fst_measurable[OF P_borel, simplified]
- apply (subst IJ.positive_integral_cong_measure[symmetric])
- apply (rule measure_fold[OF IJ fin])
- apply assumption
proof (rule P.positive_integral_vimage)
show "sigma_algebra IJ.P" by default
show "(\<lambda>(x, y). merge I x J y) \<in> measurable P.P IJ.P" by (rule measurable_merge[OF IJ])
- show "f \<in> borel_measurable IJ.P" using f .
+ show "f \<in> borel_measurable IJ.P" using f by simp
+ next
+ fix A assume "A \<in> sets (Pi\<^isub>M (I \<union> J) M)"
+ then show "IJ.\<mu> A = P.\<mu> ((\<lambda>(x,y). merge I x J y) -` A \<inter> space P.P)"
+ using measure_fold[OF IJ fin] by simp
qed
qed
lemma (in product_sigma_finite) product_positive_integral_singleton:
assumes f: "f \<in> borel_measurable (M i)"
- shows "product_positive_integral {i} (\<lambda>x. f (x i)) = M.positive_integral i f"
+ shows "integral\<^isup>P (Pi\<^isub>M {i} M) (\<lambda>x. f (x i)) = integral\<^isup>P (M i) f"
proof -
- interpret I: finite_product_sigma_finite M \<mu> "{i}" by default simp
- have T: "(\<lambda>x. x i) \<in> measurable (sigma (product_algebra M {i})) (M i)"
- using measurable_component_singleton[of "\<lambda>x. x" i]
- measurable_ident[unfolded id_def] by auto
- show "I.positive_integral (\<lambda>x. f (x i)) = M.positive_integral i f"
- unfolding I.positive_integral_vimage[OF sigma_algebras T f, symmetric]
- proof (rule positive_integral_cong_measure)
- fix A let ?P = "(\<lambda>x. x i) -` A \<inter> space (sigma (product_algebra M {i}))"
+ interpret I: finite_product_sigma_finite M "{i}" by default simp
+ show ?thesis
+ proof (rule I.positive_integral_vimage[symmetric])
+ show "sigma_algebra (M i)" by (rule sigma_algebras)
+ show "(\<lambda>x. x i) \<in> measurable (Pi\<^isub>M {i} M) (M i)" by (rule measurable_component_singleton) auto
+ show "f \<in> borel_measurable (M i)" by fact
+ next
+ fix A let ?P = "(\<lambda>x. x i) -` A \<inter> space I.P"
assume A: "A \<in> sets (M i)"
then have *: "?P = {i} \<rightarrow>\<^isub>E A" using sets_into_space by auto
- show "product_measure {i} ?P = \<mu> i A" unfolding *
+ show "M.\<mu> i A = I.\<mu> ?P" unfolding *
using A I.measure_times[of "\<lambda>_. A"] by auto
qed
qed
lemma (in product_sigma_finite) product_positive_integral_insert:
assumes [simp]: "finite I" "i \<notin> I"
- and f: "f \<in> borel_measurable (sigma (product_algebra M (insert i I)))"
- shows "product_positive_integral (insert i I) f
- = product_positive_integral I (\<lambda>x. M.positive_integral i (\<lambda>y. f (x(i:=y))))"
+ and f: "f \<in> borel_measurable (Pi\<^isub>M (insert i I) M)"
+ shows "integral\<^isup>P (Pi\<^isub>M (insert i I) M) f = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^isub>M I M))"
proof -
- interpret I: finite_product_sigma_finite M \<mu> I by default auto
- interpret i: finite_product_sigma_finite M \<mu> "{i}" by default auto
+ interpret I: finite_product_sigma_finite M I by default auto
+ interpret i: finite_product_sigma_finite M "{i}" by default auto
interpret P: pair_sigma_algebra I.P i.P ..
- have IJ: "I \<inter> {i} = {}" by auto
+ have IJ: "I \<inter> {i} = {}" and insert: "I \<union> {i} = insert i I"
+ using f by auto
show ?thesis
- unfolding product_positive_integral_fold[OF IJ, simplified, OF f]
+ unfolding product_positive_integral_fold[OF IJ, unfolded insert, simplified, OF f]
proof (rule I.positive_integral_cong, subst product_positive_integral_singleton)
fix x assume x: "x \<in> space I.P"
let "?f y" = "f (restrict (x(i := y)) (insert i I))"
have f'_eq: "\<And>y. ?f y = f (x(i := y))"
using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def)
- note fP = f[unfolded measurable_merge_iff[OF IJ, simplified]]
- show "?f \<in> borel_measurable (M i)"
- using P.measurable_pair_image_snd[OF fP x]
- unfolding measurable_singleton f'_eq by (simp add: f'_eq)
- show "M.positive_integral i ?f = M.positive_integral i (\<lambda>y. f (x(i := y)))"
+ show "?f \<in> borel_measurable (M i)" unfolding f'_eq
+ using measurable_comp[OF measurable_component_update f] x `i \<notin> I`
+ by (simp add: comp_def)
+ show "integral\<^isup>P (M i) ?f = \<integral>\<^isup>+ y. f (x(i:=y)) \<partial>M i"
unfolding f'_eq by simp
qed
qed
@@ -1808,24 +1719,21 @@
lemma (in product_sigma_finite) product_positive_integral_setprod:
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> pextreal"
assumes "finite I" and borel: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
- shows "product_positive_integral I (\<lambda>x. (\<Prod>i\<in>I. f i (x i))) =
- (\<Prod>i\<in>I. M.positive_integral i (f i))"
+ shows "(\<integral>\<^isup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>P (M i) (f i))"
using assms proof induct
case empty
- interpret finite_product_sigma_finite M \<mu> "{}" by default auto
+ interpret finite_product_sigma_finite M "{}" by default auto
then show ?case by simp
next
case (insert i I)
note `finite I`[intro, simp]
- interpret I: finite_product_sigma_finite M \<mu> I by default auto
+ interpret I: finite_product_sigma_finite M I by default auto
have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
using insert by (auto intro!: setprod_cong)
- have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow>
- (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (sigma (product_algebra M J))"
+ have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^isub>M J M)"
using sets_into_space insert
- by (intro sigma_algebra.borel_measurable_pextreal_setprod
- sigma_algebra_sigma product_algebra_sets_into_space
- measurable_component)
+ by (intro sigma_algebra.borel_measurable_pextreal_setprod sigma_algebra_product_algebra
+ measurable_comp[OF measurable_component_singleton, unfolded comp_def])
auto
show ?case
by (simp add: product_positive_integral_insert[OF insert(1,2) prod])
@@ -1834,67 +1742,51 @@
lemma (in product_sigma_finite) product_integral_singleton:
assumes f: "f \<in> borel_measurable (M i)"
- shows "product_integral {i} (\<lambda>x. f (x i)) = M.integral i f"
+ shows "(\<integral>x. f (x i) \<partial>Pi\<^isub>M {i} M) = integral\<^isup>L (M i) f"
proof -
- interpret I: finite_product_sigma_finite M \<mu> "{i}" by default simp
+ interpret I: finite_product_sigma_finite M "{i}" by default simp
have *: "(\<lambda>x. Real (f x)) \<in> borel_measurable (M i)"
"(\<lambda>x. Real (- f x)) \<in> borel_measurable (M i)"
using assms by auto
show ?thesis
- unfolding I.integral_def integral_def
- unfolding *[THEN product_positive_integral_singleton] ..
+ unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] ..
qed
lemma (in product_sigma_finite) product_integral_fold:
assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
- and f: "measure_space.integrable (sigma (product_algebra M (I \<union> J))) (product_measure (I \<union> J)) f"
- shows "product_integral (I \<union> J) f =
- product_integral I (\<lambda>x. product_integral J (\<lambda>y. f (merge I x J y)))"
+ and f: "integrable (Pi\<^isub>M (I \<union> J) M) f"
+ shows "integral\<^isup>L (Pi\<^isub>M (I \<union> J) M) f = (\<integral>x. (\<integral>y. f (merge I x J y) \<partial>Pi\<^isub>M J M) \<partial>Pi\<^isub>M I M)"
proof -
- interpret I: finite_product_sigma_finite M \<mu> I by default fact
- interpret J: finite_product_sigma_finite M \<mu> J by default fact
+ interpret I: finite_product_sigma_finite M I by default fact
+ interpret J: finite_product_sigma_finite M J by default fact
have "finite (I \<union> J)" using fin by auto
- interpret IJ: finite_product_sigma_finite M \<mu> "I \<union> J" by default fact
- interpret P: pair_sigma_finite I.P I.measure J.P J.measure by default
- let ?f = "\<lambda>(x,y). f (merge I x J y)"
- have f_borel: "f \<in> borel_measurable IJ.P"
- "(\<lambda>x. Real (f x)) \<in> borel_measurable IJ.P"
- "(\<lambda>x. Real (- f x)) \<in> borel_measurable IJ.P"
- using f unfolding integrable_def by auto
- have f_restrict: "(\<lambda>x. f (restrict x (I \<union> J))) \<in> borel_measurable IJ.P"
- by (rule measurable_cong[THEN iffD2, OF _ f_borel(1)])
- (auto intro!: arg_cong[where f=f] simp: extensional_restrict)
- then have f'_borel:
- "(\<lambda>x. Real (?f x)) \<in> borel_measurable P.P"
- "(\<lambda>x. Real (- ?f x)) \<in> borel_measurable P.P"
- unfolding measurable_restrict_iff[OF IJ]
- by simp_all
- have PI:
- "P.positive_integral (\<lambda>x. Real (?f x)) = IJ.positive_integral (\<lambda>x. Real (f x))"
- "P.positive_integral (\<lambda>x. Real (- ?f x)) = IJ.positive_integral (\<lambda>x. Real (- f x))"
- using f'_borel[THEN P.positive_integral_fst_measurable(2)]
- using f_borel(2,3)[THEN product_positive_integral_fold[OF assms(1-3)]]
- by simp_all
- have "P.integrable ?f" using `IJ.integrable f`
- unfolding P.integrable_def IJ.integrable_def
- unfolding measurable_restrict_iff[OF IJ]
- using f_restrict PI by simp_all
+ interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
+ interpret P: pair_sigma_finite I.P J.P by default
+ let ?M = "\<lambda>(x, y). merge I x J y"
+ let ?f = "\<lambda>x. f (?M x)"
show ?thesis
- unfolding P.integrable_fst_measurable(2)[OF `P.integrable ?f`, simplified]
- unfolding IJ.integral_def P.integral_def
- unfolding PI by simp
+ proof (subst P.integrable_fst_measurable(2)[of ?f, simplified])
+ have 1: "sigma_algebra IJ.P" by default
+ have 2: "?M \<in> measurable P.P IJ.P" using measurable_merge[OF IJ] .
+ have 3: "\<And>A. A \<in> sets IJ.P \<Longrightarrow> IJ.\<mu> A = P.\<mu> (?M -` A \<inter> space P.P)"
+ by (rule measure_fold[OF IJ fin, symmetric])
+ have 4: "integrable (Pi\<^isub>M (I \<union> J) M) f" by fact
+ show "integrable P.P ?f"
+ by (rule P.integral_vimage[where f=f, OF 1 2 3 4])
+ show "integral\<^isup>L IJ.P f = integral\<^isup>L P.P ?f"
+ by (rule P.integral_vimage[where f=f, OF 1 2 3 4])
+ qed
qed
lemma (in product_sigma_finite) product_integral_insert:
assumes [simp]: "finite I" "i \<notin> I"
- and f: "measure_space.integrable (sigma (product_algebra M (insert i I))) (product_measure (insert i I)) f"
- shows "product_integral (insert i I) f
- = product_integral I (\<lambda>x. M.integral i (\<lambda>y. f (x(i:=y))))"
+ and f: "integrable (Pi\<^isub>M (insert i I) M) f"
+ shows "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = (\<integral>x. (\<integral>y. f (x(i:=y)) \<partial>M i) \<partial>Pi\<^isub>M I M)"
proof -
- interpret I: finite_product_sigma_finite M \<mu> I by default auto
- interpret I': finite_product_sigma_finite M \<mu> "insert i I" by default auto
- interpret i: finite_product_sigma_finite M \<mu> "{i}" by default auto
- interpret P: pair_sigma_algebra I.P i.P ..
+ interpret I: finite_product_sigma_finite M I by default auto
+ interpret I': finite_product_sigma_finite M "insert i I" by default auto
+ interpret i: finite_product_sigma_finite M "{i}" by default auto
+ interpret P: pair_sigma_finite I.P i.P ..
have IJ: "I \<inter> {i} = {}" by auto
show ?thesis
unfolding product_integral_fold[OF IJ, simplified, OF f]
@@ -1903,39 +1795,40 @@
let "?f y" = "f (restrict (x(i := y)) (insert i I))"
have f'_eq: "\<And>y. ?f y = f (x(i := y))"
using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def)
- have "f \<in> borel_measurable I'.P" using f unfolding I'.integrable_def by auto
- note fP = this[unfolded measurable_merge_iff[OF IJ, simplified]]
+ have f: "f \<in> borel_measurable I'.P" using f unfolding integrable_def by auto
show "?f \<in> borel_measurable (M i)"
- using P.measurable_pair_image_snd[OF fP x]
- unfolding measurable_singleton f'_eq by (simp add: f'_eq)
- show "M.integral i ?f = M.integral i (\<lambda>y. f (x(i := y)))"
+ unfolding measurable_cong[OF f'_eq]
+ using measurable_comp[OF measurable_component_update f] x `i \<notin> I`
+ by (simp add: comp_def)
+ show "integral\<^isup>L (M i) ?f = integral\<^isup>L (M i) (\<lambda>y. f (x(i := y)))"
unfolding f'_eq by simp
qed
qed
lemma (in product_sigma_finite) product_integrable_setprod:
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real"
- assumes [simp]: "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> M.integrable i (f i)"
- shows "product_integrable I (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "product_integrable I ?f")
+ assumes [simp]: "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
+ shows "integrable (Pi\<^isub>M I M) (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "integrable _ ?f")
proof -
- interpret finite_product_sigma_finite M \<mu> I by default fact
+ interpret finite_product_sigma_finite M I by default fact
have f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
- using integrable unfolding M.integrable_def by auto
+ using integrable unfolding integrable_def by auto
then have borel: "?f \<in> borel_measurable P"
- by (intro borel_measurable_setprod measurable_component) auto
- moreover have "integrable (\<lambda>x. \<bar>\<Prod>i\<in>I. f i (x i)\<bar>)"
+ using measurable_comp[OF measurable_component_singleton f]
+ by (auto intro!: borel_measurable_setprod simp: comp_def)
+ moreover have "integrable (Pi\<^isub>M I M) (\<lambda>x. \<bar>\<Prod>i\<in>I. f i (x i)\<bar>)"
proof (unfold integrable_def, intro conjI)
show "(\<lambda>x. abs (?f x)) \<in> borel_measurable P"
using borel by auto
- have "positive_integral (\<lambda>x. Real (abs (?f x))) = positive_integral (\<lambda>x. \<Prod>i\<in>I. Real (abs (f i (x i))))"
+ have "(\<integral>\<^isup>+x. Real (abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. (\<Prod>i\<in>I. Real (abs (f i (x i)))) \<partial>P)"
by (simp add: Real_setprod abs_setprod)
- also have "\<dots> = (\<Prod>i\<in>I. M.positive_integral i (\<lambda>x. Real (abs (f i x))))"
+ also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^isup>+x. Real (abs (f i x)) \<partial>M i))"
using f by (subst product_positive_integral_setprod) auto
also have "\<dots> < \<omega>"
using integrable[THEN M.integrable_abs]
- unfolding pextreal_less_\<omega> setprod_\<omega> M.integrable_def by simp
- finally show "positive_integral (\<lambda>x. Real (abs (?f x))) \<noteq> \<omega>" by auto
- show "positive_integral (\<lambda>x. Real (- abs (?f x))) \<noteq> \<omega>" by simp
+ unfolding pextreal_less_\<omega> setprod_\<omega> integrable_def by simp
+ finally show "(\<integral>\<^isup>+x. Real (abs (?f x)) \<partial>P) \<noteq> \<omega>" by auto
+ show "(\<integral>\<^isup>+x. Real (- abs (?f x)) \<partial>P) \<noteq> \<omega>" by simp
qed
ultimately show ?thesis
by (rule integrable_abs_iff[THEN iffD1])
@@ -1943,8 +1836,8 @@
lemma (in product_sigma_finite) product_integral_setprod:
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real"
- assumes "finite I" "I \<noteq> {}" and integrable: "\<And>i. i \<in> I \<Longrightarrow> M.integrable i (f i)"
- shows "product_integral I (\<lambda>x. (\<Prod>i\<in>I. f i (x i))) = (\<Prod>i\<in>I. M.integral i (f i))"
+ assumes "finite I" "I \<noteq> {}" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
+ shows "(\<integral>x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>L (M i) (f i))"
using assms proof (induct rule: finite_ne_induct)
case (singleton i)
then show ?case by (simp add: product_integral_singleton integrable_def)
@@ -1952,9 +1845,9 @@
case (insert i I)
then have iI: "finite (insert i I)" by auto
then have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow>
- product_integrable J (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))"
+ integrable (Pi\<^isub>M J M) (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))"
by (intro product_integrable_setprod insert(5)) (auto intro: finite_subset)
- interpret I: finite_product_sigma_finite M \<mu> I by default fact
+ interpret I: finite_product_sigma_finite M I by default fact
have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
using `i \<notin> I` by (auto intro!: setprod_cong)
show ?case
@@ -1964,9 +1857,9 @@
section "Products on finite spaces"
-lemma sigma_sets_pair_algebra_finite:
+lemma sigma_sets_pair_measure_generator_finite:
assumes "finite A" and "finite B"
- shows "sigma_sets (A \<times> B) ((\<lambda>(x,y). x \<times> y) ` (Pow A \<times> Pow B)) = Pow (A \<times> B)"
+ shows "sigma_sets (A \<times> B) { a \<times> b | a b. a \<in> Pow A \<and> b \<in> Pow B} = Pow (A \<times> B)"
(is "sigma_sets ?prod ?sets = _")
proof safe
have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product)
@@ -1978,7 +1871,7 @@
next
case (insert a x)
hence "{a} \<in> sigma_sets ?prod ?sets"
- by (auto simp: pair_algebra_def intro!: sigma_sets.Basic)
+ by (auto simp: pair_measure_generator_def intro!: sigma_sets.Basic)
moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto
ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un)
qed
@@ -1989,26 +1882,28 @@
show "a \<in> A" and "b \<in> B" by auto
qed
-locale pair_finite_sigma_algebra = M1: finite_sigma_algebra M1 + M2: finite_sigma_algebra M2 for M1 M2
+locale pair_finite_sigma_algebra = M1: finite_sigma_algebra M1 + M2: finite_sigma_algebra M2
+ for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
sublocale pair_finite_sigma_algebra \<subseteq> pair_sigma_algebra by default
-lemma (in pair_finite_sigma_algebra) finite_pair_sigma_algebra[simp]:
- shows "P = (| space = space M1 \<times> space M2, sets = Pow (space M1 \<times> space M2) |)"
+lemma (in pair_finite_sigma_algebra) finite_pair_sigma_algebra:
+ shows "P = \<lparr> space = space M1 \<times> space M2, sets = Pow (space M1 \<times> space M2), \<dots> = algebra.more P \<rparr>"
proof -
- show ?thesis using M1.finite_space M2.finite_space
- by (simp add: sigma_def space_pair_algebra sets_pair_algebra
- sigma_sets_pair_algebra_finite M1.sets_eq_Pow M2.sets_eq_Pow)
+ show ?thesis
+ using sigma_sets_pair_measure_generator_finite[OF M1.finite_space M2.finite_space]
+ by (intro algebra.equality) (simp_all add: pair_measure_def pair_measure_generator_def sigma_def)
qed
sublocale pair_finite_sigma_algebra \<subseteq> finite_sigma_algebra P
-proof
- show "finite (space P)" "sets P = Pow (space P)"
- using M1.finite_space M2.finite_space by auto
-qed
+ apply default
+ using M1.finite_space M2.finite_space
+ apply (subst finite_pair_sigma_algebra) apply simp
+ apply (subst (1 2) finite_pair_sigma_algebra) apply simp
+ done
-locale pair_finite_space = M1: finite_measure_space M1 \<mu>1 + M2: finite_measure_space M2 \<mu>2
- for M1 \<mu>1 M2 \<mu>2
+locale pair_finite_space = M1: finite_measure_space M1 + M2: finite_measure_space M2
+ for M1 M2
sublocale pair_finite_space \<subseteq> pair_finite_sigma_algebra
by default
@@ -2016,19 +1911,11 @@
sublocale pair_finite_space \<subseteq> pair_sigma_finite
by default
-lemma (in pair_finite_space) finite_pair_sigma_algebra[simp]:
- shows "P = (| space = space M1 \<times> space M2, sets = Pow (space M1 \<times> space M2) |)"
-proof -
- show ?thesis using M1.finite_space M2.finite_space
- by (simp add: sigma_def space_pair_algebra sets_pair_algebra
- sigma_sets_pair_algebra_finite M1.sets_eq_Pow M2.sets_eq_Pow)
-qed
-
lemma (in pair_finite_space) pair_measure_Pair[simp]:
assumes "a \<in> space M1" "b \<in> space M2"
- shows "pair_measure {(a, b)} = \<mu>1 {a} * \<mu>2 {b}"
+ shows "\<mu> {(a, b)} = M1.\<mu> {a} * M2.\<mu> {b}"
proof -
- have "pair_measure ({a}\<times>{b}) = \<mu>1 {a} * \<mu>2 {b}"
+ have "\<mu> ({a}\<times>{b}) = M1.\<mu> {a} * M2.\<mu> {b}"
using M1.sets_eq_Pow M2.sets_eq_Pow assms
by (subst pair_measure_times) auto
then show ?thesis by simp
@@ -2036,15 +1923,10 @@
lemma (in pair_finite_space) pair_measure_singleton[simp]:
assumes "x \<in> space M1 \<times> space M2"
- shows "pair_measure {x} = \<mu>1 {fst x} * \<mu>2 {snd x}"
+ shows "\<mu> {x} = M1.\<mu> {fst x} * M2.\<mu> {snd x}"
using pair_measure_Pair assms by (cases x) auto
-sublocale pair_finite_space \<subseteq> finite_measure_space P pair_measure
- by default auto
-
-lemma (in pair_finite_space) finite_measure_space_finite_prod_measure_alterantive:
- "finite_measure_space \<lparr>space = space M1 \<times> space M2, sets = Pow (space M1 \<times> space M2)\<rparr> pair_measure"
- unfolding finite_pair_sigma_algebra[symmetric]
- by default
+sublocale pair_finite_space \<subseteq> finite_measure_space P
+ by default (auto simp: space_pair_measure)
end
\ No newline at end of file
--- a/src/HOL/Probability/Radon_Nikodym.thy Wed Feb 02 10:35:41 2011 +0100
+++ b/src/HOL/Probability/Radon_Nikodym.thy Wed Feb 02 12:34:45 2011 +0100
@@ -11,7 +11,7 @@
qed auto
lemma (in sigma_finite_measure) Ex_finite_integrable_function:
- shows "\<exists>h\<in>borel_measurable M. positive_integral h \<noteq> \<omega> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<omega>)"
+ shows "\<exists>h\<in>borel_measurable M. integral\<^isup>P M h \<noteq> \<omega> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<omega>)"
proof -
obtain A :: "nat \<Rightarrow> 'a set" where
range: "range A \<subseteq> sets M" and
@@ -42,7 +42,7 @@
proof (safe intro!: bexI[of _ ?h] del: notI)
have "\<And>i. A i \<in> sets M"
using range by fastsimp+
- then have "positive_integral ?h = (\<Sum>\<^isub>\<infinity> i. n i * \<mu> (A i))"
+ then have "integral\<^isup>P M ?h = (\<Sum>\<^isub>\<infinity> i. n i * \<mu> (A i))"
by (simp add: positive_integral_psuminf positive_integral_cmult_indicator)
also have "\<dots> \<le> (\<Sum>\<^isub>\<infinity> i. Real ((1 / 2)^Suc i))"
proof (rule psuminf_le)
@@ -56,7 +56,7 @@
qed
also have "\<dots> = Real 1"
by (rule suminf_imp_psuminf, rule power_half_series, auto)
- finally show "positive_integral ?h \<noteq> \<omega>" by auto
+ finally show "integral\<^isup>P M ?h \<noteq> \<omega>" by auto
next
fix x assume "x \<in> space M"
then obtain i where "x \<in> A i" using space[symmetric] by auto
@@ -75,46 +75,47 @@
"absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: pextreal))"
lemma (in sigma_finite_measure) absolutely_continuous_AE:
- assumes "measure_space M \<nu>" "absolutely_continuous \<nu>" "AE x. P x"
- shows "measure_space.almost_everywhere M \<nu> P"
+ assumes "measure_space M'" and [simp]: "sets M' = sets M" "space M' = space M"
+ and "absolutely_continuous (measure M')" "AE x. P x"
+ shows "measure_space.almost_everywhere M' P"
proof -
- interpret \<nu>: measure_space M \<nu> by fact
+ interpret \<nu>: measure_space M' by fact
from `AE x. P x` obtain N where N: "N \<in> null_sets" and "{x\<in>space M. \<not> P x} \<subseteq> N"
unfolding almost_everywhere_def by auto
show "\<nu>.almost_everywhere P"
proof (rule \<nu>.AE_I')
- show "{x\<in>space M. \<not> P x} \<subseteq> N" by fact
- from `absolutely_continuous \<nu>` show "N \<in> \<nu>.null_sets"
+ show "{x\<in>space M'. \<not> P x} \<subseteq> N" by simp fact
+ from `absolutely_continuous (measure M')` show "N \<in> \<nu>.null_sets"
using N unfolding absolutely_continuous_def by auto
qed
qed
lemma (in finite_measure_space) absolutely_continuousI:
- assumes "finite_measure_space M \<nu>"
+ assumes "finite_measure_space (M\<lparr> measure := \<nu>\<rparr>)" (is "finite_measure_space ?\<nu>")
assumes v: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0"
shows "absolutely_continuous \<nu>"
proof (unfold absolutely_continuous_def sets_eq_Pow, safe)
fix N assume "\<mu> N = 0" "N \<subseteq> space M"
- interpret v: finite_measure_space M \<nu> by fact
- have "\<nu> N = \<nu> (\<Union>x\<in>N. {x})" by simp
- also have "\<dots> = (\<Sum>x\<in>N. \<nu> {x})"
+ interpret v: finite_measure_space ?\<nu> by fact
+ have "\<nu> N = measure ?\<nu> (\<Union>x\<in>N. {x})" by simp
+ also have "\<dots> = (\<Sum>x\<in>N. measure ?\<nu> {x})"
proof (rule v.measure_finitely_additive''[symmetric])
show "finite N" using `N \<subseteq> space M` finite_space by (auto intro: finite_subset)
show "disjoint_family_on (\<lambda>i. {i}) N" unfolding disjoint_family_on_def by auto
- fix x assume "x \<in> N" thus "{x} \<in> sets M" using `N \<subseteq> space M` sets_eq_Pow by auto
+ fix x assume "x \<in> N" thus "{x} \<in> sets ?\<nu>" using `N \<subseteq> space M` sets_eq_Pow by auto
qed
also have "\<dots> = 0"
proof (safe intro!: setsum_0')
fix x assume "x \<in> N"
hence "\<mu> {x} \<le> \<mu> N" using sets_eq_Pow `N \<subseteq> space M` by (auto intro!: measure_mono)
hence "\<mu> {x} = 0" using `\<mu> N = 0` by simp
- thus "\<nu> {x} = 0" using v[of x] `x \<in> N` `N \<subseteq> space M` by auto
+ thus "measure ?\<nu> {x} = 0" using v[of x] `x \<in> N` `N \<subseteq> space M` by auto
qed
- finally show "\<nu> N = 0" .
+ finally show "\<nu> N = 0" by simp
qed
lemma (in measure_space) density_is_absolutely_continuous:
- assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x)"
+ assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
shows "absolutely_continuous \<nu>"
using assms unfolding absolutely_continuous_def
by (simp add: positive_integral_null_set)
@@ -123,13 +124,13 @@
lemma (in finite_measure) Radon_Nikodym_aux_epsilon:
fixes e :: real assumes "0 < e"
- assumes "finite_measure M s"
- shows "\<exists>A\<in>sets M. real (\<mu> (space M)) - real (s (space M)) \<le>
- real (\<mu> A) - real (s A) \<and>
- (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> - e < real (\<mu> B) - real (s B))"
+ assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)"
+ shows "\<exists>A\<in>sets M. real (\<mu> (space M)) - real (\<nu> (space M)) \<le>
+ real (\<mu> A) - real (\<nu> A) \<and>
+ (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> - e < real (\<mu> B) - real (\<nu> B))"
proof -
- let "?d A" = "real (\<mu> A) - real (s A)"
- interpret M': finite_measure M s by fact
+ let "?d A" = "real (\<mu> A) - real (\<nu> A)"
+ interpret M': finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact
let "?A A" = "if (\<forall>B\<in>sets M. B \<subseteq> space M - A \<longrightarrow> -e < ?d B)
then {}
else (SOME B. B \<in> sets M \<and> B \<subseteq> space M - A \<and> ?d B \<le> -e)"
@@ -216,21 +217,24 @@
qed
lemma (in finite_measure) Radon_Nikodym_aux:
- assumes "finite_measure M s"
- shows "\<exists>A\<in>sets M. real (\<mu> (space M)) - real (s (space M)) \<le>
- real (\<mu> A) - real (s A) \<and>
- (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> 0 \<le> real (\<mu> B) - real (s B))"
+ assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" (is "finite_measure ?M'")
+ shows "\<exists>A\<in>sets M. real (\<mu> (space M)) - real (\<nu> (space M)) \<le>
+ real (\<mu> A) - real (\<nu> A) \<and>
+ (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> 0 \<le> real (\<mu> B) - real (\<nu> B))"
proof -
- let "?d A" = "real (\<mu> A) - real (s A)"
+ let "?d A" = "real (\<mu> A) - real (\<nu> A)"
let "?P A B n" = "A \<in> sets M \<and> A \<subseteq> B \<and> ?d B \<le> ?d A \<and> (\<forall>C\<in>sets M. C \<subseteq> A \<longrightarrow> - 1 / real (Suc n) < ?d C)"
- interpret M': finite_measure M s by fact
+ interpret M': finite_measure ?M' where
+ "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \<nu>" by fact auto
let "?r S" = "restricted_space S"
{ fix S n
assume S: "S \<in> sets M"
hence **: "\<And>X. X \<in> op \<inter> S ` sets M \<longleftrightarrow> X \<in> sets M \<and> X \<subseteq> S" by auto
- from M'.restricted_finite_measure[of S] restricted_finite_measure[of S] S
- have "finite_measure (?r S) \<mu>" "0 < 1 / real (Suc n)"
- "finite_measure (?r S) s" by auto
+ have [simp]: "(restricted_space S\<lparr>measure := \<nu>\<rparr>) = M'.restricted_space S"
+ by (cases M) simp
+ from M'.restricted_finite_measure[of S] restricted_finite_measure[OF S] S
+ have "finite_measure (?r S)" "0 < 1 / real (Suc n)"
+ "finite_measure (?r S\<lparr>measure := \<nu>\<rparr>)" by auto
from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X ..
hence "?P X S n"
proof (simp add: **, safe)
@@ -287,12 +291,14 @@
qed
lemma (in finite_measure) Radon_Nikodym_finite_measure:
- assumes "finite_measure M \<nu>"
+ assumes "finite_measure (M\<lparr> measure := \<nu>\<rparr>)" (is "finite_measure ?M'")
assumes "absolutely_continuous \<nu>"
- shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x)"
+ shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
proof -
- interpret M': finite_measure M \<nu> using assms(1) .
- def G \<equiv> "{g \<in> borel_measurable M. \<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x) \<le> \<nu> A}"
+ interpret M': finite_measure ?M'
+ where "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \<nu>"
+ using assms(1) by auto
+ def G \<equiv> "{g \<in> borel_measurable M. \<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> \<nu> A}"
have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto
hence "G \<noteq> {}" by auto
{ fix f g assume f: "f \<in> G" and g: "g \<in> G"
@@ -308,16 +314,16 @@
have "\<And>x. x \<in> space M \<Longrightarrow> max (g x) (f x) * indicator A x =
g x * indicator (?A \<inter> A) x + f x * indicator ((space M - ?A) \<inter> A) x"
by (auto simp: indicator_def max_def)
- hence "(\<integral>\<^isup>+x. max (g x) (f x) * indicator A x) =
- (\<integral>\<^isup>+x. g x * indicator (?A \<inter> A) x) +
- (\<integral>\<^isup>+x. f x * indicator ((space M - ?A) \<inter> A) x)"
+ hence "(\<integral>\<^isup>+x. max (g x) (f x) * indicator A x \<partial>M) =
+ (\<integral>\<^isup>+x. g x * indicator (?A \<inter> A) x \<partial>M) +
+ (\<integral>\<^isup>+x. f x * indicator ((space M - ?A) \<inter> A) x \<partial>M)"
using f g sets unfolding G_def
by (auto cong: positive_integral_cong intro!: positive_integral_add borel_measurable_indicator)
also have "\<dots> \<le> \<nu> (?A \<inter> A) + \<nu> ((space M - ?A) \<inter> A)"
using f g sets unfolding G_def by (auto intro!: add_mono)
also have "\<dots> = \<nu> A"
using M'.measure_additive[OF sets] union by auto
- finally show "(\<integral>\<^isup>+x. max (g x) (f x) * indicator A x) \<le> \<nu> A" .
+ finally show "(\<integral>\<^isup>+x. max (g x) (f x) * indicator A x \<partial>M) \<le> \<nu> A" .
qed }
note max_in_G = this
{ fix f g assume "f \<up> g" and f: "\<And>i. f i \<in> G"
@@ -331,30 +337,30 @@
hence "\<And>i. (\<lambda>x. f i x * indicator A x) \<in> borel_measurable M"
using f_borel by (auto intro!: borel_measurable_indicator)
from positive_integral_isoton[OF isoton_indicator[OF `f \<up> g`] this]
- have SUP: "(\<integral>\<^isup>+x. g x * indicator A x) =
- (SUP i. (\<integral>\<^isup>+x. f i x * indicator A x))"
+ have SUP: "(\<integral>\<^isup>+x. g x * indicator A x \<partial>M) =
+ (SUP i. (\<integral>\<^isup>+x. f i x * indicator A x \<partial>M))"
unfolding isoton_def by simp
- show "(\<integral>\<^isup>+x. g x * indicator A x) \<le> \<nu> A" unfolding SUP
+ show "(\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> \<nu> A" unfolding SUP
using f `A \<in> sets M` unfolding G_def by (auto intro!: SUP_leI)
qed }
note SUP_in_G = this
- let ?y = "SUP g : G. positive_integral g"
+ let ?y = "SUP g : G. integral\<^isup>P M g"
have "?y \<le> \<nu> (space M)" unfolding G_def
proof (safe intro!: SUP_leI)
- fix g assume "\<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x) \<le> \<nu> A"
- from this[THEN bspec, OF top] show "positive_integral g \<le> \<nu> (space M)"
+ fix g assume "\<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> \<nu> A"
+ from this[THEN bspec, OF top] show "integral\<^isup>P M g \<le> \<nu> (space M)"
by (simp cong: positive_integral_cong)
qed
hence "?y \<noteq> \<omega>" using M'.finite_measure_of_space by auto
from SUPR_countable_SUPR[OF this `G \<noteq> {}`] guess ys .. note ys = this
- hence "\<forall>n. \<exists>g. g\<in>G \<and> positive_integral g = ys n"
+ hence "\<forall>n. \<exists>g. g\<in>G \<and> integral\<^isup>P M g = ys n"
proof safe
- fix n assume "range ys \<subseteq> positive_integral ` G"
- hence "ys n \<in> positive_integral ` G" by auto
- thus "\<exists>g. g\<in>G \<and> positive_integral g = ys n" by auto
+ fix n assume "range ys \<subseteq> integral\<^isup>P M ` G"
+ hence "ys n \<in> integral\<^isup>P M ` G" by auto
+ thus "\<exists>g. g\<in>G \<and> integral\<^isup>P M g = ys n" by auto
qed
- from choice[OF this] obtain gs where "\<And>i. gs i \<in> G" "\<And>n. positive_integral (gs n) = ys n" by auto
- hence y_eq: "?y = (SUP i. positive_integral (gs i))" using ys by auto
+ from choice[OF this] obtain gs where "\<And>i. gs i \<in> G" "\<And>n. integral\<^isup>P M (gs n) = ys n" by auto
+ hence y_eq: "?y = (SUP i. integral\<^isup>P M (gs i))" using ys by auto
let "?g i x" = "Max ((\<lambda>n. gs n x) ` {..i})"
def f \<equiv> "SUP i. ?g i"
have gs_not_empty: "\<And>i. (\<lambda>n. gs n x) ` {..i} \<noteq> {}" by auto
@@ -372,53 +378,53 @@
hence isoton_g: "?g \<up> f" by (simp add: isoton_def le_fun_def f_def)
from SUP_in_G[OF this g_in_G] have "f \<in> G" .
hence [simp, intro]: "f \<in> borel_measurable M" unfolding G_def by auto
- have "(\<lambda>i. positive_integral (?g i)) \<up> positive_integral f"
+ have "(\<lambda>i. integral\<^isup>P M (?g i)) \<up> integral\<^isup>P M f"
using isoton_g g_in_G by (auto intro!: positive_integral_isoton simp: G_def f_def)
- hence "positive_integral f = (SUP i. positive_integral (?g i))"
+ hence "integral\<^isup>P M f = (SUP i. integral\<^isup>P M (?g i))"
unfolding isoton_def by simp
also have "\<dots> = ?y"
proof (rule antisym)
- show "(SUP i. positive_integral (?g i)) \<le> ?y"
+ show "(SUP i. integral\<^isup>P M (?g i)) \<le> ?y"
using g_in_G by (auto intro!: exI Sup_mono simp: SUPR_def)
- show "?y \<le> (SUP i. positive_integral (?g i))" unfolding y_eq
+ show "?y \<le> (SUP i. integral\<^isup>P M (?g i))" unfolding y_eq
by (auto intro!: SUP_mono positive_integral_mono Max_ge)
qed
- finally have int_f_eq_y: "positive_integral f = ?y" .
- let "?t A" = "\<nu> A - (\<integral>\<^isup>+x. f x * indicator A x)"
- have "finite_measure M ?t"
- proof
- show "?t {} = 0" by simp
- show "?t (space M) \<noteq> \<omega>" using M'.finite_measure by simp
- show "countably_additive M ?t" unfolding countably_additive_def
- proof safe
- fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "disjoint_family A"
- have "(\<Sum>\<^isub>\<infinity> n. (\<integral>\<^isup>+x. f x * indicator (A n) x))
- = (\<integral>\<^isup>+x. (\<Sum>\<^isub>\<infinity>n. f x * indicator (A n) x))"
- using `range A \<subseteq> sets M`
- by (rule_tac positive_integral_psuminf[symmetric]) (auto intro!: borel_measurable_indicator)
- also have "\<dots> = (\<integral>\<^isup>+x. f x * indicator (\<Union>n. A n) x)"
- apply (rule positive_integral_cong)
- apply (subst psuminf_cmult_right)
- unfolding psuminf_indicator[OF `disjoint_family A`] ..
- finally have "(\<Sum>\<^isub>\<infinity> n. (\<integral>\<^isup>+x. f x * indicator (A n) x))
- = (\<integral>\<^isup>+x. f x * indicator (\<Union>n. A n) x)" .
- moreover have "(\<Sum>\<^isub>\<infinity>n. \<nu> (A n)) = \<nu> (\<Union>n. A n)"
- using M'.measure_countably_additive A by (simp add: comp_def)
- moreover have "\<And>i. (\<integral>\<^isup>+x. f x * indicator (A i) x) \<le> \<nu> (A i)"
- using A `f \<in> G` unfolding G_def by auto
- moreover have v_fin: "\<nu> (\<Union>i. A i) \<noteq> \<omega>" using M'.finite_measure A by (simp add: countable_UN)
- moreover {
- have "(\<integral>\<^isup>+x. f x * indicator (\<Union>i. A i) x) \<le> \<nu> (\<Union>i. A i)"
- using A `f \<in> G` unfolding G_def by (auto simp: countable_UN)
- also have "\<nu> (\<Union>i. A i) < \<omega>" using v_fin by (simp add: pextreal_less_\<omega>)
- finally have "(\<integral>\<^isup>+x. f x * indicator (\<Union>i. A i) x) \<noteq> \<omega>"
- by (simp add: pextreal_less_\<omega>) }
- ultimately
- show "(\<Sum>\<^isub>\<infinity> n. ?t (A n)) = ?t (\<Union>i. A i)"
- apply (subst psuminf_minus) by simp_all
- qed
+ finally have int_f_eq_y: "integral\<^isup>P M f = ?y" .
+ let "?t A" = "\<nu> A - (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
+ let ?M = "M\<lparr> measure := ?t\<rparr>"
+ interpret M: sigma_algebra ?M
+ by (intro sigma_algebra_cong) auto
+ have fmM: "finite_measure ?M"
+ proof (default, simp_all add: countably_additive_def, safe)
+ fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "disjoint_family A"
+ have "(\<Sum>\<^isub>\<infinity> n. (\<integral>\<^isup>+x. f x * indicator (A n) x \<partial>M))
+ = (\<integral>\<^isup>+x. (\<Sum>\<^isub>\<infinity>n. f x * indicator (A n) x) \<partial>M)"
+ using `range A \<subseteq> sets M`
+ by (rule_tac positive_integral_psuminf[symmetric]) (auto intro!: borel_measurable_indicator)
+ also have "\<dots> = (\<integral>\<^isup>+x. f x * indicator (\<Union>n. A n) x \<partial>M)"
+ apply (rule positive_integral_cong)
+ apply (subst psuminf_cmult_right)
+ unfolding psuminf_indicator[OF `disjoint_family A`] ..
+ finally have "(\<Sum>\<^isub>\<infinity> n. (\<integral>\<^isup>+x. f x * indicator (A n) x \<partial>M))
+ = (\<integral>\<^isup>+x. f x * indicator (\<Union>n. A n) x \<partial>M)" .
+ moreover have "(\<Sum>\<^isub>\<infinity>n. \<nu> (A n)) = \<nu> (\<Union>n. A n)"
+ using M'.measure_countably_additive A by (simp add: comp_def)
+ moreover have "\<And>i. (\<integral>\<^isup>+x. f x * indicator (A i) x \<partial>M) \<le> \<nu> (A i)"
+ using A `f \<in> G` unfolding G_def by auto
+ moreover have v_fin: "\<nu> (\<Union>i. A i) \<noteq> \<omega>" using M'.finite_measure A by (simp add: countable_UN)
+ moreover {
+ have "(\<integral>\<^isup>+x. f x * indicator (\<Union>i. A i) x \<partial>M) \<le> \<nu> (\<Union>i. A i)"
+ using A `f \<in> G` unfolding G_def by (auto simp: countable_UN)
+ also have "\<nu> (\<Union>i. A i) < \<omega>" using v_fin by (simp add: pextreal_less_\<omega>)
+ finally have "(\<integral>\<^isup>+x. f x * indicator (\<Union>i. A i) x \<partial>M) \<noteq> \<omega>"
+ by (simp add: pextreal_less_\<omega>) }
+ ultimately
+ show "(\<Sum>\<^isub>\<infinity> n. ?t (A n)) = ?t (\<Union>i. A i)"
+ apply (subst psuminf_minus) by simp_all
qed
- then interpret M: finite_measure M ?t .
+ then interpret M: finite_measure ?M
+ where "space ?M = space M" and "sets ?M = sets M" and "measure ?M = ?t"
+ by (simp_all add: fmM)
have ac: "absolutely_continuous ?t" using `absolutely_continuous \<nu>` unfolding absolutely_continuous_def by auto
have upper_bound: "\<forall>A\<in>sets M. ?t A \<le> 0"
proof (rule ccontr)
@@ -433,23 +439,21 @@
hence pos_M: "0 < \<mu> (space M)"
using ac top unfolding absolutely_continuous_def by auto
moreover
- have "(\<integral>\<^isup>+x. f x * indicator (space M) x) \<le> \<nu> (space M)"
+ have "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<le> \<nu> (space M)"
using `f \<in> G` unfolding G_def by auto
- hence "(\<integral>\<^isup>+x. f x * indicator (space M) x) \<noteq> \<omega>"
+ hence "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<noteq> \<omega>"
using M'.finite_measure_of_space by auto
moreover
def b \<equiv> "?t (space M) / \<mu> (space M) / 2"
ultimately have b: "b \<noteq> 0" "b \<noteq> \<omega>"
using M'.finite_measure_of_space
by (auto simp: pextreal_inverse_eq_0 finite_measure_of_space)
- have "finite_measure M (\<lambda>A. b * \<mu> A)" (is "finite_measure M ?b")
- proof
- show "?b {} = 0" by simp
- show "?b (space M) \<noteq> \<omega>" using finite_measure_of_space b by auto
- show "countably_additive M ?b"
- unfolding countably_additive_def psuminf_cmult_right
- using measure_countably_additive by auto
- qed
+ let ?Mb = "?M\<lparr>measure := \<lambda>A. b * \<mu> A\<rparr>"
+ interpret b: sigma_algebra ?Mb by (intro sigma_algebra_cong) auto
+ have "finite_measure ?Mb"
+ by default
+ (insert finite_measure_of_space b measure_countably_additive,
+ auto simp: psuminf_cmult_right countably_additive_def)
from M.Radon_Nikodym_aux[OF this]
obtain A0 where "A0 \<in> sets M" and
space_less_A0: "real (?t (space M)) - real (b * \<mu> (space M)) \<le> real (?t A0) - real (b * \<mu> A0)" and
@@ -462,30 +466,30 @@
let "?f0 x" = "f x + b * indicator A0 x"
{ fix A assume A: "A \<in> sets M"
hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
- have "(\<integral>\<^isup>+x. ?f0 x * indicator A x) =
- (\<integral>\<^isup>+x. f x * indicator A x + b * indicator (A \<inter> A0) x)"
+ have "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) =
+ (\<integral>\<^isup>+x. f x * indicator A x + b * indicator (A \<inter> A0) x \<partial>M)"
by (auto intro!: positive_integral_cong simp: field_simps indicator_inter_arith)
- hence "(\<integral>\<^isup>+x. ?f0 x * indicator A x) =
- (\<integral>\<^isup>+x. f x * indicator A x) + b * \<mu> (A \<inter> A0)"
+ hence "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) =
+ (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + b * \<mu> (A \<inter> A0)"
using `A0 \<in> sets M` `A \<inter> A0 \<in> sets M` A
by (simp add: borel_measurable_indicator positive_integral_add positive_integral_cmult_indicator) }
note f0_eq = this
{ fix A assume A: "A \<in> sets M"
hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
- have f_le_v: "(\<integral>\<^isup>+x. f x * indicator A x) \<le> \<nu> A"
+ have f_le_v: "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) \<le> \<nu> A"
using `f \<in> G` A unfolding G_def by auto
note f0_eq[OF A]
- also have "(\<integral>\<^isup>+x. f x * indicator A x) + b * \<mu> (A \<inter> A0) \<le>
- (\<integral>\<^isup>+x. f x * indicator A x) + ?t (A \<inter> A0)"
+ also have "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + b * \<mu> (A \<inter> A0) \<le>
+ (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + ?t (A \<inter> A0)"
using bM_le_t[OF `A \<inter> A0 \<in> sets M`] `A \<in> sets M` `A0 \<in> sets M`
by (auto intro!: add_left_mono)
- also have "\<dots> \<le> (\<integral>\<^isup>+x. f x * indicator A x) + ?t A"
+ also have "\<dots> \<le> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + ?t A"
using M.measure_mono[simplified, OF _ `A \<inter> A0 \<in> sets M` `A \<in> sets M`]
by (auto intro!: add_left_mono)
also have "\<dots> \<le> \<nu> A"
using f_le_v M'.finite_measure[simplified, OF `A \<in> sets M`]
- by (cases "(\<integral>\<^isup>+x. f x * indicator A x)", cases "\<nu> A", auto)
- finally have "(\<integral>\<^isup>+x. ?f0 x * indicator A x) \<le> \<nu> A" . }
+ by (cases "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M)", cases "\<nu> A", auto)
+ finally have "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) \<le> \<nu> A" . }
hence "?f0 \<in> G" using `A0 \<in> sets M` unfolding G_def
by (auto intro!: borel_measurable_indicator borel_measurable_pextreal_add borel_measurable_pextreal_times)
have real: "?t (space M) \<noteq> \<omega>" "?t A0 \<noteq> \<omega>"
@@ -494,7 +498,7 @@
finite_measure[of A0] M.finite_measure[of A0]
finite_measure_of_space M.finite_measure_of_space
by auto
- have int_f_finite: "positive_integral f \<noteq> \<omega>"
+ have int_f_finite: "integral\<^isup>P M f \<noteq> \<omega>"
using M'.finite_measure_of_space pos_t unfolding pextreal_zero_less_diff_iff
by (auto cong: positive_integral_cong)
have "?t (space M) > b * \<mu> (space M)" unfolding b_def
@@ -514,22 +518,22 @@
using `A0 \<in> sets M` by auto
hence "0 < b * \<mu> A0" using b by auto
from int_f_finite this
- have "?y + 0 < positive_integral f + b * \<mu> A0" unfolding int_f_eq_y
+ have "?y + 0 < integral\<^isup>P M f + b * \<mu> A0" unfolding int_f_eq_y
by (rule pextreal_less_add)
- also have "\<dots> = positive_integral ?f0" using f0_eq[OF top] `A0 \<in> sets M` sets_into_space
+ also have "\<dots> = integral\<^isup>P M ?f0" using f0_eq[OF top] `A0 \<in> sets M` sets_into_space
by (simp cong: positive_integral_cong)
- finally have "?y < positive_integral ?f0" by simp
- moreover from `?f0 \<in> G` have "positive_integral ?f0 \<le> ?y" by (auto intro!: le_SUPI)
+ finally have "?y < integral\<^isup>P M ?f0" by simp
+ moreover from `?f0 \<in> G` have "integral\<^isup>P M ?f0 \<le> ?y" by (auto intro!: le_SUPI)
ultimately show False by auto
qed
show ?thesis
proof (safe intro!: bexI[of _ f])
fix A assume "A\<in>sets M"
- show "\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x)"
+ show "\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
proof (rule antisym)
- show "(\<integral>\<^isup>+x. f x * indicator A x) \<le> \<nu> A"
+ show "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) \<le> \<nu> A"
using `f \<in> G` `A \<in> sets M` unfolding G_def by auto
- show "\<nu> A \<le> (\<integral>\<^isup>+x. f x * indicator A x)"
+ show "\<nu> A \<le> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
using upper_bound[THEN bspec, OF `A \<in> sets M`]
by (simp add: pextreal_zero_le_diff)
qed
@@ -537,13 +541,15 @@
qed
lemma (in finite_measure) split_space_into_finite_sets_and_rest:
- assumes "measure_space M \<nu>"
+ assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
assumes ac: "absolutely_continuous \<nu>"
shows "\<exists>\<Omega>0\<in>sets M. \<exists>\<Omega>::nat\<Rightarrow>'a set. disjoint_family \<Omega> \<and> range \<Omega> \<subseteq> sets M \<and> \<Omega>0 = space M - (\<Union>i. \<Omega> i) \<and>
(\<forall>A\<in>sets M. A \<subseteq> \<Omega>0 \<longrightarrow> (\<mu> A = 0 \<and> \<nu> A = 0) \<or> (\<mu> A > 0 \<and> \<nu> A = \<omega>)) \<and>
(\<forall>i. \<nu> (\<Omega> i) \<noteq> \<omega>)"
proof -
- interpret v: measure_space M \<nu> by fact
+ interpret v: measure_space ?N
+ where "space ?N = space M" and "sets ?N = sets M" and "measure ?N = \<nu>"
+ by fact auto
let ?Q = "{Q\<in>sets M. \<nu> Q \<noteq> \<omega>}"
let ?a = "SUP Q:?Q. \<mu> Q"
have "{} \<in> ?Q" using v.empty_measure by auto
@@ -667,11 +673,13 @@
qed
lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite:
- assumes "measure_space M \<nu>"
+ assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
assumes "absolutely_continuous \<nu>"
- shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x)"
+ shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
proof -
- interpret v: measure_space M \<nu> by fact
+ interpret v: measure_space ?N
+ where "space ?N = space M" and "sets ?N = sets M" and "measure ?N = \<nu>"
+ by fact auto
from split_space_into_finite_sets_and_rest[OF assms]
obtain Q0 and Q :: "nat \<Rightarrow> 'a set"
where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
@@ -680,39 +688,38 @@
and Q_fin: "\<And>i. \<nu> (Q i) \<noteq> \<omega>" by force
from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
have "\<forall>i. \<exists>f. f\<in>borel_measurable M \<and> (\<forall>A\<in>sets M.
- \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x))"
+ \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x \<partial>M))"
proof
fix i
have indicator_eq: "\<And>f x A. (f x :: pextreal) * indicator (Q i \<inter> A) x * indicator (Q i) x
= (f x * indicator (Q i) x) * indicator A x"
unfolding indicator_def by auto
- have fm: "finite_measure (restricted_space (Q i)) \<mu>"
- (is "finite_measure ?R \<mu>") by (rule restricted_finite_measure[OF Q_sets[of i]])
+ have fm: "finite_measure (restricted_space (Q i))"
+ (is "finite_measure ?R") by (rule restricted_finite_measure[OF Q_sets[of i]])
then interpret R: finite_measure ?R .
- have fmv: "finite_measure ?R \<nu>"
+ have fmv: "finite_measure (restricted_space (Q i) \<lparr> measure := \<nu>\<rparr>)" (is "finite_measure ?Q")
unfolding finite_measure_def finite_measure_axioms_def
proof
- show "measure_space ?R \<nu>"
+ show "measure_space ?Q"
using v.restricted_measure_space Q_sets[of i] by auto
- show "\<nu> (space ?R) \<noteq> \<omega>"
- using Q_fin by simp
+ show "measure ?Q (space ?Q) \<noteq> \<omega>" using Q_fin by simp
qed
have "R.absolutely_continuous \<nu>"
using `absolutely_continuous \<nu>` `Q i \<in> sets M`
by (auto simp: R.absolutely_continuous_def absolutely_continuous_def)
- from finite_measure.Radon_Nikodym_finite_measure[OF fm fmv this]
+ from R.Radon_Nikodym_finite_measure[OF fmv this]
obtain f where f: "(\<lambda>x. f x * indicator (Q i) x) \<in> borel_measurable M"
- and f_int: "\<And>A. A\<in>sets M \<Longrightarrow> \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. (f x * indicator (Q i) x) * indicator A x)"
+ and f_int: "\<And>A. A\<in>sets M \<Longrightarrow> \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. (f x * indicator (Q i) x) * indicator A x \<partial>M)"
unfolding Bex_def borel_measurable_restricted[OF `Q i \<in> sets M`]
positive_integral_restricted[OF `Q i \<in> sets M`] by (auto simp: indicator_eq)
then show "\<exists>f. f\<in>borel_measurable M \<and> (\<forall>A\<in>sets M.
- \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x))"
+ \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x \<partial>M))"
by (fastsimp intro!: exI[of _ "\<lambda>x. f x * indicator (Q i) x"] positive_integral_cong
simp: indicator_def)
qed
from choice[OF this] obtain f where borel: "\<And>i. f i \<in> borel_measurable M"
and f: "\<And>A i. A \<in> sets M \<Longrightarrow>
- \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f i x * indicator (Q i \<inter> A) x)"
+ \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M)"
by auto
let "?f x" =
"(\<Sum>\<^isub>\<infinity> i. f i x * indicator (Q i) x) + \<omega> * indicator Q0 x"
@@ -728,7 +735,7 @@
f i x * indicator (Q i \<inter> A) x"
"\<And>x i. (indicator A x * indicator Q0 x :: pextreal) =
indicator (Q0 \<inter> A) x" by (auto simp: indicator_def)
- have "(\<integral>\<^isup>+x. ?f x * indicator A x) =
+ have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) =
(\<Sum>\<^isub>\<infinity> i. \<nu> (Q i \<inter> A)) + \<omega> * \<mu> (Q0 \<inter> A)"
unfolding f[OF `A \<in> sets M`]
apply (simp del: pextreal_times(2) add: field_simps *)
@@ -755,27 +762,29 @@
using Q_sets `A \<in> sets M` Q0(1) by (auto intro!: countable_UN)
moreover have "((\<Union>i. Q i) \<inter> A) \<union> (Q0 \<inter> A) = A" "((\<Union>i. Q i) \<inter> A) \<inter> (Q0 \<inter> A) = {}"
using `A \<in> sets M` sets_into_space Q0 by auto
- ultimately show "\<nu> A = (\<integral>\<^isup>+x. ?f x * indicator A x)"
+ ultimately show "\<nu> A = (\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M)"
using v.measure_additive[simplified, of "(\<Union>i. Q i) \<inter> A" "Q0 \<inter> A"]
by simp
qed
qed
lemma (in sigma_finite_measure) Radon_Nikodym:
- assumes "measure_space M \<nu>"
+ assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
assumes "absolutely_continuous \<nu>"
- shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x)"
+ shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
proof -
from Ex_finite_integrable_function
- obtain h where finite: "positive_integral h \<noteq> \<omega>" and
+ obtain h where finite: "integral\<^isup>P M h \<noteq> \<omega>" and
borel: "h \<in> borel_measurable M" and
pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x" and
"\<And>x. x \<in> space M \<Longrightarrow> h x < \<omega>" by auto
- let "?T A" = "(\<integral>\<^isup>+x. h x * indicator A x)"
+ let "?T A" = "(\<integral>\<^isup>+x. h x * indicator A x \<partial>M)"
+ let ?MT = "M\<lparr> measure := ?T \<rparr>"
from measure_space_density[OF borel] finite
- interpret T: finite_measure M ?T
+ interpret T: finite_measure ?MT
+ where "space ?MT = space M" and "sets ?MT = sets M" and "measure ?MT = ?T"
unfolding finite_measure_def finite_measure_axioms_def
- by (simp cong: positive_integral_cong)
+ by (simp_all cong: positive_integral_cong)
have "\<And>N. N \<in> sets M \<Longrightarrow> {x \<in> space M. h x \<noteq> 0 \<and> indicator N x \<noteq> (0::pextreal)} = N"
using sets_into_space pos by (force simp: indicator_def)
then have "T.absolutely_continuous \<nu>" using assms(2) borel
@@ -783,7 +792,8 @@
by (fastsimp simp: borel_measurable_indicator positive_integral_0_iff)
from T.Radon_Nikodym_finite_measure_infinite[simplified, OF assms(1) this]
obtain f where f_borel: "f \<in> borel_measurable M" and
- fT: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = T.positive_integral (\<lambda>x. f x * indicator A x)" by auto
+ fT: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>?MT)"
+ by (auto simp: measurable_def)
show ?thesis
proof (safe intro!: bexI[of _ "\<lambda>x. h x * f x"])
show "(\<lambda>x. h x * f x) \<in> borel_measurable M"
@@ -792,7 +802,7 @@
then have "(\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
using f_borel by (auto intro: borel_measurable_pextreal_times borel_measurable_indicator)
from positive_integral_translated_density[OF borel this]
- show "\<nu> A = (\<integral>\<^isup>+x. h x * f x * indicator A x)"
+ show "\<nu> A = (\<integral>\<^isup>+x. h x * f x * indicator A x \<partial>M)"
unfolding fT[OF `A \<in> sets M`] by (simp add: field_simps)
qed
qed
@@ -801,8 +811,8 @@
lemma (in measure_space) finite_density_unique:
assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
- and fin: "positive_integral f < \<omega>"
- shows "(\<forall>A\<in>sets M. (\<integral>\<^isup>+x. f x * indicator A x) = (\<integral>\<^isup>+x. g x * indicator A x))
+ and fin: "integral\<^isup>P M f < \<omega>"
+ shows "(\<forall>A\<in>sets M. (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. g x * indicator A x \<partial>M))
\<longleftrightarrow> (AE x. f x = g x)"
(is "(\<forall>A\<in>sets M. ?P f A = ?P g A) \<longleftrightarrow> _")
proof (intro iffI ballI)
@@ -812,18 +822,18 @@
next
assume eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
from this[THEN bspec, OF top] fin
- have g_fin: "positive_integral g < \<omega>" by (simp cong: positive_integral_cong)
+ have g_fin: "integral\<^isup>P M g < \<omega>" by (simp cong: positive_integral_cong)
{ fix f g assume borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
- and g_fin: "positive_integral g < \<omega>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
+ and g_fin: "integral\<^isup>P M g < \<omega>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
let ?N = "{x\<in>space M. g x < f x}"
have N: "?N \<in> sets M" using borel by simp
- have "?P (\<lambda>x. (f x - g x)) ?N = (\<integral>\<^isup>+x. f x * indicator ?N x - g x * indicator ?N x)"
+ have "?P (\<lambda>x. (f x - g x)) ?N = (\<integral>\<^isup>+x. f x * indicator ?N x - g x * indicator ?N x \<partial>M)"
by (auto intro!: positive_integral_cong simp: indicator_def)
also have "\<dots> = ?P f ?N - ?P g ?N"
proof (rule positive_integral_diff)
show "(\<lambda>x. f x * indicator ?N x) \<in> borel_measurable M" "(\<lambda>x. g x * indicator ?N x) \<in> borel_measurable M"
using borel N by auto
- have "?P g ?N \<le> positive_integral g"
+ have "?P g ?N \<le> integral\<^isup>P M g"
by (auto intro!: positive_integral_mono simp: indicator_def)
then show "?P g ?N \<noteq> \<omega>" using g_fin by auto
fix x assume "x \<in> space M"
@@ -848,17 +858,17 @@
lemma (in finite_measure) density_unique_finite_measure:
assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"
- assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. f x * indicator A x) = (\<integral>\<^isup>+x. f' x * indicator A x)"
+ assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)"
(is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
shows "AE x. f x = f' x"
proof -
let "?\<nu> A" = "?P f A" and "?\<nu>' A" = "?P f' A"
let "?f A x" = "f x * indicator A x" and "?f' A x" = "f' x * indicator A x"
- interpret M: measure_space M ?\<nu>
- using borel(1) by (rule measure_space_density)
+ interpret M: measure_space "M\<lparr> measure := ?\<nu>\<rparr>"
+ using borel(1) by (rule measure_space_density) simp
have ac: "absolutely_continuous ?\<nu>"
using f by (rule density_is_absolutely_continuous)
- from split_space_into_finite_sets_and_rest[OF `measure_space M ?\<nu>` ac]
+ from split_space_into_finite_sets_and_rest[OF `measure_space (M\<lparr> measure := ?\<nu>\<rparr>)` ac]
obtain Q0 and Q :: "nat \<Rightarrow> 'a set"
where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)"
@@ -876,13 +886,13 @@
have 2: "AE x. ?f Q0 x = ?f' Q0 x"
proof (rule AE_I')
{ fix f :: "'a \<Rightarrow> pextreal" assume borel: "f \<in> borel_measurable M"
- and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x)"
+ and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
let "?A i" = "Q0 \<inter> {x \<in> space M. f x < of_nat i}"
have "(\<Union>i. ?A i) \<in> null_sets"
proof (rule null_sets_UN)
fix i have "?A i \<in> sets M"
using borel Q0(1) by auto
- have "?\<nu> (?A i) \<le> (\<integral>\<^isup>+x. of_nat i * indicator (?A i) x)"
+ have "?\<nu> (?A i) \<le> (\<integral>\<^isup>+x. of_nat i * indicator (?A i) x \<partial>M)"
unfolding eq[OF `?A i \<in> sets M`]
by (auto intro!: positive_integral_mono simp: indicator_def)
also have "\<dots> = of_nat i * \<mu> (?A i)"
@@ -912,63 +922,72 @@
lemma (in sigma_finite_measure) density_unique:
assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"
- assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. f x * indicator A x) = (\<integral>\<^isup>+x. f' x * indicator A x)"
+ assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)"
(is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
shows "AE x. f x = f' x"
proof -
obtain h where h_borel: "h \<in> borel_measurable M"
- and fin: "positive_integral h \<noteq> \<omega>" and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<omega>"
+ and fin: "integral\<^isup>P M h \<noteq> \<omega>" and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<omega>"
using Ex_finite_integrable_function by auto
- interpret h: measure_space M "\<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x)"
- using h_borel by (rule measure_space_density)
- interpret h: finite_measure M "\<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x)"
+ interpret h: measure_space "M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M) \<rparr>"
+ using h_borel by (rule measure_space_density) simp
+ interpret h: finite_measure "M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M) \<rparr>"
by default (simp cong: positive_integral_cong add: fin)
- interpret f: measure_space M "\<lambda>A. (\<integral>\<^isup>+x. f x * indicator A x)"
- using borel(1) by (rule measure_space_density)
- interpret f': measure_space M "\<lambda>A. (\<integral>\<^isup>+x. f' x * indicator A x)"
- using borel(2) by (rule measure_space_density)
+ let ?fM = "M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)\<rparr>"
+ interpret f: measure_space ?fM
+ using borel(1) by (rule measure_space_density) simp
+ let ?f'M = "M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)\<rparr>"
+ interpret f': measure_space ?f'M
+ using borel(2) by (rule measure_space_density) simp
{ fix A assume "A \<in> sets M"
then have " {x \<in> space M. h x \<noteq> 0 \<and> indicator A x \<noteq> (0::pextreal)} = A"
using pos sets_into_space by (force simp: indicator_def)
- then have "(\<integral>\<^isup>+x. h x * indicator A x) = 0 \<longleftrightarrow> A \<in> null_sets"
+ then have "(\<integral>\<^isup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets"
using h_borel `A \<in> sets M` by (simp add: positive_integral_0_iff) }
note h_null_sets = this
{ fix A assume "A \<in> sets M"
- have "(\<integral>\<^isup>+x. h x * (f x * indicator A x)) =
- f.positive_integral (\<lambda>x. h x * indicator A x)"
+ have "(\<integral>\<^isup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?fM)"
+ using `A \<in> sets M` h_borel borel
+ by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong)
+ also have "\<dots> = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?f'M)"
+ by (rule f'.positive_integral_cong_measure) (simp_all add: f)
+ also have "\<dots> = (\<integral>\<^isup>+x. h x * (f' x * indicator A x) \<partial>M)"
using `A \<in> sets M` h_borel borel
by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong)
- also have "\<dots> = f'.positive_integral (\<lambda>x. h x * indicator A x)"
- by (rule f'.positive_integral_cong_measure) (rule f)
- also have "\<dots> = (\<integral>\<^isup>+x. h x * (f' x * indicator A x))"
- using `A \<in> sets M` h_borel borel
- by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong)
- finally have "(\<integral>\<^isup>+x. h x * (f x * indicator A x)) = (\<integral>\<^isup>+x. h x * (f' x * indicator A x))" . }
+ finally have "(\<integral>\<^isup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^isup>+x. h x * (f' x * indicator A x) \<partial>M)" . }
then have "h.almost_everywhere (\<lambda>x. f x = f' x)"
using h_borel borel
- by (intro h.density_unique_finite_measure[OF borel])
- (simp add: positive_integral_translated_density)
+ apply (intro h.density_unique_finite_measure)
+ apply (simp add: measurable_def)
+ apply (simp add: measurable_def)
+ by (simp add: positive_integral_translated_density)
then show "AE x. f x = f' x"
unfolding h.almost_everywhere_def almost_everywhere_def
by (auto simp add: h_null_sets)
qed
lemma (in sigma_finite_measure) sigma_finite_iff_density_finite:
- assumes \<nu>: "measure_space M \<nu>" and f: "f \<in> borel_measurable M"
- and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x)"
- shows "sigma_finite_measure M \<nu> \<longleftrightarrow> (AE x. f x \<noteq> \<omega>)"
+ assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
+ and f: "f \<in> borel_measurable M"
+ and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
+ shows "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>) \<longleftrightarrow> (AE x. f x \<noteq> \<omega>)"
proof
- assume "sigma_finite_measure M \<nu>"
- then interpret \<nu>: sigma_finite_measure M \<nu> .
+ assume "sigma_finite_measure ?N"
+ then interpret \<nu>: sigma_finite_measure ?N
+ where "borel_measurable ?N = borel_measurable M" and "measure ?N = \<nu>"
+ and "sets ?N = sets M" and "space ?N = space M" by (auto simp: measurable_def)
from \<nu>.Ex_finite_integrable_function obtain h where
- h: "h \<in> borel_measurable M" "\<nu>.positive_integral h \<noteq> \<omega>"
+ h: "h \<in> borel_measurable M" "integral\<^isup>P ?N h \<noteq> \<omega>"
and fin: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<omega>" by auto
have "AE x. f x * h x \<noteq> \<omega>"
proof (rule AE_I')
- have "\<nu>.positive_integral h = (\<integral>\<^isup>+x. f x * h x)"
- by (simp add: \<nu>.positive_integral_cong_measure[symmetric, OF eq[symmetric]])
- (intro positive_integral_translated_density f h)
- then have "(\<integral>\<^isup>+x. f x * h x) \<noteq> \<omega>"
+ have "integral\<^isup>P ?N h = (\<integral>\<^isup>+x. f x * h x \<partial>M)"
+ apply (subst \<nu>.positive_integral_cong_measure[symmetric,
+ of "M\<lparr> measure := \<lambda> A. \<integral>\<^isup>+x. f x * indicator A x \<partial>M \<rparr>"])
+ apply (simp_all add: eq)
+ apply (rule positive_integral_translated_density)
+ using f h by auto
+ then have "(\<integral>\<^isup>+x. f x * h x \<partial>M) \<noteq> \<omega>"
using h(2) by simp
then show "(\<lambda>x. f x * h x) -` {\<omega>} \<inter> space M \<in> null_sets"
using f h(1) by (auto intro!: positive_integral_omega borel_measurable_vimage)
@@ -981,7 +1000,9 @@
next
assume AE: "AE x. f x \<noteq> \<omega>"
from sigma_finite guess Q .. note Q = this
- interpret \<nu>: measure_space M \<nu> by fact
+ interpret \<nu>: measure_space ?N
+ where "borel_measurable ?N = borel_measurable M" and "measure ?N = \<nu>"
+ and "sets ?N = sets M" and "space ?N = space M" using \<nu> by (auto simp: measurable_def)
def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<omega>} | Suc n \<Rightarrow> {.. of_nat (Suc n)}) \<inter> space M"
{ fix i j have "A i \<inter> Q j \<in> sets M"
unfolding A_def using f Q
@@ -989,11 +1010,11 @@
by (cases i) (auto intro: measurable_sets[OF f]) }
note A_in_sets = this
let "?A n" = "case prod_decode n of (i,j) \<Rightarrow> A i \<inter> Q j"
- show "sigma_finite_measure M \<nu>"
+ show "sigma_finite_measure ?N"
proof (default, intro exI conjI subsetI allI)
fix x assume "x \<in> range ?A"
then obtain n where n: "x = ?A n" by auto
- then show "x \<in> sets M" using A_in_sets by (cases "prod_decode n") auto
+ then show "x \<in> sets ?N" using A_in_sets by (cases "prod_decode n") auto
next
have "(\<Union>i. ?A i) = (\<Union>i j. A i \<inter> Q j)"
proof safe
@@ -1014,16 +1035,16 @@
then show ?thesis using x preal unfolding A_def by (auto intro!: exI[of _ "Suc n"])
qed
qed (auto simp: A_def)
- finally show "(\<Union>i. ?A i) = space M" by simp
+ finally show "(\<Union>i. ?A i) = space ?N" by simp
next
fix n obtain i j where
[simp]: "prod_decode n = (i, j)" by (cases "prod_decode n") auto
- have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x) \<noteq> \<omega>"
+ have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<noteq> \<omega>"
proof (cases i)
case 0
have "AE x. f x * indicator (A i \<inter> Q j) x = 0"
using AE by (rule AE_mp) (auto intro!: AE_cong simp: A_def `i = 0`)
- then have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x) = 0"
+ then have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) = 0"
using A_in_sets f
apply (subst positive_integral_0_iff)
apply fast
@@ -1034,8 +1055,8 @@
then show ?thesis by simp
next
case (Suc n)
- then have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x) \<le>
- (\<integral>\<^isup>+x. of_nat (Suc n) * indicator (Q j) x)"
+ then have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<le>
+ (\<integral>\<^isup>+x. of_nat (Suc n) * indicator (Q j) x \<partial>M)"
by (auto intro!: positive_integral_mono simp: indicator_def A_def)
also have "\<dots> = of_nat (Suc n) * \<mu> (Q j)"
using Q by (auto intro!: positive_integral_cmult_indicator)
@@ -1043,33 +1064,34 @@
using Q by auto
finally show ?thesis by simp
qed
- then show "\<nu> (?A n) \<noteq> \<omega>"
+ then show "measure ?N (?A n) \<noteq> \<omega>"
using A_in_sets Q eq by auto
qed
qed
section "Radon-Nikodym derivative"
-definition (in sigma_finite_measure)
- "RN_deriv \<nu> \<equiv> SOME f. f \<in> borel_measurable M \<and>
- (\<forall>A \<in> sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x))"
+definition
+ "RN_deriv M \<nu> \<equiv> SOME f. f \<in> borel_measurable M \<and>
+ (\<forall>A \<in> sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M))"
lemma (in sigma_finite_measure) RN_deriv_cong:
- assumes cong: "\<And>A. A \<in> sets M \<Longrightarrow> \<mu>' A = \<mu> A" "\<And>A. A \<in> sets M \<Longrightarrow> \<nu>' A = \<nu> A"
- shows "sigma_finite_measure.RN_deriv M \<mu>' \<nu>' x = RN_deriv \<nu> x"
+ assumes cong: "\<And>A. A \<in> sets M \<Longrightarrow> measure M' A = \<mu> A" "sets M' = sets M" "space M' = space M"
+ and \<nu>: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu>' A = \<nu> A"
+ shows "RN_deriv M' \<nu>' x = RN_deriv M \<nu> x"
proof -
- interpret \<mu>': sigma_finite_measure M \<mu>'
- using cong(1) by (rule sigma_finite_measure_cong)
+ interpret \<mu>': sigma_finite_measure M'
+ using cong by (rule sigma_finite_measure_cong)
show ?thesis
- unfolding RN_deriv_def \<mu>'.RN_deriv_def
- by (simp add: cong positive_integral_cong_measure[OF cong(1)])
+ unfolding RN_deriv_def
+ by (simp add: cong \<nu> positive_integral_cong_measure[OF cong] measurable_def)
qed
lemma (in sigma_finite_measure) RN_deriv:
- assumes "measure_space M \<nu>"
+ assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)"
assumes "absolutely_continuous \<nu>"
- shows "RN_deriv \<nu> \<in> borel_measurable M" (is ?borel)
- and "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. RN_deriv \<nu> x * indicator A x)"
+ shows "RN_deriv M \<nu> \<in> borel_measurable M" (is ?borel)
+ and "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)"
(is "\<And>A. _ \<Longrightarrow> ?int A")
proof -
note Ex = Radon_Nikodym[OF assms, unfolded Bex_def]
@@ -1080,87 +1102,92 @@
qed
lemma (in sigma_finite_measure) RN_deriv_positive_integral:
- assumes \<nu>: "measure_space M \<nu>" "absolutely_continuous \<nu>"
+ assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>"
and f: "f \<in> borel_measurable M"
- shows "measure_space.positive_integral M \<nu> f = (\<integral>\<^isup>+x. RN_deriv \<nu> x * f x)"
+ shows "integral\<^isup>P (M\<lparr>measure := \<nu>\<rparr>) f = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * f x \<partial>M)"
proof -
- interpret \<nu>: measure_space M \<nu> by fact
- have "\<nu>.positive_integral f =
- measure_space.positive_integral M (\<lambda>A. (\<integral>\<^isup>+x. RN_deriv \<nu> x * indicator A x)) f"
- by (intro \<nu>.positive_integral_cong_measure[symmetric] RN_deriv(2)[OF \<nu>, symmetric])
- also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv \<nu> x * f x)"
- by (intro positive_integral_translated_density RN_deriv[OF \<nu>] f)
+ interpret \<nu>: measure_space "M\<lparr>measure := \<nu>\<rparr>" by fact
+ have "integral\<^isup>P (M\<lparr>measure := \<nu>\<rparr>) f =
+ integral\<^isup>P (M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)\<rparr>) f"
+ by (intro \<nu>.positive_integral_cong_measure[symmetric])
+ (simp_all add: RN_deriv(2)[OF \<nu>, symmetric])
+ also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * f x \<partial>M)"
+ by (intro positive_integral_translated_density)
+ (simp_all add: RN_deriv[OF \<nu>] f)
finally show ?thesis .
qed
lemma (in sigma_finite_measure) RN_deriv_unique:
- assumes \<nu>: "measure_space M \<nu>" "absolutely_continuous \<nu>"
+ assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>"
and f: "f \<in> borel_measurable M"
- and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x)"
- shows "AE x. f x = RN_deriv \<nu> x"
+ and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
+ shows "AE x. f x = RN_deriv M \<nu> x"
proof (rule density_unique[OF f RN_deriv(1)[OF \<nu>]])
fix A assume A: "A \<in> sets M"
- show "(\<integral>\<^isup>+x. f x * indicator A x) = (\<integral>\<^isup>+x. RN_deriv \<nu> x * indicator A x)"
+ show "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)"
unfolding eq[OF A, symmetric] RN_deriv(2)[OF \<nu> A, symmetric] ..
qed
-
lemma (in sigma_finite_measure) RN_deriv_finite:
- assumes sfm: "sigma_finite_measure M \<nu>" and ac: "absolutely_continuous \<nu>"
- shows "AE x. RN_deriv \<nu> x \<noteq> \<omega>"
+ assumes sfm: "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>)" and ac: "absolutely_continuous \<nu>"
+ shows "AE x. RN_deriv M \<nu> x \<noteq> \<omega>"
proof -
- interpret \<nu>: sigma_finite_measure M \<nu> by fact
- have \<nu>: "measure_space M \<nu>" by default
+ interpret \<nu>: sigma_finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact
+ have \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
from sfm show ?thesis
using sigma_finite_iff_density_finite[OF \<nu> RN_deriv[OF \<nu> ac]] by simp
qed
lemma (in sigma_finite_measure)
- assumes \<nu>: "sigma_finite_measure M \<nu>" "absolutely_continuous \<nu>"
+ assumes \<nu>: "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>"
and f: "f \<in> borel_measurable M"
- shows RN_deriv_integral: "measure_space.integral M \<nu> f = (\<integral>x. real (RN_deriv \<nu> x) * f x)" (is ?integral)
- and RN_deriv_integrable: "measure_space.integrable M \<nu> f \<longleftrightarrow> integrable (\<lambda>x. real (RN_deriv \<nu> x) * f x)" (is ?integrable)
+ shows RN_deriv_integrable: "integrable (M\<lparr>measure := \<nu>\<rparr>) f \<longleftrightarrow>
+ integrable M (\<lambda>x. real (RN_deriv M \<nu> x) * f x)" (is ?integrable)
+ and RN_deriv_integral: "integral\<^isup>L (M\<lparr>measure := \<nu>\<rparr>) f =
+ (\<integral>x. real (RN_deriv M \<nu> x) * f x \<partial>M)" (is ?integral)
proof -
- interpret \<nu>: sigma_finite_measure M \<nu> by fact
- have ms: "measure_space M \<nu>" by default
+ interpret \<nu>: sigma_finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact
+ have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
have minus_cong: "\<And>A B A' B'::pextreal. A = A' \<Longrightarrow> B = B' \<Longrightarrow> real A - real B = real A' - real B'" by simp
have f': "(\<lambda>x. - f x) \<in> borel_measurable M" using f by auto
- { fix f :: "'a \<Rightarrow> real" assume "f \<in> borel_measurable M"
- { fix x assume *: "RN_deriv \<nu> x \<noteq> \<omega>"
- have "Real (real (RN_deriv \<nu> x)) * Real (f x) = Real (real (RN_deriv \<nu> x) * f x)"
+ have Nf: "f \<in> borel_measurable (M\<lparr>measure := \<nu>\<rparr>)" using f unfolding measurable_def by auto
+ { fix f :: "'a \<Rightarrow> real"
+ { fix x assume *: "RN_deriv M \<nu> x \<noteq> \<omega>"
+ have "Real (real (RN_deriv M \<nu> x)) * Real (f x) = Real (real (RN_deriv M \<nu> x) * f x)"
by (simp add: mult_le_0_iff)
- then have "RN_deriv \<nu> x * Real (f x) = Real (real (RN_deriv \<nu> x) * f x)"
+ then have "RN_deriv M \<nu> x * Real (f x) = Real (real (RN_deriv M \<nu> x) * f x)"
using * by (simp add: Real_real) }
note * = this
- have "(\<integral>\<^isup>+x. RN_deriv \<nu> x * Real (f x)) = (\<integral>\<^isup>+x. Real (real (RN_deriv \<nu> x) * f x))"
+ have "(\<integral>\<^isup>+x. RN_deriv M \<nu> x * Real (f x) \<partial>M) = (\<integral>\<^isup>+x. Real (real (RN_deriv M \<nu> x) * f x) \<partial>M)"
apply (rule positive_integral_cong_AE)
apply (rule AE_mp[OF RN_deriv_finite[OF \<nu>]])
by (auto intro!: AE_cong simp: *) }
- with this[OF f] this[OF f'] f f'
+ with this this f f' Nf
show ?integral ?integrable
- unfolding \<nu>.integral_def integral_def \<nu>.integrable_def integrable_def
- by (auto intro!: RN_deriv(1)[OF ms \<nu>(2)] minus_cong simp: RN_deriv_positive_integral[OF ms \<nu>(2)])
+ unfolding lebesgue_integral_def integrable_def
+ by (auto intro!: RN_deriv(1)[OF ms \<nu>(2)] minus_cong
+ simp: RN_deriv_positive_integral[OF ms \<nu>(2)])
qed
lemma (in sigma_finite_measure) RN_deriv_singleton:
- assumes "measure_space M \<nu>"
+ assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)"
and ac: "absolutely_continuous \<nu>"
and "{x} \<in> sets M"
- shows "\<nu> {x} = RN_deriv \<nu> x * \<mu> {x}"
+ shows "\<nu> {x} = RN_deriv M \<nu> x * \<mu> {x}"
proof -
note deriv = RN_deriv[OF assms(1, 2)]
from deriv(2)[OF `{x} \<in> sets M`]
- have "\<nu> {x} = (\<integral>\<^isup>+w. RN_deriv \<nu> x * indicator {x} w)"
+ have "\<nu> {x} = (\<integral>\<^isup>+w. RN_deriv M \<nu> x * indicator {x} w \<partial>M)"
by (auto simp: indicator_def intro!: positive_integral_cong)
thus ?thesis using positive_integral_cmult_indicator[OF `{x} \<in> sets M`]
by auto
qed
theorem (in finite_measure_space) RN_deriv_finite_measure:
- assumes "measure_space M \<nu>"
+ assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)"
and ac: "absolutely_continuous \<nu>"
and "x \<in> space M"
- shows "\<nu> {x} = RN_deriv \<nu> x * \<mu> {x}"
+ shows "\<nu> {x} = RN_deriv M \<nu> x * \<mu> {x}"
proof -
have "{x} \<in> sets M" using sets_eq_Pow `x \<in> space M` by auto
from RN_deriv_singleton[OF assms(1,2) this] show ?thesis .
--- a/src/HOL/Probability/Sigma_Algebra.thy Wed Feb 02 10:35:41 2011 +0100
+++ b/src/HOL/Probability/Sigma_Algebra.thy Wed Feb 02 12:34:45 2011 +0100
@@ -29,7 +29,7 @@
sets :: "'a set set"
locale algebra =
- fixes M :: "'a algebra"
+ fixes M :: "('a, 'b) algebra_scheme"
assumes space_closed: "sets M \<subseteq> Pow (space M)"
and empty_sets [iff]: "{} \<in> sets M"
and compl_sets [intro]: "!!a. a \<in> sets M \<Longrightarrow> space M - a \<in> sets M"
@@ -104,7 +104,7 @@
section {* Restricted algebras *}
abbreviation (in algebra)
- "restricted_space A \<equiv> \<lparr> space = A, sets = (\<lambda>S. (A \<inter> S)) ` sets M \<rparr>"
+ "restricted_space A \<equiv> \<lparr> space = A, sets = (\<lambda>S. (A \<inter> S)) ` sets M, \<dots> = more M \<rparr>"
lemma (in algebra) restricted_algebra:
assumes "A \<in> sets M" shows "algebra (restricted_space A)"
@@ -222,7 +222,7 @@
| Union: "(\<And>i::nat. a i \<in> sigma_sets sp A) \<Longrightarrow> (\<Union>i. a i) \<in> sigma_sets sp A"
definition
- "sigma M = \<lparr> space = space M, sets = sigma_sets (space M) (sets M) \<rparr>"
+ "sigma M = \<lparr> space = space M, sets = sigma_sets (space M) (sets M), \<dots> = more M \<rparr>"
lemma (in sigma_algebra) sigma_sets_subset:
assumes a: "a \<subseteq> sets M"
@@ -354,12 +354,12 @@
qed
lemma sigma_sets_Int:
- assumes "A \<in> sigma_sets sp st"
- shows "op \<inter> A ` sigma_sets sp st = sigma_sets (A \<inter> sp) (op \<inter> A ` st)"
+ assumes "A \<in> sigma_sets sp st" "A \<subseteq> sp"
+ shows "op \<inter> A ` sigma_sets sp st = sigma_sets A (op \<inter> A ` st)"
proof (intro equalityI subsetI)
fix x assume "x \<in> op \<inter> A ` sigma_sets sp st"
then obtain y where "y \<in> sigma_sets sp st" "x = y \<inter> A" by auto
- then show "x \<in> sigma_sets (A \<inter> sp) (op \<inter> A ` st)"
+ then have "x \<in> sigma_sets (A \<inter> sp) (op \<inter> A ` st)"
proof (induct arbitrary: x)
case (Compl a)
then show ?case
@@ -370,13 +370,15 @@
by (auto intro!: sigma_sets.Union
simp add: UN_extend_simps simp del: UN_simps)
qed (auto intro!: sigma_sets.intros)
+ then show "x \<in> sigma_sets A (op \<inter> A ` st)"
+ using `A \<subseteq> sp` by (simp add: Int_absorb2)
next
- fix x assume "x \<in> sigma_sets (A \<inter> sp) (op \<inter> A ` st)"
+ fix x assume "x \<in> sigma_sets A (op \<inter> A ` st)"
then show "x \<in> op \<inter> A ` sigma_sets sp st"
proof induct
case (Compl a)
then obtain x where "a = A \<inter> x" "x \<in> sigma_sets sp st" by auto
- then show ?case
+ then show ?case using `A \<subseteq> sp`
by (force simp add: image_iff intro!: bexI[of _ "sp - x"] sigma_sets.Compl)
next
case (Union a)
@@ -707,7 +709,7 @@
subsection {* Sigma algebra generated by function preimages *}
definition (in sigma_algebra)
- "vimage_algebra S f = \<lparr> space = S, sets = (\<lambda>A. f -` A \<inter> S) ` sets M \<rparr>"
+ "vimage_algebra S f = \<lparr> space = S, sets = (\<lambda>A. f -` A \<inter> S) ` sets M, \<dots> = more M \<rparr>"
lemma (in sigma_algebra) in_vimage_algebra[simp]:
"A \<in> sets (vimage_algebra S f) \<longleftrightarrow> (\<exists>B\<in>sets M. A = f -` B \<inter> S)"
@@ -870,7 +872,7 @@
section {* Conditional space *}
definition (in algebra)
- "image_space X = \<lparr> space = X`space M, sets = (\<lambda>S. X`S) ` sets M \<rparr>"
+ "image_space X = \<lparr> space = X`space M, sets = (\<lambda>S. X`S) ` sets M, \<dots> = more M \<rparr>"
definition (in algebra)
"conditional_space X A = algebra.image_space (restricted_space A) X"
@@ -1158,7 +1160,7 @@
section {* Dynkin systems *}
locale dynkin_system =
- fixes M :: "'a algebra"
+ fixes M :: "('a, 'b) algebra_scheme"
assumes space_closed: "sets M \<subseteq> Pow (space M)"
and space: "space M \<in> sets M"
and compl[intro!]: "\<And>A. A \<in> sets M \<Longrightarrow> space M - A \<in> sets M"
@@ -1239,12 +1241,12 @@
subsection "Smallest Dynkin systems"
-definition dynkin :: "'a algebra \<Rightarrow> 'a algebra" where
+definition dynkin where
"dynkin M = \<lparr> space = space M,
- sets = \<Inter>{D. dynkin_system \<lparr> space = space M, sets = D\<rparr> \<and> sets M \<subseteq> D}\<rparr>"
+ sets = \<Inter>{D. dynkin_system \<lparr> space = space M, sets = D \<rparr> \<and> sets M \<subseteq> D},
+ \<dots> = more M \<rparr>"
lemma dynkin_system_dynkin:
- fixes M :: "'a algebra"
assumes "sets M \<subseteq> Pow (space M)"
shows "dynkin_system (dynkin M)"
proof (rule dynkin_systemI)
@@ -1311,21 +1313,17 @@
qed
lemma (in dynkin_system) dynkin_subset:
- fixes N :: "'a algebra"
assumes "sets N \<subseteq> sets M"
assumes "space N = space M"
shows "sets (dynkin N) \<subseteq> sets M"
proof -
- have *: "\<lparr>space = space N, sets = sets M\<rparr> = M"
- unfolding `space N = space M` by simp
have "dynkin_system M" by default
- then have "dynkin_system \<lparr>space = space N, sets = sets M\<rparr>"
- using assms unfolding * by simp
+ then have "dynkin_system \<lparr>space = space N, sets = sets M \<rparr>"
+ using assms unfolding dynkin_system_def by simp
with `sets N \<subseteq> sets M` show ?thesis by (auto simp add: dynkin_def)
qed
lemma sigma_eq_dynkin:
- fixes M :: "'a algebra"
assumes sets: "sets M \<subseteq> Pow (space M)"
assumes "Int_stable M"
shows "sigma M = dynkin M"
@@ -1367,7 +1365,7 @@
have "sigma_sets (space M) (sets M) \<subseteq> sets (dynkin M)" by auto
ultimately have "sigma_sets (space M) (sets M) = sets (dynkin M)" by auto
then show ?thesis
- by (intro algebra.equality) (simp_all add: sigma_def)
+ by (auto intro!: algebra.equality simp: sigma_def dynkin_def)
qed
lemma (in dynkin_system) dynkin_idem:
@@ -1381,23 +1379,22 @@
by (intro dynkin_subset) auto
qed
then show ?thesis
- by (auto intro!: algebra.equality)
+ by (auto intro!: algebra.equality simp: dynkin_def)
qed
lemma (in dynkin_system) dynkin_lemma:
- fixes E :: "'a algebra"
- assumes "Int_stable E" and E: "sets E \<subseteq> sets M" "space E = space M"
- and "sets M \<subseteq> sets (sigma E)"
- shows "sigma E = M"
+ assumes "Int_stable E"
+ and E: "sets E \<subseteq> sets M" "space E = space M" "sets M \<subseteq> sets (sigma E)"
+ shows "sets (sigma E) = sets M"
proof -
have "sets E \<subseteq> Pow (space E)"
- using E sets_into_space by auto
+ using E sets_into_space by force
then have "sigma E = dynkin E"
using `Int_stable E` by (rule sigma_eq_dynkin)
moreover then have "sets (dynkin E) = sets M"
- using assms dynkin_subset[OF E] by simp
+ using assms dynkin_subset[OF E(1,2)] by simp
ultimately show ?thesis
- using E by simp
+ using assms by (auto intro!: algebra.equality simp: dynkin_def)
qed
subsection "Sigma algebras on finite sets"
@@ -1467,6 +1464,7 @@
lemma vimage_algebra_sigma:
assumes bi: "bij_inv (space (sigma F)) (space (sigma E)) f e"
and "sets E \<subseteq> Pow (space E)" and F: "sets F \<subseteq> Pow (space F)"
+ and "more E = more F"
and "f \<in> measurable F E" "e \<in> measurable E F"
shows "sigma_algebra.vimage_algebra (sigma E) (space (sigma F)) f = sigma F"
proof -
@@ -1486,8 +1484,8 @@
using `f \<in> measurable F E` unfolding measurable_def by auto
qed
show "vimage_algebra (space (sigma F)) f = sigma F"
- unfolding vimage_algebra_def
- using assms by (auto simp: bij_inv_def eq sigma_sets_vimage[symmetric] sigma_def)
+ unfolding vimage_algebra_def using assms
+ by (auto simp: bij_inv_def eq sigma_sets_vimage[symmetric] sigma_def)
qed
lemma measurable_sigma_sigma:
--- a/src/HOL/Probability/ex/Dining_Cryptographers.thy Wed Feb 02 10:35:41 2011 +0100
+++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy Wed Feb 02 12:34:45 2011 +0100
@@ -7,25 +7,25 @@
assumes finite[simp]: "finite S"
and not_empty[simp]: "S \<noteq> {}"
-definition (in finite_space) "M = \<lparr> space = S, sets = Pow S \<rparr>"
-definition (in finite_space) \<mu>_def[simp]: "\<mu> A = (of_nat (card A) / of_nat (card S) :: pextreal)"
+definition (in finite_space) "M = \<lparr> space = S, sets = Pow S,
+ measure = (\<lambda>A. of_nat (card A) / of_nat (card S) :: pextreal) \<rparr>"
lemma (in finite_space)
shows space_M[simp]: "space M = S"
and sets_M[simp]: "sets M = Pow S"
by (simp_all add: M_def)
-sublocale finite_space \<subseteq> finite_measure_space M \<mu>
+sublocale finite_space \<subseteq> finite_measure_space M
proof (rule finite_measure_spaceI)
fix A B :: "'a set" assume "A \<inter> B = {}" "A \<subseteq> space M" "B \<subseteq> space M"
- then show "\<mu> (A \<union> B) = \<mu> A + \<mu> B"
- by (simp add: inverse_eq_divide field_simps Real_real
+ then show "measure M (A \<union> B) = measure M A + measure M B"
+ by (simp add: inverse_eq_divide field_simps Real_real M_def
divide_le_0_iff zero_le_divide_iff
card_Un_disjoint finite_subset[OF _ finite])
-qed auto
+qed (auto simp: M_def)
-sublocale finite_space \<subseteq> information_space M \<mu> 2
- by default simp_all
+sublocale finite_space \<subseteq> information_space M 2
+ by default (simp_all add: M_def)
lemma set_of_list_extend:
"{xs. length xs = Suc n \<and> (\<forall>x\<in>set xs. x \<in> A)} =
@@ -506,7 +506,8 @@
moreover from x have "length x = n" by (auto simp: inversion_def_raw dc_crypto)
ultimately
have "real (?dIP {(x, z)}) = 2 / (real n * 2^n)" using x
- by (simp add: distribution_def card_dc_crypto card_payer_and_inversion
+ unfolding distribution_def
+ by (simp add: M_def card_dc_crypto card_payer_and_inversion
inverse_eq_divide mult_le_0_iff zero_le_mult_iff power_le_zero_eq)
moreover
from z have "payer -` {z} \<inter> dc_crypto = {z} \<times> {xs. length xs = n}"
@@ -514,8 +515,8 @@
hence "card (payer -` {z} \<inter> dc_crypto) = 2^n"
using card_list_length[where A="UNIV::bool set"]
by (simp add: card_cartesian_product_singleton)
- hence "real (?dP {z}) = 1 / real n"
- by (simp add: distribution_def card_dc_crypto field_simps inverse_eq_divide
+ hence "real (?dP {z}) = 1 / real n" unfolding distribution_def
+ by (simp add: card_dc_crypto field_simps inverse_eq_divide M_def
mult_le_0_iff zero_le_mult_iff power_le_zero_eq)
ultimately
show "real (?dIP {(x,z)}) * log 2 (real (?dIP {(x,z)}) / real (?dP {z})) =
@@ -535,8 +536,8 @@
fix x assume x_inv: "x \<in> inversion ` dc_crypto"
hence "length x = n" by (auto simp: inversion_def_raw dc_crypto)
moreover have "inversion -` {x} \<inter> dc_crypto = {dc \<in> dc_crypto. inversion dc = x}" by auto
- ultimately have "?dI {x} = 2 / 2^n" using `0 < n`
- by (simp add: distribution_def card_inversion[OF x_inv] card_dc_crypto
+ ultimately have "?dI {x} = 2 / 2^n" using `0 < n` unfolding distribution_def
+ by (simp add: card_inversion[OF x_inv] card_dc_crypto M_def
mult_le_0_iff zero_le_mult_iff power_le_zero_eq)
thus "real (?dI {x}) * log 2 (real (?dI {x})) = 2 * (1 - real n) / 2^n"
by (simp add: log_divide log_nat_power power_le_zero_eq inverse_eq_divide)
--- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Wed Feb 02 10:35:41 2011 +0100
+++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Wed Feb 02 12:34:45 2011 +0100
@@ -1,6 +1,6 @@
-(* Author: Johannes Hoelzl, TU Muenchen *)
+(* Author: Johannes Hölzl, TU München *)
-header {* Formalization of a Countermeasure by Koepf & Duermuth 2009 *}
+header {* Formalization of a Countermeasure by Koepf \& Duermuth 2009 *}
theory Koepf_Duermuth_Countermeasure
imports Information "~~/src/HOL/Library/Permutation"
@@ -80,9 +80,7 @@
shows "bij_betw (ordered_variable_partition X) {0..<card (X`space M)} (X`space M)"
and "\<And>i j. \<lbrakk> i < card (X`space M) ; j < card (X`space M) ; i \<le> j \<rbrakk> \<Longrightarrow>
distribution X {ordered_variable_partition X i} \<le> distribution X {ordered_variable_partition X j}"
-proof -
-
-qed
+ oops
definition (in prob_space)
"order_distribution X i = real (distribution X {ordered_variable_partition X i})"
@@ -90,12 +88,10 @@
definition (in prob_space)
"guessing_entropy b X = (\<Sum>i<card(X`space M). real i * log b (order_distribution X i))"
-abbreviation (in finite_information_space)
+abbreviation (in information_space)
finite_guessing_entropy ("\<G>'(_')") where
"\<G>(X) \<equiv> guessing_entropy b X"
-
-
lemma zero_notin_Suc_image[simp]: "0 \<notin> Suc ` A"
by auto
@@ -103,7 +99,7 @@
"extensional d A = {f. \<forall>x. x \<notin> A \<longrightarrow> f x = d}"
lemma extensional_empty[simp]: "extensional d {} = {\<lambda>x. d}"
- unfolding extensional_def by (simp add: expand_set_eq expand_fun_eq)
+ unfolding extensional_def by (simp add: set_eq_iff fun_eq_iff)
lemma funset_eq_UN_fun_upd_I:
assumes *: "\<And>f. f \<in> F (insert a A) \<Longrightarrow> f(a := d) \<in> F A"
@@ -138,7 +134,7 @@
using assms by induct auto
lemma fun_upd_eq_iff: "f(a := b) = g(a := b') \<longleftrightarrow> b = b' \<and> f(a := d) = g(a := d)"
- by (auto simp: expand_fun_eq)
+ by (auto simp: fun_eq_iff)
lemma card_funcset:
assumes "finite A" "finite B"
@@ -205,24 +201,14 @@
lemma (in finite_information) positive_p_sum[simp]: "0 \<le> setsum p X"
by (auto intro!: setsum_nonneg)
-sublocale finite_information \<subseteq> finite_information_space "\<lparr> space = \<Omega>, sets = Pow \<Omega> \<rparr>" "\<lambda>S. Real (setsum p S)" b
-proof -
- show "finite_information_space \<lparr> space = \<Omega>, sets = Pow \<Omega> \<rparr> (\<lambda>S. Real (setsum p S)) b"
- unfolding finite_information_space_def finite_information_space_axioms_def
- unfolding finite_prob_space_def prob_space_def prob_space_axioms_def
- unfolding finite_measure_space_def finite_measure_space_axioms_def
- by (force intro!: sigma_algebra.finite_additivity_sufficient
- simp: additive_def sigma_algebra_Pow positive_def Real_eq_Real
- setsum.union_disjoint finite_subset)
-qed
+sublocale finite_information \<subseteq> finite_measure_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. Real (setsum p S)\<rparr>"
+ by (rule finite_measure_spaceI) (simp_all add: setsum_Un_disjoint finite_subset)
-lemma (in prob_space) prob_space_subalgebra:
- assumes "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)"
- shows "prob_space (M\<lparr> sets := N \<rparr>) \<mu>" sorry
+sublocale finite_information \<subseteq> finite_prob_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. Real (setsum p S)\<rparr>"
+ by default simp
-lemma (in measure_space) measure_space_subalgebra:
- assumes "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)"
- shows "measure_space (M\<lparr> sets := N \<rparr>) \<mu>" sorry
+sublocale finite_information \<subseteq> information_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. Real (setsum p S)\<rparr>" b
+ by default simp
locale koepf_duermuth = K: finite_information keys K b + M: finite_information messages M b
for b :: real
@@ -258,12 +244,11 @@
next
case (Suc n)
then show ?case
- by (simp add: comp_def set_of_list_extend
- lessThan_eq_Suc_image setsum_reindex setprod_reindex)
+ by (simp add: comp_def set_of_list_extend lessThan_Suc_eq_insert_0
+ setsum_reindex setprod_reindex)
qed
then show "setsum P msgs = 1"
unfolding msgs_def P_def by simp
-
fix x
have "\<And> A f. 0 \<le> (\<Prod>x\<in>A. M (f x))" by (auto simp: setprod_nonneg)
then show "0 \<le> P x"
@@ -303,7 +288,7 @@
show ?case unfolding *
using Suc[of "\<lambda>i. P (Suc i)"]
by (simp add: setsum_reindex split_conv setsum_cartesian_product'
- lessThan_eq_Suc_image setprod_reindex setsum_left_distrib[symmetric] setsum_right_distrib[symmetric])
+ lessThan_Suc_eq_insert_0 setprod_reindex setsum_left_distrib[symmetric] setsum_right_distrib[symmetric])
qed
context koepf_duermuth
@@ -347,22 +332,22 @@
"\<P>(X|Y) x \<equiv> \<P>(X, Y) x / \<P>(Y) (snd`x)"
notation
- finite_entropy ("\<H>'( _ ')")
+ entropy_Pow ("\<H>'( _ ')")
notation
- finite_conditional_entropy ("\<H>'( _ | _ ')")
+ conditional_entropy_Pow ("\<H>'( _ | _ ')")
notation
- finite_mutual_information ("\<I>'( _ ; _ ')")
+ mutual_information_Pow ("\<I>'( _ ; _ ')")
lemma t_eq_imp_bij_func:
assumes "t xs = t ys"
shows "\<exists>f. bij_betw f {..<length xs} {..<length ys} \<and> (\<forall>i<length xs. xs ! i = ys ! (f i))"
proof -
have "count (multiset_of xs) = count (multiset_of ys)"
- using assms by (simp add: expand_fun_eq count_multiset_of t_def)
+ using assms by (simp add: fun_eq_iff count_multiset_of t_def)
then have "xs <~~> ys" unfolding multiset_of_eq_perm count_inject .
- then show ?thesis by (rule permutation_Ex_func)
+ then show ?thesis by (rule permutation_Ex_bij)
qed
lemma \<P>_k: assumes "k \<in> keys" shows "\<P>(fst) {k} = K k"
@@ -471,8 +456,6 @@
let "?H obs" = "(\<Sum>k\<in>keys. \<P>(fst|OB) {(k, obs)} * log b (\<P>(fst|OB) {(k, obs)})) :: real"
let "?Ht obs" = "(\<Sum>k\<in>keys. \<P>(fst|t\<circ>OB) {(k, obs)} * log b (\<P>(fst|t\<circ>OB) {(k, obs)})) :: real"
- note [[simproc del: finite_information_space.mult_log]]
-
{ fix obs assume obs: "obs \<in> OB`msgs"
have "?H obs = (\<Sum>k\<in>keys. \<P>(fst|t\<circ>OB) {(k, t obs)} * log b (\<P>(fst|t\<circ>OB) {(k, t obs)}))"
using CP_T_eq_CP_O[OF _ obs]
@@ -495,7 +478,8 @@
txt {* Lemma 3 *}
have "\<H>(fst | OB) = -(\<Sum>obs\<in>OB`msgs. \<P>(OB) {obs} * ?Ht (t obs))"
- unfolding conditional_entropy_eq_ce_with_hypothesis using * by simp
+ unfolding conditional_entropy_eq_ce_with_hypothesis[OF
+ simple_function_finite simple_function_finite] using * by simp
also have "\<dots> = -(\<Sum>obs\<in>t`OB`msgs. \<P>(t\<circ>OB) {obs} * ?Ht obs)"
apply (subst SIGMA_image_vimage[symmetric, of "OB`msgs" t])
apply (subst setsum_reindex)
@@ -509,20 +493,22 @@
by (simp add: setsum_divide_distrib[symmetric] field_simps **
setsum_right_distrib[symmetric] setsum_left_distrib[symmetric])
also have "\<dots> = \<H>(fst | t\<circ>OB)"
- unfolding conditional_entropy_eq_ce_with_hypothesis
+ unfolding conditional_entropy_eq_ce_with_hypothesis[OF
+ simple_function_finite simple_function_finite]
by (simp add: comp_def image_image[symmetric])
finally show ?thesis .
qed
theorem "\<I>(fst ; OB) \<le> real (card observations) * log b (real n + 1)"
proof -
+ from simple_function_finite simple_function_finite
have "\<I>(fst ; OB) = \<H>(fst) - \<H>(fst | OB)"
- using mutual_information_eq_entropy_conditional_entropy .
+ by (rule mutual_information_eq_entropy_conditional_entropy)
also have "\<dots> = \<H>(fst) - \<H>(fst | t\<circ>OB)"
unfolding ce_OB_eq_ce_t ..
also have "\<dots> = \<H>(t\<circ>OB) - \<H>(t\<circ>OB | fst)"
- unfolding entropy_chain_rule[symmetric] sign_simps
- by (subst entropy_commute) simp
+ unfolding entropy_chain_rule[symmetric, OF simple_function_finite simple_function_finite] sign_simps
+ by (subst entropy_commute[OF simple_function_finite simple_function_finite]) simp
also have "\<dots> \<le> \<H>(t\<circ>OB)"
using conditional_entropy_positive[of "t\<circ>OB" fst] by simp
also have "\<dots> \<le> log b (real (card ((t\<circ>OB)`msgs)))"
@@ -537,4 +523,4 @@
end
-end
\ No newline at end of file
+end