| 43556 |      1 | (*  Title:      HOL/Probability/Conditional_Probability.thy
 | 
|  |      2 |     Author:     Johannes Hölzl, TU München
 | 
|  |      3 | *)
 | 
|  |      4 | 
 | 
|  |      5 | header {*Conditional probability*}
 | 
|  |      6 | 
 | 
|  |      7 | theory Conditional_Probability
 | 
|  |      8 | imports Probability_Measure Radon_Nikodym
 | 
|  |      9 | begin
 | 
|  |     10 | 
 | 
|  |     11 | section "Conditional Expectation and Probability"
 | 
|  |     12 | 
 | 
|  |     13 | definition (in prob_space)
 | 
|  |     14 |   "conditional_expectation N X = (SOME Y. Y\<in>borel_measurable N \<and> (\<forall>x. 0 \<le> Y x)
 | 
|  |     15 |     \<and> (\<forall>C\<in>sets N. (\<integral>\<^isup>+x. Y x * indicator C x\<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x\<partial>M)))"
 | 
|  |     16 | 
 | 
|  |     17 | lemma (in prob_space) conditional_expectation_exists:
 | 
| 43920 |     18 |   fixes X :: "'a \<Rightarrow> ereal" and N :: "('a, 'b) measure_space_scheme"
 | 
| 43556 |     19 |   assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
 | 
|  |     20 |   and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
 | 
|  |     21 |   shows "\<exists>Y\<in>borel_measurable N. (\<forall>x. 0 \<le> Y x) \<and> (\<forall>C\<in>sets N.
 | 
|  |     22 |       (\<integral>\<^isup>+x. Y x * indicator C x \<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x \<partial>M))"
 | 
|  |     23 | proof -
 | 
|  |     24 |   note N(4)[simp]
 | 
|  |     25 |   interpret P: prob_space N
 | 
|  |     26 |     using prob_space_subalgebra[OF N] .
 | 
|  |     27 | 
 | 
|  |     28 |   let "?f A" = "\<lambda>x. X x * indicator A x"
 | 
|  |     29 |   let "?Q A" = "integral\<^isup>P M (?f A)"
 | 
|  |     30 | 
 | 
|  |     31 |   from measure_space_density[OF borel]
 | 
|  |     32 |   have Q: "measure_space (N\<lparr> measure := ?Q \<rparr>)"
 | 
|  |     33 |     apply (rule measure_space.measure_space_subalgebra[of "M\<lparr> measure := ?Q \<rparr>"])
 | 
|  |     34 |     using N by (auto intro!: P.sigma_algebra_cong)
 | 
|  |     35 |   then interpret Q: measure_space "N\<lparr> measure := ?Q \<rparr>" .
 | 
|  |     36 | 
 | 
|  |     37 |   have "P.absolutely_continuous ?Q"
 | 
|  |     38 |     unfolding P.absolutely_continuous_def
 | 
|  |     39 |   proof safe
 | 
|  |     40 |     fix A assume "A \<in> sets N" "P.\<mu> A = 0"
 | 
|  |     41 |     then have f_borel: "?f A \<in> borel_measurable M" "AE x. x \<notin> A"
 | 
|  |     42 |       using borel N by (auto intro!: borel_measurable_indicator AE_not_in)
 | 
|  |     43 |     then show "?Q A = 0"
 | 
|  |     44 |       by (auto simp add: positive_integral_0_iff_AE)
 | 
|  |     45 |   qed
 | 
|  |     46 |   from P.Radon_Nikodym[OF Q this]
 | 
|  |     47 |   obtain Y where Y: "Y \<in> borel_measurable N" "\<And>x. 0 \<le> Y x"
 | 
|  |     48 |     "\<And>A. A \<in> sets N \<Longrightarrow> ?Q A =(\<integral>\<^isup>+x. Y x * indicator A x \<partial>N)"
 | 
|  |     49 |     by blast
 | 
