author | paulson <lp15@cam.ac.uk> |
Fri, 13 Apr 2018 15:58:27 +0100 | |
changeset 67977 | 557ea2740125 |
parent 67226 | ec32cdaab97b |
child 69712 | dc85b5b3a532 |
permissions | -rw-r--r-- |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1 |
(* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
2 |
License: BSD |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
3 |
*) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
4 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
5 |
section \<open>Conditional Expectation\<close> |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
6 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
7 |
theory Conditional_Expectation |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
8 |
imports Probability_Measure |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
9 |
begin |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
10 |
|
67226 | 11 |
subsection \<open>Restricting a measure to a sub-sigma-algebra\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
12 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
13 |
definition subalgebra::"'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
14 |
"subalgebra M F = ((space F = space M) \<and> (sets F \<subseteq> sets M))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
15 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
16 |
lemma sub_measure_space: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
17 |
assumes i: "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
18 |
shows "measure_space (space M) (sets F) (emeasure M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
19 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
20 |
have "sigma_algebra (space M) (sets F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
21 |
by (metis i measure_space measure_space_def subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
22 |
moreover have "positive (sets F) (emeasure M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
23 |
using Sigma_Algebra.positive_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
24 |
moreover have "countably_additive (sets F) (emeasure M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
25 |
by (meson countably_additive_def emeasure_countably_additive i order_trans subalgebra_def subsetCE) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
26 |
ultimately show ?thesis unfolding measure_space_def by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
27 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
28 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
29 |
definition restr_to_subalg::"'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
30 |
"restr_to_subalg M F = measure_of (space M) (sets F) (emeasure M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
31 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
32 |
lemma space_restr_to_subalg: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
33 |
"space (restr_to_subalg M F) = space M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
34 |
unfolding restr_to_subalg_def by (simp add: space_measure_of_conv) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
35 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
36 |
lemma sets_restr_to_subalg [measurable_cong]: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
37 |
assumes "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
38 |
shows "sets (restr_to_subalg M F) = sets F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
39 |
unfolding restr_to_subalg_def by (metis sets.sets_measure_of_eq assms subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
40 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
41 |
lemma emeasure_restr_to_subalg: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
42 |
assumes "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
43 |
"A \<in> sets F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
44 |
shows "emeasure (restr_to_subalg M F) A = emeasure M A" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
45 |
unfolding restr_to_subalg_def |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
46 |
by (metis assms subalgebra_def emeasure_measure_of_conv sub_measure_space sets.sigma_sets_eq) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
47 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
48 |
lemma null_sets_restr_to_subalg: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
49 |
assumes "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
50 |
shows "null_sets (restr_to_subalg M F) = null_sets M \<inter> sets F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
51 |
proof |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
52 |
have "x \<in> null_sets M \<inter> sets F" if "x \<in> null_sets (restr_to_subalg M F)" for x |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
53 |
by (metis that Int_iff assms emeasure_restr_to_subalg null_setsD1 null_setsD2 null_setsI |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
54 |
sets_restr_to_subalg subalgebra_def subsetD) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
55 |
then show "null_sets (restr_to_subalg M F) \<subseteq> null_sets M \<inter> sets F" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
56 |
next |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
57 |
have "x \<in> null_sets (restr_to_subalg M F)" if "x \<in> null_sets M \<inter> sets F" for x |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
58 |
by (metis that Int_iff assms null_setsD1 null_setsI sets_restr_to_subalg emeasure_restr_to_subalg[OF assms]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
59 |
then show "null_sets M \<inter> sets F \<subseteq> null_sets (restr_to_subalg M F)" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
60 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
61 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
62 |
lemma AE_restr_to_subalg: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
63 |
assumes "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
64 |
"AE x in (restr_to_subalg M F). P x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
65 |
shows "AE x in M. P x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
66 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
67 |
obtain A where A: "\<And>x. x \<in> space (restr_to_subalg M F) - A \<Longrightarrow> P x" "A \<in> null_sets (restr_to_subalg M F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
68 |
using AE_E3[OF assms(2)] by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
69 |
then have "A \<in> null_sets M" using null_sets_restr_to_subalg[OF assms(1)] by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
70 |
moreover have "\<And>x. x \<in> space M - A \<Longrightarrow> P x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
71 |
using space_restr_to_subalg A(1) by fastforce |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
72 |
ultimately show ?thesis |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
73 |
unfolding eventually_ae_filter by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
74 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
75 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
76 |
lemma AE_restr_to_subalg2: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
77 |
assumes "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
78 |
"AE x in M. P x" and [measurable]: "P \<in> measurable F (count_space UNIV)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
79 |
shows "AE x in (restr_to_subalg M F). P x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
80 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
81 |
define U where "U = {x \<in> space M. \<not>(P x)}" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
82 |
then have *: "U = {x \<in> space F. \<not>(P x)}" using assms(1) by (simp add: subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
83 |
then have "U \<in> sets F" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
84 |
then have "U \<in> sets M" using assms(1) by (meson subalgebra_def subsetD) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
85 |
then have "U \<in> null_sets M" unfolding U_def using assms(2) using AE_iff_measurable by blast |
67226 | 86 |
then have "U \<in> null_sets (restr_to_subalg M F)" using null_sets_restr_to_subalg[OF assms(1)] \<open>U \<in> sets F\<close> by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
87 |
then show ?thesis using * by (metis (no_types, lifting) Collect_mono U_def eventually_ae_filter space_restr_to_subalg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
88 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
89 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
90 |
lemma prob_space_restr_to_subalg: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
91 |
assumes "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
92 |
"prob_space M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
93 |
shows "prob_space (restr_to_subalg M F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
94 |
by (metis (no_types, lifting) assms(1) assms(2) emeasure_restr_to_subalg prob_space.emeasure_space_1 prob_spaceI |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
95 |
sets.top space_restr_to_subalg subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
96 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
97 |
lemma finite_measure_restr_to_subalg: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
98 |
assumes "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
99 |
"finite_measure M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
100 |
shows "finite_measure (restr_to_subalg M F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
101 |
by (metis (no_types, lifting) assms emeasure_restr_to_subalg finite_measure.finite_emeasure_space |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
102 |
finite_measureI sets.top space_restr_to_subalg subalgebra_def infinity_ennreal_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
103 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
104 |
lemma measurable_in_subalg: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
105 |
assumes "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
106 |
"f \<in> measurable F N" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
107 |
shows "f \<in> measurable (restr_to_subalg M F) N" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
108 |
by (metis measurable_cong_sets assms(2) sets_restr_to_subalg[OF assms(1)]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
109 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
110 |
lemma measurable_in_subalg': |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
111 |
assumes "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
112 |
"f \<in> measurable (restr_to_subalg M F) N" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
113 |
shows "f \<in> measurable F N" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
114 |
by (metis measurable_cong_sets assms(2) sets_restr_to_subalg[OF assms(1)]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
115 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
116 |
lemma measurable_from_subalg: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
117 |
assumes "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
118 |
"f \<in> measurable F N" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
119 |
shows "f \<in> measurable M N" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
120 |
using assms unfolding measurable_def subalgebra_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
121 |
|
67226 | 122 |
text\<open>The following is the direct transposition of \verb+nn_integral_subalgebra+ |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
123 |
(from \verb+Nonnegative_Lebesgue_Integration+) in the |
67226 | 124 |
current notations, with the removal of the useless assumption $f \geq 0$.\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
125 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
126 |
lemma nn_integral_subalgebra2: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
127 |
assumes "subalgebra M F" and [measurable]: "f \<in> borel_measurable F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
128 |
shows "(\<integral>\<^sup>+ x. f x \<partial>(restr_to_subalg M F)) = (\<integral>\<^sup>+ x. f x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
129 |
proof (rule nn_integral_subalgebra) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
130 |
show "f \<in> borel_measurable (restr_to_subalg M F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
131 |
by (rule measurable_in_subalg[OF assms(1)]) simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
132 |
show "sets (restr_to_subalg M F) \<subseteq> sets M" by (metis sets_restr_to_subalg[OF assms(1)] assms(1) subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
133 |
fix A assume "A \<in> sets (restr_to_subalg M F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
134 |
then show "emeasure (restr_to_subalg M F) A = emeasure M A" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
135 |
by (metis sets_restr_to_subalg[OF assms(1)] emeasure_restr_to_subalg[OF assms(1)]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
136 |
qed (auto simp add: assms space_restr_to_subalg sets_restr_to_subalg[OF assms(1)]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
137 |
|
67226 | 138 |
text\<open>The following is the direct transposition of \verb+integral_subalgebra+ |
139 |
(from \verb+Bochner_Integration+) in the current notations.\<close> |
|
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
140 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
141 |
lemma integral_subalgebra2: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
142 |
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
143 |
assumes "subalgebra M F" and |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
144 |
[measurable]: "f \<in> borel_measurable F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
145 |
shows "(\<integral>x. f x \<partial>(restr_to_subalg M F)) = (\<integral>x. f x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
146 |
by (rule integral_subalgebra, |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
147 |
metis measurable_in_subalg[OF assms(1)] assms(2), |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
148 |
auto simp add: assms space_restr_to_subalg sets_restr_to_subalg emeasure_restr_to_subalg, |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
149 |
meson assms(1) subalgebra_def subset_eq) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
150 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
151 |
lemma integrable_from_subalg: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
152 |
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
153 |
assumes "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
154 |
"integrable (restr_to_subalg M F) f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
155 |
shows "integrable M f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
156 |
proof (rule integrableI_bounded) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
157 |
have [measurable]: "f \<in> borel_measurable F" using assms by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
158 |
then show "f \<in> borel_measurable M" using assms(1) measurable_from_subalg by blast |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
159 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
160 |
have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>(restr_to_subalg M F))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
161 |
by (rule nn_integral_subalgebra2[symmetric], auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
162 |
also have "... < \<infinity>" using integrable_iff_bounded assms by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
163 |
finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
164 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
165 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
166 |
lemma integrable_in_subalg: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
167 |
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
168 |
assumes [measurable]: "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
169 |
"f \<in> borel_measurable F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
170 |
"integrable M f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
171 |
shows "integrable (restr_to_subalg M F) f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
172 |
proof (rule integrableI_bounded) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
173 |
show "f \<in> borel_measurable (restr_to_subalg M F)" using assms(2) assms(1) by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
174 |
have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>(restr_to_subalg M F)) = (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
175 |
by (rule nn_integral_subalgebra2, auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
176 |
also have "... < \<infinity>" using integrable_iff_bounded assms by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
177 |
finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>(restr_to_subalg M F)) < \<infinity>" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
178 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
179 |
|
67226 | 180 |
subsection \<open>Nonnegative conditional expectation\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
181 |
|
67226 | 182 |
text \<open>The conditional expectation of a function $f$, on a measure space $M$, with respect to a |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
183 |
sub sigma algebra $F$, should be a function $g$ which is $F$-measurable whose integral on any |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
184 |
$F$-set coincides with the integral of $f$. |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
185 |
Such a function is uniquely defined almost everywhere. |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
186 |
The most direct construction is to use the measure $f dM$, restrict it to the sigma-algebra $F$, |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
187 |
and apply the Radon-Nikodym theorem to write it as $g dM_{|F}$ for some $F$-measurable function $g$. |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
188 |
Another classical construction for $L^2$ functions is done by orthogonal projection on $F$-measurable |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
189 |
functions, and then extending by density to $L^1$. The Radon-Nikodym point of view avoids the $L^2$ |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
190 |
machinery, and works for all positive functions. |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
191 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
192 |
In this paragraph, we develop the definition and basic properties for nonnegative functions, |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
193 |
as the basics of the general case. As in the definition of integrals, the nonnegative case is done |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
194 |
with ennreal-valued functions, without any integrability assumption. |
67226 | 195 |
\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
196 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
197 |
definition nn_cond_exp :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ('a \<Rightarrow> ennreal)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
198 |
where |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
199 |
"nn_cond_exp M F f = |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
200 |
(if f \<in> borel_measurable M \<and> subalgebra M F |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
201 |
then RN_deriv (restr_to_subalg M F) (restr_to_subalg (density M f) F) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
202 |
else (\<lambda>_. 0))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
203 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
204 |
lemma |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
205 |
shows borel_measurable_nn_cond_exp [measurable]: "nn_cond_exp M F f \<in> borel_measurable F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
206 |
and borel_measurable_nn_cond_exp2 [measurable]: "nn_cond_exp M F f \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
207 |
by (simp_all add: nn_cond_exp_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
208 |
(metis borel_measurable_RN_deriv borel_measurable_subalgebra sets_restr_to_subalg space_restr_to_subalg subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
209 |
|
67226 | 210 |
text \<open>The good setting for conditional expectations is the situation where the subalgebra $F$ |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
211 |
gives rise to a sigma-finite measure space. To see what goes wrong if it is not sigma-finite, |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
212 |
think of $\mathbb{R}$ with the trivial sigma-algebra $\{\emptyset, \mathbb{R}\}$. In this case, |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
213 |
conditional expectations have to be constant functions, so they have integral $0$ or $\infty$. |
67226 | 214 |
This means that a positive integrable function can have no meaningful conditional expectation.\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
215 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
216 |
locale sigma_finite_subalgebra = |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
217 |
fixes M F::"'a measure" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
218 |
assumes subalg: "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
219 |
and sigma_fin_subalg: "sigma_finite_measure (restr_to_subalg M F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
220 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
221 |
lemma sigma_finite_subalgebra_is_sigma_finite: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
222 |
assumes "sigma_finite_subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
223 |
shows "sigma_finite_measure M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
224 |
proof |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
225 |
have subalg: "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
226 |
and sigma_fin_subalg: "sigma_finite_measure (restr_to_subalg M F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
227 |
using assms unfolding sigma_finite_subalgebra_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
228 |
obtain A where Ap: "countable A \<and> A \<subseteq> sets (restr_to_subalg M F) \<and> \<Union>A = space (restr_to_subalg M F) \<and> (\<forall>a\<in>A. emeasure (restr_to_subalg M F) a \<noteq> \<infinity>)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
229 |
using sigma_finite_measure.sigma_finite_countable[OF sigma_fin_subalg] by fastforce |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
230 |
have "A \<subseteq> sets F" using Ap sets_restr_to_subalg[OF subalg] by fastforce |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
231 |
then have "A \<subseteq> sets M" using subalg subalgebra_def by force |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
232 |
moreover have "\<Union>A = space M" using Ap space_restr_to_subalg by simp |
67226 | 233 |
moreover have "\<forall>a\<in>A. emeasure M a \<noteq> \<infinity>" by (metis subsetD emeasure_restr_to_subalg[OF subalg] \<open>A \<subseteq> sets F\<close> Ap) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
234 |
ultimately show "\<exists>A. countable A \<and> A \<subseteq> sets M \<and> \<Union>A = space M \<and> (\<forall>a\<in>A. emeasure M a \<noteq> \<infinity>)" using Ap by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
235 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
236 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
237 |
sublocale sigma_finite_subalgebra \<subseteq> sigma_finite_measure |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
238 |
using sigma_finite_subalgebra_is_sigma_finite sigma_finite_subalgebra_axioms by blast |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
239 |
|
67226 | 240 |
text \<open>Conditional expectations are very often used in probability spaces. This is a special case |
241 |
of the previous one, as we prove now.\<close> |
|
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
242 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
243 |
locale finite_measure_subalgebra = finite_measure + |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
244 |
fixes F::"'a measure" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
245 |
assumes subalg: "subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
246 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
247 |
lemma finite_measure_subalgebra_is_sigma_finite: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
248 |
assumes "finite_measure_subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
249 |
shows "sigma_finite_subalgebra M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
250 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
251 |
interpret finite_measure_subalgebra M F using assms by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
252 |
have "finite_measure (restr_to_subalg M F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
253 |
using finite_measure_restr_to_subalg subalg finite_emeasure_space finite_measureI unfolding infinity_ennreal_def by blast |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
254 |
then have "sigma_finite_measure (restr_to_subalg M F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
255 |
unfolding finite_measure_def by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
256 |
then show "sigma_finite_subalgebra M F" unfolding sigma_finite_subalgebra_def using subalg by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
257 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
258 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
259 |
sublocale finite_measure_subalgebra \<subseteq> sigma_finite_subalgebra |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
260 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
261 |
have "finite_measure (restr_to_subalg M F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
262 |
using finite_measure_restr_to_subalg subalg finite_emeasure_space finite_measureI unfolding infinity_ennreal_def by blast |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
263 |
then have "sigma_finite_measure (restr_to_subalg M F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
264 |
unfolding finite_measure_def by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
265 |
then show "sigma_finite_subalgebra M F" unfolding sigma_finite_subalgebra_def using subalg by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
266 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
267 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
268 |
context sigma_finite_subalgebra |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
269 |
begin |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
270 |
|
67226 | 271 |
text\<open>The next lemma is arguably the most fundamental property of conditional expectation: |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
272 |
when computing an expectation against an $F$-measurable function, it is equivalent to work |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
273 |
with a function or with its $F$-conditional expectation. |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
274 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
275 |
This property (even for bounded test functions) characterizes conditional expectations, as the second lemma below shows. |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
276 |
From this point on, we will only work with it, and forget completely about |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
277 |
the definition using Radon-Nikodym derivatives. |
67226 | 278 |
\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
279 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
280 |
lemma nn_cond_exp_intg: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
281 |
assumes [measurable]: "f \<in> borel_measurable F" "g \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
282 |
shows "(\<integral>\<^sup>+ x. f x * nn_cond_exp M F g x \<partial>M) = (\<integral>\<^sup>+ x. f x * g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
283 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
284 |
have [measurable]: "f \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
285 |
by (meson assms subalg borel_measurable_subalgebra subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
286 |
have ac: "absolutely_continuous (restr_to_subalg M F) (restr_to_subalg (density M g) F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
287 |
unfolding absolutely_continuous_def |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
288 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
289 |
have "null_sets (restr_to_subalg M F) = null_sets M \<inter> sets F" by (rule null_sets_restr_to_subalg[OF subalg]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
290 |
moreover have "null_sets M \<subseteq> null_sets (density M g)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
291 |
by (rule absolutely_continuousI_density[unfolded absolutely_continuous_def]) auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
292 |
ultimately have "null_sets (restr_to_subalg M F) \<subseteq> null_sets (density M g) \<inter> sets F" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
293 |
moreover have "null_sets (density M g) \<inter> sets F = null_sets (restr_to_subalg (density M g) F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
294 |
by (rule null_sets_restr_to_subalg[symmetric]) (metis subalg sets_density space_density subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
295 |
ultimately show "null_sets (restr_to_subalg M F) \<subseteq> null_sets (restr_to_subalg (density M g) F)" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
296 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
297 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
298 |
have "(\<integral>\<^sup>+ x. f x * nn_cond_exp M F g x \<partial>M) = (\<integral>\<^sup>+ x. f x * nn_cond_exp M F g x \<partial>(restr_to_subalg M F))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
299 |
by (rule nn_integral_subalgebra2[symmetric]) (simp_all add: assms subalg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
300 |
also have "... = (\<integral>\<^sup>+ x. f x * RN_deriv (restr_to_subalg M F) (restr_to_subalg (density M g) F) x \<partial>(restr_to_subalg M F))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
301 |
unfolding nn_cond_exp_def using assms subalg by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
302 |
also have "... = (\<integral>\<^sup>+ x. RN_deriv (restr_to_subalg M F) (restr_to_subalg (density M g) F) x * f x \<partial>(restr_to_subalg M F))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
303 |
by (simp add: mult.commute) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
304 |
also have "... = (\<integral>\<^sup>+ x. f x \<partial>(restr_to_subalg (density M g) F))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
305 |
proof (rule sigma_finite_measure.RN_deriv_nn_integral[symmetric]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
306 |
show "sets (restr_to_subalg (density M g) F) = sets (restr_to_subalg M F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
307 |
by (metis subalg restr_to_subalg_def sets.sets_measure_of_eq space_density subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
308 |
qed (auto simp add: assms measurable_restrict ac measurable_in_subalg subalg sigma_fin_subalg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
309 |
also have "... = (\<integral>\<^sup>+ x. f x \<partial>(density M g))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
310 |
by (metis nn_integral_subalgebra2 subalg assms(1) sets_density space_density subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
311 |
also have "... = (\<integral>\<^sup>+ x. g x * f x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
312 |
by (rule nn_integral_density) (simp_all add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
313 |
also have "... = (\<integral>\<^sup>+ x. f x * g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
314 |
by (simp add: mult.commute) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
315 |
finally show ?thesis by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
316 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
317 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
318 |
lemma nn_cond_exp_charact: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
319 |
assumes "\<And>A. A \<in> sets F \<Longrightarrow> (\<integral>\<^sup>+ x \<in> A. f x \<partial>M) = (\<integral>\<^sup>+ x \<in> A. g x \<partial>M)" and |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
320 |
[measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
321 |
shows "AE x in M. g x = nn_cond_exp M F f x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
322 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
323 |
let ?MF = "restr_to_subalg M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
324 |
{ |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
325 |
fix A assume "A \<in> sets ?MF" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
326 |
then have [measurable]: "A \<in> sets F" using sets_restr_to_subalg[OF subalg] by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
327 |
have "(\<integral>\<^sup>+ x \<in> A. g x \<partial> ?MF) = (\<integral>\<^sup>+ x \<in> A. g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
328 |
by (simp add: nn_integral_subalgebra2 subalg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
329 |
also have "... = (\<integral>\<^sup>+ x \<in> A. f x \<partial>M)" using assms(1) by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
330 |
also have "... = (\<integral>\<^sup>+ x. indicator A x * f x \<partial>M)" by (simp add: mult.commute) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
331 |
also have "... = (\<integral>\<^sup>+ x. indicator A x * nn_cond_exp M F f x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
332 |
by (rule nn_cond_exp_intg[symmetric]) (auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
333 |
also have "... = (\<integral>\<^sup>+ x \<in> A. nn_cond_exp M F f x \<partial>M)" by (simp add: mult.commute) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
334 |
also have "... = (\<integral>\<^sup>+ x \<in> A. nn_cond_exp M F f x \<partial> ?MF)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
335 |
by (simp add: nn_integral_subalgebra2 subalg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
336 |
finally have "(\<integral>\<^sup>+ x \<in> A. g x \<partial> ?MF) = (\<integral>\<^sup>+ x \<in> A. nn_cond_exp M F f x \<partial> ?MF)" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
337 |
} note * = this |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
338 |
have "AE x in ?MF. g x = nn_cond_exp M F f x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
339 |
by (rule sigma_finite_measure.density_unique2) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
340 |
(auto simp add: assms subalg sigma_fin_subalg AE_restr_to_subalg2 *) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
341 |
then show ?thesis using AE_restr_to_subalg[OF subalg] by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
342 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
343 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
344 |
lemma nn_cond_exp_F_meas: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
345 |
assumes "f \<in> borel_measurable F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
346 |
shows "AE x in M. f x = nn_cond_exp M F f x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
347 |
by (rule nn_cond_exp_charact) (auto simp add: assms measurable_from_subalg[OF subalg]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
348 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
349 |
lemma nn_cond_exp_prod: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
350 |
assumes [measurable]: "f \<in> borel_measurable F" "g \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
351 |
shows "AE x in M. f x * nn_cond_exp M F g x = nn_cond_exp M F (\<lambda>x. f x * g x) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
352 |
proof (rule nn_cond_exp_charact) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
353 |
have [measurable]: "f \<in> borel_measurable M" by (rule measurable_from_subalg[OF subalg assms(1)]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
354 |
show "(\<lambda>x. f x * g x) \<in> borel_measurable M" by measurable |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
355 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
356 |
fix A assume "A \<in> sets F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
357 |
then have [measurable]: "(\<lambda>x. f x * indicator A x) \<in> borel_measurable F" by measurable |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
358 |
have "\<integral>\<^sup>+x\<in>A. (f x * g x) \<partial>M = \<integral>\<^sup>+x. (f x * indicator A x) * g x \<partial>M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
359 |
by (simp add: mult.commute mult.left_commute) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
360 |
also have "... = \<integral>\<^sup>+x. (f x * indicator A x) * nn_cond_exp M F g x \<partial>M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
361 |
by (rule nn_cond_exp_intg[symmetric]) (auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
362 |
also have "... = \<integral>\<^sup>+x\<in>A. (f x * nn_cond_exp M F g x)\<partial>M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
363 |
by (simp add: mult.commute mult.left_commute) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
364 |
finally show "\<integral>\<^sup>+x\<in>A. (f x * g x) \<partial>M = \<integral>\<^sup>+x\<in>A. (f x * nn_cond_exp M F g x)\<partial>M" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
365 |
qed (auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
366 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
367 |
lemma nn_cond_exp_sum: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
368 |
assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
369 |
shows "AE x in M. nn_cond_exp M F f x + nn_cond_exp M F g x = nn_cond_exp M F (\<lambda>x. f x + g x) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
370 |
proof (rule nn_cond_exp_charact) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
371 |
fix A assume [measurable]: "A \<in> sets F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
372 |
then have "A \<in> sets M" by (meson subalg subalgebra_def subsetD) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
373 |
have "\<integral>\<^sup>+x\<in>A. (nn_cond_exp M F f x + nn_cond_exp M F g x)\<partial>M = (\<integral>\<^sup>+x\<in>A. nn_cond_exp M F f x \<partial>M) + (\<integral>\<^sup>+x\<in>A. nn_cond_exp M F g x \<partial>M)" |
67226 | 374 |
by (rule nn_set_integral_add) (auto simp add: assms \<open>A \<in> sets M\<close>) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
375 |
also have "... = (\<integral>\<^sup>+x. indicator A x * nn_cond_exp M F f x \<partial>M) + (\<integral>\<^sup>+x. indicator A x * nn_cond_exp M F g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
376 |
by (metis (no_types, lifting) mult.commute nn_integral_cong) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
377 |
also have "... = (\<integral>\<^sup>+x. indicator A x * f x \<partial>M) + (\<integral>\<^sup>+x. indicator A x * g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
378 |
by (simp add: nn_cond_exp_intg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
379 |
also have "... = (\<integral>\<^sup>+x\<in>A. f x \<partial>M) + (\<integral>\<^sup>+x\<in>A. g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
380 |
by (metis (no_types, lifting) mult.commute nn_integral_cong) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
381 |
also have "... = \<integral>\<^sup>+x\<in>A. (f x + g x)\<partial>M" |
67226 | 382 |
by (rule nn_set_integral_add[symmetric]) (auto simp add: assms \<open>A \<in> sets M\<close>) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
383 |
finally show "\<integral>\<^sup>+x\<in>A. (f x + g x)\<partial>M = \<integral>\<^sup>+x\<in>A. (nn_cond_exp M F f x + nn_cond_exp M F g x)\<partial>M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
384 |
by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
385 |
qed (auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
386 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
387 |
lemma nn_cond_exp_cong: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
388 |
assumes "AE x in M. f x = g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
389 |
and [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
390 |
shows "AE x in M. nn_cond_exp M F f x = nn_cond_exp M F g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
391 |
proof (rule nn_cond_exp_charact) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
392 |
fix A assume [measurable]: "A \<in> sets F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
393 |
have "\<integral>\<^sup>+x\<in>A. nn_cond_exp M F f x \<partial>M = \<integral>\<^sup>+x. indicator A x * nn_cond_exp M F f x \<partial>M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
394 |
by (simp add: mult.commute) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
395 |
also have "... = \<integral>\<^sup>+x. indicator A x * f x \<partial>M" by (simp add: nn_cond_exp_intg assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
396 |
also have "... = \<integral>\<^sup>+x\<in>A. f x \<partial>M" by (simp add: mult.commute) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
397 |
also have "... = \<integral>\<^sup>+x\<in>A. g x \<partial>M" by (rule nn_set_integral_cong[OF assms(1)]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
398 |
finally show "\<integral>\<^sup>+x\<in>A. g x \<partial>M = \<integral>\<^sup>+x\<in>A. nn_cond_exp M F f x \<partial>M" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
399 |
qed (auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
400 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
401 |
lemma nn_cond_exp_mono: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
402 |
assumes "AE x in M. f x \<le> g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
403 |
and [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
404 |
shows "AE x in M. nn_cond_exp M F f x \<le> nn_cond_exp M F g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
405 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
406 |
define h where "h = (\<lambda>x. g x - f x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
407 |
have [measurable]: "h \<in> borel_measurable M" unfolding h_def by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
408 |
have *: "AE x in M. g x = f x + h x" unfolding h_def using assms(1) by (auto simp: ennreal_ineq_diff_add) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
409 |
have "AE x in M. nn_cond_exp M F g x = nn_cond_exp M F (\<lambda>x. f x + h x) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
410 |
by (rule nn_cond_exp_cong) (auto simp add: * assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
411 |
moreover have "AE x in M. nn_cond_exp M F f x + nn_cond_exp M F h x = nn_cond_exp M F (\<lambda>x. f x + h x) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
412 |
by (rule nn_cond_exp_sum) (auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
413 |
ultimately have "AE x in M. nn_cond_exp M F g x = nn_cond_exp M F f x + nn_cond_exp M F h x" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
414 |
then show ?thesis by force |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
415 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
416 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
417 |
lemma nested_subalg_is_sigma_finite: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
418 |
assumes "subalgebra M G" "subalgebra G F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
419 |
shows "sigma_finite_subalgebra M G" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
420 |
unfolding sigma_finite_subalgebra_def |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
421 |
proof (auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
422 |
have "\<exists>A. countable A \<and> A \<subseteq> sets (restr_to_subalg M F) \<and> \<Union>A = space (restr_to_subalg M F) \<and> (\<forall>a\<in>A. emeasure (restr_to_subalg M F) a \<noteq> \<infinity>)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
423 |
using sigma_fin_subalg sigma_finite_measure_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
424 |
then obtain A where A:"countable A \<and> A \<subseteq> sets (restr_to_subalg M F) \<and> \<Union>A = space (restr_to_subalg M F) \<and> (\<forall>a\<in>A. emeasure (restr_to_subalg M F) a \<noteq> \<infinity>)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
425 |
by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
426 |
have "sets F \<subseteq> sets M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
427 |
by (meson assms order_trans subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
428 |
then have "countable A \<and> A \<subseteq> sets (restr_to_subalg M G) \<and> \<Union>A = space (restr_to_subalg M F) \<and> (\<forall>a\<in>A. emeasure (restr_to_subalg M G) a \<noteq> \<infinity>)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
429 |
by (metis (no_types) A assms basic_trans_rules(31) emeasure_restr_to_subalg order_trans sets_restr_to_subalg subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
430 |
then show "sigma_finite_measure (restr_to_subalg M G)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
431 |
by (metis sigma_finite_measure.intro space_restr_to_subalg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
432 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
433 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
434 |
lemma nn_cond_exp_nested_subalg: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
435 |
assumes "subalgebra M G" "subalgebra G F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
436 |
and [measurable]: "f \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
437 |
shows "AE x in M. nn_cond_exp M F f x = nn_cond_exp M F (nn_cond_exp M G f) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
438 |
proof (rule nn_cond_exp_charact, auto) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
439 |
interpret G: sigma_finite_subalgebra M G by (rule nested_subalg_is_sigma_finite[OF assms(1) assms(2)]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
440 |
fix A assume [measurable]: "A \<in> sets F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
441 |
then have [measurable]: "A \<in> sets G" using assms(2) by (meson set_mp subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
442 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
443 |
have "set_nn_integral M A (nn_cond_exp M G f) = (\<integral>\<^sup>+ x. indicator A x * nn_cond_exp M G f x\<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
444 |
by (metis (no_types) mult.commute) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
445 |
also have "... = (\<integral>\<^sup>+ x. indicator A x * f x \<partial>M)" by (rule G.nn_cond_exp_intg, auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
446 |
also have "... = (\<integral>\<^sup>+ x. indicator A x * nn_cond_exp M F f x \<partial>M)" by (rule nn_cond_exp_intg[symmetric], auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
447 |
also have "... = set_nn_integral M A (nn_cond_exp M F f)" by (metis (no_types) mult.commute) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
448 |
finally show "set_nn_integral M A (nn_cond_exp M G f) = set_nn_integral M A (nn_cond_exp M F f)". |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
449 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
450 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
451 |
end |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
452 |
|
67226 | 453 |
subsection \<open>Real conditional expectation\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
454 |
|
67226 | 455 |
text \<open>Once conditional expectations of positive functions are defined, the definition for real-valued functions |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
456 |
follows readily, by taking the difference of positive and negative parts. |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
457 |
One could also define a conditional expectation of vector-space valued functions, as in |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
458 |
\verb+Bochner_Integral+, but since the real-valued case is the most important, and quicker to formalize, I |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
459 |
concentrate on it. (It is also essential for the case of the most general Pettis integral.) |
67226 | 460 |
\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
461 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
462 |
definition real_cond_exp :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> ('a \<Rightarrow> real)" where |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
463 |
"real_cond_exp M F f = |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
464 |
(\<lambda>x. enn2real(nn_cond_exp M F (\<lambda>x. ennreal (f x)) x) - enn2real(nn_cond_exp M F (\<lambda>x. ennreal (-f x)) x))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
465 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
466 |
lemma |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
467 |
shows borel_measurable_cond_exp [measurable]: "real_cond_exp M F f \<in> borel_measurable F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
468 |
and borel_measurable_cond_exp2 [measurable]: "real_cond_exp M F f \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
469 |
unfolding real_cond_exp_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
470 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
471 |
context sigma_finite_subalgebra |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
472 |
begin |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
473 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
474 |
lemma real_cond_exp_abs: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
475 |
assumes [measurable]: "f \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
476 |
shows "AE x in M. abs(real_cond_exp M F f x) \<le> nn_cond_exp M F (\<lambda>x. ennreal (abs(f x))) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
477 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
478 |
define fp where "fp = (\<lambda>x. ennreal (f x))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
479 |
define fm where "fm = (\<lambda>x. ennreal (- f x))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
480 |
have [measurable]: "fp \<in> borel_measurable M" "fm \<in> borel_measurable M" unfolding fp_def fm_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
481 |
have eq: "\<And>x. ennreal \<bar>f x\<bar> = fp x + fm x" unfolding fp_def fm_def by (simp add: abs_real_def ennreal_neg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
482 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
483 |
{ |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
484 |
fix x assume H: "nn_cond_exp M F fp x + nn_cond_exp M F fm x = nn_cond_exp M F (\<lambda>x. fp x + fm x) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
485 |
have "\<bar>real_cond_exp M F f x\<bar> \<le> \<bar>enn2real(nn_cond_exp M F fp x)\<bar> + \<bar>enn2real(nn_cond_exp M F fm x)\<bar>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
486 |
unfolding real_cond_exp_def fp_def fm_def by (auto intro: abs_triangle_ineq4 simp del: enn2real_nonneg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
487 |
from ennreal_leI[OF this] |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
488 |
have "abs(real_cond_exp M F f x) \<le> nn_cond_exp M F fp x + nn_cond_exp M F fm x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
489 |
by simp (metis add.commute ennreal_enn2real le_iff_add not_le top_unique) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
490 |
also have "... = nn_cond_exp M F (\<lambda>x. fp x + fm x) x" using H by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
491 |
finally have "abs(real_cond_exp M F f x) \<le> nn_cond_exp M F (\<lambda>x. fp x + fm x) x" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
492 |
} |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
493 |
moreover have "AE x in M. nn_cond_exp M F fp x + nn_cond_exp M F fm x = nn_cond_exp M F (\<lambda>x. fp x + fm x) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
494 |
by (rule nn_cond_exp_sum) (auto simp add: fp_def fm_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
495 |
ultimately have "AE x in M. abs(real_cond_exp M F f x) \<le> nn_cond_exp M F (\<lambda>x. fp x + fm x) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
496 |
by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
497 |
then show ?thesis using eq by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
498 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
499 |
|
67226 | 500 |
text\<open>The next lemma shows that the conditional expectation |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
501 |
is an $F$-measurable function whose average against an $F$-measurable |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
502 |
function $f$ coincides with the average of the original function against $f$. |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
503 |
It is obtained as a consequence of the same property for the positive conditional |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
504 |
expectation, taking the difference of the positive and the negative part. The proof |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
505 |
is given first assuming $f \geq 0$ for simplicity, and then extended to the general case in |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
506 |
the subsequent lemma. The idea of the proof is essentially trivial, but the implementation |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
507 |
is slightly tedious as one should check all the integrability properties of the different parts, and go |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
508 |
back and forth between positive integral and signed integrals, and between real-valued |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
509 |
functions and ennreal-valued functions. |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
510 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
511 |
Once this lemma is available, we will use it to characterize the conditional expectation, |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
512 |
and never come back to the original technical definition, as we did in the case of the |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
513 |
nonnegative conditional expectation. |
67226 | 514 |
\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
515 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
516 |
lemma real_cond_exp_intg_fpos: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
517 |
assumes "integrable M (\<lambda>x. f x * g x)" and f_pos[simp]: "\<And>x. f x \<ge> 0" and |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
518 |
[measurable]: "f \<in> borel_measurable F" "g \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
519 |
shows "integrable M (\<lambda>x. f x * real_cond_exp M F g x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
520 |
"(\<integral> x. f x * real_cond_exp M F g x \<partial>M) = (\<integral> x. f x * g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
521 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
522 |
have [measurable]: "f \<in> borel_measurable M" by (rule measurable_from_subalg[OF subalg assms(3)]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
523 |
define gp where "gp = (\<lambda>x. ennreal (g x))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
524 |
define gm where "gm = (\<lambda>x. ennreal (- g x))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
525 |
have [measurable]: "gp \<in> borel_measurable M" "gm \<in> borel_measurable M" unfolding gp_def gm_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
526 |
define h where "h = (\<lambda>x. ennreal(abs(g x)))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
527 |
have hgpgm: "\<And>x. h x = gp x + gm x" unfolding gp_def gm_def h_def by (simp add: abs_real_def ennreal_neg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
528 |
have [measurable]: "h \<in> borel_measurable M" unfolding h_def by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
529 |
have pos[simp]: "\<And>x. h x \<ge> 0" "\<And>x. gp x \<ge> 0" "\<And>x. gm x \<ge> 0" unfolding h_def gp_def gm_def by simp_all |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
530 |
have gp_real: "\<And>x. enn2real(gp x) = max (g x) 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
531 |
unfolding gp_def by (simp add: max_def ennreal_neg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
532 |
have gm_real: "\<And>x. enn2real(gm x) = max (-g x) 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
533 |
unfolding gm_def by (simp add: max_def ennreal_neg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
534 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
535 |
have "(\<integral>\<^sup>+ x. norm(f x * max (g x) 0) \<partial>M) \<le> (\<integral>\<^sup>+ x. norm(f x * g x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
536 |
by (simp add: nn_integral_mono) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
537 |
also have "... < \<infinity>" using assms(1) by (simp add: integrable_iff_bounded) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
538 |
finally have "(\<integral>\<^sup>+ x. norm(f x * max (g x) 0) \<partial>M) < \<infinity>" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
539 |
then have int1: "integrable M (\<lambda>x. f x * max (g x) 0)" by (simp add: integrableI_bounded) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
540 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
541 |
have "(\<integral>\<^sup>+ x. norm(f x * max (-g x) 0) \<partial>M) \<le> (\<integral>\<^sup>+ x. norm(f x * g x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
542 |
by (simp add: nn_integral_mono) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
543 |
also have "... < \<infinity>" using assms(1) by (simp add: integrable_iff_bounded) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
544 |
finally have "(\<integral>\<^sup>+ x. norm(f x * max (-g x) 0) \<partial>M) < \<infinity>" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
545 |
then have int2: "integrable M (\<lambda>x. f x * max (-g x) 0)" by (simp add: integrableI_bounded) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
546 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
547 |
have "(\<integral>\<^sup>+x. f x * nn_cond_exp M F h x \<partial>M) = (\<integral>\<^sup>+x. f x * h x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
548 |
by (rule nn_cond_exp_intg) auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
549 |
also have "\<dots> = \<integral>\<^sup>+ x. ennreal (f x * max (g x) 0 + f x * max (- g x) 0) \<partial>M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
550 |
unfolding h_def |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
551 |
by (intro nn_integral_cong)(auto simp: ennreal_mult[symmetric] abs_mult split: split_max) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
552 |
also have "... < \<infinity>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
553 |
using Bochner_Integration.integrable_add[OF int1 int2, THEN integrableD(2)] by (auto simp add: less_top) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
554 |
finally have *: "(\<integral>\<^sup>+x. f x * nn_cond_exp M F h x \<partial>M) < \<infinity>" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
555 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
556 |
have "(\<integral>\<^sup>+x. norm(f x * real_cond_exp M F g x) \<partial>M) = (\<integral>\<^sup>+x. f x * abs(real_cond_exp M F g x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
557 |
by (simp add: abs_mult) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
558 |
also have "... \<le> (\<integral>\<^sup>+x. f x * nn_cond_exp M F h x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
559 |
proof (rule nn_integral_mono_AE) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
560 |
{ |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
561 |
fix x assume *: "abs(real_cond_exp M F g x) \<le> nn_cond_exp M F h x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
562 |
have "ennreal (f x * \<bar>real_cond_exp M F g x\<bar>) = f x * ennreal(\<bar>real_cond_exp M F g x\<bar>)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
563 |
by (simp add: ennreal_mult) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
564 |
also have "... \<le> f x * nn_cond_exp M F h x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
565 |
using * by (auto intro!: mult_left_mono) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
566 |
finally have "ennreal (f x * \<bar>real_cond_exp M F g x\<bar>) \<le> f x * nn_cond_exp M F h x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
567 |
by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
568 |
} |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
569 |
then show "AE x in M. ennreal (f x * \<bar>real_cond_exp M F g x\<bar>) \<le> f x * nn_cond_exp M F h x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
570 |
using real_cond_exp_abs[OF assms(4)] h_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
571 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
572 |
finally have **: "(\<integral>\<^sup>+x. norm(f x * real_cond_exp M F g x) \<partial>M) < \<infinity>" using * by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
573 |
show "integrable M (\<lambda>x. f x * real_cond_exp M F g x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
574 |
using ** by (intro integrableI_bounded) auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
575 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
576 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
577 |
have "(\<integral>\<^sup>+x. f x * nn_cond_exp M F gp x \<partial>M) \<le> (\<integral>\<^sup>+x. f x * nn_cond_exp M F h x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
578 |
proof (rule nn_integral_mono_AE) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
579 |
have "AE x in M. nn_cond_exp M F gp x \<le> nn_cond_exp M F h x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
580 |
by (rule nn_cond_exp_mono) (auto simp add: hgpgm) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
581 |
then show "AE x in M. f x * nn_cond_exp M F gp x \<le> f x * nn_cond_exp M F h x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
582 |
by (auto simp: mult_left_mono) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
583 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
584 |
then have a: "(\<integral>\<^sup>+x. f x * nn_cond_exp M F gp x \<partial>M) < \<infinity>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
585 |
using * by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
586 |
have "ennreal(norm(f x * enn2real(nn_cond_exp M F gp x))) \<le> f x * nn_cond_exp M F gp x" for x |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
587 |
by (auto simp add: ennreal_mult intro!: mult_left_mono) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
588 |
(metis enn2real_ennreal enn2real_nonneg le_cases le_ennreal_iff) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
589 |
then have "(\<integral>\<^sup>+x. norm(f x * enn2real(nn_cond_exp M F gp x)) \<partial>M) \<le> (\<integral>\<^sup>+x. f x * nn_cond_exp M F gp x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
590 |
by (simp add: nn_integral_mono) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
591 |
then have "(\<integral>\<^sup>+x. norm(f x * enn2real(nn_cond_exp M F gp x)) \<partial>M) < \<infinity>" using a by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
592 |
then have gp_int: "integrable M (\<lambda>x. f x * enn2real(nn_cond_exp M F gp x))" by (simp add: integrableI_bounded) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
593 |
have gp_fin: "AE x in M. f x * nn_cond_exp M F gp x \<noteq> \<infinity>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
594 |
apply (rule nn_integral_PInf_AE) using a by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
595 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
596 |
have "(\<integral> x. f x * enn2real(nn_cond_exp M F gp x) \<partial>M) = enn2real (\<integral>\<^sup>+ x. f x * enn2real(nn_cond_exp M F gp x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
597 |
by (rule integral_eq_nn_integral) auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
598 |
also have "... = enn2real(\<integral>\<^sup>+ x. ennreal(f x * enn2real(gp x)) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
599 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
600 |
{ |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
601 |
fix x assume "f x * nn_cond_exp M F gp x \<noteq> \<infinity>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
602 |
then have "ennreal (f x * enn2real (nn_cond_exp M F gp x)) = ennreal (f x) * nn_cond_exp M F gp x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
603 |
by (auto simp add: ennreal_mult ennreal_mult_eq_top_iff less_top intro!: ennreal_mult_left_cong) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
604 |
} |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
605 |
then have "AE x in M. ennreal (f x * enn2real (nn_cond_exp M F gp x)) = ennreal (f x) * nn_cond_exp M F gp x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
606 |
using gp_fin by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
607 |
then have "(\<integral>\<^sup>+ x. f x * enn2real(nn_cond_exp M F gp x) \<partial>M) = (\<integral>\<^sup>+ x. f x * nn_cond_exp M F gp x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
608 |
by (rule nn_integral_cong_AE) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
609 |
also have "... = (\<integral>\<^sup>+ x. f x * gp x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
610 |
by (rule nn_cond_exp_intg) (auto simp add: gp_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
611 |
also have "... = (\<integral>\<^sup>+ x. ennreal(f x * enn2real(gp x)) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
612 |
by (rule nn_integral_cong_AE) (auto simp: ennreal_mult gp_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
613 |
finally have "(\<integral>\<^sup>+ x. f x * enn2real(nn_cond_exp M F gp x) \<partial>M) = (\<integral>\<^sup>+ x. ennreal(f x * enn2real(gp x)) \<partial>M)" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
614 |
then show ?thesis by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
615 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
616 |
also have "... = (\<integral> x. f x * enn2real(gp x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
617 |
by (rule integral_eq_nn_integral[symmetric]) (auto simp add: gp_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
618 |
finally have gp_expr: "(\<integral> x. f x * enn2real(nn_cond_exp M F gp x) \<partial>M) = (\<integral> x. f x * enn2real(gp x) \<partial>M)" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
619 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
620 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
621 |
have "(\<integral>\<^sup>+x. f x * nn_cond_exp M F gm x \<partial>M) \<le> (\<integral>\<^sup>+x. f x * nn_cond_exp M F h x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
622 |
proof (rule nn_integral_mono_AE) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
623 |
have "AE x in M. nn_cond_exp M F gm x \<le> nn_cond_exp M F h x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
624 |
by (rule nn_cond_exp_mono) (auto simp add: hgpgm) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
625 |
then show "AE x in M. f x * nn_cond_exp M F gm x \<le> f x * nn_cond_exp M F h x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
626 |
by (auto simp: mult_left_mono) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
627 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
628 |
then have a: "(\<integral>\<^sup>+x. f x * nn_cond_exp M F gm x \<partial>M) < \<infinity>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
629 |
using * by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
630 |
have "\<And>x. ennreal(norm(f x * enn2real(nn_cond_exp M F gm x))) \<le> f x * nn_cond_exp M F gm x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
631 |
by (auto simp add: ennreal_mult intro!: mult_left_mono) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
632 |
(metis enn2real_ennreal enn2real_nonneg le_cases le_ennreal_iff) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
633 |
then have "(\<integral>\<^sup>+x. norm(f x * enn2real(nn_cond_exp M F gm x)) \<partial>M) \<le> (\<integral>\<^sup>+x. f x * nn_cond_exp M F gm x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
634 |
by (simp add: nn_integral_mono) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
635 |
then have "(\<integral>\<^sup>+x. norm(f x * enn2real(nn_cond_exp M F gm x)) \<partial>M) < \<infinity>" using a by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
636 |
then have gm_int: "integrable M (\<lambda>x. f x * enn2real(nn_cond_exp M F gm x))" by (simp add: integrableI_bounded) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
637 |
have gm_fin: "AE x in M. f x * nn_cond_exp M F gm x \<noteq> \<infinity>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
638 |
apply (rule nn_integral_PInf_AE) using a by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
639 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
640 |
have "(\<integral> x. f x * enn2real(nn_cond_exp M F gm x) \<partial>M) = enn2real (\<integral>\<^sup>+ x. f x * enn2real(nn_cond_exp M F gm x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
641 |
by (rule integral_eq_nn_integral) auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
642 |
also have "... = enn2real(\<integral>\<^sup>+ x. ennreal(f x * enn2real(gm x)) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
643 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
644 |
{ |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
645 |
fix x assume "f x * nn_cond_exp M F gm x \<noteq> \<infinity>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
646 |
then have "ennreal (f x * enn2real (nn_cond_exp M F gm x)) = ennreal (f x) * nn_cond_exp M F gm x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
647 |
by (auto simp add: ennreal_mult ennreal_mult_eq_top_iff less_top intro!: ennreal_mult_left_cong) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
648 |
} |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
649 |
then have "AE x in M. ennreal (f x * enn2real (nn_cond_exp M F gm x)) = ennreal (f x) * nn_cond_exp M F gm x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
650 |
using gm_fin by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
651 |
then have "(\<integral>\<^sup>+ x. f x * enn2real(nn_cond_exp M F gm x) \<partial>M) = (\<integral>\<^sup>+ x. f x * nn_cond_exp M F gm x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
652 |
by (rule nn_integral_cong_AE) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
653 |
also have "... = (\<integral>\<^sup>+ x. f x * gm x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
654 |
by (rule nn_cond_exp_intg) (auto simp add: gm_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
655 |
also have "... = (\<integral>\<^sup>+ x. ennreal(f x * enn2real(gm x)) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
656 |
by (rule nn_integral_cong_AE) (auto simp: ennreal_mult gm_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
657 |
finally have "(\<integral>\<^sup>+ x. f x * enn2real(nn_cond_exp M F gm x) \<partial>M) = (\<integral>\<^sup>+ x. ennreal(f x * enn2real(gm x)) \<partial>M)" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
658 |
then show ?thesis by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
659 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
660 |
also have "... = (\<integral> x. f x * enn2real(gm x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
661 |
by (rule integral_eq_nn_integral[symmetric]) (auto simp add: gm_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
662 |
finally have gm_expr: "(\<integral> x. f x * enn2real(nn_cond_exp M F gm x) \<partial>M) = (\<integral> x. f x * enn2real(gm x) \<partial>M)" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
663 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
664 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
665 |
have "(\<integral> x. f x * real_cond_exp M F g x \<partial>M) = (\<integral> x. f x * enn2real(nn_cond_exp M F gp x) - f x * enn2real(nn_cond_exp M F gm x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
666 |
unfolding real_cond_exp_def gp_def gm_def by (simp add: right_diff_distrib) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
667 |
also have "... = (\<integral> x. f x * enn2real(nn_cond_exp M F gp x) \<partial>M) - (\<integral> x. f x * enn2real(nn_cond_exp M F gm x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
668 |
by (rule Bochner_Integration.integral_diff) (simp_all add: gp_int gm_int) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
669 |
also have "... = (\<integral> x. f x * enn2real(gp x) \<partial>M) - (\<integral> x. f x * enn2real(gm x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
670 |
using gp_expr gm_expr by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
671 |
also have "... = (\<integral> x. f x * max (g x) 0 \<partial>M) - (\<integral> x. f x * max (-g x) 0 \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
672 |
using gp_real gm_real by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
673 |
also have "... = (\<integral> x. f x * max (g x) 0 - f x * max (-g x) 0 \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
674 |
by (rule Bochner_Integration.integral_diff[symmetric]) (simp_all add: int1 int2) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
675 |
also have "... = (\<integral>x. f x * g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
676 |
by (metis (mono_tags, hide_lams) diff_0 diff_zero eq_iff max.cobounded2 max_def minus_minus neg_le_0_iff_le right_diff_distrib) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
677 |
finally show "(\<integral> x. f x * real_cond_exp M F g x \<partial>M) = (\<integral>x. f x * g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
678 |
by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
679 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
680 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
681 |
lemma real_cond_exp_intg: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
682 |
assumes "integrable M (\<lambda>x. f x * g x)" and |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
683 |
[measurable]: "f \<in> borel_measurable F" "g \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
684 |
shows "integrable M (\<lambda>x. f x * real_cond_exp M F g x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
685 |
"(\<integral> x. f x * real_cond_exp M F g x \<partial>M) = (\<integral> x. f x * g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
686 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
687 |
have [measurable]: "f \<in> borel_measurable M" by (rule measurable_from_subalg[OF subalg assms(2)]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
688 |
define fp where "fp = (\<lambda>x. max (f x) 0)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
689 |
define fm where "fm = (\<lambda>x. max (-f x) 0)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
690 |
have [measurable]: "fp \<in> borel_measurable M" "fm \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
691 |
unfolding fp_def fm_def by simp_all |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
692 |
have [measurable]: "fp \<in> borel_measurable F" "fm \<in> borel_measurable F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
693 |
unfolding fp_def fm_def by simp_all |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
694 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
695 |
have "(\<integral>\<^sup>+ x. norm(fp x * g x) \<partial>M) \<le> (\<integral>\<^sup>+ x. norm(f x * g x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
696 |
by (simp add: fp_def nn_integral_mono) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
697 |
also have "... < \<infinity>" using assms(1) by (simp add: integrable_iff_bounded) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
698 |
finally have "(\<integral>\<^sup>+ x. norm(fp x * g x) \<partial>M) < \<infinity>" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
699 |
then have intp: "integrable M (\<lambda>x. fp x * g x)" by (simp add: integrableI_bounded) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
700 |
moreover have "\<And>x. fp x \<ge> 0" unfolding fp_def by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
701 |
ultimately have Rp: "integrable M (\<lambda>x. fp x * real_cond_exp M F g x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
702 |
"(\<integral> x. fp x * real_cond_exp M F g x \<partial>M) = (\<integral> x. fp x * g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
703 |
using real_cond_exp_intg_fpos by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
704 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
705 |
have "(\<integral>\<^sup>+ x. norm(fm x * g x) \<partial>M) \<le> (\<integral>\<^sup>+ x. norm(f x * g x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
706 |
by (simp add: fm_def nn_integral_mono) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
707 |
also have "... < \<infinity>" using assms(1) by (simp add: integrable_iff_bounded) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
708 |
finally have "(\<integral>\<^sup>+ x. norm(fm x * g x) \<partial>M) < \<infinity>" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
709 |
then have intm: "integrable M (\<lambda>x. fm x * g x)" by (simp add: integrableI_bounded) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
710 |
moreover have "\<And>x. fm x \<ge> 0" unfolding fm_def by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
711 |
ultimately have Rm: "integrable M (\<lambda>x. fm x * real_cond_exp M F g x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
712 |
"(\<integral> x. fm x * real_cond_exp M F g x \<partial>M) = (\<integral> x. fm x * g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
713 |
using real_cond_exp_intg_fpos by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
714 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
715 |
have "integrable M (\<lambda>x. fp x * real_cond_exp M F g x - fm x * real_cond_exp M F g x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
716 |
using Rp(1) Rm(1) integrable_diff by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
717 |
moreover have *: "\<And>x. f x * real_cond_exp M F g x = fp x * real_cond_exp M F g x - fm x * real_cond_exp M F g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
718 |
unfolding fp_def fm_def by (simp add: max_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
719 |
ultimately show "integrable M (\<lambda>x. f x * real_cond_exp M F g x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
720 |
by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
721 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
722 |
have "(\<integral> x. f x * real_cond_exp M F g x \<partial>M) = (\<integral> x. fp x * real_cond_exp M F g x - fm x * real_cond_exp M F g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
723 |
using * by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
724 |
also have "... = (\<integral> x. fp x * real_cond_exp M F g x \<partial>M) - (\<integral> x. fm x * real_cond_exp M F g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
725 |
using Rp(1) Rm(1) by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
726 |
also have "... = (\<integral> x. fp x * g x \<partial>M) - (\<integral> x. fm x * g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
727 |
using Rp(2) Rm(2) by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
728 |
also have "... = (\<integral> x. fp x * g x - fm x * g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
729 |
using intm intp by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
730 |
also have "... = (\<integral> x. f x * g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
731 |
unfolding fp_def fm_def by (metis (no_types, hide_lams) diff_0 diff_zero max.commute |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
732 |
max_def minus_minus mult.commute neg_le_iff_le right_diff_distrib) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
733 |
finally show "(\<integral> x. f x * real_cond_exp M F g x \<partial>M) = (\<integral> x. f x * g x \<partial>M)" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
734 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
735 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
736 |
lemma real_cond_exp_intA: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
737 |
assumes [measurable]: "integrable M f" "A \<in> sets F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
738 |
shows "(\<integral> x \<in> A. f x \<partial>M) = (\<integral>x \<in> A. real_cond_exp M F f x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
739 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
740 |
have "A \<in> sets M" by (meson assms(2) subalg subalgebra_def subsetD) |
67226 | 741 |
have "integrable M (\<lambda>x. indicator A x * f x)" using integrable_mult_indicator[OF \<open>A \<in> sets M\<close> assms(1)] by auto |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
742 |
then show ?thesis using real_cond_exp_intg(2)[where ?f = "indicator A" and ?g = f, symmetric] |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
743 |
unfolding set_lebesgue_integral_def by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
744 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
745 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
746 |
lemma real_cond_exp_int [intro]: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
747 |
assumes "integrable M f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
748 |
shows "integrable M (real_cond_exp M F f)" "(\<integral>x. real_cond_exp M F f x \<partial>M) = (\<integral>x. f x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
749 |
using real_cond_exp_intg[where ?f = "\<lambda>x. 1" and ?g = f] assms by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
750 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
751 |
lemma real_cond_exp_charact: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
752 |
assumes "\<And>A. A \<in> sets F \<Longrightarrow> (\<integral> x \<in> A. f x \<partial>M) = (\<integral> x \<in> A. g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
753 |
and [measurable]: "integrable M f" "integrable M g" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
754 |
"g \<in> borel_measurable F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
755 |
shows "AE x in M. real_cond_exp M F f x = g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
756 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
757 |
let ?MF = "restr_to_subalg M F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
758 |
have "AE x in ?MF. real_cond_exp M F f x = g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
759 |
proof (rule AE_symmetric[OF density_unique_real]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
760 |
fix A assume "A \<in> sets ?MF" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
761 |
then have [measurable]: "A \<in> sets F" using sets_restr_to_subalg[OF subalg] by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
762 |
then have a [measurable]: "A \<in> sets M" by (meson subalg subalgebra_def subsetD) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
763 |
have "(\<integral>x \<in> A. g x \<partial> ?MF) = (\<integral>x \<in> A. g x \<partial>M)" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
764 |
unfolding set_lebesgue_integral_def by (simp add: integral_subalgebra2 subalg) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
765 |
also have "... = (\<integral>x \<in> A. f x \<partial>M)" using assms(1) by simp |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
766 |
also have "... = (\<integral>x. indicator A x * f x \<partial>M)" by (simp add: mult.commute set_lebesgue_integral_def) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
767 |
also have "... = (\<integral>x. indicator A x * real_cond_exp M F f x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
768 |
apply (rule real_cond_exp_intg(2)[symmetric]) using integrable_mult_indicator[OF a assms(2)] by (auto simp add: assms) |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
769 |
also have "... = (\<integral>x \<in> A. real_cond_exp M F f x \<partial>M)" by (simp add: mult.commute set_lebesgue_integral_def) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
770 |
also have "... = (\<integral>x \<in> A. real_cond_exp M F f x \<partial> ?MF)" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
771 |
by (simp add: integral_subalgebra2 subalg set_lebesgue_integral_def) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
772 |
finally show "(\<integral>x \<in> A. g x \<partial> ?MF) = (\<integral>x \<in> A. real_cond_exp M F f x \<partial> ?MF)" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
773 |
next |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
774 |
have "integrable M (real_cond_exp M F f)" by (rule real_cond_exp_int(1)[OF assms(2)]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
775 |
then show "integrable ?MF (real_cond_exp M F f)" by (metis borel_measurable_cond_exp integrable_in_subalg[OF subalg]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
776 |
show "integrable (restr_to_subalg M F) g" by (simp add: assms(3) integrable_in_subalg[OF subalg]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
777 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
778 |
then show ?thesis using AE_restr_to_subalg[OF subalg] by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
779 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
780 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
781 |
lemma real_cond_exp_F_meas [intro, simp]: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
782 |
assumes "integrable M f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
783 |
"f \<in> borel_measurable F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
784 |
shows "AE x in M. real_cond_exp M F f x = f x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
785 |
by (rule real_cond_exp_charact, auto simp add: assms measurable_from_subalg[OF subalg]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
786 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
787 |
lemma real_cond_exp_mult: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
788 |
assumes [measurable]:"f \<in> borel_measurable F" "g \<in> borel_measurable M" "integrable M (\<lambda>x. f x * g x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
789 |
shows "AE x in M. real_cond_exp M F (\<lambda>x. f x * g x) x = f x * real_cond_exp M F g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
790 |
proof (rule real_cond_exp_charact) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
791 |
fix A assume "A \<in> sets F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
792 |
then have [measurable]: "(\<lambda>x. f x * indicator A x) \<in> borel_measurable F" by measurable |
67226 | 793 |
have [measurable]: "A \<in> sets M" using subalg by (meson \<open>A \<in> sets F\<close> subalgebra_def subsetD) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
794 |
have "\<integral>x\<in>A. (f x * g x) \<partial>M = \<integral>x. (f x * indicator A x) * g x \<partial>M" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
795 |
by (simp add: mult.commute mult.left_commute set_lebesgue_integral_def) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
796 |
also have "... = \<integral>x. (f x * indicator A x) * real_cond_exp M F g x \<partial>M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
797 |
apply (rule real_cond_exp_intg(2)[symmetric], auto simp add: assms) |
67226 | 798 |
using integrable_mult_indicator[OF \<open>A \<in> sets M\<close> assms(3)] by (simp add: mult.commute mult.left_commute) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
799 |
also have "... = \<integral>x\<in>A. (f x * real_cond_exp M F g x)\<partial>M" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
800 |
by (simp add: mult.commute mult.left_commute set_lebesgue_integral_def) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
801 |
finally show "\<integral>x\<in>A. (f x * g x) \<partial>M = \<integral>x\<in>A. (f x * real_cond_exp M F g x)\<partial>M" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
802 |
qed (auto simp add: real_cond_exp_intg(1) assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
803 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
804 |
lemma real_cond_exp_add [intro]: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
805 |
assumes [measurable]: "integrable M f" "integrable M g" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
806 |
shows "AE x in M. real_cond_exp M F (\<lambda>x. f x + g x) x = real_cond_exp M F f x + real_cond_exp M F g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
807 |
proof (rule real_cond_exp_charact) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
808 |
have "integrable M (real_cond_exp M F f)" "integrable M (real_cond_exp M F g)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
809 |
using real_cond_exp_int(1) assms by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
810 |
then show "integrable M (\<lambda>x. real_cond_exp M F f x + real_cond_exp M F g x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
811 |
by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
812 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
813 |
fix A assume [measurable]: "A \<in> sets F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
814 |
then have "A \<in> sets M" by (meson subalg subalgebra_def subsetD) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
815 |
have intAf: "integrable M (\<lambda>x. indicator A x * f x)" |
67226 | 816 |
using integrable_mult_indicator[OF \<open>A \<in> sets M\<close> assms(1)] by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
817 |
have intAg: "integrable M (\<lambda>x. indicator A x * g x)" |
67226 | 818 |
using integrable_mult_indicator[OF \<open>A \<in> sets M\<close> assms(2)] by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
819 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
820 |
have "\<integral>x\<in>A. (real_cond_exp M F f x + real_cond_exp M F g x)\<partial>M = (\<integral>x\<in>A. real_cond_exp M F f x \<partial>M) + (\<integral>x\<in>A. real_cond_exp M F g x \<partial>M)" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
821 |
apply (rule set_integral_add, auto simp add: assms set_integrable_def) |
67226 | 822 |
using integrable_mult_indicator[OF \<open>A \<in> sets M\<close> real_cond_exp_int(1)[OF assms(1)]] |
823 |
integrable_mult_indicator[OF \<open>A \<in> sets M\<close> real_cond_exp_int(1)[OF assms(2)]] by simp_all |
|
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
824 |
also have "... = (\<integral>x. indicator A x * real_cond_exp M F f x \<partial>M) + (\<integral>x. indicator A x * real_cond_exp M F g x \<partial>M)" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
825 |
unfolding set_lebesgue_integral_def by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
826 |
also have "... = (\<integral>x. indicator A x * f x \<partial>M) + (\<integral>x. indicator A x * g x \<partial>M)" |
67226 | 827 |
using real_cond_exp_intg(2) assms \<open>A \<in> sets F\<close> intAf intAg by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
828 |
also have "... = (\<integral>x\<in>A. f x \<partial>M) + (\<integral>x\<in>A. g x \<partial>M)" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
829 |
unfolding set_lebesgue_integral_def by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
830 |
also have "... = \<integral>x\<in>A. (f x + g x)\<partial>M" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
831 |
by (rule set_integral_add(2)[symmetric]) (auto simp add: assms set_integrable_def \<open>A \<in> sets M\<close> intAf intAg) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
832 |
finally show "\<integral>x\<in>A. (f x + g x)\<partial>M = \<integral>x\<in>A. (real_cond_exp M F f x + real_cond_exp M F g x)\<partial>M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
833 |
by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
834 |
qed (auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
835 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
836 |
lemma real_cond_exp_cong: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
837 |
assumes ae: "AE x in M. f x = g x" and [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
838 |
shows "AE x in M. real_cond_exp M F f x = real_cond_exp M F g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
839 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
840 |
have "AE x in M. nn_cond_exp M F (\<lambda>x. ennreal (f x)) x = nn_cond_exp M F (\<lambda>x. ennreal (g x)) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
841 |
apply (rule nn_cond_exp_cong) using assms by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
842 |
moreover have "AE x in M. nn_cond_exp M F (\<lambda>x. ennreal (-f x)) x = nn_cond_exp M F (\<lambda>x. ennreal(-g x)) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
843 |
apply (rule nn_cond_exp_cong) using assms by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
844 |
ultimately show "AE x in M. real_cond_exp M F f x = real_cond_exp M F g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
845 |
unfolding real_cond_exp_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
846 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
847 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
848 |
lemma real_cond_exp_cmult [intro, simp]: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
849 |
fixes c::real |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
850 |
assumes "integrable M f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
851 |
shows "AE x in M. real_cond_exp M F (\<lambda>x. c * f x) x = c * real_cond_exp M F f x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
852 |
by (rule real_cond_exp_mult[where ?f = "\<lambda>x. c" and ?g = f], auto simp add: assms borel_measurable_integrable) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
853 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
854 |
lemma real_cond_exp_cdiv [intro, simp]: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
855 |
fixes c::real |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
856 |
assumes "integrable M f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
857 |
shows "AE x in M. real_cond_exp M F (\<lambda>x. f x / c) x = real_cond_exp M F f x / c" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
858 |
using real_cond_exp_cmult[of _ "1/c", OF assms] by (auto simp add: divide_simps) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
859 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
860 |
lemma real_cond_exp_diff [intro, simp]: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
861 |
assumes [measurable]: "integrable M f" "integrable M g" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
862 |
shows "AE x in M. real_cond_exp M F (\<lambda>x. f x - g x) x = real_cond_exp M F f x - real_cond_exp M F g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
863 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
864 |
have "AE x in M. real_cond_exp M F (\<lambda>x. f x + (- g x)) x = real_cond_exp M F f x + real_cond_exp M F (\<lambda>x. -g x) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
865 |
using real_cond_exp_add[where ?f = f and ?g = "\<lambda>x. - g x"] assms by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
866 |
moreover have "AE x in M. real_cond_exp M F (\<lambda>x. -g x) x = - real_cond_exp M F g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
867 |
using real_cond_exp_cmult[where ?f = g and ?c = "-1"] assms(2) by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
868 |
ultimately show ?thesis by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
869 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
870 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
871 |
lemma real_cond_exp_pos [intro]: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
872 |
assumes "AE x in M. f x \<ge> 0" and [measurable]: "f \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
873 |
shows "AE x in M. real_cond_exp M F f x \<ge> 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
874 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
875 |
define g where "g = (\<lambda>x. max (f x) 0)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
876 |
have "AE x in M. f x = g x" using assms g_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
877 |
then have *: "AE x in M. real_cond_exp M F f x = real_cond_exp M F g x" using real_cond_exp_cong g_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
878 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
879 |
have "\<And>x. g x \<ge> 0" unfolding g_def by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
880 |
then have "(\<lambda>x. ennreal(-g x)) = (\<lambda>x. 0)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
881 |
by (simp add: ennreal_neg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
882 |
moreover have "AE x in M. 0 = nn_cond_exp M F (\<lambda>x. 0) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
883 |
by (rule nn_cond_exp_F_meas, auto) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
884 |
ultimately have "AE x in M. nn_cond_exp M F (\<lambda>x. ennreal(-g x)) x = 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
885 |
by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
886 |
then have "AE x in M. real_cond_exp M F g x = enn2real(nn_cond_exp M F (\<lambda>x. ennreal (g x)) x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
887 |
unfolding real_cond_exp_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
888 |
then have "AE x in M. real_cond_exp M F g x \<ge> 0" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
889 |
then show ?thesis using * by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
890 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
891 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
892 |
lemma real_cond_exp_mono: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
893 |
assumes "AE x in M. f x \<le> g x" and [measurable]: "integrable M f" "integrable M g" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
894 |
shows "AE x in M. real_cond_exp M F f x \<le> real_cond_exp M F g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
895 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
896 |
have "AE x in M. real_cond_exp M F g x - real_cond_exp M F f x = real_cond_exp M F (\<lambda>x. g x - f x) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
897 |
by (rule AE_symmetric[OF real_cond_exp_diff], auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
898 |
moreover have "AE x in M. real_cond_exp M F (\<lambda>x. g x - f x) x \<ge> 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
899 |
by (rule real_cond_exp_pos, auto simp add: assms(1)) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
900 |
ultimately have "AE x in M. real_cond_exp M F g x - real_cond_exp M F f x \<ge> 0" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
901 |
then show ?thesis by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
902 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
903 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
904 |
lemma (in -) measurable_P_restriction [measurable (raw)]: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
905 |
assumes [measurable]: "Measurable.pred M P" "A \<in> sets M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
906 |
shows "{x \<in> A. P x} \<in> sets M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
907 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
908 |
have "A \<subseteq> space M" using sets.sets_into_space[OF assms(2)]. |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
909 |
then have "{x \<in> A. P x} = A \<inter> {x \<in> space M. P x}" by blast |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
910 |
then show ?thesis by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
911 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
912 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
913 |
lemma real_cond_exp_gr_c: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
914 |
assumes [measurable]: "integrable M f" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
915 |
and AE: "AE x in M. f x > c" |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
916 |
shows "AE x in M. real_cond_exp M F f x > c" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
917 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
918 |
define X where "X = {x \<in> space M. real_cond_exp M F f x \<le> c}" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
919 |
have [measurable]: "X \<in> sets F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
920 |
unfolding X_def apply measurable by (metis sets.top subalg subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
921 |
then have [measurable]: "X \<in> sets M" using sets_restr_to_subalg subalg subalgebra_def by blast |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
922 |
have "emeasure M X = 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
923 |
proof (rule ccontr) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
924 |
assume "\<not>(emeasure M X) = 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
925 |
have "emeasure (restr_to_subalg M F) X = emeasure M X" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
926 |
by (simp add: emeasure_restr_to_subalg subalg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
927 |
then have "emeasure (restr_to_subalg M F) X > 0" |
67226 | 928 |
using \<open>\<not>(emeasure M X) = 0\<close> gr_zeroI by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
929 |
then obtain A where "A \<in> sets (restr_to_subalg M F)" "A \<subseteq> X" "emeasure (restr_to_subalg M F) A > 0" "emeasure (restr_to_subalg M F) A < \<infinity>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
930 |
using sigma_fin_subalg by (metis emeasure_notin_sets ennreal_0 infinity_ennreal_def le_less_linear neq_top_trans |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
931 |
not_gr_zero order_refl sigma_finite_measure.approx_PInf_emeasure_with_finite) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
932 |
then have [measurable]: "A \<in> sets F" using subalg sets_restr_to_subalg by blast |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
933 |
then have [measurable]: "A \<in> sets M" using sets_restr_to_subalg subalg subalgebra_def by blast |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
934 |
have Ic: "set_integrable M A (\<lambda>x. c)" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
935 |
unfolding set_integrable_def |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
936 |
using \<open>emeasure (restr_to_subalg M F) A < \<infinity>\<close> emeasure_restr_to_subalg subalg by fastforce |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
937 |
have If: "set_integrable M A f" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
938 |
unfolding set_integrable_def |
67226 | 939 |
by (rule integrable_mult_indicator, auto simp add: \<open>integrable M f\<close>) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
940 |
have "AE x in M. indicator A x * c = indicator A x * f x" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
941 |
proof (rule integral_ineq_eq_0_then_AE) |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
942 |
have "(\<integral>x\<in>A. c \<partial>M) = (\<integral>x\<in>A. f x \<partial>M)" |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
943 |
proof (rule antisym) |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
944 |
show "(\<integral>x\<in>A. c \<partial>M) \<le> (\<integral>x\<in>A. f x \<partial>M)" |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
945 |
apply (rule set_integral_mono_AE) using Ic If assms(2) by auto |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
946 |
have "(\<integral>x\<in>A. f x \<partial>M) = (\<integral>x\<in>A. real_cond_exp M F f x \<partial>M)" |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
947 |
by (rule real_cond_exp_intA, auto simp add: \<open>integrable M f\<close>) |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
948 |
also have "... \<le> (\<integral>x\<in>A. c \<partial>M)" |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
949 |
apply (rule set_integral_mono) |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
950 |
unfolding set_integrable_def |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
951 |
apply (rule integrable_mult_indicator, simp, simp add: real_cond_exp_int(1)[OF \<open>integrable M f\<close>]) |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
952 |
using Ic X_def \<open>A \<subseteq> X\<close> by (auto simp: set_integrable_def) |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
953 |
finally show "(\<integral>x\<in>A. f x \<partial>M) \<le> (\<integral>x\<in>A. c \<partial>M)" by simp |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
954 |
qed |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
955 |
then have "measure M A * c = LINT x|M. indicat_real A x * f x" |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
956 |
by (auto simp: set_lebesgue_integral_def) |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
957 |
then show "LINT x|M. indicat_real A x * c = LINT x|M. indicat_real A x * f x" |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
958 |
by auto |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
959 |
show "AE x in M. indicat_real A x * c \<le> indicat_real A x * f x" |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
960 |
using AE unfolding indicator_def by auto |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
961 |
qed (use Ic If in \<open>auto simp: set_integrable_def\<close>) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
962 |
then have "AE x\<in>A in M. c = f x" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
963 |
then have "AE x\<in>A in M. False" using assms(2) by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
964 |
have "A \<in> null_sets M" unfolding ae_filter_def by (meson AE_iff_null_sets \<open>A \<in> sets M\<close> \<open>AE x\<in>A in M. False\<close>) |
67226 | 965 |
then show False using \<open>emeasure (restr_to_subalg M F) A > 0\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
966 |
by (simp add: emeasure_restr_to_subalg null_setsD1 subalg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
967 |
qed |
67226 | 968 |
then show ?thesis using AE_iff_null_sets[OF \<open>X \<in> sets M\<close>] unfolding X_def by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
969 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
970 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
971 |
lemma real_cond_exp_less_c: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
972 |
assumes [measurable]: "integrable M f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
973 |
and "AE x in M. f x < c" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
974 |
shows "AE x in M. real_cond_exp M F f x < c" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
975 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
976 |
have "AE x in M. real_cond_exp M F f x = -real_cond_exp M F (\<lambda>x. -f x) x" |
67226 | 977 |
using real_cond_exp_cmult[OF \<open>integrable M f\<close>, of "-1"] by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
978 |
moreover have "AE x in M. real_cond_exp M F (\<lambda>x. -f x) x > -c" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
979 |
apply (rule real_cond_exp_gr_c) using assms by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
980 |
ultimately show ?thesis by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
981 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
982 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
983 |
lemma real_cond_exp_ge_c: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
984 |
assumes [measurable]: "integrable M f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
985 |
and "AE x in M. f x \<ge> c" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
986 |
shows "AE x in M. real_cond_exp M F f x \<ge> c" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
987 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
988 |
obtain u::"nat \<Rightarrow> real" where u: "\<And>n. u n < c" "u \<longlonglongrightarrow> c" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
989 |
using approx_from_below_dense_linorder[of "c-1" c] by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
990 |
have *: "AE x in M. real_cond_exp M F f x > u n" for n::nat |
67226 | 991 |
apply (rule real_cond_exp_gr_c) using assms \<open>u n < c\<close> by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
992 |
have "AE x in M. \<forall>n. real_cond_exp M F f x > u n" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
993 |
by (subst AE_all_countable, auto simp add: *) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
994 |
moreover have "real_cond_exp M F f x \<ge> c" if "\<forall>n. real_cond_exp M F f x > u n" for x |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
995 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
996 |
have "real_cond_exp M F f x \<ge> u n" for n using that less_imp_le by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
997 |
then show ?thesis using u(2) LIMSEQ_le_const2 by blast |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
998 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
999 |
ultimately show ?thesis by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1000 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1001 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1002 |
lemma real_cond_exp_le_c: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1003 |
assumes [measurable]: "integrable M f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1004 |
and "AE x in M. f x \<le> c" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1005 |
shows "AE x in M. real_cond_exp M F f x \<le> c" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1006 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1007 |
have "AE x in M. real_cond_exp M F f x = -real_cond_exp M F (\<lambda>x. -f x) x" |
67226 | 1008 |
using real_cond_exp_cmult[OF \<open>integrable M f\<close>, of "-1"] by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1009 |
moreover have "AE x in M. real_cond_exp M F (\<lambda>x. -f x) x \<ge> -c" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1010 |
apply (rule real_cond_exp_ge_c) using assms by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1011 |
ultimately show ?thesis by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1012 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1013 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1014 |
lemma real_cond_exp_mono_strict: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1015 |
assumes "AE x in M. f x < g x" and [measurable]: "integrable M f" "integrable M g" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1016 |
shows "AE x in M. real_cond_exp M F f x < real_cond_exp M F g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1017 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1018 |
have "AE x in M. real_cond_exp M F g x - real_cond_exp M F f x = real_cond_exp M F (\<lambda>x. g x - f x) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1019 |
by (rule AE_symmetric[OF real_cond_exp_diff], auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1020 |
moreover have "AE x in M. real_cond_exp M F (\<lambda>x. g x - f x) x > 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1021 |
by (rule real_cond_exp_gr_c, auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1022 |
ultimately have "AE x in M. real_cond_exp M F g x - real_cond_exp M F f x > 0" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1023 |
then show ?thesis by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1024 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1025 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1026 |
lemma real_cond_exp_nested_subalg [intro, simp]: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1027 |
assumes "subalgebra M G" "subalgebra G F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1028 |
and [measurable]: "integrable M f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1029 |
shows "AE x in M. real_cond_exp M F (real_cond_exp M G f) x = real_cond_exp M F f x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1030 |
proof (rule real_cond_exp_charact) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1031 |
interpret G: sigma_finite_subalgebra M G by (rule nested_subalg_is_sigma_finite[OF assms(1) assms(2)]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1032 |
show "integrable M (real_cond_exp M G f)" by (auto simp add: assms G.real_cond_exp_int(1)) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1033 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1034 |
fix A assume [measurable]: "A \<in> sets F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1035 |
then have [measurable]: "A \<in> sets G" using assms(2) by (meson set_mp subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1036 |
have "set_lebesgue_integral M A (real_cond_exp M G f) = set_lebesgue_integral M A f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1037 |
by (rule G.real_cond_exp_intA[symmetric], auto simp add: assms(3)) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1038 |
also have "... = set_lebesgue_integral M A (real_cond_exp M F f)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1039 |
by (rule real_cond_exp_intA, auto simp add: assms(3)) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1040 |
finally show "set_lebesgue_integral M A (real_cond_exp M G f) = set_lebesgue_integral M A (real_cond_exp M F f)" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1041 |
qed (auto simp add: assms real_cond_exp_int(1)) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1042 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1043 |
lemma real_cond_exp_sum [intro, simp]: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1044 |
fixes f::"'b \<Rightarrow> 'a \<Rightarrow> real" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1045 |
assumes [measurable]: "\<And>i. integrable M (f i)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1046 |
shows "AE x in M. real_cond_exp M F (\<lambda>x. \<Sum>i\<in>I. f i x) x = (\<Sum>i\<in>I. real_cond_exp M F (f i) x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1047 |
proof (rule real_cond_exp_charact) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1048 |
fix A assume [measurable]: "A \<in> sets F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1049 |
then have A_meas [measurable]: "A \<in> sets M" by (meson set_mp subalg subalgebra_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1050 |
have *: "integrable M (\<lambda>x. indicator A x * f i x)" for i |
67226 | 1051 |
using integrable_mult_indicator[OF \<open>A \<in> sets M\<close> assms(1)] by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1052 |
have **: "integrable M (\<lambda>x. indicator A x * real_cond_exp M F (f i) x)" for i |
67226 | 1053 |
using integrable_mult_indicator[OF \<open>A \<in> sets M\<close> real_cond_exp_int(1)[OF assms(1)]] by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1054 |
have inti: "(\<integral>x. indicator A x * f i x \<partial>M) = (\<integral>x. indicator A x * real_cond_exp M F (f i) x \<partial>M)" for i |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1055 |
by (rule real_cond_exp_intg(2)[symmetric], auto simp add: *) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1056 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1057 |
have "(\<integral>x\<in>A. (\<Sum>i\<in>I. f i x)\<partial>M) = (\<integral>x. (\<Sum>i\<in>I. indicator A x * f i x)\<partial>M)" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
1058 |
by (simp add: sum_distrib_left set_lebesgue_integral_def) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1059 |
also have "... = (\<Sum>i\<in>I. (\<integral>x. indicator A x * f i x \<partial>M))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1060 |
by (rule Bochner_Integration.integral_sum, simp add: *) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1061 |
also have "... = (\<Sum>i\<in>I. (\<integral>x. indicator A x * real_cond_exp M F (f i) x \<partial>M))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1062 |
using inti by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1063 |
also have "... = (\<integral>x. (\<Sum>i\<in>I. indicator A x * real_cond_exp M F (f i) x)\<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1064 |
by (rule Bochner_Integration.integral_sum[symmetric], simp add: **) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1065 |
also have "... = (\<integral>x\<in>A. (\<Sum>i\<in>I. real_cond_exp M F (f i) x)\<partial>M)" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67226
diff
changeset
|
1066 |
by (simp add: sum_distrib_left set_lebesgue_integral_def) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1067 |
finally show "(\<integral>x\<in>A. (\<Sum>i\<in>I. f i x)\<partial>M) = (\<integral>x\<in>A. (\<Sum>i\<in>I. real_cond_exp M F (f i) x)\<partial>M)" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1068 |
qed (auto simp add: assms real_cond_exp_int(1)[OF assms(1)]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1069 |
|
67226 | 1070 |
text \<open>Jensen's inequality, describing the behavior of the integral under a convex function, admits |
1071 |
a version for the conditional expectation, as follows.\<close> |
|
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1072 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1073 |
theorem real_cond_exp_jensens_inequality: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1074 |
fixes q :: "real \<Rightarrow> real" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1075 |
assumes X: "integrable M X" "AE x in M. X x \<in> I" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1076 |
assumes I: "I = {a <..< b} \<or> I = {a <..} \<or> I = {..< b} \<or> I = UNIV" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1077 |
assumes q: "integrable M (\<lambda>x. q (X x))" "convex_on I q" "q \<in> borel_measurable borel" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1078 |
shows "AE x in M. real_cond_exp M F X x \<in> I" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1079 |
"AE x in M. q (real_cond_exp M F X x) \<le> real_cond_exp M F (\<lambda>x. q (X x)) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1080 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1081 |
have "open I" using I by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1082 |
then have "interior I = I" by (simp add: interior_eq) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1083 |
have [measurable]: "I \<in> sets borel" using I by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1084 |
define phi where "phi = (\<lambda>x. Inf ((\<lambda>t. (q x - q t) / (x - t)) ` ({x<..} \<inter> I)))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1085 |
have **: "q (X x) \<ge> q (real_cond_exp M F X x) + phi (real_cond_exp M F X x) * (X x - real_cond_exp M F X x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1086 |
if "X x \<in> I" "real_cond_exp M F X x \<in> I" for x |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1087 |
unfolding phi_def apply (rule convex_le_Inf_differential) |
67226 | 1088 |
using \<open>convex_on I q\<close> that \<open>interior I = I\<close> by auto |
1089 |
text \<open>It is not clear that the function $\phi$ is measurable. We replace it by a version which |
|
1090 |
is better behaved.\<close> |
|
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1091 |
define psi where "psi = (\<lambda>x. phi x * indicator I x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1092 |
have A: "psi y = phi y" if "y \<in> I" for y unfolding psi_def indicator_def using that by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1093 |
have *: "q (X x) \<ge> q (real_cond_exp M F X x) + psi (real_cond_exp M F X x) * (X x - real_cond_exp M F X x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1094 |
if "X x \<in> I" "real_cond_exp M F X x \<in> I" for x |
67226 | 1095 |
unfolding A[OF \<open>real_cond_exp M F X x \<in> I\<close>] using ** that by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1096 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1097 |
note I |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1098 |
moreover have "AE x in M. real_cond_exp M F X x > a" if "I \<subseteq> {a <..}" for a |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1099 |
apply (rule real_cond_exp_gr_c) using X that by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1100 |
moreover have "AE x in M. real_cond_exp M F X x < b" if "I \<subseteq> {..<b}" for b |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1101 |
apply (rule real_cond_exp_less_c) using X that by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1102 |
ultimately show "AE x in M. real_cond_exp M F X x \<in> I" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1103 |
by (elim disjE) (auto simp: subset_eq) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1104 |
then have main_ineq: "AE x in M. q (X x) \<ge> q (real_cond_exp M F X x) + psi (real_cond_exp M F X x) * (X x - real_cond_exp M F X x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1105 |
using * X(2) by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1106 |
|
67226 | 1107 |
text \<open>Then, one wants to take the conditional expectation of this inequality. On the left, one gets |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1108 |
the conditional expectation of $q \circ X$. On the right, the last term vanishes, and one |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1109 |
is left with $q$ of the conditional expectation, as desired. Unfortunately, this argument only |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1110 |
works if $\psi \cdot X$ and $q(E(X | F))$ are integrable, and there is no reason why this should be true. The |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1111 |
trick is to multiply by a $F$-measurable function which is small enough to make |
67226 | 1112 |
everything integrable.\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1113 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1114 |
obtain f::"'a \<Rightarrow> real" where [measurable]: "f \<in> borel_measurable (restr_to_subalg M F)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1115 |
"integrable (restr_to_subalg M F) f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1116 |
and f: "\<And>x. f x > 0" "\<And>x. f x \<le> 1" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1117 |
using sigma_finite_measure.obtain_positive_integrable_function[OF sigma_fin_subalg] by metis |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1118 |
then have [measurable]: "f \<in> borel_measurable F" by (simp add: subalg) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1119 |
then have [measurable]: "f \<in> borel_measurable M" using measurable_from_subalg[OF subalg] by blast |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1120 |
define g where "g = (\<lambda>x. f x/(1+ \<bar>psi (real_cond_exp M F X x)\<bar> + \<bar>q (real_cond_exp M F X x)\<bar>))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1121 |
define G where "G = (\<lambda>x. g x * psi (real_cond_exp M F X x))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1122 |
have g: "g x > 0" "g x \<le> 1" for x unfolding G_def g_def using f[of x] by (auto simp add: abs_mult) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1123 |
have G: "\<bar>G x\<bar> \<le> 1" for x unfolding G_def g_def using f[of x] |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1124 |
proof (auto simp add: abs_mult) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1125 |
have "f x * \<bar>psi (real_cond_exp M F X x)\<bar> \<le> 1 * \<bar>psi (real_cond_exp M F X x)\<bar>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1126 |
apply (rule mult_mono) using f[of x] by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1127 |
also have "... \<le> 1 + \<bar>psi (real_cond_exp M F X x)\<bar> + \<bar>q (real_cond_exp M F X x)\<bar>" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1128 |
finally show "f x * \<bar>psi (real_cond_exp M F X x)\<bar> \<le> 1 + \<bar>psi (real_cond_exp M F X x)\<bar> + \<bar>q (real_cond_exp M F X x)\<bar>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1129 |
by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1130 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1131 |
have "AE x in M. g x * q (X x) \<ge> g x * (q (real_cond_exp M F X x) + psi (real_cond_exp M F X x) * (X x - real_cond_exp M F X x))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1132 |
using main_ineq g by (auto simp add: divide_simps) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1133 |
then have main_G: "AE x in M. g x * q (X x) \<ge> g x * q (real_cond_exp M F X x) + G x * (X x - real_cond_exp M F X x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1134 |
unfolding G_def by (auto simp add: algebra_simps) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1135 |
|
67226 | 1136 |
text \<open>To proceed, we need to know that $\psi$ is measurable.\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1137 |
have phi_mono: "phi x \<le> phi y" if "x \<le> y" "x \<in> I" "y \<in> I" for x y |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1138 |
proof (cases "x < y") |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1139 |
case True |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1140 |
have "q x + phi x * (y-x) \<le> q y" |
67226 | 1141 |
unfolding phi_def apply (rule convex_le_Inf_differential) using \<open>convex_on I q\<close> that \<open>interior I = I\<close> by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1142 |
then have "phi x \<le> (q x - q y) / (x - y)" |
67226 | 1143 |
using that \<open>x < y\<close> by (auto simp add: divide_simps algebra_simps) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1144 |
moreover have "(q x - q y)/(x - y) \<le> phi y" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1145 |
unfolding phi_def proof (rule cInf_greatest, auto) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1146 |
fix t assume "t \<in> I" "y < t" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1147 |
have "(q x - q y) / (x - y) \<le> (q x - q t) / (x - t)" |
67226 | 1148 |
apply (rule convex_on_diff[OF q(2)]) using \<open>y < t\<close> \<open>x < y\<close> \<open>t \<in> I\<close> \<open>x \<in> I\<close> by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1149 |
also have "... \<le> (q y - q t) / (y - t)" |
67226 | 1150 |
apply (rule convex_on_diff[OF q(2)]) using \<open>y < t\<close> \<open>x < y\<close> \<open>t \<in> I\<close> \<open>x \<in> I\<close> by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1151 |
finally show "(q x - q y) / (x - y) \<le> (q y - q t) / (y - t)" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1152 |
next |
67226 | 1153 |
obtain e where "0 < e" "ball y e \<subseteq> I" using \<open>open I\<close> \<open>y \<in> I\<close> openE by blast |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1154 |
then have "y + e/2 \<in> {y<..} \<inter> I" by (auto simp: dist_real_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1155 |
then show "{y<..} \<inter> I = {} \<Longrightarrow> False" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1156 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1157 |
ultimately show "phi x \<le> phi y" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1158 |
next |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1159 |
case False |
67226 | 1160 |
then have "x = y" using \<open>x \<le> y\<close> by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1161 |
then show ?thesis by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1162 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1163 |
have [measurable]: "psi \<in> borel_measurable borel" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1164 |
by (rule borel_measurable_piecewise_mono[of "{I, -I}"]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1165 |
(auto simp add: psi_def indicator_def phi_mono intro: mono_onI) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1166 |
have [measurable]: "q \<in> borel_measurable borel" using q by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1167 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1168 |
have [measurable]: "X \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1169 |
"real_cond_exp M F X \<in> borel_measurable F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1170 |
"g \<in> borel_measurable F" "g \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1171 |
"G \<in> borel_measurable F" "G \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1172 |
using X measurable_from_subalg[OF subalg] unfolding G_def g_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1173 |
have int1: "integrable (restr_to_subalg M F) (\<lambda>x. g x * q (real_cond_exp M F X x))" |
67226 | 1174 |
apply (rule Bochner_Integration.integrable_bound[of _ f], auto simp add: subalg \<open>integrable (restr_to_subalg M F) f\<close>) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1175 |
unfolding g_def by (auto simp add: divide_simps abs_mult algebra_simps) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1176 |
have int2: "integrable M (\<lambda>x. G x * (X x - real_cond_exp M F X x))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1177 |
apply (rule Bochner_Integration.integrable_bound[of _ "\<lambda>x. \<bar>X x\<bar> + \<bar>real_cond_exp M F X x\<bar>"]) |
67226 | 1178 |
apply (auto intro!: Bochner_Integration.integrable_add integrable_abs real_cond_exp_int \<open>integrable M X\<close> AE_I2) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1179 |
using G unfolding abs_mult by (meson abs_ge_zero abs_triangle_ineq4 dual_order.trans mult_left_le_one_le) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1180 |
have int3: "integrable M (\<lambda>x. g x * q (X x))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1181 |
apply (rule Bochner_Integration.integrable_bound[of _ "\<lambda>x. q(X x)"], auto simp add: q(1) abs_mult) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1182 |
using g by (simp add: less_imp_le mult_left_le_one_le) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1183 |
|
67226 | 1184 |
text \<open>Taking the conditional expectation of the main convexity inequality \verb+main_G+, we get |
1185 |
the following.\<close> |
|
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1186 |
have "AE x in M. real_cond_exp M F (\<lambda>x. g x * q (X x)) x \<ge> real_cond_exp M F (\<lambda>x. g x * q (real_cond_exp M F X x) + G x * (X x - real_cond_exp M F X x)) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1187 |
apply (rule real_cond_exp_mono[OF main_G]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1188 |
apply (rule Bochner_Integration.integrable_add[OF integrable_from_subalg[OF subalg int1]]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1189 |
using int2 int3 by auto |
67226 | 1190 |
text \<open>This reduces to the desired inequality thanks to the properties of conditional expectation, |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1191 |
i.e., the conditional expectation of an $F$-measurable function is this function, and one can |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1192 |
multiply an $F$-measurable function outside of conditional expectations. |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1193 |
Since all these equalities only hold almost everywhere, we formulate them separately, |
67226 | 1194 |
and then combine all of them to simplify the above equation, again almost everywhere.\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1195 |
moreover have "AE x in M. real_cond_exp M F (\<lambda>x. g x * q (X x)) x = g x * real_cond_exp M F (\<lambda>x. q (X x)) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1196 |
by (rule real_cond_exp_mult, auto simp add: int3) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1197 |
moreover have "AE x in M. real_cond_exp M F (\<lambda>x. g x * q (real_cond_exp M F X x) + G x * (X x - real_cond_exp M F X x)) x |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1198 |
= real_cond_exp M F (\<lambda>x. g x * q (real_cond_exp M F X x)) x + real_cond_exp M F (\<lambda>x. G x * (X x - real_cond_exp M F X x)) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1199 |
by (rule real_cond_exp_add, auto simp add: integrable_from_subalg[OF subalg int1] int2) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1200 |
moreover have "AE x in M. real_cond_exp M F (\<lambda>x. g x * q (real_cond_exp M F X x)) x = g x * q (real_cond_exp M F X x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1201 |
by (rule real_cond_exp_F_meas, auto simp add: integrable_from_subalg[OF subalg int1]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1202 |
moreover have "AE x in M. real_cond_exp M F (\<lambda>x. G x * (X x - real_cond_exp M F X x)) x = G x * real_cond_exp M F (\<lambda>x. (X x - real_cond_exp M F X x)) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1203 |
by (rule real_cond_exp_mult, auto simp add: int2) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1204 |
moreover have "AE x in M. real_cond_exp M F (\<lambda>x. (X x - real_cond_exp M F X x)) x = real_cond_exp M F X x - real_cond_exp M F (\<lambda>x. real_cond_exp M F X x) x" |
67226 | 1205 |
by (rule real_cond_exp_diff, auto intro!: real_cond_exp_int \<open>integrable M X\<close>) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1206 |
moreover have "AE x in M. real_cond_exp M F (\<lambda>x. real_cond_exp M F X x) x = real_cond_exp M F X x " |
67226 | 1207 |
by (rule real_cond_exp_F_meas, auto intro!: real_cond_exp_int \<open>integrable M X\<close>) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1208 |
ultimately have "AE x in M. g x * real_cond_exp M F (\<lambda>x. q (X x)) x \<ge> g x * q (real_cond_exp M F X x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1209 |
by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1210 |
then show "AE x in M. real_cond_exp M F (\<lambda>x. q (X x)) x \<ge> q (real_cond_exp M F X x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1211 |
using g(1) by (auto simp add: divide_simps) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1212 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1213 |
|
67226 | 1214 |
text \<open>Jensen's inequality does not imply that $q(E(X|F))$ is integrable, as it only proves an upper |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1215 |
bound for it. Indeed, this is not true in general, as the following counterexample shows: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1216 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1217 |
on $[1,\infty)$ with Lebesgue measure, let $F$ be the sigma-algebra generated by the intervals $[n, n+1)$ |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1218 |
for integer $n$. Let $q(x) = - \sqrt{x}$ for $x\geq 0$. Define $X$ which is equal to $1/n$ over |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1219 |
$[n, n+1/n)$ and $2^{-n}$ on $[n+1/n, n+1)$. Then $X$ is integrable as $\sum 1/n^2 < \infty$, and |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1220 |
$q(X)$ is integrable as $\sum 1/n^{3/2} < \infty$. On the other hand, $E(X|F)$ is essentially equal |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1221 |
to $1/n^2$ on $[n, n+1)$ (we neglect the term $2^{-n}$, we only put it there because $X$ should take |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1222 |
its values in $I=(0,\infty)$). Hence, $q(E(X|F))$ is equal to $-1/n$ on $[n, n+1)$, hence it is not |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1223 |
integrable. |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1224 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1225 |
However, this counterexample is essentially the only situation where this function is not |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1226 |
integrable, as shown by the next lemma. |
67226 | 1227 |
\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1228 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1229 |
lemma integrable_convex_cond_exp: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1230 |
fixes q :: "real \<Rightarrow> real" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1231 |
assumes X: "integrable M X" "AE x in M. X x \<in> I" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1232 |
assumes I: "I = {a <..< b} \<or> I = {a <..} \<or> I = {..< b} \<or> I = UNIV" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1233 |
assumes q: "integrable M (\<lambda>x. q (X x))" "convex_on I q" "q \<in> borel_measurable borel" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1234 |
assumes H: "emeasure M (space M) = \<infinity> \<Longrightarrow> 0 \<in> I" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1235 |
shows "integrable M (\<lambda>x. q (real_cond_exp M F X x))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1236 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1237 |
have [measurable]: "(\<lambda>x. q (real_cond_exp M F X x)) \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1238 |
"q \<in> borel_measurable borel" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1239 |
"X \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1240 |
using X(1) q(3) by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1241 |
have "open I" using I by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1242 |
then have "interior I = I" by (simp add: interior_eq) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1243 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1244 |
consider "emeasure M (space M) = 0" | "emeasure M (space M) > 0 \<and> emeasure M (space M) < \<infinity>" | "emeasure M (space M) = \<infinity>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1245 |
by (metis infinity_ennreal_def not_gr_zero top.not_eq_extremum) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1246 |
then show ?thesis |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1247 |
proof (cases) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1248 |
case 1 |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1249 |
show ?thesis by (subst integrable_cong_AE[of _ _ "\<lambda>x. 0"], auto intro: emeasure_0_AE[OF 1]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1250 |
next |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1251 |
case 2 |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1252 |
interpret finite_measure M using 2 by (auto intro!: finite_measureI) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1253 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1254 |
have "I \<noteq> {}" |
67226 | 1255 |
using \<open>AE x in M. X x \<in> I\<close> 2 eventually_mono integral_less_AE_space by fastforce |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1256 |
then obtain z where "z \<in> I" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1257 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1258 |
define A where "A = Inf ((\<lambda>t. (q z - q t) / (z - t)) ` ({z<..} \<inter> I))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1259 |
have "q y \<ge> q z + A * (y - z)" if "y \<in> I" for y unfolding A_def apply (rule convex_le_Inf_differential) |
67226 | 1260 |
using \<open>z \<in> I\<close> \<open>y \<in> I\<close> \<open>interior I = I\<close> q(2) by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1261 |
then have "AE x in M. q (real_cond_exp M F X x) \<ge> q z + A * (real_cond_exp M F X x - z)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1262 |
using real_cond_exp_jensens_inequality(1)[OF X I q] by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1263 |
moreover have "AE x in M. q (real_cond_exp M F X x) \<le> real_cond_exp M F (\<lambda>x. q (X x)) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1264 |
using real_cond_exp_jensens_inequality(2)[OF X I q] by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1265 |
moreover have "\<bar>a\<bar> \<le> \<bar>b\<bar> + \<bar>c\<bar>" if "b \<le> a \<and> a \<le> c" for a b c::real |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1266 |
using that by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1267 |
ultimately have *: "AE x in M. \<bar>q (real_cond_exp M F X x)\<bar> |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1268 |
\<le> \<bar>real_cond_exp M F (\<lambda>x. q (X x)) x\<bar> + \<bar>q z + A * (real_cond_exp M F X x - z)\<bar>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1269 |
by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1270 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1271 |
show "integrable M (\<lambda>x. q (real_cond_exp M F X x))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1272 |
apply (rule Bochner_Integration.integrable_bound[of _ "\<lambda>x. \<bar>real_cond_exp M F (\<lambda>x. q (X x)) x\<bar> + \<bar>q z + A * (real_cond_exp M F X x - z)\<bar>"]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1273 |
apply (auto intro!: Bochner_Integration.integrable_add integrable_abs integrable_mult_right Bochner_Integration.integrable_diff real_cond_exp_int(1)) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1274 |
using X(1) q(1) * by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1275 |
next |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1276 |
case 3 |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1277 |
then have "0 \<in> I" using H finite_measure.finite_emeasure_space by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1278 |
have "q(0) = 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1279 |
proof (rule ccontr) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1280 |
assume *: "\<not>(q(0) = 0)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1281 |
define e where "e = \<bar>q(0)\<bar> / 2" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1282 |
then have "e > 0" using * by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1283 |
have "continuous (at 0) q" |
67226 | 1284 |
using q(2) \<open>0 \<in> I\<close> \<open>open I\<close> \<open>interior I = I\<close> continuous_on_interior convex_on_continuous by blast |
1285 |
then obtain d where d: "d > 0" "\<And>y. \<bar>y - 0\<bar> < d \<Longrightarrow> \<bar>q y - q 0\<bar> < e" using \<open>e > 0\<close> |
|
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1286 |
by (metis continuous_at_real_range real_norm_def) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1287 |
then have *: "\<bar>q(y)\<bar> > e" if "\<bar>y\<bar> < d" for y |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1288 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1289 |
have "\<bar>q 0\<bar> \<le> \<bar>q 0 - q y\<bar> + \<bar>q y\<bar>" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1290 |
also have "... < e + \<bar>q y\<bar>" using d(2) that by force |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1291 |
finally have "\<bar>q y\<bar> > \<bar>q 0\<bar> - e" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1292 |
then show ?thesis unfolding e_def by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1293 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1294 |
have "emeasure M {x \<in> space M. \<bar>X x\<bar> < d} \<le> emeasure M ({x \<in> space M. 1 \<le> ennreal(1/e) * \<bar>q(X x)\<bar>} \<inter> space M)" |
67226 | 1295 |
by (rule emeasure_mono, auto simp add: * \<open>e>0\<close> less_imp_le ennreal_mult''[symmetric]) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1296 |
also have "... \<le> (1/e) * (\<integral>\<^sup>+x. ennreal(\<bar>q(X x)\<bar>) * indicator (space M) x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1297 |
by (rule nn_integral_Markov_inequality, auto) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1298 |
also have "... = (1/e) * (\<integral>\<^sup>+x. ennreal(\<bar>q(X x)\<bar>) \<partial>M)" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1299 |
also have "... = (1/e) * ennreal(\<integral>x. \<bar>q(X x)\<bar> \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1300 |
using nn_integral_eq_integral[OF integrable_abs[OF q(1)]] by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1301 |
also have "... < \<infinity>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1302 |
by (simp add: ennreal_mult_less_top) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1303 |
finally have A: "emeasure M {x \<in> space M. \<bar>X x\<bar> < d} < \<infinity>" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1304 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1305 |
have "{x \<in> space M. \<bar>X x\<bar> \<ge> d} = {x \<in> space M. 1 \<le> ennreal(1/d) * \<bar>X x\<bar>} \<inter> space M" |
67226 | 1306 |
by (auto simp add: \<open>d>0\<close> ennreal_mult''[symmetric]) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1307 |
then have "emeasure M {x \<in> space M. \<bar>X x\<bar> \<ge> d} = emeasure M ({x \<in> space M. 1 \<le> ennreal(1/d) * \<bar>X x\<bar>} \<inter> space M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1308 |
by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1309 |
also have "... \<le> (1/d) * (\<integral>\<^sup>+x. ennreal(\<bar>X x\<bar>) * indicator (space M) x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1310 |
by (rule nn_integral_Markov_inequality, auto) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1311 |
also have "... = (1/d) * (\<integral>\<^sup>+x. ennreal(\<bar>X x\<bar>) \<partial>M)" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1312 |
also have "... = (1/d) * ennreal(\<integral>x. \<bar>X x\<bar> \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1313 |
using nn_integral_eq_integral[OF integrable_abs[OF X(1)]] by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1314 |
also have "... < \<infinity>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1315 |
by (simp add: ennreal_mult_less_top) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1316 |
finally have B: "emeasure M {x \<in> space M. \<bar>X x\<bar> \<ge> d} < \<infinity>" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1317 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1318 |
have "space M = {x \<in> space M. \<bar>X x\<bar> < d} \<union> {x \<in> space M. \<bar>X x\<bar> \<ge> d}" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1319 |
then have "emeasure M (space M) = emeasure M ({x \<in> space M. \<bar>X x\<bar> < d} \<union> {x \<in> space M. \<bar>X x\<bar> \<ge> d})" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1320 |
by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1321 |
also have "... \<le> emeasure M {x \<in> space M. \<bar>X x\<bar> < d} + emeasure M {x \<in> space M. \<bar>X x\<bar> \<ge> d}" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1322 |
by (auto intro!: emeasure_subadditive) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1323 |
also have "... < \<infinity>" using A B by auto |
67226 | 1324 |
finally show False using \<open>emeasure M (space M) = \<infinity>\<close> by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1325 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1326 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1327 |
define A where "A = Inf ((\<lambda>t. (q 0 - q t) / (0 - t)) ` ({0<..} \<inter> I))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1328 |
have "q y \<ge> q 0 + A * (y - 0)" if "y \<in> I" for y unfolding A_def apply (rule convex_le_Inf_differential) |
67226 | 1329 |
using \<open>0 \<in> I\<close> \<open>y \<in> I\<close> \<open>interior I = I\<close> q(2) by auto |
1330 |
then have "q y \<ge> A * y" if "y \<in> I" for y using \<open>q 0 = 0\<close> that by auto |
|
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1331 |
then have "AE x in M. q (real_cond_exp M F X x) \<ge> A * real_cond_exp M F X x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1332 |
using real_cond_exp_jensens_inequality(1)[OF X I q] by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1333 |
moreover have "AE x in M. q (real_cond_exp M F X x) \<le> real_cond_exp M F (\<lambda>x. q (X x)) x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1334 |
using real_cond_exp_jensens_inequality(2)[OF X I q] by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1335 |
moreover have "\<bar>a\<bar> \<le> \<bar>b\<bar> + \<bar>c\<bar>" if "b \<le> a \<and> a \<le> c" for a b c::real |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1336 |
using that by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1337 |
ultimately have *: "AE x in M. \<bar>q (real_cond_exp M F X x)\<bar> |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1338 |
\<le> \<bar>real_cond_exp M F (\<lambda>x. q (X x)) x\<bar> + \<bar>A * real_cond_exp M F X x\<bar>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1339 |
by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1340 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1341 |
show "integrable M (\<lambda>x. q (real_cond_exp M F X x))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1342 |
apply (rule Bochner_Integration.integrable_bound[of _ "\<lambda>x. \<bar>real_cond_exp M F (\<lambda>x. q (X x)) x\<bar> + \<bar>A * real_cond_exp M F X x \<bar>"]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1343 |
apply (auto intro!: Bochner_Integration.integrable_add integrable_abs integrable_mult_right Bochner_Integration.integrable_diff real_cond_exp_int(1)) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1344 |
using X(1) q(1) * by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1345 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1346 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1347 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1348 |
end |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1349 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
diff
changeset
|
1350 |
end |