# HG changeset patch # User hoelzl # Date 1306424401 -7200 # Node ID d8f3fc934ff678a9355354ba019ec5438342305b # Parent 73e2d802ea417f6b74786e6dfbd8f9d6a0cbc2a5 add lemma indep_distribution_eq_measure diff -r 73e2d802ea41 -r d8f3fc934ff6 src/HOL/Probability/Finite_Product_Measure.thy --- a/src/HOL/Probability/Finite_Product_Measure.thy Thu May 26 14:12:06 2011 +0200 +++ b/src/HOL/Probability/Finite_Product_Measure.thy Thu May 26 17:40:01 2011 +0200 @@ -325,6 +325,10 @@ "\ \i. i \ I \ A i \ sets (M i) \ \ Pi\<^isub>E I A \ sets P" by (auto simp: sets_product_algebra) +lemma Int_stable_product_algebra_generator: + "(\i. i \ I \ Int_stable (M i)) \ Int_stable (product_algebra_generator I M)" + by (auto simp add: product_algebra_generator_def Int_stable_def PiE_Int Pi_iff) + section "Generating set generates also product algebra" lemma sigma_product_algebra_sigma_eq: @@ -478,6 +482,32 @@ using `A \ sets (M i)` by (auto intro!: product_algebraI) qed (insert `i \ I`, auto) +lemma (in sigma_algebra) measurable_restrict: + assumes I: "finite I" + assumes "\i. i \ I \ sets (N i) \ Pow (space (N i))" + assumes X: "\i. i \ I \ X i \ measurable M (N i)" + shows "(\x. \i\I. X i x) \ measurable M (Pi\<^isub>M I N)" + unfolding product_algebra_def +proof (simp, rule measurable_sigma) + show "sets (product_algebra_generator I N) \ Pow (space (product_algebra_generator I N))" + by (rule product_algebra_generator_sets_into_space) fact + show "(\x. \i\I. X i x) \ space M \ space (product_algebra_generator I N)" + using X by (auto simp: measurable_def) + fix E assume "E \ sets (product_algebra_generator I N)" + then obtain F where "E = Pi\<^isub>E I F" and F: "\i. i \ I \ F i \ sets (N i)" + by (auto simp: product_algebra_generator_def) + then have "(\x. \i\I. X i x) -` E \ space M = (\i\I. X i -` F i \ space M) \ space M" + by (auto simp: Pi_iff) + also have "\ \ sets M" + proof cases + assume "I = {}" then show ?thesis by simp + next + assume "I \ {}" with X F I show ?thesis + by (intro finite_INT measurable_sets Int) auto + qed + finally show "(\x. \i\I. X i x) -` E \ space M \ sets M" . +qed + locale product_sigma_finite = fixes M :: "'i \ ('a,'b) measure_space_scheme" assumes sigma_finite_measures: "\i. sigma_finite_measure (M i)" @@ -719,9 +749,8 @@ using A unfolding product_algebra_def by auto next show "Int_stable IJ.G" - by (simp add: PiE_Int Int_stable_def product_algebra_def - product_algebra_generator_def) - auto + by (rule Int_stable_product_algebra_generator) + (auto simp: Int_stable_def) show "range ?F \ sets IJ.G" using F by (simp add: image_subset_iff product_algebra_def product_algebra_generator_def) diff -r 73e2d802ea41 -r d8f3fc934ff6 src/HOL/Probability/Independent_Family.thy --- a/src/HOL/Probability/Independent_Family.thy Thu May 26 14:12:06 2011 +0200 +++ b/src/HOL/Probability/Independent_Family.thy Thu May 26 17:40:01 2011 +0200 @@ -120,19 +120,6 @@ and indep_setD_ev2: "B \ events" using indep unfolding indep_set_def indep_sets_def UNIV_bool by auto -lemma dynkin_systemI': - assumes 1: "\ A. A \ sets M \ A \ space M" - assumes empty: "{} \ sets M" - assumes Diff: "\ A. A \ sets M \ space M - A \ sets M" - assumes 2: "\ A. disjoint_family A \ range A \ sets M - \ (\i::nat. A i) \ sets M" - shows "dynkin_system M" -proof - - from Diff[OF empty] have "space M \ sets M" by auto - from 1 this Diff 2 show ?thesis - by (intro dynkin_systemI) auto -qed - lemma (in prob_space) indep_sets_dynkin: assumes indep: "indep_sets F I" shows "indep_sets (\i. sets (dynkin \ space = space M, sets = F i \)) I" @@ -714,7 +701,7 @@ and Int_stable: "\i. i \ I \ Int_stable (M' i)" and space: "\i. i \ I \ space (M' i) \ sets (M' i)" shows "indep_rv (\i. sigma (M' i)) X I \ - (\A\(\ i\I. sets (M' i)). prob (\j\I. X j -` A j \ space M) = (\j\I. distribution (X j) (A j)))" + (\A\(\ i\I. sets (M' i)). prob (\j\I. X j -` A j \ space M) = (\j\I. prob (X j -` A j \ space M)))" proof - from rv have X: "\i. i \ I \ X i \ space M \ space (M' i)" unfolding measurable_def by simp @@ -774,7 +761,138 @@ by simp qed then show ?thesis using `I \ {}` - by (simp add: rv distribution_def indep_rv_def) + by (simp add: rv indep_rv_def) +qed + +lemma (in prob_space) indep_rv_compose: + assumes "indep_rv M' X I" + assumes rv: + "\i. i \ I \ sigma_algebra (N i)" + "\i. i \ I \ Y i \ measurable (M' i) (N i)" + shows "indep_rv N (\i. Y i \ X i) I" + unfolding indep_rv_def +proof + from rv `indep_rv M' X I` + show "\i\I. random_variable (N i) (Y i \ X i)" + by (auto intro!: measurable_comp simp: indep_rv_def) + + have "indep_sets (\i. sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}) I" + using `indep_rv M' X I` by (simp add: indep_rv_def) + then show "indep_sets (\i. sigma_sets (space M) {(Y i \ X i) -` A \ space M |A. A \ sets (N i)}) I" + proof (rule indep_sets_mono_sets) + fix i assume "i \ I" + with `indep_rv M' X I` have X: "X i \ space M \ space (M' i)" + unfolding indep_rv_def measurable_def by auto + { fix A assume "A \ sets (N i)" + then have "\B. (Y i \ X i) -` A \ space M = X i -` B \ space M \ B \ sets (M' i)" + by (intro exI[of _ "Y i -` A \ space (M' i)"]) + (auto simp: vimage_compose intro!: measurable_sets rv `i \ I` funcset_mem[OF X]) } + then show "sigma_sets (space M) {(Y i \ X i) -` A \ space M |A. A \ sets (N i)} \ + sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}" + by (intro sigma_sets_subseteq) (auto simp: vimage_compose) + qed +qed + +lemma (in prob_space) indep_rvD: + assumes X: "indep_rv M' X I" + assumes I: "I \ {}" "finite I" "\i. i \ I \ A i \ sets (M' i)" + shows "prob (\i\I. X i -` A i \ space M) = (\i\I. prob (X i -` A i \ space M))" +proof (rule indep_setsD) + show "indep_sets (\i. sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}) I" + using X by (auto simp: indep_rv_def) + show "I \ I" "I \ {}" "finite I" using I by auto + show "\i\I. X i -` A i \ space M \ sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}" + using I by (auto intro: sigma_sets.Basic) +qed + +lemma (in prob_space) indep_distribution_eq_measure: + assumes I: "I \ {}" "finite I" + assumes rv: "\i. random_variable (M' i) (X i)" + shows "indep_rv M' X I \ + (\A\sets (\\<^isub>M i\I. (M' i \ measure := extreal\distribution (X i) \)). + distribution (\x. \i\I. X i x) A = + finite_measure.\' (\\<^isub>M i\I. (M' i \ measure := extreal\distribution (X i) \)) A)" + (is "_ \ (\X\_. distribution ?D X = finite_measure.\' (Pi\<^isub>M I ?M) X)") +proof - + interpret M': prob_space "?M i" for i + using rv by (rule distribution_prob_space) + interpret P: finite_product_prob_space ?M I + proof qed fact + + let ?D' = "(Pi\<^isub>M I ?M) \ measure := extreal \ distribution ?D \" + have "random_variable P.P ?D" + using `finite I` rv by (intro random_variable_restrict) auto + then interpret D: prob_space ?D' + by (rule distribution_prob_space) + + show ?thesis + proof (intro iffI ballI) + assume "indep_rv M' X I" + fix A assume "A \ sets P.P" + moreover + have "D.prob A = P.prob A" + proof (rule prob_space_unique_Int_stable) + show "prob_space ?D'" by default + show "prob_space (Pi\<^isub>M I ?M)" by default + show "Int_stable P.G" using M'.Int + by (intro Int_stable_product_algebra_generator) (simp add: Int_stable_def) + show "space P.G \ sets P.G" + using M'.top by (simp add: product_algebra_generator_def) + show "space ?D' = space P.G" "sets ?D' = sets (sigma P.G)" + by (simp_all add: product_algebra_def product_algebra_generator_def sets_sigma) + show "space P.P = space P.G" "sets P.P = sets (sigma P.G)" + by (simp_all add: product_algebra_def) + show "A \ sets (sigma P.G)" + using `A \ sets P.P` by (simp add: product_algebra_def) + + fix E assume E: "E \ sets P.G" + then have "E \ sets P.P" + by (simp add: sets_sigma sigma_sets.Basic product_algebra_def) + then have "D.prob E = distribution ?D E" + unfolding D.\'_def by simp + also + from E obtain F where "E = Pi\<^isub>E I F" and F: "\i. i \ I \ F i \ sets (M' i)" + by (auto simp: product_algebra_generator_def) + with `I \ {}` have "distribution ?D E = prob (\i\I. X i -` F i \ space M)" + using `I \ {}` by (auto intro!: arg_cong[where f=prob] simp: Pi_iff distribution_def) + also have "\ = (\i\I. prob (X i -` F i \ space M))" + using `indep_rv M' X I` I F by (rule indep_rvD) + also have "\ = P.prob E" + using F by (simp add: `E = Pi\<^isub>E I F` P.prob_times M'.\'_def distribution_def) + finally show "D.prob E = P.prob E" . + qed + ultimately show "distribution ?D A = P.prob A" + by (simp add: D.\'_def) + next + assume eq: "\A\sets P.P. distribution ?D A = P.prob A" + have [simp]: "\i. sigma (M' i) = M' i" + using rv by (intro sigma_algebra.sigma_eq) simp + have "indep_rv (\i. sigma (M' i)) X I" + proof (subst indep_rv_finite[OF I]) + fix i assume [simp]: "i \ I" + show "random_variable (sigma (M' i)) (X i)" + using rv[of i] by simp + show "Int_stable (M' i)" "space (M' i) \ sets (M' i)" + using M'.Int[of _ i] M'.top by (auto simp: Int_stable_def) + next + show "\A\\ i\I. sets (M' i). prob (\j\I. X j -` A j \ space M) = (\j\I. prob (X j -` A j \ space M))" + proof + fix A assume A: "A \ (\ i\I. sets (M' i))" + then have A_in_P: "(Pi\<^isub>E I A) \ sets P.P" + by (auto intro!: product_algebraI) + have "prob (\j\I. X j -` A j \ space M) = distribution ?D (Pi\<^isub>E I A)" + using `I \ {}`by (auto intro!: arg_cong[where f=prob] simp: Pi_iff distribution_def) + also have "\ = P.prob (Pi\<^isub>E I A)" using A_in_P eq by simp + also have "\ = (\i\I. M'.prob i (A i))" + using A by (intro P.prob_times) auto + also have "\ = (\i\I. prob (X i -` A i \ space M))" + using A by (auto intro!: setprod_cong simp: M'.\'_def Pi_iff distribution_def) + finally show "prob (\j\I. X j -` A j \ space M) = (\j\I. prob (X j -` A j \ space M))" . + qed + qed + then show "indep_rv M' X I" + by simp + qed qed end diff -r 73e2d802ea41 -r d8f3fc934ff6 src/HOL/Probability/Probability_Measure.thy --- a/src/HOL/Probability/Probability_Measure.thy Thu May 26 14:12:06 2011 +0200 +++ b/src/HOL/Probability/Probability_Measure.thy Thu May 26 17:40:01 2011 +0200 @@ -188,6 +188,37 @@ qed qed +lemma prob_space_unique_Int_stable: + fixes E :: "('a, 'b) algebra_scheme" and A :: "nat \ 'a set" + assumes E: "Int_stable E" "space E \ sets E" + and M: "prob_space M" "space M = space E" "sets M = sets (sigma E)" + and N: "prob_space N" "space N = space E" "sets N = sets (sigma E)" + and eq: "\X. X \ sets E \ finite_measure.\' M X = finite_measure.\' N X" + assumes "X \ sets (sigma E)" + shows "finite_measure.\' M X = finite_measure.\' N X" +proof - + interpret M!: prob_space M by fact + interpret N!: prob_space N by fact + have "measure M X = measure N X" + proof (rule measure_unique_Int_stable[OF `Int_stable E`]) + show "range (\i. space M) \ sets E" "incseq (\i. space M)" "(\i. space M) = space E" + using E M N by auto + show "\i. M.\ (space M) \ \" + using M.measure_space_1 by simp + show "measure_space \space = space E, sets = sets (sigma E), measure_space.measure = M.\\" + using E M N by (auto intro!: M.measure_space_cong) + show "measure_space \space = space E, sets = sets (sigma E), measure_space.measure = N.\\" + using E M N by (auto intro!: N.measure_space_cong) + { fix X assume "X \ sets E" + then have "X \ sets (sigma E)" + by (auto simp: sets_sigma sigma_sets.Basic) + with eq[OF `X \ sets E`] M N show "M.\ X = N.\ X" + by (simp add: M.finite_measure_eq N.finite_measure_eq) } + qed fact + with `X \ sets (sigma E)` M N show ?thesis + by (simp add: M.finite_measure_eq N.finite_measure_eq) +qed + lemma (in prob_space) distribution_prob_space: assumes X: "random_variable S X" shows "prob_space (S\measure := extreal \ distribution X\)" (is "prob_space ?S") @@ -359,6 +390,50 @@ shows "prob_space ((MX \\<^isub>M MY) \ measure := extreal \ joint_distribution X Y\)" using random_variable_pairI[OF assms] by (rule distribution_prob_space) + +locale finite_product_prob_space = + fixes M :: "'i \ ('a,'b) measure_space_scheme" + and I :: "'i set" + assumes prob_space: "\i. prob_space (M i)" and finite_index: "finite I" + +sublocale finite_product_prob_space \ M: prob_space "M i" for i + by (rule prob_space) + +sublocale finite_product_prob_space \ finite_product_sigma_finite M I + by default (rule finite_index) + +sublocale finite_product_prob_space \ prob_space "\\<^isub>M i\I. M i" + proof qed (simp add: measure_times M.measure_space_1 setprod_1) + +lemma (in finite_product_prob_space) prob_times: + assumes X: "\i. i \ I \ X i \ sets (M i)" + shows "prob (\\<^isub>E i\I. X i) = (\i\I. M.prob i (X i))" +proof - + have "extreal (\' (\\<^isub>E i\I. X i)) = \ (\\<^isub>E i\I. X i)" + using X by (intro finite_measure_eq[symmetric] in_P) auto + also have "\ = (\i\I. M.\ i (X i))" + using measure_times X by simp + also have "\ = extreal (\i\I. M.\' i (X i))" + using X by (simp add: M.finite_measure_eq setprod_extreal) + finally show ?thesis by simp +qed + +lemma (in prob_space) random_variable_restrict: + assumes I: "finite I" + assumes X: "\i. i \ I \ random_variable (N i) (X i)" + shows "random_variable (Pi\<^isub>M I N) (\x. \i\I. X i x)" +proof + { fix i assume "i \ I" + with X interpret N: sigma_algebra "N i" by simp + have "sets (N i) \ Pow (space (N i))" by (rule N.space_closed) } + note N_closed = this + then show "sigma_algebra (Pi\<^isub>M I N)" + by (simp add: product_algebra_def) + (intro sigma_algebra_sigma product_algebra_generator_sets_into_space) + show "(\x. \i\I. X i x) \ measurable M (Pi\<^isub>M I N)" + using X by (intro measurable_restrict[OF I N_closed]) auto +qed + section "Probability spaces on finite sets" locale finite_prob_space = prob_space + finite_measure_space diff -r 73e2d802ea41 -r d8f3fc934ff6 src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Thu May 26 14:12:06 2011 +0200 +++ b/src/HOL/Probability/Sigma_Algebra.thy Thu May 26 17:40:01 2011 +0200 @@ -1406,6 +1406,19 @@ shows "dynkin_system M" using assms by (auto simp: dynkin_system_def dynkin_system_axioms_def subset_class_def) +lemma dynkin_systemI': + assumes 1: "\ A. A \ sets M \ A \ space M" + assumes empty: "{} \ sets M" + assumes Diff: "\ A. A \ sets M \ space M - A \ sets M" + assumes 2: "\ A. disjoint_family A \ range A \ sets M + \ (\i::nat. A i) \ sets M" + shows "dynkin_system M" +proof - + from Diff[OF empty] have "space M \ sets M" by auto + from 1 this Diff 2 show ?thesis + by (intro dynkin_systemI) auto +qed + lemma dynkin_system_trivial: shows "dynkin_system \ space = A, sets = Pow A \" by (rule dynkin_systemI) auto