|  |     50 |   with N(2) show ?thesis
 | 
|  |     51 |     by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ _ N(2,3,4,1)])
 | 
|  |     52 | qed
 | 
|  |     53 | 
 | 
|  |     54 | lemma (in prob_space)
 | 
| 43920 |     55 |   fixes X :: "'a \<Rightarrow> ereal" and N :: "('a, 'b) measure_space_scheme"
 | 
| 43556 |     56 |   assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
 | 
|  |     57 |   and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
 | 
|  |     58 |   shows borel_measurable_conditional_expectation:
 | 
|  |     59 |     "conditional_expectation N X \<in> borel_measurable N"
 | 
|  |     60 |   and conditional_expectation: "\<And>C. C \<in> sets N \<Longrightarrow>
 | 
|  |     61 |       (\<integral>\<^isup>+x. conditional_expectation N X x * indicator C x \<partial>M) =
 | 
|  |     62 |       (\<integral>\<^isup>+x. X x * indicator C x \<partial>M)"
 | 
|  |     63 |    (is "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C")
 | 
|  |     64 | proof -
 | 
|  |     65 |   note CE = conditional_expectation_exists[OF assms, unfolded Bex_def]
 | 
|  |     66 |   then show "conditional_expectation N X \<in> borel_measurable N"
 | 
|  |     67 |     unfolding conditional_expectation_def by (rule someI2_ex) blast
 | 
|  |     68 | 
 | 
|  |     69 |   from CE show "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C"
 | 
|  |     70 |     unfolding conditional_expectation_def by (rule someI2_ex) blast
 | 
|  |     71 | qed
 | 
|  |     72 | 
 | 
|  |     73 | lemma (in sigma_algebra) factorize_measurable_function_pos:
 | 
| 43920 |     74 |   fixes Z :: "'a \<Rightarrow> ereal" and Y :: "'a \<Rightarrow> 'c"
 | 
| 43556 |     75 |   assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
 | 
|  |     76 |   assumes Z: "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)"
 | 
|  |     77 |   shows "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. max 0 (Z x) = g (Y x)"
 | 
|  |     78 | proof -
 | 
|  |     79 |   interpret M': sigma_algebra M' by fact
 | 
|  |     80 |   have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
 | 
|  |     81 |   from M'.sigma_algebra_vimage[OF this]
 | 
|  |     82 |   interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
 | 
|  |     83 | 
 | 
|  |     84 |   from va.borel_measurable_implies_simple_function_sequence'[OF Z] guess f . note f = this
 | 
|  |     85 | 
 | 
|  |     86 |   have "\<forall>i. \<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
 | 
|  |     87 |   proof
 | 
|  |     88 |     fix i
 | 
|  |     89 |     from f(1)[of i] have "finite (f i`space M)" and B_ex:
 | 
|  |     90 |       "\<forall>z\<in>(f i)`space M. \<exists>B. B \<in> sets M' \<and> (f i) -` {z} \<inter> space M = Y -` B \<inter> space M"
 | 
|  |     91 |       unfolding simple_function_def by auto
 | 
|  |     92 |     from B_ex[THEN bchoice] guess B .. note B = this
 | 
|  |     93 | 
 | 
|  |     94 |     let ?g = "\<lambda>x. \<Sum>z\<in>f i`space M. z * indicator (B z) x"
 | 
|  |     95 | 
 | 
|  |     96 |     show "\<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
 | 
|  |     97 |     proof (intro exI[of _ ?g] conjI ballI)
 | 
|  |     98 |       show "simple_function M' ?g" using B by auto
 | 
|  |     99 | 
 | 
|  |    100 |       fix x assume "x \<in> space M"
 | 
| 43920 |    101 |       then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::ereal)"
 | 
| 43556 |    102 |         unfolding indicator_def using B by auto
 | 
|  |    103 |       then show "f i x = ?g (Y x)" using `x \<in> space M` f(1)[of i]
 | 
