diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Probability_Measure.thy --- a/src/HOL/Probability/Probability_Measure.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Probability_Measure.thy Mon Apr 23 12:14:35 2012 +0200 @@ -6,110 +6,219 @@ header {*Probability measure*} theory Probability_Measure -imports Lebesgue_Measure + imports Lebesgue_Measure Radon_Nikodym begin +lemma funset_eq_UN_fun_upd_I: + assumes *: "\f. f \ F (insert a A) \ f(a := d) \ F A" + and **: "\f. f \ F (insert a A) \ f a \ G (f(a:=d))" + and ***: "\f x. \ f \ F A ; x \ G f \ \ f(a:=x) \ F (insert a A)" + shows "F (insert a A) = (\f\F A. fun_upd f a ` (G f))" +proof safe + fix f assume f: "f \ F (insert a A)" + show "f \ (\f\F A. fun_upd f a ` G f)" + proof (rule UN_I[of "f(a := d)"]) + show "f(a := d) \ F A" using *[OF f] . + show "f \ fun_upd (f(a:=d)) a ` G (f(a:=d))" + proof (rule image_eqI[of _ _ "f a"]) + show "f a \ G (f(a := d))" using **[OF f] . + qed simp + qed +next + fix f x assume "f \ F A" "x \ G f" + from ***[OF this] show "f(a := x) \ F (insert a A)" . +qed + +lemma extensional_funcset_insert_eq[simp]: + assumes "a \ A" + shows "extensional (insert a A) \ (insert a A \ B) = (\f \ extensional A \ (A \ B). (\b. f(a := b)) ` B)" + apply (rule funset_eq_UN_fun_upd_I) + using assms + by (auto intro!: inj_onI dest: inj_onD split: split_if_asm simp: extensional_def) + +lemma finite_extensional_funcset[simp, intro]: + assumes "finite A" "finite B" + shows "finite (extensional A \ (A \ B))" + using assms by induct auto + +lemma finite_PiE[simp, intro]: + assumes fin: "finite A" "\i. i \ A \ finite (B i)" + shows "finite (Pi\<^isub>E A B)" +proof - + have *: "(Pi\<^isub>E A B) \ extensional A \ (A \ (\i\A. B i))" by auto + show ?thesis + using fin by (intro finite_subset[OF *] finite_extensional_funcset) auto +qed + + +lemma countably_additiveI[case_names countably]: + assumes "\A. \ range A \ M ; disjoint_family A ; (\i. A i) \ M\ \ (\n. \ (A n)) = \ (\i. A i)" + shows "countably_additive M \" + using assms unfolding countably_additive_def by auto + +lemma convex_le_Inf_differential: + fixes f :: "real \ real" + assumes "convex_on I f" + assumes "x \ interior I" "y \ I" + shows "f y \ f x + Inf ((\t. (f x - f t) / (x - t)) ` ({x<..} \ I)) * (y - x)" + (is "_ \ _ + Inf (?F x) * (y - x)") +proof - + show ?thesis + proof (cases rule: linorder_cases) + assume "x < y" + moreover + have "open (interior I)" by auto + from openE[OF this `x \ interior I`] guess e . note e = this + moreover def t \ "min (x + e / 2) ((x + y) / 2)" + ultimately have "x < t" "t < y" "t \ ball x e" + by (auto simp: dist_real_def field_simps split: split_min) + with `x \ interior I` e interior_subset[of I] have "t \ I" "x \ I" by auto + + have "open (interior I)" by auto + from openE[OF this `x \ interior I`] guess e . + moreover def K \ "x - e / 2" + with `0 < e` have "K \ ball x e" "K < x" by (auto simp: dist_real_def) + ultimately have "K \ I" "K < x" "x \ I" + using interior_subset[of I] `x \ interior I` by auto + + have "Inf (?F x) \ (f x - f y) / (x - y)" + proof (rule Inf_lower2) + show "(f x - f t) / (x - t) \ ?F x" + using `t \ I` `x < t` by auto + show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" + using `convex_on I f` `x \ I` `y \ I` `x < t` `t < y` by (rule convex_on_diff) + next + fix y assume "y \ ?F x" + with order_trans[OF convex_on_diff[OF `convex_on I f` `K \ I` _ `K < x` _]] + show "(f K - f x) / (K - x) \ y" by auto + qed + then show ?thesis + using `x < y` by (simp add: field_simps) + next + assume "y < x" + moreover + have "open (interior I)" by auto + from openE[OF this `x \ interior I`] guess e . note e = this + moreover def t \ "x + e / 2" + ultimately have "x < t" "t \ ball x e" + by (auto simp: dist_real_def field_simps) + with `x \ interior I` e interior_subset[of I] have "t \ I" "x \ I" by auto + + have "(f x - f y) / (x - y) \ Inf (?F x)" + proof (rule Inf_greatest) + have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" + using `y < x` by (auto simp: field_simps) + also + fix z assume "z \ ?F x" + with order_trans[OF convex_on_diff[OF `convex_on I f` `y \ I` _ `y < x`]] + have "(f y - f x) / (y - x) \ z" by auto + finally show "(f x - f y) / (x - y) \ z" . + next + have "open (interior I)" by auto + from openE[OF this `x \ interior I`] guess e . note e = this + then have "x + e / 2 \ ball x e" by (auto simp: dist_real_def) + with e interior_subset[of I] have "x + e / 2 \ {x<..} \ I" by auto + then show "?F x \ {}" by blast + qed + then show ?thesis + using `y < x` by (simp add: field_simps) + qed simp +qed + +lemma distr_id[simp]: "distr N N (\x. x) = N" + by (rule measure_eqI) (auto simp: emeasure_distr) + locale prob_space = finite_measure + - assumes measure_space_1: "measure M (space M) = 1" + assumes emeasure_space_1: "emeasure M (space M) = 1" lemma prob_spaceI[Pure.intro!]: - assumes "measure_space M" - assumes *: "measure M (space M) = 1" + assumes *: "emeasure M (space M) = 1" shows "prob_space M" proof - interpret finite_measure M proof - show "measure_space M" by fact - show "measure M (space M) \ \" using * by simp + show "emeasure M (space M) \ \" using * by simp qed show "prob_space M" by default fact qed abbreviation (in prob_space) "events \ sets M" -abbreviation (in prob_space) "prob \ \'" -abbreviation (in prob_space) "random_variable M' X \ sigma_algebra M' \ X \ measurable M M'" +abbreviation (in prob_space) "prob \ measure M" +abbreviation (in prob_space) "random_variable M' X \ X \ measurable M M'" abbreviation (in prob_space) "expectation \ integral\<^isup>L M" -definition (in prob_space) - "distribution X A = \' (X -` A \ space M)" - -abbreviation (in prob_space) - "joint_distribution X Y \ distribution (\x. (X x, Y x))" - -lemma (in prob_space) prob_space_cong: - assumes "\A. A \ sets M \ measure N A = \ A" "space N = space M" "sets N = sets M" - shows "prob_space N" -proof - show "measure_space N" by (intro measure_space_cong assms) - show "measure N (space N) = 1" - using measure_space_1 assms by simp +lemma (in prob_space) prob_space_distr: + assumes f: "f \ measurable M M'" shows "prob_space (distr M M' f)" +proof (rule prob_spaceI) + have "f -` space M' \ space M = space M" using f by (auto dest: measurable_space) + with f show "emeasure (distr M M' f) (space (distr M M' f)) = 1" + by (auto simp: emeasure_distr emeasure_space_1) qed -lemma (in prob_space) distribution_cong: - assumes "\x. x \ space M \ X x = Y x" - shows "distribution X = distribution Y" - unfolding distribution_def fun_eq_iff - using assms by (auto intro!: arg_cong[where f="\'"]) - -lemma (in prob_space) joint_distribution_cong: - assumes "\x. x \ space M \ X x = X' x" - assumes "\x. x \ space M \ Y x = Y' x" - shows "joint_distribution X Y = joint_distribution X' Y'" - unfolding distribution_def fun_eq_iff - using assms by (auto intro!: arg_cong[where f="\'"]) - -lemma (in prob_space) distribution_id[simp]: - "N \ events \ distribution (\x. x) N = prob N" - by (auto simp: distribution_def intro!: arg_cong[where f=prob]) - lemma (in prob_space) prob_space: "prob (space M) = 1" - using measure_space_1 unfolding \'_def by (simp add: one_ereal_def) + using emeasure_space_1 unfolding measure_def by (simp add: one_ereal_def) lemma (in prob_space) prob_le_1[simp, intro]: "prob A \ 1" using bounded_measure[of A] by (simp add: prob_space) -lemma (in prob_space) distribution_positive[simp, intro]: - "0 \ distribution X A" unfolding distribution_def by auto - -lemma (in prob_space) not_zero_less_distribution[simp]: - "(\ 0 < distribution X A) \ distribution X A = 0" - using distribution_positive[of X A] by arith - -lemma (in prob_space) joint_distribution_remove[simp]: - "joint_distribution X X {(x, x)} = distribution X {x}" - unfolding distribution_def by (auto intro!: arg_cong[where f=\']) +lemma (in prob_space) not_empty: "space M \ {}" + using prob_space by auto -lemma (in prob_space) distribution_unit[simp]: "distribution (\x. ()) {()} = 1" - unfolding distribution_def using prob_space by auto - -lemma (in prob_space) joint_distribution_unit[simp]: "distribution (\x. (X x, ())) {(a, ())} = distribution X {a}" - unfolding distribution_def by (auto intro!: arg_cong[where f=\']) - -lemma (in prob_space) not_empty: "space M \ {}" - using prob_space empty_measure' by auto - -lemma (in prob_space) measure_le_1: "X \ sets M \ \ X \ 1" - unfolding measure_space_1[symmetric] - using sets_into_space - by (intro measure_mono) auto +lemma (in prob_space) measure_le_1: "emeasure M X \ 1" + using emeasure_space[of M X] by (simp add: emeasure_space_1) lemma (in prob_space) AE_I_eq_1: - assumes "\ {x\space M. P x} = 1" "{x\space M. P x} \ sets M" - shows "AE x. P x" + assumes "emeasure M {x\space M. P x} = 1" "{x\space M. P x} \ sets M" + shows "AE x in M. P x" proof (rule AE_I) - show "\ (space M - {x \ space M. P x}) = 0" - using assms measure_space_1 by (simp add: measure_compl) + show "emeasure M (space M - {x \ space M. P x}) = 0" + using assms emeasure_space_1 by (simp add: emeasure_compl) qed (insert assms, auto) -lemma (in prob_space) distribution_1: - "distribution X A \ 1" - unfolding distribution_def by simp - lemma (in prob_space) prob_compl: assumes A: "A \ events" shows "prob (space M - A) = 1 - prob A" using finite_measure_compl[OF A] by (simp add: prob_space) +lemma (in prob_space) AE_in_set_eq_1: + assumes "A \ events" shows "(AE x in M. x \ A) \ prob A = 1" +proof + assume ae: "AE x in M. x \ A" + have "{x \ space M. x \ A} = A" "{x \ space M. x \ A} = space M - A" + using `A \ events`[THEN sets_into_space] by auto + with AE_E2[OF ae] `A \ events` have "1 - emeasure M A = 0" + by (simp add: emeasure_compl emeasure_space_1) + then show "prob A = 1" + using `A \ events` by (simp add: emeasure_eq_measure one_ereal_def) +next + assume prob: "prob A = 1" + show "AE x in M. x \ A" + proof (rule AE_I) + show "{x \ space M. x \ A} \ space M - A" by auto + show "emeasure M (space M - A) = 0" + using `A \ events` prob + by (simp add: prob_compl emeasure_space_1 emeasure_eq_measure one_ereal_def) + show "space M - A \ events" + using `A \ events` by auto + qed +qed + +lemma (in prob_space) AE_False: "(AE x in M. False) \ False" +proof + assume "AE x in M. False" + then have "AE x in M. x \ {}" by simp + then show False + by (subst (asm) AE_in_set_eq_1) auto +qed simp + +lemma (in prob_space) AE_prob_1: + assumes "prob A = 1" shows "AE x in M. x \ A" +proof - + from `prob A = 1` have "A \ events" + by (metis measure_notin_sets zero_neq_one) + with AE_in_set_eq_1 assms show ?thesis by simp +qed + lemma (in prob_space) prob_space_increasing: "increasing M prob" by (auto intro!: finite_measure_mono simp: increasing_def) @@ -164,9 +273,8 @@ shows "prob (\ i :: nat. c i) = 0" proof (rule antisym) show "prob (\ i :: nat. c i) \ 0" - using finite_measure_countably_subadditive[OF assms(1)] - by (simp add: assms(2) suminf_zero summable_zero) -qed simp + using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2)) +qed (simp add: measure_nonneg) lemma (in prob_space) prob_equiprobable_finite_unions: assumes "s \ events" @@ -178,7 +286,7 @@ from someI_ex[OF this] assms have prob_some: "\ x. x \ s \ prob {x} = prob {SOME y. y \ s}" by blast have "prob s = (\ x \ s. prob {x})" - using finite_measure_finite_singleton[OF s_finite] by simp + using finite_measure_eq_setsum_singleton[OF s_finite] by simp also have "\ = (\ x \ s. prob {SOME y. y \ s})" using prob_some by auto also have "\ = real (card s) * prob {(SOME x. x \ s)}" using setsum_constant assms by (simp add: real_eq_of_nat) @@ -199,96 +307,20 @@ also have "\ = (\ x \ s. prob (e \ f x))" proof (rule finite_measure_finite_Union) show "finite s" by fact - show "\i. i \ s \ e \ f i \ events" by fact + show "(\i. e \ f i)`s \ events" using assms(2) by auto show "disjoint_family_on (\i. e \ f i) s" using disjoint by (auto simp: disjoint_family_on_def) qed finally show ?thesis . qed -lemma (in prob_space) prob_space_vimage: - assumes S: "sigma_algebra S" - assumes T: "T \ measure_preserving M S" - shows "prob_space S" -proof - interpret S: measure_space S - using S and T by (rule measure_space_vimage) - show "measure_space S" .. - - from T[THEN measure_preservingD2] - have "T -` space S \ space M = space M" - by (auto simp: measurable_def) - with T[THEN measure_preservingD, of "space S", symmetric] - show "measure S (space S) = 1" - using measure_space_1 by simp -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 := ereal \ distribution X\)" (is "prob_space ?S") -proof (rule prob_space_vimage) - show "X \ measure_preserving M ?S" - using X - unfolding measure_preserving_def distribution_def [abs_def] - by (auto simp: finite_measure_eq measurable_sets) - show "sigma_algebra ?S" using X by simp -qed - -lemma (in prob_space) AE_distribution: - assumes X: "random_variable MX X" and "AE x in MX\measure := ereal \ distribution X\. Q x" - shows "AE x. Q (X x)" -proof - - interpret X: prob_space "MX\measure := ereal \ distribution X\" using X by (rule distribution_prob_space) - obtain N where N: "N \ sets MX" "distribution X N = 0" "{x\space MX. \ Q x} \ N" - using assms unfolding X.almost_everywhere_def by auto - from X[unfolded measurable_def] N show "AE x. Q (X x)" - by (intro AE_I'[where N="X -` N \ space M"]) - (auto simp: finite_measure_eq distribution_def measurable_sets) -qed - -lemma (in prob_space) distribution_eq_integral: - "random_variable S X \ A \ sets S \ distribution X A = expectation (indicator (X -` A \ space M))" - using finite_measure_eq[of "X -` A \ space M"] - by (auto simp: measurable_sets distribution_def) - lemma (in prob_space) expectation_less: assumes [simp]: "integrable M X" assumes gt: "\x\space M. X x < b" shows "expectation X < b" proof - have "expectation X < expectation (\x. b)" - using gt measure_space_1 + using gt emeasure_space_1 by (intro integral_less_AE_space) auto then show ?thesis using prob_space by simp qed @@ -299,80 +331,11 @@ shows "a < expectation X" proof - have "expectation (\x. a) < expectation X" - using gt measure_space_1 + using gt emeasure_space_1 by (intro integral_less_AE_space) auto then show ?thesis using prob_space by simp qed -lemma convex_le_Inf_differential: - fixes f :: "real \ real" - assumes "convex_on I f" - assumes "x \ interior I" "y \ I" - shows "f y \ f x + Inf ((\t. (f x - f t) / (x - t)) ` ({x<..} \ I)) * (y - x)" - (is "_ \ _ + Inf (?F x) * (y - x)") -proof - - show ?thesis - proof (cases rule: linorder_cases) - assume "x < y" - moreover - have "open (interior I)" by auto - from openE[OF this `x \ interior I`] guess e . note e = this - moreover def t \ "min (x + e / 2) ((x + y) / 2)" - ultimately have "x < t" "t < y" "t \ ball x e" - by (auto simp: mem_ball dist_real_def field_simps split: split_min) - with `x \ interior I` e interior_subset[of I] have "t \ I" "x \ I" by auto - - have "open (interior I)" by auto - from openE[OF this `x \ interior I`] guess e . - moreover def K \ "x - e / 2" - with `0 < e` have "K \ ball x e" "K < x" by (auto simp: mem_ball dist_real_def) - ultimately have "K \ I" "K < x" "x \ I" - using interior_subset[of I] `x \ interior I` by auto - - have "Inf (?F x) \ (f x - f y) / (x - y)" - proof (rule Inf_lower2) - show "(f x - f t) / (x - t) \ ?F x" - using `t \ I` `x < t` by auto - show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" - using `convex_on I f` `x \ I` `y \ I` `x < t` `t < y` by (rule convex_on_diff) - next - fix y assume "y \ ?F x" - with order_trans[OF convex_on_diff[OF `convex_on I f` `K \ I` _ `K < x` _]] - show "(f K - f x) / (K - x) \ y" by auto - qed - then show ?thesis - using `x < y` by (simp add: field_simps) - next - assume "y < x" - moreover - have "open (interior I)" by auto - from openE[OF this `x \ interior I`] guess e . note e = this - moreover def t \ "x + e / 2" - ultimately have "x < t" "t \ ball x e" - by (auto simp: mem_ball dist_real_def field_simps) - with `x \ interior I` e interior_subset[of I] have "t \ I" "x \ I" by auto - - have "(f x - f y) / (x - y) \ Inf (?F x)" - proof (rule Inf_greatest) - have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" - using `y < x` by (auto simp: field_simps) - also - fix z assume "z \ ?F x" - with order_trans[OF convex_on_diff[OF `convex_on I f` `y \ I` _ `y < x`]] - have "(f y - f x) / (y - x) \ z" by auto - finally show "(f x - f y) / (x - y) \ z" . - next - have "open (interior I)" by auto - from openE[OF this `x \ interior I`] guess e . note e = this - then have "x + e / 2 \ ball x e" by (auto simp: mem_ball dist_real_def) - with e interior_subset[of I] have "x + e / 2 \ {x<..} \ I" by auto - then show "?F x \ {}" by blast - qed - then show ?thesis - using `y < x` by (simp add: field_simps) - qed simp -qed - lemma (in prob_space) jensens_inequality: fixes a b :: real assumes X: "integrable M X" "X ` space M \ I" @@ -410,8 +373,7 @@ fix k assume "k \ (\x. q x + ?F x * (expectation X - x)) ` I" then guess x .. note x = this have "q x + ?F x * (expectation X - x) = expectation (\w. q x + ?F x * (X w - x))" - using prob_space - by (simp add: integral_add integral_cmult integral_diff lebesgue_integral_const X) + using prob_space by (simp add: X) also have "\ \ expectation (\w. q (X w))" using `x \ I` `open I` X(2) by (intro integral_mono integral_add integral_cmult integral_diff @@ -422,31 +384,6 @@ finally show "q (expectation X) \ expectation (\x. q (X x))" . qed -lemma (in prob_space) distribution_eq_translated_integral: - assumes "random_variable S X" "A \ sets S" - shows "distribution X A = integral\<^isup>P (S\measure := ereal \ distribution X\) (indicator A)" -proof - - interpret S: prob_space "S\measure := ereal \ distribution X\" - using assms(1) by (rule distribution_prob_space) - show ?thesis - using S.positive_integral_indicator(1)[of A] assms by simp -qed - -lemma (in prob_space) finite_expectation1: - assumes f: "finite (X`space M)" and rv: "random_variable borel X" - shows "expectation X = (\r \ X ` space M. r * prob (X -` {r} \ space M))" (is "_ = ?r") -proof (subst integral_on_finite) - show "X \ borel_measurable M" "finite (X`space M)" using assms by auto - show "(\ r \ X ` space M. r * real (\ (X -` {r} \ space M))) = ?r" - "\x. \ (X -` {x} \ space M) \ \" - using finite_measure_eq[OF borel_measurable_vimage, of X] rv by auto -qed - -lemma (in prob_space) finite_expectation: - assumes "finite (X`space M)" "random_variable borel X" - shows "expectation X = (\ r \ X ` (space M). r * distribution X {r})" - using assms unfolding distribution_def using finite_expectation1 by auto - lemma (in prob_space) prob_x_eq_1_imp_prob_y_eq_0: assumes "{x} \ events" assumes "prob {x} = 1" @@ -455,119 +392,25 @@ shows "prob {y} = 0" using prob_one_inter[of "{y}" "{x}"] assms by auto -lemma (in prob_space) distribution_empty[simp]: "distribution X {} = 0" - unfolding distribution_def by simp - -lemma (in prob_space) distribution_space[simp]: "distribution X (X ` space M) = 1" -proof - - have "X -` X ` space M \ space M = space M" by auto - thus ?thesis unfolding distribution_def by (simp add: prob_space) -qed - -lemma (in prob_space) distribution_one: - assumes "random_variable M' X" and "A \ sets M'" - shows "distribution X A \ 1" -proof - - have "distribution X A \ \' (space M)" unfolding distribution_def - using assms[unfolded measurable_def] by (auto intro!: finite_measure_mono) - thus ?thesis by (simp add: prob_space) -qed - -lemma (in prob_space) distribution_x_eq_1_imp_distribution_y_eq_0: - assumes X: "random_variable \space = X ` (space M), sets = Pow (X ` (space M))\ X" - (is "random_variable ?S X") - assumes "distribution X {x} = 1" - assumes "y \ x" - shows "distribution X {y} = 0" -proof cases - { fix x have "X -` {x} \ space M \ sets M" - proof cases - assume "x \ X`space M" with X show ?thesis - by (auto simp: measurable_def image_iff) - next - assume "x \ X`space M" then have "X -` {x} \ space M = {}" by auto - then show ?thesis by auto - qed } note single = this - have "X -` {x} \ space M - X -` {y} \ space M = X -` {x} \ space M" - "X -` {y} \ space M \ (X -` {x} \ space M) = {}" - using `y \ x` by auto - with finite_measure_inter_full_set[OF single single, of x y] assms(2) - show ?thesis by (auto simp: distribution_def prob_space) -next - assume "{y} \ sets ?S" - then have "X -` {y} \ space M = {}" by auto - thus "distribution X {y} = 0" unfolding distribution_def by auto -qed - lemma (in prob_space) joint_distribution_Times_le_fst: - assumes X: "random_variable MX X" and Y: "random_variable MY Y" - and A: "A \ sets MX" and B: "B \ sets MY" - shows "joint_distribution X Y (A \ B) \ distribution X A" - unfolding distribution_def -proof (intro finite_measure_mono) - show "(\x. (X x, Y x)) -` (A \ B) \ space M \ X -` A \ space M" by force - show "X -` A \ space M \ events" - using X A unfolding measurable_def by simp - have *: "(\x. (X x, Y x)) -` (A \ B) \ space M = - (X -` A \ space M) \ (Y -` B \ space M)" by auto -qed - -lemma (in prob_space) joint_distribution_commute: - "joint_distribution X Y x = joint_distribution Y X ((\(x,y). (y,x))`x)" - unfolding distribution_def by (auto intro!: arg_cong[where f=\']) + "random_variable MX X \ random_variable MY Y \ A \ sets MX \ B \ sets MY + \ emeasure (distr M (MX \\<^isub>M MY) (\x. (X x, Y x))) (A \ B) \ emeasure (distr M MX X) A" + by (auto simp: emeasure_distr measurable_pair_iff comp_def intro!: emeasure_mono measurable_sets) lemma (in prob_space) joint_distribution_Times_le_snd: - assumes X: "random_variable MX X" and Y: "random_variable MY Y" - and A: "A \ sets MX" and B: "B \ sets MY" - shows "joint_distribution X Y (A \ B) \ distribution Y B" - using assms - by (subst joint_distribution_commute) - (simp add: swap_product joint_distribution_Times_le_fst) - -lemma (in prob_space) random_variable_pairI: - assumes "random_variable MX X" - assumes "random_variable MY Y" - shows "random_variable (MX \\<^isub>M MY) (\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 (MX \\<^isub>M MY)" by default - have sa: "sigma_algebra M" by default - show "(\x. (X x, Y x)) \ measurable M (MX \\<^isub>M MY)" - unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def) -qed - -lemma (in prob_space) joint_distribution_commute_singleton: - "joint_distribution X Y {(x, y)} = joint_distribution Y X {(y, x)}" - unfolding distribution_def by (auto intro!: arg_cong[where f=\']) - -lemma (in prob_space) joint_distribution_assoc_singleton: - "joint_distribution X (\x. (Y x, Z x)) {(x, y, z)} = - joint_distribution (\x. (X x, Y x)) Z {((x, y), z)}" - unfolding distribution_def by (auto intro!: arg_cong[where f=\']) + "random_variable MX X \ random_variable MY Y \ A \ sets MX \ B \ sets MY + \ emeasure (distr M (MX \\<^isub>M MY) (\x. (X x, Y x))) (A \ B) \ emeasure (distr M MY Y) B" + by (auto simp: emeasure_distr measurable_pair_iff comp_def intro!: emeasure_mono measurable_sets) locale pair_prob_space = pair_sigma_finite M1 M2 + M1: prob_space M1 + M2: prob_space M2 for M1 M2 -sublocale pair_prob_space \ P: prob_space P +sublocale pair_prob_space \ P: prob_space "M1 \\<^isub>M M2" proof - show "measure_space P" .. - show "measure P (space P) = 1" - by (simp add: pair_measure_times M1.measure_space_1 M2.measure_space_1 space_pair_measure) + show "emeasure (M1 \\<^isub>M M2) (space (M1 \\<^isub>M M2)) = 1" + by (simp add: emeasure_pair_measure_Times M1.emeasure_space_1 M2.emeasure_space_1 space_pair_measure) qed -lemma countably_additiveI[case_names countably]: - assumes "\A. \ range A \ sets M ; disjoint_family A ; (\i. A i) \ sets M\ \ - (\n. \ (A n)) = \ (\i. A i)" - shows "countably_additive M \" - using assms unfolding countably_additive_def by auto - -lemma (in prob_space) joint_distribution_prob_space: - assumes "random_variable MX X" "random_variable MY Y" - shows "prob_space ((MX \\<^isub>M MY) \ measure := ereal \ joint_distribution X Y\)" - using random_variable_pairI[OF assms] by (rule distribution_prob_space) - -locale product_prob_space = product_sigma_finite M for M :: "'i \ ('a, 'b) measure_space_scheme" + +locale product_prob_space = product_sigma_finite M for M :: "'i \ 'a measure" + fixes I :: "'i set" assumes prob_space: "\i. prob_space (M i)" @@ -578,648 +421,401 @@ sublocale finite_product_prob_space \ prob_space "\\<^isub>M i\I. M i" proof - show "measure_space P" .. - show "measure P (space P) = 1" - by (simp add: measure_times M.measure_space_1 setprod_1) + show "emeasure (\\<^isub>M i\I. M i) (space (\\<^isub>M i\I. M i)) = 1" + by (simp add: measure_times M.emeasure_space_1 setprod_1 space_PiM) qed 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 "ereal (\' (\\<^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))" + have "ereal (measure (\\<^isub>M i\I. M i) (\\<^isub>E i\I. X i)) = emeasure (\\<^isub>M i\I. M i) (\\<^isub>E i\I. X i)" + using X by (simp add: emeasure_eq_measure) + also have "\ = (\i\I. emeasure (M i) (X i))" using measure_times X by simp - also have "\ = ereal (\i\I. M.\' i (X i))" - using X by (simp add: M.finite_measure_eq setprod_ereal) - 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 - -abbreviation (in prob_space) "finite_random_variable M' X \ finite_sigma_algebra M' \ X \ measurable M M'" - -lemma (in prob_space) finite_random_variableD: - assumes "finite_random_variable M' X" shows "random_variable M' X" -proof - - interpret M': finite_sigma_algebra M' using assms by simp - show "random_variable M' X" using assms by simp default -qed - -lemma (in prob_space) distribution_finite_prob_space: - assumes "finite_random_variable MX X" - shows "finite_prob_space (MX\measure := ereal \ distribution X\)" -proof - - interpret X: prob_space "MX\measure := ereal \ distribution X\" - using assms[THEN finite_random_variableD] by (rule distribution_prob_space) - interpret MX: finite_sigma_algebra MX - using assms by auto - show ?thesis by default (simp_all add: MX.finite_space) -qed - -lemma (in prob_space) simple_function_imp_finite_random_variable[simp, intro]: - assumes "simple_function M X" - shows "finite_random_variable \ space = X`space M, sets = Pow (X`space M), \ = x \ 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 ?X by (rule sigma_algebra_Pow) - show "finite_sigma_algebra ?X" - by default auto - show "X \ measurable M ?X" - proof (unfold measurable_def, clarsimp) - fix A assume A: "A \ X`space M" - then have "finite A" by (rule finite_subset) simp - then have "X -` (\a\A. {a}) \ space M \ events" - unfolding vimage_UN UN_extend_simps - apply (rule finite_UN) - using A assms unfolding simple_function_def by auto - then show "X -` A \ space M \ events" by simp - qed -qed - -lemma (in prob_space) simple_function_imp_random_variable[simp, intro]: - assumes "simple_function M X" - shows "random_variable \ space = X`space M, sets = Pow (X`space M), \ = ext \ 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 M X \ (\x\X`space M. distribution X {x}) = 1" - unfolding distribution_def prob_space[symmetric] - by (subst finite_measure_finite_Union[symmetric]) - (auto simp add: disjoint_family_on_def simple_function_def - intro!: arg_cong[where f=prob]) - -lemma (in prob_space) finite_random_variable_pairI: - assumes "finite_random_variable MX X" - assumes "finite_random_variable MY Y" - shows "finite_random_variable (MX \\<^isub>M MY) (\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 (MX \\<^isub>M MY)" .. - have sa: "sigma_algebra M" by default - show "(\x. (X x, Y x)) \ measurable M (MX \\<^isub>M MY)" - unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def) -qed - -lemma (in prob_space) finite_random_variable_imp_sets: - "finite_random_variable MX X \ x \ space MX \ {x} \ sets MX" - unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp - -lemma (in prob_space) finite_random_variable_measurable: - assumes X: "finite_random_variable MX X" shows "X -` A \ space M \ events" -proof - - interpret X: finite_sigma_algebra MX using X by simp - from X have vimage: "\A. A \ space MX \ X -` A \ space M \ events" and - "X \ space M \ space MX" - by (auto simp: measurable_def) - then have *: "X -` A \ space M = X -` (A \ space MX) \ space M" - by auto - show "X -` A \ space M \ events" - unfolding * by (intro vimage) auto -qed - -lemma (in prob_space) joint_distribution_finite_Times_le_fst: - assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" - shows "joint_distribution X Y (A \ B) \ distribution X A" - unfolding distribution_def -proof (intro finite_measure_mono) - show "(\x. (X x, Y x)) -` (A \ B) \ space M \ X -` A \ space M" by force - show "X -` A \ space M \ events" - using finite_random_variable_measurable[OF X] . - have *: "(\x. (X x, Y x)) -` (A \ B) \ space M = - (X -` A \ space M) \ (Y -` B \ space M)" by auto -qed - -lemma (in prob_space) joint_distribution_finite_Times_le_snd: - assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" - shows "joint_distribution X Y (A \ B) \ distribution Y B" - using assms - by (subst joint_distribution_commute) - (simp add: swap_product joint_distribution_finite_Times_le_fst) - -lemma (in prob_space) finite_distribution_order: - fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme" - assumes "finite_random_variable MX X" "finite_random_variable MY Y" - shows "r \ joint_distribution X Y {(x, y)} \ r \ distribution X {x}" - and "r \ joint_distribution X Y {(x, y)} \ r \ distribution Y {y}" - and "r < joint_distribution X Y {(x, y)} \ r < distribution X {x}" - and "r < joint_distribution X Y {(x, y)} \ r < distribution Y {y}" - and "distribution X {x} = 0 \ joint_distribution X Y {(x, y)} = 0" - and "distribution Y {y} = 0 \ joint_distribution X Y {(x, y)} = 0" - using joint_distribution_finite_Times_le_snd[OF assms, of "{x}" "{y}"] - using joint_distribution_finite_Times_le_fst[OF assms, of "{x}" "{y}"] - by (auto intro: antisym) - -lemma (in prob_space) setsum_joint_distribution: - assumes X: "finite_random_variable MX X" - assumes Y: "random_variable MY Y" "B \ sets MY" - shows "(\a\space MX. joint_distribution X Y ({a} \ B)) = distribution Y B" - unfolding distribution_def -proof (subst finite_measure_finite_Union[symmetric]) - interpret MX: finite_sigma_algebra MX using X by auto - show "finite (space MX)" using MX.finite_space . - let ?d = "\i. (\x. (X x, Y x)) -` ({i} \ B) \ space M" - { fix i assume "i \ space MX" - moreover have "?d i = (X -` {i} \ space M) \ (Y -` B \ space M)" by auto - ultimately show "?d i \ events" - using measurable_sets[of X M MX] measurable_sets[of Y M MY B] X Y - using MX.sets_eq_Pow by auto } - show "disjoint_family_on ?d (space MX)" by (auto simp: disjoint_family_on_def) - show "\' (\i\space MX. ?d i) = \' (Y -` B \ space M)" - using X[unfolded measurable_def] by (auto intro!: arg_cong[where f=\']) -qed - -lemma (in prob_space) setsum_joint_distribution_singleton: - assumes X: "finite_random_variable MX X" - assumes Y: "finite_random_variable MY Y" "b \ space MY" - shows "(\a\space MX. joint_distribution X Y {(a, b)}) = distribution Y {b}" - using setsum_joint_distribution[OF X - finite_random_variableD[OF Y(1)] - finite_random_variable_imp_sets[OF Y]] by simp - -lemma (in prob_space) setsum_distribution: - assumes X: "finite_random_variable MX X" shows "(\a\space MX. distribution X {a}) = 1" - using setsum_joint_distribution[OF assms, of "\ space = UNIV, sets = Pow UNIV \" "\x. ()" "{()}"] - using sigma_algebra_Pow[of "UNIV::unit set" "()"] by simp - -locale pair_finite_prob_space = pair_prob_space M1 M2 + pair_finite_space M1 M2 + M1: finite_prob_space M1 + M2: finite_prob_space M2 for M1 M2 - -sublocale pair_finite_prob_space \ finite_prob_space P by default - -lemma funset_eq_UN_fun_upd_I: - assumes *: "\f. f \ F (insert a A) \ f(a := d) \ F A" - and **: "\f. f \ F (insert a A) \ f a \ G (f(a:=d))" - and ***: "\f x. \ f \ F A ; x \ G f \ \ f(a:=x) \ F (insert a A)" - shows "F (insert a A) = (\f\F A. fun_upd f a ` (G f))" -proof safe - fix f assume f: "f \ F (insert a A)" - show "f \ (\f\F A. fun_upd f a ` G f)" - proof (rule UN_I[of "f(a := d)"]) - show "f(a := d) \ F A" using *[OF f] . - show "f \ fun_upd (f(a:=d)) a ` G (f(a:=d))" - proof (rule image_eqI[of _ _ "f a"]) - show "f a \ G (f(a := d))" using **[OF f] . - qed simp - qed -next - fix f x assume "f \ F A" "x \ G f" - from ***[OF this] show "f(a := x) \ F (insert a A)" . -qed - -lemma extensional_funcset_insert_eq[simp]: - assumes "a \ A" - shows "extensional (insert a A) \ (insert a A \ B) = (\f \ extensional A \ (A \ B). (\b. f(a := b)) ` B)" - apply (rule funset_eq_UN_fun_upd_I) - using assms - by (auto intro!: inj_onI dest: inj_onD split: split_if_asm simp: extensional_def) - -lemma finite_extensional_funcset[simp, intro]: - assumes "finite A" "finite B" - shows "finite (extensional A \ (A \ B))" - using assms by induct (auto simp: extensional_funcset_insert_eq) - -lemma finite_PiE[simp, intro]: - assumes fin: "finite A" "\i. i \ A \ finite (B i)" - shows "finite (Pi\<^isub>E A B)" -proof - - have *: "(Pi\<^isub>E A B) \ extensional A \ (A \ (\i\A. B i))" by auto - show ?thesis - using fin by (intro finite_subset[OF *] finite_extensional_funcset) auto -qed - -locale finite_product_finite_prob_space = finite_product_prob_space M I for M I + - assumes finite_space: "\i. finite_prob_space (M i)" - -sublocale finite_product_finite_prob_space \ M!: finite_prob_space "M i" using finite_space . - -lemma (in finite_product_finite_prob_space) singleton_eq_product: - assumes x: "x \ space P" shows "{x} = (\\<^isub>E i\I. {x i})" -proof (safe intro!: ext[of _ x]) - fix y i assume *: "y \ (\ i\I. {x i})" "y \ extensional I" - with x show "y i = x i" - by (cases "i \ I") (auto simp: extensional_def) -qed (insert x, auto) - -sublocale finite_product_finite_prob_space \ finite_prob_space "Pi\<^isub>M I M" -proof - show "finite (space P)" - using finite_index M.finite_space by auto - - { fix x assume "x \ space P" - with this[THEN singleton_eq_product] - have "{x} \ sets P" - by (auto intro!: in_P) } - note x_in_P = this - - have "Pow (space P) \ sets P" - proof - fix X assume "X \ Pow (space P)" - moreover then have "finite X" - using `finite (space P)` by (blast intro: finite_subset) - ultimately have "(\x\X. {x}) \ sets P" - by (intro finite_UN x_in_P) auto - then show "X \ sets P" by simp - qed - with space_closed show [simp]: "sets P = Pow (space P)" .. -qed - -lemma (in finite_product_finite_prob_space) measure_finite_times: - "(\i. i \ I \ X i \ space (M i)) \ \ (\\<^isub>E i\I. X i) = (\i\I. M.\ i (X i))" - by (rule measure_times) simp - -lemma (in finite_product_finite_prob_space) measure_singleton_times: - assumes x: "x \ space P" shows "\ {x} = (\i\I. M.\ i {x i})" - unfolding singleton_eq_product[OF x] using x - by (intro measure_finite_times) auto - -lemma (in finite_product_finite_prob_space) prob_finite_times: - assumes X: "\i. i \ I \ X i \ space (M i)" - shows "prob (\\<^isub>E i\I. X i) = (\i\I. M.prob i (X i))" -proof - - have "ereal (\' (\\<^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_finite_times X by simp - also have "\ = ereal (\i\I. M.\' i (X i))" - using X by (simp add: M.finite_measure_eq setprod_ereal) + also have "\ = ereal (\i\I. measure (M i) (X i))" + using X by (simp add: M.emeasure_eq_measure setprod_ereal) finally show ?thesis by simp qed -lemma (in finite_product_finite_prob_space) prob_singleton_times: - assumes x: "x \ space P" - shows "prob {x} = (\i\I. M.prob i {x i})" - unfolding singleton_eq_product[OF x] using x - by (intro prob_finite_times) auto - -lemma (in finite_product_finite_prob_space) prob_finite_product: - "A \ space P \ prob A = (\x\A. \i\I. M.prob i {x i})" - by (auto simp add: finite_measure_singleton prob_singleton_times - simp del: space_product_algebra - intro!: setsum_cong prob_singleton_times) +section {* Distributions *} -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 ((MX \\<^isub>M MY)\ measure := ereal \ joint_distribution X Y\)" - by (intro distribution_finite_prob_space finite_random_variable_pairI X Y) - -lemma finite_prob_space_eq: - "finite_prob_space M \ finite_measure_space M \ measure M (space M) = 1" - unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def - by auto - -lemma (in finite_prob_space) sum_over_space_eq_1: "(\x\space M. \ {x}) = 1" - using measure_space_1 sum_over_space by simp +definition "distributed M N X f \ distr M N X = density N f \ + f \ borel_measurable N \ (AE x in N. 0 \ f x) \ X \ measurable M N" -lemma (in finite_prob_space) joint_distribution_restriction_fst: - "joint_distribution X Y A \ distribution X (fst ` A)" - unfolding distribution_def -proof (safe intro!: finite_measure_mono) - fix x assume "x \ space M" and *: "(X x, Y x) \ A" - show "x \ X -` fst ` A" - by (auto intro!: image_eqI[OF _ *]) -qed (simp_all add: sets_eq_Pow) - -lemma (in finite_prob_space) joint_distribution_restriction_snd: - "joint_distribution X Y A \ distribution Y (snd ` A)" - unfolding distribution_def -proof (safe intro!: finite_measure_mono) - fix x assume "x \ space M" and *: "(X x, Y x) \ A" - show "x \ Y -` snd ` A" - by (auto intro!: image_eqI[OF _ *]) -qed (simp_all add: sets_eq_Pow) - -lemma (in finite_prob_space) distribution_order: - shows "0 \ distribution X x'" - and "(distribution X x' \ 0) \ (0 < distribution X x')" - and "r \ joint_distribution X Y {(x, y)} \ r \ distribution X {x}" - and "r \ joint_distribution X Y {(x, y)} \ r \ distribution Y {y}" - and "r < joint_distribution X Y {(x, y)} \ r < distribution X {x}" - and "r < joint_distribution X Y {(x, y)} \ r < distribution Y {y}" - and "distribution X {x} = 0 \ joint_distribution X Y {(x, y)} = 0" - and "distribution Y {y} = 0 \ joint_distribution X Y {(x, y)} = 0" - using - joint_distribution_restriction_fst[of X Y "{(x, y)}"] - joint_distribution_restriction_snd[of X Y "{(x, y)}"] - by (auto intro: antisym) - -lemma (in finite_prob_space) distribution_mono: - assumes "\t. \ t \ space M ; X t \ x \ \ Y t \ y" - shows "distribution X x \ distribution Y y" - unfolding distribution_def - using assms by (auto simp: sets_eq_Pow intro!: finite_measure_mono) +lemma + shows distributed_distr_eq_density: "distributed M N X f \ distr M N X = density N f" + and distributed_measurable: "distributed M N X f \ X \ measurable M N" + and distributed_borel_measurable: "distributed M N X f \ f \ borel_measurable N" + and distributed_AE: "distributed M N X f \ (AE x in N. 0 \ f x)" + by (simp_all add: distributed_def) -lemma (in finite_prob_space) distribution_mono_gt_0: - assumes gt_0: "0 < distribution X x" - assumes *: "\t. \ t \ space M ; X t \ x \ \ Y t \ y" - shows "0 < distribution Y y" - by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *) - -lemma (in finite_prob_space) sum_over_space_distrib: - "(\x\X`space M. distribution X {x}) = 1" - unfolding distribution_def prob_space[symmetric] using finite_space - by (subst finite_measure_finite_Union[symmetric]) - (auto simp add: disjoint_family_on_def sets_eq_Pow - intro!: arg_cong[where f=\']) - -lemma (in finite_prob_space) finite_sum_over_space_eq_1: - "(\x\space M. prob {x}) = 1" - using prob_space finite_space - by (subst (asm) finite_measure_finite_singleton) auto - -lemma (in prob_space) distribution_remove_const: - shows "joint_distribution X (\x. ()) {(x, ())} = distribution X {x}" - and "joint_distribution (\x. ()) X {((), x)} = distribution X {x}" - and "joint_distribution X (\x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}" - and "joint_distribution X (\x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}" - and "distribution (\x. ()) {()} = 1" - by (auto intro!: arg_cong[where f=\'] simp: distribution_def prob_space[symmetric]) +lemma + shows distributed_real_measurable: "distributed M N X (\x. ereal (f x)) \ f \ borel_measurable N" + and distributed_real_AE: "distributed M N X (\x. ereal (f x)) \ (AE x in N. 0 \ f x)" + by (simp_all add: distributed_def borel_measurable_ereal_iff) -lemma (in finite_prob_space) setsum_distribution_gen: - assumes "Z -` {c} \ space M = (\x \ X`space M. Y -` {f x}) \ space M" - and "inj_on f (X`space M)" - shows "(\x \ X`space M. distribution Y {f x}) = distribution Z {c}" - unfolding distribution_def assms - using finite_space assms - by (subst finite_measure_finite_Union[symmetric]) - (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def - intro!: arg_cong[where f=prob]) - -lemma (in finite_prob_space) setsum_distribution_cut: - "(\x \ X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}" - "(\y \ Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}" - "(\x \ X`space M. joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}" - "(\y \ Y`space M. joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}" - "(\z \ Z`space M. joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}" - by (auto intro!: inj_onI setsum_distribution_gen) - -lemma (in finite_prob_space) uniform_prob: - assumes "x \ space M" - assumes "\ x y. \x \ space M ; y \ space M\ \ prob {x} = prob {y}" - shows "prob {x} = 1 / card (space M)" +lemma distributed_count_space: + assumes X: "distributed M (count_space A) X P" and a: "a \ A" and A: "finite A" + shows "P a = emeasure M (X -` {a} \ space M)" proof - - have prob_x: "\ y. y \ space M \ prob {y} = prob {x}" - using assms(2)[OF _ `x \ space M`] by blast - have "1 = prob (space M)" - using prob_space by auto - also have "\ = (\ x \ space M. prob {x})" - using finite_measure_finite_Union[of "space M" "\ x. {x}", simplified] - sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format] - finite_space unfolding disjoint_family_on_def prob_space[symmetric] - by (auto simp add:setsum_restrict_set) - also have "\ = (\ y \ space M. prob {x})" - using prob_x by auto - also have "\ = real_of_nat (card (space M)) * prob {x}" by simp - finally have one: "1 = real (card (space M)) * prob {x}" - using real_eq_of_nat by auto - hence two: "real (card (space M)) \ 0" by fastforce - from one have three: "prob {x} \ 0" by fastforce - thus ?thesis using one two three divide_cancel_right - by (auto simp:field_simps) + have "emeasure M (X -` {a} \ space M) = emeasure (distr M (count_space A) X) {a}" + using X a A by (simp add: distributed_measurable emeasure_distr) + also have "\ = emeasure (density (count_space A) P) {a}" + using X by (simp add: distributed_distr_eq_density) + also have "\ = (\\<^isup>+x. P a * indicator {a} x \count_space A)" + using X a by (auto simp add: emeasure_density distributed_def indicator_def intro!: positive_integral_cong) + also have "\ = P a" + using X a by (subst positive_integral_cmult_indicator) (auto simp: distributed_def one_ereal_def[symmetric] AE_count_space) + finally show ?thesis .. qed -lemma (in prob_space) prob_space_subalgebra: - assumes "sigma_algebra N" "sets N \ sets M" "space N = space M" - and "\A. A \ sets N \ measure N A = \ A" - shows "prob_space N" -proof - interpret N: measure_space N - by (rule measure_space_subalgebra[OF assms]) - show "measure_space N" .. - show "measure N (space N) = 1" - using assms(4)[OF N.top] by (simp add: assms measure_space_1) +lemma distributed_cong_density: + "(AE x in N. f x = g x) \ g \ borel_measurable N \ f \ borel_measurable N \ + distributed M N X f \ distributed M N X g" + by (auto simp: distributed_def intro!: density_cong) + +lemma subdensity: + assumes T: "T \ measurable P Q" + assumes f: "distributed M P X f" + assumes g: "distributed M Q Y g" + assumes Y: "Y = T \ X" + shows "AE x in P. g (T x) = 0 \ f x = 0" +proof - + have "{x\space Q. g x = 0} \ null_sets (distr M Q (T \ X))" + using g Y by (auto simp: null_sets_density_iff distributed_def) + also have "distr M Q (T \ X) = distr (distr M P X) Q T" + using T f[THEN distributed_measurable] by (rule distr_distr[symmetric]) + finally have "T -` {x\space Q. g x = 0} \ space P \ null_sets (distr M P X)" + using T by (subst (asm) null_sets_distr_iff) auto + also have "T -` {x\space Q. g x = 0} \ space P = {x\space P. g (T x) = 0}" + using T by (auto dest: measurable_space) + finally show ?thesis + using f g by (auto simp add: null_sets_density_iff distributed_def) qed -lemma (in prob_space) prob_space_of_restricted_space: - assumes "\ A \ 0" "A \ sets M" - shows "prob_space (restricted_space A \measure := \S. \ S / \ A\)" - (is "prob_space ?P") +lemma subdensity_real: + fixes g :: "'a \ real" and f :: "'b \ real" + assumes T: "T \ measurable P Q" + assumes f: "distributed M P X f" + assumes g: "distributed M Q Y g" + assumes Y: "Y = T \ X" + shows "AE x in P. g (T x) = 0 \ f x = 0" + using subdensity[OF T, of M X "\x. ereal (f x)" Y "\x. ereal (g x)"] assms by auto + +lemma distributed_integral: + "distributed M N X f \ g \ borel_measurable N \ (\x. f x * g x \N) = (\x. g (X x) \M)" + by (auto simp: distributed_real_measurable distributed_real_AE distributed_measurable + distributed_distr_eq_density[symmetric] integral_density[symmetric] integral_distr) + +lemma distributed_transform_integral: + assumes Px: "distributed M N X Px" + assumes "distributed M P Y Py" + assumes Y: "Y = T \ X" and T: "T \ measurable N P" and f: "f \ borel_measurable P" + shows "(\x. Py x * f x \P) = (\x. Px x * f (T x) \N)" proof - - interpret A: measure_space "restricted_space A" - using `A \ sets M` by (rule restricted_measure_space) - interpret A': sigma_algebra ?P - by (rule A.sigma_algebra_cong) auto - show "prob_space ?P" - proof - show "measure_space ?P" - proof - show "positive ?P (measure ?P)" - proof (simp add: positive_def, safe) - fix B assume "B \ events" - with real_measure[of "A \ B"] real_measure[OF `A \ events`] `A \ sets M` - show "0 \ \ (A \ B) / \ A" by (auto simp: Int) - qed - show "countably_additive ?P (measure ?P)" - proof (simp add: countably_additive_def, safe) - fix B and F :: "nat \ 'a set" - assume F: "range F \ op \ A ` events" "disjoint_family F" - { fix i - from F have "F i \ op \ A ` events" by auto - with `A \ events` have "F i \ events" by auto } - moreover then have "range F \ events" by auto - moreover have "\S. \ S / \ A = inverse (\ A) * \ S" - by (simp add: mult_commute divide_ereal_def) - moreover have "0 \ inverse (\ A)" - using real_measure[OF `A \ events`] by auto - ultimately show "(\i. \ (F i) / \ A) = \ (\i. F i) / \ A" - using measure_countably_additive[of F] F - by (auto simp: suminf_cmult_ereal) - qed - qed - show "measure ?P (space ?P) = 1" - using real_measure[OF `A \ events`] `\ A \ 0` by auto - qed -qed - -lemma finite_prob_spaceI: - assumes "finite (space M)" "sets M = Pow(space M)" - and 1: "measure M (space M) = 1" and "\x. x \ space M \ 0 \ measure M {x}" - and add: "\A B. A \ space M \ measure M A = (\x\A. measure M {x})" - shows "finite_prob_space M" -proof - - interpret finite_measure_space M - proof - show "measure M (space M) \ \" using 1 by simp - qed fact+ - show ?thesis by default fact -qed - -lemma (in finite_prob_space) distribution_eq_setsum: - "distribution X A = (\x\A \ X ` space M. distribution X {x})" -proof - - have *: "X -` A \ space M = (\x\A \ X ` space M. X -` {x} \ space M)" - by auto - then show "distribution X A = (\x\A \ X ` space M. distribution X {x})" - using finite_space unfolding distribution_def * - by (intro finite_measure_finite_Union) - (auto simp: disjoint_family_on_def) -qed - -lemma (in finite_prob_space) distribution_eq_setsum_finite: - assumes "finite A" - shows "distribution X A = (\x\A. distribution X {x})" -proof - - note distribution_eq_setsum[of X A] - also have "(\x\A \ X ` space M. distribution X {x}) = (\x\A. distribution X {x})" - proof (intro setsum_mono_zero_cong_left ballI) - fix i assume "i\A - A \ X ` space M" - then have "X -` {i} \ space M = {}" by auto - then show "distribution X {i} = 0" - by (simp add: distribution_def) - next - show "finite A" by fact - qed simp_all + have "(\x. Py x * f x \P) = (\x. f (Y x) \M)" + by (rule distributed_integral) fact+ + also have "\ = (\x. f (T (X x)) \M)" + using Y by simp + also have "\ = (\x. Px x * f (T x) \N)" + using measurable_comp[OF T f] Px by (intro distributed_integral[symmetric]) (auto simp: comp_def) finally show ?thesis . qed -lemma (in finite_prob_space) finite_measure_space: - fixes X :: "'a \ 'x" - shows "finite_measure_space \space = X ` space M, sets = Pow (X ` space M), measure = ereal \ distribution X\" - (is "finite_measure_space ?S") -proof (rule finite_measure_spaceI, simp_all) - show "finite (X ` space M)" using finite_space by simp -next - fix A assume "A \ X ` space M" - then show "distribution X A = (\x\A. distribution X {x})" - by (subst distribution_eq_setsum) (simp add: Int_absorb2) +lemma distributed_marginal_eq_joint: + assumes T: "sigma_finite_measure T" + assumes S: "sigma_finite_measure S" + assumes Px: "distributed M S X Px" + assumes Py: "distributed M T Y Py" + assumes Pxy: "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) Pxy" + shows "AE y in T. Py y = (\\<^isup>+x. Pxy (x, y) \S)" +proof (rule sigma_finite_measure.density_unique[OF T]) + interpret ST: pair_sigma_finite S T using S T unfolding pair_sigma_finite_def by simp + show "Py \ borel_measurable T" "AE y in T. 0 \ Py y" + "(\x. \\<^isup>+ xa. Pxy (xa, x) \S) \ borel_measurable T" "AE y in T. 0 \ \\<^isup>+ x. Pxy (x, y) \S" + using Pxy[THEN distributed_borel_measurable] + by (auto intro!: Py[THEN distributed_borel_measurable] Py[THEN distributed_AE] + ST.positive_integral_snd_measurable' positive_integral_positive) + + show "density T Py = density T (\x. \\<^isup>+ xa. Pxy (xa, x) \S)" + proof (rule measure_eqI) + fix A assume A: "A \ sets (density T Py)" + have *: "\x y. x \ space S \ indicator (space S \ A) (x, y) = indicator A y" + by (auto simp: indicator_def) + have "emeasure (density T Py) A = emeasure (distr M T Y) A" + unfolding Py[THEN distributed_distr_eq_density] .. + also have "\ = emeasure (distr M (S \\<^isub>M T) (\x. (X x, Y x))) (space S \ A)" + using A Px Py Pxy + by (subst (1 2) emeasure_distr) + (auto dest: measurable_space distributed_measurable intro!: arg_cong[where f="emeasure M"]) + also have "\ = emeasure (density (S \\<^isub>M T) Pxy) (space S \ A)" + unfolding Pxy[THEN distributed_distr_eq_density] .. + also have "\ = (\\<^isup>+ x. Pxy x * indicator (space S \ A) x \(S \\<^isub>M T))" + using A Pxy by (simp add: emeasure_density distributed_borel_measurable) + also have "\ = (\\<^isup>+y. \\<^isup>+x. Pxy (x, y) * indicator (space S \ A) (x, y) \S \T)" + using A Pxy + by (subst ST.positive_integral_snd_measurable) (simp_all add: emeasure_density distributed_borel_measurable) + also have "\ = (\\<^isup>+y. (\\<^isup>+x. Pxy (x, y) \S) * indicator A y \T)" + using measurable_comp[OF measurable_Pair1[OF measurable_identity] distributed_borel_measurable[OF Pxy]] + using distributed_borel_measurable[OF Pxy] distributed_AE[OF Pxy, THEN ST.AE_pair] + by (subst (asm) ST.AE_commute) (auto intro!: positive_integral_cong_AE positive_integral_multc cong: positive_integral_cong simp: * comp_def) + also have "\ = emeasure (density T (\x. \\<^isup>+ xa. Pxy (xa, x) \S)) A" + using A by (intro emeasure_density[symmetric]) (auto intro!: ST.positive_integral_snd_measurable' Pxy[THEN distributed_borel_measurable]) + finally show "emeasure (density T Py) A = emeasure (density T (\x. \\<^isup>+ xa. Pxy (xa, x) \S)) A" . + qed simp qed -lemma (in finite_prob_space) finite_prob_space_of_images: - "finite_prob_space \ space = X ` space M, sets = Pow (X ` space M), measure = ereal \ distribution X \" - by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_ereal_def) +lemma (in prob_space) distr_marginal1: + fixes Pxy :: "('b \ 'c) \ real" + assumes "sigma_finite_measure S" "sigma_finite_measure T" + assumes Pxy: "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) Pxy" + defines "Px \ \x. real (\\<^isup>+z. Pxy (x, z) \T)" + shows "distributed M S X Px" + unfolding distributed_def +proof safe + interpret S: sigma_finite_measure S by fact + interpret T: sigma_finite_measure T by fact + interpret ST: pair_sigma_finite S T by default + + have XY: "(\x. (X x, Y x)) \ measurable M (S \\<^isub>M T)" + using Pxy by (rule distributed_measurable) + then show X: "X \ measurable M S" + unfolding measurable_pair_iff by (simp add: comp_def) + from XY have Y: "Y \ measurable M T" + unfolding measurable_pair_iff by (simp add: comp_def) + + from Pxy show borel: "(\x. ereal (Px x)) \ borel_measurable S" + by (auto intro!: ST.positive_integral_fst_measurable borel_measurable_real_of_ereal dest!: distributed_real_measurable simp: Px_def) -lemma (in finite_prob_space) finite_product_measure_space: - fixes X :: "'a \ 'x" and Y :: "'a \ 'y" - assumes "finite s1" "finite s2" - shows "finite_measure_space \ space = s1 \ s2, sets = Pow (s1 \ s2), measure = ereal \ joint_distribution X Y\" - (is "finite_measure_space ?M") -proof (rule finite_measure_spaceI, simp_all) - show "finite (s1 \ s2)" - using assms by auto -next - fix A assume "A \ (s1 \ s2)" - with assms show "joint_distribution X Y A = (\x\A. joint_distribution X Y {x})" - by (intro distribution_eq_setsum_finite) (auto dest: finite_subset) -qed - -lemma (in finite_prob_space) finite_product_measure_space_of_images: - shows "finite_measure_space \ space = X ` space M \ Y ` space M, - sets = Pow (X ` space M \ Y ` space M), - measure = ereal \ joint_distribution X Y \" - using finite_space by (auto intro!: finite_product_measure_space) - -lemma (in finite_prob_space) finite_product_prob_space_of_images: - "finite_prob_space \ space = X ` space M \ Y ` space M, sets = Pow (X ` space M \ Y ` space M), - measure = ereal \ joint_distribution X Y \" - (is "finite_prob_space ?S") -proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_ereal_def) - have "X -` X ` space M \ Y -` Y ` space M \ space M = space M" by auto - thus "joint_distribution X Y (X ` space M \ Y ` space M) = 1" - by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1) + interpret Pxy: prob_space "distr M (S \\<^isub>M T) (\x. (X x, Y x))" + using XY by (rule prob_space_distr) + have "(\\<^isup>+ x. max 0 (ereal (- Pxy x)) \(S \\<^isub>M T)) = (\\<^isup>+ x. 0 \(S \\<^isub>M T))" + using Pxy + by (intro positive_integral_cong_AE) (auto simp: max_def dest: distributed_real_measurable distributed_real_AE) + then have Pxy_integrable: "integrable (S \\<^isub>M T) Pxy" + using Pxy Pxy.emeasure_space_1 + by (simp add: integrable_def emeasure_density positive_integral_max_0 distributed_def borel_measurable_ereal_iff cong: positive_integral_cong) + + show "distr M S X = density S Px" + proof (rule measure_eqI) + fix A assume A: "A \ sets (distr M S X)" + with X Y XY have "emeasure (distr M S X) A = emeasure (distr M (S \\<^isub>M T) (\x. (X x, Y x))) (A \ space T)" + by (auto simp add: emeasure_distr + intro!: arg_cong[where f="emeasure M"] dest: measurable_space) + also have "\ = emeasure (density (S \\<^isub>M T) Pxy) (A \ space T)" + using Pxy by (simp add: distributed_def) + also have "\ = \\<^isup>+ x. \\<^isup>+ y. ereal (Pxy (x, y)) * indicator (A \ space T) (x, y) \T \S" + using A borel Pxy + by (simp add: emeasure_density ST.positive_integral_fst_measurable(2)[symmetric] distributed_def) + also have "\ = \\<^isup>+ x. ereal (Px x) * indicator A x \S" + apply (rule positive_integral_cong_AE) + using Pxy[THEN distributed_real_AE, THEN ST.AE_pair] ST.integrable_fst_measurable(1)[OF Pxy_integrable] AE_space + proof eventually_elim + fix x assume "x \ space S" "AE y in T. 0 \ Pxy (x, y)" and i: "integrable T (\y. Pxy (x, y))" + moreover have eq: "\y. y \ space T \ indicator (A \ space T) (x, y) = indicator A x" + by (auto simp: indicator_def) + ultimately have "(\\<^isup>+ y. ereal (Pxy (x, y)) * indicator (A \ space T) (x, y) \T) = + (\\<^isup>+ y. ereal (Pxy (x, y)) \T) * indicator A x" + using Pxy[THEN distributed_real_measurable] by (simp add: eq positive_integral_multc measurable_Pair2 cong: positive_integral_cong) + also have "(\\<^isup>+ y. ereal (Pxy (x, y)) \T) = Px x" + using i by (simp add: Px_def ereal_real integrable_def positive_integral_positive) + finally show "(\\<^isup>+ y. ereal (Pxy (x, y)) * indicator (A \ space T) (x, y) \T) = ereal (Px x) * indicator A x" . + qed + finally show "emeasure (distr M S X) A = emeasure (density S Px) A" + using A borel Pxy by (simp add: emeasure_density) + qed simp + + show "AE x in S. 0 \ ereal (Px x)" + by (simp add: Px_def positive_integral_positive real_of_ereal_pos) qed -subsection "Borel Measure on {0 ..< 1}" - -definition pborel :: "real measure_space" where - "pborel = lborel.restricted_space {0 ..< 1}" - -lemma space_pborel[simp]: - "space pborel = {0 ..< 1}" - unfolding pborel_def by auto - -lemma sets_pborel: - "A \ sets pborel \ A \ sets borel \ A \ { 0 ..< 1}" - unfolding pborel_def by auto +definition + "simple_distributed M X f \ distributed M (count_space (X`space M)) X (\x. ereal (f x)) \ + finite (X`space M)" -lemma in_pborel[intro, simp]: - "A \ {0 ..< 1} \ A \ sets borel \ A \ sets pborel" - unfolding pborel_def by auto - -interpretation pborel: measure_space pborel - using lborel.restricted_measure_space[of "{0 ..< 1}"] - by (simp add: pborel_def) +lemma simple_distributed: + "simple_distributed M X Px \ distributed M (count_space (X`space M)) X Px" + unfolding simple_distributed_def by auto -interpretation pborel: prob_space pborel -proof - show "measure pborel (space pborel) = 1" - by (simp add: one_ereal_def pborel_def) -qed default - -lemma pborel_prob: "pborel.prob A = (if A \ sets borel \ A \ {0 ..< 1} then real (lborel.\ A) else 0)" - unfolding pborel.\'_def by (auto simp: pborel_def) +lemma simple_distributed_finite[dest]: "simple_distributed M X P \ finite (X`space M)" + by (simp add: simple_distributed_def) -lemma pborel_singelton[simp]: "pborel.prob {a} = 0" - by (auto simp: pborel_prob) - -lemma - shows pborel_atLeastAtMost[simp]: "pborel.\' {a .. b} = (if 0 \ a \ a \ b \ b < 1 then b - a else 0)" - and pborel_atLeastLessThan[simp]: "pborel.\' {a ..< b} = (if 0 \ a \ a \ b \ b \ 1 then b - a else 0)" - and pborel_greaterThanAtMost[simp]: "pborel.\' {a <.. b} = (if 0 \ a \ a \ b \ b < 1 then b - a else 0)" - and pborel_greaterThanLessThan[simp]: "pborel.\' {a <..< b} = (if 0 \ a \ a \ b \ b \ 1 then b - a else 0)" - unfolding pborel_prob - by (auto simp: atLeastAtMost_subseteq_atLeastLessThan_iff - greaterThanAtMost_subseteq_atLeastLessThan_iff greaterThanLessThan_subseteq_atLeastLessThan_iff) - -lemma pborel_lebesgue_measure: - "A \ sets pborel \ pborel.prob A = real (measure lebesgue A)" - by (simp add: sets_pborel pborel_prob) +lemma (in prob_space) distributed_simple_function_superset: + assumes X: "simple_function M X" "\x. x \ X ` space M \ P x = measure M (X -` {x} \ space M)" + assumes A: "X`space M \ A" "finite A" + defines "S \ count_space A" and "P' \ (\x. if x \ X`space M then P x else 0)" + shows "distributed M S X P'" + unfolding distributed_def +proof safe + show "(\x. ereal (P' x)) \ borel_measurable S" unfolding S_def by simp + show "AE x in S. 0 \ ereal (P' x)" + using X by (auto simp: S_def P'_def simple_distributed_def intro!: measure_nonneg) + show "distr M S X = density S P'" + proof (rule measure_eqI_finite) + show "sets (distr M S X) = Pow A" "sets (density S P') = Pow A" + using A unfolding S_def by auto + show "finite A" by fact + fix a assume a: "a \ A" + then have "a \ X`space M \ X -` {a} \ space M = {}" by auto + with A a X have "emeasure (distr M S X) {a} = P' a" + by (subst emeasure_distr) + (auto simp add: S_def P'_def simple_functionD emeasure_eq_measure + intro!: arg_cong[where f=prob]) + also have "\ = (\\<^isup>+x. ereal (P' a) * indicator {a} x \S)" + using A X a + by (subst positive_integral_cmult_indicator) + (auto simp: S_def P'_def simple_distributed_def simple_functionD measure_nonneg) + also have "\ = (\\<^isup>+x. ereal (P' x) * indicator {a} x \S)" + by (auto simp: indicator_def intro!: positive_integral_cong) + also have "\ = emeasure (density S P') {a}" + using a A by (intro emeasure_density[symmetric]) (auto simp: S_def) + finally show "emeasure (distr M S X) {a} = emeasure (density S P') {a}" . + qed + show "random_variable S X" + using X(1) A by (auto simp: measurable_def simple_functionD S_def) +qed -lemma pborel_alt: - "pborel = sigma \ - space = {0..<1}, - sets = range (\(x,y). {x.. {0..<1}), - measure = measure lborel \" (is "_ = ?R") +lemma (in prob_space) simple_distributedI: + assumes X: "simple_function M X" "\x. x \ X ` space M \ P x = measure M (X -` {x} \ space M)" + shows "simple_distributed M X P" + unfolding simple_distributed_def +proof + have "distributed M (count_space (X ` space M)) X (\x. ereal (if x \ X`space M then P x else 0))" + (is "?A") + using simple_functionD[OF X(1)] by (intro distributed_simple_function_superset[OF X]) auto + also have "?A \ distributed M (count_space (X ` space M)) X (\x. ereal (P x))" + by (rule distributed_cong_density) auto + finally show "\" . +qed (rule simple_functionD[OF X(1)]) + +lemma simple_distributed_joint_finite: + assumes X: "simple_distributed M (\x. (X x, Y x)) Px" + shows "finite (X ` space M)" "finite (Y ` space M)" proof - - have *: "{0..<1::real} \ sets borel" by auto - have **: "op \ {0..<1::real} ` range (\(x, y). {x..(x,y). {x.. {0..<1})" - unfolding image_image by (intro arg_cong[where f=range]) auto - have "pborel = algebra.restricted_space (sigma \space=UNIV, sets=range (\(a, b). {a ..< b :: real}), - measure = measure pborel\) {0 ..< 1}" - by (simp add: sigma_def lebesgue_def pborel_def borel_eq_atLeastLessThan lborel_def) - also have "\ = ?R" - by (subst restricted_sigma) - (simp_all add: sets_sigma sigma_sets.Basic ** pborel_def image_eqI[of _ _ "(0,1)"]) - finally show ?thesis . + have "finite ((\x. (X x, Y x)) ` space M)" + using X by (auto simp: simple_distributed_def simple_functionD) + then have "finite (fst ` (\x. (X x, Y x)) ` space M)" "finite (snd ` (\x. (X x, Y x)) ` space M)" + by auto + then show fin: "finite (X ` space M)" "finite (Y ` space M)" + by (auto simp: image_image) +qed + +lemma simple_distributed_joint2_finite: + assumes X: "simple_distributed M (\x. (X x, Y x, Z x)) Px" + shows "finite (X ` space M)" "finite (Y ` space M)" "finite (Z ` space M)" +proof - + have "finite ((\x. (X x, Y x, Z x)) ` space M)" + using X by (auto simp: simple_distributed_def simple_functionD) + then have "finite (fst ` (\x. (X x, Y x, Z x)) ` space M)" + "finite ((fst \ snd) ` (\x. (X x, Y x, Z x)) ` space M)" + "finite ((snd \ snd) ` (\x. (X x, Y x, Z x)) ` space M)" + by auto + then show fin: "finite (X ` space M)" "finite (Y ` space M)" "finite (Z ` space M)" + by (auto simp: image_image) qed -subsection "Bernoulli space" +lemma simple_distributed_simple_function: + "simple_distributed M X Px \ simple_function M X" + unfolding simple_distributed_def distributed_def + by (auto simp: simple_function_def) + +lemma simple_distributed_measure: + "simple_distributed M X P \ a \ X`space M \ P a = measure M (X -` {a} \ space M)" + using distributed_count_space[of M "X`space M" X P a, symmetric] + by (auto simp: simple_distributed_def measure_def) + +lemma simple_distributed_nonneg: "simple_distributed M X f \ x \ space M \ 0 \ f (X x)" + by (auto simp: simple_distributed_measure measure_nonneg) -definition "bernoulli_space p = \ space = UNIV, sets = UNIV, - measure = ereal \ setsum (\b. if b then min 1 (max 0 p) else 1 - min 1 (max 0 p)) \" +lemma (in prob_space) simple_distributed_joint: + assumes X: "simple_distributed M (\x. (X x, Y x)) Px" + defines "S \ count_space (X`space M) \\<^isub>M count_space (Y`space M)" + defines "P \ (\x. if x \ (\x. (X x, Y x))`space M then Px x else 0)" + shows "distributed M S (\x. (X x, Y x)) P" +proof - + from simple_distributed_joint_finite[OF X, simp] + have S_eq: "S = count_space (X`space M \ Y`space M)" + by (simp add: S_def pair_measure_count_space) + show ?thesis + unfolding S_eq P_def + proof (rule distributed_simple_function_superset) + show "simple_function M (\x. (X x, Y x))" + using X by (rule simple_distributed_simple_function) + fix x assume "x \ (\x. (X x, Y x)) ` space M" + from simple_distributed_measure[OF X this] + show "Px x = prob ((\x. (X x, Y x)) -` {x} \ space M)" . + qed auto +qed -interpretation bernoulli: finite_prob_space "bernoulli_space p" for p - by (rule finite_prob_spaceI) - (auto simp: bernoulli_space_def UNIV_bool one_ereal_def setsum_Un_disjoint intro!: setsum_nonneg) +lemma (in prob_space) simple_distributed_joint2: + assumes X: "simple_distributed M (\x. (X x, Y x, Z x)) Px" + defines "S \ count_space (X`space M) \\<^isub>M count_space (Y`space M) \\<^isub>M count_space (Z`space M)" + defines "P \ (\x. if x \ (\x. (X x, Y x, Z x))`space M then Px x else 0)" + shows "distributed M S (\x. (X x, Y x, Z x)) P" +proof - + from simple_distributed_joint2_finite[OF X, simp] + have S_eq: "S = count_space (X`space M \ Y`space M \ Z`space M)" + by (simp add: S_def pair_measure_count_space) + show ?thesis + unfolding S_eq P_def + proof (rule distributed_simple_function_superset) + show "simple_function M (\x. (X x, Y x, Z x))" + using X by (rule simple_distributed_simple_function) + fix x assume "x \ (\x. (X x, Y x, Z x)) ` space M" + from simple_distributed_measure[OF X this] + show "Px x = prob ((\x. (X x, Y x, Z x)) -` {x} \ space M)" . + qed auto +qed + +lemma (in prob_space) simple_distributed_setsum_space: + assumes X: "simple_distributed M X f" + shows "setsum f (X`space M) = 1" +proof - + from X have "setsum f (X`space M) = prob (\i\X`space M. X -` {i} \ space M)" + by (subst finite_measure_finite_Union) + (auto simp add: disjoint_family_on_def simple_distributed_measure simple_distributed_simple_function simple_functionD + intro!: setsum_cong arg_cong[where f="prob"]) + also have "\ = prob (space M)" + by (auto intro!: arg_cong[where f=prob]) + finally show ?thesis + using emeasure_space_1 by (simp add: emeasure_eq_measure one_ereal_def) +qed -lemma bernoulli_measure: - "0 \ p \ p \ 1 \ bernoulli.prob p B = (\b\B. if b then p else 1 - p)" - unfolding bernoulli.\'_def unfolding bernoulli_space_def by (auto intro!: setsum_cong) +lemma (in prob_space) distributed_marginal_eq_joint_simple: + assumes Px: "simple_function M X" + assumes Py: "simple_distributed M Y Py" + assumes Pxy: "simple_distributed M (\x. (X x, Y x)) Pxy" + assumes y: "y \ Y`space M" + shows "Py y = (\x\X`space M. if (x, y) \ (\x. (X x, Y x)) ` space M then Pxy (x, y) else 0)" +proof - + note Px = simple_distributedI[OF Px refl] + have *: "\f A. setsum (\x. max 0 (ereal (f x))) A = ereal (setsum (\x. max 0 (f x)) A)" + by (simp add: setsum_ereal[symmetric] zero_ereal_def) + from distributed_marginal_eq_joint[OF sigma_finite_measure_count_space_finite sigma_finite_measure_count_space_finite + simple_distributed[OF Px] simple_distributed[OF Py] simple_distributed_joint[OF Pxy], + OF Py[THEN simple_distributed_finite] Px[THEN simple_distributed_finite]] + y Px[THEN simple_distributed_finite] Py[THEN simple_distributed_finite] + Pxy[THEN simple_distributed, THEN distributed_real_AE] + show ?thesis + unfolding AE_count_space + apply (elim ballE[where x=y]) + apply (auto simp add: positive_integral_count_space_finite * intro!: setsum_cong split: split_max) + done +qed -lemma bernoulli_measure_True: "0 \ p \ p \ 1 \ bernoulli.prob p {True} = p" - and bernoulli_measure_False: "0 \ p \ p \ 1 \ bernoulli.prob p {False} = 1 - p" - unfolding bernoulli_measure by simp_all + +lemma prob_space_uniform_measure: + assumes A: "emeasure M A \ 0" "emeasure M A \ \" + shows "prob_space (uniform_measure M A)" +proof + show "emeasure (uniform_measure M A) (space (uniform_measure M A)) = 1" + using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)], of "space M"] + using sets_into_space[OF emeasure_neq_0_sets[OF A(1)]] A + by (simp add: Int_absorb2 emeasure_nonneg) +qed + +lemma prob_space_uniform_count_measure: "finite A \ A \ {} \ prob_space (uniform_count_measure A)" + by default (auto simp: emeasure_uniform_count_measure space_uniform_count_measure one_ereal_def) end