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