|  |    104 |         by (subst va.simple_function_indicator_representation) auto
 | 
|  |    105 |     qed
 | 
|  |    106 |   qed
 | 
|  |    107 |   from choice[OF this] guess g .. note g = this
 | 
|  |    108 | 
 | 
|  |    109 |   show ?thesis
 | 
|  |    110 |   proof (intro ballI bexI)
 | 
|  |    111 |     show "(\<lambda>x. SUP i. g i x) \<in> borel_measurable M'"
 | 
|  |    112 |       using g by (auto intro: M'.borel_measurable_simple_function)
 | 
|  |    113 |     fix x assume "x \<in> space M"
 | 
|  |    114 |     have "max 0 (Z x) = (SUP i. f i x)" using f by simp
 | 
|  |    115 |     also have "\<dots> = (SUP i. g i (Y x))"
 | 
|  |    116 |       using g `x \<in> space M` by simp
 | 
|  |    117 |     finally show "max 0 (Z x) = (SUP i. g i (Y x))" .
 | 
|  |    118 |   qed
 | 
|  |    119 | qed
 | 
|  |    120 | 
 | 
|  |    121 | lemma (in sigma_algebra) factorize_measurable_function:
 | 
| 43920 |    122 |   fixes Z :: "'a \<Rightarrow> ereal" and Y :: "'a \<Rightarrow> 'c"
 | 
| 43556 |    123 |   assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
 | 
|  |    124 |   shows "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)
 | 
|  |    125 |     \<longleftrightarrow> (\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x))"
 | 
|  |    126 | proof safe
 | 
|  |    127 |   interpret M': sigma_algebra M' by fact
 | 
|  |    128 |   have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
 | 
|  |    129 |   from M'.sigma_algebra_vimage[OF this]
 | 
|  |    130 |   interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
 | 
|  |    131 | 
 | 
| 43920 |    132 |   { fix g :: "'c \<Rightarrow> ereal" assume "g \<in> borel_measurable M'"
 | 
| 43556 |    133 |     with M'.measurable_vimage_algebra[OF Y]
 | 
|  |    134 |     have "g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
 | 
|  |    135 |       by (rule measurable_comp)
 | 
|  |    136 |     moreover assume "\<forall>x\<in>space M. Z x = g (Y x)"
 | 
|  |    137 |     then have "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y) \<longleftrightarrow>
 | 
|  |    138 |        g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
 | 
|  |    139 |        by (auto intro!: measurable_cong)
 | 
|  |    140 |     ultimately show "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
 | 
|  |    141 |       by simp }
 | 
|  |    142 | 
 | 
|  |    143 |   assume Z: "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
 | 
|  |    144 |   with assms have "(\<lambda>x. - Z x) \<in> borel_measurable M"
 | 
|  |    145 |     "(\<lambda>x. - Z x) \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
 | 
|  |    146 |     by auto
 | 
|  |    147 |   from factorize_measurable_function_pos[OF assms(1,2) this] guess n .. note n = this
 | 
|  |    148 |   from factorize_measurable_function_pos[OF assms Z] guess p .. note p = this
 | 
|  |    149 |   let "?g x" = "p x - n x"
 | 
|  |    150 |   show "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x)"
 | 
|  |    151 |   proof (intro bexI ballI)
 | 
|  |    152 |     show "?g \<in> borel_measurable M'" using p n by auto
 | 
|  |    153 |     fix x assume "x \<in> space M"
 | 
|  |    154 |     then have "p (Y x) = max 0 (Z x)" "n (Y x) = max 0 (- Z x)"
 | 
|  |    155 |       using p n by auto
 | 
|  |    156 |     then show "Z x = ?g (Y x)"
 | 
|  |    157 |       by (auto split: split_max)
 | 
|  |    158 |   qed
 | 
|  |    159 | qed
 | 
|  |    160 | 
 | 
|  |    161 | end |