--- a/src/HOL/Probability/Borel.thy Fri Aug 27 16:23:51 2010 +0200
+++ b/src/HOL/Probability/Borel.thy Thu Sep 02 17:12:40 2010 +0200
@@ -6,6 +6,10 @@
imports Sigma_Algebra Positive_Infinite_Real Multivariate_Analysis
begin
+lemma LIMSEQ_max:
+ "u ----> (x::real) \<Longrightarrow> (\<lambda>i. max (u i) 0) ----> max x 0"
+ by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D)
+
section "Generic Borel spaces"
definition "borel_space = sigma (UNIV::'a::topological_space set) open"
@@ -105,6 +109,53 @@
qed (auto simp add: vimage_UN)
qed
+lemma (in sigma_algebra) borel_measurable_restricted:
+ fixes f :: "'a \<Rightarrow> 'x\<Colon>{topological_space, semiring_1}" assumes "A \<in> sets M"
+ shows "f \<in> borel_measurable (restricted_space A) \<longleftrightarrow>
+ (\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
+ (is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M")
+proof -
+ interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
+ have *: "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable ?R"
+ by (auto intro!: measurable_cong)
+ show ?thesis unfolding *
+ unfolding in_borel_measurable_borel_space
+ proof (simp, safe)
+ fix S :: "'x set" assume "S \<in> sets borel_space"
+ "\<forall>S\<in>sets borel_space. ?f -` S \<inter> A \<in> op \<inter> A ` sets M"
+ then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto
+ then have f: "?f -` S \<inter> A \<in> sets M"
+ using `A \<in> sets M` sets_into_space by fastsimp
+ show "?f -` S \<inter> space M \<in> sets M"
+ proof cases
+ assume "0 \<in> S"
+ then have "?f -` S \<inter> space M = ?f -` S \<inter> A \<union> (space M - A)"
+ using `A \<in> sets M` sets_into_space by auto
+ then show ?thesis using f `A \<in> sets M` by (auto intro!: Un Diff)
+ next
+ assume "0 \<notin> S"
+ then have "?f -` S \<inter> space M = ?f -` S \<inter> A"
+ using `A \<in> sets M` sets_into_space
+ by (auto simp: indicator_def split: split_if_asm)
+ then show ?thesis using f by auto
+ qed
+ next
+ fix S :: "'x set" assume "S \<in> sets borel_space"
+ "\<forall>S\<in>sets borel_space. ?f -` S \<inter> space M \<in> sets M"
+ then have f: "?f -` S \<inter> space M \<in> sets M" by auto
+ then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M"
+ using `A \<in> sets M` sets_into_space
+ apply (simp add: image_iff)
+ apply (rule bexI[OF _ f])
+ by auto
+ qed
+qed
+
+lemma (in sigma_algebra) borel_measurable_subalgebra:
+ assumes "N \<subseteq> sets M" "f \<in> borel_measurable (M\<lparr>sets:=N\<rparr>)"
+ shows "f \<in> borel_measurable M"
+ using assms unfolding measurable_def by auto
+
section "Borel spaces on euclidean spaces"
lemma lessThan_borel[simp, intro]:
@@ -1294,4 +1345,46 @@
using assms by auto
qed
+lemma (in sigma_algebra) borel_measurable_psuminf:
+ assumes "\<And>i. f i \<in> borel_measurable M"
+ shows "(\<lambda>x. (\<Sum>\<^isub>\<infinity> i. f i x)) \<in> borel_measurable M"
+ using assms unfolding psuminf_def
+ by (auto intro!: borel_measurable_SUP[unfolded SUPR_fun_expand])
+
+section "LIMSEQ is borel measurable"
+
+lemma (in sigma_algebra) borel_measurable_LIMSEQ:
+ fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
+ assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
+ and u: "\<And>i. u i \<in> borel_measurable M"
+ shows "u' \<in> borel_measurable M"
+proof -
+ let "?pu x i" = "max (u i x) 0"
+ let "?nu x i" = "max (- u i x) 0"
+
+ { fix x assume x: "x \<in> space M"
+ have "(?pu x) ----> max (u' x) 0"
+ "(?nu x) ----> max (- u' x) 0"
+ using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus)
+ from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)]
+ have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)"
+ "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)"
+ by (simp_all add: Real_max'[symmetric]) }
+ note eq = this
+
+ have *: "\<And>x. real (Real (u' x)) - real (Real (- u' x)) = u' x"
+ by auto
+
+ have "(SUP n. INF m. (\<lambda>x. Real (u (n + m) x))) \<in> borel_measurable M"
+ "(SUP n. INF m. (\<lambda>x. Real (- u (n + m) x))) \<in> borel_measurable M"
+ using u by (auto intro: borel_measurable_SUP borel_measurable_INF borel_measurable_Real)
+ with eq[THEN measurable_cong, of M "\<lambda>x. x" borel_space]
+ have "(\<lambda>x. Real (u' x)) \<in> borel_measurable M"
+ "(\<lambda>x. Real (- u' x)) \<in> borel_measurable M"
+ unfolding SUPR_fun_expand INFI_fun_expand by auto
+ note this[THEN borel_measurable_real]
+ from borel_measurable_diff[OF this]
+ show ?thesis unfolding * .
+qed
+
end
--- a/src/HOL/Probability/Information.thy Fri Aug 27 16:23:51 2010 +0200
+++ b/src/HOL/Probability/Information.thy Thu Sep 02 17:12:40 2010 +0200
@@ -359,6 +359,48 @@
"(\<Sum>x\<in>A \<times> B. f x) = (\<Sum>x\<in>A. setsum (\<lambda>y. f (x, y)) B)"
unfolding setsum_cartesian_product by simp
+lemma (in finite_information_space) mutual_information_generic_eq:
+ assumes MX: "finite_measure_space MX (distribution X)"
+ assumes MY: "finite_measure_space MY (distribution Y)"
+ shows "mutual_information b MX MY X Y = (\<Sum> (x,y) \<in> space MX \<times> space MY.
+ real (joint_distribution X Y {(x,y)}) *
+ log b (real (joint_distribution X Y {(x,y)}) /
+ (real (distribution X {x}) * real (distribution Y {y}))))"
+proof -
+ let ?P = "prod_measure_space MX MY"
+ let ?\<mu> = "prod_measure MX (distribution X) MY (distribution Y)"
+ let ?\<nu> = "joint_distribution X Y"
+ interpret X: finite_measure_space MX "distribution X" by fact
+ moreover interpret Y: finite_measure_space MY "distribution Y" by fact
+ have fms: "finite_measure_space MX (distribution X)"
+ "finite_measure_space MY (distribution Y)" by fact+
+ have fms_P: "finite_measure_space ?P ?\<mu>"
+ by (rule X.finite_measure_space_finite_prod_measure) fact
+ then interpret P: finite_measure_space ?P ?\<mu> .
+ have fms_P': "finite_measure_space ?P ?\<nu>"
+ using finite_product_measure_space[of "space MX" "space MY"]
+ X.finite_space Y.finite_space sigma_prod_sets_finite[OF X.finite_space Y.finite_space]
+ X.sets_eq_Pow Y.sets_eq_Pow
+ by (simp add: prod_measure_space_def sigma_def)
+ then interpret P': finite_measure_space ?P ?\<nu> .
+ { fix x assume "x \<in> space ?P"
+ hence in_MX: "{fst x} \<in> sets MX" "{snd x} \<in> sets MY" using X.sets_eq_Pow Y.sets_eq_Pow
+ by (auto simp: prod_measure_space_def)
+ assume "?\<mu> {x} = 0"
+ with X.finite_prod_measure_times[OF fms(2), of "{fst x}" "{snd x}"] in_MX
+ have "distribution X {fst x} = 0 \<or> distribution Y {snd x} = 0"
+ by (simp add: prod_measure_space_def)
+ hence "joint_distribution X Y {x} = 0"
+ by (cases x) (auto simp: distribution_order) }
+ note measure_0 = this
+ show ?thesis
+ unfolding Let_def mutual_information_def
+ using measure_0 fms_P fms_P' MX MY P.absolutely_continuous_def
+ by (subst P.KL_divergence_eq_finite)
+ (auto simp add: prod_measure_space_def prod_measure_times_finite
+ finite_prob_space_eq setsum_cartesian_product' real_of_pinfreal_mult[symmetric])
+qed
+
lemma (in finite_information_space)
assumes MX: "finite_prob_space MX (distribution X)"
assumes MY: "finite_prob_space MY (distribution Y)"
--- a/src/HOL/Probability/Lebesgue_Integration.thy Fri Aug 27 16:23:51 2010 +0200
+++ b/src/HOL/Probability/Lebesgue_Integration.thy Thu Sep 02 17:12:40 2010 +0200
@@ -209,19 +209,6 @@
by (auto intro!: **)
qed
-lemma setsum_indicator_disjoint_family:
- fixes f :: "'d \<Rightarrow> 'e::semiring_1"
- assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P"
- shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j"
-proof -
- have "P \<inter> {i. x \<in> A i} = {j}"
- using d `x \<in> A j` `j \<in> P` unfolding disjoint_family_on_def
- by auto
- thus ?thesis
- unfolding indicator_def
- by (simp add: if_distrib setsum_cases[OF `finite P`])
-qed
-
lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence:
fixes u :: "'a \<Rightarrow> pinfreal"
assumes u: "u \<in> borel_measurable M"
@@ -426,6 +413,62 @@
with x show thesis by (auto intro!: that[of f])
qed
+lemma (in sigma_algebra) simple_function_eq_borel_measurable:
+ fixes f :: "'a \<Rightarrow> pinfreal"
+ shows "simple_function f \<longleftrightarrow>
+ finite (f`space M) \<and> f \<in> borel_measurable M"
+ using simple_function_borel_measurable[of f]
+ borel_measurable_simple_function[of f]
+ by (fastsimp simp: simple_function_def)
+
+lemma (in measure_space) simple_function_restricted:
+ fixes f :: "'a \<Rightarrow> pinfreal" 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")
+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)"
+ proof cases
+ assume "A = space M"
+ then have "f`A = ?f`space M" by (fastsimp simp: image_iff)
+ then show ?thesis by simp
+ next
+ assume "A \<noteq> space M"
+ then obtain x where x: "x \<in> space M" "x \<notin> A"
+ using sets_into_space `A \<in> sets M` by auto
+ have *: "?f`space M = f`A \<union> {0}"
+ proof (auto simp add: image_iff)
+ show "\<exists>x\<in>space M. f x = 0 \<or> indicator A x = 0"
+ using x by (auto intro!: bexI[of _ x])
+ next
+ fix x assume "x \<in> A"
+ then show "\<exists>y\<in>space M. f x = f y * indicator A y"
+ using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x])
+ next
+ fix x
+ assume "indicator A x \<noteq> (0::pinfreal)"
+ then have "x \<in> A" by (auto simp: indicator_def split: split_if_asm)
+ moreover assume "x \<in> space M" "\<forall>y\<in>A. ?f x \<noteq> f y"
+ ultimately show "f x = 0" by auto
+ qed
+ then show ?thesis by auto
+ qed
+ then show ?thesis
+ unfolding simple_function_eq_borel_measurable
+ R.simple_function_eq_borel_measurable
+ unfolding borel_measurable_restricted[OF `A \<in> sets M`]
+ by auto
+qed
+
+lemma (in sigma_algebra) simple_function_subalgebra:
+ assumes "sigma_algebra.simple_function (M\<lparr>sets:=N\<rparr>) f"
+ and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr>sets:=N\<rparr>)"
+ shows "simple_function f"
+ using assms
+ unfolding simple_function_def
+ unfolding sigma_algebra.simple_function_def[OF N_subalgebra(2)]
+ by auto
+
section "Simple integral"
definition (in measure_space)
@@ -668,6 +711,41 @@
qed
qed
+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 = simple_integral (\<lambda>x. f x * indicator A x)"
+ (is "_ = simple_integral ?f")
+ unfolding measure_space.simple_integral_def[OF restricted_measure_space[OF `A \<in> sets M`]]
+ unfolding simple_integral_def
+proof (simp, safe intro!: setsum_mono_zero_cong_left)
+ from sf show "finite (?f ` space M)"
+ unfolding simple_function_def by auto
+next
+ fix x assume "x \<in> A"
+ then show "f x \<in> ?f ` space M"
+ using sets_into_space `A \<in> sets M` by (auto intro!: image_eqI[of _ _ x])
+next
+ fix x assume "x \<in> space M" "?f x \<notin> f`A"
+ then have "x \<notin> A" by (auto simp: image_iff)
+ then show "?f x * \<mu> (?f -` {?f x} \<inter> space M) = 0" by simp
+next
+ fix x assume "x \<in> A"
+ then have "f x \<noteq> 0 \<Longrightarrow>
+ f -` {f x} \<inter> A = ?f -` {f x} \<inter> space M"
+ using `A \<in> sets M` sets_into_space
+ by (auto simp: indicator_def split: split_if_asm)
+ then show "f x * \<mu> (f -` {f x} \<inter> A) =
+ f x * \<mu> (?f -` {f x} \<inter> space M)"
+ unfolding pinfreal_mult_cancel_left by auto
+qed
+
+lemma (in measure_space) simple_integral_subalgebra[simp]:
+ assumes "measure_space (M\<lparr>sets := N\<rparr>) \<mu>"
+ shows "measure_space.simple_integral (M\<lparr>sets := N\<rparr>) \<mu> = simple_integral"
+ unfolding simple_integral_def_raw
+ unfolding measure_space.simple_integral_def_raw[OF assms] by simp
+
section "Continuous posititve integration"
definition (in measure_space)
@@ -1077,6 +1155,43 @@
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. positive_integral (\<lambda>x. f x * indicator A x)) g =
+ positive_integral (\<lambda>x. f x * g x)" (is "measure_space.positive_integral M ?T _ = _")
+proof -
+ from measure_space_density[OF assms(1)]
+ interpret T: measure_space M ?T .
+ 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
+ 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" .
+ { 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)"
+ using G T.positive_integral_eq_simple_integral by simp
+ also have "\<dots> = positive_integral (\<lambda>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)
+ apply (subst positive_integral_cmult[symmetric])
+ using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage)
+ apply (subst positive_integral_setsum[symmetric])
+ using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage)
+ by (simp add: setsum_right_distrib field_simps)
+ also have "\<dots> = positive_integral (\<lambda>x. f x * G i x)"
+ by (auto intro!: positive_integral_cong
+ simp: indicator_def if_distrib setsum_cases)
+ finally have "T.positive_integral (G i) = positive_integral (\<lambda>x. f x * G i x)" . }
+ with * have eq_Tg: "(\<lambda>i. positive_integral (\<lambda>x. f x * G i x)) \<up> T.positive_integral 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. positive_integral (\<lambda>x. f x * G i x)) \<up> positive_integral (\<lambda>x. f x * g x)"
+ using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pinfreal_times)
+ with eq_Tg show "T.positive_integral g = positive_integral (\<lambda>x. f x * g x)"
+ unfolding isoton_def by simp
+qed
+
lemma (in measure_space) positive_integral_null_set:
assumes borel: "u \<in> borel_measurable M" and "N \<in> null_sets"
shows "positive_integral (\<lambda>x. u x * indicator N x) = 0" (is "?I = 0")
@@ -1222,6 +1337,58 @@
finally show ?thesis by simp
qed
+lemma (in measure_space) positive_integral_restricted:
+ assumes "A \<in> sets M"
+ shows "measure_space.positive_integral (restricted_space A) \<mu> f = positive_integral (\<lambda>x. f x * indicator A x)"
+ (is "measure_space.positive_integral ?R \<mu> f = positive_integral ?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 saR: "sigma_algebra ?R" by fact
+ have *: "R.positive_integral f = R.positive_integral ?f"
+ by (auto intro!: R.positive_integral_cong)
+ show ?thesis
+ unfolding * R.positive_integral_def 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: image_iff simple_integral_restricted[OF `A \<in> sets M`])
+ fix g assume "simple_function (\<lambda>x. g x * indicator A x)"
+ "g \<le> f" "\<forall>x\<in>A. \<omega> \<noteq> g x"
+ then show "\<exists>x. simple_function x \<and> x \<le> (\<lambda>x. f x * indicator A x) \<and> (\<forall>y\<in>space M. \<omega> \<noteq> x y) \<and>
+ simple_integral (\<lambda>x. g x * indicator A x) = simple_integral 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)"
+ "\<forall>x\<in>space M. \<omega> \<noteq> g 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 expand_fun_eq 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> (\<forall>xa\<in>A. \<omega> \<noteq> x xa) \<and>
+ simple_integral g = simple_integral (\<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 *)
+ qed
+qed
+
+lemma (in measure_space) positive_integral_subalgebra[simp]:
+ assumes borel: "f \<in> borel_measurable (M\<lparr>sets := N\<rparr>)"
+ and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr>sets := N\<rparr>)"
+ shows "measure_space.positive_integral (M\<lparr>sets := N\<rparr>) \<mu> f = positive_integral f"
+proof -
+ note msN = measure_space_subalgebra[OF N_subalgebra]
+ then interpret N: measure_space "M\<lparr>sets:=N\<rparr>" \<mu> .
+ 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_subalgebra] 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 simple_integral_subalgebra[OF msN] isoton_def by simp
+qed
+
section "Lebesgue Integral"
definition (in measure_space) integrable where
@@ -1629,44 +1796,6 @@
by (simp add: real_of_pinfreal_eq_0)
qed
-lemma LIMSEQ_max:
- "u ----> (x::real) \<Longrightarrow> (\<lambda>i. max (u i) 0) ----> max x 0"
- by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D)
-
-lemma (in sigma_algebra) borel_measurable_LIMSEQ:
- fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
- assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
- and u: "\<And>i. u i \<in> borel_measurable M"
- shows "u' \<in> borel_measurable M"
-proof -
- let "?pu x i" = "max (u i x) 0"
- let "?nu x i" = "max (- u i x) 0"
-
- { fix x assume x: "x \<in> space M"
- have "(?pu x) ----> max (u' x) 0"
- "(?nu x) ----> max (- u' x) 0"
- using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus)
- from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)]
- have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)"
- "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)"
- by (simp_all add: Real_max'[symmetric]) }
- note eq = this
-
- have *: "\<And>x. real (Real (u' x)) - real (Real (- u' x)) = u' x"
- by auto
-
- have "(SUP n. INF m. (\<lambda>x. Real (u (n + m) x))) \<in> borel_measurable M"
- "(SUP n. INF m. (\<lambda>x. Real (- u (n + m) x))) \<in> borel_measurable M"
- using u by (auto intro: borel_measurable_SUP borel_measurable_INF borel_measurable_Real)
- with eq[THEN measurable_cong, of M "\<lambda>x. x" borel_space]
- have "(\<lambda>x. Real (u' x)) \<in> borel_measurable M"
- "(\<lambda>x. Real (- u' x)) \<in> borel_measurable M"
- unfolding SUPR_fun_expand INFI_fun_expand by auto
- note this[THEN borel_measurable_real]
- from borel_measurable_diff[OF this]
- show ?thesis unfolding * .
-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"
@@ -1926,41 +2055,11 @@
by (simp_all add: integral_cmul_indicator borel_measurable_vimage)
qed
-lemma sigma_algebra_cong:
- fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme"
- assumes *: "sigma_algebra M"
- and cong: "space M = space M'" "sets M = sets M'"
- shows "sigma_algebra M'"
-using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong .
-
-lemma finite_Pow_additivity_sufficient:
- assumes "finite (space M)" and "sets M = Pow (space M)"
- and "positive \<mu>" and "additive M \<mu>"
- and "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
- shows "finite_measure_space M \<mu>"
-proof -
- have "sigma_algebra M"
- using assms by (auto intro!: sigma_algebra_cong[OF sigma_algebra_Pow])
-
- have "measure_space M \<mu>"
- by (rule sigma_algebra.finite_additivity_sufficient) (fact+)
- thus ?thesis
- unfolding finite_measure_space_def finite_measure_space_axioms_def
- using assms by simp
-qed
-
-lemma finite_measure_spaceI:
- assumes "measure_space M \<mu>" and "finite (space M)" and "sets M = Pow (space M)"
- and "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
- shows "finite_measure_space M \<mu>"
- unfolding finite_measure_space_def finite_measure_space_axioms_def
- using assms by simp
+lemma (in finite_measure_space) simple_function_finite[simp, intro]: "simple_function f"
+ unfolding simple_function_def sets_eq_Pow using finite_space by auto
lemma (in finite_measure_space) borel_measurable_finite[intro, simp]: "f \<in> borel_measurable M"
- unfolding measurable_def sets_eq_Pow by auto
-
-lemma (in finite_measure_space) simple_function_finite: "simple_function f"
- unfolding simple_function_def sets_eq_Pow using finite_space by auto
+ 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})"
@@ -1979,10 +2078,8 @@
"positive_integral (\<lambda>x. Real (f x)) = (\<Sum>x \<in> space M. Real (f x) * \<mu> {x})"
"positive_integral (\<lambda>x. Real (- f x)) = (\<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
by (simp add: setsum_\<omega> integrable_def sets_eq_Pow)
-
show ?I using finite_measure
apply (simp add: integral_def sets_eq_Pow real_of_pinfreal_setsum[symmetric]
real_of_pinfreal_mult[symmetric] setsum_subtractf[symmetric])
--- a/src/HOL/Probability/Measure.thy Fri Aug 27 16:23:51 2010 +0200
+++ b/src/HOL/Probability/Measure.thy Thu Sep 02 17:12:40 2010 +0200
@@ -414,6 +414,19 @@
finally show ?thesis .
qed
+lemma (in measure_space) measure_finitely_subadditive:
+ assumes "finite I" "A ` I \<subseteq> sets M"
+ shows "\<mu> (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu> (A i))"
+using assms proof induct
+ case (insert i I)
+ then have "(\<Union>i\<in>I. A i) \<in> sets M" by (auto intro: finite_UN)
+ then have "\<mu> (\<Union>i\<in>insert i I. A i) \<le> \<mu> (A i) + \<mu> (\<Union>i\<in>I. A i)"
+ using insert by (simp add: measure_subadditive)
+ also have "\<dots> \<le> (\<Sum>i\<in>insert i I. \<mu> (A i))"
+ using insert by (auto intro!: add_left_mono)
+ finally show ?case .
+qed simp
+
lemma (in measure_space) measurable_countably_subadditive:
assumes "range f \<subseteq> sets M"
shows "\<mu> (\<Union>i. f i) \<le> (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
@@ -432,9 +445,34 @@
finally show ?thesis .
qed
+lemma (in measure_space) measure_inter_full_set:
+ assumes "S \<in> sets M" "T \<in> sets M" and not_\<omega>: "\<mu> (T - S) \<noteq> \<omega>"
+ assumes T: "\<mu> T = \<mu> (space M)"
+ shows "\<mu> (S \<inter> T) = \<mu> S"
+proof (rule antisym)
+ show " \<mu> (S \<inter> T) \<le> \<mu> S"
+ using assms by (auto intro!: measure_mono)
+
+ show "\<mu> S \<le> \<mu> (S \<inter> T)"
+ proof (rule ccontr)
+ assume contr: "\<not> ?thesis"
+ have "\<mu> (space M) = \<mu> ((T - S) \<union> (S \<inter> T))"
+ unfolding T[symmetric] by (auto intro!: arg_cong[where f="\<mu>"])
+ also have "\<dots> \<le> \<mu> (T - S) + \<mu> (S \<inter> T)"
+ using assms by (auto intro!: measure_subadditive)
+ also have "\<dots> < \<mu> (T - S) + \<mu> S"
+ by (rule pinfreal_less_add[OF not_\<omega>]) (insert contr, auto)
+ also have "\<dots> = \<mu> (T \<union> S)"
+ using assms by (subst measure_additive) auto
+ also have "\<dots> \<le> \<mu> (space M)"
+ using assms sets_into_space by (auto intro!: measure_mono)
+ finally show False ..
+ qed
+qed
+
lemma (in measure_space) restricted_measure_space:
assumes "S \<in> sets M"
- shows "measure_space (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>"
+ shows "measure_space (restricted_space S) \<mu>"
(is "measure_space ?r \<mu>")
unfolding measure_space_def measure_space_axioms_def
proof safe
@@ -477,6 +515,20 @@
qed
qed
+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>"
+proof -
+ interpret N: sigma_algebra "M\<lparr> sets := N \<rparr>" by fact
+ show ?thesis
+ proof
+ show "countably_additive (M\<lparr>sets := N\<rparr>) \<mu>"
+ using `N \<subseteq> sets M`
+ by (auto simp add: countably_additive_def
+ intro!: measure_countably_additive)
+ qed simp
+qed
+
section "@{text \<sigma>}-finite Measures"
locale sigma_finite_measure = measure_space +
@@ -484,7 +536,7 @@
lemma (in sigma_finite_measure) restricted_sigma_finite_measure:
assumes "S \<in> sets M"
- shows "sigma_finite_measure (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>"
+ shows "sigma_finite_measure (restricted_space S) \<mu>"
(is "sigma_finite_measure ?r _")
unfolding sigma_finite_measure_def sigma_finite_measure_axioms_def
proof safe
@@ -512,6 +564,25 @@
qed
qed
+lemma (in sigma_finite_measure) disjoint_sigma_finite:
+ "\<exists>A::nat\<Rightarrow>'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and>
+ (\<forall>i. \<mu> (A i) \<noteq> \<omega>) \<and> disjoint_family A"
+proof -
+ obtain A :: "nat \<Rightarrow> 'a set" where
+ range: "range A \<subseteq> sets M" and
+ space: "(\<Union>i. A i) = space M" and
+ measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+ using sigma_finite by auto
+ note range' = range_disjointed_sets[OF range] range
+ { fix i
+ have "\<mu> (disjointed A i) \<le> \<mu> (A i)"
+ using range' disjointed_subset[of A i] by (auto intro!: measure_mono)
+ then have "\<mu> (disjointed A i) \<noteq> \<omega>"
+ using measure[of i] by auto }
+ with disjoint_family_disjointed UN_disjointed_eq[of A] space range'
+ show ?thesis by (auto intro!: exI[of _ "disjointed A"])
+qed
+
section "Real measure values"
lemma (in measure_space) real_measure_Union:
@@ -630,7 +701,7 @@
using finite_measure_of_space by (auto intro!: exI[of _ "\<lambda>x. space M"])
qed
-lemma (in finite_measure) finite_measure:
+lemma (in finite_measure) finite_measure[simp, intro]:
assumes "A \<in> sets M"
shows "\<mu> A \<noteq> \<omega>"
proof -
@@ -645,7 +716,7 @@
lemma (in finite_measure) restricted_finite_measure:
assumes "S \<in> sets M"
- shows "finite_measure (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>"
+ shows "finite_measure (restricted_space S) \<mu>"
(is "finite_measure ?r _")
unfolding finite_measure_def finite_measure_axioms_def
proof (safe del: notI)
@@ -733,6 +804,13 @@
shows "\<mu> (space M - s) = \<mu> (space M) - \<mu> s"
using measure_compl[OF s, OF finite_measure, OF s] .
+lemma (in finite_measure) finite_measure_inter_full_set:
+ assumes "S \<in> sets M" "T \<in> sets M"
+ assumes T: "\<mu> T = \<mu> (space M)"
+ shows "\<mu> (S \<inter> T) = \<mu> S"
+ using measure_inter_full_set[OF assms(1,2) finite_measure assms(3)] assms
+ by auto
+
section {* Measure preserving *}
definition "measure_preserving A \<mu> B \<nu> =
@@ -843,10 +921,51 @@
and sets_eq_Pow: "sets M = Pow (space M)"
and finite_single_measure: "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
+lemma (in finite_measure_space) sets_image_space_eq_Pow:
+ "sets (image_space X) = Pow (space (image_space X))"
+proof safe
+ fix x S assume "S \<in> sets (image_space X)" "x \<in> S"
+ then show "x \<in> space (image_space X)"
+ using sets_into_space by (auto intro!: imageI simp: image_space_def)
+next
+ fix S assume "S \<subseteq> space (image_space X)"
+ then obtain S' where "S = X`S'" "S'\<in>sets M"
+ by (auto simp: subset_image_iff sets_eq_Pow image_space_def)
+ then show "S \<in> sets (image_space X)"
+ by (auto simp: image_space_def)
+qed
+
lemma (in finite_measure_space) sum_over_space: "(\<Sum>x\<in>space M. \<mu> {x}) = \<mu> (space M)"
using measure_finitely_additive''[of "space M" "\<lambda>i. {i}"]
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>"
+ unfolding finite_measure_space_def finite_measure_space_axioms_def
+proof (safe del: notI)
+ show "measure_space M \<mu>"
+ proof (rule sigma_algebra.finite_additivity_sufficient)
+ show "sigma_algebra M"
+ apply (rule sigma_algebra_cong)
+ apply (rule sigma_algebra_Pow[of "space M"])
+ using assms by simp_all
+ 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
+ qed
+ show "finite (space M)" by fact
+ { fix A x assume "A \<in> sets M" "x \<in> A" then show "x \<in> space M"
+ using assms by auto }
+ { fix A assume "A \<subseteq> space M" then show "A \<in> sets M"
+ using assms by auto }
+ { fix x assume *: "x \<in> space M"
+ with add[of "{x}" "space M - {x}"] space
+ show "\<mu> {x} \<noteq> \<omega>" by (auto simp: insert_absorb[OF *] Diff_subset) }
+qed
+
sublocale finite_measure_space < finite_measure
proof
show "\<mu> (space M) \<noteq> \<omega>"
@@ -854,6 +973,22 @@
using finite_space finite_single_measure by auto
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)"
+ (is "_ = ?rhs")
+proof (intro iffI)
+ assume "finite_measure_space M \<mu>"
+ then interpret finite_measure_space M \<mu> .
+ 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>"
+ by (auto intro!: finite_measure_spaceI)
+qed
+
lemma psuminf_cmult_indicator:
assumes "disjoint_family A" "x \<in> A i"
shows "(\<Sum>\<^isub>\<infinity> n. f n * indicator (A n) x) = f i"
--- a/src/HOL/Probability/Positive_Infinite_Real.thy Fri Aug 27 16:23:51 2010 +0200
+++ b/src/HOL/Probability/Positive_Infinite_Real.thy Thu Sep 02 17:12:40 2010 +0200
@@ -411,6 +411,10 @@
lemma pinfreal_less_\<omega>: "x < \<omega> \<longleftrightarrow> x \<noteq> \<omega>"
by (cases x) auto
+lemma pinfreal_0_less_mult_iff[simp]:
+ fixes x y :: pinfreal shows "0 < x * y \<longleftrightarrow> 0 < x \<and> 0 < y"
+ by (cases x, cases y) (auto simp: zero_less_mult_iff)
+
subsection {* @{text "x - y"} on @{typ pinfreal} *}
instantiation pinfreal :: minus
--- a/src/HOL/Probability/Probability_Space.thy Fri Aug 27 16:23:51 2010 +0200
+++ b/src/HOL/Probability/Probability_Space.thy Thu Sep 02 17:12:40 2010 +0200
@@ -2,38 +2,6 @@
imports Lebesgue_Integration Radon_Nikodym
begin
-lemma (in measure_space) measure_inter_full_set:
- assumes "S \<in> sets M" "T \<in> sets M" and not_\<omega>: "\<mu> (T - S) \<noteq> \<omega>"
- assumes T: "\<mu> T = \<mu> (space M)"
- shows "\<mu> (S \<inter> T) = \<mu> S"
-proof (rule antisym)
- show " \<mu> (S \<inter> T) \<le> \<mu> S"
- using assms by (auto intro!: measure_mono)
-
- show "\<mu> S \<le> \<mu> (S \<inter> T)"
- proof (rule ccontr)
- assume contr: "\<not> ?thesis"
- have "\<mu> (space M) = \<mu> ((T - S) \<union> (S \<inter> T))"
- unfolding T[symmetric] by (auto intro!: arg_cong[where f="\<mu>"])
- also have "\<dots> \<le> \<mu> (T - S) + \<mu> (S \<inter> T)"
- using assms by (auto intro!: measure_subadditive)
- also have "\<dots> < \<mu> (T - S) + \<mu> S"
- by (rule pinfreal_less_add[OF not_\<omega>]) (insert contr, auto)
- also have "\<dots> = \<mu> (T \<union> S)"
- using assms by (subst measure_additive) auto
- also have "\<dots> \<le> \<mu> (space M)"
- using assms sets_into_space by (auto intro!: measure_mono)
- finally show False ..
- qed
-qed
-
-lemma (in finite_measure) finite_measure_inter_full_set:
- assumes "S \<in> sets M" "T \<in> sets M"
- assumes T: "\<mu> T = \<mu> (space M)"
- shows "\<mu> (S \<inter> T) = \<mu> S"
- using measure_inter_full_set[OF assms(1,2) finite_measure assms(3)] assms
- by auto
-
locale prob_space = measure_space +
assumes measure_space_1: "\<mu> (space M) = 1"
@@ -75,10 +43,6 @@
finally show ?thesis .
qed
-lemma measure_finite[simp, intro]:
- assumes "A \<in> events" shows "\<mu> A \<noteq> \<omega>"
- using measure_le_1[OF assms] by auto
-
lemma prob_compl:
assumes "A \<in> events"
shows "prob (space M - A) = 1 - prob A"
@@ -361,51 +325,64 @@
joint_distribution_restriction_snd[of X Y "{(x, y)}"]
by auto
-lemma (in finite_prob_space) finite_product_measure_space:
- 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")
-proof (rule finite_Pow_additivity_sufficient)
- show "positive ?D"
- unfolding positive_def using assms sets_eq_Pow
- by (simp add: distribution_def)
+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)"
+ by (simp add: finite_prob_space_eq finite_measure_space)
- show "additive ?M ?D" unfolding additive_def
- proof safe
- fix x y
- have A: "((\<lambda>x. (X x, Y x)) -` x) \<inter> space M \<in> sets M" using assms sets_eq_Pow by auto
- have B: "((\<lambda>x. (X x, Y x)) -` y) \<inter> space M \<in> sets M" using assms sets_eq_Pow by auto
- assume "x \<inter> y = {}"
- hence "(\<lambda>x. (X x, Y x)) -` x \<inter> space M \<inter> ((\<lambda>x. (X x, Y x)) -` y \<inter> space M) = {}"
- by auto
- from additive[unfolded additive_def, rule_format, OF A B] this
- finite_measure[OF A] finite_measure[OF B]
- show "?D (x \<union> y) = ?D x + ?D y"
- apply (simp add: distribution_def)
- apply (subst Int_Un_distrib2)
- by (auto simp: real_of_pinfreal_add)
- qed
+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 _")
+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"
+ by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1)
+qed
- show "finite (space ?M)"
- using assms by auto
-
- show "sets ?M = Pow (space ?M)"
- by simp
-
- { fix x assume "x \<in> space ?M" thus "?D {x} \<noteq> \<omega>"
- unfolding distribution_def by (auto intro!: finite_measure simp: sets_eq_Pow) }
+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>"
+proof -
+ interpret N: measure_space "M\<lparr> sets := N \<rparr>" \<mu>
+ using measure_space_subalgebra[OF assms] .
+ show ?thesis
+ proof qed (simp add: measure_space_1)
qed
-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)"
- using finite_space by (auto intro!: finite_product_measure_space)
+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: pinfreal_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>
+ using `A \<in> sets M` by (rule restricted_measure_space)
+ show "measure_space (restricted_space A) (\<lambda>S. \<mu> S / \<mu> A)"
+ 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 *
+ 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>"
+ unfolding finite_prob_space_eq
+proof
+ show "finite_measure_space M \<mu>" using assms
+ by (auto intro!: finite_measure_spaceI)
+ show "\<mu> (space M) = 1" by fact
+qed
lemma (in finite_prob_space) finite_measure_space:
shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)"
(is "finite_measure_space ?S _")
-proof (rule finite_Pow_additivity_sufficient, simp_all)
+proof (rule finite_measure_spaceI, simp_all)
show "finite (X ` space M)" using finite_space by simp
show "positive (distribution X)"
@@ -431,69 +408,6 @@
unfolding distribution_def by (auto intro!: finite_measure simp: sets_eq_Pow) }
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)"
- by (simp add: finite_prob_space_eq finite_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 _")
-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"
- by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1)
-qed
-
-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
-
-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
-
-lemma pinfreal_0_less_mult_iff[simp]:
- fixes x y :: pinfreal shows "0 < x * y \<longleftrightarrow> 0 < x \<and> 0 < y"
- by (cases x, cases y) (auto simp: zero_less_mult_iff)
-
-lemma (in sigma_algebra) simple_function_subalgebra:
- assumes "sigma_algebra.simple_function (M\<lparr>sets:=N\<rparr>) f"
- and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr>sets:=N\<rparr>)"
- shows "simple_function f"
- using assms
- unfolding simple_function_def
- unfolding sigma_algebra.simple_function_def[OF N_subalgebra(2)]
- by auto
-
-lemma (in measure_space) simple_integral_subalgebra[simp]:
- assumes "measure_space (M\<lparr>sets := N\<rparr>) \<mu>"
- shows "measure_space.simple_integral (M\<lparr>sets := N\<rparr>) \<mu> = simple_integral"
- unfolding simple_integral_def_raw
- unfolding measure_space.simple_integral_def_raw[OF assms] by simp
-
-lemma (in sigma_algebra) borel_measurable_subalgebra:
- assumes "N \<subseteq> sets M" "f \<in> borel_measurable (M\<lparr>sets:=N\<rparr>)"
- shows "f \<in> borel_measurable M"
- using assms unfolding measurable_def by auto
-
-lemma (in measure_space) positive_integral_subalgebra[simp]:
- assumes borel: "f \<in> borel_measurable (M\<lparr>sets := N\<rparr>)"
- and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr>sets := N\<rparr>)"
- shows "measure_space.positive_integral (M\<lparr>sets := N\<rparr>) \<mu> f = positive_integral f"
-proof -
- note msN = measure_space_subalgebra[OF N_subalgebra]
- then interpret N: measure_space "M\<lparr>sets:=N\<rparr>" \<mu> .
-
- 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_subalgebra] 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 simple_integral_subalgebra[OF msN] isoton_def by simp
-qed
-
section "Conditional Expectation and Probability"
lemma (in prob_space) conditional_expectation_exists:
@@ -541,7 +455,7 @@
\<and> (\<forall>C\<in>N. positive_integral (\<lambda>x. Y x * indicator C x) = positive_integral (\<lambda>x. X x * indicator C x)))"
abbreviation (in prob_space)
- "conditional_probabiltiy N A \<equiv> conditional_expectation N (indicator A)"
+ "conditional_prob N A \<equiv> conditional_expectation N (indicator A)"
lemma (in prob_space)
fixes X :: "'a \<Rightarrow> pinfreal"
--- a/src/HOL/Probability/Product_Measure.thy Fri Aug 27 16:23:51 2010 +0200
+++ b/src/HOL/Probability/Product_Measure.thy Thu Sep 02 17:12:40 2010 +0200
@@ -403,4 +403,45 @@
unfolding finite_prod_measure_space[OF N, symmetric]
using finite_measure_space_finite_prod_measure[OF N] .
+lemma (in finite_prob_space) finite_product_measure_space:
+ 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")
+proof (rule finite_Pow_additivity_sufficient)
+ show "positive ?D"
+ unfolding positive_def using assms sets_eq_Pow
+ by (simp add: distribution_def)
+
+ show "additive ?M ?D" unfolding additive_def
+ proof safe
+ fix x y
+ have A: "((\<lambda>x. (X x, Y x)) -` x) \<inter> space M \<in> sets M" using assms sets_eq_Pow by auto
+ have B: "((\<lambda>x. (X x, Y x)) -` y) \<inter> space M \<in> sets M" using assms sets_eq_Pow by auto
+ assume "x \<inter> y = {}"
+ hence "(\<lambda>x. (X x, Y x)) -` x \<inter> space M \<inter> ((\<lambda>x. (X x, Y x)) -` y \<inter> space M) = {}"
+ by auto
+ from additive[unfolded additive_def, rule_format, OF A B] this
+ finite_measure[OF A] finite_measure[OF B]
+ show "?D (x \<union> y) = ?D x + ?D y"
+ apply (simp add: distribution_def)
+ apply (subst Int_Un_distrib2)
+ by (auto simp: real_of_pinfreal_add)
+ qed
+
+ show "finite (space ?M)"
+ using assms by auto
+
+ show "sets ?M = Pow (space ?M)"
+ by simp
+
+ { fix x assume "x \<in> space ?M" thus "?D {x} \<noteq> \<omega>"
+ unfolding distribution_def by (auto intro!: finite_measure simp: sets_eq_Pow) }
+qed
+
+lemma (in finite_measure_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)"
+ using finite_space by (auto intro!: finite_product_measure_space)
+
end
\ No newline at end of file
--- a/src/HOL/Probability/Radon_Nikodym.thy Fri Aug 27 16:23:51 2010 +0200
+++ b/src/HOL/Probability/Radon_Nikodym.thy Thu Sep 02 17:12:40 2010 +0200
@@ -2,201 +2,6 @@
imports Lebesgue_Integration
begin
-lemma (in measure_space) measure_finitely_subadditive:
- assumes "finite I" "A ` I \<subseteq> sets M"
- shows "\<mu> (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu> (A i))"
-using assms proof induct
- case (insert i I)
- then have "(\<Union>i\<in>I. A i) \<in> sets M" by (auto intro: finite_UN)
- then have "\<mu> (\<Union>i\<in>insert i I. A i) \<le> \<mu> (A i) + \<mu> (\<Union>i\<in>I. A i)"
- using insert by (simp add: measure_subadditive)
- also have "\<dots> \<le> (\<Sum>i\<in>insert i I. \<mu> (A i))"
- using insert by (auto intro!: add_left_mono)
- finally show ?case .
-qed simp
-
-lemma (in sigma_algebra) borel_measurable_restricted:
- fixes f :: "'a \<Rightarrow> pinfreal" assumes "A \<in> sets M"
- shows "f \<in> borel_measurable (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) \<longleftrightarrow>
- (\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
- (is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M")
-proof -
- interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
- have *: "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable ?R"
- by (auto intro!: measurable_cong)
- show ?thesis unfolding *
- unfolding in_borel_measurable_borel_space
- proof (simp, safe)
- fix S :: "pinfreal set" assume "S \<in> sets borel_space"
- "\<forall>S\<in>sets borel_space. ?f -` S \<inter> A \<in> op \<inter> A ` sets M"
- then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto
- then have f: "?f -` S \<inter> A \<in> sets M"
- using `A \<in> sets M` sets_into_space by fastsimp
- show "?f -` S \<inter> space M \<in> sets M"
- proof cases
- assume "0 \<in> S"
- then have "?f -` S \<inter> space M = ?f -` S \<inter> A \<union> (space M - A)"
- using `A \<in> sets M` sets_into_space by auto
- then show ?thesis using f `A \<in> sets M` by (auto intro!: Un Diff)
- next
- assume "0 \<notin> S"
- then have "?f -` S \<inter> space M = ?f -` S \<inter> A"
- using `A \<in> sets M` sets_into_space
- by (auto simp: indicator_def split: split_if_asm)
- then show ?thesis using f by auto
- qed
- next
- fix S :: "pinfreal set" assume "S \<in> sets borel_space"
- "\<forall>S\<in>sets borel_space. ?f -` S \<inter> space M \<in> sets M"
- then have f: "?f -` S \<inter> space M \<in> sets M" by auto
- then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M"
- using `A \<in> sets M` sets_into_space
- apply (simp add: image_iff)
- apply (rule bexI[OF _ f])
- by auto
- qed
-qed
-
-lemma (in sigma_algebra) simple_function_eq_borel_measurable:
- fixes f :: "'a \<Rightarrow> pinfreal"
- shows "simple_function f \<longleftrightarrow>
- finite (f`space M) \<and> f \<in> borel_measurable M"
- using simple_function_borel_measurable[of f]
- borel_measurable_simple_function[of f]
- by (fastsimp simp: simple_function_def)
-
-lemma (in measure_space) simple_function_restricted:
- fixes f :: "'a \<Rightarrow> pinfreal" assumes "A \<in> sets M"
- shows "sigma_algebra.simple_function (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) f \<longleftrightarrow> simple_function (\<lambda>x. f x * indicator A x)"
- (is "sigma_algebra.simple_function ?R f \<longleftrightarrow> simple_function ?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)"
- proof cases
- assume "A = space M"
- then have "f`A = ?f`space M" by (fastsimp simp: image_iff)
- then show ?thesis by simp
- next
- assume "A \<noteq> space M"
- then obtain x where x: "x \<in> space M" "x \<notin> A"
- using sets_into_space `A \<in> sets M` by auto
- have *: "?f`space M = f`A \<union> {0}"
- proof (auto simp add: image_iff)
- show "\<exists>x\<in>space M. f x = 0 \<or> indicator A x = 0"
- using x by (auto intro!: bexI[of _ x])
- next
- fix x assume "x \<in> A"
- then show "\<exists>y\<in>space M. f x = f y * indicator A y"
- using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x])
- next
- fix x
- assume "indicator A x \<noteq> (0::pinfreal)"
- then have "x \<in> A" by (auto simp: indicator_def split: split_if_asm)
- moreover assume "x \<in> space M" "\<forall>y\<in>A. ?f x \<noteq> f y"
- ultimately show "f x = 0" by auto
- qed
- then show ?thesis by auto
- qed
- then show ?thesis
- unfolding simple_function_eq_borel_measurable
- R.simple_function_eq_borel_measurable
- unfolding borel_measurable_restricted[OF `A \<in> sets M`]
- by auto
-qed
-
-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 (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) \<mu> f = simple_integral (\<lambda>x. f x * indicator A x)"
- (is "_ = simple_integral ?f")
- unfolding measure_space.simple_integral_def[OF restricted_measure_space[OF `A \<in> sets M`]]
- unfolding simple_integral_def
-proof (simp, safe intro!: setsum_mono_zero_cong_left)
- from sf show "finite (?f ` space M)"
- unfolding simple_function_def by auto
-next
- fix x assume "x \<in> A"
- then show "f x \<in> ?f ` space M"
- using sets_into_space `A \<in> sets M` by (auto intro!: image_eqI[of _ _ x])
-next
- fix x assume "x \<in> space M" "?f x \<notin> f`A"
- then have "x \<notin> A" by (auto simp: image_iff)
- then show "?f x * \<mu> (?f -` {?f x} \<inter> space M) = 0" by simp
-next
- fix x assume "x \<in> A"
- then have "f x \<noteq> 0 \<Longrightarrow>
- f -` {f x} \<inter> A = ?f -` {f x} \<inter> space M"
- using `A \<in> sets M` sets_into_space
- by (auto simp: indicator_def split: split_if_asm)
- then show "f x * \<mu> (f -` {f x} \<inter> A) =
- f x * \<mu> (?f -` {f x} \<inter> space M)"
- unfolding pinfreal_mult_cancel_left by auto
-qed
-
-lemma (in measure_space) positive_integral_restricted:
- assumes "A \<in> sets M"
- shows "measure_space.positive_integral (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) \<mu> f = positive_integral (\<lambda>x. f x * indicator A x)"
- (is "measure_space.positive_integral ?R \<mu> f = positive_integral ?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 saR: "sigma_algebra ?R" by fact
- have *: "R.positive_integral f = R.positive_integral ?f"
- by (auto intro!: R.positive_integral_cong)
- show ?thesis
- unfolding * R.positive_integral_def 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: image_iff simple_integral_restricted[OF `A \<in> sets M`])
- fix g assume "simple_function (\<lambda>x. g x * indicator A x)"
- "g \<le> f" "\<forall>x\<in>A. \<omega> \<noteq> g x"
- then show "\<exists>x. simple_function x \<and> x \<le> (\<lambda>x. f x * indicator A x) \<and> (\<forall>y\<in>space M. \<omega> \<noteq> x y) \<and>
- simple_integral (\<lambda>x. g x * indicator A x) = simple_integral 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)"
- "\<forall>x\<in>space M. \<omega> \<noteq> g 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 expand_fun_eq 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> (\<forall>xa\<in>A. \<omega> \<noteq> x xa) \<and>
- simple_integral g = simple_integral (\<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 *)
- qed
-qed
-
-lemma (in sigma_algebra) borel_measurable_psuminf:
- assumes "\<And>i. f i \<in> borel_measurable M"
- shows "(\<lambda>x. (\<Sum>\<^isub>\<infinity> i. f i x)) \<in> borel_measurable M"
- using assms unfolding psuminf_def
- by (auto intro!: borel_measurable_SUP[unfolded SUPR_fun_expand])
-
-lemma (in sigma_finite_measure) disjoint_sigma_finite:
- "\<exists>A::nat\<Rightarrow>'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and>
- (\<forall>i. \<mu> (A i) \<noteq> \<omega>) \<and> disjoint_family A"
-proof -
- obtain A :: "nat \<Rightarrow> 'a set" where
- range: "range A \<subseteq> sets M" and
- space: "(\<Union>i. A i) = space M" and
- measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
- using sigma_finite by auto
-
- note range' = range_disjointed_sets[OF range] range
-
- { fix i
- have "\<mu> (disjointed A i) \<le> \<mu> (A i)"
- using range' disjointed_subset[of A i] by (auto intro!: measure_mono)
- then have "\<mu> (disjointed A i) \<noteq> \<omega>"
- using measure[of i] by auto }
- with disjoint_family_disjointed UN_disjointed_eq[of A] space range'
- show ?thesis by (auto intro!: exI[of _ "disjointed A"])
-qed
-
lemma (in sigma_finite_measure) 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>)"
proof -
@@ -206,7 +11,6 @@
measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>" and
disjoint: "disjoint_family A"
using disjoint_sigma_finite by auto
-
let "?B i" = "2^Suc i * \<mu> (A i)"
have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)"
proof
@@ -225,20 +29,22 @@
qed
from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
"\<And>i. n i < inverse (2^Suc i * \<mu> (A i))" by auto
-
let "?h x" = "\<Sum>\<^isub>\<infinity> i. n i * indicator (A i) x"
show ?thesis
proof (safe intro!: bexI[of _ ?h] del: notI)
- have "positive_integral ?h = (\<Sum>\<^isub>\<infinity> i. n i * \<mu> (A i))"
- apply (subst positive_integral_psuminf)
- using range apply (fastsimp intro!: borel_measurable_pinfreal_times borel_measurable_const borel_measurable_indicator)
- apply (subst positive_integral_cmult_indicator)
- using range by auto
+ 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))"
+ 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)
fix N show "n N * \<mu> (A N) \<le> Real ((1 / 2) ^ Suc N)"
using measure[of N] n[of N]
- by (cases "n N") (auto simp: pinfreal_noteq_omega_Ex field_simps zero_le_mult_iff mult_le_0_iff mult_less_0_iff power_less_zero_eq power_le_zero_eq inverse_eq_divide less_divide_eq power_divide split: split_if_asm)
+ by (cases "n N")
+ (auto simp: pinfreal_noteq_omega_Ex field_simps zero_le_mult_iff
+ mult_le_0_iff mult_less_0_iff power_less_zero_eq
+ power_le_zero_eq inverse_eq_divide less_divide_eq
+ power_divide split: split_if_asm)
qed
also have "\<dots> = Real 1"
by (rule suminf_imp_psuminf, rule power_half_series, auto)
@@ -251,7 +57,7 @@
then show "0 < ?h x" and "?h x < \<omega>" using n[of i] by auto
next
show "?h \<in> borel_measurable M" using range
- by (auto intro!: borel_measurable_psuminf borel_measurable_pinfreal_times borel_measurable_indicator)
+ by (auto intro!: borel_measurable_psuminf borel_measurable_pinfreal_times)
qed
qed
@@ -370,7 +176,7 @@
interpret M': finite_measure M s by fact
- let "?r S" = "M\<lparr> space := S, sets := (\<lambda>C. S \<inter> C)`sets M\<rparr>"
+ let "?r S" = "restricted_space S"
{ fix S n
assume S: "S \<in> sets M"
@@ -838,7 +644,7 @@
= (f x * indicator (Q i) x) * indicator A x"
unfolding indicator_def by auto
- have fm: "finite_measure (M\<lparr>space := Q i, sets := op \<inter> (Q i) ` sets M\<rparr>) \<mu>"
+ have fm: "finite_measure (restricted_space (Q i)) \<mu>"
(is "finite_measure ?R \<mu>") by (rule restricted_finite_measure[OF Q_sets[of i]])
then interpret R: finite_measure ?R .
have fmv: "finite_measure ?R \<nu>"
@@ -935,47 +741,6 @@
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. positive_integral (\<lambda>x. f x * indicator A x)) g =
- positive_integral (\<lambda>x. f x * g x)" (is "measure_space.positive_integral M ?T _ = _")
-proof -
- from measure_space_density[OF assms(1)]
- interpret T: measure_space M ?T .
-
- 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
- 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" .
-
- { 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)"
- using G T.positive_integral_eq_simple_integral by simp
- also have "\<dots> = positive_integral (\<lambda>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)
- apply (subst positive_integral_cmult[symmetric])
- using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage)
- apply (subst positive_integral_setsum[symmetric])
- using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage)
- by (simp add: setsum_right_distrib field_simps)
- also have "\<dots> = positive_integral (\<lambda>x. f x * G i x)"
- by (auto intro!: positive_integral_cong
- simp: indicator_def if_distrib setsum_cases)
- finally have "T.positive_integral (G i) = positive_integral (\<lambda>x. f x * G i x)" . }
- with * have eq_Tg: "(\<lambda>i. positive_integral (\<lambda>x. f x * G i x)) \<up> T.positive_integral 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. positive_integral (\<lambda>x. f x * G i x)) \<up> positive_integral (\<lambda>x. f x * g x)"
- using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pinfreal_times)
- with eq_Tg show "T.positive_integral g = positive_integral (\<lambda>x. f x * g x)"
- unfolding isoton_def by simp
-qed
-
lemma (in sigma_finite_measure) Radon_Nikodym:
assumes "measure_space M \<nu>"
assumes "absolutely_continuous \<nu>"
--- a/src/HOL/Probability/Sigma_Algebra.thy Fri Aug 27 16:23:51 2010 +0200
+++ b/src/HOL/Probability/Sigma_Algebra.thy Thu Sep 02 17:12:40 2010 +0200
@@ -6,7 +6,7 @@
header {* Sigma Algebras *}
-theory Sigma_Algebra imports Main Countable FuncSet begin
+theory Sigma_Algebra imports Main Countable FuncSet Indicator_Function begin
text {* Sigma algebras are an elementary concept in measure
theory. To measure --- that is to integrate --- functions, we first have
@@ -95,10 +95,13 @@
lemma (in algebra) Int_space_eq2 [simp]: "x \<in> sets M \<Longrightarrow> x \<inter> space M = x"
by (metis Int_absorb2 sets_into_space)
+section {* Restricted algebras *}
+
+abbreviation (in algebra)
+ "restricted_space A \<equiv> \<lparr> space = A, sets = (\<lambda>S. (A \<inter> S)) ` sets M \<rparr>"
+
lemma (in algebra) restricted_algebra:
- assumes "S \<in> sets M"
- shows "algebra (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>)"
- (is "algebra ?r")
+ assumes "A \<in> sets M" shows "algebra (restricted_space A)"
using assms by unfold_locales auto
subsection {* Sigma Algebras *}
@@ -107,6 +110,13 @@
assumes countable_nat_UN [intro]:
"!!A. range A \<subseteq> sets M \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
+lemma sigma_algebra_cong:
+ fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme"
+ assumes *: "sigma_algebra M"
+ and cong: "space M = space M'" "sets M = sets M'"
+ shows "sigma_algebra M'"
+using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong .
+
lemma countable_UN_eq:
fixes A :: "'i::countable \<Rightarrow> 'a set"
shows "(range A \<subseteq> sets M \<longrightarrow> (\<Union>i. A i) \<in> sets M) \<longleftrightarrow>
@@ -320,15 +330,14 @@
lemma (in sigma_algebra) restricted_sigma_algebra:
assumes "S \<in> sets M"
- shows "sigma_algebra (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>)"
- (is "sigma_algebra ?r")
+ shows "sigma_algebra (restricted_space S)"
unfolding sigma_algebra_def sigma_algebra_axioms_def
proof safe
- show "algebra ?r" using restricted_algebra[OF assms] .
+ show "algebra (restricted_space S)" using restricted_algebra[OF assms] .
next
- fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets ?r"
+ fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets (restricted_space S)"
from restriction_in_sets[OF assms this[simplified]]
- show "(\<Union>i. A i) \<in> sets ?r" by simp
+ show "(\<Union>i. A i) \<in> sets (restricted_space S)" by simp
qed
section {* Measurable functions *}
@@ -560,6 +569,19 @@
(metis insert_absorb insert_subset le_SucE le_antisym not_leE)
qed
+lemma setsum_indicator_disjoint_family:
+ fixes f :: "'d \<Rightarrow> 'e::semiring_1"
+ assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P"
+ shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j"
+proof -
+ have "P \<inter> {i. x \<in> A i} = {j}"
+ using d `x \<in> A j` `j \<in> P` unfolding disjoint_family_on_def
+ by auto
+ thus ?thesis
+ unfolding indicator_def
+ by (simp add: if_distrib setsum_cases[OF `finite P`])
+qed
+
definition disjointed :: "(nat \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a set "
where "disjointed A n = A n - (\<Union>i\<in>{0..<n}. A i)"
@@ -671,6 +693,22 @@
shows "f \<in> measurable (vimage_algebra S f) M"
unfolding measurable_def using assms by force
+section {* Conditional space *}
+
+definition (in algebra)
+ "image_space X = \<lparr> space = X`space M, sets = (\<lambda>S. X`S) ` sets M \<rparr>"
+
+definition (in algebra)
+ "conditional_space X A = algebra.image_space (restricted_space A) X"
+
+lemma (in algebra) space_conditional_space:
+ assumes "A \<in> sets M" shows "space (conditional_space X A) = X`A"
+proof -
+ interpret r: algebra "restricted_space A" using assms by (rule restricted_algebra)
+ show ?thesis unfolding conditional_space_def r.image_space_def
+ by simp
+qed
+
subsection {* A Two-Element Series *}
definition binaryset :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> 'a set "